diff --git a/.gitignore b/.gitignore index 08e40751d..67dab47bd 100644 --- a/.gitignore +++ b/.gitignore @@ -412,3 +412,9 @@ erasure/src/signature.ml erasure/src/signature.mli erasure/src/eOptimizePropDiscr.ml erasure/src/eOptimizePropDiscr.mli +erasure/src/pCUICErrors.ml +erasure/src/pCUICErrors.mli +erasure/src/pCUICTypeChecker.ml +erasure/src/pCUICTypeChecker.mli +erasure/src/pCUICPrimitive.ml +erasure/src/pCUICPrimitive.mli diff --git a/INSTALL.md b/INSTALL.md index cc7c8a978..f007ab5e5 100644 --- a/INSTALL.md +++ b/INSTALL.md @@ -25,7 +25,6 @@ MetaCoq is split into multiple packages that get all installed using the `coq-metacoq` meta-package: - `coq-metacoq-template` for the Template Monad and quoting plugin - - `coq-metacoq-checker` for the UNverified checker of template-coq terms - `coq-metacoq-pcuic` for the PCUIC development and proof of the Template-Coq -> PCUIC translation - `coq-metacoq-safechecker` for the verified checker on PCUIC terms @@ -103,7 +102,7 @@ the sources directory. Then use: -- `make` to compile the `template-coq` plugin, the `checker`, the `pcuic` +- `make` to compile the `template-coq` plugin, the `pcuic` development and the `safechecker` and `erasure` plugins. You can also selectively build each target. diff --git a/Makefile b/Makefile index ad2c5d82a..61343cfbb 100644 --- a/Makefile +++ b/Makefile @@ -1,12 +1,11 @@ TIMED ?= -all: template-coq checker pcuic safechecker erasure examples +all: template-coq pcuic safechecker erasure examples -.PHONY: all template-coq checker pcuic erasure install html clean mrproper .merlin test-suite translations +.PHONY: all template-coq pcuic erasure install html clean mrproper .merlin test-suite translations install: all $(MAKE) -C template-coq install - $(MAKE) -C checker install $(MAKE) -C pcuic install $(MAKE) -C safechecker install $(MAKE) -C erasure install @@ -14,7 +13,6 @@ install: all uninstall: all $(MAKE) -C template-coq uninstall - $(MAKE) -C checker uninstall $(MAKE) -C pcuic uninstall $(MAKE) -C safechecker uninstall $(MAKE) -C erasure uninstall @@ -23,7 +21,6 @@ uninstall: all html: all "coqdoc" -toc -utf8 -interpolate -l -html \ -R template-coq/theories MetaCoq.Template \ - -R checker/theories MetaCoq.Checker \ -R pcuic/theories MetaCoq.PCUIC \ -R safechecker/theories MetaCoq.SafeChecker \ -R erasure/theories MetaCoq.Erasure \ @@ -32,7 +29,6 @@ html: all clean: $(MAKE) -C template-coq clean - $(MAKE) -C checker clean $(MAKE) -C pcuic clean $(MAKE) -C safechecker clean $(MAKE) -C erasure clean @@ -45,7 +41,6 @@ mrproper: $(MAKE) -C pcuic mrproper $(MAKE) -C safechecker mrproper $(MAKE) -C erasure mrproper - $(MAKE) -C checker mrproper $(MAKE) -C examples mrproper $(MAKE) -C test-suite mrproper $(MAKE) -C translations mrproper @@ -55,7 +50,6 @@ mrproper: $(MAKE) -C pcuic .merlin $(MAKE) -C safechecker .merlin $(MAKE) -C erasure .merlin - $(MAKE) -C checker .merlin template-coq: $(MAKE) -C template-coq @@ -69,28 +63,23 @@ safechecker: template-coq pcuic erasure: template-coq safechecker pcuic $(MAKE) -C erasure -checker: template-coq - $(MAKE) -C checker - -examples: checker safechecker +examples: checker safechecker erasure $(MAKE) -C examples -test-suite: template-coq checker safechecker erasure +test-suite: template-coq safechecker erasure $(MAKE) -C test-suite -translations: template-coq checker +translations: template-coq $(MAKE) -C translations cleanplugins: $(MAKE) -C template-coq cleanplugin - $(MAKE) -C pcuic cleanplugin - $(MAKE) -C checker cleanplugin $(MAKE) -C safechecker cleanplugin $(MAKE) -C erasure cleanplugin ci-local: ./configure.sh local - $(MAKE) all test-suite TIMED=pretty-timed + $(MAKE) all test-suite ci-opam: # Use -v so that regular output is produced diff --git a/README.md b/README.md index 01d1e91c2..4f6225cb8 100644 --- a/README.md +++ b/README.md @@ -68,20 +68,13 @@ In addition to this representation of terms, Template Coq includes: checker, and inserting them in the global environment, in the style of MTac. +- A formalisation of the expected typing rules reflecting the ones of Coq -### [Checker](https://github.com/MetaCoq/metacoq/tree/coq-8.12/checker) - -A partial type-checker for the Calculus of Inductive Constructions, -whose extraction to ML is runnable as a plugin (using command `MetaCoq -Check foo`). This checker uses _fuel_, so it must be passed a number -of maximal reduction steps to perform when calling conversion, and is -NOT verified. - -### [PCUIC](https://github.com/MetaCoq/metacoq/tree/coq-8.12/pcuic) +### [PCUIC](https://github.com/MetaCoq/metacoq/tree/coq-8.11/pcuic) PCUIC, the Polymorphic Cumulative Calculus of Inductive Constructions is a cleaned up version of the term language of Coq and its associated -type system, equivalent to the one of Coq. This version of the +type system, shown equivalent to the one of Coq. This version of the calculus has proofs of standard metatheoretical results: - Weakening for global declarations, weakening and substitution for diff --git a/TODO.md b/TODO.md index 0c158d854..8596882d8 100644 --- a/TODO.md +++ b/TODO.md @@ -29,11 +29,16 @@ - Clean `Derive`s: always derive `Siganture`, `NoConf`, ... directly after the definition of the inductive. (To avoid doing it several times.) + (Mostly done) + +- Finish the PCUICSigmaCalculus proofs. -- Remove `Program` from everywhere. - - +# Medium Projects +- Change Template-PCUIC translations to translate casts to applications of + identity functions (vm_cast, default_cast etc) to make the back and forth + the identity and derive weakening/substitution/etc.. from the PCUIC theorems. + Is that really better than identity functions? # Big projects - Refine the longest-simple-path algorithm on universes with the @@ -49,6 +54,9 @@ be definable. Probably not very useful though because if the elimination is restricted then it means some Type is in the constructor and won't be projectable. +- Verify the substitution calculus of P.M Pédrot using skewed lists at + https://github.com/coq/coq/pull/13537 and try to use it to implement efficient explicit + substitutions. ## Website diff --git a/checker/Makefile b/checker/Makefile deleted file mode 100644 index d5ddec678..000000000 --- a/checker/Makefile +++ /dev/null @@ -1,44 +0,0 @@ -all: coq # plugin - -coq: Makefile.coq - $(MAKE) -f Makefile.coq - -_CoqProject: _CoqProject.in metacoq-config - cat metacoq-config > _CoqProject - cat _CoqProject.in >> _CoqProject - -_PluginProject: _PluginProject.in metacoq-config - cat metacoq-config > _PluginProject - cat _PluginProject.in >> _PluginProject - -Makefile.coq: _CoqProject - coq_makefile -f _CoqProject -o Makefile.coq - -Makefile.plugin: _PluginProject - coq_makefile -f _PluginProject -o Makefile.plugin - -plugin: Makefile.plugin coq - $(MAKE) -f Makefile.plugin - -install: coq # plugin - $(MAKE) -f Makefile.coq install - # $(MAKE) -f Makefile.plugin install - -uninstall: coq # plugin - $(MAKE) -f Makefile.coq uninstall - # $(MAKE) -f Makefile.plugin uninstall - -.PHONY: plugin - -clean: Makefile.coq Makefile.plugin - $(MAKE) -f Makefile.coq clean - $(MAKE) -f Makefile.plugin clean - -mrproper: - rm -f Makefile.coq Makefile.plugin _CoqProject _PluginProject - -.merlin: - make -f Makefile.plugin .merlin - -cleanplugin: Makefile.plugin - make -f Makefile.plugin clean diff --git a/checker/Makefile.coq.local b/checker/Makefile.coq.local deleted file mode 100644 index bb5386f49..000000000 --- a/checker/Makefile.coq.local +++ /dev/null @@ -1,4 +0,0 @@ -CAMLFLAGS+=-w -33 # Unused opens - -post-all:: - ./update_plugin.sh diff --git a/checker/Makefile.local b/checker/Makefile.local deleted file mode 100644 index 0ea15a69b..000000000 --- a/checker/Makefile.local +++ /dev/null @@ -1,3 +0,0 @@ -COQDOCFLAGS=--no-index --interpolate --parse-comments --utf8 $(COQLIBS_NOML) -INSTALLDEFAULTROOT=MetaCoq/Checker -CAMLFLAGS+=-w -33 # Unused opens \ No newline at end of file diff --git a/checker/Makefile.plugin.local b/checker/Makefile.plugin.local deleted file mode 100644 index c8b2e9962..000000000 --- a/checker/Makefile.plugin.local +++ /dev/null @@ -1,16 +0,0 @@ -CAMLFLAGS+=-open Metacoq_template_plugin # Loads the extractable monad -CAMLFLAGS+=-w -33 # Unused opens -CAMLFLAGS+=-w -32 # Unused value -CAMLFLAGS+=-w -39 # Unused rec flag -CAMLFLAGS+=-w -26 # Unused variable -CAMLFLAGS+=-w -34 # Unused type -CAMLFLAGS+=-w -20 # Unused argument -CAMLFLAGS+=-bin-annot # For merlin - -# Clean the source directory of all extracted filess -cleansrc: - rm -f src/*.ml src/*.mli src/*.d src/*.o src/*.cm* - -.PHONY: cleansrc - -theories/Loader.vo: src/metacoq_checker_plugin.cmxs diff --git a/checker/_CoqProject.in b/checker/_CoqProject.in deleted file mode 100644 index 4cc8ab28f..000000000 --- a/checker/_CoqProject.in +++ /dev/null @@ -1,17 +0,0 @@ --I ../template-coq/build --I src --R theories MetaCoq.Checker - -theories/Generation.v -theories/WeakeningEnv.v -theories/Closed.v -theories/Weakening.v -theories/Substitution.v -theories/Checker.v -theories/WcbvEval.v -theories/Retyping.v -theories/Normal.v -theories/All.v - -# Extraction -# theories/Extraction.v diff --git a/checker/_PluginProject.in b/checker/_PluginProject.in deleted file mode 100644 index 56a34318c..000000000 --- a/checker/_PluginProject.in +++ /dev/null @@ -1,99 +0,0 @@ --I src --Q src MetaCoq.Checker --R theories MetaCoq.Checker - -src/all_Forall.ml -src/all_Forall.mli -src/ascii.ml -src/ascii.mli -src/ast0.ml -src/ast0.mli -src/astUtils.ml -src/astUtils.mli -src/basicAst.ml -src/basicAst.mli -src/basics.ml -src/basics.mli -src/bool.ml -src/bool.mli -src/checker0.ml -src/checker0.mli -src/classes.ml -src/classes.mli -src/compare_dec.ml -src/compare_dec.mli -src/config0.ml -src/config0.mli -src/datatypes.ml -src/datatypes.mli -src/environment.ml -src/environment.mli -src/eqdepFacts.ml -src/eqdepFacts.mli -src/equalities.ml -src/equalities.mli -src/init.ml -src/init.mli -src/liftSubst.ml -src/liftSubst.mli -src/list0.ml -src/list0.mli -src/mCCompare.ml -src/mCCompare.mli -src/mCList.ml -src/mCList.mli -src/mCProd.ml -src/mCProd.mli -src/mSetDecide.ml -src/mSetDecide.mli -src/mSetFacts.ml -src/mSetFacts.mli -src/mSetInterface.ml -src/mSetInterface.mli -src/mSetList.ml -src/mSetList.mli -src/mSetProperties.ml -src/mSetProperties.mli -src/mSetWeakList.ml -src/mSetWeakList.mli -src/monad_utils.ml -src/monad_utils.mli -src/nat0.ml -src/nat0.mli -src/orders.ml -src/orders.mli -src/ordersFacts.ml -src/ordersFacts.mli -src/ordersLists.ml -src/ordersLists.mli -src/ordersTac.ml -src/ordersTac.mli -src/peanoNat.ml -src/peanoNat.mli -src/retyping0.ml -src/retyping0.mli -src/signature.ml -src/signature.mli -src/specif.ml -src/specif.mli -src/string0.ml -src/string0.mli -src/typing0.ml -src/typing0.mli -src/uGraph0.ml -src/uGraph0.mli -src/univSubst0.ml -src/univSubst0.mli -src/universes0.ml -src/universes0.mli -src/utils0.ml -src/utils0.mli -src/wGraph.ml -src/wGraph.mli -src/wf0.ml -src/wf0.mli - -src/g_metacoq_checker.mlg -src/metacoq_checker_plugin.mlpack - -theories/Loader.v diff --git a/checker/demo.v b/checker/demo.v deleted file mode 100644 index 59a75c78d..000000000 --- a/checker/demo.v +++ /dev/null @@ -1,9 +0,0 @@ -(* Distributed under the terms of the MIT license. *) -Require Import MetaCoq.Checker.Loader. - - -MetaCoq Check (3 + 9). - -Require Import Reals. - -MetaCoq Check Rplus. diff --git a/checker/src/.gitignore b/checker/src/.gitignore deleted file mode 100644 index bd41c435d..000000000 --- a/checker/src/.gitignore +++ /dev/null @@ -1,3 +0,0 @@ -*.mli -*.ml -to-lower.sh \ No newline at end of file diff --git a/checker/src/g_metacoq_checker.mlg b/checker/src/g_metacoq_checker.mlg deleted file mode 100644 index e2d1da052..000000000 --- a/checker/src/g_metacoq_checker.mlg +++ /dev/null @@ -1,36 +0,0 @@ - -DECLARE PLUGIN "metacoq_checker" - -{ - -open Stdarg -open Pp -open PeanoNat.Nat -open Checker0 -let pr_char c = str (Char.escaped c) - -let pr_char_list = prlist_with_sep mt pr_char - -let check env evm c = - (* Feedback.msg_debug (str"Quoting"); *) - let term = Ast_quoter.quote_term_rec env (EConstr.to_constr evm c) in - (* Feedback.msg_debug (str"Finished quoting.. checking."); *) - let fuel = pow two (pow two (pow two two)) in - let checker_flags = Config0.default_checker_flags in - match Checker0.typecheck_program checker_flags fuel term with - | CorrectDecl t -> - Feedback.msg_info (str "Successfully checked of type: " ++ pr_char_list (AstUtils.string_of_term t)) - | EnvError (AlreadyDeclared id) -> - CErrors.user_err ~hdr:"metacoq" (str "Already declared: " ++ pr_char_list id) - | EnvError (IllFormedDecl (id, e)) -> - CErrors.user_err ~hdr:"metacoq" (pr_char_list (Checker0.string_of_type_error e) ++ str ", while checking " ++ pr_char_list id) -} - -VERNAC COMMAND EXTEND MetaCoqCheck CLASSIFIED AS QUERY -| #[ poly = polymorphic ] [ "MetaCoq" "Check" constr(c) ] -> { - let env = Global.env () in - let evm = Evd.from_env env in - let (c, _) = Constrintern.interp_constr env evm c in - check env evm c - } -END diff --git a/checker/src/metacoq_checker_plugin.mlpack b/checker/src/metacoq_checker_plugin.mlpack deleted file mode 100644 index 34f9bd2fb..000000000 --- a/checker/src/metacoq_checker_plugin.mlpack +++ /dev/null @@ -1,14 +0,0 @@ -EqdepFacts -Wf -WGraph -Monad_utils -Utils -MSetWeakList -UGraph0 -EnvironmentTyping -Typing0 -Checker0 -Retyping0 -TypingWf - -G_metacoq_checker diff --git a/checker/src/to-lower.sh b/checker/src/to-lower.sh new file mode 100755 index 000000000..93bb3e362 --- /dev/null +++ b/checker/src/to-lower.sh @@ -0,0 +1,9 @@ +for i in *.ml* +do + newi=`echo $i | cut -b 1 | tr '[:upper:]' '[:lower:]'``echo $i | cut -b 2-`; + if [ $i != $newi ] + then + echo "Moving " $i "to" $newi; + mv $i $newi; + fi +done diff --git a/checker/theories/All.v b/checker/theories/All.v deleted file mode 100644 index 136a08cb0..000000000 --- a/checker/theories/All.v +++ /dev/null @@ -1,8 +0,0 @@ -(* Distributed under the terms of the MIT license. *) - -Require Export MetaCoq.Template.All. -From MetaCoq.Checker Require Export - Weakening (* Weakening lemmas *) - Substitution (* Weakening lemmas *) - Checker (* Partial typechecker implementation *) - Retyping (* Fast retyping judgment *). diff --git a/checker/theories/Closed.v b/checker/theories/Closed.v deleted file mode 100644 index 672c5ef80..000000000 --- a/checker/theories/Closed.v +++ /dev/null @@ -1,688 +0,0 @@ -(* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import config utils Ast AstUtils Reflect Induction - LiftSubst UnivSubst Typing TypingWf. -From MetaCoq.Checker Require Import WeakeningEnv. - -Require Import ssreflect. -From Equations Require Import Equations. - -(** * Lemmas about the [closedn] predicate *) - - -Definition closed_decl n d := - option_default (closedn n) d.(decl_body) true && closedn n d.(decl_type). - -Definition closedn_ctx n ctx := - forallb id (mapi (fun k d => closed_decl (n + k) d) (List.rev ctx)). - -Lemma closed_decl_upwards k d : closed_decl k d -> forall k', k <= k' -> closed_decl k' d. -Proof. - case: d => na [body|] ty; rewrite /closed_decl /=. - move/andP => [cb cty] k' lek'. do 2 rewrite (@closed_upwards k) //. - move=> cty k' lek'; rewrite (@closed_upwards k) //. -Qed. - -Notation closed_ctx ctx := (closedn_ctx 0 ctx). - -Lemma closedn_lift n k k' t : closedn k t -> closedn (k + n) (lift n k' t). -Proof. - revert k. - induction t in n, k' |- * using term_forall_list_ind; intros; - simpl in *; rewrite -> ?andb_and in *; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc; - simpl closed in *; solve_all; - unfold test_def, test_snd in *; - try solve [simpl lift; simpl closed; f_equal; auto; repeat (rtoProp; solve_all)]; try easy. - - - elim (Nat.leb_spec k' n0); intros. simpl. - elim (Nat.ltb_spec); auto. apply Nat.ltb_lt in H. lia. - simpl. elim (Nat.ltb_spec); auto. intros. - apply Nat.ltb_lt in H. lia. -Qed. - -Lemma closedn_lift_inv n k k' t : k <= k' -> - closedn (k' + n) (lift n k t) -> - closedn k' t. -Proof. - induction t in n, k, k' |- * using term_forall_list_ind; intros; - simpl in *; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length, ?Nat.add_assoc in *; - simpl closed in *; repeat (rtoProp; solve_all); try change_Sk; - unfold test_def, on_snd, test_snd in *; simpl in *; eauto with all. - - - revert H0. - elim (Nat.leb_spec k n0); intros. simpl in *. - elim (Nat.ltb_spec); auto. apply Nat.ltb_lt in H1. intros. lia. - revert H1. simpl. intro. repeat toProp. lia. - - specialize (IHt2 n (S k) (S k')). eauto with all. - - specialize (IHt2 n (S k) (S k')). eauto with all. - - specialize (IHt3 n (S k) (S k')). eauto with all. - - rtoProp. solve_all. specialize (b0 n (#|m| + k) (#|m| + k')). eauto with all. - - rtoProp. solve_all. specialize (b0 n (#|m| + k) (#|m| + k')). eauto with all. -Qed. - -Lemma closedn_mkApps k f u: - closedn k f -> forallb (closedn k) u -> - closedn k (mkApps f u). -Proof. - intros Hf Hu. - induction u; simpl; auto. - destruct f; simpl in *; try rewrite Hf Hu; auto. - rewrite forallb_app. simpl. rewrite Hu. - rewrite andb_assoc. now rewrite Hf. -Qed. - -Lemma closedn_mkApps_inv k f u: - closedn k (mkApps f u) -> - closedn k f && forallb (closedn k) u. -Proof. - induction u; simpl; auto. - - now rewrite andb_true_r. - - intros. destruct f; simpl in *; auto. - rewrite forallb_app in H. simpl in H. - now rewrite andb_assoc in H. -Qed. - -Lemma closedn_subst s k k' t : - forallb (closedn k) s -> closedn (k + k' + #|s|) t -> - closedn (k + k') (subst s k' t). -Proof. - intros Hs. solve_all. revert H. - induction t in k' |- * using term_forall_list_ind; intros; - simpl in *; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; - simpl closed in *; try change_Sk; repeat (rtoProp; solve_all); - unfold test_def, on_snd, test_snd in *; simpl in *; eauto with all. - - - elim (Nat.leb_spec k' n); intros. simpl. - apply Nat.ltb_lt in H. - destruct nth_error eqn:Heq. - -- eapply closedn_lift. - now eapply nth_error_all in Heq; simpl; eauto; simpl in *. - -- simpl. elim (Nat.ltb_spec); auto. intros. - apply nth_error_None in Heq. lia. - -- simpl. apply Nat.ltb_lt in H0. - apply Nat.ltb_lt. apply Nat.ltb_lt in H0. lia. - - - specialize (IHt2 (S k')). - rewrite <- Nat.add_succ_comm in IHt2. eauto. - - specialize (IHt2 (S k')). - rewrite <- Nat.add_succ_comm in IHt2. eauto. - - specialize (IHt3 (S k')). - rewrite <- Nat.add_succ_comm in IHt3. eauto. - - rewrite closedn_mkApps; solve_all. - - rtoProp; solve_all. rewrite -> !Nat.add_assoc in *. - specialize (b0 (#|m| + k')). unfold is_true. rewrite <- b0. f_equal. lia. - unfold is_true. rewrite <- H0. f_equal. lia. - - rtoProp; solve_all. rewrite -> !Nat.add_assoc in *. - specialize (b0 (#|m| + k')). unfold is_true. rewrite <- b0. f_equal. lia. - unfold is_true. rewrite <- H0. f_equal. lia. -Qed. - -Lemma closedn_subst0 s k t : - forallb (closedn k) s -> closedn (k + #|s|) t -> - closedn k (subst0 s t). -Proof. - intros. - generalize (closedn_subst s k 0 t H). - rewrite Nat.add_0_r. eauto. -Qed. - -Lemma subst_closedn s k t : Ast.wf t -> - closedn k t -> subst s k t = t. -Proof. - intros Hcl. - pose proof (simpl_subst_rec _ Hcl s 0 k k). - intros. pose proof (lift_closed (#|s| + 0) _ _ H0). - do 2 (forward H; auto). rewrite H1 in H. - rewrite H. now apply lift_closed. -Qed. - -Lemma closedn_subst_instance_constr k t u : - closedn k (subst_instance_constr u t) = closedn k t. -Proof. - revert k. - induction t in |- * using term_forall_list_ind; intros; - simpl in *; rewrite -> ?andb_and in *; - rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def; - try solve [repeat (f_equal; eauto)]; simpl closed in *; - try rewrite ?map_length; intuition auto. - - - rewrite forallb_map; eapply Forall_forallb_eq_forallb; eauto. - - rewrite forallb_map. f_equal; eauto using Forall_forallb_eq_forallb. - - red in X. rewrite forallb_map. f_equal. - + f_equal; auto. - + eapply Forall_forallb_eq_forallb. - eapply All_Forall; tea. - intros [] XX; apply XX. - - red in X. rewrite forallb_map. - eapply Forall_forallb_eq_forallb. eapply All_Forall; eauto. - unfold test_def, map_def. simpl. - do 3 (f_equal; intuition eauto). - - red in X. rewrite forallb_map. - eapply Forall_forallb_eq_forallb. eapply All_Forall; eauto. - unfold test_def, map_def. simpl. - do 3 (f_equal; intuition eauto). -Qed. - - -Lemma destArity_spec ctx T : - match destArity ctx T with - | Some (ctx', s) => it_mkProd_or_LetIn ctx T = it_mkProd_or_LetIn ctx' (tSort s) - | None => True - end. -Proof. - induction T in ctx |- *; simpl; try easy. - specialize (IHT2 (ctx,, vass na T1)). now destruct destArity. - specialize (IHT3 (ctx,, vdef na T1 T2)). now destruct destArity. -Qed. - -Lemma closedn_All_local_closed: - forall (cf : checker_flags) (Σ : global_env_ext) (ctx : list context_decl) - (wfΓ' : wf_local Σ ctx), - All_local_env_over typing - (fun (Σ0 : global_env_ext) (Γ0 : context) (_ : wf_local Σ0 Γ0) (t T : term) (_ : Σ0;;; Γ0 |- t : T) => - closedn #|Γ0| t && closedn #|Γ0| T) Σ ctx wfΓ' -> - closedn_ctx 0 ctx. -Proof. - intros cf Σ Γ al. - induction 1. constructor. - rewrite /closedn_ctx /= mapi_app forallb_app /= [forallb _ _]IHX /id /closed_decl /=. - now rewrite Nat.add_0_r List.rev_length p. - rewrite /closedn_ctx /= mapi_app forallb_app /= [forallb _ _]IHX /id /closed_decl /=. - now rewrite Nat.add_0_r List.rev_length p. -Qed. - -Lemma closedn_ctx_cons n d Γ : closedn_ctx n (d :: Γ) -> closedn_ctx n Γ && closed_decl (n + #|Γ|) d. -Proof. - unfold closedn_ctx. - simpl. rewrite mapi_app. rewrite forallb_app. - move/andP => [] -> /=. now rewrite Nat.add_0_r List.rev_length andb_true_r. -Qed. - -Lemma closedn_ctx_app n Γ Γ' : - closedn_ctx n (Γ ,,, Γ') = - closedn_ctx n Γ && closedn_ctx (n + #|Γ|) Γ'. -Proof. - rewrite /closedn_ctx /app_context /= List.rev_app_distr mapi_app forallb_app /=. - f_equal. - rewrite List.rev_length. - f_equal. eapply mapi_ext. intros. - f_equal. lia. -Qed. - -Lemma closedn_mkProd_or_LetIn (Γ : context) d T : - closed_decl #|Γ| d -> - closedn (S #|Γ|) T -> closedn #|Γ| (mkProd_or_LetIn d T). -Proof. - destruct d as [na [b|] ty]; simpl in *. unfold closed_decl. - simpl. now move/andP => [] -> ->. - simpl. now move/andP => [] /= _ -> ->. -Qed. - -Lemma closedn_mkLambda_or_LetIn (Γ : context) d T : - closed_decl #|Γ| d -> - closedn (S #|Γ|) T -> closedn #|Γ| (mkLambda_or_LetIn d T). -Proof. - destruct d as [na [b|] ty]; simpl in *. unfold closed_decl. - simpl. now move/andP => [] -> ->. - simpl. now move/andP => [] /= _ -> ->. -Qed. - -Lemma closedn_it_mkProd_or_LetIn - (Γ : context) (ctx : list context_decl) T : - closedn_ctx #|Γ| ctx -> - closedn (#|Γ| + #|ctx|) T -> closedn #|Γ| (it_mkProd_or_LetIn ctx T). -Proof. - induction ctx in Γ, T |- *. simpl. - - now rewrite Nat.add_0_r. - - move/closedn_ctx_cons/andP => [] cctx ca cT. - apply (IHctx Γ (mkProd_or_LetIn a T) cctx). - simpl in cT. rewrite <- app_length. - eapply closedn_mkProd_or_LetIn; - now rewrite app_length // plus_n_Sm. -Qed. - -Lemma closedn_it_mkLambda_or_LetIn - (Γ : context) (ctx : list context_decl) T : - closedn_ctx #|Γ| ctx -> - closedn (#|Γ| + #|ctx|) T -> closedn #|Γ| (it_mkLambda_or_LetIn ctx T). -Proof. - induction ctx in Γ, T |- *. simpl. - - now rewrite Nat.add_0_r. - - move/closedn_ctx_cons/andP => [] cctx ca cT. - apply (IHctx Γ (mkLambda_or_LetIn a T) cctx). - simpl in cT. rewrite <- app_length. - eapply closedn_mkLambda_or_LetIn; - now rewrite app_length // plus_n_Sm. -Qed. - -Lemma type_local_ctx_All_local_env {cf:checker_flags} P Σ Γ Δ s : - All_local_env (lift_typing P Σ) Γ -> - type_local_ctx (lift_typing P) Σ Γ Δ s -> - All_local_env (lift_typing P Σ) (Γ ,,, Δ). -Proof. - induction Δ; simpl; auto. - destruct a as [na [b|] ty]; - intros wfΓ wfctx; constructor; intuition auto. - exists s; auto. -Qed. - -Definition Pclosed := - (fun (_ : global_env_ext) (Γ : context) (t T : term) => - closedn #|Γ| t && closedn #|Γ| T). - -Lemma type_local_ctx_Pclosed Σ Γ Δ s : - type_local_ctx (lift_typing Pclosed) Σ Γ Δ s -> - Alli (fun i d => closed_decl (#|Γ| + i) d) 0 (List.rev Δ). -Proof. - induction Δ; simpl; auto; try constructor. - destruct a as [? [] ?]; intuition auto. - - apply Alli_app_inv; auto. constructor. simpl. - rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in b0. simpl. - rewrite app_context_length in b0. now rewrite Nat.add_comm. - - apply Alli_app_inv; auto. constructor. simpl. - rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in b. simpl. - rewrite app_context_length in b. rewrite Nat.add_comm. - now rewrite andb_true_r in b. -Qed. - -Lemma sorts_local_ctx_Pclosed Σ Γ Δ s : - sorts_local_ctx (lift_typing Pclosed) Σ Γ Δ s -> - Alli (fun i d => closed_decl (#|Γ| + i) d) 0 (List.rev Δ). -Proof. - induction Δ in s |- *; simpl; auto; try constructor. - destruct a as [? [] ?]; intuition auto. - - apply Alli_app_inv; eauto. constructor. simpl. - rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in b0. simpl. - rewrite app_context_length in b0. now rewrite Nat.add_comm. - - destruct s; auto. destruct X. - apply Alli_app_inv; eauto. constructor. simpl. - rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in i. simpl. - rewrite app_context_length in i. rewrite Nat.add_comm. - now rewrite andb_true_r in i. -Qed. - -Lemma All_local_env_Pclosed Σ Γ : - All_local_env ( lift_typing Pclosed Σ) Γ -> - Alli (fun i d => closed_decl i d) 0 (List.rev Γ). -Proof. - induction Γ; simpl; auto; try constructor. - intros all; depelim all; intuition auto. - - apply Alli_app_inv; auto. constructor. simpl. - rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed in l. simpl. red in l. - destruct l as [s H]. - now rewrite andb_true_r in H. - - apply Alli_app_inv; auto. constructor. simpl. - rewrite List.rev_length. 2:constructor. - unfold closed_decl. unfold Pclosed, lift_typing in *. now simpl. -Qed. - -Lemma closed_subst_context (Γ Δ Δ' : context) t : - closedn (#|Γ| + #|Δ|) t -> - Alli (fun i d => closed_decl (#|Γ| + S #|Δ| + i) d) 0 (List.rev Δ') -> - Alli (fun i d => closed_decl (#|Γ| + #|Δ| + i) d) 0 (List.rev (subst_context [t] 0 Δ')). -Proof. - induction Δ' in Δ |- *. - intros. simpl. auto. constructor. - destruct a as [? [] ?]; simpl. - - intros. eapply Alli_app in X as [X X']. - rewrite subst_context_snoc. simpl. eapply Alli_app_inv. - eapply IHΔ'; eauto. constructor; [|constructor]. - simpl. - rewrite /closed_decl /map_decl /= Nat.add_0_r List.rev_length subst_context_length. - inv X'. unfold closed_decl in H0. simpl in H0. - rewrite List.rev_length Nat.add_0_r in H0. - move/andP: H0 => [Hl Hr]. - rewrite !closedn_subst /= ?H //; eapply closed_upwards; eauto; try lia. - - intros. eapply Alli_app in X as [X X']. - rewrite subst_context_snoc. simpl. eapply Alli_app_inv. - eapply IHΔ'; eauto. constructor; [|constructor]. - simpl. - rewrite /closed_decl /map_decl /= Nat.add_0_r List.rev_length subst_context_length. - inv X'. unfold closed_decl in H0. simpl in H0. - rewrite List.rev_length Nat.add_0_r in H0. - rewrite !closedn_subst /= ?H //; eapply closed_upwards; eauto; try lia. -Qed. - -Lemma closed_smash_context_gen (Γ Δ Δ' : context) : - Alli (fun i d => closed_decl (#|Γ| + i) d) 0 (List.rev Δ) -> - Alli (fun i d => closed_decl (#|Γ| + #|Δ| + i) d) 0 (List.rev Δ') -> - Alli (fun i d => closed_decl (#|Γ| + i) d) 0 (List.rev (smash_context Δ' Δ)). -Proof. - induction Δ in Δ' |- *. - intros. simpl. auto. - simpl. eapply (Alli_impl _ X0). simpl. now rewrite Nat.add_0_r. - destruct a as [? [] ?]; simpl. - - intros. eapply Alli_app in X as [X X']. - eapply IHΔ; eauto. - inv X'. unfold closed_decl in H. simpl in H. - rewrite List.rev_length Nat.add_0_r in H. - clear X1. eapply closed_subst_context; auto. - now move/andP: H => []. - - intros. eapply Alli_app in X as [X X']. - eapply IHΔ; eauto. - inv X'. unfold closed_decl in H. simpl in H. - rewrite List.rev_length Nat.add_0_r in H. - clear X1. rewrite List.rev_app_distr. - eapply Alli_app_inv; repeat constructor. - now rewrite /closed_decl Nat.add_0_r /=. - rewrite List.rev_length /=. clear -X0. - apply Alli_shift, (Alli_impl _ X0). intros. - eapply closed_decl_upwards; eauto; lia. -Qed. - -Lemma closed_smash_context (Γ Δ : context) : - Alli (fun i d => closed_decl (#|Γ| + i) d) 0 (List.rev Δ) -> - Alli (fun i d => closed_decl (#|Γ| + i) d) 0 (List.rev (smash_context [] Δ)). -Proof. - intros; apply (closed_smash_context_gen Γ _ []); auto. constructor. -Qed. - -Lemma context_assumptions_app Γ Δ : context_assumptions (Γ ++ Δ) = - context_assumptions Γ + context_assumptions Δ. -Proof. - induction Γ as [|[? [] ?] ?]; simpl; auto. -Qed. - -Lemma weaken_env_prop_closed {cf:checker_flags} : - weaken_env_prop (lift_typing (fun (_ : global_env_ext) (Γ : context) (t T : term) => - closedn #|Γ| t && closedn #|Γ| T)). -Proof. repeat red. intros. destruct t; red in X0; eauto. Qed. - -Lemma declared_projection_closed_ind {cf:checker_flags} {Σ : global_env_ext} {mdecl idecl p pdecl} : - wf Σ.1 -> - declared_projection Σ.1 mdecl idecl p pdecl -> - Forall_decls_typing - (fun (_ : TemplateEnvironment.global_env_ext) (Γ : TemplateEnvironment.context) (t T : term) => - closedn #|Γ| t && closedn #|Γ| T) Σ.1 -> - closedn (S (ind_npars mdecl)) pdecl.2. -Proof. - intros wfΣ isdecl X0. - pose proof (declared_projection_inv weaken_env_prop_closed wfΣ X0 isdecl) as onp. - set (declared_inductive_inv _ wfΣ X0 isdecl.p1) as oib in *. - clearbody oib. - have onpars := onParams (declared_minductive_inv weaken_env_prop_closed wfΣ X0 isdecl.p1.p1). - have parslen := onNpars (declared_minductive_inv weaken_env_prop_closed wfΣ X0 isdecl.p1.p1). - simpl in onp. destruct (ind_cshapes oib) as [|? []] eqn:Heq; try contradiction. - red in onp. - destruct (nth_error (smash_context [] _) _) eqn:Heq'; try contradiction. - destruct onp as [onna onp]. - rewrite {}onp. - pose proof (onConstructors oib) as onc. - red in onc. rewrite Heq in onc. inv onc. clear X1. - eapply on_cargs in X. - simpl. - replace (S (ind_npars mdecl)) with (1 + ind_npars mdecl) by lia. - eapply (closedn_subst _ 0). - { clear. unfold inds. generalize (TemplateEnvironment.ind_bodies mdecl). - induction l; simpl; auto. } - rewrite inds_length. - eapply closedn_subst0. - { clear. unfold projs. induction p.2; simpl; auto. } - rewrite projs_length /=. - eapply (@closed_upwards (ind_npars mdecl + #|ind_bodies mdecl| + p.2 + 1)). - 2:lia. - eapply closedn_lift. - clear -parslen isdecl Heq' onpars X. - rename X into args. - apply sorts_local_ctx_Pclosed in args. - red in onpars. - eapply All_local_env_Pclosed in onpars. - eapply (Alli_impl (Q:=fun i d => closed_decl (#|ind_params mdecl| + i + #|arities_context (ind_bodies mdecl)|) d)) in args. - 2:{ intros n x. rewrite app_context_length. - intros H; eapply closed_decl_upwards; eauto. lia. } - eapply (Alli_shiftn (P:=fun i d => closed_decl (i + #|arities_context (ind_bodies mdecl)|) d)) in args. - rewrite Nat.add_0_r in args. - eapply (Alli_impl (Q:=fun i d => closed_decl (i + #|arities_context (ind_bodies mdecl)|) d)) in onpars. - 2:{ intros n x d. eapply closed_decl_upwards; eauto; lia. } - epose proof (Alli_app_inv onpars). rewrite List.rev_length /= in X. - specialize (X args). rewrite -List.rev_app_distr in X. - clear onpars args. - eapply (Alli_impl (Q:=fun i d => closed_decl (#|arities_context (ind_bodies mdecl)| + i) d)) in X. - 2:intros; eapply closed_decl_upwards; eauto; lia. - eapply closed_smash_context in X. - eapply Alli_rev_nth_error in X; eauto. - rewrite smash_context_length in X. simpl in X. - eapply nth_error_Some_length in Heq'. rewrite smash_context_length in Heq'. - simpl in Heq'. unfold closed_decl in X. - move/andP: X => [] _ cl. - eapply closed_upwards. eauto. rewrite arities_context_length. - rewrite context_assumptions_app parslen. - rewrite context_assumptions_app in Heq'. lia. -Qed. - -Lemma typecheck_closed `{cf : checker_flags} : - env_prop (fun Σ Γ t T => - closedn #|Γ| t && closedn #|Γ| T). -Proof. - assert (X := weaken_env_prop_closed). - apply typing_ind_env; intros; simpl; rewrite -> ?andb_and in *; try solve [intuition auto]. - - pose proof (nth_error_Some_length H). - elim (Nat.ltb_spec n #|Γ|); intuition. clear H0. - eapply (nth_error_All_local_env_over H) in X0 as [HΓ Hdecl]. - destruct lookup_wf_local_decl; cbn in *. - destruct decl as [na [b|] ty]; cbn in *. - -- move/andb_and: Hdecl => [] _. - rewrite skipn_length; try lia. - move/(closedn_lift (S n)). - now have->: #|Γ| - S n + S n = #|Γ| by lia. - -- rewrite andb_true_r in Hdecl. - move/(closedn_lift (S n)): Hdecl. - rewrite skipn_length; try lia. - now have->: #|Γ| - S n + S n = #|Γ| by lia. - - - destruct H. rewrite and_assoc. split. auto. - clear X0 H H0 H1. - induction X1; simpl. intuition auto. - rewrite -> andb_and in *. destruct p0. rewrite H. - rewrite <- andb_and. simpl. - apply IHX1. unfold subst. - pose proof (closedn_subst [hd] #|Γ| 0). rewrite Nat.add_0_r in H1. - apply H1. simpl. now rewrite H. - simpl. simpl in p. rewrite -> andb_and in p. intuition auto. - rewrite Nat.add_1_r. auto. - - - rewrite closedn_subst_instance_constr. - eapply lookup_on_global_env in H; eauto. - destruct H as [Σ' [HΣ' IH]]. - repeat red in IH. destruct decl, cst_body. simpl in *. - rewrite -> andb_and in IH. intuition. - eauto using closed_upwards with arith. - simpl in *. - repeat red in IH. destruct IH as [s Hs]. - rewrite -> andb_and in Hs. intuition. - eauto using closed_upwards with arith. - - - rewrite closedn_subst_instance_constr. - eapply declared_inductive_inv in X0; eauto. - apply onArity in X0. repeat red in X0. - destruct X0 as [s Hs]. rewrite -> andb_and in Hs. - intuition eauto using closed_upwards with arith. - - - destruct isdecl as [Hidecl Hcdecl]. - apply closedn_subst0. - + unfold inds. clear. simpl. induction #|ind_bodies mdecl|. constructor. - simpl. now rewrite IHn. - + rewrite inds_length. - rewrite closedn_subst_instance_constr. - eapply declared_inductive_inv in X0; eauto. - pose proof X0.(onConstructors) as XX. - eapply All2_nth_error_Some in Hcdecl; eauto. - destruct Hcdecl as [? [? ?]]. cbn in *. - destruct o.(on_ctype) as [s Hs]. rewrite -> andb_and in Hs. - apply proj1 in Hs. - unfold arities_context in Hs. - rewrite rev_map_length in Hs. - eauto using closed_upwards with arith. - - - intuition auto. solve_all. unfold test_snd. simpl in *. - rtoProp; eauto. - apply closedn_mkApps; auto. - rewrite forallb_app. simpl. rewrite H1. - rewrite forallb_skipn; auto. - now apply closedn_mkApps_inv in H7. - - - intuition auto. - apply closedn_subst0. - simpl. apply closedn_mkApps_inv in H2. - rewrite forallb_rev H1. apply H2. - simpl. rewrite List.rev_length. rewrite H0. - rewrite closedn_subst_instance_constr. - eapply closed_upwards. eapply declared_projection_closed_ind; eauto. - lia. - - - split. solve_all. - destruct x; simpl in *. - unfold test_def. simpl. rtoProp. - split. - rewrite -> app_context_length in *. rewrite -> Nat.add_comm in *. - eapply closedn_lift_inv in H2; eauto. lia. - subst types. - now rewrite app_context_length fix_context_length in H1. - eapply nth_error_all in H0; eauto. simpl in H0. intuition. rtoProp. - subst types. rewrite app_context_length in H1. - rewrite Nat.add_comm in H1. - now eapply closedn_lift_inv in H1. - - - split. solve_all. destruct x; simpl in *. - unfold test_def. simpl. rtoProp. - split. - rewrite -> app_context_length in *. rewrite -> Nat.add_comm in *. - eapply closedn_lift_inv in H1; eauto. lia. - subst types. - now rewrite -> app_context_length, fix_context_length in H0. - eapply (nth_error_all) in H; eauto. simpl in *. - intuition. rtoProp. - subst types. rewrite app_context_length in H0. - rewrite Nat.add_comm in H0. - now eapply closedn_lift_inv in H0. - - - destruct X2; intuition eauto. - + destruct i as [[u [ctx [Heq Hi]]] Hwfi]. simpl in Hwfi. - generalize (destArity_spec [] B). rewrite Heq. - simpl; intros ->. - apply closedn_All_local_closed in Hwfi. - rewrite closedn_ctx_app in Hwfi. - move/andP: Hwfi => [] clΓ clctx. - apply closedn_it_mkProd_or_LetIn => //. - + destruct s. rewrite andb_true_r in p. intuition auto. -Qed. - -Lemma declared_decl_closed `{checker_flags} {Σ : global_env} {cst decl} : - wf Σ -> - lookup_env Σ cst = Some decl -> - on_global_decl (fun Σ Γ b t => closedn #|Γ| b && option_default (closedn #|Γ|) t true) - (Σ, universes_decl_of_decl decl) cst decl. -Proof. - intros. - eapply weaken_lookup_on_global_env; try red; eauto. - eapply on_global_env_impl; cycle 1. - apply (env_prop_sigma _ typecheck_closed _ X). - red; intros. unfold lift_typing in *. destruct T; intuition auto with wf. - destruct X1 as [s0 Hs0]. simpl. rtoProp; intuition. -Qed. - -Definition closed_inductive_body mdecl idecl := - closedn 0 idecl.(ind_type) && - forallb (fun cdecl => - closedn (#|arities_context mdecl.(ind_bodies)|) (cdecl_type cdecl)) idecl.(ind_ctors) && - forallb (fun x => closedn (S (ind_npars mdecl)) x.2) idecl.(ind_projs). - -Definition closed_inductive_decl mdecl := - closed_ctx (ind_params mdecl) && - forallb (closed_inductive_body mdecl) (ind_bodies mdecl). - -Lemma closedn_All_local_env (ctx : list context_decl) : - All_local_env - (fun (Γ : context) (b : term) (t : option term) => - closedn #|Γ| b && option_default (closedn #|Γ|) t true) ctx -> - closedn_ctx 0 ctx. -Proof. - induction 1. constructor. - rewrite /closedn_ctx /= mapi_app forallb_app /= [forallb _ _]IHX /id /closed_decl /=. - now rewrite Nat.add_0_r List.rev_length t0. - rewrite /closedn_ctx /= mapi_app forallb_app /= [forallb _ _]IHX /id /closed_decl /=. - now rewrite Nat.add_0_r List.rev_length t1. -Qed. - -Lemma skipn_0 {A} (l : list A) : skipn 0 l = l. -Proof. reflexivity. Qed. - -Lemma declared_projection_closed {cf:checker_flags} {Σ : global_env_ext} {mdecl idecl p pdecl} : - wf Σ.1 -> - declared_projection Σ.1 mdecl idecl p pdecl -> - closedn (S (ind_npars mdecl)) pdecl.2. -Proof. - intros; eapply declared_projection_closed_ind; eauto. - eapply typecheck_closed; eauto. - constructor. eapply type_Prop. -Qed. - -Lemma declared_inductive_closed {cf:checker_flags} {Σ : global_env_ext} {mdecl mind} : - wf Σ.1 -> - declared_minductive Σ.1 mind mdecl -> - closed_inductive_decl mdecl. -Proof. - intros wfΣ decl. - pose proof (declared_decl_closed wfΣ decl) as decl'. - red in decl'. - unfold closed_inductive_decl. - apply andb_and. split. apply onParams in decl'. - now apply closedn_All_local_env in decl'; auto. - apply onInductives in decl'. - eapply All_forallb. - - assert (Alli (fun i => declared_inductive Σ.1 mdecl {| inductive_mind := mind; inductive_ind := i |}) - 0 (ind_bodies mdecl)). - { eapply forall_nth_error_Alli. intros. - split; auto. } - eapply Alli_mix in decl'; eauto. clear X. - clear decl. - eapply Alli_All; eauto. - intros n x [decli oib]. - rewrite /closed_inductive_body. - apply andb_and; split. apply andb_and. split. - - apply onArity in oib. hnf in oib. - now move/andP: oib => []. - - pose proof (onConstructors oib). - red in X. eapply All_forallb. eapply All2_All_left; eauto. - firstorder auto. - eapply on_ctype in X0. - now move/andP: X0 => []. - - eapply All_forallb. - pose proof (onProjections oib). - destruct (eq_dec (ind_projs x) []) as [->|eq]; try constructor. - specialize (X eq). clear eq. - destruct (ind_cshapes oib) as [|? []]; try contradiction. - apply on_projs in X. - assert (Alli (fun i pdecl => declared_projection Σ.1 mdecl x - (({| inductive_mind := mind; inductive_ind := n |}, mdecl.(ind_npars)), i) pdecl) - 0 (ind_projs x)). - { eapply forall_nth_error_Alli. - intros. split; auto. } - eapply (Alli_All X0). intros. - now eapply declared_projection_closed in H. -Qed. - -Lemma rev_subst_instance_context u Γ : - List.rev (subst_instance_context u Γ) = subst_instance_context u (List.rev Γ). -Proof. - unfold subst_instance_context, map_context. - now rewrite map_rev. -Qed. - -Lemma closedn_subst_instance_context k Γ u : - closedn_ctx k (subst_instance_context u Γ) = closedn_ctx k Γ. -Proof. - unfold closedn_ctx; f_equal. - rewrite rev_subst_instance_context. - rewrite mapi_map. apply mapi_ext. - intros n [? [] ?]; unfold closed_decl; cbn. - all: now rewrite !closedn_subst_instance_constr. -Qed. diff --git a/checker/theories/Extraction.v b/checker/theories/Extraction.v deleted file mode 100644 index cc159c287..000000000 --- a/checker/theories/Extraction.v +++ /dev/null @@ -1,48 +0,0 @@ -(* Distributed under the terms of the MIT license. *) -Require Import MetaCoq.Template.utils. -Require Import FSets ExtrOcamlBasic ExtrOcamlString ExtrOcamlZInt. - -(** * Extraction setup for template-coq. - - Any extracted code planning to link with the plugin's OCaml reifier - should use these same directives for consistency. -*) - - -(* Ignore [Decimal.int] before the extraction issue is solved: - https://github.com/coq/coq/issues/7017. *) -Extract Inductive Decimal.int => unit [ "(fun _ -> ())" "(fun _ -> ())" ] "(fun _ _ _ -> assert false)". -Extract Inductive Hexadecimal.int => unit [ "(fun _ -> ())" "(fun _ -> ())" ] "(fun _ _ _ -> assert false)". -Extract Inductive Numeral.int => unit [ "(fun _ -> ())" "(fun _ -> ())" ] "(fun _ _ _ -> assert false)". - -Extract Constant ascii_compare => - "fun x y -> match Char.compare x y with 0 -> Eq | x when x < 0 -> Lt | _ -> Gt". - -Extraction Blacklist config uGraph Universes Ast String List Nat Int - UnivSubst Typing Checker Retyping OrderedType Logic Common Equality utils. -Set Warnings "-extraction-opaque-accessed". -Set Warnings "-extraction-reserved-identifier". - -Require Export MetaCoq.Template.Ast. -From MetaCoq.Checker Require All. - -Cd "src". - -From Equations Require Import Equations. - -(* Should be in Equations *) -Extraction Inline Equations.Prop.Classes.noConfusion. -Extraction Inline Equations.Prop.Logic.eq_elim. -Extraction Inline Equations.Prop.Logic.eq_elim_r. -Extraction Inline Equations.Prop.Logic.transport. -Extraction Inline Equations.Prop.Logic.transport_r. -Extraction Inline Equations.Prop.Logic.False_rect_dep. -Extraction Inline Equations.Prop.Logic.True_rect_dep. -Extraction Inline Equations.Init.pr1. -Extraction Inline Equations.Init.pr2. -Extraction Inline Equations.Init.hidebody. -Extraction Inline Equations.Prop.DepElim.solution_left. - -Separate Extraction Checker.typecheck_program Checker.infer' Retyping.sort_of. - -Cd "..". diff --git a/checker/theories/Generation.v b/checker/theories/Generation.v deleted file mode 100644 index db01d128f..000000000 --- a/checker/theories/Generation.v +++ /dev/null @@ -1,53 +0,0 @@ -(* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import config utils Ast AstUtils Reflect LiftSubst Typing. -Require Import Equations.Prop.DepElim. - -(** * Substitution lemmas for typing derivations. *) - -Derive Signature for typing. - -Lemma invert_type_App `{checker_flags} Σ Γ f u T : - Σ ;;; Γ |- tApp f u : T -> - { T' : term & { U' & ((Σ ;;; Γ |- f : T') * typing_spine Σ Γ T' u U' * - (isApp f = false) * (u <> []) * - (Σ ;;; Γ |- U' <= T))%type } }. -Proof. - intros Hty. - dependent induction Hty. - - exists t_ty, t'. intuition. - - destruct IHHty as [T' [U' [H' H'']]]. - exists T', U'. split; auto. - eapply cumul_trans; eauto. -Qed. - -Lemma type_mkApp `{checker_flags} Σ Γ f u T T' : - Σ ;;; Γ |- f : T -> - typing_spine Σ Γ T [u] T' -> - Σ ;;; Γ |- mkApp f u : T'. -Proof. - intros Hf Hu. inv Hu. - simpl. destruct (isApp f) eqn:Happ. - destruct f; try discriminate. simpl. - eapply invert_type_App in Hf. - destruct Hf as (T'' & U' & (((Hf & HU) & Happf) & Hunil) & Hcumul). - eapply type_App; eauto. intro. destruct args; discriminate. - inv X2. clear Happ Hf Hunil. - induction HU. simpl. econstructor; eauto. - eapply cumul_trans; eauto. - econstructor. econstructor. eapply t. eauto. eauto. - apply IHHU; eauto. - rewrite -> mkApp_tApp; eauto. - econstructor; eauto. congruence. - econstructor; eauto. -Qed. - -Lemma type_mkApps `{checker_flags} Σ Γ f u T T' : - Σ ;;; Γ |- f : T -> - typing_spine Σ Γ T u T' -> - Σ ;;; Γ |- mkApps f u : T'. -Proof. - intros Hf Hu. induction Hu in f, Hf |- *. auto. - rewrite <- mkApps_mkApp. - eapply IHHu. eapply type_mkApp. eauto. - econstructor; eauto. constructor. -Qed. diff --git a/checker/theories/Loader.v b/checker/theories/Loader.v deleted file mode 100644 index 481ef62fc..000000000 --- a/checker/theories/Loader.v +++ /dev/null @@ -1,6 +0,0 @@ -(* Distributed under the terms of the MIT license. *) -Require Import String. -From MetaCoq.Template Require Import BasicAst Ast Loader. - -Declare ML Module "metacoq_template_plugin". -Declare ML Module "metacoq_checker_plugin". diff --git a/checker/theories/Retyping.v b/checker/theories/Retyping.v deleted file mode 100644 index b91fcc9e6..000000000 --- a/checker/theories/Retyping.v +++ /dev/null @@ -1,108 +0,0 @@ -(* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import config utils Ast LiftSubst. -From MetaCoq.Checker Require Import Checker. - -(** * Retyping - - The [type_of] function provides a fast (and loose) type inference - function which assumes that the provided term is already well-typed in - the given context and recomputes its type. Only reduction (for - head-reducing types to uncover dependent products or inductives) and - substitution are costly here. No universe checking or conversion is done - in particular. *) - - -Local Existing Instance default_checker_flags. - -Section TypeOf. - Context `{F : Fuel}. - Context (Σ : global_env). - - Fixpoint infer_spine (Γ : context) (ty : term) (l : list term) - {struct l} : typing_result term := - match l with - | nil => ret ty - | cons x xs => - pi <- reduce_to_prod Σ Γ ty ;; - let '(a1, b1) := pi in - infer_spine Γ (subst10 x b1) xs - end. - - Section SortOf. - Context (type_of : context -> term -> typing_result term). - - Definition type_of_as_sort Γ t := - tx <- type_of Γ t ;; - reduce_to_sort Σ Γ tx. - - End SortOf. - - Fixpoint type_of (Γ : context) (t : term) : typing_result term := - match t with - | tRel n => - match nth_error Γ n with - | Some d => ret (lift0 (S n) d.(decl_type)) - | None => raise (UnboundRel n) - end - - | tVar n => raise (UnboundVar n) - | tEvar ev args => raise (UnboundEvar ev) - - | tSort s => ret (tSort (Universe.super s)) - - | tCast c k t => ret t - - | tProd n t b => - s1 <- type_of_as_sort type_of Γ t ;; - s2 <- type_of_as_sort type_of (Γ ,, vass n t) b ;; - ret (tSort (Universe.sort_of_product s1 s2)) - - | tLambda n t b => - t2 <- type_of (Γ ,, vass n t) b ;; - ret (tProd n t t2) - - | tLetIn n b b_ty b' => - b'_ty <- type_of (Γ ,, vdef n b b_ty) b' ;; - ret (tLetIn n b b_ty b'_ty) - - | tApp t l => - t_ty <- type_of Γ t ;; - infer_spine Γ t_ty l - - | tConst cst u => lookup_constant_type Σ cst u - - | tInd (mkInd ind i) u => lookup_ind_type Σ ind i u - - | tConstruct (mkInd ind i) k u => - lookup_constructor_type Σ ind i k u - - | tCase ((ind, par), rel) p c brs => - ty <- type_of Γ c ;; - indargs <- reduce_to_ind Σ Γ ty ;; - let '(ind', u, args) := indargs in - ret (tApp p (List.skipn par args ++ [c])) - - | tProj p c => - ty <- type_of Γ c ;; - indargs <- reduce_to_ind Σ Γ ty ;; - (* FIXME *) - ret ty - - | tFix mfix n => - match nth_error mfix n with - | Some f => ret f.(dtype) - | None => raise (IllFormedFix mfix n) - end - - | tCoFix mfix n => - match nth_error mfix n with - | Some f => ret f.(dtype) - | None => raise (IllFormedFix mfix n) - end - end. - - Definition sort_of (Γ : context) (t : term) : typing_result Universe.t := - ty <- type_of Γ t;; - type_of_as_sort type_of Γ ty. - -End TypeOf. diff --git a/checker/theories/Substitution.v b/checker/theories/Substitution.v deleted file mode 100644 index d8fa43330..000000000 --- a/checker/theories/Substitution.v +++ /dev/null @@ -1,1923 +0,0 @@ -(* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import config utils Ast AstUtils Induction LiftSubst - UnivSubst Typing TypingWf. -From MetaCoq.Checker Require Import Generation Closed WeakeningEnv Weakening. - -From Equations Require Import Equations. -Require Import ssreflect. - -(** * Substitution lemmas for typing derivations. *) - - -Generalizable Variables Σ Γ t T. - -Derive Signature for Ast.wf. - -Definition subst_decl s k (d : context_decl) := map_decl (subst s k) d. - -Definition subst_context n k (Γ : context) : context := - fold_context (fun k' => subst n (k' + k)) Γ. - -(** Well-typed substitution into a context with *no* let-ins *) - -Inductive subs `{cf : checker_flags} Σ (Γ : context) : list term -> context -> Type := -| emptys : subs Σ Γ [] [] -| cons_ass Δ s na t T : subs Σ Γ s Δ -> - Σ ;;; Γ |- t : subst0 s T -> - subs Σ Γ (t :: s) (Δ ,, vass na T). -(* | cons_let Δ s na t T : subs Σ Γ s Δ -> *) -(* Σ ;;; Γ |- subst0 s t : subst0 s T -> *) -(* subs Σ Γ (subst0 s t :: s) (Δ ,, vdef na t T). *) - -(** Linking a context (with let-ins), an instance (reversed substitution) - for its assumptions and a well-formed substitution for it. *) - -Inductive context_subst : context -> list term -> list term -> Type := -| context_subst_nil : context_subst [] [] [] -| context_subst_ass Γ args s na t a : - context_subst Γ args s -> - context_subst (vass na t :: Γ) (args ++ [a]) (a :: s) -| context_subst_def Γ args s na b t : - context_subst Γ args s -> - context_subst (vdef na b t :: Γ) args (subst s 0 b :: s). - -(** Promoting a substitution for the non-let declarations of ctx into a - substitution for the whole context *) - -Fixpoint make_context_subst ctx args s := - match ctx with - | [] => match args with - | [] => Some s - | a :: args => None - end - | d :: ctx => - match d.(decl_body) with - | Some body => make_context_subst ctx args (subst0 s body :: s) - | None => match args with - | a :: args => make_context_subst ctx args (a :: s) - | [] => None - end - end - end. - -Lemma subst_decl0 Γ k d : on_local_decl wf_decl_pred Γ d -> map_decl (subst [] k) d = d. -Proof. - unfold wf_decl_pred; intros Hd; destruct d; destruct decl_body; - unfold on_local_decl in Hd; unfold subst_decl, map_decl; simpl in *; - f_equal; simpl; rewrite subst_empty; intuition trivial. -Qed. - -Lemma subst0_context k Γ : All_local_env wf_decl_pred Γ -> subst_context [] k Γ = Γ. -Proof. - unfold subst_context, fold_context. - rewrite rev_mapi. rewrite List.rev_involutive. - unfold mapi. generalize 0. generalize #|List.rev Γ|. - induction Γ; intros; simpl; trivial. - eapply All_local_env_wf_decl_inv in X as [Ha HΓ]. - erewrite subst_decl0; f_equal; eauto. -Qed. - -Lemma fold_context_length f Γ : #|fold_context f Γ| = #|Γ|. -Proof. - unfold fold_context. now rewrite !List.rev_length mapi_length List.rev_length. -Qed. - -Lemma subst_context_length s k Γ : #|subst_context s k Γ| = #|Γ|. -Proof. - unfold subst_context. apply fold_context_length. -Qed. -Hint Rewrite subst_context_length : subst. - -Lemma subst_context_snoc s k Γ d : subst_context s k (d :: Γ) = subst_context s k Γ ,, subst_decl s (#|Γ| + k) d. -Proof. - unfold subst_context, fold_context. - rewrite !rev_mapi !rev_involutive /mapi mapi_rec_eqn /snoc. - f_equal. now rewrite Nat.sub_0_r List.rev_length. - rewrite mapi_rec_Sk. simpl. apply mapi_rec_ext. intros. - rewrite app_length !List.rev_length. simpl. f_equal. f_equal. lia. -Qed. -Hint Rewrite subst_context_snoc : subst. - -Lemma subst_context_snoc0 s Γ d : subst_context s 0 (Γ ,, d) = subst_context s 0 Γ ,, subst_decl s #|Γ| d. -Proof. - unfold snoc. now rewrite subst_context_snoc Nat.add_0_r. -Qed. -Hint Rewrite subst_context_snoc : subst. - -Lemma subst_context_alt s k Γ : - subst_context s k Γ = - mapi (fun k' d => subst_decl s (Nat.pred #|Γ| - k' + k) d) Γ. -Proof. - unfold subst_context, fold_context. rewrite rev_mapi. rewrite List.rev_involutive. - apply mapi_ext. intros. f_equal. now rewrite List.rev_length. -Qed. - -Lemma subst_context_app s k Γ Δ : - subst_context s k (Γ ,,, Δ) = subst_context s k Γ ,,, subst_context s (#|Γ| + k) Δ. -Proof. - unfold subst_context, fold_context, app_context. - rewrite List.rev_app_distr. - rewrite mapi_app. rewrite <- List.rev_app_distr. f_equal. f_equal. - apply mapi_ext. intros. f_equal. rewrite List.rev_length. f_equal. lia. -Qed. - - -Lemma subst_decl_closed n k d : wf_decl d -> closed_decl k d -> subst_decl n k d = d. -Proof. - case: d => na [body|] ty; rewrite /closed_decl /subst_decl /map_decl /wf_decl /= => [[wfb wfty]]. - move/andP => [cb cty]. rewrite !subst_closedn //. - move=> cty; now rewrite !subst_closedn //. -Qed. - -Lemma closed_ctx_subst n k ctx : All wf_decl ctx -> closed_ctx ctx = true -> subst_context n k ctx = ctx. -Proof. - induction ctx in n, k |- *; auto. - unfold closed_ctx, id. intros wfctx. inv wfctx. - rewrite mapi_app forallb_app List.rev_length /= Nat.add_0_r. - move/andb_and => /= [Hctx /andb_and [Ha _]]. - rewrite subst_context_snoc /snoc /= IHctx // subst_decl_closed //. - now apply: closed_decl_upwards. -Qed. - -Lemma map_vass_map_def g l n k : - (mapi (fun i (d : def term) => vass (dname d) (lift0 i (dtype d))) - (map (map_def (subst n k) g) l)) = - (mapi (fun i d => map_decl (subst n (i + k)) d) (mapi (fun i (d : def term) => vass (dname d) (lift0 i (dtype d))) l)). -Proof. - rewrite mapi_mapi mapi_map. apply mapi_ext. - intros. unfold map_decl, vass; simpl; f_equal. - rewrite commut_lift_subst_rec. lia. f_equal; lia. -Qed. - -Lemma All_local_env_subst (P Q : context -> term -> option term -> Type) c n k : - All_local_env Q c -> - (forall Γ t T, - Q Γ t T -> - P (subst_context n k Γ) (subst n (#|Γ| + k) t) - (option_map (subst n (#|Γ| + k)) T) - ) -> - All_local_env P (subst_context n k c). -Proof. - intros Hq Hf. - induction Hq in |- *; try econstructor; eauto; - simpl; unfold snoc; rewrite subst_context_snoc; econstructor; eauto. - - simpl. eapply (Hf _ _ None). eauto. - - simpl. eapply (Hf _ _ None). eauto. - - simpl. eapply (Hf _ _ (Some t)). eauto. -Qed. - -Lemma nth_error_subst_context (Γ' : context) s (v : nat) k : - nth_error (subst_context s k Γ') v = - option_map (subst_decl s (#|Γ'| - S v + k)) (nth_error Γ' v). -Proof. - induction Γ' in v |- *; intros. - - simpl. unfold subst_context, fold_context; simpl; rewrite nth_error_nil. easy. - - simpl. destruct v; rewrite subst_context_snoc. - + simpl. repeat f_equal; try lia. - + simpl. rewrite IHΓ'; simpl in *; (lia || congruence). -Qed. - -Lemma subst_length `{checker_flags} Σ Γ s Γ' : subs Σ Γ s Γ' -> #|s| = #|Γ'|. -Proof. - induction 1; simpl; auto with arith. -Qed. - -Hint Rewrite @app_context_length @subst_context_length : wf. - -Lemma subs_nth_error_ge `{checker_flags} Σ Γ Γ' Γ'' v s : - subs Σ Γ s Γ' -> - #|Γ' ,,, Γ''| <= v -> - nth_error (Γ ,,, Γ' ,,, Γ'') v = - nth_error (Γ ,,, subst_context s 0 Γ'') (v - #|Γ'|). -Proof. - simpl. - intros. rewrite app_context_length in H0. - rewrite !nth_error_app_ge; autorewrite with wf; f_equal; try lia. -Qed. - -Lemma subs_nth_error_lt `{checker_flags} Σ Γ Γ' Γ'' v s : - subs Σ Γ s Γ' -> - v < #|Γ''| -> - nth_error (Γ ,,, subst_context s 0 Γ'') v = - option_map (map_decl (subst s (#|Γ''| - S v))) (nth_error (Γ ,,, Γ' ,,, Γ'') v). -Proof. - simpl. intros Hs Hv. - rewrite !nth_error_app_lt; autorewrite with wf; f_equal; try lia. - erewrite nth_error_subst_context. f_equal. unfold subst_decl. rewrite Nat.add_0_r. reflexivity. -Qed. - -Lemma subst_iota_red s k pars c args brs : - subst s k (iota_red pars c args brs) = - iota_red pars c (List.map (subst s k) args) (List.map (on_snd (subst s k)) brs). -Proof. - unfold iota_red. rewrite !subst_mkApps. f_equal; auto using map_skipn. - rewrite nth_map; simpl; auto. -Qed. - -Lemma subst_unfold_fix n k mfix idx narg fn : - All Ast.wf n -> - unfold_fix mfix idx = Some (narg, fn) -> - unfold_fix (map (map_def (subst n k) (subst n (#|mfix| + k))) mfix) idx = Some (narg, subst n k fn). -Proof. - unfold unfold_fix. intros wfn. - rewrite nth_error_map. destruct (nth_error mfix idx) eqn:Hdef; try congruence. - case_eq (isLambda (dbody d)); [|discriminate]. - move=> Hlam [= <- <-] /=. rewrite isLambda_subst; tas. - f_equal. f_equal. - solve_all. - erewrite (distr_subst_rec _ _ _ wfn k 0). - rewrite fix_subst_length. simpl. f_equal. - unfold fix_subst. rewrite !map_length. - generalize #|mfix| at 2 3. induction n0; auto. simpl. - f_equal. apply IHn0. -Qed. -Hint Resolve subst_unfold_fix : core. - -Lemma subst_unfold_cofix n k mfix idx narg fn : - All Ast.wf n -> - unfold_cofix mfix idx = Some (narg, fn) -> - unfold_cofix (map (map_def (subst n k) (subst n (#|mfix| + k))) mfix) idx = Some (narg, subst n k fn). -Proof. - unfold unfold_cofix. intros wfn. - rewrite nth_error_map. destruct (nth_error mfix idx) eqn:Hdef; try congruence. - intros [= <- <-]. simpl. do 2 f_equal. solve_all. - erewrite (distr_subst_rec _ _ _ wfn k 0). - rewrite cofix_subst_length. simpl. f_equal. - unfold cofix_subst. rewrite !map_length. - generalize #|mfix| at 2 3. induction n0; auto. simpl. - f_equal. apply IHn0. -Qed. -Hint Resolve subst_unfold_cofix : core. - -Lemma subst_is_constructor: - forall (args : list term) (narg : nat) n k, - is_constructor narg args = true -> is_constructor narg (map (subst n k) args) = true. -Proof. - intros args narg. - unfold is_constructor; intros. - rewrite nth_error_map. destruct nth_error; try discriminate. simpl. intros. - destruct t; try discriminate || reflexivity. - destruct t; try discriminate || reflexivity. simpl. - destruct args0; auto. -Qed. -Hint Resolve subst_is_constructor : core. -Hint Constructors All_local_env : core. - -Lemma typed_subst `{checker_flags} Σ Γ t T n k : - wf Σ.1 -> k >= #|Γ| -> - Σ ;;; Γ |- t : T -> subst n k T = T /\ subst n k t = t. -Proof. - intros wfΣ Hk Hty. - pose proof (typing_wf_local Hty). - pose proof (typing_wf _ wfΣ _ _ _ Hty) as [wft wfT]. - apply typecheck_closed in Hty; eauto. - destruct Hty as [_ Hcl]. - rewrite -> andb_and in Hcl. destruct Hcl as [clb clty]. - pose proof (closed_upwards k clb). - simpl in *. forward H0 by lia. - pose proof (closed_upwards k clty). - simpl in *. forward H1 by lia. - apply (subst_closedn n) in H0; apply (subst_closedn n) in H1; auto. -Qed. - -Lemma subst_wf_local `{checker_flags} Σ Γ n k : - wf Σ.1 -> - wf_local Σ Γ -> - subst_context n k Γ = Γ. -Proof. - intros wfΣ. - induction 1; auto; unfold subst_context, snoc; rewrite fold_context_snoc0; - auto; unfold snoc; - f_equal; auto; unfold map_decl; simpl. - - destruct t0 as [s Hs]. unfold vass. simpl. f_equal. - eapply typed_subst; eauto. lia. - - unfold vdef. - f_equal. - + f_equal. eapply typed_subst; eauto. lia. - + eapply typed_subst in t1 as [Ht HT]; eauto. lia. -Qed. - -(* TODO MOVE *) -Lemma wf_Forall_decls_typing_wf `{checker_flags} : - forall Σ, - wf Σ -> - Forall_decls_typing (fun (_ : global_env_ext) (_ : context) (t T : term) => - Ast.wf t /\ Ast.wf T - ) Σ. -Proof. - intros Σ h. - apply typing_wf_sigma in h as ?. - eapply on_global_env_impl. 2: eassumption. - intros Σ' Γ t [T|] h1 h2. all: cbn in *. - - hnf in h2. assumption. - - hnf in h2. exists fresh_universe. intuition eauto. - constructor. -Qed. - -Lemma subst_declared_constant `{H:checker_flags} Σ cst decl n k u : - wf Σ -> - declared_constant Σ cst decl -> - map_constant_body (subst n k) (map_constant_body (subst_instance_constr u) decl) = - map_constant_body (subst_instance_constr u) decl. -Proof. - intros wΣ h. - apply declared_constant_wf in h as w. - 2: eapply wf_Forall_decls_typing_wf. 2: assumption. - destruct w as [wty wbo]. - eapply declared_decl_closed in h ; eauto. - unfold map_constant_body. - do 2 red in h. destruct decl as [ty [body|] univs]. all: simpl in *. - - rewrite -> andb_and in h. destruct h as [h1 h2]. - rewrite <- (closedn_subst_instance_constr 0 body u) in h1. - rewrite <- (closedn_subst_instance_constr 0 ty u) in h2. - f_equal. - + apply subst_closedn. - * apply wf_subst_instance_constr. assumption. - * eauto using closed_upwards with arith wf. - + f_equal. apply subst_closedn. - * apply wf_subst_instance_constr. assumption. - * eauto using closed_upwards with arith wf. - - red in h. f_equal. - simpl in *. - rewrite <- (closedn_subst_instance_constr 0 ty u) in h. - rewrite andb_true_r in h. - eapply subst_closedn. - + apply wf_subst_instance_constr. assumption. - + eauto using closed_upwards with arith wf. -Qed. - -Definition subst_mutual_inductive_body n k m := - map_mutual_inductive_body (fun k' => subst n (k' + k)) m. - -(* Damnit we also need Ast.wf arguments in Template-coq... to use closedness *) -Lemma subst_declared_minductive {cf:checker_flags} Σ cst decl n k : - wf Σ -> - declared_minductive Σ cst decl -> - subst_mutual_inductive_body n k decl = decl. -Proof. - intros wfΣ Hdecl. - pose proof (on_declared_minductive wfΣ Hdecl). apply onNpars in X. - unshelve epose proof (declared_inductive_closed (Σ:=(empty_ext Σ)) _ Hdecl). auto. - unshelve epose proof (declared_minductive_wf (Σ:=(empty_ext Σ)) _ Hdecl) as [wfpars wfbodies]. auto. - move: H. - rewrite /closed_inductive_decl /lift_mutual_inductive_body. - destruct decl; simpl in *. - move/andP => [clpar clbodies]. f_equal. - rewrite [fold_context _ _]closed_ctx_subst; try easy. - eapply Forall_All; eauto. - eapply forallb_All in clbodies. apply Forall_All in wfbodies. - eapply (All_mix wfbodies) in clbodies. clear wfbodies. - eapply Alli_mapi_id. - eapply (All_Alli clbodies). intros i oib [wfb clb]; eauto. - 2:{ intros. eapply X0. } - simpl. - move: clb. rewrite /closed_inductive_body. - destruct oib; simpl. move=> /andP[/andP [ci ct] cp]. - destruct wfb; simpl in *. - f_equal. rewrite subst_closedn; eauto. - eapply closed_upwards; eauto; lia. - eapply All_map_id. eapply forallb_All in ct. - eapply Forall_All in wf_ind_ctors. - eapply (All_mix wf_ind_ctors) in ct. - eapply (All_impl ct). intros x [wf cl]. - destruct x as [[id ty] arg]; unfold on_pi2; simpl; repeat f_equal. - apply subst_closedn. apply wf. unfold cdecl_type in cl; simpl in cl. - eapply closed_upwards; eauto; lia. - simpl in X. rewrite -X in cp. - eapply forallb_All in cp. - eapply Forall_All in wf_ind_projs. apply (All_mix wf_ind_projs) in cp. - eapply All_map_id; eauto. - eapply (All_impl cp); firstorder auto. - destruct x; unfold on_snd; simpl; f_equal. - apply subst_closedn; auto. rewrite context_assumptions_fold. - eapply closed_upwards; eauto; lia. -Qed. - -Lemma subst_declared_inductive `{checker_flags} Σ ind mdecl idecl n k : - wf Σ -> - declared_inductive Σ mdecl ind idecl -> - map_one_inductive_body (context_assumptions mdecl.(ind_params)) - (length (arities_context mdecl.(ind_bodies))) - (fun k' => subst n (k' + k)) (inductive_ind ind) idecl = idecl. -Proof. - unfold declared_inductive. intros wfΣ [Hmdecl Hidecl]. - eapply (subst_declared_minductive _ _ _ n k) in Hmdecl. - unfold subst_mutual_inductive_body in Hmdecl. - destruct mdecl. simpl in *. - injection Hmdecl. intros Heq. - clear Hmdecl. - pose proof Hidecl as Hidecl'. - rewrite <- Heq in Hidecl'. - rewrite nth_error_mapi in Hidecl'. - clear Heq. - unfold option_map in Hidecl'. rewrite Hidecl in Hidecl'. - congruence. auto. -Qed. - -Lemma subst0_inds_subst ind u mdecl n k t : - All Ast.wf n -> - (subst0 (inds (inductive_mind ind) u (ind_bodies mdecl)) - (subst n (#|arities_context (ind_bodies mdecl)| + k) t)) = - subst n k (subst0 (inds (inductive_mind ind) u (ind_bodies mdecl)) t). -Proof. - intros wfn. - pose proof (distr_subst_rec t (inds (inductive_mind ind) u (ind_bodies mdecl)) n wfn k 0). - simpl in H. rewrite H. - unfold arities_context. rewrite rev_map_length inds_length. - f_equal. generalize (ind_bodies mdecl). - clear. intros. - induction l; unfold inds; simpl; auto. f_equal. auto. -Qed. - -Lemma subst_declared_constructor `{cf:checker_flags} Σ c u mdecl idecl cdecl n k : - wf Σ -> All Ast.wf n -> - declared_constructor Σ mdecl idecl c cdecl -> - subst (map (UnivSubst.subst_instance_constr u) n) k (type_of_constructor mdecl cdecl c u) = (type_of_constructor mdecl cdecl c u). -Proof. - unfold declared_constructor. destruct c as [i ci]. - intros wfΣ wfn [Hidecl Hcdecl]. - eapply (subst_declared_inductive _ _ _ _ n k) in Hidecl; eauto. - unfold type_of_constructor. destruct cdecl as [[id t'] arity]. - destruct idecl; simpl in *. - injection Hidecl. intros. - pose Hcdecl as Hcdecl'. - rewrite <- H0 in Hcdecl'. - rewrite nth_error_map in Hcdecl'. rewrite Hcdecl in Hcdecl'. - simpl in Hcdecl'. injection Hcdecl'. - intros. - rewrite <- H2 at 2. - rewrite <- subst0_inds_subst. f_equal. - now rewrite <- subst_subst_instance_constr. - apply All_map. eapply All_impl; eauto. - intros. now apply wf_subst_instance_constr. -Qed. - -(* TODO MOVE *) -Lemma Forall_impl' : - forall A (P Q : A -> Prop) l, - @Forall A (fun x => P x -> Q x) l -> - Forall P l -> - Forall Q l. -Proof. - intros A P Q l h1 h2. - induction h2 in h1 |- *. - - constructor. - - inversion h1. subst. constructor. - + eauto. - + eauto. -Qed. - -Lemma subst_declared_projection {cf:checker_flags} Σ c mdecl idecl pdecl n k : - wf Σ -> - declared_projection Σ mdecl idecl c pdecl -> - on_snd (subst n (S (ind_npars mdecl + k))) pdecl = pdecl. -Proof. - intros wfΣ Hd. - eapply declared_projection_wf in Hd as w. - 2: eapply wf_Forall_decls_typing_wf. 2: assumption. - eapply (declared_projection_closed (Σ:=empty_ext Σ)) in Hd. 2:auto. - destruct pdecl; unfold on_snd; simpl; f_equal. apply subst_closedn; auto. - eapply closed_upwards; eauto. lia. -Qed. - -Lemma subst_fix_context: - forall (mfix : list (def term)) n (k : nat), - fix_context (map (map_def (subst n k) (subst n (#|mfix| + k))) mfix) = - subst_context n k (fix_context mfix). -Proof. - intros mfix n k. unfold fix_context. - rewrite map_vass_map_def rev_mapi. - fold (fix_context mfix). - rewrite (subst_context_alt n k (fix_context mfix)). - now rewrite /subst_decl mapi_length fix_context_length. -Qed. - -Lemma subst_destArity ctx t n k : - Ast.wf t -> - match destArity ctx t with - | Some (args, s) => - destArity (subst_context n k ctx) (subst n (#|ctx| + k) t) = Some (subst_context n k args, s) - | None => True - end. -Proof. - intros wf; revert ctx. - induction wf in n, k |- * using term_wf_forall_list_ind; intros ctx; simpl; trivial. - - - specialize (IHwf0 n k (ctx,, vass n0 t)). unfold snoc in IHwf0; rewrite subst_context_snoc in IHwf0. - simpl in IHwf0. unfold subst_decl, map_decl in IHwf0. unfold vass in *. simpl in IHwf0. - destruct destArity. destruct p. simpl in *. auto. exact I. - - specialize (IHwf1 n k (ctx,, vdef n0 t t0)). - unfold snoc in IHwf1; rewrite subst_context_snoc in IHwf1. - unfold vdef, subst_decl, map_decl in *. simpl in *. - destruct destArity as [[args s]|]. apply IHwf1. exact I. -Qed. - -Lemma decompose_prod_n_assum0 ctx t : decompose_prod_n_assum ctx 0 t = Some (ctx, t). -Proof. destruct t; simpl; reflexivity. Qed. - -Definition strip_outer_cast_tProd_tLetIn_morph f := - forall t k, - match strip_outer_cast t with - | tProd na A B => - strip_outer_cast (f k t) = tProd na (f k A) (f (S k) B) - | tLetIn na b ty b' => - strip_outer_cast (f k t) = tLetIn na (f k b) (f k ty) (f (S k) b') - | _ => True - end. - -Lemma strip_outer_cast_tProd_tLetIn_subst n : - strip_outer_cast_tProd_tLetIn_morph (subst n). -Proof. - unfold strip_outer_cast_tProd_tLetIn_morph. intros t k. - induction t; simpl in *; auto. -Qed. - -Lemma strip_outer_cast_subst_instance_constr u t : - strip_outer_cast (subst_instance_constr u t) = - subst_instance_constr u (strip_outer_cast t). -Proof. induction t; simpl; auto. Qed. - -Lemma strip_outer_cast_tProd_tLetIn_subst_instance_constr u : - strip_outer_cast_tProd_tLetIn_morph (fun _ => subst_instance_constr u). -Proof. - red. intros. - destruct (strip_outer_cast t) eqn:Heq; try auto. - rewrite strip_outer_cast_subst_instance_constr. now rewrite Heq. - rewrite strip_outer_cast_subst_instance_constr. now rewrite Heq. -Qed. - -Lemma subst_instantiate_params_subst n k params args s t : - All Ast.wf n -> forall s' t', - instantiate_params_subst params args s t = Some (s', t') -> - instantiate_params_subst - (mapi_rec (fun k' decl => subst_decl n (k' + k) decl) params #|s|) - (map (subst n k) args) (map (subst n k) s) (subst n (#|s| + k) t) = - Some (map (subst n k) s', subst n (#|s| + k + #|params|) t'). -Proof. - intros Hn. - induction params in args, t, n, Hn, k, s |- *; intros ctx' t'. - - destruct args; simpl; rewrite ?Nat.add_0_r; try congruence. - - simpl. - destruct a as [na [body|] ty]; simpl; try congruence; - destruct t; try congruence. - -- intros Ht'. - specialize (IHparams _ k _ _ _ Hn _ _ Ht'). - simpl in IHparams. - replace (#|s| + k + S #|params|) with (S (#|s| + k + #|params|)) by lia. - rewrite <- IHparams. f_equal. - now rewrite distr_subst. - -- intros Ht'. destruct args; try congruence. simpl. - specialize (IHparams _ k _ _ _ Hn _ _ Ht'). - simpl in IHparams. - replace (#|s| + k + S #|params|) with (S (#|s| + k + #|params|)) by lia. - now rewrite <- IHparams. -Qed. - -Lemma closed_tele_subst n k ctx : - All wf_decl ctx -> closed_ctx ctx -> - mapi (fun (k' : nat) (decl : context_decl) => subst_decl n (k' + k) decl) (List.rev ctx) = List.rev ctx. -Proof. - rewrite /closedn_ctx /mapi. simpl. generalize 0. - induction ctx using rev_ind; try easy. - move=> n0. - rewrite /closedn_ctx !rev_app_distr /id /=. - intro Hwf. move /andP => [closedx Hctx]. - apply All_app in Hwf as [? X]; inv X. - rewrite subst_decl_closed //. - rewrite (closed_decl_upwards n0) //; lia. - f_equal. now rewrite IHctx. -Qed. - -Lemma decompose_prod_n_assum_extend_ctx {ctx n t ctx' t'} ctx'' : - decompose_prod_n_assum ctx n t = Some (ctx', t') -> - decompose_prod_n_assum (ctx ++ ctx'') n t = Some (ctx' ++ ctx'', t'). -Proof. - induction n in ctx, t, ctx', t', ctx'' |- *. simpl. intros [= -> ->]. eauto. - simpl. - destruct (strip_outer_cast t); simpl; try congruence. - intros H. eapply (IHn _ _ _ _ ctx'' H). - intros H. eapply (IHn _ _ _ _ ctx'' H). -Qed. - -Lemma subst_it_mkProd_or_LetIn n k ctx t : - subst n k (it_mkProd_or_LetIn ctx t) = - it_mkProd_or_LetIn (subst_context n k ctx) (subst n (length ctx + k) t). -Proof. - induction ctx in n, k, t |- *; simpl; try congruence. - pose (subst_context_snoc n k ctx a). unfold snoc in e. rewrite e. clear e. - simpl. rewrite -> IHctx. - pose (subst_context_snoc n k ctx a). simpl. now destruct a as [na [b|] ty]. -Qed. - -Lemma to_extended_list_k_subst n k c k' : - to_extended_list_k (subst_context n k c) k' = to_extended_list_k c k'. -Proof. - unfold to_extended_list_k. revert k'. - generalize (@nil term) at 1 2. - induction c in n, k |- *; simpl; intros. reflexivity. - rewrite subst_context_snoc. unfold snoc. simpl. - destruct a. destruct decl_body. unfold subst_decl, map_decl. simpl. - now rewrite IHc. simpl. apply IHc. -Qed. - -Lemma to_extended_list_k_map_subst: - forall n (k : nat) (c : context) k', - #|c| + k' <= k -> - to_extended_list_k c k' = map (subst n k) (to_extended_list_k c k'). -Proof. - intros n k c k'. - pose proof (to_extended_list_k_spec c k'). - symmetry. solve_all. - destruct H as [x' [-> Hx']]. intuition. simpl. - destruct (leb_spec_Set k x'). lia. reflexivity. -Qed. - -Lemma subst_instantiate_params n k params args t ty : - All wf_decl params -> All Ast.wf n -> All Ast.wf args -> Ast.wf t -> - closed_ctx params -> - instantiate_params params args t = Some ty -> - instantiate_params params (map (subst n k) args) (subst n k t) = Some (subst n k ty). -Proof. - unfold instantiate_params. - move=> wfparams wfn wfargs wft. - move/(closed_tele_subst n k params)=> Heq. - rewrite -{2}Heq //. - specialize (subst_instantiate_params_subst n k (List.rev params) args [] t wfn). - move=> /= Heq'. - case E: (instantiate_params_subst (List.rev params) args)=> [[l' t']|] /= //. - specialize (Heq' _ _ E). rewrite Heq'. move=> [= <-]. f_equal. - rewrite distr_subst //. - move/instantiate_params_subst_length: E => -> /=. do 2 f_equal. lia. -Qed. -Hint Rewrite subst_instantiate_params : lift. - -Lemma wf_arities_context `{checker_flags} Σ mind mdecl : wf Σ -> - declared_minductive Σ mind mdecl -> wf_local (Σ, ind_universes mdecl) (arities_context mdecl.(ind_bodies)). -Proof. - intros wfΣ Hdecl. - eapply declared_minductive_inv in Hdecl. 2:apply weaken_env_prop_typing. all:eauto. - apply onInductives in Hdecl. - unfold arities_context. - revert Hdecl. - induction (ind_bodies mdecl) using rev_ind. constructor. - intros Ha. - rewrite rev_map_app. - simpl. apply Alli_app in Ha as [Hl Hx]. - inv Hx. clear X0. - apply onArity in X as [s Hs]. - specialize (IHl Hl). - econstructor; eauto. - fold (arities_context l) in *. - unshelve epose proof (weakening (Σ, ind_universes mdecl) [] (arities_context l) _ _ wfΣ _ Hs). - now rewrite app_context_nil_l. - simpl in X. - eapply (env_prop_typing _ typecheck_closed) in Hs; eauto. - rewrite -> andb_and in Hs. destruct Hs as [Hs Ht]. - simpl in Hs. apply (lift_closed #|arities_context l|) in Hs. - rewrite -> Hs, app_context_nil_l in X. simpl. exists s. - apply X. -Qed. - -Lemma on_constructor_closed_wf `{checker_flags} {Σ mind mdecl u idecl indices cdecl cs} : - wf Σ -> - on_constructor (lift_typing typing) (Σ, ind_universes mdecl) mdecl (inductive_ind mind) idecl indices cdecl cs -> - let cty := subst0 (inds (inductive_mind mind) u (ind_bodies mdecl)) - (subst_instance_constr u (snd (fst cdecl))) - in closed cty /\ Ast.wf cty. -Proof. - intros wfΣ [cs' ? Hparams [s Hs]]. - pose proof (typing_wf_local Hs). - apply typing_wf in Hs as w. 2: assumption. - destruct w as [w _]. - destruct cdecl as [[id cty] car]. - eapply (env_prop_typing _ typecheck_closed) in Hs; eauto. - rewrite arities_context_length in Hs. - rewrite -> andb_and in *. simpl in *. - destruct Hs as [Hs _]. split. - - eapply (closedn_subst _ 0 0). - + clear. unfold inds. - induction #|ind_bodies mdecl|; simpl; try constructor; auto. - + simpl. now rewrite -> inds_length, closedn_subst_instance_constr. - - apply wf_subst. - + auto with wf. - + apply wf_subst_instance_constr. assumption. -Qed. -(* -Lemma on_projection_closed_wf `{checker_flags} {Σ mind mdecl u i idecl pdecl} : - wf Σ -> mdecl.(ind_npars) = context_assumptions mdecl.(ind_params) -> - on_projection (lift_typing typing) (Σ, ind_universes mdecl) (inductive_mind mind) mdecl (inductive_ind mind) idecl i pdecl -> - let pty := subst_instance_constr u (snd pdecl) in - closedn (S (ind_npars mdecl)) pty = true /\ Ast.wf pty. -Proof. - intros wfΣ Hpar. - unfold on_projection. - intros [s Hs]. - pose proof (typing_wf_local Hs). - destruct pdecl as [id cty]. - eapply (env_prop_typing _ typing_wf_gen) in Hs as Hs'; eauto. - eapply (env_prop_typing _ typecheck_closed) in Hs; eauto. - rewrite -> andb_and in *. simpl in *. - destruct Hs as [Hs _]. - rewrite smash_context_length in Hs. simpl in *. - rewrite - Hpar in Hs. rewrite -> closedn_subst_instance_constr. - split; auto with wf. now apply wf_subst_instance_constr. -Qed. -*) -Lemma context_subst_length Γ a s : context_subst Γ a s -> #|Γ| = #|s|. -Proof. induction 1; simpl; congruence. Qed. - -Lemma context_subst_assumptions_length Γ a s : context_subst Γ a s -> context_assumptions Γ = #|a|. -Proof. induction 1; simpl; try congruence. rewrite app_length /=. lia. Qed. - -(* Lemma context_subst_app `{cf:checker_flags} Σ Γ Γ' a s : *) -(* All_local_env (fun _ _ t T => Ast.wf t /\ Ast.wf T) Σ Γ' -> *) -(* All Ast.wf a -> All Ast.wf s -> *) -(* context_subst (Γ' ++ Γ) a s -> *) -(* { a0 & { a1 & { s0 & { s1 & (context_subst Γ a0 s0 * context_subst (subst_context s0 0 Γ') a1 s1 *) -(* * (a = a0 ++ a1) * (s = s1 ++ s0))%type } } } }. *) -(* Proof. *) -(* induction Γ' in Γ, a, s |- *. simpl. *) -(* exists a, [], s, []. rewrite app_nil_r; intuition. *) - -(* simpl. intros wfΓ wfa wfs Hs. *) -(* revert wfΓ wfa wfs. inv Hs; intros wfΓ wfa wfs; inv wfΓ. *) -(* - eapply All_app in wfa as [Hargs Ha1]; *) -(* eapply (All_app (l:=[a1])) in wfs as [Ha'' Hs0]. *) -(* specialize (IHΓ' _ _ _ H0 Hargs Hs0 H). *) -(* destruct IHΓ' as (a0' & a1' & s1 & s2 & ((sa0 & sa1) & eqargs) & eqs0). *) -(* subst. exists a0', (a1' ++ [a1]), s1, (a1 :: s2). intuition eauto. *) -(* rewrite subst_context_snoc. constructor. auto. now rewrite app_assoc. *) -(* - eapply (All_app (l:=[subst0 s0 b])) in wfs as [Ha'' Hs0]. *) -(* specialize (IHΓ' _ _ _ H0 wfa Hs0 H). *) -(* destruct IHΓ' as (a0' & a1' & s1 & s2 & ((sa0 & sa1) & eqargs) & eqs0). *) -(* subst. exists a0', a1', s1, (subst s2 0 (subst s1 #|Γ'| b) :: s2). intuition eauto. *) -(* rewrite -> subst_context_snoc, Nat.add_0_r. *) -(* unfold subst_decl; simpl. unfold map_decl. simpl. *) -(* econstructor. auto. simpl. f_equal. *) -(* rewrite -> subst_app_simpl; auto. simpl. *) -(* pose proof(context_subst_length _ _ _ sa1) as Hs1. *) -(* rewrite subst_context_length in Hs1. rewrite -> Hs1. auto. auto. *) -(* apply All_app in Hs0. solve_all. *) -(* apply All_app in Hs0. solve_all. *) -(* Qed. *) - -Lemma make_context_subst_rec_spec ctx args s tele args' s' : - context_subst ctx args s -> - make_context_subst tele args' s = Some s' -> - context_subst (List.rev tele ++ ctx) (args ++ args') s'. -Proof. - induction tele in ctx, args, s, args', s' |- *. - - move=> /= Hc. case: args'. move => [= <-]. - now rewrite app_nil_r. - move=> a l //. - - move=> Hc /=. case: a => [na [body|] ty] /=. - -- specialize (IHtele (vdef na body ty :: ctx) args (subst0 s body :: s) args' s'). - move=> /=. rewrite <- app_assoc. - move/(IHtele _). move=> H /=. apply H. - constructor. auto. - -- case: args' => [|a args']; try congruence. - specialize (IHtele (vass na ty :: ctx) (args ++ [a]) (a :: s) args' s'). - move=> /=. rewrite <- app_assoc. - move/(IHtele _). move=> H /=. simpl in H. rewrite <- app_assoc in H. apply H. - constructor. auto. -Qed. - -Lemma make_context_subst_spec tele args s' : - make_context_subst tele args [] = Some s' -> - context_subst (List.rev tele) args s'. -Proof. - move/(make_context_subst_rec_spec [] [] [] _ _ _ context_subst_nil). - rewrite app_nil_r /= //. -Qed. - -Lemma instantiate_params_subst_make_context_subst ctx args s ty s' ty' : - instantiate_params_subst ctx args s ty = Some (s', ty') -> - exists ctx'', - make_context_subst ctx args s = Some s' /\ - decompose_prod_n_assum [] (length ctx) ty = Some (ctx'', ty'). -Proof. - induction ctx in args, s, ty, s' |- *; simpl. - case: args => [|a args'] // [= <- <-]. exists []; intuition congruence. - case: a => [na [body|] ty''] /=. - - destruct ty; try congruence. - intros. move: (IHctx _ _ _ _ H) => [ctx'' [Hmake Hdecomp]]. - eapply (decompose_prod_n_assum_extend_ctx [vdef na0 ty1 ty2]) in Hdecomp. - unfold snoc. eexists; intuition eauto. - - destruct ty; try congruence. - case: args => [|a args']; try congruence. - move=> H. move: (IHctx _ _ _ _ H) => [ctx'' [Hmake Hdecomp]]. - eapply (decompose_prod_n_assum_extend_ctx [vass na0 ty1]) in Hdecomp. - unfold snoc. eexists; intuition eauto. -Qed. - -Lemma instantiate_params_make_context_subst ctx args ty ty' : - instantiate_params ctx args ty = Some ty' -> - exists ctx' ty'' s', - decompose_prod_n_assum [] (length ctx) ty = Some (ctx', ty'') /\ - make_context_subst (List.rev ctx) args [] = Some s' /\ ty' = subst0 s' ty''. -Proof. - unfold instantiate_params. - case E: instantiate_params_subst => // [[s ty'']]. - move=> [= <-]. - eapply instantiate_params_subst_make_context_subst in E. - destruct E as [ctx'' [Hs Hty'']]. - exists ctx'', ty'', s. split; auto. - now rewrite -> List.rev_length in Hty''. -Qed. - -Lemma subst_cstr_concl_head ind u mdecl (arity : context) args : - let head := tRel (#|ind_bodies mdecl| - S (inductive_ind ind) + #|ind_params mdecl| + #|arity|) in - let s := (inds (inductive_mind ind) u (ind_bodies mdecl)) in - inductive_ind ind < #|ind_bodies mdecl| -> - subst s (#|arity| + #|ind_params mdecl|) - (subst_instance_constr u (mkApps head (to_extended_list_k (ind_params mdecl) #|arity| ++ args))) - = mkApps (tInd ind u) (to_extended_list_k (ind_params mdecl) #|arity| ++ - map (subst s (#|arity| + #|ind_params mdecl|)) (map (subst_instance_constr u) args)). -Proof. - intros. - rewrite subst_instance_constr_mkApps subst_mkApps. - f_equal. - - subst head. simpl subst_instance_constr. - rewrite (subst_rel_eq _ _ (#|ind_bodies mdecl| - S (inductive_ind ind)) (tInd ind u)) //; try lia. - subst s. rewrite inds_spec rev_mapi nth_error_mapi /=. - elim nth_error_spec. intros. simpl. - f_equal. destruct ind; simpl. f_equal. f_equal. simpl in H. lia. - rewrite List.rev_length. lia. - - rewrite !map_app. f_equal. - -- rewrite map_subst_instance_constr_to_extended_list_k. - erewrite to_extended_list_k_map_subst at 2. - now rewrite Nat.add_comm. lia. -Qed. - -Lemma subst_to_extended_list_k s k args ctx : - All Ast.wf s -> - make_context_subst (List.rev ctx) args [] = Some s -> - map (subst s k) (to_extended_list_k ctx k) = map (lift0 k) args. -Proof. - move=> wfs. - move/make_context_subst_spec. rewrite List.rev_involutive. - move=> H; induction H; simpl; auto. - - inv wfs. - rewrite map_app -IHcontext_subst //. - rewrite to_extended_list_k_cons /= !map_app; unf_term. - f_equal. - rewrite (lift_to_extended_list_k _ _ 1) map_map_compose. - pose proof (to_extended_list_k_spec Γ k); unf_term. - solve_all. destruct H2 as [n [-> Hn]]. - rewrite /lift (subst_app_decomp [a] s k); auto with wf. - rewrite subst_rel_gt. simpl; lia. - repeat (f_equal; simpl; try lia). - cbn -[subst]. f_equal. apply (subst_rel_eq _ _ 0 a); cbnr; lia. - - inv wfs. - rewrite -IHcontext_subst // to_extended_list_k_cons /=. - rewrite (lift_to_extended_list_k _ _ 1) map_map_compose. - pose proof (to_extended_list_k_spec Γ k). - solve_all. destruct H2 as [n [-> Hn]]. - rewrite /lift (subst_app_decomp [subst0 s b] s k); auto with wf. - rewrite subst_rel_gt. simpl; lia. - repeat (unf_term; f_equal; simpl; try lia). -Qed. - -Lemma wf_context_subst ctx args s : - All wf_decl ctx -> All Ast.wf args -> - context_subst ctx args s -> All Ast.wf s. -Proof. - move=> wfctx wfargs. - induction 1. constructor. - eapply All_app in wfargs as [wfargs wfa]. inv wfa. inv wfctx. - constructor; intuition auto. - inv wfctx. red in H. constructor; solve_all. apply wf_subst. solve_all. intuition auto with wf. -Qed. - -Lemma wf_make_context_subst ctx args s' : - All wf_decl ctx -> All Ast.wf args -> - make_context_subst (List.rev ctx) args [] = Some s' -> All Ast.wf s'. -Proof. - move=> wfctx wfargs. - move/make_context_subst_spec. rewrite rev_involutive. - eauto using wf_context_subst. -Qed. - -Hint Resolve All_Forall : wf. - -(* TODO Duplicate of PCUICSubstitution *) -Lemma Alli_map_option_out_mapi_Some_spec' {A B} (f g : nat -> A -> option B) h l t P : - All P l -> - (forall i x t, P x -> f i x = Some t -> g i x = Some (h t)) -> - map_option_out (mapi f l) = Some t -> - map_option_out (mapi g l) = Some (map h t). -Proof. - unfold mapi. generalize 0. - move => n Ha Hfg. move: t. - induction Ha in n |- *; try constructor; auto. - destruct t; cbnr; discriminate. - move=> t /=. case E: (f n x) => [b|]; try congruence. - rewrite (Hfg n _ _ p E). - case E' : map_option_out => [b'|]; try congruence. - move=> [= <-]. now rewrite (IHHa _ _ E'). -Qed. - -Lemma on_constructor_closed {cf:checker_flags} {Σ mind mdecl u idecl indices cdecl cs} : - wf Σ -> - on_constructor (lift_typing typing) (Σ, ind_universes mdecl) mdecl (inductive_ind mind) idecl indices cdecl cs -> - let cty := subst0 (inds (inductive_mind mind) u (ind_bodies mdecl)) - (subst_instance_constr u (snd (fst cdecl))) - in closed cty. -Proof. - intros wfΣ [cs' ? Hparams [s Hs] _]. - pose proof (typing_wf_local Hs). - destruct cdecl as [[id cty] car]. - eapply (env_prop_typing _ typecheck_closed) in Hs; eauto. - rewrite arities_context_length in Hs. - rewrite -> andb_and in *. simpl in *. - destruct Hs as [Hs _]. - eapply (closedn_subst _ 0 0). - clear. unfold inds. induction #|ind_bodies mdecl|; simpl; try constructor; auto. - simpl. now rewrite -> inds_length, closedn_subst_instance_constr. -Qed. - -(* TODO Duplicate of PCUICSubstitution? *) -Lemma map_subst_instance_constr_to_extended_list_k u ctx k : - to_extended_list_k (subst_instance_context u ctx) k - = to_extended_list_k ctx k. -Proof. - unfold to_extended_list_k. - cut (map (subst_instance_constr u) [] = []); [|reflexivity]. - generalize (@nil term); intros l Hl. - induction ctx in k, l, Hl |- *; cbnr. - destruct a as [? [] ?]; cbnr; eauto. apply IHctx. - simpl. now rewrite Hl. -Qed. - -Lemma subst_instance_context_assumptions u ctx : - context_assumptions (subst_instance_context u ctx) - = context_assumptions ctx. -Proof. - induction ctx; cbnr. - destruct (decl_body a); cbn; now rewrite IHctx. -Qed. - - -Lemma subst_build_case_predicate_type ind mdecl idecl u params ps pty n k : - closed_ctx (subst_instance_context u (ind_params mdecl)) -> - closed (ind_type idecl) -> - build_case_predicate_type ind mdecl idecl params u ps = Some pty -> - build_case_predicate_type ind mdecl - (map_one_inductive_body (context_assumptions mdecl.(ind_params)) - (length (arities_context (ind_bodies mdecl))) (fun k' => subst n (k' + k)) - (inductive_ind ind) idecl) - (map (subst n k) params) u ps - = Some (subst n k pty). -Proof. -(* intros closedparams closedtype. *) -(* unfold build_case_predicate_type; simpl. *) -(* case_eq (instantiate_params (subst_instance_context u (ind_params mdecl)) params *) -(* (subst_instance_constr u (ind_type idecl))); *) -(* [|discriminate ]. *) -(* intros ipars Hipars. *) -(* apply (subst_instantiate_params n k) in Hipars; tas. *) -(* rewrite ind_type_map. simpl. *) -(* rewrite subst_closedn in Hipars. *) -(* { rewrite closedn_subst_instance_constr; eapply closed_upwards; tea; lia. } *) -(* rewrite subst_closedn; [eapply closed_upwards; tea; lia|]. *) -(* rewrite Hipars. *) -(* specialize (subst_destArity [] ipars n k); *) -(* destruct (destArity [] ipars) as [[ctx s]|]; [|discriminate]. *) -(* simpl. cbn. move => -> /some_inj-HH. simpl. f_equal. *) -(* etransitivity. *) -(* 2: exact (f_equal (subst n k) HH). *) -(* rewrite subst_it_mkProd_or_LetIn. simpl. f_equal. f_equal. *) -(* { destruct idecl; reflexivity. } *) -(* rewrite subst_mkApps; simpl. f_equal. *) -(* rewrite map_app; f_equal. *) -(* - rewrite !map_map; apply map_ext; clear; intro. *) -(* rewrite subst_context_length. *) -(* rewrite commut_lift_subst_rec. lia. *) -(* f_equal. lia. *) -(* - rewrite /to_extended_list to_extended_list_k_subst. *) -(* rewrite <- to_extended_list_k_map_subst. reflexivity. lia. *) -(* Qed. *) -Admitted. - -Lemma subst_build_branches_type {cf:checker_flags} - n k Σ ind mdecl idecl indices args u p brs cs : - wf Σ -> - All Ast.wf args -> - Ast.wf (ind_type idecl) -> - All Ast.wf n -> - declared_inductive Σ mdecl ind idecl -> - closed_ctx (subst_instance_context u (ind_params mdecl)) -> - on_inductive (lift_typing typing) (Σ, ind_universes mdecl) - (inductive_mind ind) mdecl -> - on_constructors (lift_typing typing) (Σ, ind_universes mdecl) - mdecl (inductive_ind ind) idecl indices (ind_ctors idecl) cs -> - map_option_out (build_branches_type ind mdecl idecl args u p) = Some brs -> - map_option_out (build_branches_type ind mdecl - idecl (map (subst n k) args) u (subst n k p)) = - Some (map (on_snd (subst n k)) brs). -Proof. - intros wfΣ wargs widecl wn wfidecl closedparams onmind Honcs. - assert (wc : Forall wf_decl (subst_instance_context u (ind_params mdecl))). - { assert (h : Forall wf_decl (ind_params mdecl)). - { eapply typing_all_wf_decl ; revgoals. - - apply onParams in onmind. - unfold on_context in onmind. - change TemplateEnvironment.ind_params with ind_params in onmind. - eassumption. - - simpl. assumption. - } - clear - h. induction h. 1: constructor. - simpl. constructor. 2: assumption. - destruct x as [na [bo|] ty]. - all: unfold map_decl. all: unfold wf_decl in *. all: simpl in *. - all: intuition eauto with wf. - } - rewrite !build_branches_type_. cbn. - eapply Alli_map_option_out_mapi_Some_spec'. - eapply All2_All_left; tea. intros x y u'; exact (y; u'). - clear Honcs brs. - intros j [[id t] i] [t' k'] [cs' Honc]. - case_eq (instantiate_params (subst_instance_context u (ind_params mdecl)) args - (subst0 (inds (inductive_mind ind) u - (ind_bodies mdecl)) - (subst_instance_constr u t))); - [|discriminate]. - intros ty Heq; cbn. - pose proof (on_constructor_closed_wf wfΣ (u:=u) Honc) as [clt wt]. - cbn in clt. cbn in wt. - eapply (closed_upwards k) in clt; try lia. - remember (subst0 (inds (inductive_mind ind) u (ind_bodies mdecl)) - (subst_instance_constr u t)) as ty'. - apply (subst_instantiate_params n k) in Heq as Heq'. - all: auto using Forall_All with wf. - erewrite subst_closedn in Heq'; tas. - rewrite Heq'. - eapply instantiate_params_make_context_subst in Heq. - destruct Heq as [ctx' [ty'' [s' [Hty [Hsubst Ht0]]]]]. - subst ty; simpl. - rewrite Heqty' in Hty. - destruct Honc as [Hcshape_args ? cshape_eq Hty' _ _]; unfold cdecl_type in *; simpl in *. - rewrite cshape_eq in Hty. - rewrite -> !subst_instance_constr_it_mkProd_or_LetIn in Hty. - rewrite !subst_it_mkProd_or_LetIn in Hty. - assert (H0: #|subst_context (inds (inductive_mind ind) u (ind_bodies mdecl)) 0 - (subst_instance_context u (ind_params mdecl))| - = #|subst_instance_context u (ind_params mdecl)|). - { now rewrite subst_context_length. } - rewrite <- H0 in Hty. - rewrite decompose_prod_n_assum_it_mkProd in Hty. - injection Hty. clear Hty. - intros Heqty'' <-. revert Heqty''. - rewrite !subst_instance_context_length Nat.add_0_r. - rewrite subst_context_length subst_instance_context_length. - rewrite (subst_cstr_concl_head ind u mdecl cs'.(cshape_args) cs'.(cshape_indices)). - - destruct wfidecl as [Hmdecl Hnth]. - now apply nth_error_Some_length in Hnth. - - intros <-. rewrite !subst_it_mkProd_or_LetIn !subst_mkApps /=. - rewrite !decompose_prod_assum_it_mkProd /=; - try by rewrite is_ind_app_head_mkApps. - rewrite !decompose_app_mkApps; try by reflexivity. - simpl. rewrite !map_app !subst_context_length - !subst_instance_context_length Nat.add_0_r. - eapply subst_to_extended_list_k with (k:=#|cs'.(cshape_args)|) in Hsubst as XX. - 2:{ eapply wf_make_context_subst. 3: eassumption. - all: auto using Forall_All with wf. - } - rewrite map_subst_instance_constr_to_extended_list_k in XX. - rewrite !XX; clear XX. - apply make_context_subst_spec in Hsubst as Hsubst'. - rewrite rev_involutive in Hsubst'. - pose proof (context_subst_assumptions_length _ _ _ Hsubst') as H1. - case E: chop => [l l']. - have chopm := (chop_map _ _ _ _ _ E). - move: E chopm. - rewrite chop_n_app ?map_length. - { rewrite <- H1. apply onNpars in onmind. - now rewrite subst_instance_context_assumptions. - } - move=> [= <- <-] chopm. - move: {chopm}(chopm _ (subst n (#|cs'.(cshape_args)| + k))). - rewrite map_app. - move=> chopm; rewrite {}chopm /=. - inversion 1; subst. f_equal. unfold on_snd; cbn; f_equal. - rewrite !app_nil_r /on_snd /=. - rewrite subst_it_mkProd_or_LetIn !subst_context_length !subst_mkApps - !subst_instance_context_length !map_app. - f_equal. f_equal. - -- rewrite -> commut_lift_subst_rec. arith_congr. lia. - -- f_equal. simpl. f_equal. - rewrite !subst_mkApps /= !map_app. f_equal. - f_equal. - rewrite /to_extended_list -to_extended_list_k_map_subst. - rewrite !subst_context_length subst_instance_context_length. lia. - now rewrite to_extended_list_k_subst. -Qed. - - -Hint Unfold subst1 : subst. -Hint Rewrite subst_mkApps distr_subst: subst. - -Lemma subs_nth_error `{checker_flags} Σ Γ s Δ decl n t : - subs Σ Γ s Δ -> - nth_error Δ n = Some decl -> - nth_error s n = Some t -> - match decl_body decl return Type with - | Some t' => False - | None => - let ty := subst0 (skipn (S n) s) (decl_type decl) in - Σ ;;; Γ |- t : ty - end. -Proof. - induction 1 in n |- *; simpl; auto; destruct n; simpl; try congruence. - - intros [= <-]. intros [= ->]. - simpl. exact t1. - - apply IHX. -Qed. - -Lemma red1_tApp_mkApps_l Σ Γ M1 N1 M2 : - red1 Σ Γ M1 N1 -> red1 Σ Γ (tApp M1 M2) (mkApps N1 M2). -Proof. constructor. auto. Qed. - -Lemma red1_tApp_mkApp Σ Γ M1 N1 M2 : - red1 Σ Γ M1 N1 -> red1 Σ Γ (tApp M1 [M2]) (mkApp N1 M2). -Proof. - intros. - change (mkApp N1 M2) with (mkApps N1 [M2]). - apply app_red_l. auto. -Qed. - -Lemma red1_mkApp Σ Γ M1 N1 M2 : - Ast.wf M1 -> - red1 Σ Γ M1 N1 -> red1 Σ Γ (mkApp M1 M2) (mkApp N1 M2). -Proof. - intros wfM1 H. - destruct (isApp M1) eqn:Heq. - destruct M1; try discriminate. simpl. - revert wfM1. inv H. simpl. intros. - rewrite mkApp_mkApps. constructor. - - intros. - econstructor; eauto. - inv wfM1. simpl. - clear -H1. - unfold is_constructor in *. - destruct (nth_error args narg) eqn:Heq. - eapply nth_error_app_left in Heq. now rewrite -> Heq. discriminate. - - intros. rewrite mkApp_mkApps. now constructor. - - intros. simpl. - constructor. clear -X. induction X; constructor; auto. - - rewrite mkApp_tApp; auto. - now apply red1_tApp_mkApp. -Qed. - -Lemma red1_mkApps_l Σ Γ M1 N1 M2 : - Ast.wf M1 -> All Ast.wf M2 -> - red1 Σ Γ M1 N1 -> red1 Σ Γ (mkApps M1 M2) (mkApps N1 M2). -Proof. - induction M2 in M1, N1 |- *. simpl; auto. - intros. specialize (IHM2 (mkApp M1 a) (mkApp N1 a)). - inv X. - forward IHM2. apply wf_mkApp; auto. - forward IHM2. auto. - rewrite <- !mkApps_mkApp; auto. - apply IHM2. - apply red1_mkApp; auto. -Qed. - -Lemma red1_mkApps_r Σ Γ M1 M2 N2 : - Ast.wf M1 -> All Ast.wf M2 -> - OnOne2 (red1 Σ Γ) M2 N2 -> red1 Σ Γ (mkApps M1 M2) (mkApps M1 N2). -Proof. - intros. induction X0 in M1, H, X |- *. - inv X. - destruct (isApp M1) eqn:Heq. destruct M1; try discriminate. - simpl. constructor. apply OnOne2_app. constructor. auto. - rewrite mkApps_tApp; try congruence. - rewrite mkApps_tApp; try congruence. - constructor. constructor. auto. - inv X. - specialize (IHX0 (mkApp M1 hd)). forward IHX0. - apply wf_mkApp; auto. forward IHX0; auto. - now rewrite !mkApps_mkApp in IHX0. -Qed. - -Lemma wf_fix : - forall (mfix : list (def term)) (k : nat), Ast.wf (tFix mfix k) -> - Forall - (fun def : def term => Ast.wf (dtype def) /\ Ast.wf (dbody def) /\ isLambda (dbody def) = true) - mfix. -Proof. - intros. inv H. auto. -Defined. - -Lemma wf_cofix : - forall (mfix : list (def term)) (k : nat), Ast.wf (tCoFix mfix k) -> - Forall - (fun def : def term => Ast.wf (dtype def) /\ Ast.wf (dbody def)) - mfix. -Proof. - intros. inv H. auto. -Defined. - -Local Set Keyed Unification. - -Lemma substitution_red1 `{CF:checker_flags} Σ Γ Γ' Γ'' s M N : - wf Σ.1 -> All Ast.wf s -> subs Σ Γ s Γ' -> wf_local Σ Γ -> Ast.wf M -> - red1 (fst Σ) (Γ ,,, Γ' ,,, Γ'') M N -> - red1 (fst Σ) (Γ ,,, subst_context s 0 Γ'') (subst s #|Γ''| M) (subst s #|Γ''| N). -Proof. - intros wfΣ wfs Hs wfΓ wfM H. - remember (Γ ,,, Γ' ,,, Γ'') as Γ0. revert Γ Γ' Γ'' HeqΓ0 wfΓ Hs. - induction H using red1_ind_all in |- *; intros Γ0 Γ' Γ'' HeqΓ0 wfΓ Hs; try subst Γ; cbn -[iota_red]; - match goal with - |- context [iota_red _ _ _ _] => idtac - | |- _ => autorewrite with subst - end; auto; - try solve [ econstructor; ((forward IHred1; try inv wfM; eauto) || eauto) ]. - - (* remember (Γ ,,, Γ' ,,, Γ'') as Γ0. revert Γ Γ' Γ'' HeqΓ0 wfΓ Hs. *) - (* induction H using red1_ind_all in |- *; intros Γ0 Γ' Γ'' HeqΓ0 wfΓ Hs; try subst Γ; simpl; *) - (* autorewrite with subst; auto; *) - (* try solve [ econstructor; try eapply IHred1; try inv wfM; eauto; eauto ]. *) - (* - autorewrite with subst. rewrite distr_subst; auto. - eapply red_beta. - - - autorewrite with subst. rewrite distr_subst; auto. - eapply red_zeta. *) - - - pose proof (subst_length _ _ _ _ Hs). - elim (leb_spec_Set); intros Hn. - + destruct (nth_error s) eqn:Heq. - ++ pose proof (nth_error_Some_length Heq). - rewrite -> nth_error_app_context_ge in H by lia. - rewrite -> nth_error_app_context_lt in H by lia. - destruct nth_error eqn:HΓ'. - eapply subs_nth_error in Heq; eauto. - simpl in H. destruct decl_body; try contradiction. - discriminate. discriminate. - ++ apply nth_error_None in Heq. - assert(S i = #|s| + (S (i - #|s|))) by lia. - rewrite H1. rewrite -> simpl_subst; try lia. - econstructor. - rewrite nth_error_app_context_ge // in H. - rewrite nth_error_app_context_ge // in H. lia. - rewrite -> nth_error_app_context_ge. 2:(autorewrite with wf; lia). - rewrite <- H. f_equal. f_equal. autorewrite with wf. lia. - rewrite -> !nth_error_app_context_ge in H by lia. - destruct nth_error eqn:Hi. - eapply nth_error_All_local_env in wfΓ. red in wfΓ. - rewrite -> Hi in wfΓ. simpl in H. - destruct c; simpl in H; try discriminate. - injection H. intros ->. red in wfΓ. cbn in *. apply typing_wf in wfΓ. intuition eauto. eauto. - apply nth_error_Some_length in Hi. lia. discriminate. - + rewrite -> nth_error_app_context_lt in H by lia. - pose (commut_lift_subst_rec body s (S i) (#|Γ''| - S i) 0). - assert(eq:#|Γ''| - S i + S i = #|Γ''|) by lia. - rewrite -> eq in e. rewrite <- e by lia. - constructor. - rewrite -> nth_error_app_context_lt by (autorewrite with wf; lia). - rewrite -> nth_error_subst_context. - unfold subst_decl; now rewrite -> option_map_decl_body_map_decl, H, Nat.add_0_r. - - - rewrite subst_iota_red. - autorewrite with subst. - constructor. - - - rewrite mkApps_tApp; simpl; auto with wf. - inv wfM. auto with wf. rewrite -> mkApps_tApp; simpl. - + eapply red_fix. erewrite subst_unfold_fix; eauto. - now apply subst_is_constructor. - + inv wfM. inv H3. - unfold unfold_fix in H. - destruct nth_error eqn:Heq. - destruct (isLambda (dbody d)); [|discriminate]. - injection H. intros <- <-. - eapply nth_error_forall in H5; eauto. - destruct d as [na b ty]; simpl in *. - destruct H5 as [_ [_ Hty]]. - destruct ty; try discriminate. reflexivity. - discriminate. - + apply not_empty_map. now inv wfM. - - - pose proof (subst_declared_constant _ _ _ s #|Γ''| u wfΣ H). - apply (f_equal cst_body) in H1. - rewrite <- !map_cst_body in H1. rewrite H0 in H1. simpl in H1. - injection H1. intros ->. - econstructor. eauto. eauto. - - - simpl. constructor. - inv wfM. - now rewrite nth_error_map H. - - - constructor. - forward IHred1; try now inv wfM. - specialize (IHred1 Γ0 Γ' (Γ'' ,, _) eq_refl). - now rewrite subst_context_snoc0 in IHred1. - - - constructor. - forward IHred1; try now inv wfM. - specialize (IHred1 Γ0 Γ' (Γ'' ,, _) eq_refl). - now rewrite subst_context_snoc0 in IHred1. - - - constructor. - induction X. all: simpl. all: constructor. - + intuition auto. simpl. eapply b0. all: auto. - inversion wfM. subst. inversion H5. subst. auto. - + eapply IHX. inversion wfM. subst. - constructor. all: auto. - inversion H5. subst. assumption. - - - forward IHred1; try now inv wfM. - assert(All (Ast.wf ∘ subst s #|Γ''|) M2). - eapply (Forall_All (fun x => Ast.wf (subst s #|Γ''| x))). inv wfM; solve_all. - apply wf_subst; auto. solve_all. - specialize (IHred1 _ _ _ eq_refl). - specialize (IHred1 wfΓ Hs). - apply red1_mkApps_l; auto. inv wfM. apply wf_subst; auto. solve_all. - solve_all. - - - assert(Ast.wf (subst s #|Γ''| M1)). inv wfM. apply wf_subst; auto with wf. - assert(All (Ast.wf ∘ subst s #|Γ''|) M2). - { apply (Forall_All (fun x => Ast.wf (subst s #|Γ''| x))). - inv wfM. - eapply Forall_impl; eauto. simpl; eauto with wf. } - apply red1_mkApps_r; auto with wf. - now apply All_map. - assert(Ast.wf M1). now inv wfM. - assert(All Ast.wf M2). eapply Forall_All. now inv wfM. - clear -X H0 X1 wfΓ Hs. - induction X; constructor; auto. - intuition. - eapply b; eauto. now inv X1. - apply IHX. now inv X1. - - - forward IHred1. now inv wfM. - constructor. specialize (IHred1 _ _ (Γ'' ,, vass na M1) eq_refl). - now rewrite subst_context_snoc0 in IHred1. - - - constructor. - induction X; constructor; auto. - intuition. eapply b; eauto. - inv wfM. now inv H. - apply IHX. inv wfM. inv H; now constructor. - - - constructor. - rewrite -> (OnOne2_length X). generalize (#|mfix1|). - induction X. all: simpl. all: constructor. all: simpl. - + intuition eauto. - * eapply b0. all: eauto. - inv wfM. inv H. intuition eauto. - * inversion b. auto. - + eapply IHX. constructor. - apply wf_fix in wfM. inv wfM. assumption. - - - apply wf_fix in wfM. - apply fix_red_body. rewrite !subst_fix_context. - solve_all. apply (OnOne2_All_mix_left wfM) in X. (* clear wfM. *) - rewrite <- (OnOne2_length X). - eapply OnOne2_map. unfold on_Trel. solve_all. - + specialize (X Γ0 Γ' (Γ'' ,,, fix_context mfix0)). - rewrite app_context_assoc in X. specialize (X eq_refl). - rewrite -> app_context_length, fix_context_length in *. - rewrite -> subst_context_app in *. - rewrite -> app_context_assoc, Nat.add_0_r in *. - auto. - + inversion b0. auto. - - - constructor. - rewrite -> (OnOne2_length X). generalize (#|mfix1|). - induction X. all: simpl. all: constructor. all: simpl. - + intuition eauto. - * eapply b0. all: eauto. - inv wfM. inv H. intuition eauto. - * inversion b. auto. - + eapply IHX. constructor. - apply wf_cofix in wfM. inv wfM. assumption. - - - apply wf_cofix in wfM. - apply cofix_red_body. rewrite !subst_fix_context. - solve_all. apply (OnOne2_All_mix_left wfM) in X. (* clear wfM. *) - rewrite <- (OnOne2_length X). - eapply OnOne2_map. unfold on_Trel. solve_all. - + specialize (X Γ0 Γ' (Γ'' ,,, fix_context mfix0)). - rewrite app_context_assoc in X. specialize (X eq_refl). - rewrite -> app_context_length, fix_context_length in *. - rewrite -> subst_context_app in *. - rewrite -> app_context_assoc, Nat.add_0_r in *. - auto. - + inversion b0. auto. -Qed. - -Lemma eq_term_upto_univ_refl Re Rle : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - forall t, eq_term_upto_univ Re Rle t t. -Proof. - intros hRe hRle. - induction t in Rle, hRle |- * using term_forall_list_rect; simpl; - try constructor; try apply Forall_Forall2; try apply All_All2 ; try easy; - try now (try apply Forall_All ; apply Forall_True). - - eapply All_All2. 1: eassumption. - intros. easy. - - eapply All_All2. 1: eassumption. - intros. easy. - - destruct p. constructor; try easy. - eapply All_All2. 1: eassumption. - intros. split ; auto. - - eapply All_All2. 1: eassumption. - intros x [? ?]. repeat split ; auto. - - eapply All_All2. 1: eassumption. - intros x [? ?]. repeat split ; auto. -Qed. - -Lemma eq_term_refl `{checker_flags} φ t : eq_term φ t t. -Proof. - apply eq_term_upto_univ_refl. - - intro; apply eq_universe_refl. - - intro; apply eq_universe_refl. -Qed. - - -Lemma leq_term_refl `{checker_flags} φ t : leq_term φ t t. -Proof. - apply eq_term_upto_univ_refl. - - intro; apply eq_universe_refl. - - intro; apply leq_universe_refl. -Qed. - - -Lemma eq_term_leq_term `{checker_flags} φ t u : eq_term φ t u -> leq_term φ t u. -Proof. - induction t in u |- * using term_forall_list_rect; simpl; inversion 1; - subst; constructor; try (now unfold eq_term, leq_term in * ); - try eapply Forall2_impl' ; try eapply All2_impl' ; try easy. - now apply eq_universe_leq_universe. - all: try (apply Forall_True, eq_universe_leq_universe). -Qed. - -Lemma eq_term_upto_univ_App `{checker_flags} Re Rle f f' : - eq_term_upto_univ Re Rle f f' -> - isApp f = isApp f'. -Proof. - inversion 1; reflexivity. -Qed. - -Lemma eq_term_App `{checker_flags} φ f f' : - eq_term φ f f' -> - isApp f = isApp f'. -Proof. - inversion 1; reflexivity. -Qed. - -Lemma eq_term_upto_univ_mkApps `{checker_flags} Re Rle f l f' l' : - eq_term_upto_univ Re Rle f f' -> - All2 (eq_term_upto_univ Re Re) l l' -> - eq_term_upto_univ Re Rle (mkApps f l) (mkApps f' l'). -Proof. - induction l in f, f' |- *; intro e; inversion_clear 1. - - assumption. - - pose proof (eq_term_upto_univ_App _ _ _ _ e). - case_eq (isApp f). - + intro X; rewrite X in H0. - destruct f; try discriminate. - destruct f'; try discriminate. - cbn. inversion_clear e. constructor. assumption. - apply All2_app. assumption. - now constructor. - + intro X; rewrite X in H0. - rewrite !mkApps_tApp; eauto. - intro; discriminate. - intro; discriminate. - constructor. assumption. - now constructor. -Qed. - -Lemma leq_term_mkApps `{checker_flags} φ f l f' l' : - leq_term φ f f' -> - All2 (eq_term φ) l l' -> - leq_term φ (mkApps f l) (mkApps f' l'). -Proof. - eapply eq_term_upto_univ_mkApps. -Qed. - -Lemma leq_term_App `{checker_flags} φ f f' : - leq_term φ f f' -> - isApp f = isApp f'. -Proof. - inversion 1; reflexivity. -Qed. - -Lemma subst_eq_term_upto_univ `{checker_flags} Re Rle n k T U : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - eq_term_upto_univ Re Rle T U -> - eq_term_upto_univ Re Rle (subst n k T) (subst n k U). -Proof. - intros hRe hRle h. - induction T in n, k, U, h, Rle, hRle |- * using term_forall_list_rect. - all: simpl ; inversion h ; simpl. - all: try (eapply eq_term_upto_univ_refl ; easy). - all: try (constructor ; easy). - - constructor. - eapply All2_map. eapply All2_impl' ; try eassumption. - eapply All_impl ; try eassumption. - cbn. intros x HH y HH'. now eapply HH. - - subst. eapply eq_term_upto_univ_mkApps. all: try easy. - eapply All2_map. eapply All2_impl' ; try eassumption. - eapply All_impl ; try eassumption. - cbn. intros x HH y HH'. now eapply HH. - - constructor ; try easy. - eapply All2_map. - eapply All2_impl'. eassumption. - eapply All_impl. easy. - cbn. intros x HH y [? HH']. split ; [assumption | now apply HH]. - - constructor. - eapply All2_map. - eapply All2_impl'. eassumption. - eapply All_impl. eassumption. - cbn. apply All2_length in X0 as e. rewrite e. - intros x [] y []; now split. - - constructor. - eapply All2_map. - eapply All2_impl'. eassumption. - eapply All_impl. eassumption. - cbn. apply All2_length in X0 as e. rewrite e. - intros x [] y []; now split. -Qed. - -Lemma subst_eq_term `{checker_flags} ϕ n k T U : - eq_term ϕ T U -> - eq_term ϕ (subst n k T) (subst n k U). -Proof. - intros Hleq. - eapply subst_eq_term_upto_univ. - - intro. eapply eq_universe_refl. - - intro. eapply eq_universe_refl. - - assumption. -Qed. - -Lemma subst_leq_term `{checker_flags} ϕ n k T U : - leq_term ϕ T U -> - leq_term ϕ (subst n k T) (subst n k U). -Proof. - intros Hleq. - eapply subst_eq_term_upto_univ. - - intro. eapply eq_universe_refl. - - intro. eapply leq_universe_refl. - - assumption. -Qed. - - -Lemma subst_eq_decl `{checker_flags} ϕ l k d d' : - eq_decl ϕ d d' -> eq_decl ϕ (subst_decl l k d) (subst_decl l k d'). -Proof. - destruct d, d', decl_body, decl_body0; - unfold eq_decl, map_decl; cbn; intuition auto using subst_eq_term. -Qed. - -Lemma subst_eq_context `{checker_flags} φ l l' n k : - eq_context φ l l' -> - eq_context φ (subst_context n k l) (subst_context n k l'). -Proof. - induction l in l', n, k |- *; inversion 1. constructor. - rewrite !subst_context_snoc. constructor. - apply All2_length in X1 as e. rewrite e. - now apply subst_eq_decl. - now apply IHl. -Qed. - -Lemma subs_wf `{checker_flags} Σ Γ s Δ : wf Σ.1 -> subs Σ Γ s Δ -> All Ast.wf s. -Proof. - induction 2; constructor. - apply typing_wf in t0; intuition auto with wf. - eapply typing_wf_local in t0; eauto. -Qed. - -Lemma wf_red1_wf `{checker_flags} Σ Γ M N : wf Σ.1 -> wf_local Σ Γ -> Ast.wf M -> red1 (fst Σ) Γ M N -> Ast.wf N. -Proof. - intros wfΣ wfΓ wfM Hr. - apply wf_red1 in Hr; eauto. - apply typing_wf_sigma; auto. - apply typing_all_wf_decl in wfΓ; eauto. -Qed. - -(** The cumulativity relation is substitutive, yay! *) - -Lemma substitution_cumul `{CF:checker_flags} Σ Γ Γ' Γ'' s M N : - wf Σ.1 -> wf_local Σ (Γ ,,, Γ' ,,, Γ'') -> subs Σ Γ s Γ' -> Ast.wf M -> Ast.wf N -> - Σ ;;; Γ ,,, Γ' ,,, Γ'' |- M <= N -> - Σ ;;; Γ ,,, subst_context s 0 Γ'' |- subst s #|Γ''| M <= subst s #|Γ''| N. -Proof. - intros wfΣ wfΓ Hs wfM wfN. induction 1. - constructor. - - now apply subst_leq_term. - - pose proof r. - apply wf_red1_wf in X0; eauto. - eapply substitution_red1 in r. 4:eauto. all:auto. - econstructor 2; eauto. - eauto using subs_wf. - eauto with wf. - - pose proof r. - apply wf_red1_wf in X0; eauto. - eapply substitution_red1 in r. 4:eauto. - all:eauto using subs_wf with wf. - now econstructor 3. -Qed. - - -Lemma substitution_All_local_env_over `{cf : checker_flags} {Σ Γ Γ' Δ s} : - wf Σ.1 -> - forall wfΓ0 : wf_local Σ (Γ ,,, Γ' ,,, Δ), - All_local_env_over typing - (fun (Σ : global_env_ext) (Γ : context) (_ : wf_local Σ Γ) - (t T : term) (_ : Σ;;; Γ |- t : T) => - forall (Γ0 Γ' Δ : context) (s : list term), - subs Σ Γ0 s Γ' -> - Γ = Γ0 ,,, Γ' ,,, Δ -> - Σ;;; Γ0 ,,, subst_context s 0 Δ |- subst s #|Δ| t : subst s #|Δ| T) Σ - (Γ ,,, Γ' ,,, Δ) wfΓ0 -> - subs Σ Γ s Γ' -> - wf_local Σ (Γ ,,, subst_context s 0 Δ). -Proof. - intros wfΣ wfΓ. - induction Δ; simpl; intros Hwf. - intros _. clear Hwf. simpl in wfΓ. now eapply All_local_env_app in wfΓ. - depelim Hwf; - rewrite subst_context_snoc; simpl; constructor. - eapply IHΔ; eauto. red. exists (tu.π1). simpl. - rewrite Nat.add_0_r. eapply t0; eauto. - eapply IHΔ; eauto. red. exists (tu.π1). simpl. - rewrite Nat.add_0_r. eapply t1; eauto. - simpl. rewrite Nat.add_0_r. eapply t0; eauto. -Qed. - -Theorem substitution `{checker_flags} Σ Γ Γ' s Δ (t : term) T : - wf Σ.1 -> subs Σ Γ s Γ' -> - Σ ;;; Γ ,,, Γ' ,,, Δ |- t : T -> - Σ ;;; Γ ,,, subst_context s 0 Δ |- subst s #|Δ| t : subst s #|Δ| T. -Proof. - intros HΣ Hs Ht. - pose proof (typing_wf_local Ht). - generalize_eq Γ0 (Γ ,,, Γ' ,,, Δ). - revert Γ Γ' Δ s Hs. revert Σ HΣ Γ0 X t T Ht. - apply (typing_ind_env (fun Σ Γ0 t T => - forall (Γ Γ' Δ : context) (s : list term) - (sub : subs Σ Γ s Γ') (eqΓ0 : Γ0 = Γ ,,, Γ' ,,, Δ), - Σ ;;; Γ ,,, subst_context s 0 Δ |- subst s #|Δ| t : subst s #|Δ| T)); - intros Σ wfΣ Γ0 wfΓ0; intros; subst Γ0; simpl in *; - try ((epose proof (substitution_All_local_env_over wfΣ wfΓ0 X sub)) - || (epose proof (substitution_All_local_env_over wfΣ wfΓ0 X0 sub))); - try solve [econstructor; eauto]. - - - assert (wfcdecl : Ast.wf (decl_type decl)). - { clear X. apply typing_all_wf_decl in wfΓ0; eauto. - eapply (nth_error_forall) in wfΓ0; eauto. red in wfΓ0. intuition. } - elim (leb_spec_Set); intros Hn. - elim nth_error_spec. - + intros x Heq Hlt. - pose proof (subst_length _ _ _ _ sub). - rewrite -> nth_error_app_context_ge in H0 by lia. - rewrite -> nth_error_app_context_lt in H0 by lia. - eapply subs_nth_error in Heq; eauto. - destruct decl_body; try contradiction. - cbn -[skipn] in Heq. - eapply refine_type. - eapply (weakening _ _ (subst_context s 0 Δ)) in Heq; eauto with wf. - rewrite subst_context_length in Heq. eapply Heq. - rewrite -> commut_lift_subst_rec by lia. - rewrite <- (firstn_skipn (S (n - #|Δ|)) s) at 2. - rewrite -> subst_app_decomp. f_equal. - replace (S n) with ((S n - #|Δ|) + #|Δ|) by lia. - assert (eq:#|(map (lift0 #|skipn (S (n - #|Δ|)) s|) (firstn (S (n - #|Δ|)) s))| = S n - #|Δ|). - rewrite map_length. rewrite -> firstn_length by lia. lia. - rewrite <- eq. rewrite -> simpl_subst_rec; auto; try lia. - auto with wf. eapply subs_wf in sub; eauto. - now apply All_Forall, All_firstn. - - + intros Hs. - pose proof (subst_length _ _ _ _ sub). - rewrite H1 in Hs. - assert (S n = #|s| + (S (n - #|s|))) by lia. - rewrite H2. rewrite simpl_subst; auto; try lia. - constructor; auto. - rewrite -> nth_error_app_context_ge; try lia; rewrite subst_context_length. - rewrite -> 2!nth_error_app_context_ge in H0 by lia. - rewrite <- H0. f_equal. lia. - lia. - - + eapply subs_nth_error_lt in sub; eauto. - rewrite H0 in sub. simpl in sub. - eapply refine_type. constructor; eauto. - rewrite <- map_decl_type. - rewrite -> commut_lift_subst_rec by lia. - f_equal. lia. - - - econstructor; auto. eapply X1; eauto. - specialize (X1 Γ Γ' Δ s sub eq_refl). - specialize (X3 Γ Γ' (Δ,, vass n t) s sub eq_refl). - now rewrite subst_context_snoc0 in X3. - - - econstructor; auto. eapply X1; eauto. - specialize (X1 Γ Γ' Δ s sub eq_refl). - specialize (X3 Γ Γ' (Δ,, vass n t) s sub eq_refl). - now rewrite subst_context_snoc0 in X3. - - - specialize (X1 Γ Γ' Δ s sub eq_refl). - specialize (X3 Γ Γ' Δ s sub eq_refl). - specialize (X5 Γ Γ' (Δ,, vdef n b b_ty) s sub eq_refl). - rewrite subst_context_snoc0 in X5. - econstructor; eauto. - - - specialize (X0 Γ Γ' Δ s0 sub eq_refl). - eapply type_mkApps; eauto. - eapply typing_wf in X; eauto. destruct X. - clear X0 H2 H0 H1. - induction X1; try constructor; eauto. - specialize (p Γ Γ' Δ s0 sub eq_refl). - specialize (p0 Γ Γ' Δ s0 sub eq_refl). - simpl in *. econstructor; eauto. - change (tProd na (subst s0 #|Δ| A) (subst s0 (S #|Δ|) B)) - with (subst s0 #|Δ| (tProd na A B)). - eapply substitution_cumul; eauto. - eapply typing_wf in typrod; eauto. intuition. - unfold subst1 in IHX1 |- *. rewrite -> distr_subst in IHX1. - apply IHX1. - apply wf_subst. constructor; auto. - now apply typing_wf in ty. - apply typing_wf in typrod; eauto. - intuition. now inv H0. - eauto using subs_wf. - - - eapply refine_type. constructor; eauto. - rewrite !map_cst_type. eapply subst_declared_constant in H0 as ->; eauto. - - - eapply refine_type. econstructor; eauto. - eapply on_declared_inductive in isdecl as [on_mind on_ind]; auto. - apply onArity in on_ind as [s' Hindty]. - pose proof (typing_wf (Σ.1, ind_universes mdecl) wfΣ _ _ _ Hindty) as [clty _]. - apply typecheck_closed in Hindty as [_ Hindty]; eauto. symmetry. - move/andP/proj1: Hindty. rewrite -(closedn_subst_instance_constr _ _ u) => Hty. - apply: (subst_closedn s #|Δ|); auto with wf. - eapply closed_upwards. eauto. simpl; lia. - - - eapply refine_type. econstructor; eauto. - symmetry. - destruct (on_declared_constructor wfΣ isdecl) as [? [cs [? onc]]]. - eapply on_constructor_closed_wf in onc as [clty wty]; auto. - unfold type_of_constructor. - apply subst_closedn. 1: eauto with wf. - eapply closed_upwards; eauto. lia. - - - rewrite subst_mkApps map_app map_skipn. - specialize (X2 Γ Γ' Δ s sub eq_refl). - specialize (X4 Γ Γ' Δ s sub eq_refl). - assert (Hclos: closed_ctx (ind_params mdecl)). { - destruct isdecl as [Hmdecl Hidecl]. - eapply on_declared_minductive in Hmdecl; eauto. - eapply onParams in Hmdecl. - eapply closed_wf_local in Hmdecl; eauto. } - assert (Hclos': closed (ind_type idecl)). { - eapply on_declared_inductive in isdecl; eauto. - destruct isdecl as [_ oind]. clear -oind wfΣ. - apply onArity in oind; destruct oind as [s HH]; cbn in *. - apply typecheck_closed in HH; eauto. - apply snd in HH. apply andP in HH; apply HH. } - simpl. econstructor. all: eauto. - -- eapply subst_build_case_predicate_type in H1; tea. - simpl in *. subst params. rewrite firstn_map. - etransitivity; [|eapply H1; eauto]. f_equal. - now erewrite subst_declared_inductive. - now rewrite closedn_subst_instance_context. - -- now rewrite !subst_mkApps in X4. - -- simpl. - destruct (on_declared_inductive wfΣ isdecl) as [oind obod]. - pose obod.(onConstructors) as onc. - eapply (subst_build_branches_type s #|Δ|) in H3; eauto. - subst params. rewrite firstn_map. exact H3. - 4: now rewrite closedn_subst_instance_context. - all: eauto with wf. - admit. admit. - -- solve_all. - - - specialize (X2 Γ Γ' Δ s sub eq_refl). - eapply refine_type. econstructor. - eauto. - rewrite subst_mkApps in X2. eauto. - rewrite map_length; eauto. - rewrite <- (Nat.add_0_l #|Δ|). - pose proof (subs_wf _ _ _ _ wfΣ sub). - erewrite distr_subst_rec. simpl. - rewrite map_rev. subst ty. - f_equal. - pose proof (declared_projection_closed wfΣ isdecl). - pose proof (declared_projection_wf _ _ _ _ _ isdecl). - pose proof (typing_wf_gen _ wfΣ _ localenv_nil _ _ (type_Prop _)) as [X' _]. - specialize (H2 X'). clear X'. - symmetry; apply subst_closedn; eauto. now apply wf_subst_instance_constr. - rewrite List.rev_length H0. rewrite closedn_subst_instance_constr. - eapply closed_upwards; eauto. lia. auto. - - - rewrite -> (map_dtype _ (subst s (#|mfix| + #|Δ|))). - eapply type_Fix; auto. - * eapply fix_guard_subst ; eauto. - * now rewrite -> nth_error_map, H1. - * eapply All_map. - eapply (All_impl X0); simpl. - intros x [u [Hs Hs']]; exists u. - now specialize (Hs' _ _ _ _ sub eq_refl). - * eapply All_map. - eapply (All_impl X1); simpl. - intros x [[Hb Hlam] IH]. - rewrite subst_fix_context. - specialize (IH Γ Γ' (Δ ,,, (fix_context mfix)) _ sub). - rewrite app_context_assoc in IH. specialize (IH eq_refl). - split; auto. - rewrite subst_context_app Nat.add_0_r app_context_assoc in IH. - rewrite app_context_length fix_context_length in IH. - rewrite subst_context_length fix_context_length. - rewrite commut_lift_subst_rec; try lia. now rewrite (Nat.add_comm #|Δ|). - now rewrite isLambda_subst. - - - rewrite -> (map_dtype _ (subst s (#|mfix| + #|Δ|))). - eapply type_CoFix; auto. - * now rewrite -> nth_error_map, H0. - * eapply All_map. - eapply (All_impl X0); simpl. - intros x [u [Hs Hs']]; exists u. - now specialize (Hs' _ _ _ _ sub eq_refl). - * eapply All_map. - eapply (All_impl X1); simpl. - intros x [Hb IH]. - rewrite subst_fix_context. - specialize (IH Γ Γ' (Δ ,,, (fix_context mfix)) _ sub). - rewrite app_context_assoc in IH. specialize (IH eq_refl). - rewrite subst_context_app Nat.add_0_r app_context_assoc in IH. - rewrite app_context_length fix_context_length in IH. - rewrite subst_context_length fix_context_length. - rewrite commut_lift_subst_rec; try lia. now rewrite (Nat.add_comm #|Δ|). - - - assert(wfB : Ast.wf B). - { destruct X2 as [Bs|[u Hu]]. - destruct Bs as [[ctx [u [Hd IH]]]]. simpl in *. - generalize (destArity_spec [] B); rewrite Hd. - move => /= ->. clear -wfΣ IH. - eapply (it_mkProd_or_LetIn_wf (Γ ,,, Γ' ,,, Δ)). - rewrite -it_mkProd_or_LetIn_app. - eapply (wf_it_mkProd_or_LetIn _ _ IH); auto. - induction IH; constructor; auto. - destruct t0. split; try constructor. - eapply typing_wf in t0. intuition. auto. - eapply typing_wf in t1. intuition. auto. - eapply typing_wf in t1. intuition. auto. constructor. - destruct Hu. apply typing_wf in t0; intuition auto. } - econstructor; eauto. - destruct X2 as [Bs|[u Hu]]. - + left. destruct Bs as [[ctx [u [Hd IH]]]]. simpl in *. - exists (subst_context s #|Δ| ctx), u. - pose proof (subst_destArity [] B s #|Δ|). rewrite Hd in H0. - rewrite {}H0; auto. - split; auto. - apply All_local_env_app_inv; intuition auto. - clear -sub a. - induction ctx; try constructor; depelim a. - -- rewrite subst_context_snoc. - unfold snoc. - econstructor; auto. - eapply IHctx. eapply a. - simpl. destruct tu as [u tu]. exists u. - specialize (t0 _ _ (Δ ,,, ctx) _ sub). forward t0. - now rewrite app_context_assoc. simpl in t0. - now rewrite subst_context_app Nat.add_0_r app_context_assoc app_length in t0. - -- rewrite subst_context_snoc. - constructor; auto. - ++ eapply IHctx. eapply a. - ++ simpl. - specialize (t1 _ _ (Δ ,,, ctx) _ sub). - forward t1. 1: now rewrite app_context_assoc. - simpl in t1. - now rewrite subst_context_app Nat.add_0_r app_context_assoc app_length in t1. - ++ simpl. - specialize (t0 _ _ (Δ ,,, ctx) _ sub). forward t0. - 1: now rewrite app_context_assoc. simpl in t0. - now rewrite subst_context_app Nat.add_0_r app_context_assoc app_length in t0. - + right; exists u; intuition eauto. - + eapply substitution_cumul; eauto. - now eapply typing_wf in X0. -Admitted. - -Lemma substitution0 `{checker_flags} Σ Γ n u U (t : term) T : - wf Σ.1 -> - Σ ;;; Γ ,, vass n U |- t : T -> Σ ;;; Γ |- u : U -> - Σ ;;; Γ |- t {0 := u} : T {0 := u}. -Proof. - intros HΣ Ht Hu. - assert (wfΓ : wf_local Σ Γ). - apply typing_wf_local in Hu; eauto. - pose proof (substitution Σ Γ [vass n U] [u] [] t T HΣ) as thm. - forward thm. constructor. constructor. rewrite subst_empty; auto. - apply typing_wf in Hu; intuition. - now apply (thm Ht). -Qed. diff --git a/checker/theories/Weakening.v b/checker/theories/Weakening.v deleted file mode 100644 index 5a6031bfb..000000000 --- a/checker/theories/Weakening.v +++ /dev/null @@ -1,1150 +0,0 @@ -(* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import config utils Ast AstUtils Reflect Induction - LiftSubst UnivSubst Typing TypingWf LibHypsNaming. -From MetaCoq.Checker Require Import WeakeningEnv Closed. - -Require Import ssreflect. -From Equations Require Import Equations. - -(** * Weakening lemmas for typing derivations. - - [weakening_*] proves weakening of typing, reduction etc... w.r.t. the *local* - environment. *) - - -Derive Signature for Ast.wf Forall. - -Generalizable Variables Σ Γ t T. - -Lemma typed_liftn `{checker_flags} Σ Γ t T n k : - wf Σ.1 -> wf_local Σ Γ -> k >= #|Γ| -> - Σ ;;; Γ |- t : T -> lift n k T = T /\ lift n k t = t. -Proof. - intros wfΣ wfΓ Hk Hty. - apply typecheck_closed in Hty; eauto. - destruct Hty as [_ Hcl]. - rewrite -> andb_and in Hcl. destruct Hcl as [clb clty]. - pose proof (closed_upwards k clb). - pose proof (closed_upwards k clty). - simpl in *. forward H0 by lia. - apply (lift_closed n) in H0. - simpl in *. forward H1 by lia. - now apply (lift_closed n) in H1. -Qed. - -Lemma lift_decl0 k d : map_decl (lift 0 k) d = d. -Proof. - destruct d; destruct decl_body; unfold map_decl; simpl; - f_equal; now rewrite ?lift0_id. -Qed. - -Lemma lift0_context k Γ : lift_context 0 k Γ = Γ. -Proof. - unfold lift_context, fold_context. - rewrite rev_mapi. rewrite List.rev_involutive. - unfold mapi. generalize 0 at 2. generalize #|List.rev Γ|. - induction Γ; intros; simpl; trivial. - rewrite lift_decl0; f_equal; auto. -Qed. - -Lemma lift_context_length n k Γ : #|lift_context n k Γ| = #|Γ|. -Proof. apply fold_context_length. Qed. -Hint Rewrite lift_context_length : lift. - -Definition lift_context_snoc0 n k Γ d : lift_context n k (d :: Γ) = lift_context n k Γ ,, lift_decl n (#|Γ| + k) d. -Proof. unfold lift_context. now rewrite fold_context_snoc0. Qed. -Hint Rewrite lift_context_snoc0 : lift. - -Lemma lift_context_snoc n k Γ d : lift_context n k (Γ ,, d) = lift_context n k Γ ,, lift_decl n (#|Γ| + k) d. -Proof. - unfold snoc. apply lift_context_snoc0. -Qed. -Hint Rewrite lift_context_snoc : lift. - -Lemma lift_context_alt n k Γ : - lift_context n k Γ = - mapi (fun k' d => lift_decl n (Nat.pred #|Γ| - k' + k) d) Γ. -Proof. - unfold lift_context. apply fold_context_alt. -Qed. - -Lemma lift_context_app n k Γ Δ : - lift_context n k (Γ ,,, Δ) = lift_context n k Γ ,,, lift_context n (#|Γ| + k) Δ. -Proof. - unfold lift_context, fold_context, app_context. - rewrite List.rev_app_distr. - rewrite mapi_app. rewrite <- List.rev_app_distr. f_equal. f_equal. - apply mapi_ext. intros. f_equal. rewrite List.rev_length. f_equal. lia. -Qed. - -Lemma nth_error_lift_context: - forall (Γ' Γ'' : context) (v : nat), - v < #|Γ'| -> forall nth k, - nth_error Γ' v = Some nth -> - nth_error (lift_context #|Γ''| k Γ') v = Some (lift_decl #|Γ''| (#|Γ'| - S v + k) nth). -Proof. - induction Γ'; intros. - - easy. - - simpl. destruct v; rewrite lift_context_snoc0. - + simpl. repeat f_equal; try lia. simpl in *. congruence. - + simpl. apply IHΓ'; simpl in *; (lia || congruence). -Qed. - -Lemma nth_error_lift_context_eq: - forall (Γ' Γ'' : context) (v : nat) k, - nth_error (lift_context #|Γ''| k Γ') v = - option_map (lift_decl #|Γ''| (#|Γ'| - S v + k)) (nth_error Γ' v). -Proof. - induction Γ'; intros. - - simpl. unfold lift_context, fold_context; simpl. now rewrite nth_error_nil. - - simpl. destruct v; rewrite lift_context_snoc0. - + simpl. repeat f_equal; try lia. - + simpl. apply IHΓ'; simpl in *; (lia || congruence). -Qed. - -Lemma weaken_nth_error_ge {Γ Γ' v Γ''} : #|Γ'| <= v -> - nth_error (Γ ,,, Γ') v = - nth_error (Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ') (#|Γ''| + v). -Proof. - intros Hv. - rewrite -> !nth_error_app_context_ge, ?lift_context_length. f_equal. lia. - rewrite lift_context_length. lia. - rewrite lift_context_length. lia. auto. -Qed. - -Lemma weaken_nth_error_lt {Γ Γ' Γ'' v} : v < #|Γ'| -> - nth_error (Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ') v = - option_map (lift_decl #|Γ''| (#|Γ'| - S v)) (nth_error (Γ ,,, Γ') v). -Proof. - simpl. intros Hv. - rewrite -> !nth_error_app_context_lt. - rewrite nth_error_lift_context_eq. - do 2 f_equal. lia. lia. now rewrite lift_context_length. -Qed. - -Lemma lift_simpl {Γ'' Γ' : context} {i t} : i < #|Γ'| -> - lift0 (S i) (lift #|Γ''| (#|Γ'| - S i) t) = lift #|Γ''| #|Γ'| (lift0 (S i) t). -Proof. - intros. assert (#|Γ'| = S i + (#|Γ'| - S i)) by easy. - rewrite -> H0 at 2. - rewrite permute_lift; try easy. -Qed. - - -Lemma lift_decl_closed n k d : closed_decl k d -> lift_decl n k d = d. -Proof. - case: d => na [body|] ty; rewrite /closed_decl /lift_decl /map_decl /=. - move/andP => [cb cty]. now rewrite !lift_closed //. - move=> cty; now rewrite !lift_closed //. -Qed. - -Lemma closed_decl_upwards k d : closed_decl k d -> forall k', k <= k' -> closed_decl k' d. -Proof. - case: d => na [body|] ty; rewrite /closed_decl /=. - move/andP => [cb cty] k' lek'. do 2 rewrite (@closed_upwards k) //. - move=> cty k' lek'; rewrite (@closed_upwards k) //. -Qed. - -Lemma closed_ctx_lift n k ctx : closed_ctx ctx -> lift_context n k ctx = ctx. -Proof. - induction ctx in n, k |- *; auto. - unfold closed_ctx, id. simpl. - rewrite mapi_app forallb_app List.rev_length /= lift_context_snoc0 /snoc Nat.add_0_r. - move/andb_and => /= [Hctx /andb_and [Ha _]]. - rewrite IHctx // lift_decl_closed //. now apply: closed_decl_upwards. -Qed. - - -Lemma lift_iota_red n k pars c args brs : - lift n k (iota_red pars c args brs) = - iota_red pars c (List.map (lift n k) args) (List.map (on_snd (lift n k)) brs). -Proof. - unfold iota_red. rewrite !lift_mkApps. f_equal; auto using map_skipn. - rewrite nth_map; simpl; auto. -Qed. - -Lemma parsubst_empty k a : Ast.wf a -> subst [] k a = a. -Proof. - induction 1 in k |- * using term_wf_forall_list_ind; simpl; try congruence; - try solve [f_equal; eauto; solve_all; eauto]. - - - elim (Nat.compare_spec k n); destruct (Nat.leb_spec k n); intros; try easy. - subst. rewrite Nat.sub_diag. simpl. rewrite Nat.sub_0_r. reflexivity. - assert (n - k > 0) by lia. - assert (exists n', n - k = S n'). exists (Nat.pred (n - k)). lia. - destruct H2. rewrite H2. simpl. now rewrite Nat.sub_0_r. - - rewrite IHwf. rewrite mkApps_tApp; eauto with wf. - f_equal; solve_all. -Qed. - -Lemma lift_unfold_fix n k mfix idx narg fn : - unfold_fix mfix idx = Some (narg, fn) -> - unfold_fix (map (map_def (lift n k) (lift n (#|mfix| + k))) mfix) idx = Some (narg, lift n k fn). -Proof. - unfold unfold_fix. - rewrite nth_error_map. destruct (nth_error mfix idx) eqn:Hdef; try congruence. - case_eq (isLambda (dbody d)); [|discriminate]. - intros Hlam [= <- <-]. simpl. - rewrite isLambda_lift; tas. - repeat f_equal. - rewrite (distr_lift_subst_rec _ _ n 0 k). - rewrite fix_subst_length. f_equal. - unfold fix_subst. rewrite !map_length. - generalize #|mfix| at 2 3. induction n0; auto. simpl. - f_equal. apply IHn0. -Qed. -Hint Resolve lift_unfold_fix : core. - -Lemma lift_unfold_cofix n k mfix idx narg fn : - unfold_cofix mfix idx = Some (narg, fn) -> - unfold_cofix (map (map_def (lift n k) (lift n (#|mfix| + k))) mfix) idx = Some (narg, lift n k fn). -Proof. - unfold unfold_cofix. - rewrite nth_error_map. destruct (nth_error mfix idx) eqn:Hdef; try congruence. - intros [= <- <-]. simpl. repeat f_equal. - rewrite (distr_lift_subst_rec _ _ n 0 k). - rewrite cofix_subst_length. f_equal. - unfold cofix_subst. rewrite !map_length. - generalize #|mfix| at 2 3. induction n0; auto. simpl. - f_equal. apply IHn0. -Qed. -Hint Resolve lift_unfold_cofix : core. - -Lemma lift_is_constructor: - forall (args : list term) (narg : nat) n k, - is_constructor narg args = true -> is_constructor narg (map (lift n k) args) = true. -Proof. - intros args narg. - unfold is_constructor; intros. - rewrite nth_error_map. destruct nth_error; try discriminate. simpl. intros. - destruct t; try discriminate || reflexivity. - destruct t; try discriminate || reflexivity. -Qed. -Hint Resolve lift_is_constructor : core. - -Hint Rewrite UnivSubst.lift_subst_instance_constr : lift. -Hint Rewrite lift_mkApps : lift. -Hint Rewrite distr_lift_subst distr_lift_subst10 : lift. -Hint Rewrite lift_iota_red : lift. - -Lemma lift_declared_constant `{checker_flags} Σ cst decl n k : - wf Σ -> - declared_constant Σ cst decl -> - decl = map_constant_body (lift n k) decl. -Proof. - unfold declared_constant. - intros. - eapply lookup_on_global_env in H0; eauto. - destruct H0 as [Σ' [wfΣ' decl']]. - red in decl'. red in decl'. - destruct decl. simpl in *. destruct cst_body. unfold map_constant_body. simpl. - pose proof decl' as declty. - apply typecheck_closed in declty; eauto. - destruct declty as [declty Hcl]. - rewrite -> andb_and in Hcl. destruct Hcl as [clb clty]. - pose proof (closed_upwards k clb). - pose proof (closed_upwards k clty). - simpl in *. forward H0 by lia. forward H1 by lia. - apply (lift_closed n k) in H0. - apply (lift_closed n k) in H1. rewrite H0 H1. reflexivity. - constructor. - red in decl'. - destruct decl'. - apply typecheck_closed in t. destruct t as [_ ccst]. - rewrite -> andb_and in ccst. destruct ccst as [ccst _]. - eapply closed_upwards in ccst; simpl. - apply (lift_closed n k) in ccst. unfold map_constant_body. simpl. - rewrite ccst. reflexivity. lia. apply wfΣ'. constructor. -Qed. - -Definition lift_mutual_inductive_body n k m := - map_mutual_inductive_body (fun k' => lift n (k' + k)) m. - -Lemma lift_wf_local `{checker_flags} Σ Γ n k : wf Σ.1 -> wf_local Σ Γ -> lift_context n k Γ = Γ. -Proof. - intros wfΣ. - induction 1; auto; unfold lift_context, snoc; rewrite fold_context_snoc0; auto; unfold snoc; - f_equal; auto; unfold map_decl; simpl. unfold vass. simpl. f_equal. - destruct t0. - eapply typed_liftn; eauto. lia. unfold vdef. - f_equal. f_equal. eapply typed_liftn; eauto. lia. - destruct t0. - eapply typed_liftn in t0 as [Ht HT]; eauto. lia. -Qed. - -Lemma closed_wf_local `{checker_flags} Σ Γ : - wf Σ.1 -> - wf_local Σ Γ -> - closed_ctx Γ. -Proof. - intros wfΣ. unfold closed_ctx, mapi. intros wf. generalize 0. - induction wf; intros n; auto; unfold closed_ctx, snoc; simpl; - rewrite mapi_rec_app forallb_app; unfold id, closed_decl. - - simpl. - destruct t0 as [s Hs]. - eapply typecheck_closed in Hs as [closedΓ tcl]; eauto. - rewrite -> andb_and in *. simpl in *; rewrite andb_true_r in tcl |- *. - simpl in *. intuition auto. apply IHwf. auto. - erewrite closed_upwards; eauto. rewrite List.rev_length. lia. - - simpl. eapply typecheck_closed in t1 as [closedΓ tcl]; eauto. - rewrite -> andb_and in *. intuition auto. apply IHwf. - erewrite closed_upwards; eauto. - * erewrite closed_upwards; eauto. - rewrite List.rev_length. lia. - * rewrite List.rev_length. lia. -Qed. - -Lemma lift_declared_minductive `{checker_flags} {Σ : global_env} cst decl n k : - wf Σ -> - declared_minductive Σ cst decl -> - lift_mutual_inductive_body n k decl = decl. -Proof. - intros wfΣ Hdecl. - pose proof (on_declared_minductive wfΣ Hdecl). apply onNpars in X. - apply (declared_inductive_closed (Σ:=(empty_ext Σ))) in Hdecl. auto. - move: Hdecl. - rewrite /closed_inductive_decl /lift_mutual_inductive_body. - destruct decl; simpl. - move/andP => [clpar clbodies]. f_equal. - now rewrite [fold_context _ _]closed_ctx_lift. - eapply forallb_All in clbodies. - eapply Alli_mapi_id. - eapply (All_Alli clbodies). intros; eauto. - 2:{ intros. eapply X0. } - simpl; intros. - move: H0. rewrite /closed_inductive_body. - destruct x; simpl. move=> /andP[/andP [ci ct] cp]. - f_equal. rewrite lift_closed; eauto. - eapply closed_upwards; eauto; lia. - eapply All_map_id. eapply forallb_All in ct. - eapply (All_impl ct). intros x. - destruct x as [[id ty] arg]; unfold on_pi2; intros c; simpl; repeat f_equal. - apply lift_closed. unfold cdecl_type in c; simpl in c. - eapply closed_upwards; eauto; lia. - simpl in X. rewrite -X in cp. - eapply forallb_All in cp. eapply All_map_id; eauto. - eapply (All_impl cp); firstorder auto. - destruct x; unfold on_snd; simpl; f_equal. - apply lift_closed. rewrite context_assumptions_fold. - eapply closed_upwards; eauto; lia. - simpl. auto. -Qed. - -Lemma lift_declared_inductive `{checker_flags} Σ ind mdecl idecl n k : - wf Σ -> - declared_inductive Σ mdecl ind idecl -> - map_one_inductive_body (context_assumptions mdecl.(ind_params)) - (length (arities_context mdecl.(ind_bodies))) - (fun k' => lift n (k' + k)) - (inductive_ind ind) idecl = idecl. -Proof. - unfold declared_inductive. intros wfΣ [Hmdecl Hidecl]. - eapply (lift_declared_minductive _ _ n k) in Hmdecl. - unfold lift_mutual_inductive_body in Hmdecl. - destruct mdecl. simpl in *. - injection Hmdecl. intros Heq. - clear Hmdecl. - pose proof Hidecl as Hidecl'. - rewrite <- Heq in Hidecl'. - rewrite nth_error_mapi in Hidecl'. - clear Heq. - unfold option_map in Hidecl'. rewrite Hidecl in Hidecl'. - congruence. auto. -Qed. - -Lemma subst0_inds_lift ind u mdecl n k t : - (subst0 (inds (inductive_mind ind) u (ind_bodies mdecl)) - (lift n (#|arities_context (ind_bodies mdecl)| + k) t)) = - lift n k (subst0 (inds (inductive_mind ind) u (ind_bodies mdecl)) t). -Proof. - rewrite (distr_lift_subst_rec _ _ n 0 k). simpl. - unfold arities_context. rewrite rev_map_length inds_length. - f_equal. generalize (ind_bodies mdecl). - clear. intros. - induction l; unfold inds; simpl; auto. f_equal. auto. -Qed. - -Lemma lift_declared_constructor `{checker_flags} Σ c u mdecl idecl cdecl n k : - wf Σ -> - declared_constructor Σ mdecl idecl c cdecl -> - lift n k (type_of_constructor mdecl cdecl c u) = (type_of_constructor mdecl cdecl c u). -Proof. - unfold declared_constructor. destruct c as [i ci]. intros wfΣ [Hidecl Hcdecl]. - eapply (lift_declared_inductive _ _ _ _ n k) in Hidecl; eauto. - unfold type_of_constructor. destruct cdecl as [[id t'] arity]. - destruct idecl; simpl in *. - injection Hidecl. - intros. - pose Hcdecl as Hcdecl'. - rewrite <- H1 in Hcdecl'. - rewrite nth_error_map in Hcdecl'. rewrite Hcdecl in Hcdecl'. - simpl in Hcdecl'. injection Hcdecl'. - intros. - rewrite <- H3 at 2. - rewrite <- lift_subst_instance_constr. - now rewrite subst0_inds_lift. -Qed. - -Lemma lift_declared_projection `{checker_flags} Σ c mdecl idecl pdecl n k : - wf Σ -> - declared_projection Σ mdecl idecl c pdecl -> - on_snd (lift n (S (ind_npars mdecl + k))) pdecl = pdecl. -Proof. - intros. - eapply (declared_projection_closed (Σ:=empty_ext Σ)) in H0; auto. - unfold on_snd. simpl. - rewrite lift_closed. eapply closed_upwards; eauto; try lia. - destruct pdecl; reflexivity. -Qed. - -Lemma lift_fix_context: - forall (mfix : list (def term)) (n k : nat), - fix_context (map (map_def (lift n k) (lift n (#|mfix| + k))) mfix) = lift_context n k (fix_context mfix). -Proof. - intros mfix n k. unfold fix_context. - rewrite map_vass_map_def rev_mapi. - fold (fix_context mfix). - rewrite (lift_context_alt n k (fix_context mfix)). - unfold lift_decl. now rewrite mapi_length fix_context_length. -Qed. - -Hint Rewrite <- lift_fix_context : lift. - -Lemma All_local_env_lift `{checker_flags} - (P Q : context -> term -> option term -> Type) c n k : - All_local_env Q c -> - (forall Γ t T, - Q Γ t T -> - P (lift_context n k Γ) (lift n (#|Γ| + k) t) - (option_map (lift n (#|Γ| + k)) T) - ) -> - All_local_env P (lift_context n k c). -Proof. - intros Hq Hf. - induction Hq in |- *; try econstructor; eauto; - simpl; rewrite lift_context_snoc; econstructor; eauto. - - simpl. eapply (Hf _ _ None). eauto. - - simpl. eapply (Hf _ _ None). eauto. - - simpl. eapply (Hf _ _ (Some t)). eauto. -Qed. - -Lemma lift_destArity ctx t n k : Ast.wf t -> - destArity (lift_context n k ctx) (lift n (#|ctx| + k) t) = - match destArity ctx t with - | Some (args, s) => Some (lift_context n k args, s) - | None => None - end. -Proof. - intros wf; revert ctx. - induction wf in n, k |- * using term_wf_forall_list_ind; intros ctx; simpl; trivial. - specialize (IHwf0 n k (ctx,, vass n0 t)). rewrite lift_context_snoc in IHwf0. - simpl in IHwf0. unfold lift_decl, map_decl in IHwf0. unfold vass. simpl in IHwf0. rewrite IHwf0. - reflexivity. - specialize (IHwf1 n k (ctx,, vdef n0 t t0)). rewrite lift_context_snoc in IHwf1. - unfold vdef, lift_decl, map_decl in *. simpl in *. rewrite IHwf1. reflexivity. -Qed. - -Lemma lift_strip_outer_cast n k t : lift n k (strip_outer_cast t) = strip_outer_cast (lift n k t). -Proof. - induction t; simpl; try reflexivity. - now rewrite IHt1. -Qed. - -Definition on_pair {A B C D} (f : A -> B) (g : C -> D) (x : A * C) := - (f (fst x), g (snd x)). - -Lemma lift_instantiate_params_subst n k params args s t : - instantiate_params_subst (mapi_rec (fun k' decl => lift_decl n (k' + k) decl) params #|s|) - (map (lift n k) args) (map (lift n k) s) (lift n (#|s| + k) t) = - option_map (on_pair (map (lift n k)) (lift n (#|s| + k + #|params|))) (instantiate_params_subst params args s t). -Proof. - induction params in args, t, n, k, s |- *. - - destruct args; simpl; rewrite ?Nat.add_0_r; reflexivity. - - simpl. simpl. - destruct a as [na [body|] ty]; simpl; try congruence. - destruct t; simpl; try congruence. - -- specialize (IHparams n k args (subst0 s body :: s) t3). - rewrite <- Nat.add_succ_r. simpl in IHparams. - rewrite Nat.add_succ_r. - replace (#|s| + k + S #|params|) with (S (#|s| + k + #|params|)) by lia. - rewrite <- IHparams. - rewrite distr_lift_subst. reflexivity. - -- destruct t; simpl; try congruence. - destruct args; simpl; try congruence. - specialize (IHparams n k args (t :: s) t2). simpl in IHparams. - replace (#|s| + k + S #|params|) with (S (#|s| + k + #|params|)) by lia. - rewrite <- IHparams. auto. -Qed. - -Lemma instantiate_params_subst_length params args s t ctx t' : - instantiate_params_subst params args s t = Some (ctx, t') -> - #|ctx| = #|s| + #|params|. -Proof. - induction params in args, s, t, ctx, t' |- * ; destruct args; simpl; auto; try congruence. - rewrite Nat.add_0_r. congruence. - destruct decl_body. simpl. - destruct t; simpl; try congruence. - intros. erewrite IHparams; eauto. simpl. lia. - destruct t; simpl; try congruence. - destruct decl_body. simpl. - destruct t; simpl; try congruence. - intros. erewrite IHparams; eauto. simpl. lia. - destruct t; simpl; try congruence. - intros. erewrite IHparams; eauto. simpl. lia. -Qed. - -Lemma closed_tele_lift n k ctx : - closed_ctx ctx -> - mapi (fun (k' : nat) (decl : context_decl) => lift_decl n (k' + k) decl) (List.rev ctx) = List.rev ctx. -Proof. - rewrite /closedn_ctx /mapi. simpl. generalize 0. - induction ctx using rev_ind. move=> //. - move=> n0. - rewrite /closedn_ctx !rev_app_distr /id /=. - move/andP => [closedx Hctx]. - rewrite lift_decl_closed. rewrite (@closed_decl_upwards n0) //; try lia. - f_equal. now rewrite IHctx. -Qed. - -Lemma lift_instantiate_params n k params args t : - closed_ctx params -> - option_map (lift n k) (instantiate_params params args t) = - instantiate_params params (map (lift n k) args) (lift n k t). -Proof. - unfold instantiate_params. - move/(closed_tele_lift n k params)=> Heq. - rewrite -{2}Heq. - specialize (lift_instantiate_params_subst n k (List.rev params) args [] t). - move=> /= Heq'; rewrite Heq'. - case E: (instantiate_params_subst (List.rev params) args)=> [[l' t']|] /= //. - rewrite distr_lift_subst. - move/instantiate_params_subst_length: E => -> /=. do 3 f_equal. lia. -Qed. -Hint Rewrite lift_instantiate_params : lift. - -(* bug eauto with let .. in hypothesis failing *) -Lemma lift_decompose_prod_assum_rec ctx t n k : - let (ctx', t') := decompose_prod_assum ctx t in - (lift_context n k ctx', lift n (length ctx' + k) t') = - decompose_prod_assum (lift_context n k ctx) (lift n (length ctx + k) t). -Proof. - induction t in n, k, ctx |- *; simpl; - try rewrite -> Nat.sub_diag, Nat.add_0_r; try (eauto; congruence). - - eapply IHt1. - - specialize (IHt2 (ctx ,, vass na t1) n k). - destruct decompose_prod_assum. rewrite IHt2. simpl. - rewrite lift_context_snoc. reflexivity. - - specialize (IHt3 (ctx ,, vdef na t1 t2) n k). - destruct decompose_prod_assum. rewrite IHt3. simpl. - rewrite lift_context_snoc. reflexivity. -Qed. - -Lemma lift_decompose_prod_assum t n k : - let (ctx', t') := decompose_prod_assum [] t in - (lift_context n k ctx', lift n (length ctx' + k) t') = - decompose_prod_assum [] (lift n k t). -Proof. apply lift_decompose_prod_assum_rec. Qed. - -Lemma decompose_app_lift n k t f a : - decompose_app t = (f, a) -> decompose_app (lift n k t) = (lift n k f, map (lift n k) a). -Proof. destruct t; simpl; intros [= <- <-]; try reflexivity. Qed. -Hint Rewrite decompose_app_lift using auto : lift. - -Lemma lift_it_mkProd_or_LetIn n k ctx t : - lift n k (it_mkProd_or_LetIn ctx t) = - it_mkProd_or_LetIn (lift_context n k ctx) (lift n (length ctx + k) t). -Proof. - induction ctx in n, k, t |- *; simpl; try congruence. - pose (lift_context_snoc n k ctx a). unfold snoc in e. rewrite -> e. clear e. - simpl. rewrite -> IHctx. - pose (lift_context_snoc n k ctx a). - now destruct a as [na [b|] ty]. -Qed. -Hint Rewrite lift_it_mkProd_or_LetIn : lift. - -Lemma to_extended_list_lift n k c : - to_extended_list (lift_context n k c) = to_extended_list c. -Proof. - unfold to_extended_list, to_extended_list_k. generalize 0. - generalize (@nil term) at 1 2. - induction c in n, k |- *; simpl; intros. reflexivity. - rewrite -> lift_context_snoc0. unfold snoc. simpl. - destruct a. destruct decl_body. unfold lift_decl, map_decl. simpl. - now rewrite -> IHc. simpl. apply IHc. -Qed. - -Lemma to_extended_list_map_lift: - forall (n k : nat) (c : context), to_extended_list c = map (lift n (#|c| + k)) (to_extended_list c). -Proof. - intros n k c. - pose proof (to_extended_list_lift_above c). - symmetry. solve_all. - destruct H as [x' [-> Hx]]. simpl. - destruct (leb_spec_Set (#|c| + k) x'); unf_term. - f_equal. lia. reflexivity. -Qed. - -Lemma weakening_red1 `{CF:checker_flags} Σ Γ Γ' Γ'' M N : - wf Σ -> - red1 Σ (Γ ,,, Γ') M N -> - red1 Σ (Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ') (lift #|Γ''| #|Γ'| M) (lift #|Γ''| #|Γ'| N). -Proof. - intros wfΣ H. - remember (Γ ,,, Γ') as Γ0. revert Γ Γ' Γ'' HeqΓ0. - induction H using red1_ind_all in |- *; intros Γ0 Γ' Γ'' HeqΓ0; try subst Γ; simpl; - autorewrite with lift; - try solve [ econstructor; eauto ]. - - - elim (leb_spec_Set); intros Hn. - + rewrite -> simpl_lift; try lia. rewrite -> Nat.add_succ_r. - econstructor; eauto. - erewrite (weaken_nth_error_ge Hn) in H. eauto. - - + rewrite <- lift_simpl by easy. - econstructor. - rewrite -> (weaken_nth_error_lt Hn). - now unfold lift_decl; rewrite -> option_map_decl_body_map_decl, H. - - - econstructor; eauto. - rewrite H0. f_equal. - eapply (lookup_on_global_env _ _ _ _ wfΣ) in H. - destruct H as [Σ' [wfΣ' decl']]. - red in decl'. red in decl'. - rewrite -> H0 in decl'. - apply typecheck_closed in decl'; eauto. destruct decl'. - rewrite -> andb_and in i. destruct i as [Hclosed _]. - simpl in Hclosed. - pose proof (closed_upwards #|Γ'| Hclosed). - forward H by lia. - apply (lift_closed #|Γ''| #|Γ'|) in H. auto. constructor. - - - simpl. constructor. - now rewrite -> nth_error_map, H. - - - constructor. - specialize (IHred1 Γ0 (Γ' ,, vass na N) Γ'' eq_refl). - rewrite -> lift_context_snoc, Nat.add_0_r in IHred1. apply IHred1. - - - constructor. - specialize (IHred1 Γ0 (Γ' ,, vdef na b t) Γ'' eq_refl). - rewrite -> lift_context_snoc, Nat.add_0_r in IHred1. apply IHred1. - - - constructor. - induction X; constructor; auto. - intuition; eauto. - - - constructor. - induction X; constructor; auto. - intuition; eauto. - - - constructor. - specialize (IHred1 Γ0 (Γ' ,, vass na M1) Γ'' eq_refl). - rewrite -> lift_context_snoc, Nat.add_0_r in IHred1. apply IHred1. - - - constructor. - induction X; constructor; auto. - intuition; eauto. - - - constructor. - rewrite -> (OnOne2_length X). generalize (#|mfix1|). - induction X. all: simpl. all: constructor. 2: solve [ eauto ]. - simpl. intuition eauto. - inversion b. eauto. - - - apply fix_red_body. rewrite !lift_fix_context. - rewrite <- (OnOne2_length X). - eapply OnOne2_map. unfold on_Trel. solve_all. - + specialize (b0 Γ0 (Γ' ,,, fix_context mfix0)). - rewrite app_context_assoc in b0. specialize (b0 Γ'' eq_refl). - rewrite -> app_context_length, fix_context_length in *. - rewrite -> lift_context_app in *. - rewrite -> app_context_assoc, Nat.add_0_r in *. - auto. - + inversion b. eauto. - - - constructor. - rewrite -> (OnOne2_length X). generalize (#|mfix1|). - induction X. all: simpl. all: constructor. 2: solve [ eauto ]. - simpl. intuition eauto. - inversion b. eauto. - - - apply cofix_red_body. rewrite !lift_fix_context. - rewrite <- (OnOne2_length X). - eapply OnOne2_map. unfold on_Trel. solve_all. - + specialize (b0 Γ0 (Γ' ,,, fix_context mfix0)). - rewrite app_context_assoc in b0. specialize (b0 Γ'' eq_refl). - rewrite -> app_context_length, fix_context_length in *. - rewrite -> lift_context_app in *. - rewrite -> app_context_assoc, Nat.add_0_r in *. - auto. - + inversion b. eauto. -Qed. - -Lemma lift_eq_term_upto_univ Re Rl n k T U : - eq_term_upto_univ Re Rl T U -> - eq_term_upto_univ Re Rl (lift n k T) (lift n k U). -Proof. - induction T in n, k, U, Rl |- * using term_forall_list_rect; - inversion 1; simpl; try (now constructor). - - constructor. subst. clear - X X1. - induction l in X, args', X1 |- *. - + inversion X1; constructor. - + inversion X1. inversion X. subst. - simpl. constructor. all: easy. - - constructor. easy. clear - X X2. - induction l in X, args', X2 |- *. - + inversion X2; constructor. - + inversion X2; inversion X; subst. - now constructor. - - constructor; try easy. clear - X X3. - induction l in X, brs', X3 |- *. - + inversion X3; constructor. - + inversion X3; inversion X; subst. - constructor. cbn; easy. - easy. - - constructor; try easy. clear - X X1. - assert (XX:forall k k', All2 - (fun x y => eq_term_upto_univ Re Re (dtype x) (dtype y) × - eq_term_upto_univ Re Re (dbody x) (dbody y) × - rarg x = rarg y) - (map (map_def (lift n k) (lift n (#|m| + k'))) m) - (map (map_def (lift n k) (lift n (#|mfix'| + k'))) mfix')); - [|now apply XX]. clear k. - induction m in X, mfix', X1 |- *. - + inversion X1; constructor. - + inversion X1; inversion X; subst. - simpl. constructor. split. cbn; easy. - cbn; erewrite All2_length by eassumption. - easy. - unfold tFixProp in IHm. cbn. - rewrite !plus_n_Sm. now apply IHm. - - constructor; try easy. clear - X X1. - assert (XX:forall k k', All2 - (fun x y => eq_term_upto_univ Re Re (dtype x) (dtype y) × - eq_term_upto_univ Re Re (dbody x) (dbody y) × - rarg x = rarg y) - (map (map_def (lift n k) (lift n (#|m| + k'))) m) - (map (map_def (lift n k) (lift n (#|mfix'| + k'))) mfix')); - [|now apply XX]. clear k. - induction m in X, mfix', X1 |- *. - + inversion X1; constructor. - + inversion X1; inversion X; subst. - simpl. constructor. split. cbn; easy. - cbn; erewrite All2_length by eassumption. - easy. - unfold tFixProp in IHm. cbn. - rewrite !plus_n_Sm. now apply IHm. -Qed. - - -Lemma lift_eq_term `{checker_flags} ϕ n k T U : - eq_term ϕ T U -> eq_term ϕ (lift n k T) (lift n k U). -Proof. - apply lift_eq_term_upto_univ. -Qed. - -Lemma lift_leq_term `{checker_flags} ϕ n k T U : - leq_term ϕ T U -> leq_term ϕ (lift n k T) (lift n k U). -Proof. - apply lift_eq_term_upto_univ. -Qed. - -Lemma weakening_cumul `{CF:checker_flags} Σ Γ Γ' Γ'' M N : - wf Σ.1 -> - Σ ;;; Γ ,,, Γ' |- M <= N -> - Σ ;;; Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ' |- lift #|Γ''| #|Γ'| M <= lift #|Γ''| #|Γ'| N. -Proof. - intros wfΣ. induction 1. - constructor. - - now apply lift_leq_term. - - eapply weakening_red1 in r; auto. - econstructor 2; eauto. - - eapply weakening_red1 in r; auto. - econstructor 3; eauto. -Qed. - -Lemma lift_eq_decl `{checker_flags} ϕ n k d d' : - eq_decl ϕ d d' -> eq_decl ϕ (lift_decl n k d) (lift_decl n k d'). -Proof. - destruct d, d', decl_body, decl_body0; - unfold eq_decl, map_decl; cbn; intuition auto using lift_eq_term. -Qed. - -Lemma lift_eq_context `{checker_flags} φ l l' n k : - eq_context φ l l' -> - eq_context φ (lift_context n k l) (lift_context n k l'). -Proof. - induction l in l', n, k |- *; intros; destruct l'; rewrite -> ?lift_context_snoc0. - - constructor. - - inversion X. - - inversion X. - - inversion X; subst. constructor. - + apply All2_length in X1. rewrite X1. - now apply lift_eq_decl. - + now apply IHl. -Qed. - - -Lemma lift_build_case_predicate_type ind mdecl idecl u params ps n k : - closed_ctx (subst_instance_context u (ind_params mdecl)) -> - build_case_predicate_type ind mdecl - (map_one_inductive_body (context_assumptions mdecl.(ind_params)) - (length (arities_context (ind_bodies mdecl))) (fun k' => lift n (k' + k)) - (inductive_ind ind) idecl) - (map (lift n k) params) u ps - = option_map (lift n k) (build_case_predicate_type ind mdecl idecl params u ps). -Proof. -(* intros closedpars. unfold build_case_predicate_type. *) -(* rewrite -> ind_type_map. simpl. *) -(* epose proof (lift_instantiate_params n k _ params (subst_instance_constr u (ind_type idecl))) as H. *) -(* rewrite <- lift_subst_instance_constr. *) -(* erewrite <- H; trivial. clear H. *) -(* case_eq (instantiate_params (subst_instance_context u (ind_params mdecl)) params (subst_instance_constr u (ind_type idecl))) ; cbnr. *) -(* intros ity eq. *) -(* pose proof (lift_destArity [] ity n k) as H; cbn in H. rewrite H; clear H. *) -(* destruct destArity as [[ctx s] | ]; [|reflexivity]. simpl. f_equal. *) -(* rewrite lift_it_mkProd_or_LetIn; cbn. f_equal. f_equal. *) -(* - destruct idecl; reflexivity. *) -(* - rewrite lift_mkApps; cbn; f_equal. rewrite map_app. f_equal. *) -(* + rewrite !map_map lift_context_length; apply map_ext. clear. *) -(* intro. now rewrite -> permute_lift by lia. *) -(* + now rewrite -> to_extended_list_lift, <- to_extended_list_map_lift. *) -(* Qed. *) -Admitted. - -Lemma lift_build_branches_type ind mdecl idecl u p params n k : - closed_ctx (subst_instance_context u (ind_params mdecl)) -> - build_branches_type ind mdecl - (map_one_inductive_body (context_assumptions mdecl.(ind_params)) - #|arities_context (ind_bodies mdecl)| (fun k' => lift n (k' + k)) - (inductive_ind ind) idecl) - (map (lift n k) params) u (lift n k p) - = map (option_map (on_snd (lift n k))) - (build_branches_type ind mdecl idecl params u p). -Proof. - intros closedpars. unfold build_branches_type. - rewrite -> ind_ctors_map. - rewrite -> mapi_map, map_mapi. eapply mapi_ext. intros i x. - destruct x as [[id t] arity]. simpl. - rewrite <- lift_subst_instance_constr. - rewrite subst0_inds_lift. - rewrite <- lift_instantiate_params ; trivial. - match goal with - | |- context [ option_map _ (instantiate_params ?x ?y ?z) ] => - destruct (instantiate_params x y z) eqn:Heqip; cbnr - end. - epose proof (lift_decompose_prod_assum t0 n k). - destruct (decompose_prod_assum [] t0). - rewrite <- H. - destruct (decompose_app t1) as [fn arg] eqn:?. - rewrite (decompose_app_lift _ _ _ fn arg); auto. simpl. - destruct (chop _ arg) eqn:Heqchop. - eapply chop_map in Heqchop. - rewrite -> Heqchop. clear Heqchop. - unfold on_snd. simpl. f_equal. - rewrite -> lift_it_mkProd_or_LetIn, !lift_mkApps, map_app; simpl. - rewrite -> !lift_mkApps, !map_app, lift_context_length. - rewrite -> permute_lift by lia. arith_congr. - now rewrite -> to_extended_list_lift, <- to_extended_list_map_lift. -Qed. - - -Lemma wf_ind_type `{checker_flags} Σ mdecl ind idecl : - wf Σ -> declared_inductive Σ mdecl ind idecl -> Ast.wf (ind_type idecl). -Proof. - intros wfΣ Hidecl. - eapply typing_wf_sigma in wfΣ. - destruct Hidecl. red in H0. - eapply (lookup_on_global_env _ _ _ _ wfΣ) in H0 as [Σ' [wfΣ' H0]]; eauto. - red in H0. apply onInductives in H0. - eapply (nth_error_alli H1) in H0. apply onArity in H0 as [Hs Hpars]. wf. -Qed. -Hint Resolve wf_ind_type : wf. - - -Lemma destArity_it_mkProd_or_LetIn ctx ctx' t : - destArity ctx (it_mkProd_or_LetIn ctx' t) = - destArity (ctx ,,, ctx') t. -Proof. - induction ctx' in ctx, t |- *; simpl; auto. - rewrite IHctx'. destruct a as [na [b|] ty]; reflexivity. -Qed. - -Derive Signature for All_local_env_over. - -Lemma weakening_All_local_env_over `{cf : checker_flags} {Σ Γ Γ' Γ''} : - wf Σ.1 -> - wf_local Σ (Γ ,,, Γ'') -> - forall wfΓ0 : wf_local Σ (Γ ,,, Γ'), - All_local_env_over typing - (fun (Σ : global_env_ext) (Γ : context) (_ : wf_local Σ Γ) - (t T : term) (_ : Σ;;; Γ |- t : T) => - forall Γ0 Γ' Γ'' : context, - wf_local Σ (Γ0 ,,, Γ'') -> - Γ = Γ0 ,,, Γ' -> - Σ;;; Γ0 ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ' |- - lift #|Γ''| #|Γ'| t : lift #|Γ''| #|Γ'| T) Σ - (Γ ,,, Γ') wfΓ0 -> - wf_local Σ (Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ'). -Proof. - intros wfΣ wfΓ. - induction Γ'; simpl; intros wf Hwf. - induction Hwf; simpl; auto. - depelim Hwf; - rewrite lift_context_snoc; simpl; constructor. - eapply IHΓ'; eauto. red. exists (tu.π1). simpl. - rewrite Nat.add_0_r. apply t0; auto. - eapply IHΓ'; eauto. red. exists (tu.π1). simpl. - rewrite Nat.add_0_r. apply t1; auto. - simpl. rewrite Nat.add_0_r. apply t0; auto. -Qed. - - -Lemma weakening_typing `{cf : checker_flags} Σ Γ Γ' Γ'' (t : term) : - wf Σ.1 -> - wf_local Σ (Γ ,,, Γ') -> - wf_local Σ (Γ ,,, Γ'') -> - `(Σ ;;; Γ ,,, Γ' |- t : T -> - Σ ;;; Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ' |- - lift #|Γ''| #|Γ'| t : lift #|Γ''| #|Γ'| T). -Proof. - intros HΣ HΓΓ' HΓ'' * H. - generalize_eq Γ0 (Γ ,,, Γ'). intros eqw. - revert Γ Γ' Γ'' HΓ'' eqw. - revert Σ HΣ Γ0 HΓΓ' t T H. - apply (typing_ind_env (fun Σ Γ0 t T => forall Γ Γ' Γ'' : context, - wf_local Σ (Γ ,,, Γ'') -> - Γ0 = Γ ,,, Γ' -> - Σ;;; Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ' |- lift #|Γ''| #|Γ'| t : lift #|Γ''| #|Γ'| T)); - intros Σ wfΣ Γ0; !!intros; subst Γ0; simpl in *; - try ((epose proof (weakening_All_local_env_over wfΣ wf _ X); eauto) || - (epose proof (weakening_All_local_env_over wfΣ wf _ X0); eauto)); - try solve [econstructor; eauto]. - - - elim (leb_spec_Set); intros Hn. - + rewrite -> simpl_lift; try lia. rewrite -> Nat.add_succ_r. - constructor. auto. - now rewrite <- (weaken_nth_error_ge Hn). - + assert (forall t, lift0 (S n) (lift #|Γ''| (#|Γ'| - S n) t) = lift #|Γ''| #|Γ'| (lift0 (S n) t)). - intros. assert (H:#|Γ'| = S n + (#|Γ'| - S n)) by easy. - rewrite -> H at 2. - rewrite permute_lift; try easy. - rewrite <- H. - rewrite -> map_decl_type. constructor; auto. - now rewrite -> (weaken_nth_error_lt Hn), heq_nth_error. - - - econstructor; auto. - specialize (IHb Γ (Γ' ,, vass n t) Γ''). - forward IHb; auto. - rewrite -> lift_context_snoc, plus_0_r in IHb. - now eapply IHb. - - - econstructor; auto. - simpl. - specialize (IHb Γ (Γ' ,, vass n t) Γ''). - specialize (IHb wf eq_refl). - rewrite -> lift_context_snoc, plus_0_r in IHb. - eapply IHb. - - - econstructor; auto. - specialize (IHb Γ Γ' Γ'' wf eq_refl). simpl. - specialize (IHb' Γ (Γ' ,, vdef n b b_ty) Γ'' wf eq_refl). - rewrite -> lift_context_snoc, plus_0_r in IHb'. - apply IHb'. - - - econstructor; auto. - now apply lift_isApp. - now apply map_non_nil. - clear IHt heq_isApp hneq_l typet. - induction X1. econstructor; eauto. - eapply type_spine_cons. - simpl in p. apply p; auto. - change (tProd na (lift #|Γ''| #|Γ'| A) (lift #|Γ''| (S #|Γ'|) B)) - with (lift #|Γ''| #|Γ'| (tProd na A B)). - eapply weakening_cumul; eauto. auto. - rewrite -> distr_lift_subst10 in IHX1. apply IHX1. - - - autorewrite with lift. - rewrite -> map_cst_type. constructor; auto. - erewrite <- lift_declared_constant; eauto. - - - autorewrite with lift. - erewrite <- (ind_type_map (fun k' => lift #|Γ''| (k' + #|Γ'|))). - pose proof isdecl as isdecl'. - destruct isdecl. intuition auto. - eapply lift_declared_inductive in isdecl'. - rewrite -> isdecl'. - econstructor; try red; intuition eauto. - auto. - - - rewrite (lift_declared_constructor _ (ind, i) u mdecl idecl cdecl _ _ wfΣ isdecl). - econstructor; eauto. - - - rewrite -> lift_mkApps, map_app, map_skipn. - specialize (IHc _ _ _ wf eq_refl). - specialize (IHp _ _ _ wf eq_refl). - assert (Hclos: closed_ctx (subst_instance_context u (ind_params mdecl))). { - destruct isdecl as [Hmdecl Hidecl]. - eapply on_declared_minductive in Hmdecl; eauto. - eapply onParams in Hmdecl. - rewrite -> closedn_subst_instance_context. - eapply closed_wf_local in Hmdecl; eauto. } - simpl. econstructor. - 7:{ cbn. rewrite -> firstn_map. - erewrite lift_build_branches_type; tea. - rewrite map_option_out_map_option_map. - subst params. erewrite heq_map_option_out. reflexivity. } - all: eauto. - -- erewrite -> lift_declared_inductive; eauto. - -- simpl. erewrite firstn_map, lift_build_case_predicate_type; tea. - subst params. erewrite heq_build_case_predicate_type; reflexivity. - -- destruct idecl; simpl in *; auto. - -- now rewrite -> !lift_mkApps in IHc. - -- solve_all. - - - simpl. - erewrite (distr_lift_subst_rec _ _ _ 0 #|Γ'|). - simpl. rewrite -> map_rev. - subst ty. - rewrite -> List.rev_length, lift_subst_instance_constr. - replace (lift #|Γ''| (S (#|args| + #|Γ'|)) (snd pdecl)) - with (snd (on_snd (lift #|Γ''| (S (#|args| + #|Γ'|))) pdecl)) by now destruct pdecl. - econstructor. - red. split. apply (proj1 isdecl). - split. rewrite -> (proj1 (proj2 isdecl)). f_equal. - rewrite -> heq_length. - symmetry; eapply lift_declared_projection; eauto. - apply (proj2 (proj2 isdecl)). - specialize (IHc _ _ _ wf eq_refl). - rewrite -> lift_mkApps in *. eapply IHc. - now rewrite -> map_length. - - - rewrite -> (map_dtype _ (lift #|Γ''| (#|mfix| + #|Γ'|))). - eapply type_Fix; auto. - * eapply fix_guard_lift ; eauto. - * rewrite -> nth_error_map, heq_nth_error. reflexivity. - * eapply All_map. - eapply (All_impl X0); simpl. - intros x [s [Hs Hs']]; exists s. - now specialize (Hs' _ _ _ wf eq_refl). - * eapply All_map. - eapply (All_impl X1); simpl. - intros x [[Hb Hlam] IH]. - rewrite lift_fix_context. - specialize (IH Γ (Γ' ,,, (fix_context mfix)) Γ'' wf). - rewrite app_context_assoc in IH. specialize (IH eq_refl). - split; auto. - rewrite lift_context_app Nat.add_0_r app_context_assoc in IH. - rewrite app_context_length fix_context_length in IH. - rewrite lift_context_length fix_context_length. - rewrite permute_lift; try lia. now rewrite (Nat.add_comm #|Γ'|). - now rewrite isLambda_lift. - - - rewrite -> (map_dtype _ (lift #|Γ''| (#|mfix| + #|Γ'|))). - eapply type_CoFix; auto. - * rewrite -> nth_error_map, heq_nth_error. reflexivity. - * eapply All_map. - eapply (All_impl X0); simpl. - intros x [s [Hs Hs']]; exists s. - now specialize (Hs' _ _ _ wf eq_refl). - * eapply All_map. - eapply (All_impl X1); simpl. - intros x [Hb IH]. - rewrite lift_fix_context. - specialize (IH Γ (Γ' ,,, (fix_context mfix)) Γ'' wf). - rewrite app_context_assoc in IH. specialize (IH eq_refl). - rewrite lift_context_app Nat.add_0_r app_context_assoc in IH. - rewrite app_context_length fix_context_length in IH. - rewrite lift_context_length fix_context_length. - rewrite permute_lift; try lia. now rewrite (Nat.add_comm #|Γ'|). - - - econstructor; eauto. - destruct IHB. - + left. destruct i as [[ctx [u [Hd IH]]]]. simpl in *. - exists (lift_context #|Γ''| #|Γ'| ctx), u. - rewrite (lift_destArity [] B #|Γ''| #|Γ'|) ?Hd. - { generalize (destArity_spec [] B); rewrite Hd. - move => /= ->. clear -wfΣ IH. - eapply (it_mkProd_or_LetIn_wf (Γ ,,, Γ')). - rewrite -it_mkProd_or_LetIn_app. - eapply (wf_it_mkProd_or_LetIn _ _ IH); auto. - induction IH; constructor; auto. - destruct t0. split; try constructor. - eapply typing_wf in t0. intuition. auto. - eapply typing_wf in t1. intuition. auto. - eapply typing_wf in t1. intuition. auto. constructor. } - split; auto. - apply All_local_env_app_inv; intuition auto. - clear -wf a. - induction ctx; try constructor; depelim a. - -- rewrite lift_context_snoc. - constructor; auto. - eapply IHctx. eapply a. - simpl. destruct tu as [u tu]. exists u. - specialize (t0 Γ (Γ' ,,, ctx) Γ''). forward t0. auto. - rewrite app_context_assoc in t0. - specialize (t0 eq_refl). simpl in t0. - rewrite app_context_length lift_context_app app_context_assoc Nat.add_0_r in t0. apply t0. - -- rewrite lift_context_snoc. - constructor; auto. - ++ eapply IHctx. eapply a. - ++ simpl. - specialize (t1 Γ (Γ' ,,, ctx) Γ''). forward t1 by auto. - rewrite app_context_assoc in t1. - specialize (t1 eq_refl). simpl in t1. - rewrite app_context_length lift_context_app app_context_assoc Nat.add_0_r in t1. - eexists. apply t1. - ++ simpl. - specialize (t0 Γ (Γ' ,,, ctx) Γ'' wf). - rewrite app_context_assoc in t0. - specialize (t0 eq_refl). simpl in t0. - rewrite app_context_length lift_context_app app_context_assoc Nat.add_0_r in t0. - apply t0. - + right. destruct s as [u Hu]; exists u. intuition; now eapply weakening_cumul. - + now eapply weakening_cumul. -Qed. - -Lemma weakening `{cf : checker_flags} Σ Γ Γ' (t : term) T : - wf Σ.1 -> wf_local Σ (Γ ,,, Γ') -> - Σ ;;; Γ |- t : T -> - Σ ;;; Γ ,,, Γ' |- lift0 #|Γ'| t : lift0 #|Γ'| T. -Proof. - intros HΣ HΓΓ' * H. - pose (weakening_typing Σ Γ [] Γ' t). - forward t0; eauto. - forward t0; eauto. now eapply wf_local_app in HΓΓ'. -Qed. - -Definition fix_context_gen k mfix := - (List.rev - (mapi_rec - (fun (i : nat) (d : def term) => - vass (dname d) (lift0 i (dtype d))) mfix k)). - -Corollary All_mfix_wf {cf:checker_flags} Σ Γ mfix : - wf Σ.1 -> wf_local Σ Γ -> - All (fun d : def term => isType Σ Γ (dtype d)) mfix -> - wf_local Σ (Γ ,,, fix_context mfix). -Proof. - move=> wfΣ wf a; move: wf. - change (fix_context mfix) with (fix_context_gen #|@nil context_decl| mfix). - change Γ with (Γ ,,, []). - generalize (@nil context_decl) as Δ. - rewrite /fix_context_gen. - intros Δ wfΔ. - eapply All_local_env_app_inv. split; auto. - induction a in Δ, wfΔ |- *; simpl; auto. constructor. - simpl. - eapply All_local_env_app_inv. split; auto. - constructor. constructor. simpl. - destruct p as [s Hs]. - exists s. eapply (weakening Σ Γ Δ _ (tSort s)); auto. - specialize (IHa (Δ ,,, [vass (dname x) (lift0 #|Δ| (dtype x))])). - rewrite app_length in IHa. simpl in IHa. - forward IHa. simpl. constructor; auto. - destruct p as [s Hs]. - exists s. eapply (weakening Σ Γ Δ _ (tSort s)); auto. - eapply All_local_env_impl; eauto. - simpl; intros. - rewrite app_context_assoc. apply X. -Qed. diff --git a/checker/theories/WeakeningEnv.v b/checker/theories/WeakeningEnv.v deleted file mode 100644 index 2b212f77c..000000000 --- a/checker/theories/WeakeningEnv.v +++ /dev/null @@ -1,668 +0,0 @@ -(* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import config utils Ast AstUtils - LibHypsNaming Typing. -Require Import ssreflect. - -(** * Weakening lemmas w.r.t. the global environment *) - - -Generalizable Variables Σ Γ t T. - -Definition extends (Σ Σ' : global_env) := { Σ'' & Σ' = Σ'' ++ Σ }. - -Lemma weakening_env_global_ext_levels Σ Σ' φ (H : extends Σ Σ') l - : LevelSet.In l (global_ext_levels (Σ, φ)) - -> LevelSet.In l (global_ext_levels (Σ', φ)). -Proof. - unfold global_ext_levels; simpl. - intros Hl. apply LevelSet.union_spec in Hl. - apply LevelSet.union_spec. - destruct Hl as [Hl|Hl]; [now left|right]. clear φ. - destruct H as [Σ'' eq]; subst. - induction Σ'' in l, Hl |- *. - - now rewrite app_nil_l. - - simpl. apply LevelSet.union_spec. right; eauto. -Qed. -Hint Resolve weakening_env_global_ext_levels : extends. - -Lemma weakening_env_global_ext_levels' Σ Σ' φ (H : extends Σ Σ') l - : LevelSet.mem l (global_ext_levels (Σ, φ)) - -> LevelSet.mem l (global_ext_levels (Σ', φ)). -Proof. - intro HH. apply LevelSet.mem_spec in HH. - now eapply LevelSet.mem_spec, weakening_env_global_ext_levels. -Qed. - -Lemma lookup_env_Some_fresh Σ c decl : - lookup_env Σ c = Some decl -> ~ (fresh_global c Σ). -Proof. - induction Σ; cbn. congruence. - unfold eq_kername. destruct (kername_eq_dec c a.1). - - subst c. intros [= <-]. intros H2. inv H2. contradiction. - - intros H1 H2. apply IHΣ; tas. - now inv H2. -Qed. - -Lemma extends_lookup `{checker_flags} Σ Σ' c decl : - wf Σ' -> extends Σ Σ' -> lookup_env Σ c = Some decl -> lookup_env Σ' c = Some decl. -Proof. - intros wfΣ' [Σ'' ->]. simpl. - induction Σ'' in wfΣ', c, decl |- *. simpl. auto. - specialize (IHΣ'' c decl). forward IHΣ''. - inv wfΣ'. simpl in X0. apply X. - intros HΣ. specialize (IHΣ'' HΣ). - inv wfΣ'. simpl in *. - unfold eq_kername; destruct kername_eq_dec; subst. - eapply lookup_env_Some_fresh in IHΣ''; eauto. contradiction. - assumption. -Qed. -Hint Resolve extends_lookup : extends. - -Lemma extends_wf_local `{checker_flags} Σ Γ (wfΓ : wf_local Σ Γ) : - All_local_env_over typing - (fun Σ0 Γ0 wfΓ (t T : term) ty => - forall Σ' : global_env, - wf Σ' -> - extends Σ0 Σ' -> - (Σ', Σ0.2);;; Γ0 |- t : T - ) Σ Γ wfΓ -> - forall Σ' : global_env, wf Σ' -> extends Σ Σ' -> wf_local (Σ', Σ.2) Γ. -Proof. - intros X0 Σ' H0. - induction X0 in H0 |- *; try econstructor; simpl in *; intuition auto. - - destruct tu as [u Hu]; exists u; auto. - - destruct tu as [u Hu]; exists u; auto. -Qed. -Hint Resolve extends_wf_local : extends. - -Lemma weakening_env_red1 `{CF:checker_flags} Σ Σ' Γ M N : - wf Σ' -> - extends Σ Σ' -> - red1 Σ Γ M N -> - red1 Σ' Γ M N. -Proof. - induction 3 using red1_ind_all; - try solve [econstructor; eauto; - eapply (OnOne2_impl X1); simpl; intuition eauto]. - - eapply extends_lookup in X0; eauto. - econstructor; eauto. -Qed. - -Lemma global_ext_constraints_app Σ Σ' φ - : ConstraintSet.Subset (global_ext_constraints (Σ, φ)) - (global_ext_constraints (Σ' ++ Σ, φ)). -Proof. - unfold global_ext_constraints; simpl. - intros ctr Hc. apply ConstraintSet.union_spec in Hc. - apply ConstraintSet.union_spec. - destruct Hc as [Hc|Hc]; [now left|right]. clear φ. - induction Σ' in ctr, Hc |- *. - - now rewrite app_nil_l. - - simpl. apply ConstraintSet.union_spec. right; eauto. -Qed. - -Lemma satisfies_subset φ φ' val : - ConstraintSet.Subset φ φ' -> - satisfies val φ' -> - satisfies val φ. -Proof. - intros sub sat ? isin. - apply sat, sub; auto. -Qed. - -Lemma leq_universe_subset {cf:checker_flags} ctrs ctrs' t u - : ConstraintSet.Subset ctrs ctrs' - -> leq_universe ctrs t u -> leq_universe ctrs' t u. -Proof. - intros Hctrs H. unfold leq_universe in *. - destruct check_univs; [|trivial]. - intros v Hv. apply H. - eapply satisfies_subset; eauto. -Qed. - -Lemma eq_universe_subset {cf:checker_flags} ctrs ctrs' t u - : ConstraintSet.Subset ctrs ctrs' - -> eq_universe ctrs t u -> eq_universe ctrs' t u. -Proof. - intros Hctrs H. unfold eq_universe in *. - destruct check_univs; [|trivial]. - intros v Hv. apply H. - eapply satisfies_subset; eauto. -Qed. - - -Lemma eq_term_upto_univ_morphism0 (Re Re' : _ -> _ -> Prop) - (Hre : forall t u, Re t u -> Re' t u) - : forall t u, eq_term_upto_univ Re Re t u -> eq_term_upto_univ Re' Re' t u. -Proof. - fix aux 3. - destruct 1; constructor; eauto. - all: unfold R_universe_instance in *. - all: match goal with - | H : All2 _ _ _ |- _ => induction H; constructor; eauto - | H : Forall2 _ _ _ |- _ => induction H; constructor; eauto - end. - all: intuition eauto. -Qed. - -Lemma eq_term_upto_univ_morphism (Re Re' Rle Rle' : _ -> _ -> Prop) - (Hre : forall t u, Re t u -> Re' t u) - (Hrle : forall t u, Rle t u -> Rle' t u) - : forall t u, eq_term_upto_univ Re Rle t u -> eq_term_upto_univ Re' Rle' t u. -Proof. - fix aux 3. - destruct 1; constructor; eauto using eq_term_upto_univ_morphism0. - all: unfold R_universe_instance in *. - all: match goal with - | H : Forall2 _ _ _ |- _ => induction H; constructor; - eauto using eq_term_upto_univ_morphism0 - | H : All2 _ _ _ |- _ => induction H; constructor; - eauto using eq_term_upto_univ_morphism0 - end. - - destruct r. split; eauto using eq_term_upto_univ_morphism0. - - destruct r as [? [? ?]]. - repeat split; eauto using eq_term_upto_univ_morphism0. - - destruct r as [? [? ?]]. - repeat split; eauto using eq_term_upto_univ_morphism0. -Qed. - -Lemma eq_term_subset {cf:checker_flags} φ φ' t t' - : ConstraintSet.Subset φ φ' - -> eq_term φ t t' -> eq_term φ' t t'. -Proof. - intro H. apply eq_term_upto_univ_morphism. - all: intros u u'; eapply eq_universe_subset; assumption. -Qed. - -Lemma leq_term_subset {cf:checker_flags} ctrs ctrs' t u - : ConstraintSet.Subset ctrs ctrs' -> leq_term ctrs t u -> leq_term ctrs' t u. -Proof. - intro H. apply eq_term_upto_univ_morphism. - intros t' u'; eapply eq_universe_subset; assumption. - intros t' u'; eapply leq_universe_subset; assumption. -Qed. - - -Lemma weakening_env_conv `{CF:checker_flags} Σ Σ' φ Γ M N : - wf Σ' -> - extends Σ Σ' -> - conv (Σ, φ) Γ M N -> - conv (Σ', φ) Γ M N. -Proof. - intros wfΣ [Σ'' ->]. - induction 1; simpl. - - econstructor. eapply eq_term_subset. - + eapply global_ext_constraints_app. - + assumption. - - econstructor 2; eauto. eapply weakening_env_red1; eauto. exists Σ''; eauto. - - econstructor 3; eauto. eapply weakening_env_red1; eauto. exists Σ''; eauto. -Qed. - -Lemma weakening_env_cumul `{CF:checker_flags} Σ Σ' φ Γ M N : - wf Σ' -> extends Σ Σ' -> - cumul (Σ, φ) Γ M N -> cumul (Σ', φ) Γ M N. -Proof. - intros wfΣ [Σ'' ->]. - induction 1; simpl. - - econstructor. eapply leq_term_subset. eapply global_ext_constraints_app. - assumption. - - econstructor 2; eauto. eapply weakening_env_red1; eauto. exists Σ''; eauto. - - econstructor 3; eauto. eapply weakening_env_red1; eauto. exists Σ''; eauto. -Qed. - -Lemma weakening_env_is_allowed_elimination `{CF:checker_flags} Σ Σ' φ u allowed : - wf Σ' -> extends Σ Σ' -> - is_allowed_elimination (global_ext_constraints (Σ, φ)) u allowed -> - is_allowed_elimination (global_ext_constraints (Σ', φ)) u allowed. -Proof. - intros wfΣ [Σ'' ->] al. - unfold is_allowed_elimination in *. - destruct check_univs; auto. - intros val sat. - unshelve epose proof (al val _) as al. - { eapply satisfies_subset; eauto. - apply global_ext_constraints_app. } - destruct allowed; auto; cbn in *; destruct ?; auto. -Qed. -Hint Resolve weakening_env_is_allowed_elimination : extends. - -Lemma weakening_env_declared_constant `{CF:checker_flags}: - forall (Σ : global_env) cst (decl : constant_body), - declared_constant Σ cst decl -> - forall Σ' : global_env, wf Σ' -> extends Σ Σ' -> declared_constant Σ' cst decl. -Proof. - intros Σ cst decl H0 Σ' X2 H2. - eapply extends_lookup; eauto. -Qed. -Hint Resolve weakening_env_declared_constant : extends. - -Lemma weakening_env_declared_minductive `{CF:checker_flags}: - forall (Σ : global_env) ind decl, - declared_minductive Σ ind decl -> - forall Σ' : global_env, wf Σ' -> extends Σ Σ' -> declared_minductive Σ' ind decl. -Proof. - intros Σ cst decl H0 Σ' X2 H2. - eapply extends_lookup; eauto. -Qed. -Hint Resolve weakening_env_declared_minductive : extends. - -Lemma weakening_env_declared_inductive: - forall (H : checker_flags) (Σ : global_env) ind mdecl decl, - declared_inductive Σ ind mdecl decl -> - forall Σ' : global_env, wf Σ' -> extends Σ Σ' -> declared_inductive Σ' ind mdecl decl. -Proof. - intros H Σ cst decl H0 [Hmdecl Hidecl] Σ' X2 H2. split; eauto with extends. -Qed. -Hint Resolve weakening_env_declared_inductive : extends. - -Lemma weakening_env_declared_constructor : - forall (H : checker_flags) (Σ : global_env) ind mdecl idecl decl, - declared_constructor Σ ind mdecl idecl decl -> - forall Σ' : global_env, wf Σ' -> extends Σ Σ' -> - declared_constructor Σ' ind mdecl idecl decl. -Proof. - intros H Σ cst mdecl idecl cdecl [Hidecl Hcdecl] Σ' X2 H2. - split; eauto with extends. -Qed. -Hint Resolve weakening_env_declared_constructor : extends. - -Lemma weakening_env_declared_projection : - forall (H : checker_flags) (Σ : global_env) ind mdecl idecl decl, - declared_projection Σ ind mdecl idecl decl -> - forall Σ' : global_env, wf Σ' -> extends Σ Σ' -> - declared_projection Σ' ind mdecl idecl decl. -Proof. - intros H Σ cst mdecl idecl cdecl [Hidecl Hcdecl] Σ' X2 H2. - split; eauto with extends. -Qed. -Hint Resolve weakening_env_declared_projection : extends. - -Lemma weakening_All_local_env_impl `{checker_flags} - (P Q : context -> term -> option term -> Type) l : - All_local_env P l -> - (forall Γ t T, P Γ t T -> Q Γ t T) -> - All_local_env Q l. -Proof. - induction 1; intros; simpl; econstructor; eauto. -Qed. - -Lemma weakening_env_global_ext_constraints Σ Σ' φ (H : extends Σ Σ') - : ConstraintSet.Subset (global_ext_constraints (Σ, φ)) - (global_ext_constraints (Σ', φ)). -Proof. - destruct H as [Σ'' eq]. subst. - apply global_ext_constraints_app. -Qed. - -Lemma eq_decl_subset {cf:checker_flags} φ φ' d d' - : ConstraintSet.Subset φ φ' - -> eq_decl φ d d' -> eq_decl φ' d d'. -Proof. - intros Hφ [H1 H2]. split; [|eapply eq_term_subset; eauto]. - destruct d as [na [bd|] ty], d' as [na' [bd'|] ty']; cbn in *; trivial. - eapply eq_term_subset; eauto. -Qed. - -Lemma eq_context_subset {cf:checker_flags} φ φ' Γ Γ' - : ConstraintSet.Subset φ φ' - -> eq_context φ Γ Γ' -> eq_context φ' Γ Γ'. -Proof. - intros Hφ. induction 1; constructor. - eapply eq_decl_subset; eassumption. assumption. -Qed. - -Lemma weakening_env_consistent_instance {cf:checker_flags} : - forall Σ Σ' φ ctrs u, - extends Σ Σ' -> - consistent_instance_ext (Σ, φ) ctrs u -> - consistent_instance_ext (Σ', φ) ctrs u. -Proof. - intros Σ Σ' φ ctrs u he. - unfold consistent_instance_ext, consistent_instance. - intros hc. - destruct ctrs=> //. - destruct cst as [nas cst]; simpl in *. - intuition auto. - eapply forallb_Forall in H. eapply forallb_Forall, Forall_impl; tea. - intros ? ? ; now eapply weakening_env_global_ext_levels'. - eapply valid_subset; tea; - now eapply weakening_env_global_ext_constraints. -Qed. -Hint Resolve weakening_env_consistent_instance : extends. - - -Lemma global_levels_Set Σ : - LevelSet.In Level.lSet (global_levels Σ). -Proof. - unfold global_levels. - induction Σ; simpl; auto. - - now apply LevelSet.singleton_spec. - - apply LevelSet.union_spec. right; auto. -Qed. - -Lemma global_levels_set Σ : - LevelSet.Equal (LevelSet.union (LevelSet.singleton Level.lSet) (global_levels Σ)) - (global_levels Σ). -Proof. - apply LevelSetProp.union_subset_equal. - intros x hin. eapply LevelSet.singleton_spec in hin; subst x. - apply global_levels_Set. -Qed. - -Lemma global_levels_ext {Σ Σ'} : - LevelSet.Equal (global_levels (Σ ++ Σ')) (LevelSet.union (global_levels Σ) (global_levels Σ')). -Proof. - unfold global_levels at 1. - induction Σ; simpl. - - rewrite global_levels_set. reflexivity. - - rewrite IHΣ. lsets. -Qed. - -Lemma extends_wf_universe {cf:checker_flags} {Σ : global_env_ext} Σ' u : extends Σ Σ' -> wf Σ' -> - wf_universe Σ u -> wf_universe (Σ', Σ.2) u. -Proof. - destruct Σ as [Σ univ]; cbn. - intros [Σ'' eq] wf. - destruct u; simpl; auto. - intros Hl. - intros l inl; specialize (Hl l inl). - cbn. rewrite eq /=. - unfold global_ext_levels. - eapply LevelSet.union_spec; simpl. - apply LevelSet.union_spec in Hl as [Hl|Hl]; cbn in Hl. - - simpl. simpl in Hl. now left. - - right. rewrite global_levels_ext. - eapply LevelSet.union_spec. right. - apply Hl. -Qed. - - -Ltac typing_my_rename_hyp h th := - match th with - | (wf ?E) => fresh "wf" E - | (typing _ _ ?t _) => fresh "type" t - | (@cumul _ _ _ ?t _) => fresh "cumul" t - | (conv _ _ ?t _) => fresh "conv" t - | (All_local_env (lift_typing (@typing _) _) ?G) => fresh "wf" G - | (All_local_env (lift_typing (@typing _) _) _) => fresh "wf" - | (All_local_env _ _ ?G) => fresh "H" G - | context [typing _ _ (_ ?t) _] => fresh "IH" t - end. - -Ltac my_rename_hyp h th := - match th with - | (extends ?t _) => fresh "ext" t - | (extends ?t.1 _) => fresh "ext" t - | (extends _ _) => fresh "ext" - | _ => typing_my_rename_hyp h th - end. - -Ltac rename_hyp h ht ::= my_rename_hyp h ht. - -Lemma weakening_env `{checker_flags} : - env_prop (fun Σ Γ t T => - forall Σ', wf Σ' -> extends Σ.1 Σ' -> (Σ', Σ.2) ;;; Γ |- t : T). -Proof. - apply typing_ind_env; intros; - rename_all_hyps; try solve [econstructor; eauto 2 with extends]. - - - econstructor; eauto 2 with extends. - intuition auto. - destruct H1 as [l' [inl' eq]]. right; right. - exists l'. split; eauto 2 with extends. - - econstructor; eauto 2 with extends. - destruct Σ as [Σ φ]. - clear typet heq_isApp forall_Σ' hneq_l. - induction X1. constructor. econstructor; eauto with extends. - eapply weakening_env_cumul in cumul; eauto. - - econstructor; eauto 2 with extends. - close_Forall. intros; intuition eauto with extends. - - econstructor; eauto with extends. - eapply (All_impl X0); intuition eauto. - destruct X2 as [s Hs]; exists s; intuition eauto. - eapply All_impl; eauto; simpl; intuition eauto with extends. - - econstructor; eauto with extends. auto. - eapply (All_impl X0); intuition eauto. - destruct X2 as [s Hs]; exists s; intuition eauto. - eapply All_impl; eauto; simpl; intuition eauto with extends. - - econstructor. eauto. - destruct X2 as [isB|[u [Hu Ps]]]. - + left; auto. destruct isB. destruct x as [ctx [u [Heq Hu]]]. - exists ctx, u. split; eauto with extends. - + right. exists u. eapply Ps; auto. - + destruct Σ as [Σ φ]. eapply weakening_env_cumul in cumulA; eauto. -Qed. - -Definition weaken_env_prop `{checker_flags} - (P : global_env_ext -> context -> term -> option term -> Type) := - forall Σ Σ' φ, wf Σ' -> extends Σ Σ' -> forall Γ t T, P (Σ, φ) Γ t T -> P (Σ', φ) Γ t T. - -Lemma weakening_on_global_decl `{checker_flags} P Σ Σ' φ kn decl : - weaken_env_prop P -> - wf Σ' -> extends Σ Σ' -> - on_global_decl P (Σ, φ) kn decl -> - on_global_decl P (Σ', φ) kn decl. -Proof. - unfold weaken_env_prop. - intros HPΣ wfΣ' Hext Hdecl. - destruct decl. destruct c. destruct cst_body. simpl in *. - red in Hdecl |- *. simpl in *. - eapply HPΣ; eauto. - eapply HPΣ; eauto. - simpl in *. - destruct Hdecl as [onI onP onnP]; constructor; eauto. - - eapply Alli_impl; eauto. intros. - destruct X. unshelve econstructor; eauto. - -- unfold on_type in *; intuition eauto. - -- unfold on_constructors in *. eapply All2_impl; eauto. - intros decl cs []. unshelve econstructor; eauto. - red in on_ctype |- *. eauto. - clear on_cindices cstr_eq cstr_args_length cstr_concl_head. - revert on_cargs. - generalize (cshape_sorts cs). - induction (cshape_args cs); destruct l; simpl in *; auto. - ** destruct a as [na [b|] ty]; simpl in *; intuition eauto. - ** destruct a as [na [b|] ty]; simpl in *; intuition eauto. - ** clear on_ctype on_cargs. - revert on_cindices. - generalize (List.rev (LiftSubst.lift_context #|cshape_args cs| 0 ind_indices)). - generalize (cshape_indices cs). induction 1; constructor; eauto. - ** simpl. - intros v indv. specialize (on_ctype_variance v indv). - simpl in *. move: on_ctype_variance. - unfold cstr_respects_variance. destruct variance_universes as [[[univs u] u']|]; auto. - intros [args idxs]. split. - eapply (context_relation_impl args); intros. - inversion X; constructor; auto. - eapply weakening_env_cumul; eauto. - eapply weakening_env_conv; eauto. - eapply weakening_env_cumul; eauto. - eapply (All2_impl idxs); intros. - eapply weakening_env_conv; eauto. - -- unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. - split; [apply fst in ind_sorts|apply snd in ind_sorts]. - eapply Forall_impl; tea; cbn. - intros. eapply Forall_impl; eauto; simpl. - intros; eapply leq_universe_subset; tea. - apply weakening_env_global_ext_constraints; tea. - destruct indices_matter; [|trivial]. clear -ind_sorts HPΣ wfΣ' Hext. - induction ind_indices; simpl in *; auto. - ** eapply (extends_wf_universe (Σ:=(Σ,φ)) Σ'); auto. - ** destruct a as [na [b|] ty]; simpl in *; intuition eauto. - -- intros v onv. - move: (onIndices v onv). unfold ind_respects_variance. - destruct variance_universes as [[[univs u] u']|] => //. - intros idx; eapply (context_relation_impl idx); simpl. - intros par par' t t'. - intros d. inv d; constructor; auto. - eapply weakening_env_cumul; eauto. - eapply weakening_env_conv; eauto. - eapply weakening_env_cumul; eauto. - - red in onP |- *. eapply All_local_env_impl; eauto. -Qed. - -Lemma weakening_env_lookup_on_global_env `{checker_flags} P Σ Σ' c decl : - weaken_env_prop P -> - wf Σ' -> extends Σ Σ' -> on_global_env P Σ -> - lookup_env Σ c = Some decl -> - on_global_decl P (Σ', universes_decl_of_decl decl) c decl. -Proof. - intros HP wfΣ Hext HΣ. - induction HΣ; simpl. congruence. - assert (HH: extends Σ Σ'). { - destruct Hext as [Σ'' HΣ'']. - exists (Σ'' ++ [(kn, d)]). now rewrite <- app_assoc. } - unfold eq_kername; destruct kername_eq_dec; subst. - - intros [= ->]. - clear Hext; eapply weakening_on_global_decl; eauto. - - now apply IHHΣ. -Qed. - -Lemma weaken_lookup_on_global_env `{checker_flags} P Σ c decl : - weaken_env_prop P -> - wf Σ -> on_global_env P Σ -> - lookup_env Σ c = Some decl -> - on_global_decl P (Σ, universes_decl_of_decl decl) c decl. -Proof. - intros. eapply weakening_env_lookup_on_global_env; eauto. - exists []; simpl; destruct Σ; eauto. -Qed. - -Lemma declared_minductive_inv `{checker_flags} {Σ P ind mdecl} : - weaken_env_prop (lift_typing P) -> - wf Σ -> Forall_decls_typing P Σ -> - declared_minductive Σ ind mdecl -> - on_inductive (lift_typing P) (Σ, ind_universes mdecl) ind mdecl. -Proof. - intros. - eapply weaken_lookup_on_global_env in X1; eauto. apply X1. -Qed. - -Lemma declared_inductive_inv `{checker_flags} {Σ P ind mdecl idecl} : - weaken_env_prop (lift_typing P) -> - wf Σ -> Forall_decls_typing P Σ -> - declared_inductive Σ mdecl ind idecl -> - on_ind_body (lift_typing P) (Σ, ind_universes mdecl) (inductive_mind ind) mdecl (inductive_ind ind) idecl. -Proof. - intros. - destruct H0 as [Hmdecl Hidecl]. - eapply declared_minductive_inv in Hmdecl; eauto. - apply onInductives in Hmdecl. - eapply nth_error_alli in Hidecl; eauto. - apply Hidecl. -Qed. - -Lemma declared_constructor_inv `{checker_flags} {Σ P mdecl idecl ref cdecl} - (HP : weaken_env_prop (lift_typing P)) - (wfΣ : wf Σ) - (HΣ : Forall_decls_typing P Σ) - (Hdecl : declared_constructor Σ mdecl idecl ref cdecl) : - ∑ cs, - let onib := declared_inductive_inv HP wfΣ HΣ (proj1 Hdecl) in - nth_error onib.(ind_cshapes) ref.2 = Some cs - × on_constructor (lift_typing P) (Σ, ind_universes mdecl) mdecl - (inductive_ind ref.1) idecl onib.(ind_indices) cdecl cs. -Proof. - intros. - destruct Hdecl as [Hidecl Hcdecl]. - set (declared_inductive_inv HP wfΣ HΣ (proj1 (conj Hidecl Hcdecl))) as HH. - clearbody HH. pose proof HH.(onConstructors) as HH'. - eapply All2_nth_error_Some in Hcdecl; tea. -Qed. - -Lemma declared_projection_inv `{checker_flags} {Σ P mdecl idecl ref pdecl} : - forall (HP : weaken_env_prop (lift_typing P)) - (wfΣ : wf Σ) - (HΣ : Forall_decls_typing P Σ) - (Hdecl : declared_projection Σ mdecl idecl ref pdecl), - let oib := declared_inductive_inv HP wfΣ HΣ (proj1 Hdecl) in - match oib.(ind_cshapes) return Type with - | [cs] => on_projection mdecl (inductive_mind ref.1.1) (inductive_ind ref.1.1) cs (snd ref) pdecl - | _ => False - end. -Proof. - intros. - destruct (declared_inductive_inv HP wfΣ HΣ (proj1 Hdecl)) in *. - destruct Hdecl as [Hidecl [Hcdecl Hnpar]]. simpl. clearbody oib. - forward onProjections. - { eapply nth_error_Some_length in Hcdecl. - destruct (ind_projs idecl); simpl in *. lia. congruence. } - destruct ind_cshapes as [|? []]; try contradiction. - destruct onProjections. - eapply nth_error_alli in Hcdecl; eauto. eapply Hcdecl. -Qed. - -Lemma wf_extends `{checker_flags} {Σ Σ'} : wf Σ' -> extends Σ Σ' -> wf Σ. -Proof. - intros HΣ' [Σ'' ->]. simpl in *. - induction Σ''. auto. - inv HΣ'. auto. -Qed. - -Lemma weaken_env_prop_typing `{checker_flags} : weaken_env_prop (lift_typing typing). -Proof. - red. intros * wfΣ' Hext *. - destruct T; simpl. - intros Ht. pose proof (wf_extends wfΣ' Hext). - eapply (weakening_env (_, _)); eauto. eapply typing_wf_local in Ht; eauto. - intros [s Ht]. pose proof (wf_extends wfΣ' Hext). exists s. - eapply (weakening_env (_, _)); eauto. eapply typing_wf_local in Ht; eauto. -Qed. - -Hint Unfold weaken_env_prop : pcuic. - -Lemma on_declared_minductive `{checker_flags} {Σ ref decl} : - wf Σ -> - declared_minductive Σ ref decl -> - on_inductive (lift_typing typing) (Σ, ind_universes decl) ref decl. -Proof. - intros wfΣ Hdecl. - apply (declared_minductive_inv weaken_env_prop_typing wfΣ wfΣ Hdecl). -Qed. - -Lemma on_declared_inductive `{checker_flags} {Σ ref mdecl idecl} : - wf Σ -> - declared_inductive Σ mdecl ref idecl -> - on_inductive (lift_typing typing) (Σ, ind_universes mdecl) (inductive_mind ref) mdecl * - on_ind_body (lift_typing typing) (Σ, ind_universes mdecl) (inductive_mind ref) mdecl (inductive_ind ref) idecl. -Proof. - intros wfΣ Hdecl. - split. destruct Hdecl as [Hmdecl _]. now apply on_declared_minductive in Hmdecl. - apply (declared_inductive_inv weaken_env_prop_typing wfΣ wfΣ Hdecl). -Qed. - -Lemma on_declared_constructor `{checker_flags} {Σ ref mdecl idecl cdecl} - (wfΣ : wf Σ) - (Hdecl : declared_constructor Σ mdecl idecl ref cdecl) : - on_inductive (lift_typing typing) (Σ, ind_universes mdecl) - (inductive_mind (fst ref)) mdecl * - on_ind_body (lift_typing typing) (Σ, ind_universes mdecl) - (inductive_mind (fst ref)) mdecl (inductive_ind (fst ref)) idecl * - ∑ ind_ctor_sort, - let onib := declared_inductive_inv weaken_env_prop_typing wfΣ wfΣ (proj1 Hdecl) in - nth_error (ind_cshapes onib) ref.2 = Some ind_ctor_sort - × on_constructor (lift_typing typing) (Σ, ind_universes mdecl) - mdecl (inductive_ind (fst ref)) - idecl onib.(ind_indices) cdecl ind_ctor_sort. -Proof. - split. destruct Hdecl as [Hidecl Hcdecl]. - now apply on_declared_inductive in Hidecl. - apply (declared_constructor_inv weaken_env_prop_typing wfΣ wfΣ Hdecl). -Qed. - -Lemma on_declared_projection `{checker_flags} {Σ ref mdecl idecl pdecl} : - forall (wfΣ : wf Σ) (Hdecl : declared_projection Σ mdecl idecl ref pdecl), - on_inductive (lift_typing typing) (Σ, ind_universes mdecl) (inductive_mind (fst (fst ref))) mdecl * - let oib := declared_inductive_inv weaken_env_prop_typing wfΣ wfΣ (proj1 Hdecl) in - match oib.(ind_cshapes) return Type with - | [cs] => on_projection mdecl (inductive_mind ref.1.1) (inductive_ind ref.1.1) cs (snd ref) pdecl - | _ => False - end. -Proof. - intros wfΣ Hdecl. - split. destruct Hdecl as [Hidecl Hcdecl]. now apply on_declared_inductive in Hidecl. - apply (declared_projection_inv weaken_env_prop_typing wfΣ wfΣ Hdecl). -Qed. diff --git a/checker/update_plugin.sh b/checker/update_plugin.sh deleted file mode 100755 index 305d6dcb4..000000000 --- a/checker/update_plugin.sh +++ /dev/null @@ -1,12 +0,0 @@ -#!/bin/bash -TEMPLATE_LIB=../template-coq - -if [[ "src" -ot "${TEMPLATE_LIB}/gen-src" || ! -f "src/metacoq_checker_plugin.cmxa" \ - || "src/metacoq_checker_plugin.cmxa" -ot "gen-src/Extract.v" ]] -then - echo "Renaming extracted files" - cp -r ${TEMPLATE_LIB}/gen-src/to-lower.sh src - (cd src; ./to-lower.sh) -else - echo "Extracted code is up-to-date" -fi \ No newline at end of file diff --git a/configure.sh b/configure.sh index 4666804ec..4a5e4579b 100755 --- a/configure.sh +++ b/configure.sh @@ -6,8 +6,6 @@ make -f Makefile mrproper # Dependencies for local or global builds. # When building the packages separately, dependencies are not set as everything # should already be available in $(COQMF_LIB)/user-contrib/MetaCoq/* -# checker is treated specially: due to code generation, we rebuild the template-coq module locally -# when building the checker package # For local builds, we set specific dependencies of each subproject in */metacoq-config # CWD=`pwd` @@ -16,42 +14,32 @@ if command -v coqc >/dev/null 2>&1 then COQLIB=` coqc -where | tr -d '\r' | tr '\\\\' '/'` - if [ "$1" == "local" ] - then + if [ "$1" = "local" ]; then echo "Building MetaCoq locally" - CHECKER_DEPS="-R ../template-coq/theories MetaCoq.Template -I ../template-coq/build" - PCUIC_MLDEPS="-I ../checker/src" - PCUIC_VDEPS="-R ../checker/theories MetaCoq.Checker" - SAFECHECKER_DEPS="${PCUIC_VDEPS} -R ../pcuic/theories MetaCoq.PCUIC" + PCUIC_DEPS="-I ../template-coq/build -R ../template-coq/theories MetaCoq.Template" + SAFECHECKER_DEPS="-R ../pcuic/theories MetaCoq.PCUIC" ERASURE_DEPS="-R ../safechecker/theories MetaCoq.SafeChecker" - TRANSLATIONS_DEPS="${PCUIC_VDEPS}" + TRANSLATIONS_DEPS="" else echo "Building MetaCoq globally (default)" - # To find the metacoq template plugin - CHECKER_DEPS="-I ${COQLIB}/user-contrib/MetaCoq/Template" - # The pcuic plugin depends on the checker plugin - # The safechecker and erasure plugins are self-contained + # The safechecker and erasure plugins depend on the extractable template-coq plugin # These dependencies should not be necessary when separate linking of ocaml object # files is supported by coq_makefile - PCUIC_MLDEPS="-I ${COQLIB}/user-contrib/MetaCoq/Checker" - PCUIC_VDEPS="" + PCUIC_DEPS="-I ${COQLIB}/user-contrib/MetaCoq/Template" SAFECHECKER_DEPS="" ERASURE_DEPS="" TRANSLATIONS_DEPS="" fi - echo "# DO NOT EDIT THIS FILE: autogenerated from ./configure.sh" > checker/metacoq-config echo "# DO NOT EDIT THIS FILE: autogenerated from ./configure.sh" > pcuic/metacoq-config echo "# DO NOT EDIT THIS FILE: autogenerated from ./configure.sh" > safechecker/metacoq-config echo "# DO NOT EDIT THIS FILE: autogenerated from ./configure.sh" > erasure/metacoq-config echo "# DO NOT EDIT THIS FILE: autogenerated from ./configure.sh" > translations/metacoq-config - echo ${CHECKER_DEPS} >> checker/metacoq-config - # echo "OCAMLPATH = \"${CWD}/template-coq\"" >> checker/metacoq-config - echo ${CHECKER_DEPS} ${PCUIC_MLDEPS} ${PCUIC_VDEPS} >> pcuic/metacoq-config - echo ${CHECKER_DEPS} ${SAFECHECKER_DEPS} >> safechecker/metacoq-config - echo ${CHECKER_DEPS} ${SAFECHECKER_DEPS} ${ERASURE_DEPS} >> erasure/metacoq-config - echo ${CHECKER_DEPS} ${TRANSLATIONS_DEPS} >> translations/metacoq-config + echo ${PCUIC_DEPS} >> pcuic/metacoq-config + echo ${PCUIC_DEPS} ${SAFECHECKER_DEPS} >> safechecker/metacoq-config + echo ${PCUIC_DEPS} ${SAFECHECKER_DEPS} ${ERASURE_DEPS} >> erasure/metacoq-config + echo ${PCUIC_DEPS} ${TRANSLATIONS_DEPS} >> translations/metacoq-config else echo "Error: coqc not found in path" fi diff --git a/coq-metacoq-checker.opam b/coq-metacoq-checker.opam deleted file mode 100644 index dbbcf1c6f..000000000 --- a/coq-metacoq-checker.opam +++ /dev/null @@ -1,39 +0,0 @@ -opam-version: "2.0" -maintainer: "matthieu.sozeau@inria.fr" -homepage: "https://metacoq.github.io/metacoq" -dev-repo: "git+https://github.com/MetaCoq/metacoq.git#coq-8.11" -bug-reports: "https://github.com/MetaCoq/metacoq/issues" -authors: ["Abhishek Anand " - "Simon Boulier " - "Cyril Cohen " - "Yannick Forster " - "Fabian Kunze " - "Gregory Malecha " - "Matthieu Sozeau " - "Nicolas Tabareau " - "Théo Winterhalter " -] -license: "MIT" -build: [ - ["sh" "./configure.sh"] - [make "-j" "%{jobs}%" "checker"] -] -install: [ - [make "-C" "checker" "install"] -] -depends: [ - "ocaml" {>= "4.07.1" & < "4.12~"} - "coq" { >= "8.12~" & < "8.13~" } - "coq-equations" {= "1.2.3+8.12" } - "coq-metacoq-template" {= version} -] -synopsis: "Specification of Coq's type theory and reference checker implementation" -description: """ -MetaCoq is a meta-programming framework for Coq. - -The Checker module provides a complete specification of Coq's typing and conversion -relation along with a reference type-checker that is extracted to a pluging. - -This provides a command: `MetaCoq Check [global_reference]` that can be used -to typecheck a Coq definition using the verified type-checker. -""" diff --git a/coq-metacoq-erasure.opam b/coq-metacoq-erasure.opam index b54f733f6..d72cb7459 100644 --- a/coq-metacoq-erasure.opam +++ b/coq-metacoq-erasure.opam @@ -25,7 +25,6 @@ depends: [ "ocaml" {>= "4.07.1" & < "4.12~"} "coq" { >= "8.12~" & < "8.13~" } "coq-metacoq-template" {= version} - "coq-metacoq-checker" {= version} "coq-metacoq-pcuic" {= version} "coq-metacoq-safechecker" {= version} ] diff --git a/coq-metacoq-pcuic.opam b/coq-metacoq-pcuic.opam index 87a0ab2a2..1ea9f0cc0 100644 --- a/coq-metacoq-pcuic.opam +++ b/coq-metacoq-pcuic.opam @@ -26,7 +26,6 @@ depends: [ "coq" { >= "8.12~" & < "8.13~" } "coq-equations" {= "1.2.3+8.12" } "coq-metacoq-template" {= version} - "coq-metacoq-checker" {= version} ] synopsis: "A type system equivalent to Coq's and its metatheory" description: """ diff --git a/coq-metacoq-safechecker.opam b/coq-metacoq-safechecker.opam index febce6520..d9a012586 100644 --- a/coq-metacoq-safechecker.opam +++ b/coq-metacoq-safechecker.opam @@ -25,7 +25,6 @@ depends: [ "ocaml" {>= "4.07.1" & < "4.12~"} "coq" { >= "8.12~" & < "8.13~" } "coq-metacoq-template" {= version} - "coq-metacoq-checker" {= version} "coq-metacoq-pcuic" {= version} ] synopsis: "Implementation and verification of safe conversion and typechecking algorithms for Coq" diff --git a/coq-metacoq-translations.opam b/coq-metacoq-translations.opam index 70ea8159b..22037edea 100644 --- a/coq-metacoq-translations.opam +++ b/coq-metacoq-translations.opam @@ -21,7 +21,6 @@ depends: [ "ocaml" {>= "4.07.1" & < "4.12~"} "coq" { >= "8.12~" & < "8.13~" } "coq-metacoq-template" {= version} - "coq-metacoq-checker" {= version} ] synopsis: "Translations built on top of MetaCoq" description: """ diff --git a/coq-metacoq.opam b/coq-metacoq.opam index 8b6a049e9..fccde589e 100644 --- a/coq-metacoq.opam +++ b/coq-metacoq.opam @@ -18,7 +18,6 @@ depends: [ "ocaml" {>= "4.07.1" & < "4.12~"} "coq" { >= "8.12~" & < "8.13~" } "coq-metacoq-template" {= version} - "coq-metacoq-checker" {= version} "coq-metacoq-pcuic" {= version} "coq-metacoq-safechecker" {= version} "coq-metacoq-erasure" {= version} diff --git a/dependency-graph/A tour of metacoq.png b/dependency-graph/A tour of metacoq.png new file mode 100644 index 000000000..bee4f97f0 Binary files /dev/null and b/dependency-graph/A tour of metacoq.png differ diff --git a/dependency-graph/dependency-graph.md b/dependency-graph/dependency-graph.md index 4dfabf0d7..a5332c4f8 100644 --- a/dependency-graph/dependency-graph.md +++ b/dependency-graph/dependency-graph.md @@ -11,7 +11,7 @@ cd dependency-graph # Generate the dependency graph between files by hand (obsolete) -1. In each folder (template-coq + checker + pcuic + safechecker + erasure), generate the dot file with: +1. In each folder (template-coq + pcuic + safechecker + erasure), generate the dot file with: ``` coqdep -f _CoqProject -dumpgraph plop.dot > /dev/null ``` diff --git a/dependency-graph/depgraph-2020-11-28.dot b/dependency-graph/depgraph-2020-11-28.dot new file mode 100644 index 000000000..c103181dd --- /dev/null +++ b/dependency-graph/depgraph-2020-11-28.dot @@ -0,0 +1,346 @@ +digraph dependencies { +node[style=filled] +"safechecker/Extraction"[label="Extraction", color=paleturquoise1] +"safechecker/PCUICSafeRetyping"[label="PCUICSafeRetyping", color=paleturquoise1] +"safechecker/SafeTemplateChecker"[label="SafeTemplateChecker", color=paleturquoise1] +"safechecker/PCUICSafeChecker"[label="PCUICSafeChecker", color=paleturquoise1] +"safechecker/PCUICTypeChecker"[label="PCUICTypeChecker", color=paleturquoise1] +"safechecker/PCUICWfEnv"[label="PCUICWfEnv", color=paleturquoise1] +"safechecker/PCUICWfReduction"[label="PCUICWfReduction", color=paleturquoise1] +"safechecker/PCUICSafeConversion"[label="PCUICSafeConversion", color=paleturquoise1] +"safechecker/PCUICEqualityDec"[label="PCUICEqualityDec", color=paleturquoise1] +"safechecker/PCUICSafeReduce"[label="PCUICSafeReduce", color=paleturquoise1] +"safechecker/PCUICErrors"[label="PCUICErrors", color=paleturquoise1] +"pcuic/PCUICValidity" -> "pcuic/PCUICAlpha" +"pcuic/PCUICWfUniverses" -> "pcuic/PCUICArities" +"template-coq/EnvironmentTyping" -> "pcuic/PCUICAst" +"pcuic/PCUICSize" -> "pcuic/PCUICAstUtils" +"template-coq/Reflect" -> "pcuic/PCUICAstUtils" +"template-coq/common/uGraph" -> "pcuic/PCUICAstUtils" +"pcuic/PCUICClosed" -> "pcuic/PCUICCSubst" +"pcuic/PCUICElimination" -> "pcuic/PCUICCanonicity" +"pcuic/PCUICSN" -> "pcuic/PCUICCanonicity" +"pcuic/PCUICWcbvEval" -> "pcuic/PCUICCanonicity" +"pcuic/PCUICWeakeningEnv" -> "pcuic/PCUICClosed" +"pcuic/PCUICParallelReductionConfluence" -> "pcuic/PCUICConfluence" +"pcuic/PCUICConfluence" -> "pcuic/PCUICContextConversion" +"pcuic/PCUICContextRelation" -> "pcuic/PCUICContextConversion" +"pcuic/PCUICTyping" -> "pcuic/PCUICContextRelation" +"pcuic/PCUICCtxShape" -> "pcuic/PCUICContexts" +"pcuic/PCUICCumulProp" -> "pcuic/PCUICConvCumInversion" +"pcuic/PCUICContextConversion" -> "pcuic/PCUICConversion" +"pcuic/PCUICInversion" -> "pcuic/PCUICCtxShape" +"pcuic/PCUICNameless" -> "pcuic/PCUICCtxShape" +"template-coq/UnivSubst" -> "pcuic/PCUICCtxShape" +"pcuic/PCUICSafeLemmata" -> "pcuic/PCUICCumulProp" +"pcuic/PCUICReduction" -> "pcuic/PCUICCumulativity" +"pcuic/PCUICCumulProp" -> "pcuic/PCUICElimination" +"pcuic/PCUICLiftSubst" -> "pcuic/PCUICEquality" +"pcuic/PCUICReflect" -> "pcuic/PCUICEquality" +"pcuic/PCUICTyping" -> "pcuic/PCUICGeneration" +"pcuic/PCUICTyping" -> "pcuic/PCUICGlobalEnv" +"pcuic/PCUICAstUtils" -> "pcuic/PCUICInduction" +"pcuic/PCUICValidity" -> "pcuic/PCUICInductiveInversion" +"pcuic/PCUICSpine" -> "pcuic/PCUICInductives" +"pcuic/PCUICConversion" -> "pcuic/PCUICInversion" +"pcuic/PCUICInduction" -> "pcuic/PCUICLiftSubst" +"pcuic/PCUICTyping" -> "pcuic/PCUICNameless" +"pcuic/PCUICContextRelation" -> "pcuic/PCUICNormal" +"template-coq/UnivSubst" -> "pcuic/PCUICNormal" +"pcuic/PCUICSubstitution" -> "pcuic/PCUICParallelReduction" +"pcuic/PCUICParallelReduction" -> "pcuic/PCUICParallelReductionConfluence" +"pcuic/PCUICEquality" -> "pcuic/PCUICPosition" +"pcuic/PCUICLiftSubst" -> "pcuic/PCUICPretty" +"pcuic/PCUICCumulProp" -> "pcuic/PCUICPrincipality" +"pcuic/PCUICPosition" -> "pcuic/PCUICReduction" +"pcuic/PCUICRelations" -> "pcuic/PCUICReduction" +"pcuic/PCUICUnivSubst" -> "pcuic/PCUICReduction" +"pcuic/PCUICInduction" -> "pcuic/PCUICReflect" +"pcuic/PCUICSafeLemmata" -> "pcuic/PCUICSN" +"pcuic/PCUICAlpha" -> "pcuic/PCUICSR" +"pcuic/PCUICInductiveInversion" -> "pcuic/PCUICSR" +"pcuic/PCUICNormal" -> "pcuic/PCUICSafeLemmata" +"pcuic/PCUICSR" -> "pcuic/PCUICSafeLemmata" +"pcuic/PCUICWeakening" -> "pcuic/PCUICSigmaCalculus" +"pcuic/PCUICAst" -> "pcuic/PCUICSize" +"pcuic/PCUICArities" -> "pcuic/PCUICSpine" +"pcuic/PCUICSigmaCalculus" -> "pcuic/PCUICSubstitution" +"pcuic/PCUICUnivSubstitution" -> "pcuic/PCUICSubstitution" +"pcuic/PCUICCumulativity" -> "pcuic/PCUICTyping" +"pcuic/PCUICUtils" -> "pcuic/PCUICTyping" +"template-coq/utils/LibHypsNaming" -> "pcuic/PCUICTyping" +"pcuic/PCUICLiftSubst" -> "pcuic/PCUICUnivSubst" +"pcuic/PCUICWeakening" -> "pcuic/PCUICUnivSubstitution" +"pcuic/PCUICAstUtils" -> "pcuic/PCUICUtils" +"pcuic/PCUICInductives" -> "pcuic/PCUICValidity" +"pcuic/PCUICCSubst" -> "pcuic/PCUICWcbvEval" +"pcuic/PCUICSR" -> "pcuic/PCUICWcbvEval" +"pcuic/PCUICClosed" -> "pcuic/PCUICWeakening" +"pcuic/PCUICGeneration" -> "pcuic/PCUICWeakening" +"pcuic/PCUICTyping" -> "pcuic/PCUICWeakeningEnv" +"pcuic/PCUICContexts" -> "pcuic/PCUICWfUniverses" +"pcuic/PCUICAst" -> "pcuic/TemplateToPCUIC" +"template-coq/Environment" -> "template-coq/Ast" +"template-coq/Ast" -> "template-coq/AstUtils" +"template-coq/utils" -> "template-coq/BasicAst" +"template-coq/Universes" -> "template-coq/Environment" +"template-coq/AstUtils" -> "template-coq/EnvironmentTyping" +"template-coq/AstUtils" -> "template-coq/Induction" +"template-coq/Induction" -> "template-coq/LiftSubst" +"template-coq/Induction" -> "template-coq/Reflect" +"template-coq/EnvironmentTyping" -> "template-coq/Typing" +"template-coq/UnivSubst" -> "template-coq/Typing" +"template-coq/LiftSubst" -> "template-coq/UnivSubst" +"template-coq/BasicAst" -> "template-coq/Universes" +"template-coq/config" -> "template-coq/Universes" +"template-coq/Universes" -> "template-coq/common/uGraph" +"template-coq/utils/wGraph" -> "template-coq/common/uGraph" +"template-coq/utils/All_Forall" -> "template-coq/monad_utils" +"template-coq/monad_utils" -> "template-coq/utils" +"template-coq/utils/MCUtils" -> "template-coq/utils" +"template-coq/utils/MCOption" -> "template-coq/utils/All_Forall" +"template-coq/utils/MCSquash" -> "template-coq/utils/All_Forall" +"template-coq/utils/MCPrelude" -> "template-coq/utils/MCList" +"template-coq/utils/MCRelations" -> "template-coq/utils/MCList" +"template-coq/utils/MCList" -> "template-coq/utils/MCOption" +"template-coq/utils/MCProd" -> "template-coq/utils/MCOption" +"template-coq/utils/MCCompare" -> "template-coq/utils/MCString" +"template-coq/utils/All_Forall" -> "template-coq/utils/MCUtils" +"template-coq/utils/MCArith" -> "template-coq/utils/MCUtils" +"template-coq/utils/MCEquality" -> "template-coq/utils/MCUtils" +"template-coq/utils/MCString" -> "template-coq/utils/MCUtils" +"template-coq/utils/MCUtils" -> "template-coq/utils/wGraph" +"safechecker/SafeTemplateChecker" -> "safechecker/Extraction" +"pcuic/PCUICGlobalEnv" -> "safechecker/PCUICEqualityDec" +"pcuic/PCUICPretty" -> "safechecker/PCUICErrors" +"safechecker/PCUICTypeChecker" -> "safechecker/PCUICSafeChecker" +"pcuic/PCUICConvCumInversion" -> "safechecker/PCUICSafeConversion" +"pcuic/PCUICPrincipality" -> "safechecker/PCUICSafeConversion" +"safechecker/PCUICEqualityDec" -> "safechecker/PCUICSafeConversion" +"safechecker/PCUICSafeReduce" -> "safechecker/PCUICSafeConversion" +"pcuic/PCUICCanonicity" -> "safechecker/PCUICSafeReduce" +"safechecker/PCUICErrors" -> "safechecker/PCUICSafeReduce" +"safechecker/PCUICSafeReduce" -> "safechecker/PCUICSafeRetyping" +"safechecker/PCUICWfEnv" -> "safechecker/PCUICTypeChecker" +"safechecker/PCUICWfReduction" -> "safechecker/PCUICTypeChecker" +"pcuic/PCUICGlobalEnv" -> "safechecker/PCUICWfEnv" +"safechecker/PCUICSafeConversion" -> "safechecker/PCUICWfReduction" +"pcuic/TemplateToPCUIC" -> "safechecker/SafeTemplateChecker" +"template-coq/Typing" -> "safechecker/SafeTemplateChecker" +"safechecker/PCUICSafeChecker" -> "safechecker/SafeTemplateChecker" +"template-coq/Extraction"[label="Extraction", color=aquamarine] +"template-coq/Constants"[label="Constants", color=aquamarine] +"template-coq/monad_utils"[label="monad_utils", color=aquamarine] +"template-coq/TemplateMonad/Extractable"[label="Extractable", color=aquamarine] +"template-coq/TemplateMonad/Core"[label="Core", color=aquamarine] +"template-coq/TemplateMonad/Common"[label="Common", color=aquamarine] +"template-coq/TemplateMonad"[label="TemplateMonad", color=aquamarine] +"template-coq/TypingWf"[label="TypingWf", color=aquamarine] +"template-coq/Typing"[label="Typing", color=aquamarine] +"template-coq/WfInv"[label="WfInv", color=aquamarine] +"template-coq/EnvironmentTyping"[label="EnvironmentTyping", color=aquamarine] +"template-coq/Pretty"[label="Pretty", color=aquamarine] +"template-coq/UnivSubst"[label="UnivSubst", color=aquamarine] +"template-coq/LiftSubst"[label="LiftSubst", color=aquamarine] +"template-coq/Induction"[label="Induction", color=aquamarine] +"template-coq/Kernames"[label="Kernames", color=aquamarine] +"template-coq/Reflect"[label="Reflect", color=aquamarine] +"template-coq/AstUtils"[label="AstUtils", color=aquamarine] +"template-coq/Ast"[label="Ast", color=aquamarine] +"template-coq/Environment"[label="Environment", color=aquamarine] +"template-coq/BasicAst"[label="BasicAst", color=aquamarine] +"template-coq/Universes"[label="Universes", color=aquamarine] +"template-coq/config"[label="config", color=aquamarine] +"template-coq/utils"[label="utils", color=aquamarine] +"template-coq/common/uGraph"[label="uGraph", color=aquamarine] +"template-coq/utils/MCUtils"[label="MCUtils", color=aquamarine] +"template-coq/utils/wGraph"[label="wGraph", color=aquamarine] +"template-coq/utils/MCString"[label="MCString", color=aquamarine] +"template-coq/utils/MCSquash"[label="MCSquash", color=aquamarine] +"template-coq/utils/MCRelations"[label="MCRelations", color=aquamarine] +"template-coq/utils/MCProd"[label="MCProd", color=aquamarine] +"template-coq/utils/MCOption"[label="MCOption", color=aquamarine] +"template-coq/utils/MCList"[label="MCList", color=aquamarine] +"template-coq/utils/LibHypsNaming"[label="LibHypsNaming", color=aquamarine] +"template-coq/utils/MCEquality"[label="MCEquality", color=aquamarine] +"template-coq/utils/MCCompare"[label="MCCompare", color=aquamarine] +"template-coq/utils/MCArith"[label="MCArith", color=aquamarine] +"template-coq/utils/All_Forall"[label="All_Forall", color=aquamarine] +"template-coq/utils/MCPrelude"[label="MCPrelude", color=aquamarine] +"template-coq/TemplateMonad" -> "template-coq/Constants" +"template-coq/TemplateMonad/Extractable" -> "template-coq/Constants" +"template-coq/common/uGraph" -> "template-coq/Constants" +"template-coq/Pretty" -> "template-coq/Extraction" +"template-coq/Reflect" -> "template-coq/Extraction" +"template-coq/TemplateMonad/Extractable" -> "template-coq/Extraction" +"template-coq/UnivSubst" -> "template-coq/Extraction" +"template-coq/AstUtils" -> "template-coq/Kernames" +"template-coq/LiftSubst" -> "template-coq/Pretty" +"template-coq/TemplateMonad/Core" -> "template-coq/TemplateMonad" +"template-coq/Ast" -> "template-coq/TemplateMonad/Common" +"template-coq/AstUtils" -> "template-coq/TemplateMonad/Core" +"template-coq/TemplateMonad/Common" -> "template-coq/TemplateMonad/Core" +"template-coq/AstUtils" -> "template-coq/TemplateMonad/Extractable" +"template-coq/TemplateMonad/Common" -> "template-coq/TemplateMonad/Extractable" +"template-coq/Typing" -> "template-coq/TypingWf" +"template-coq/AstUtils" -> "template-coq/WfInv" +"checker/All"[label="All", color=seagreen3] +"checker/Normal"[label="Normal", color=seagreen3] +"checker/Retyping"[label="Retyping", color=seagreen3] +"checker/WcbvEval"[label="WcbvEval", color=seagreen3] +"checker/Checker"[label="Checker", color=seagreen3] +"checker/Substitution"[label="Substitution", color=seagreen3] +"checker/Weakening"[label="Weakening", color=seagreen3] +"checker/Closed"[label="Closed", color=seagreen3] +"checker/WeakeningEnv"[label="WeakeningEnv", color=seagreen3] +"checker/Generation"[label="Generation", color=seagreen3] +"template-coq/Loader" -> "template-coq/All" +"template-coq/Typing" -> "template-coq/All" +"template-coq/Constants" -> "template-coq/Loader" +"template-coq/All" -> "checker/All" +"checker/Retyping" -> "checker/All" +"checker/Substitution" -> "checker/All" +"template-coq/Typing" -> "checker/Checker" +"template-coq/common/uGraph" -> "checker/Checker" +"template-coq/Reflect" -> "checker/Closed" +"template-coq/TypingWf" -> "checker/Closed" +"checker/WeakeningEnv" -> "checker/Closed" +"template-coq/Reflect" -> "checker/Generation" +"template-coq/Typing" -> "checker/Generation" +"template-coq/Typing" -> "checker/Normal" +"checker/Checker" -> "checker/Retyping" +"checker/Generation" -> "checker/Substitution" +"checker/Weakening" -> "checker/Substitution" +"template-coq/Reflect" -> "checker/WcbvEval" +"template-coq/Typing" -> "checker/WcbvEval" +"template-coq/WfInv" -> "checker/WcbvEval" +"checker/Closed" -> "checker/Weakening" +"template-coq/Typing" -> "checker/WeakeningEnv" +"template-coq/utils/LibHypsNaming" -> "checker/WeakeningEnv" +"erasure/Erasure"[label="Erasure", color=tan] +"erasure/EOptimizePropDiscr"[label="EOptimizePropDiscr", color=tan] +"erasure/ErasureFunction"[label="ErasureFunction", color=tan] +"erasure/ErasureCorrectness"[label="ErasureCorrectness", color=tan] +"erasure/EArities"[label="EArities", color=tan] +"erasure/EInversion"[label="EInversion", color=tan] +"erasure/ESubstitution"[label="ESubstitution", color=tan] +"erasure/Prelim"[label="Prelim", color=tan] +"erasure/Extraction"[label="Extraction", color=tan] +"erasure/EAll"[label="EAll", color=tan] +"erasure/EDeps"[label="EDeps", color=tan] +"erasure/Extract"[label="Extract", color=tan] +"erasure/ETyping"[label="ETyping", color=tan] +"erasure/EWndEval"[label="EWndEval", color=tan] +"erasure/EWcbvEval"[label="EWcbvEval", color=tan] +"erasure/ECSubst"[label="ECSubst", color=tan] +"erasure/EPretty"[label="EPretty", color=tan] +"erasure/EReflect"[label="EReflect", color=tan] +"erasure/ELiftSubst"[label="ELiftSubst", color=tan] +"erasure/EInduction"[label="EInduction", color=tan] +"erasure/EAstUtils"[label="EAstUtils", color=tan] +"erasure/EAst"[label="EAst", color=tan] +"erasure/EWcbvEval" -> "erasure/EAll" +"erasure/EWndEval" -> "erasure/EAll" +"erasure/Extract" -> "erasure/EAll" +"pcuic/PCUICCanonicity" -> "erasure/EArities" +"pcuic/PCUICPrincipality" -> "erasure/EArities" +"erasure/Extract" -> "erasure/EArities" +"template-coq/Universes" -> "erasure/EAst" +"template-coq/Kernames" -> "erasure/EAstUtils" +"erasure/EAst" -> "erasure/EAstUtils" +"erasure/ELiftSubst" -> "erasure/ECSubst" +"erasure/ESubstitution" -> "erasure/EDeps" +"erasure/EAst" -> "erasure/EInduction" +"erasure/Prelim" -> "erasure/EInversion" +"erasure/EInduction" -> "erasure/ELiftSubst" +"erasure/ErasureFunction" -> "erasure/EOptimizePropDiscr" +"erasure/ETyping" -> "erasure/EPretty" +"pcuic/PCUICReflect" -> "erasure/EReflect" +"erasure/EInduction" -> "erasure/EReflect" +"erasure/Prelim" -> "erasure/ESubstitution" +"erasure/EAstUtils" -> "erasure/ETyping" +"erasure/ELiftSubst" -> "erasure/ETyping" +"erasure/EReflect" -> "erasure/ETyping" +"pcuic/PCUICWcbvEval" -> "erasure/EWcbvEval" +"erasure/ECSubst" -> "erasure/EWcbvEval" +"erasure/ETyping" -> "erasure/EWcbvEval" +"erasure/ETyping" -> "erasure/EWndEval" +"pcuic/TemplateToPCUIC" -> "erasure/Erasure" +"template-coq/Pretty" -> "erasure/Erasure" +"erasure/EOptimizePropDiscr" -> "erasure/Erasure" +"erasure/EPretty" -> "erasure/Erasure" +"erasure/EDeps" -> "erasure/ErasureCorrectness" +"erasure/EInversion" -> "erasure/ErasureCorrectness" +"safechecker/PCUICSafeRetyping" -> "erasure/ErasureFunction" +"erasure/ErasureCorrectness" -> "erasure/ErasureFunction" +"pcuic/PCUICElimination" -> "erasure/Extract" +"erasure/ETyping" -> "erasure/Extract" +"erasure/Erasure" -> "erasure/Extraction" +"safechecker/PCUICErrors" -> "erasure/Prelim" +"erasure/EArities" -> "erasure/Prelim" +"erasure/EWcbvEval" -> "erasure/Prelim" +"pcuic/PCUICToTemplateCorrectness"[label="PCUICToTemplateCorrectness", color=lemonchiffon1] +"pcuic/PCUICToTemplate"[label="PCUICToTemplate", color=lemonchiffon1] +"pcuic/TemplateToPCUICCorrectness"[label="TemplateToPCUICCorrectness", color=lemonchiffon1] +"pcuic/TemplateToPCUIC"[label="TemplateToPCUIC", color=lemonchiffon1] +"pcuic/PCUICSafeLemmata"[label="PCUICSafeLemmata", color=lemonchiffon1] +"pcuic/PCUICSigmaCalculus"[label="PCUICSigmaCalculus", color=lemonchiffon1] +"pcuic/PCUICPrincipality"[label="PCUICPrincipality", color=lemonchiffon1] +"pcuic/PCUICSN"[label="PCUICSN", color=lemonchiffon1] +"pcuic/PCUICElimination"[label="PCUICElimination", color=lemonchiffon1] +"pcuic/PCUICCumulProp"[label="PCUICCumulProp", color=lemonchiffon1] +"pcuic/PCUICPretty"[label="PCUICPretty", color=lemonchiffon1] +"pcuic/PCUICWcbvEval"[label="PCUICWcbvEval", color=lemonchiffon1] +"pcuic/PCUICCSubst"[label="PCUICCSubst", color=lemonchiffon1] +"pcuic/PCUICCanonicity"[label="PCUICCanonicity", color=lemonchiffon1] +"pcuic/PCUICSR"[label="PCUICSR", color=lemonchiffon1] +"pcuic/PCUICInductiveInversion"[label="PCUICInductiveInversion", color=lemonchiffon1] +"pcuic/PCUICValidity"[label="PCUICValidity", color=lemonchiffon1] +"pcuic/PCUICInductives"[label="PCUICInductives", color=lemonchiffon1] +"pcuic/PCUICSpine"[label="PCUICSpine", color=lemonchiffon1] +"pcuic/PCUICWfUniverses"[label="PCUICWfUniverses", color=lemonchiffon1] +"pcuic/PCUICArities"[label="PCUICArities", color=lemonchiffon1] +"pcuic/PCUICContexts"[label="PCUICContexts", color=lemonchiffon1] +"pcuic/PCUICCtxShape"[label="PCUICCtxShape", color=lemonchiffon1] +"pcuic/PCUICAlpha"[label="PCUICAlpha", color=lemonchiffon1] +"pcuic/PCUICGeneration"[label="PCUICGeneration", color=lemonchiffon1] +"pcuic/PCUICConvCumInversion"[label="PCUICConvCumInversion", color=lemonchiffon1] +"pcuic/PCUICConversion"[label="PCUICConversion", color=lemonchiffon1] +"pcuic/PCUICContextConversion"[label="PCUICContextConversion", color=lemonchiffon1] +"pcuic/PCUICContextRelation"[label="PCUICContextRelation", color=lemonchiffon1] +"pcuic/PCUICConfluence"[label="PCUICConfluence", color=lemonchiffon1] +"pcuic/PCUICParallelReductionConfluence"[label="PCUICParallelReductionConfluence", color=lemonchiffon1] +"pcuic/PCUICParallelReduction"[label="PCUICParallelReduction", color=lemonchiffon1] +"pcuic/PCUICReduction"[label="PCUICReduction", color=lemonchiffon1] +"pcuic/PCUICCumulativity"[label="PCUICCumulativity", color=lemonchiffon1] +"pcuic/PCUICSubstitution"[label="PCUICSubstitution", color=lemonchiffon1] +"pcuic/PCUICUnivSubstitution"[label="PCUICUnivSubstitution", color=lemonchiffon1] +"pcuic/PCUICWeakening"[label="PCUICWeakening", color=lemonchiffon1] +"pcuic/PCUICClosed"[label="PCUICClosed", color=lemonchiffon1] +"pcuic/PCUICWeakeningEnv"[label="PCUICWeakeningEnv", color=lemonchiffon1] +"pcuic/PCUICEquality"[label="PCUICEquality", color=lemonchiffon1] +"pcuic/PCUICNameless"[label="PCUICNameless", color=lemonchiffon1] +"pcuic/PCUICNormal"[label="PCUICNormal", color=lemonchiffon1] +"pcuic/PCUICPosition"[label="PCUICPosition", color=lemonchiffon1] +"pcuic/PCUICInversion"[label="PCUICInversion", color=lemonchiffon1] +"pcuic/PCUICGlobalEnv"[label="PCUICGlobalEnv", color=lemonchiffon1] +"pcuic/PCUICTyping"[label="PCUICTyping", color=lemonchiffon1] +"pcuic/PCUICUnivSubst"[label="PCUICUnivSubst", color=lemonchiffon1] +"pcuic/PCUICLiftSubst"[label="PCUICLiftSubst", color=lemonchiffon1] +"pcuic/PCUICReflect"[label="PCUICReflect", color=lemonchiffon1] +"pcuic/PCUICInduction"[label="PCUICInduction", color=lemonchiffon1] +"pcuic/PCUICAstUtils"[label="PCUICAstUtils", color=lemonchiffon1] +"pcuic/PCUICSize"[label="PCUICSize", color=lemonchiffon1] +"pcuic/PCUICAst"[label="PCUICAst", color=lemonchiffon1] +"pcuic/PCUICRelations"[label="PCUICRelations", color=lemonchiffon1] +"pcuic/PCUICUtils"[label="PCUICUtils", color=lemonchiffon1] +"pcuic/PCUICAstUtils" -> "pcuic/PCUICToTemplate" +"template-coq/TypingWf" -> "pcuic/PCUICToTemplateCorrectness" +"template-coq/WfInv" -> "pcuic/PCUICToTemplateCorrectness" +"pcuic/PCUICGeneration" -> "pcuic/PCUICToTemplateCorrectness" +"pcuic/PCUICToTemplate" -> "pcuic/PCUICToTemplateCorrectness" +"template-coq/TypingWf" -> "pcuic/TemplateToPCUICCorrectness" +"template-coq/WfInv" -> "pcuic/TemplateToPCUICCorrectness" +"pcuic/PCUICSubstitution" -> "pcuic/TemplateToPCUICCorrectness" +"pcuic/TemplateToPCUIC" -> "pcuic/TemplateToPCUICCorrectness" +} diff --git a/dependency-graph/depgraph-2020-11-28.png b/dependency-graph/depgraph-2020-11-28.png new file mode 100644 index 000000000..bee4f97f0 Binary files /dev/null and b/dependency-graph/depgraph-2020-11-28.png differ diff --git a/dependency-graph/depgraph-2020-11-28.svg b/dependency-graph/depgraph-2020-11-28.svg new file mode 100644 index 000000000..3aa7bef2a --- /dev/null +++ b/dependency-graph/depgraph-2020-11-28.svg @@ -0,0 +1,6059 @@ + + + + + + image/svg+xml + + + + + + + + + + dependencies + + + safechecker/Extraction + + Extraction + + + + safechecker/PCUICSafeRetyping + + PCUICSafeRetyping + + + + erasure/ErasureFunction + + ErasureFunction + + + + safechecker/PCUICSafeRetyping->erasure/ErasureFunction + + + + + + safechecker/SafeTemplateChecker + + SafeTemplateChecker + + + + safechecker/SafeTemplateChecker->safechecker/Extraction + + + + + + safechecker/PCUICSafeChecker + + PCUICSafeChecker + + + + safechecker/PCUICSafeChecker->safechecker/SafeTemplateChecker + + + + + + safechecker/PCUICTypeChecker + + PCUICTypeChecker + + + + safechecker/PCUICTypeChecker->safechecker/PCUICSafeChecker + + + + + + safechecker/PCUICWfEnv + + PCUICWfEnv + + + + safechecker/PCUICWfEnv->safechecker/PCUICTypeChecker + + + + + + safechecker/PCUICWfReduction + + PCUICWfReduction + + + + safechecker/PCUICWfReduction->safechecker/PCUICTypeChecker + + + + + + safechecker/PCUICSafeConversion + + PCUICSafeConversion + + + + safechecker/PCUICSafeConversion->safechecker/PCUICWfReduction + + + + + + safechecker/PCUICEqualityDec + + PCUICEqualityDec + + + + safechecker/PCUICEqualityDec->safechecker/PCUICSafeConversion + + + + + + safechecker/PCUICSafeReduce + + PCUICSafeReduce + + + + safechecker/PCUICSafeReduce->safechecker/PCUICSafeRetyping + + + + + + safechecker/PCUICSafeReduce->safechecker/PCUICSafeConversion + + + + + + safechecker/PCUICErrors + + PCUICErrors + + + + safechecker/PCUICErrors->safechecker/PCUICSafeReduce + + + + + + erasure/Prelim + + Prelim + + + + safechecker/PCUICErrors->erasure/Prelim + + + + + + pcuic/PCUICValidity + + PCUICValidity + + + + pcuic/PCUICAlpha + + PCUICAlpha + + + + pcuic/PCUICValidity->pcuic/PCUICAlpha + + + + + + pcuic/PCUICInductiveInversion + + PCUICInductiveInversion + + + + pcuic/PCUICValidity->pcuic/PCUICInductiveInversion + + + + + + pcuic/PCUICSR + + PCUICSR + + + + pcuic/PCUICAlpha->pcuic/PCUICSR + + + + + + pcuic/PCUICWfUniverses + + PCUICWfUniverses + + + + pcuic/PCUICArities + + PCUICArities + + + + pcuic/PCUICWfUniverses->pcuic/PCUICArities + + + + + + pcuic/PCUICSpine + + PCUICSpine + + + + pcuic/PCUICArities->pcuic/PCUICSpine + + + + + + template-coq/EnvironmentTyping + + EnvironmentTyping + + + + pcuic/PCUICAst + + PCUICAst + + + + template-coq/EnvironmentTyping->pcuic/PCUICAst + + + + + + template-coq/Typing + + Typing + + + + template-coq/EnvironmentTyping->template-coq/Typing + + + + + + pcuic/PCUICSize + + PCUICSize + + + + + pcuic/TemplateToPCUIC + + TemplateToPCUIC + + + + + pcuic/PCUICAstUtils + + PCUICAstUtils + + + + pcuic/PCUICSize->pcuic/PCUICAstUtils + + + + + pcuic/PCUICInduction + + PCUICInduction + + + + pcuic/PCUICAstUtils->pcuic/PCUICInduction + + + + + + pcuic/PCUICUtils + + PCUICUtils + + + + pcuic/PCUICAstUtils->pcuic/PCUICUtils + + + + + + pcuic/PCUICToTemplate + + PCUICToTemplate + + + + pcuic/PCUICAstUtils->pcuic/PCUICToTemplate + + + + + + template-coq/Reflect + + Reflect + + + + template-coq/Reflect->pcuic/PCUICAstUtils + + + + + template-coq/Extraction + + Extraction + + + + template-coq/Reflect->template-coq/Extraction + + + + + + + + + + + + template-coq/common/uGraph + + uGraph + + + + + template-coq/Constants + + Constants + + + + template-coq/common/uGraph->template-coq/Constants + + + + + + + + pcuic/PCUICClosed + + PCUICClosed + + + + pcuic/PCUICCSubst + + PCUICCSubst + + + + pcuic/PCUICClosed->pcuic/PCUICCSubst + + + + + + pcuic/PCUICWeakening + + PCUICWeakening + + + + pcuic/PCUICClosed->pcuic/PCUICWeakening + + + + + + pcuic/PCUICWcbvEval + + PCUICWcbvEval + + + + pcuic/PCUICCSubst->pcuic/PCUICWcbvEval + + + + + + pcuic/PCUICElimination + + PCUICElimination + + + + pcuic/PCUICCanonicity + + PCUICCanonicity + + + + pcuic/PCUICElimination->pcuic/PCUICCanonicity + + + + + + erasure/Extract + + Extract + + + + pcuic/PCUICElimination->erasure/Extract + + + + + + pcuic/PCUICCanonicity->safechecker/PCUICSafeReduce + + + + + + erasure/EArities + + EArities + + + + pcuic/PCUICCanonicity->erasure/EArities + + + + + + pcuic/PCUICSN + + PCUICSN + + + + pcuic/PCUICSN->pcuic/PCUICCanonicity + + + + + + pcuic/PCUICWcbvEval->pcuic/PCUICCanonicity + + + + + + erasure/EWcbvEval + + EWcbvEval + + + + pcuic/PCUICWcbvEval->erasure/EWcbvEval + + + + + + pcuic/PCUICWeakeningEnv + + PCUICWeakeningEnv + + + + pcuic/PCUICWeakeningEnv->pcuic/PCUICClosed + + + + + + pcuic/PCUICParallelReductionConfluence + + PCUICParallelReductionConfluence + + + + pcuic/PCUICConfluence + + PCUICConfluence + + + + pcuic/PCUICParallelReductionConfluence->pcuic/PCUICConfluence + + + + + + pcuic/PCUICContextConversion + + PCUICContextConversion + + + + pcuic/PCUICConfluence->pcuic/PCUICContextConversion + + + + + + pcuic/PCUICConversion + + PCUICConversion + + + + pcuic/PCUICContextConversion->pcuic/PCUICConversion + + + + + + pcuic/PCUICContextRelation + + PCUICContextRelation + + + + pcuic/PCUICContextRelation->pcuic/PCUICContextConversion + + + + + + pcuic/PCUICNormal + + PCUICNormal + + + + pcuic/PCUICContextRelation->pcuic/PCUICNormal + + + + + + pcuic/PCUICTyping + + PCUICTyping + + + + pcuic/PCUICTyping->pcuic/PCUICWeakeningEnv + + + + + + pcuic/PCUICTyping->pcuic/PCUICContextRelation + + + + + + pcuic/PCUICNameless + + PCUICNameless + + + + pcuic/PCUICTyping->pcuic/PCUICNameless + + + + + + pcuic/PCUICGeneration + + PCUICGeneration + + + + pcuic/PCUICTyping->pcuic/PCUICGeneration + + + + + + pcuic/PCUICGlobalEnv + + PCUICGlobalEnv + + + + pcuic/PCUICTyping->pcuic/PCUICGlobalEnv + + + + + + pcuic/PCUICCtxShape + + PCUICCtxShape + + + + pcuic/PCUICContexts + + PCUICContexts + + + + pcuic/PCUICCtxShape->pcuic/PCUICContexts + + + + + + pcuic/PCUICContexts->pcuic/PCUICWfUniverses + + + + + + pcuic/PCUICCumulProp + + PCUICCumulProp + + + + pcuic/PCUICCumulProp->pcuic/PCUICElimination + + + + + + pcuic/PCUICConvCumInversion + + PCUICConvCumInversion + + + + pcuic/PCUICCumulProp->pcuic/PCUICConvCumInversion + + + + + + pcuic/PCUICPrincipality + + PCUICPrincipality + + + + pcuic/PCUICCumulProp->pcuic/PCUICPrincipality + + + + + + pcuic/PCUICConvCumInversion->safechecker/PCUICSafeConversion + + + + + + pcuic/PCUICInversion + + PCUICInversion + + + + pcuic/PCUICConversion->pcuic/PCUICInversion + + + + + + pcuic/PCUICInversion->pcuic/PCUICCtxShape + + + + + + pcuic/PCUICNameless->pcuic/PCUICCtxShape + + + + + + template-coq/UnivSubst + + UnivSubst + + + + template-coq/UnivSubst->pcuic/PCUICCtxShape + + + + + + template-coq/UnivSubst->pcuic/PCUICNormal + + + + + + template-coq/UnivSubst->template-coq/Typing + + + + + + template-coq/UnivSubst->template-coq/Extraction + + + + + + pcuic/PCUICSafeLemmata + + PCUICSafeLemmata + + + + pcuic/PCUICSafeLemmata->pcuic/PCUICSN + + + + + + pcuic/PCUICSafeLemmata->pcuic/PCUICCumulProp + + + + + + pcuic/PCUICReduction + + PCUICReduction + + + + pcuic/PCUICCumulativity + + PCUICCumulativity + + + + pcuic/PCUICReduction->pcuic/PCUICCumulativity + + + + + + pcuic/PCUICCumulativity->pcuic/PCUICTyping + + + + + + pcuic/PCUICLiftSubst + + PCUICLiftSubst + + + + pcuic/PCUICEquality + + PCUICEquality + + + + pcuic/PCUICLiftSubst->pcuic/PCUICEquality + + + + + + pcuic/PCUICPretty + + PCUICPretty + + + + pcuic/PCUICLiftSubst->pcuic/PCUICPretty + + + + + + pcuic/PCUICUnivSubst + + PCUICUnivSubst + + + + pcuic/PCUICLiftSubst->pcuic/PCUICUnivSubst + + + + + + pcuic/PCUICPosition + + PCUICPosition + + + + pcuic/PCUICEquality->pcuic/PCUICPosition + + + + + + pcuic/PCUICReflect + + PCUICReflect + + + + pcuic/PCUICReflect->pcuic/PCUICEquality + + + + + + erasure/EReflect + + EReflect + + + + pcuic/PCUICReflect->erasure/EReflect + + + + + + pcuic/PCUICGeneration->pcuic/PCUICWeakening + + + + + + pcuic/PCUICToTemplateCorrectness + + PCUICToTemplateCorrectness + + + + pcuic/PCUICGeneration->pcuic/PCUICToTemplateCorrectness + + + + + + pcuic/PCUICGlobalEnv->safechecker/PCUICWfEnv + + + + + + pcuic/PCUICGlobalEnv->safechecker/PCUICEqualityDec + + + + + + pcuic/PCUICInduction->pcuic/PCUICLiftSubst + + + + + + pcuic/PCUICInduction->pcuic/PCUICReflect + + + + + + pcuic/PCUICInductiveInversion->pcuic/PCUICSR + + + + + + pcuic/PCUICInductives + + PCUICInductives + + + + pcuic/PCUICSpine->pcuic/PCUICInductives + + + + + + pcuic/PCUICInductives->pcuic/PCUICValidity + + + + + + pcuic/PCUICNormal->pcuic/PCUICSafeLemmata + + + + + + pcuic/PCUICSubstitution + + PCUICSubstitution + + + + pcuic/PCUICParallelReduction + + PCUICParallelReduction + + + + pcuic/PCUICSubstitution->pcuic/PCUICParallelReduction + + + + + + pcuic/TemplateToPCUICCorrectness + + TemplateToPCUICCorrectness + + + + pcuic/PCUICSubstitution->pcuic/TemplateToPCUICCorrectness + + + + + + pcuic/PCUICParallelReduction->pcuic/PCUICParallelReductionConfluence + + + + + + pcuic/PCUICPosition->pcuic/PCUICReduction + + + + + + pcuic/PCUICPretty->safechecker/PCUICErrors + + + + + + pcuic/PCUICPrincipality->safechecker/PCUICSafeConversion + + + + + + pcuic/PCUICPrincipality->erasure/EArities + + + + + + pcuic/PCUICRelations + + PCUICRelations + + + + pcuic/PCUICRelations->pcuic/PCUICReduction + + + + + + pcuic/PCUICUnivSubst->pcuic/PCUICReduction + + + + + + pcuic/PCUICSR->pcuic/PCUICWcbvEval + + + + + + pcuic/PCUICSR->pcuic/PCUICSafeLemmata + + + + + + pcuic/PCUICSigmaCalculus + + PCUICSigmaCalculus + + + + pcuic/PCUICWeakening->pcuic/PCUICSigmaCalculus + + + + + + pcuic/PCUICUnivSubstitution + + PCUICUnivSubstitution + + + + pcuic/PCUICWeakening->pcuic/PCUICUnivSubstitution + + + + + + pcuic/PCUICSigmaCalculus->pcuic/PCUICSubstitution + + + + + + pcuic/PCUICUnivSubstitution->pcuic/PCUICSubstitution + + + + + + pcuic/PCUICUtils->pcuic/PCUICTyping + + + + + + template-coq/utils/LibHypsNaming + + LibHypsNaming + + + + template-coq/utils/LibHypsNaming->pcuic/PCUICTyping + + + + + + + + pcuic/TemplateToPCUIC->safechecker/SafeTemplateChecker + + + + + + erasure/Erasure + + Erasure + + + + pcuic/TemplateToPCUIC->erasure/Erasure + + + + + + pcuic/TemplateToPCUIC->pcuic/TemplateToPCUICCorrectness + + + + + + template-coq/Environment + + Environment + + + + template-coq/Ast + + Ast + + + + template-coq/Environment->template-coq/Ast + + + + + + template-coq/AstUtils + + AstUtils + + + + template-coq/Ast->template-coq/AstUtils + + + + + + template-coq/TemplateMonad/Common + + Common + + + + template-coq/Ast->template-coq/TemplateMonad/Common + + + + + + template-coq/AstUtils->template-coq/EnvironmentTyping + + + + + + template-coq/Induction + + Induction + + + + template-coq/AstUtils->template-coq/Induction + + + + + + template-coq/TemplateMonad/Extractable + + Extractable + + + + template-coq/AstUtils->template-coq/TemplateMonad/Extractable + + + + + + template-coq/TemplateMonad/Core + + Core + + + + template-coq/AstUtils->template-coq/TemplateMonad/Core + + + + + + template-coq/WfInv + + WfInv + + + + template-coq/AstUtils->template-coq/WfInv + + + + + + template-coq/Kernames + + Kernames + + + + template-coq/AstUtils->template-coq/Kernames + + + + + + template-coq/utils + + utils + + + + template-coq/BasicAst + + BasicAst + + + + template-coq/utils->template-coq/BasicAst + + + + + + template-coq/Universes + + Universes + + + + template-coq/BasicAst->template-coq/Universes + + + + + + template-coq/Universes->template-coq/common/uGraph + + + + + + template-coq/Universes->template-coq/Environment + + + + + + erasure/EAst + + EAst + + + + template-coq/Universes->erasure/EAst + + + + + + template-coq/Induction->template-coq/Reflect + + + + + + template-coq/LiftSubst + + LiftSubst + + + + template-coq/Induction->template-coq/LiftSubst + + + + + + template-coq/LiftSubst->template-coq/UnivSubst + + + + + + template-coq/Pretty + + Pretty + + + + template-coq/LiftSubst->template-coq/Pretty + + + + + + template-coq/Typing->safechecker/SafeTemplateChecker + + + + + + template-coq/TypingWf + + TypingWf + + + + template-coq/Typing->template-coq/TypingWf + + + + + + + + + + + + template-coq/All + + template-coq/All + + + + template-coq/Typing->template-coq/All + + + + + + template-coq/config + + config + + + + template-coq/config->template-coq/Universes + + + + + + template-coq/utils/wGraph + + wGraph + + + + template-coq/utils/wGraph->template-coq/common/uGraph + + + + + + template-coq/utils/All_Forall + + All_Forall + + + + template-coq/monad_utils + + monad_utils + + + + template-coq/utils/All_Forall->template-coq/monad_utils + + + + + + template-coq/utils/MCUtils + + MCUtils + + + + template-coq/utils/All_Forall->template-coq/utils/MCUtils + + + + + + template-coq/monad_utils->template-coq/utils + + + + + + template-coq/utils/MCUtils->template-coq/utils + + + + + + template-coq/utils/MCUtils->template-coq/utils/wGraph + + + + + + template-coq/utils/MCOption + + MCOption + + + + template-coq/utils/MCOption->template-coq/utils/All_Forall + + + + + + template-coq/utils/MCSquash + + MCSquash + + + + template-coq/utils/MCSquash->template-coq/utils/All_Forall + + + + + + template-coq/utils/MCPrelude + + MCPrelude + + + + template-coq/utils/MCList + + MCList + + + + template-coq/utils/MCPrelude->template-coq/utils/MCList + + + + + + template-coq/utils/MCList->template-coq/utils/MCOption + + + + + + template-coq/utils/MCRelations + + MCRelations + + + + template-coq/utils/MCRelations->template-coq/utils/MCList + + + + + + template-coq/utils/MCProd + + MCProd + + + + template-coq/utils/MCProd->template-coq/utils/MCOption + + + + + + template-coq/utils/MCCompare + + MCCompare + + + + template-coq/utils/MCString + + MCString + + + + template-coq/utils/MCCompare->template-coq/utils/MCString + + + + + + template-coq/utils/MCString->template-coq/utils/MCUtils + + + + + + template-coq/utils/MCArith + + MCArith + + + + template-coq/utils/MCArith->template-coq/utils/MCUtils + + + + + + template-coq/utils/MCEquality + + MCEquality + + + + template-coq/utils/MCEquality->template-coq/utils/MCUtils + + + + + + template-coq/Loader + + template-coq/Loader + + + + template-coq/Constants->template-coq/Loader + + + + + + template-coq/TemplateMonad/Extractable->template-coq/Extraction + + + + + + template-coq/TemplateMonad/Extractable->template-coq/Constants + + + + + + template-coq/TemplateMonad + + TemplateMonad + + + + template-coq/TemplateMonad/Core->template-coq/TemplateMonad + + + + + + template-coq/TemplateMonad/Common->template-coq/TemplateMonad/Extractable + + + + + + template-coq/TemplateMonad/Common->template-coq/TemplateMonad/Core + + + + + + template-coq/TemplateMonad->template-coq/Constants + + + + + + + template-coq/TypingWf->pcuic/PCUICToTemplateCorrectness + + + + + + template-coq/TypingWf->pcuic/TemplateToPCUICCorrectness + + + + + + + template-coq/WfInv->pcuic/PCUICToTemplateCorrectness + + + + + + template-coq/WfInv->pcuic/TemplateToPCUICCorrectness + + + + + + template-coq/Pretty->template-coq/Extraction + + + + + + template-coq/Pretty->erasure/Erasure + + + + + + erasure/EAstUtils + + EAstUtils + + + + template-coq/Kernames->erasure/EAstUtils + + + + + + + + + + + + + + + + + template-coq/Loader->template-coq/All + + + + + + + erasure/Extraction + + Extraction + + + + erasure/Erasure->erasure/Extraction + + + + + + erasure/EOptimizePropDiscr + + EOptimizePropDiscr + + + + erasure/EOptimizePropDiscr->erasure/Erasure + + + + + + erasure/ErasureFunction->erasure/EOptimizePropDiscr + + + + + + erasure/ErasureCorrectness + + ErasureCorrectness + + + + erasure/ErasureCorrectness->erasure/ErasureFunction + + + + + + erasure/EArities->erasure/Prelim + + + + + + erasure/EInversion + + EInversion + + + + erasure/EInversion->erasure/ErasureCorrectness + + + + + + erasure/ESubstitution + + ESubstitution + + + + erasure/EDeps + + EDeps + + + + erasure/ESubstitution->erasure/EDeps + + + + + + erasure/Prelim->erasure/EInversion + + + + + + erasure/Prelim->erasure/ESubstitution + + + + + + erasure/EAll + + EAll + + + + erasure/EDeps->erasure/ErasureCorrectness + + + + + + erasure/Extract->erasure/EArities + + + + + + erasure/Extract->erasure/EAll + + + + + + erasure/ETyping + + ETyping + + + + erasure/ETyping->erasure/Extract + + + + + + erasure/EWndEval + + EWndEval + + + + erasure/ETyping->erasure/EWndEval + + + + + + erasure/ETyping->erasure/EWcbvEval + + + + + + erasure/EPretty + + EPretty + + + + erasure/ETyping->erasure/EPretty + + + + + + erasure/EWndEval->erasure/EAll + + + + + + erasure/EWcbvEval->erasure/Prelim + + + + + + erasure/EWcbvEval->erasure/EAll + + + + + + erasure/ECSubst + + ECSubst + + + + erasure/ECSubst->erasure/EWcbvEval + + + + + + erasure/EPretty->erasure/Erasure + + + + + + erasure/EReflect->erasure/ETyping + + + + + + erasure/ELiftSubst + + ELiftSubst + + + + erasure/ELiftSubst->erasure/ETyping + + + + + + erasure/ELiftSubst->erasure/ECSubst + + + + + + erasure/EInduction + + EInduction + + + + erasure/EInduction->erasure/EReflect + + + + + + erasure/EInduction->erasure/ELiftSubst + + + + + + erasure/EAstUtils->erasure/ETyping + + + + + + erasure/EAst->erasure/EInduction + + + + + + erasure/EAst->erasure/EAstUtils + + + + + + pcuic/PCUICToTemplate->pcuic/PCUICToTemplateCorrectness + + + + + + diff --git a/dependency-graph/generate-depgraph.sh b/dependency-graph/generate-depgraph.sh index 92c666d4b..ccb7497db 100755 --- a/dependency-graph/generate-depgraph.sh +++ b/dependency-graph/generate-depgraph.sh @@ -1,4 +1,8 @@ -#!/bin/bash +#!/usr/local/bin/bash + +SED=`which gsed || which sed` + +echo "sed is " $SED ############################################################### # @@ -10,6 +14,8 @@ # cd dependency-graph # ./generate-depgraph.sh depgraph-2020-09-24 # +# Requires: graphviz for .dot to .png/.svg generation, +# a recent bash (not the one shipped with OS X Catalina for example) ############################################################### @@ -27,23 +33,23 @@ folders[erasure]=tan # Two first lines echo "digraph dependencies {" > $dot_file echo "node[style=filled]" >> $dot_file - for folder in "${!folders[@]}" do cd ../$folder + echo `pwd` coqdep -f _CoqProject -dumpgraph ../dependency-graph/$folder.dot > /dev/null cd ../dependency-graph # remove the first and last lines - sed -i '1d' $folder.dot - sed -i '$d' $folder.dot + $SED -i '1d' $folder.dot + $SED -i '$d' $folder.dot # change a bit the names of the nodes for otherfolder in "${!folders[@]}" do - sed -i "s@../$otherfolder/theories@$otherfolder@g" $folder.dot + $SED -i "s@../$otherfolder/theories@$otherfolder@g" $folder.dot done - sed -i "s/theories/$folder/g" $folder.dot + $SED -i "s/theories/$folder/g" $folder.dot # change the color of the nodes - sed -i "s/]/, color=${folders[$folder]}]/g" $folder.dot + $SED -i "s/]/, color=${folders[$folder]}]/g" $folder.dot # concatenate cat $folder.dot >> $dot_file rm -f $folder.dot diff --git a/erasure/_PluginProject.in b/erasure/_PluginProject.in index e47650f5d..f46d585c2 100644 --- a/erasure/_PluginProject.in +++ b/erasure/_PluginProject.in @@ -8,8 +8,8 @@ src/init.ml # src/classes0.ml # From Coq's stdlib -src/mSetWeakList.mli -src/mSetWeakList.ml +# src/mSetWeakList.mli +# src/mSetWeakList.ml # src/eqdepFacts.mli # src/eqdepFacts.ml src/orderedTypeAlt.mli @@ -27,10 +27,12 @@ src/kernames.mli src/kernames.ml # src/reflect.mli # src/reflect.ml -src/typing0.mli -src/typing0.ml +# src/typing0.mli +# src/typing0.ml # From PCUIC +src/pCUICPrimitive.mli +src/pCUICPrimitive.ml src/pCUICAst.ml src/pCUICAst.mli src/pCUICAstUtils.ml @@ -45,8 +47,8 @@ src/pCUICTyping.ml src/pCUICTyping.mli src/pCUICUnivSubst.ml src/pCUICUnivSubst.mli -src/pCUICCumulativity.mli -src/pCUICCumulativity.ml +# src/pCUICCumulativity.mli +# src/pCUICCumulativity.ml src/pCUICPosition.mli src/pCUICPosition.ml src/pCUICNormal.mli @@ -77,16 +79,14 @@ src/templateToPCUIC.mli src/templateToPCUIC.ml # From SafeChecker +src/pCUICErrors.mli +src/pCUICErrors.ml src/pCUICSafeReduce.mli src/pCUICSafeReduce.ml -src/pCUICSafeConversion.mli -src/pCUICSafeConversion.ml -src/pCUICSafeChecker.mli -src/pCUICSafeChecker.ml +# src/pCUICSafeConversion.mli +# src/pCUICSafeConversion.ml src/pCUICSafeRetyping.ml src/pCUICSafeRetyping.mli -src/safeTemplateChecker.mli -src/safeTemplateChecker.ml # src/ssreflect.mli # src/ssreflect.ml diff --git a/erasure/src/g_metacoq_erasure.mlg b/erasure/src/g_metacoq_erasure.mlg index cb88f4803..2a545d0a5 100644 --- a/erasure/src/g_metacoq_erasure.mlg +++ b/erasure/src/g_metacoq_erasure.mlg @@ -7,7 +7,6 @@ open Stdarg open Pp open PeanoNat.Nat open Datatypes -open PCUICSafeChecker let pr_char c = str (Char.escaped c) diff --git a/erasure/src/metacoq_erasure_plugin.mlpack b/erasure/src/metacoq_erasure_plugin.mlpack index 70bc3cb2e..986078d2b 100644 --- a/erasure/src/metacoq_erasure_plugin.mlpack +++ b/erasure/src/metacoq_erasure_plugin.mlpack @@ -6,16 +6,15 @@ Utils WGraph UGraph0 -Typing0 OrderedTypeAlt Kernames -Reflect Init Classes0 Logic1 Relation Relation_Properties +PCUICPrimitive PCUICAst PCUICAstUtils PCUICUnivSubst @@ -30,14 +29,11 @@ PCUICNormal PCUICPosition PCUICPretty TemplateToPCUIC -PCUICCumulativity +PCUICErrors PCUICUnivSubst PCUICSafeReduce -PCUICSafeConversion -PCUICSafeChecker PCUICSafeRetyping -SafeTemplateChecker EAst EAstUtils diff --git a/erasure/theories/EArities.v b/erasure/theories/EArities.v index 9637f483f..7918dda63 100644 --- a/erasure/theories/EArities.v +++ b/erasure/theories/EArities.v @@ -548,7 +548,6 @@ Qed. Lemma Is_type_eval (Σ : global_env_ext) t v: wf Σ -> - axiom_free Σ -> eval Σ t v -> isErasable Σ [] t -> isErasable Σ [] v. @@ -563,13 +562,12 @@ Qed. Lemma Is_type_eval_inv (Σ : global_env_ext) t v: wf_ext Σ -> - axiom_free Σ -> PCUICSafeLemmata.welltyped Σ [] t -> PCUICWcbvEval.eval Σ t v -> isErasable Σ [] v -> ∥ isErasable Σ [] t ∥. Proof. - intros wfΣ axfree [T HT] ev [vt [Ht Hp]]. + intros wfΣ [T HT] ev [vt [Ht Hp]]. eapply wcbeval_red in ev; eauto. pose proof (subject_reduction _ _ _ _ _ wfΣ.1 HT ev). pose proof (common_typing _ wfΣ Ht X) as [P [Pvt [Pt vP]]]. diff --git a/erasure/theories/EAst.v b/erasure/theories/EAst.v index 187913650..164fb08c4 100644 --- a/erasure/theories/EAst.v +++ b/erasure/theories/EAst.v @@ -1,7 +1,6 @@ (* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Export BasicAst Universes. -From MetaCoq.Template Require Import utils. - +From MetaCoq.Template Require Export utils BasicAst Universes. +From MetaCoq.PCUIC Require Import PCUICPrimitive. (** * Extracted terms These are the terms produced by extraction: compared to kernel terms, @@ -38,7 +37,8 @@ Inductive term : Set := term (* discriminee *) -> list (nat * term) (* branches *) -> term | tProj : projection -> term -> term | tFix : mfixpoint term -> nat -> term -| tCoFix : mfixpoint term -> nat -> term. +| tCoFix : mfixpoint term -> nat -> term +| tPrim : prim_val term -> term. Fixpoint mkApps t us := match us with diff --git a/erasure/theories/EAstUtils.v b/erasure/theories/EAstUtils.v index a3b246653..d3e5f2acf 100644 --- a/erasure/theories/EAstUtils.v +++ b/erasure/theories/EAstUtils.v @@ -1,6 +1,7 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Template Require Import utils Kernames. From MetaCoq.Erasure Require Import EAst. +From MetaCoq.PCUIC Require PCUICPrimitive. Require Import ssreflect ssrbool. @@ -267,6 +268,7 @@ Fixpoint string_of_term (t : term) : string := ^ string_of_term c ^ ")" | tFix l n => "Fix(" ^ (string_of_list (string_of_def string_of_term) l) ^ "," ^ string_of_nat n ^ ")" | tCoFix l n => "CoFix(" ^ (string_of_list (string_of_def string_of_term) l) ^ "," ^ string_of_nat n ^ ")" + | tPrim p => "Prim(" ^ PCUICPrimitive.string_of_prim string_of_term p ^ ")" end. (** Compute all the global environment dependencies of the term *) diff --git a/erasure/theories/EDeps.v b/erasure/theories/EDeps.v index 4854d38b4..982764c75 100644 --- a/erasure/theories/EDeps.v +++ b/erasure/theories/EDeps.v @@ -3,10 +3,12 @@ From Equations Require Import Equations. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping PCUICInversion PCUICWeakeningEnv. +Set Warnings "-notation-overridden". From MetaCoq.Erasure Require Import EAst EAstUtils ECSubst EInduction ELiftSubst ESubstitution ETyping Extract EWcbvEval Prelim. +Set Warnings "+notation-overridden". From MetaCoq.Template Require Import config utils monad_utils. Derive NoConfusion for term. @@ -563,8 +565,7 @@ Proof. depelim typ. depelim er. depelim all_deps. - destruct p. - destruct p0 as (?&?&?). + destruct p as (?&?&?). now constructor; eauto. - constructor. apply inversion_CoFix in wt as (?&?&?&?&?&?&?); eauto. diff --git a/erasure/theories/EInduction.v b/erasure/theories/EInduction.v index 88c7ea97a..661e6df55 100644 --- a/erasure/theories/EInduction.v +++ b/erasure/theories/EInduction.v @@ -29,6 +29,7 @@ Lemma term_forall_list_ind : (forall (s : projection) (t : term), P t -> P (tProj s t)) -> (forall (m : mfixpoint term) (n : nat), All (fun x => P (dbody x)) m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), All (fun x => P (dbody x)) m -> P (tCoFix m n)) -> + (forall p, P (tPrim p)) -> forall t : term, P t. Proof. intros until t. revert t. diff --git a/erasure/theories/ELiftSubst.v b/erasure/theories/ELiftSubst.v index 940d650e4..14147e571 100644 --- a/erasure/theories/ELiftSubst.v +++ b/erasure/theories/ELiftSubst.v @@ -32,9 +32,9 @@ Fixpoint lift n k t : term := | tVar _ => t | tConst _ => t | tConstruct _ _ => t + | tPrim _ => t end. - Notation lift0 n := (lift n 0). Definition up := lift 1 0. diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 1ffc1fdc1..4176ab788 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -8,7 +8,6 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICCumulativity PCUICSR PCUICSafeLemmata PCUICValidity PCUICPrincipality PCUICElimination PCUICSN. -From MetaCoq.SafeChecker Require Import PCUICSafeReduce PCUICSafeChecker PCUICSafeRetyping. From MetaCoq.Erasure Require Import EAstUtils EArities Extract Prelim ErasureCorrectness EDeps ErasureFunction ELiftSubst. @@ -28,7 +27,9 @@ Ltac introdep := let H := fresh in intros H; depelim H. Hint Constructors Ee.eval : core. +Set Warnings "-notation-overridden". Import E. +Set Warnings "+notation-overridden". Section optimize. Context (Σ : global_context). @@ -65,6 +66,7 @@ Section optimize. | tVar _ => t | tConst _ => t | tConstruct _ _ => t + | tPrim _ => t end. Lemma optimize_mkApps f l : optimize (mkApps f l) = mkApps (optimize f) (map optimize l). @@ -230,17 +232,15 @@ Proof. Qed. Lemma erasable_tBox_value (wfl := Ee.default_wcbv_flags) (Σ : global_env_ext) (wfΣ : wf_ext Σ) t T v : - axiom_free Σ.1 -> forall wt : Σ ;;; [] |- t : T, Σ |-p t ▷ v -> erases Σ [] v tBox -> ∥ isErasable Σ [] t ∥. Proof. intros. - depind H0. + depind H. eapply Is_type_eval_inv; eauto. eexists; eauto. Qed. Lemma erase_eval_to_box (wfl := Ee.default_wcbv_flags) {Σ : global_env_ext} {wfΣ : wf_ext Σ} {t v Σ' t' deps} : - axiom_free Σ.1 -> forall wt : welltyped Σ [] t, erase Σ (sq wfΣ) [] t wt = t' -> KernameSet.subset (term_global_deps t') deps -> @@ -248,9 +248,9 @@ Lemma erase_eval_to_box (wfl := Ee.default_wcbv_flags) {Σ : global_env_ext} {w PCUICWcbvEval.eval Σ t v -> @Ee.eval Ee.default_wcbv_flags Σ' t' tBox -> ∥ isErasable Σ [] t ∥. Proof. - intros axiomfree [T wt]. + intros [T wt]. intros. - destruct (erase_correct Σ wfΣ _ _ _ _ _ axiomfree _ H H0 H1 X) as [ev [eg [eg']]]. + destruct (erase_correct Σ wfΣ _ _ _ _ _ _ H H0 H1 X) as [ev [eg [eg']]]. pose proof (Ee.eval_deterministic H2 eg'). subst. eapply erasable_tBox_value; eauto. Qed. @@ -455,7 +455,6 @@ Proof. Qed. Lemma erase_opt_correct (wfl := Ee.default_wcbv_flags) (Σ : global_env_ext) (wfΣ : wf_ext Σ) t v Σ' t' : - axiom_free Σ.1 -> forall wt : welltyped Σ [] t, erase Σ (sq wfΣ) [] t wt = t' -> erase_global (term_global_deps t') Σ (sq wfΣ.1) = Σ' -> @@ -463,10 +462,9 @@ Lemma erase_opt_correct (wfl := Ee.default_wcbv_flags) (Σ : global_env_ext) (wf ∃ v' : term, Σ;;; [] |- v ⇝ℇ v' ∧ ∥ @Ee.eval Ee.opt_wcbv_flags (optimize_env Σ') (optimize Σ' t') (optimize Σ' v') ∥. Proof. - intros axiomfree wt. + intros wt. generalize (sq wfΣ.1) as swfΣ. intros swfΣ HΣ' Ht' ev. - assert (extraction_pre Σ) by now constructor. pose proof (erases_erase (wfΣ := sq wfΣ) wt); eauto. rewrite HΣ' in H. destruct wt as [T wt]. @@ -483,5 +481,3 @@ Proof. eapply KernameSet.subset_spec. intros x hin; auto. Qed. - -Print Assumptions erase_opt_correct. \ No newline at end of file diff --git a/erasure/theories/EPretty.v b/erasure/theories/EPretty.v index 3dd6f2161..e14febe0f 100644 --- a/erasure/theories/EPretty.v +++ b/erasure/theories/EPretty.v @@ -2,10 +2,10 @@ From Coq Require Import Program. From MetaCoq.Template Require Import utils. From MetaCoq.Erasure Require Import EAst EAstUtils ETyping. +From MetaCoq.PCUIC Require Import PCUICPrimitive. (** * Pretty printing *) - Section print_term. Context (Σ : global_context). @@ -160,6 +160,8 @@ Section print_term. | tCoFix l n => parens top ("let cofix " ^ print_defs print_term Γ l ^ nl ^ " in " ^ List.nth_default (string_of_nat n) (map (string_of_name ∘ dname) l) n) + | tPrim p => + parens top (string_of_prim (print_term Γ false false) p) end. End print_term. diff --git a/erasure/theories/EReflect.v b/erasure/theories/EReflect.v index db01a0f4d..4b3424fda 100644 --- a/erasure/theories/EReflect.v +++ b/erasure/theories/EReflect.v @@ -1,6 +1,6 @@ From MetaCoq.Template Require Import utils Reflect. From MetaCoq.Erasure Require Import EAst EInduction. -From MetaCoq.PCUIC Require Import PCUICReflect. +From MetaCoq.PCUIC Require Import PCUICReflect PCUICPrimitive. From Equations Require Import Equations. Local Ltac finish := @@ -23,6 +23,7 @@ Local Ltac term_dec_tac term_dec := | i : kername, i' : kername |- _ => fcase (kername_eq_dec i i') | i : string, i' : kername |- _ => fcase (string_dec i i') | n : name, n' : name |- _ => fcase (eq_dec n n') + | i : prim_val _, i' : prim_val _ |- _ => fcase (eq_dec i i') | i : inductive, i' : inductive |- _ => fcase (eq_dec i i') | x : inductive * nat, y : inductive * nat |- _ => fcase (eq_dec x y) diff --git a/erasure/theories/ESubstitution.v b/erasure/theories/ESubstitution.v index b3c8f6801..87d7f5c81 100644 --- a/erasure/theories/ESubstitution.v +++ b/erasure/theories/ESubstitution.v @@ -108,7 +108,7 @@ Proof. - econstructor. eapply All2_All_mix_left in X1; eauto. eapply All2_impl. exact X1. - intros ? ? [[[]] [? []]]. + intros ? ? [[?] [? []]]. split; eauto. - econstructor. eapply All2_All_mix_left in X1; eauto. @@ -129,7 +129,7 @@ Lemma Is_type_weakening: Proof. intros. destruct X2 as (T & ? & ?). eexists. split. eapply weakening_typing; eauto. - now apply All_local_env_app in X1. + now apply All_local_env_app_inv in X1. destruct s as [? | [u []]]. - left. clear - i. generalize (#|Γ''|), (#|Γ'|). induction T; cbn in *; intros; try now inv i. @@ -137,7 +137,7 @@ Proof. + now eapply IHT3. - right. exists u. split; eauto. eapply weakening_typing in t1; eauto. - now apply All_local_env_app in X1. + now apply All_local_env_app_inv in X1. Qed. Require Import MetaCoq.PCUIC.PCUICInversion. @@ -183,7 +183,7 @@ Proof. rewrite lift_context_snoc0, <- plus_n_O in *. eapply H1; eauto. cbn. econstructor. eauto. cbn. exists s1. eapply weakening_typing with (T := tSort s1); eauto. - now apply All_local_env_app in X2. + now apply All_local_env_app_inv in X2. - econstructor. + eapply H0; eauto. + pose proof (H1 Γ (PCUICAst.vdef n b b_ty :: Γ') Γ''). @@ -191,8 +191,8 @@ Proof. eapply H2; eauto. cbn. econstructor. eauto. cbn. 2: cbn; eapply weakening_typing; eauto. eapply weakening_typing in X0; eauto. - now apply All_local_env_app in X3. - now apply All_local_env_app in X3. + now apply All_local_env_app_inv in X3. + now apply All_local_env_app_inv in X3. - econstructor. + eauto. + eapply H4; eauto. @@ -206,11 +206,12 @@ Proof. split; eauto. - assert (HT : Σ;;; Γ ,,, Γ' |- PCUICAst.tFix mfix n : (decl.(dtype))). econstructor; eauto. eapply All_impl. eassumption. intros. - destruct X4; cbn in *; firstorder. - eapply (All_impl X1). firstorder. + destruct X4; cbn in *; pcuicfo. + exists x0; auto. + eapply (All_impl X1). pcuicfo. eapply weakening_typing in HT; auto. - 2:{ apply All_local_env_app in X2 as [X2 _]. eapply X2. } + 2:{ apply All_local_env_app_inv in X2 as [X2 _]. eapply X2. } cbn in HT. eapply inversion_Fix in HT as (? & ? & ? & ? & ? & ? & ?) ; auto. clear a0 c. @@ -238,11 +239,12 @@ Proof. - assert (HT : Σ;;; Γ ,,, Γ' |- PCUICAst.tCoFix mfix n : (decl.(dtype))). econstructor; eauto. eapply All_impl. eassumption. intros. - destruct X4; cbn in *; firstorder. - eapply (All_impl X1). firstorder. + destruct X4; cbn in *; pcuicfo. + now exists x0. + eapply (All_impl X1). pcuicfo. eapply weakening_typing in HT; auto. - 2:{ apply All_local_env_app in X2 as [X2 _]. eapply X2. } + 2:{ apply All_local_env_app_inv in X2 as [X2 _]. eapply X2. } cbn in HT. eapply inversion_CoFix in HT as (? & ? & ? & ? & ? & ? & ?) ; auto. clear a0 c. @@ -397,10 +399,10 @@ Proof. eapply substitution; eauto. + econstructor. eapply is_type_subst; eauto. - - inv H1. + - inv H2. + cbn. econstructor. - eapply H; eauto. eapply H0; eauto. + eapply H1; eauto. + econstructor. eapply is_type_subst; eauto. - inv H1. @@ -457,8 +459,8 @@ Proof. * cbn. now rewrite app_context_length, fix_context_length. * cbn. now erewrite app_context_length, fix_context_length, All2_length. * pose proof (substitution Σ Γ Γ' s (Δ ,,, PCUICLiftSubst.fix_context mfix)). - rewrite app_context_assoc in *. destruct Hs. - eapply X1 in t; eauto. + rewrite app_context_assoc in *. + eapply X1 in Hs; eauto. eapply typing_wf_local. eassumption. + econstructor. eapply is_type_subst; eauto. diff --git a/erasure/theories/Erasure.v b/erasure/theories/Erasure.v index 04df96f6a..b3a5baf10 100644 --- a/erasure/theories/Erasure.v +++ b/erasure/theories/Erasure.v @@ -1,139 +1,22 @@ (* Distributed under the terms of the MIT license. *) From Coq Require Import Program. -From MetaCoq.Template Require Import config utils uGraph Pretty. +From MetaCoq.Template Require Import config utils uGraph Pretty Environment Typing. +Set Warnings "-notation-overridden". From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping - TemplateToPCUIC. -From MetaCoq.SafeChecker Require Import PCUICSafeReduce PCUICSafeChecker - SafeTemplateChecker. + TemplateToPCUIC TemplateToPCUICCorrectness. +Set Warnings "+notation-overridden". +From MetaCoq.SafeChecker Require Import PCUICErrors. From MetaCoq.Erasure Require Import EAstUtils ErasureFunction EPretty. From MetaCoq.Erasure Require ErasureFunction EOptimizePropDiscr. -Existing Instance envcheck_monad. Existing Instance extraction_checker_flags. -(** This is a hack to avoid having to handle template polymorphism and make - erasure fast: we actually admit the proof of wf Σ and never build it. *) - -Definition assume_wf_decl {cf : checker_flags} (Σ : global_env_ext) : - ∥ wf Σ ∥ -> - ∥ on_udecl Σ.1 Σ.2 ∥ -> - forall G : universes_graph, - is_graph_of_uctx G (global_ext_uctx Σ) -> - forall kn (d : global_decl), EnvCheck (∥ on_global_decl (lift_typing typing) Σ kn d ∥). -Proof. - intros. apply CorrectDecl. constructor. todo "assumed correct global declaration". -Defined. - -Definition assume_fresh id env : EnvCheck (∥ fresh_global id env ∥). -Proof. - left. todo "assumed fresh". -Defined. - -Program Definition compute_udecl (id : string) (Σ : global_env) (HΣ : ∥ wf Σ ∥) G - (HG : is_graph_of_uctx G (global_uctx Σ)) (udecl : universes_decl) - : EnvCheck (∑ uctx', gc_of_uctx (uctx_of_udecl udecl) = Some uctx' /\ - ∥ on_udecl Σ udecl ∥) := - match gc_of_uctx (uctx_of_udecl udecl) with - | Some uctx => ret (uctx; conj _ _) - | None => raise (empty_ext Σ, IllFormedDecl id (Msg "constraints not satisfiable")) - end. - Next Obligation. - constructor. todo "assume udecl is ok". - Defined. - -Program Fixpoint check_wf_env_only_univs (Σ : global_env) - : EnvCheck (∑ G, (is_graph_of_uctx G (global_uctx Σ) /\ ∥ wf Σ ∥)) := - match Σ with - | nil => ret (init_graph; _) - | d :: Σ => - G <- check_wf_env_only_univs Σ ;; - assume_fresh d.1 Σ ;; - let udecl := universes_decl_of_decl d.2 in - uctx <- compute_udecl (string_of_kername d.1) Σ _ G.π1 (proj1 G.π2) udecl ;; - let G' := add_uctx uctx.π1 G.π1 in - assume_wf_decl (Σ, udecl) _ _ G' _ d.1 d.2 ;; - match udecl with - | Monomorphic_ctx _ => ret (G'; _) - | Polymorphic_ctx _ => ret (G.π1; _) - end - end. - Next Obligation. - repeat constructor. - Qed. - Next Obligation. - sq. unfold is_graph_of_uctx, gc_of_uctx; simpl. - unfold gc_of_uctx in e. simpl in e. - case_eq (gc_of_constraints (constraints_of_udecl (universes_decl_of_decl g))); - [|intro HH; rewrite HH in e; discriminate e]. - intros ctrs' Hctrs'. rewrite Hctrs' in e. - cbn in e. inversion e; subst; clear e. - unfold global_ext_constraints; simpl. - pose proof (gc_of_constraints_union - (constraints_of_udecl (universes_decl_of_decl g)) (global_constraints Σ)). - rewrite Hctrs' in H0. simpl in H0. - red in i. unfold gc_of_uctx in i; simpl in i. - case_eq (gc_of_constraints (global_constraints Σ)); - [|intro HH; rewrite HH in i; cbn in i; contradiction i]. - intros Σctrs HΣctrs; rewrite HΣctrs in H0, i; simpl in *. - destruct (gc_of_constraints (ConstraintSet.union _ _)). - simpl in H0. - subst G. unfold global_ext_levels; simpl. - symmetry. rewrite add_uctx_make_graph. - apply graph_eq. simpl. reflexivity. - simpl. now rewrite H0. simpl. reflexivity. - now simpl in H0. - Qed. - Next Obligation. - split; sq. 2: constructor; tas. - unfold is_graph_of_uctx, gc_of_uctx; simpl. - unfold gc_of_uctx in e. simpl in e. - case_eq (gc_of_constraints (constraints_of_udecl (universes_decl_of_decl g))); - [|intro HH; rewrite HH in e; discriminate e]. - intros ctrs' Hctrs'. rewrite Hctrs' in e. - cbn in e. inversion e; subst; clear e. - unfold global_ext_constraints; simpl. - pose proof (gc_of_constraints_union - (constraints_of_udecl (universes_decl_of_decl g)) (global_constraints Σ)). - rewrite Hctrs' in H1; simpl in H1. - red in i. unfold gc_of_uctx in i; simpl in i. - assert (eq: monomorphic_constraints_decl g - = constraints_of_udecl (universes_decl_of_decl g)). { - destruct g. destruct c, cst_universes; try discriminate; reflexivity. - destruct m, ind_universes; try discriminate; reflexivity. } - rewrite eq; clear eq. - case_eq (gc_of_constraints (global_constraints Σ)); - [|intro HH; rewrite HH in i; cbn in i; contradiction i]. - intros Σctrs HΣctrs; rewrite HΣctrs in H1, i; simpl in *. - destruct (gc_of_constraints (ConstraintSet.union _ _)). - simpl in H1. - subst G. unfold global_ext_levels; simpl. - assert (eq: monomorphic_levels_decl g - = levels_of_udecl (universes_decl_of_decl g)). { - destruct g. destruct c, cst_universes; try discriminate; reflexivity. - destruct m, ind_universes; try discriminate; reflexivity. } - rewrite eq. simpl. rewrite add_uctx_make_graph. - apply graph_eq; try reflexivity. - simpl. now rewrite H1. - now simpl in H1. - Qed. - Next Obligation. - split; sq. 2: constructor; tas. - unfold global_uctx; simpl. - assert (eq1: monomorphic_levels_decl g = LevelSet.empty). { - destruct g. destruct c, cst_universes; try discriminate; reflexivity. - destruct m, ind_universes; try discriminate; reflexivity. } - rewrite eq1; clear eq1. - assert (eq1: monomorphic_constraints_decl g = ConstraintSet.empty). { - destruct g. destruct c, cst_universes; try discriminate; reflexivity. - destruct m, ind_universes; try discriminate; reflexivity. } - rewrite eq1; clear eq1. - now rewrite levelset_union_empty, constraintset_union_empty. - Qed. - (* This is the total erasure function + the optimization that removes all pattern-matches on propositions. *) Program Definition erase_template_program (p : Ast.program) + (wfΣ : ∥ Typing.wf_ext (Ast.empty_ext p.1) ∥) + (wt : ∥ ∑ T, Typing.typing (Ast.empty_ext p.1) [] p.2 T ∥) : (EAst.global_context * EAst.term) := let Σ := (trans_global (Ast.empty_ext p.1)).1 in let t := ErasureFunction.erase (empty_ext Σ) _ nil (trans p.2) _ in @@ -141,25 +24,28 @@ Program Definition erase_template_program (p : Ast.program) (EOptimizePropDiscr.optimize_env Σ', EOptimizePropDiscr.optimize Σ' t). Next Obligation. - unfold trans_global. - simpl. unfold wf_ext, empty_ext. simpl. - unfold on_global_env_ext. constructor. todo "assuming wf environment". -Defined. + sq. + apply (template_to_pcuic_env_ext (Ast.empty_ext p.1) wfΣ). +Qed. Next Obligation. - unfold trans_global. - simpl. unfold wf_ext, empty_ext. simpl. - unfold on_global_env_ext. todo "assuming well-typedness". + sq. destruct wt as [T Ht]. exists (trans T). + change (@nil context_decl) with (trans_local []). + change (empty_ext (trans_global_decls p.1)) with (trans_global (Ast.empty_ext p.1)). + eapply template_to_pcuic_typing; simpl. apply wfΣ. + apply Ht. Defined. Next Obligation. - constructor. todo "assuming wf environment". + sq. apply (template_to_pcuic_env_ext (Ast.empty_ext p.1) wfΣ). Defined. Local Open Scope string_scope. -(** This uses the retyping-based erasure *) +(** This uses the retyping-based erasure and assumes that the global environment and term + are welltyped (for speed). As such this should only be used for testing, or when we know that + the environment is wellformed and the term well-typed (e.g. when it comes directly from a + Coq definition). *) Program Definition erase_and_print_template_program {cf : checker_flags} (p : Ast.program) : string := - let p := fix_program_universes p in - let (Σ', t) := erase_template_program p in - "Environment is well-formed and " ^ Pretty.print_term (Ast.empty_ext p.1) [] true p.2 ^ + let (Σ', t) := erase_template_program p (todo "wf_env") (todo "welltyped") in + Pretty.print_term (Ast.empty_ext p.1) [] true p.2 ^ nl ^ " erases to: " ^ nl ^ print_term Σ' [] true false t. diff --git a/erasure/theories/ErasureCorrectness.v b/erasure/theories/ErasureCorrectness.v index 1087a20be..2112a51a2 100644 --- a/erasure/theories/ErasureCorrectness.v +++ b/erasure/theories/ErasureCorrectness.v @@ -86,17 +86,17 @@ Proof. - econstructor. - econstructor; eauto. cbn in *. destruct t0. exists x. eapply PCUICContextConversion.context_conversion with (Γ ,,, Γ0); eauto. - * eapply PCUICSafeChecker.wf_local_app_inv; eauto. + * eapply wf_local_app; eauto. * eapply conv_context_app; eauto. eapply typing_wf_local; eauto. - econstructor; eauto. + cbn in *. destruct t0. exists x. eapply PCUICContextConversion.context_conversion with (Γ ,,, Γ0); eauto. - * eapply PCUICSafeChecker.wf_local_app_inv; eauto. + * eapply wf_local_app; eauto. * eapply conv_context_app; eauto. eapply typing_wf_local; eauto. + cbn in *. eapply PCUICContextConversion.context_conversion with (Γ ,,, Γ0); eauto. - * eapply PCUICSafeChecker.wf_local_app_inv; eauto. + * eapply wf_local_app; eauto. * eapply conv_context_app; eauto. eapply typing_wf_local; eauto. Qed. @@ -131,7 +131,8 @@ Proof. eapply All2_All_left in X3. 2:{ idtac. intros ? ? [[? e] ?]. exact e. } eapply All2_impl. eapply All2_All_mix_left. - all: firstorder eauto with *. + eauto. eauto. + all: pcuicfo. - econstructor. eapply All2_impl. eapply All2_All_mix_left. eassumption. eassumption. @@ -140,11 +141,11 @@ Proof. eapply b0. subst types. eapply conv_context_app; auto. eapply typing_wf_local; eassumption. - eapply typing_wf_local in a1. subst types. + eapply typing_wf_local in a0. subst types. 2:eauto. - eapply All_local_env_app_inv. - eapply All_local_env_app in a1. intuition auto. + eapply All_local_env_app. + eapply All_local_env_app_inv in a0. intuition auto. (* clear -wfΣ X2 a2 b4 X1. *) eapply All_local_env_impl; eauto. simpl; intros. @@ -152,18 +153,18 @@ Proof. eapply PCUICContextConversion.context_conversion with (Γ ,,, Γ0); eauto. 2:{ eapply conv_context_app; auto. eapply typing_wf_local; eauto. } eapply typing_wf_local in X3. - eapply PCUICSafeChecker.wf_local_app_inv. + eapply wf_local_app. eauto. eapply wf_local_rel_local in X3. - eapply wf_local_rel_app in X3 as []. rewrite app_context_nil_l in w0. + eapply wf_local_rel_app_inv in X3 as []. rewrite app_context_nil_l in w0. eapply wf_local_rel_conv; eauto. destruct X3. exists x0. eapply PCUICContextConversion.context_conversion with (Γ ,,, Γ0); eauto. 2:{ eapply conv_context_app; auto. eapply typing_wf_local; eauto. } eapply typing_wf_local in t0. - eapply PCUICSafeChecker.wf_local_app_inv. + eapply wf_local_app. eauto. eapply wf_local_rel_local in t0. - eapply wf_local_rel_app in t0 as []. rewrite app_context_nil_l in w0. + eapply wf_local_rel_app_inv in t0 as []. rewrite app_context_nil_l in w0. eapply wf_local_rel_conv; eauto. - econstructor. @@ -176,8 +177,8 @@ Proof. eapply typing_wf_local in a0. subst types. 2:eauto. - eapply All_local_env_app_inv. - eapply All_local_env_app in a0. intuition auto. + eapply All_local_env_app. + eapply All_local_env_app_inv in a0. intuition auto. (* clear -wfΣ X2 a2 b4 X1. *) eapply All_local_env_impl; eauto. simpl; intros. @@ -185,9 +186,9 @@ Proof. eapply PCUICContextConversion.context_conversion with (Γ ,,, Γ0); eauto. 2:{ eapply conv_context_app; auto. eapply typing_wf_local; eauto. } eapply typing_wf_local in X3. - eapply PCUICSafeChecker.wf_local_app_inv. + eapply wf_local_app. eauto. eapply wf_local_rel_local in X3. - eapply wf_local_rel_app in X3 as []. rewrite app_context_nil_l in w0. + eapply wf_local_rel_app_inv in X3 as []. rewrite app_context_nil_l in w0. eapply wf_local_rel_conv; eauto. @@ -196,9 +197,9 @@ Proof. 2:{ eapply conv_context_app; auto. eapply typing_wf_local; eauto. } eapply typing_wf_local in t0. - eapply PCUICSafeChecker.wf_local_app_inv. + eapply wf_local_app. eauto. eapply wf_local_rel_local in t0. - eapply wf_local_rel_app in t0 as []. rewrite app_context_nil_l in w0. + eapply wf_local_rel_app_inv in t0 as []. rewrite app_context_nil_l in w0. eapply wf_local_rel_conv; eauto. Qed. @@ -258,7 +259,8 @@ Proof. - assert (Hw : wf_local (Σ.1, univs) (subst_instance_context u (Γ ,,, types))). { (* rewrite subst_instance_context_app. *) assert(All (fun d => isType Σ Γ (dtype d)) mfix). - eapply (All_impl X0); firstorder. + eapply (All_impl X0); pcuicfo. + now destruct X5 as [s [Hs ?]]; exists s. eapply All_mfix_wf in X5; auto. subst types. revert X5. clear - wfΣ wfΓ H2 H3 X2 X3. @@ -281,7 +283,7 @@ Proof. eapply All2_impl. eapply All2_All_mix_left. eapply X1. exact X4. intros; cbn in *. destruct X5. destruct p0. destruct p0. - destruct p. destruct p. repeat split; eauto. + destruct p. repeat split; eauto. eapply e2 in e1. unfold PCUICUnivSubst.subst_instance_context in *. unfold map_context in *. rewrite ->map_app in *. subst types. 2:eauto. @@ -292,7 +294,8 @@ Proof. - assert (Hw : wf_local (Σ.1, univs) (subst_instance_context u (Γ ,,, types))). { (* rewrite subst_instance_context_app. *) assert(All (fun d => isType Σ Γ (dtype d)) mfix). - eapply (All_impl X0); firstorder. + eapply (All_impl X0); pcuicfo. + destruct X5 as [s [Hs ?]]; now exists s. eapply All_mfix_wf in X5; auto. subst types. revert X5. clear - wfΣ wfΓ H2 H3 X2 X3. @@ -426,14 +429,7 @@ Qed. (** ** The correctness proof *) -Record extraction_pre (Σ : global_env_ext) : Type - := Build_extraction_pre - { extr_env_axiom_free' : axiom_free (fst Σ); - extr_env_wf' : wf_ext Σ }. - Hint Constructors PCUICWcbvEval.eval erases : core. -Arguments extr_env_wf' {Σ}. -Arguments extr_env_axiom_free' {Σ}. Definition EisConstruct_app := fun t => match (EAstUtils.decompose_app t).1 with @@ -679,16 +675,16 @@ Proof. Qed. Lemma erases_correct (wfl := default_wcbv_flags) Σ t T t' v Σ' : - extraction_pre Σ -> + wf_ext Σ -> Σ;;; [] |- t : T -> Σ;;; [] |- t ⇝ℇ t' -> erases_deps Σ Σ' t' -> Σ |-p t ▷ v -> exists v', Σ;;; [] |- v ⇝ℇ v' /\ ∥ Σ' ⊢ t' ▷ v' ∥. Proof. - intros pre Hty He Hed H. + intros wfΣ Hty He Hed H. revert T Hty t' He Hed. - induction H; intros T Hty t' He Hed; destruct pre as [axfree wfΣ]. + induction H; intros T Hty t' He Hed. - assert (Hty' := Hty). assert (eval Σ (PCUICAst.tApp f a) res) by eauto. eapply inversion_App in Hty as (? & ? & ? & ? & ? & ?). @@ -802,10 +798,6 @@ Proof. + exists EAst.tBox. split. econstructor. eapply Is_type_eval. 3: eassumption. eauto. eauto. eauto. constructor. econstructor. eauto. - - destruct Σ as (Σ, univs). - cbn in *. - eapply axfree in isdecl. congruence. - - assert (Hty' := Hty). assert (Σ |-p tCase (ind, pars) p discr brs ▷ res) by eauto. eapply inversion_Case in Hty' as [u' [args' [mdecl [idecl [ps [pty [btys @@ -903,7 +895,7 @@ Proof. eapply subject_reduction. eauto. exact Hty. etransitivity. eapply PCUICReduction.red_case_c. eapply wcbeval_red. eauto. - now eapply PCUICClosed.subject_closed in t0. eauto. eauto. + eauto. eauto. etransitivity. constructor. constructor. unfold iota_red. rewrite <- nth_default_eq. unfold nth_default. @@ -946,7 +938,7 @@ Proof. eapply subject_reduction. eauto. exact Hty. etransitivity. eapply PCUICReduction.red_case_c. eapply wcbeval_red. eauto. - now eapply PCUICClosed.subject_closed in t0; eauto. eauto. eauto. + eauto. eauto. etransitivity. constructor. constructor. unfold iota_red. rewrite <- nth_default_eq. reflexivity. @@ -1059,7 +1051,7 @@ Proof. eapply erases_deps_mkApps_inv in Hty_vc' as (? & ?). now eapply nth_error_forall in H1; eauto. + exists EAst.tBox. split. econstructor. - eapply Is_type_eval. 4: eassumption. all:eauto. constructor. econstructor. eauto. + eapply Is_type_eval. 3: eassumption. all:eauto. constructor. econstructor. eauto. - assert (Hty' := Hty). assert (Hunf := H). @@ -1077,7 +1069,7 @@ Proof. + exists EAst.tBox. split; [|now constructor; constructor]. econstructor. - eapply Is_type_eval. 4:eapply X. eauto. eauto. + eapply Is_type_eval. 3:eapply X. eauto. eapply eval_fix; eauto. rewrite /cunfold_fix e0 //. congruence. + depelim Hed. @@ -1094,7 +1086,7 @@ Proof. destruct H2. eapply (Is_type_app _ _ _ (x5 ++ [av])) in X as []; eauto; first last. - rewrite mkApps_nested app_assoc mkApps_snoc. - eapply type_App; eauto. + eapply PCUICValidity.type_App'; eauto. eapply subject_reduction; eauto. eapply wcbeval_red; eauto. - eapply erases_box. @@ -1146,11 +1138,13 @@ Proof. subst. unfold is_constructor. rewrite nth_error_snoc. lia. assert(Σ ;;; [] |- mkApps (tFix mfix idx) (argsv ++ [av]) : PCUICLiftSubst.subst [av] 0 x1). - { rewrite -mkApps_nested. eapply type_App; eauto. eapply subject_reduction_eval;eauto. } - epose proof (fix_app_is_constructor Σ (args:=argsv ++ [av]) axfree X). + { rewrite -mkApps_nested. eapply PCUICValidity.type_App'; eauto. + eapply subject_reduction_eval; eauto. } + epose proof (fix_app_is_constructor Σ (args:=argsv ++ [av]) X). rewrite /unfold_fix e0 in X0. specialize (X0 eq_refl). simpl in X0. rewrite nth_error_snoc in X0. auto. apply X0. + eapply value_axiom_free, eval_to_value; eauto. eapply value_whnf; eauto. eapply eval_closed; eauto. now eapply PCUICClosed.subject_closed in t0. eapply eval_to_value; eauto. } @@ -1186,7 +1180,7 @@ Proof. simpl in b. rewrite fix_context_length in b. now rewrite Nat.add_0_r. - unfold test_def in a. apply andP in a as [_ Hbod]. + unfold test_def in a. apply andb_and in a as [_ Hbod]. rewrite fix_context_length. now rewrite Nat.add_0_r in Hbod. eauto with pcuic. @@ -1210,7 +1204,7 @@ Proof. -- exists EAst.tBox. split. ++ econstructor. - eapply Is_type_eval. 4:eauto. all:eauto. + eapply Is_type_eval. 3:eauto. all:eauto. rewrite -mkApps_nested. eapply eval_fix; eauto. 1-2:eapply value_final, eval_to_value; eauto. @@ -1220,7 +1214,7 @@ Proof. eauto. -- eauto. -- rewrite mkApps_snoc. - eapply type_App; eauto. + eapply PCUICValidity.type_App'; eauto. eapply subject_reduction; eauto. eapply wcbeval_red; eauto. @@ -1228,7 +1222,7 @@ Proof. destruct Hty' as (? & ? & ? & ? & ? & ?). eapply subject_reduction in t as typ_stuck_fix; [|eauto|]; first last. - { eapply wcbeval_red. 4:eauto. all:eauto. } + { eapply wcbeval_red. 3:eauto. all:eauto. } eapply erases_App in He as He'; [|eauto]. destruct He' as [(-> & [])|(? & ? & -> & ? & ?)]. @@ -1238,12 +1232,12 @@ Proof. eapply Is_type_red. * eauto. * eapply PCUICReduction.red_app. - -- eapply wcbeval_red; [eauto|eauto| |eauto]. eauto. - -- eapply wcbeval_red; [eauto|eauto| |eauto]. eauto. + -- eapply wcbeval_red; [eauto| |eauto]. eauto. + -- eapply wcbeval_red; [eauto| |eauto]. eauto. * eauto. + depelim Hed. eapply subject_reduction in t0 as typ_arg; [|eauto|]; first last. - { eapply wcbeval_red; [eauto|eauto| |eauto]. eauto. } + { eapply wcbeval_red; [eauto| |eauto]. eauto. } eapply IHeval1 in H1 as (? & ? & [?]); [|now eauto|now eauto]. eapply IHeval2 in H2 as (? & ? & [?]); [|now eauto|now eauto]. @@ -1261,7 +1255,7 @@ Proof. -- eauto. -- eauto. -- cbn. - eapply type_App; eauto. + eapply PCUICValidity.type_App'; eauto. * depelim H1. -- exists (E.tApp (E.mkApps (E.tFix mfix' idx) x7) x5). split; [eauto using erases_tApp, erases_mkApps|]. @@ -1286,7 +1280,7 @@ Proof. ++ eauto. ++ eauto. ++ rewrite mkApps_snoc. - eapply type_App; eauto. + eapply PCUICValidity.type_App'; eauto. - destruct ip. assert (Hty' := Hty). diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 035b9fa4f..247fd006d 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -1,15 +1,15 @@ (* Distributed under the terms of the MIT license. *) From Coq Require Import Program. From MetaCoq.Template Require Import config utils Kernames. -From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICPrimitive PCUICReflect PCUICWeakeningEnv PCUICTyping PCUICInversion PCUICGeneration PCUICConfluence PCUICConversion PCUICCumulativity PCUICSR PCUICSafeLemmata PCUICValidity PCUICPrincipality PCUICElimination PCUICSN. -From MetaCoq.SafeChecker Require Import PCUICSafeReduce PCUICSafeChecker PCUICSafeRetyping. -From MetaCoq.Erasure Require Import EAstUtils EArities Extract Prelim ErasureCorrectness EDeps. +From MetaCoq.SafeChecker Require Import PCUICErrors PCUICSafeReduce PCUICSafeRetyping. +From MetaCoq.Erasure Require Import EAstUtils EArities Extract Prelim EDeps ErasureCorrectness. Local Open Scope string_scope. Set Asymmetric Patterns. @@ -20,7 +20,7 @@ Set Equations Transparent. Local Set Keyed Unification. Require Import ssreflect. -Derive Signature for on_global_env. +Derive Signature for PCUICTyping.on_global_env. (* To debug performance issues *) (* @@ -366,6 +366,11 @@ Section Erase. End EraseMfix. + Equations erase_prim (ep : prim_val term) : PCUICPrimitive.prim_val E.term := + erase_prim (_; primIntModel i) := (_; primIntModel i); + erase_prim (_; primFloatModel f) := (_; primFloatModel f). + + Equations? (noeqns noind) erase (Γ : context) (t : term) (Ht : welltyped Σ Γ t) : E.term by struct t := erase Γ t Ht with (is_erasable Σ HΣ Γ t Ht) := @@ -403,7 +408,8 @@ Section Erase. E.tFix mfix' n; erase Γ (tCoFix mfix n) Ht _ := let mfix' := erase_mfix (erase) Γ mfix _ in - E.tCoFix mfix' n + E.tCoFix mfix' n; + erase Γ (tPrim p) Ht _ := E.tPrim (erase_prim p) }. Proof. all:try clear b'; try clear f'; try clear brs'; try clear erase. @@ -431,7 +437,7 @@ Section Erase. eexists; eauto. - eapply inversion_Fix in Ht as (? & ? & ? & ? & ? & ?); auto. eapply All_In in a0; eauto. - destruct a0 as [[a0 _]]. eexists; eauto. + destruct a0 as [a0]. eexists; eauto. - eapply inversion_CoFix in Ht as (? & ? & ? & ? & ? & ?); auto. eapply All_In in a0; eauto. destruct a0 as [a0]. eexists; eauto. @@ -794,7 +800,7 @@ Proof. eapply Forall2_All2 in H. eapply All2_All_mix_left in H; eauto. eapply All2_In_right in H; eauto. - destruct H as [[def [[Hty _] Hdef]]]. + destruct H as [[def [Hty Hdef]]]. eapply Hdef; eauto. - apply inversion_CoFix in wt as (?&?&?&?&?&?&?); eauto. @@ -1117,6 +1123,7 @@ Proof. sq. destruct X0 as [bod [bodty [[Hbod Hebod] Heqdeps]]]. eapply (erase_global_erases_deps (Σ := (Σ, cst_universes c))); simpl; auto. constructor; simpl; auto. + red in w. depelim w. auto. eauto. eauto. eapply IHΣ. @@ -1180,7 +1187,6 @@ Proof. Qed. Lemma erase_correct (wfl := Ee.default_wcbv_flags) (Σ : global_env_ext) (wfΣ : wf_ext Σ) t v Σ' t' deps : - axiom_free Σ.1 -> forall wt : welltyped Σ [] t, erase Σ (sq wfΣ) [] t wt = t' -> KernameSet.subset (term_global_deps t') deps -> @@ -1188,10 +1194,9 @@ Lemma erase_correct (wfl := Ee.default_wcbv_flags) (Σ : global_env_ext) (wfΣ : Σ |-p t ▷ v -> exists v', Σ;;; [] |- v ⇝ℇ v' /\ ∥ Σ' ⊢ t' ▷ v' ∥. Proof. - intros axiomfree wt. + intros wt. generalize (sq wfΣ.1) as swfΣ. intros swfΣ HΣ' Hsub Ht' ev. - assert (extraction_pre Σ) by now constructor. pose proof (erases_erase (wfΣ := sq wfΣ) wt); eauto. rewrite HΣ' in H. destruct wt as [T wt]. diff --git a/erasure/theories/Extraction.v b/erasure/theories/Extraction.v index ef2f47566..df5ad6dd9 100644 --- a/erasure/theories/Extraction.v +++ b/erasure/theories/Extraction.v @@ -1,5 +1,6 @@ (* Distributed under the terms of the MIT license. *) -Require Import FSets ExtrOcamlBasic ExtrOcamlString ExtrOcamlZInt. +From Coq Require Import FSets ExtrOcamlBasic ExtrOcamlString ExtrOcamlZInt ExtrOCamlFloats. +From MetaCoq.Template Require Import MC_ExtrOCamlInt63 (*b/c nameclash with `comparion` *). (** * Extraction setup for the erasure phase of template-coq. @@ -38,10 +39,7 @@ Extract Constant Equations.Init.pr1 => "fst". Extract Constant Equations.Init.pr2 => "snd". Extraction Inline Equations.Init.pr1 Equations.Init.pr2. -Extract Constant PCUICTyping.fix_guard => "(fun x -> true)". -Extract Constant PCUICTyping.cofix_guard => "(fun x -> true)". -Extract Constant PCUICTyping.ind_guard => "(fun x -> true)". -(* Extract Constant erase_mfix_obligation_1 => "(fun _ _ _ _ => ret typing_monad __)". *) +Extract Constant PCUICTyping.guard_checking => "{ fix_guard = (fun _ _ _ -> true); cofix_guard = (fun _ _ _ -> true) }". Cd "src". diff --git a/erasure/theories/Prelim.v b/erasure/theories/Prelim.v index 75dc978cb..042748fdb 100644 --- a/erasure/theories/Prelim.v +++ b/erasure/theories/Prelim.v @@ -5,7 +5,7 @@ From MetaCoq.PCUIC Require Import PCUICTyping PCUICAst PCUICAstUtils PCUICSubstitution PCUICLiftSubst PCUICClosed PCUICWcbvEval PCUICSR PCUICInversion PCUICGeneration PCUICContextConversion PCUICCanonicity. -From MetaCoq.SafeChecker Require Import PCUICSafeReduce PCUICSafeChecker. +From MetaCoq.SafeChecker Require Import PCUICErrors. From Coq Require Import Program ssreflect. @@ -16,8 +16,6 @@ Module P := PCUICWcbvEval. Ltac inv H := inversion H; subst; clear H. -Existing Class axiom_free. - Lemma nth_error_app_inv X (x : X) n l1 l2 : nth_error (l1 ++ l2) n = Some x -> (n < #|l1| /\ nth_error l1 n = Some x) \/ (n >= #|l1| /\ nth_error l2 (n - List.length l1) = Some x). Proof. @@ -121,7 +119,7 @@ Proof. now exists A. Qed. -Theorem subject_reduction_eval {Σ :global_env_ext} {t u T} {wfΣ : wf Σ} {axfree : axiom_free Σ} : +Theorem subject_reduction_eval {Σ :global_env_ext} {t u T} {wfΣ : wf Σ} : Σ ;;; [] |- t : T -> PCUICWcbvEval.eval Σ t u -> Σ ;;; [] |- u : T. Proof. intros Hty Hred. @@ -133,7 +131,7 @@ Lemma typing_spine_eval: (X : All2 (PCUICWcbvEval.eval Σ) args args') (bla : wf Σ) (T x x0 : PCUICAst.term) (t0 : typing_spine Σ [] x args x0) (c : Σ;;; [] |- x0 <= T) (x1 : PCUICAst.term) - (c0 : Σ;;; [] |- x1 <= x), axiom_free Σ -> isType Σ [] T -> typing_spine Σ [] x1 args' T. + (c0 : Σ;;; [] |- x1 <= x), isType Σ [] T -> typing_spine Σ [] x1 args' T. Proof. intros. eapply typing_spine_red; eauto. eapply typing_spine_wt in t0; auto. @@ -267,7 +265,7 @@ Proof. rewrite emkApps_snoc in i. inv i. - destruct (EAstUtils.mkApps_elim t l). EAstUtils.solve_discr. rewrite Ee.value_head_spec in i. - move/andP: i => [H H']. + move/andb_and: i => [H H']. eapply Ee.atom_mkApps in H' as [H1 _]. destruct n, L; discriminate. - unfold Ee.isStuckFix in i. destruct f; try now inversion i. diff --git a/examples/_CoqProject b/examples/_CoqProject index 1c889d23d..c68a2627f 100644 --- a/examples/_CoqProject +++ b/examples/_CoqProject @@ -1,11 +1,15 @@ -I ../template-coq/build +-I ../safechecker/src +-I ../erasure/src -Q ../template-coq/theories MetaCoq.Template --Q ../checker/theories MetaCoq.Checker -Q ../pcuic/theories MetaCoq.PCUIC -Q ../safechecker/theories MetaCoq.SafeChecker +-Q ../erasure/theories MetaCoq.Erasure -R . MetaCoq.Examples demo.v add_constructor.v tauto.v -typing_correctness.v \ No newline at end of file +typing_correctness.v +metacoq_tour_prelude.v +metacoq_tour.v \ No newline at end of file diff --git a/examples/metacoq_tour.v b/examples/metacoq_tour.v new file mode 100644 index 000000000..7231d35ed --- /dev/null +++ b/examples/metacoq_tour.v @@ -0,0 +1,149 @@ +From Coq Require Import String. +From MetaCoq.Template Require config utils All. +From MetaCoq.Template Require Import TemplateMonad. +From MetaCoq.PCUIC Require Import PCUICAst PCUICReduction PCUICCumulativity PCUICTyping PCUICSafeLemmata. + +Local Open Scope string_scope. + +(** MetaCoq is: + + - The "template-coq" monad, dealing with reification of terms + and environments. + - The PCUIC development of the syntactic metatheory of Coq. + - The SafeChecker package implementing reduction, conversion + and typechecking in a sound and complete way. + - The Erasure package implementing verified extraction to + untyped lambda-calculus +*) + +(* The Template monad *) + +Print Ast.term. + + +MetaCoq Quote Definition reifx := (fun x : nat => x). +Print reifx. +MetaCoq Unquote Definition x := + (Ast.tConstruct (mkInd (MPfile ["Datatypes"; "Init"; "Coq"], "nat") 0) 0 []). + +MetaCoq Run (tmBind (tmQuote (3 + 3)) tmPrint). + +MetaCoq Run (t <- tmLemma "foo44" nat ;; + qt <- tmQuote t ;; + t <- tmEval all t ;; + tmPrint qt ;; tmPrint t). +Next Obligation. + exact (3+2). +Defined. + +(* PCUIC: + + + Universes (with full algebraic universes supported everywhere) + + Universe polymorphism + + Standard type theory with dependent products and let-ins + + Inductive types and coinductive types + + Constant and axiom declarations in a global environment *) + +Print typing. + +Check type_Sort. +Check type_LetIn. +Check type_Const. + +From MetaCoq.PCUIC Require PCUICSR. + +Check PCUICSR.subject_reduction. + +(** Verified conversion and type-checking *) + +From MetaCoq.SafeChecker Require Import PCUICErrors PCUICTypeChecker PCUICSafeChecker PCUICSafeRetyping Loader. +From MetaCoq.Erasure Require Import Erasure Loader. + +Check PCUICSafeConversion.isconv_term_sound. +Check PCUICSafeConversion.isconv_term_complete. + +Check PCUICSafeChecker.infer_wf_env. +(** Proof of completeness is near completion. *) + +(** Verified retyping: from a term that is known to be well-typeable, + compute its principal type. Very common in tactics to avoid retypechecking + the whole term. *) +Check type_of. +Check type_of_subtype. + +(* Running the safe checker inside Coq *) +From MetaCoq.Examples Require Import metacoq_tour_prelude. + +Check check_inh. + +(** We construct a proof of typing entirely within Coq, calling the typechecker to produce the derivation *) +Lemma identity_typing (u := Universe.make univ): + inh gctx_wf_env [] (tProd (bNamed "s") (tSort u) (tImpl (tRel 0) (tRel 0))). +Proof. + (* We construct a term *) + set (impl := tLambda (bNamed "s") (tSort u) (tLambda bAnon (tRel 0) (tRel 0))). + (* Show that the empty context is well-formed *) + assert (wfΓ : ∥ wf_local gctx_wf_env [] ∥) by do 2 constructor. + fill_inh impl. +Qed. + +(** The extracted typechecker also runs in OCaml *) +MetaCoq SafeCheck (fun x : nat => x + 1). + +(** Erasure *) + +(** Running erasure live in Coq *) +Definition test (p : Ast.program) : string := + erase_and_print_template_program p. + +MetaCoq Quote Recursively Definition zero := 0. + +Definition zerocst := Eval lazy in test zero. +Print zerocst. + +MetaCoq Quote Recursively Definition singleton_elim := + ((fun (X : Set) (x : X) (e : x = x) => + match e in eq _ x' return bool with + | eq_refl => true + end)). +Eval lazy in (test singleton_elim). (* Optimized to remove match on Props *) + +MetaCoq Erase singleton_elim. + +(** Conclusion: Status of MetaCoq + + - 1.0 expected by Christmas with correctness and completeness of the typechecker. + + - All metatheory proofs are finished. Compared to Coq's implementation: + + - full (max (i + k, j + l)) universe support (including a naïve acyclicity checking + algorithm) + + - partial support for SProp (in programs but not yet formalized typing rules) + + - approximate cumulative inductive types checking (not yet up to reduction) + + - missing eta-conversion: the plan is to use contravariant subtyping and eta-reduction, + as in Coq CEP #47 + + - missing template-polymorphism: we're studying a sort polymorphism extension + (with G. Gilbert, K. Maillard and N. Tabareau) to subsume it completely. + + - missing modules and fast conversion machines. + + - Much work to come on the template-coq side to ease meta-programming. + + - Relation to CertiCoq: fast and verified correct erasure, not depending on type-checking + (only retyping). + + + CertiCoq needs to have all constructors eta-expanded, a proof of the + syntactic translation expanding constructors is in progress. + + + Otherwise the front-end part of CertiCoq is complete with proofs. + + + Future work: handling of primitive types (ints, floats, arrays, ...) + +*) + + + diff --git a/examples/metacoq_tour_prelude.v b/examples/metacoq_tour_prelude.v new file mode 100644 index 000000000..a4083e74e --- /dev/null +++ b/examples/metacoq_tour_prelude.v @@ -0,0 +1,77 @@ +(* Distributed under the terms of the MIT license. *) +From MetaCoq.Template Require Import config Universes Loader. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping PCUICLiftSubst. +From MetaCoq.SafeChecker Require Import PCUICErrors PCUICTypeChecker PCUICSafeChecker. +From Equations Require Import Equations. + +Existing Instance default_checker_flags. + +(* ********************************************************* *) +(* In this file we define a small plugin which proves *) +(* the identity theorem for any sort using the safe checker. *) +(* ********************************************************* *) + +Definition bAnon := {| binder_name := nAnon; binder_relevance := Relevant |}. +Definition bNamed s := {| binder_name := nNamed s; binder_relevance := Relevant |}. + +Definition tImpl X Y := tProd bAnon X (lift0 1 Y). + +Definition univ := Level.Level "s". + +(* TODO move to SafeChecker *) + +Definition gctx : global_env_ext := + ([], Monomorphic_ctx (LevelSet.singleton univ, ConstraintSet.empty)). + +(** We use the environment checker to produce the proof that gctx, which is a singleton with only + universe "s" declared is well-formed. *) + +Program Definition check_wf_env_ext {cf:checker_flags} (Σ : global_env) id (ext : universes_decl) : + EnvCheck ({ Σ' : wf_env_ext | Σ'.(wf_env_ext_env) = (Σ, ext)}) := + '(G; pf) <- check_wf_env Σ ;; + '(G'; pf') <- check_wf_env_ext Σ id _ G _ ext ;; + ret (exist {| wf_env_ext_env := (Σ, ext) ; + wf_env_ext_wf := _ ; + wf_env_ext_graph := G' ; + wf_env_ext_graph_wf := _ |} eq_refl). + +Definition kername_of_string (s : string) : kername := + (MPfile [], s). + +Definition make_wf_env_ext {cf:checker_flags} (Σ : global_env_ext) : EnvCheck wf_env_ext := + '(exist Σ' pf) <- check_wf_env_ext Σ.1 (kername_of_string "toplevel") Σ.2 ;; + ret Σ'. + +Definition gctx_wf_env : wf_env_ext. +Proof. + let wf_proof := eval hnf in (make_wf_env_ext gctx) in + match wf_proof with + | CorrectDecl ?x => exact x + | _ => fail "Couldn't prove the global environment is well-formed" + end. +Defined. + + + +(** There is always a proof of `forall x : Sort s, x -> x` *) + +Definition inh {cf:checker_flags} (Σ : wf_env_ext) Γ T := (∑ t, ∥ typing Σ Γ t T ∥). + +Definition check_inh {cf:checker_flags} (Σ : wf_env_ext) Γ (wfΓ : ∥ wf_local Σ Γ ∥) t {T} : typing_result (inh Σ Γ T) := + prf <- check_type_wf_env_fast Σ Γ wfΓ t (T := T) ;; + ret (t; prf). + +Ltac fill_inh t := + lazymatch goal with + [ wfΓ : ∥ wf_local _ ?Γ ∥ |- inh ?Σ ?Γ ?T ] => + let t := uconstr:(check_inh Σ Γ wfΓ t (T:=T)) in + let proof := eval lazy in t in + match proof with + | Checked ?d => exact_no_check d + | TypeError ?e => + let str := eval cbn in (string_of_type_error Σ e) in + fail "Failed to inhabit " T " : " str + | _ => fail "Anomaly: unexpected return value: " proof + end + | [ |- inh _ ?Γ _ ] => fail "Missing local wellformedness assumption for" Γ + end. diff --git a/examples/typing_correctness.v b/examples/typing_correctness.v index 3d7eb9a77..e48d11e25 100644 --- a/examples/typing_correctness.v +++ b/examples/typing_correctness.v @@ -1,9 +1,8 @@ (* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import config Universes All. +From MetaCoq.Template Require Import config Universes Loader. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping PCUICLiftSubst. -From MetaCoq.SafeChecker Require Import PCUICSafeChecker. +From MetaCoq.SafeChecker Require Import PCUICErrors PCUICTypeChecker PCUICSafeChecker. From Equations Require Import Equations. - Existing Instance default_checker_flags. (* ********************************************************* *) @@ -51,9 +50,11 @@ Proof. end. Defined. + + (** There is always a proof of `forall x : Sort s, x -> x` *) -Definition inh {cf:checker_flags} (Σ : wf_env_ext) Γ T := ∑ t, ∥ Σ ;;; Γ |- t : T ∥. +Definition inh {cf:checker_flags} (Σ : wf_env_ext) Γ T := ∑ t, ∥ typing Σ Γ t T ∥. Definition check_inh {cf:checker_flags} (Σ : wf_env_ext) Γ (wfΓ : ∥ wf_local Σ Γ ∥) t {T} : typing_result (inh Σ Γ T) := prf <- check_type_wf_env_fast Σ Γ wfΓ t (T := T) ;; @@ -78,8 +79,7 @@ Lemma identity_typing (u := Universe.make univ): inh gctx_wf_env [] (tProd (bNam Proof. set (impl := tLambda (bNamed "s") (tSort u) (tLambda bAnon (tRel 0) (tRel 0))). assert (wfΓ : ∥ wf_local gctx_wf_env [] ∥) by do 2 constructor. - (** All immediate now *) Time fill_inh impl. Time Qed. -Print identity_typing. +Print identity_typing. (* Still a huge wf_env proof unfolded *) diff --git a/pcuic/_CoqProject.in b/pcuic/_CoqProject.in index 4ce3a6f5c..a78d52b8f 100644 --- a/pcuic/_CoqProject.in +++ b/pcuic/_CoqProject.in @@ -1,7 +1,7 @@ -R theories MetaCoq.PCUIC theories/PCUICUtils.v -theories/PCUICRelations.v +theories/PCUICPrimitive.v theories/PCUICAst.v theories/PCUICSize.v theories/PCUICAstUtils.v @@ -10,6 +10,7 @@ theories/PCUICReflect.v theories/PCUICLiftSubst.v theories/PCUICUnivSubst.v theories/PCUICTyping.v +theories/PCUICGlobalEnv.v theories/PCUICInversion.v theories/PCUICPosition.v theories/PCUICNormal.v diff --git a/pcuic/theories/PCUICAlpha.v b/pcuic/theories/PCUICAlpha.v index d901b0600..d0161dca0 100644 --- a/pcuic/theories/PCUICAlpha.v +++ b/pcuic/theories/PCUICAlpha.v @@ -348,13 +348,14 @@ Section Alpha. constructor. assumption. all: try (eapply upto_names_impl_eq_term ; assumption). all: eapply eq_term_refl. - - intros t na A B u ih ht iht hu ihu v e; invs e. + - intros t na A B s u ih hty ihty ht iht hu ihu v e; invs e. eapply type_Cumul'. - + econstructor. + + econstructor; cycle 1. * eapply iht. eapply eq_term_upto_univ_empty_impl in X; eauto. all:typeclasses eauto. * eapply ihu. assumption. + * eapply hty. + eapply validity_term ; eauto. econstructor ; eauto. + constructor. @@ -460,7 +461,7 @@ Section Alpha. { now rewrite !fix_context_length, (All2_length _ _ X). } eapply type_Cumul'. + econstructor. - * eapply (fix_guard_eq_term _ _ n); eauto. + * eapply (fix_guard_eq_term _ _ _ _ n); eauto. constructor. assumption. * eassumption. * assumption. @@ -470,24 +471,21 @@ Section Alpha. destruct r as [[s [Hs IH]] [[[eqty eqann] eqbod] eqrarg]]. exists s; apply IH; eauto. * solve_all. - ** destruct a0 as [s [Hs IH]]. - specialize (IH _ a1). - specialize (b _ b3). - eapply context_conversion; eauto. - eapply (type_Cumul' (lift0 #|fix_context mfix| (dtype x))); auto. - exists s. rewrite <-H. - eapply (weakening _ _ _ _ (tSort _)); eauto. now eapply typing_wf_local in b. - apply cumul_refl. rewrite <- H. - eapply eq_term_upto_univ_lift. - eapply eq_term_upto_univ_empty_impl. - 4: intuition eauto. - all: intros ? ? []. - *** eapply eq_universe_refl. - *** eapply leq_universe_refl. - *** eapply leq_universe_refl. - ** eapply isLambda_eq_term_l. - --- eassumption. - --- intuition eauto. + destruct a0 as [s [Hs IH]]. + specialize (IH _ a). + specialize (b _ b2). + eapply context_conversion; eauto. + eapply (type_Cumul' (lift0 #|fix_context mfix| (dtype x))); auto. + exists s. rewrite <-H. + eapply (weakening _ _ _ _ (tSort _)); eauto. now eapply typing_wf_local in b. + apply cumul_refl. rewrite <- H. + eapply eq_term_upto_univ_lift. + eapply eq_term_upto_univ_empty_impl. + 4: intuition eauto. + all: intros ? ? []. + *** eapply eq_universe_refl. + *** eapply leq_universe_refl. + *** eapply leq_universe_refl. * revert wffix. unfold wf_fixpoint. enough (map check_one_fix mfix = map check_one_fix mfix') as ->; auto. @@ -532,7 +530,7 @@ Section Alpha. { now rewrite !fix_context_length, (All2_length _ _ X). } eapply type_Cumul'. + econstructor. - * eapply (cofix_guard_eq_term _ _ n) ; eauto. + * eapply (cofix_guard_eq_term _ _ _ _ n) ; eauto. constructor. assumption. * eassumption. * eassumption. diff --git a/pcuic/theories/PCUICArities.v b/pcuic/theories/PCUICArities.v index c6d27e8d6..7b029062e 100644 --- a/pcuic/theories/PCUICArities.v +++ b/pcuic/theories/PCUICArities.v @@ -87,7 +87,7 @@ Proof. constructor; pcuic. eapply context_relation_app in convctx as [_ convctx]. unshelve eapply (context_relation_impl convctx). - simpl; firstorder. destruct X. constructor; auto. + simpl; pcuicfo. destruct X. constructor; auto. eapply conv_conv_ctx; eauto. eapply context_relation_app_inv. constructor; pcuic. constructor; pcuic. constructor; pcuic. now symmetry. @@ -267,7 +267,7 @@ Proof. eapply (type_mkProd_or_LetIn _ _ {| decl_body := None |}) => /=; eauto. econstructor; eauto with pcuic. eapply typing_wf_local in Ht. - depelim Ht; eapply All_local_env_app in Ht; intuition auto. + depelim Ht; eapply All_local_env_app_inv in Ht; intuition auto. now rewrite sort_of_product_twice. Qed. @@ -319,13 +319,13 @@ Lemma subslet_app_closed {cf:checker_flags} Σ Γ s s' Δ Δ' : subslet Σ Γ (s ++ s') (Δ' ,,, Δ). Proof. induction 1 in s', Δ'; simpl; auto; move=> sub'; - rewrite closedn_ctx_snoc => /andP [clctx clt]; + rewrite closedn_ctx_snoc => /andb_and [clctx clt]; try constructor; auto. - pose proof (subslet_length X). rewrite Nat.add_0_r in clt. rewrite /closed_decl /= -H in clt. rewrite subst_app_simpl /= (subst_closedn s') //. - pose proof (subslet_length X). rewrite Nat.add_0_r in clt. - rewrite /closed_decl /= -H in clt. move/andP: clt => [clt clT]. + rewrite /closed_decl /= -H in clt. move/andb_and: clt => [clt clT]. replace (subst0 s t) with (subst0 (s ++ s') t). + constructor; auto. rewrite !subst_app_simpl /= !(subst_closedn s') //. @@ -413,7 +413,7 @@ Proof. rewrite (subst_instance_ind_type_id Σ _ {| inductive_mind := inductive_mind ind; inductive_ind := i |}); eauto. destruct isdecl. split; eauto. reflexivity. } clear oind. - revert X. clear onNpars onGuard. + revert X. clear onNpars. generalize (le_n #|ind_bodies mdecl|). generalize (ind_bodies mdecl) at 1 3 4 5. induction l using rev_ind; simpl; first constructor. @@ -448,7 +448,7 @@ Proof. { apply forall_nth_error_Alli. econstructor; eauto. split; eauto. } clear oind. - revert X. clear onNpars onGuard. + revert X. clear onNpars. generalize (le_n #|ind_bodies mdecl|). generalize (ind_bodies mdecl) at 1 3 4 5. induction l using rev_ind; simpl; first constructor. @@ -594,17 +594,17 @@ Proof. + rewrite it_mkProd_or_LetIn_app. destruct x as [na [b|] ty]; cbn; move=> H. * apply inversion_LetIn in H as (s1 & A & H0 & H1 & H2 & H3); auto. - eapply All_local_env_app_inv; split; pcuic. - eapply All_local_env_app_inv. split. repeat constructor. now exists s1. + eapply All_local_env_app; split; pcuic. + eapply All_local_env_app. split. repeat constructor. now exists s1. auto. apply IHΔ in H2. - eapply All_local_env_app in H2. intuition auto. + eapply All_local_env_app_inv in H2. intuition auto. eapply All_local_env_impl; eauto. simpl. intros. now rewrite app_context_assoc. * apply inversion_Prod in H as (s1 & A & H0 & H1 & H2); auto. - eapply All_local_env_app_inv; split; pcuic. - eapply All_local_env_app_inv. split. repeat constructor. now exists s1. + eapply All_local_env_app; split; pcuic. + eapply All_local_env_app. split. repeat constructor. now exists s1. apply IHΔ in H1. - eapply All_local_env_app in H1. intuition auto. + eapply All_local_env_app_inv in H1. intuition auto. eapply All_local_env_impl; eauto. simpl. intros. now rewrite app_context_assoc. Qed. diff --git a/pcuic/theories/PCUICAst.v b/pcuic/theories/PCUICAst.v index 5d7f4d5a9..789a49970 100644 --- a/pcuic/theories/PCUICAst.v +++ b/pcuic/theories/PCUICAst.v @@ -1,6 +1,7 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Template Require Export utils Universes BasicAst Environment EnvironmentTyping. +From MetaCoq.PCUIC Require Export PCUICPrimitive. From Equations Require Import Equations. (** * AST of the Polymorphic Cumulative Calculus of Inductive Constructions @@ -9,11 +10,17 @@ From Equations Require Import Equations. In particular, it has binary applications and all terms are well-formed. Casts are absent as well. *) - Declare Scope pcuic. Delimit Scope pcuic with pcuic. Open Scope pcuic. +(** DO NOT USE firstorder, since the introduction of Ints and Floats, it became unusuable. *) +Ltac pcuicfo_gen tac := + simpl in *; intuition (simpl; intuition tac). + +Tactic Notation "pcuicfo" := pcuicfo_gen auto. +Tactic Notation "pcuicfo" tactic(tac) := pcuicfo_gen tac. + Inductive term := | tRel (n : nat) | tVar (i : ident) (* For free variables (e.g. in a goal) *) @@ -29,10 +36,14 @@ Inductive term := | tCase (indn : inductive * nat) (p c : term) (brs : list (nat * term)) (* # of parameters/type info/discriminee/branches *) | tProj (p : projection) (c : term) | tFix (mfix : mfixpoint term) (idx : nat) -| tCoFix (mfix : mfixpoint term) (idx : nat). +| tCoFix (mfix : mfixpoint term) (idx : nat) +(** We use faithful models of primitive type values in PCUIC *) +| tPrim (prim : prim_val term). Derive NoConfusion for term. +Notation prim_val := (prim_val term). + Fixpoint mkApps t us := match us with | nil => t diff --git a/pcuic/theories/PCUICAstUtils.v b/pcuic/theories/PCUICAstUtils.v index c8744d6ba..f6afefa10 100644 --- a/pcuic/theories/PCUICAstUtils.v +++ b/pcuic/theories/PCUICAstUtils.v @@ -4,6 +4,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICSize. Require Import ssreflect. From Equations Require Import Equations. +Set Equations Transparent. Derive Signature for All All2. @@ -35,6 +36,7 @@ Fixpoint string_of_term (t : term) := ^ string_of_term c ^ ")" | tFix l n => "Fix(" ^ (string_of_list (string_of_def string_of_term) l) ^ "," ^ string_of_nat n ^ ")" | tCoFix l n => "CoFix(" ^ (string_of_list (string_of_def string_of_term) l) ^ "," ^ string_of_nat n ^ ")" + | tPrim i => "Int(" ^ string_of_prim string_of_term i ^ ")" end. Lemma lookup_env_nil c s : lookup_env [] c = Some s -> False. @@ -68,6 +70,19 @@ Fixpoint decompose_app_rec (t : term) l := Definition decompose_app t := decompose_app_rec t []. +Definition mkApps_decompose_app_rec t l : + mkApps t l = mkApps (fst (decompose_app_rec t l)) (snd (decompose_app_rec t l)). +Proof. + revert l; induction t; try reflexivity. + intro l; cbn in *. + transitivity (mkApps t1 ((t2 ::l))). reflexivity. + now rewrite IHt1. +Qed. + +Definition mkApps_decompose_app t : + t = mkApps (fst (decompose_app t)) (snd (decompose_app t)) + := mkApps_decompose_app_rec t []. + Lemma decompose_app_rec_mkApps f l l' : decompose_app_rec (mkApps f l) l' = decompose_app_rec f (l ++ l'). Proof. @@ -286,16 +301,6 @@ Qed. Hint Rewrite context_assumptions_map context_assumptions_mapi : len. -Lemma mapi_rec_compose {A B C} (g : nat -> B -> C) (f : nat -> A -> B) k l : - mapi_rec g (mapi_rec f l k) k = mapi_rec (fun k x => g k (f k x)) l k. -Proof. - induction l in k |- *; simpl; auto. now rewrite IHl. -Qed. - -Lemma mapi_compose {A B C} (g : nat -> B -> C) (f : nat -> A -> B) l : - mapi g (mapi f l) = mapi (fun k x => g k (f k x)) l. -Proof. apply mapi_rec_compose. Qed. - Lemma compose_map_decl f g x : map_decl f (map_decl g x) = map_decl (f ∘ g) x. Proof. destruct x as [? [?|] ?]; reflexivity. @@ -371,6 +376,34 @@ Proof. eapply decompose_app_rec_notApp. eassumption. Qed. + +Lemma decompose_app_rec_inv {t l' f l} : + decompose_app_rec t l' = (f, l) -> + mkApps t l' = mkApps f l. +Proof. + induction t in f, l', l |- *; try intros [= <- <-]; try reflexivity. + simpl. apply/IHt1. +Qed. + +Lemma decompose_app_inv {t f l} : + decompose_app t = (f, l) -> t = mkApps f l. +Proof. by apply/decompose_app_rec_inv. Qed. + +Lemma decompose_app_nonnil t f l : + isApp t -> + decompose_app t = (f, l) -> l <> []. +Proof. + intros isApp. + destruct t; simpl => //. + intros da. + pose proof (decompose_app_notApp _ _ _ da). + apply decompose_app_inv in da. + destruct l using rev_ind. + unfold decompose_app => /=. + destruct f => //. + destruct l => //. +Qed. + Fixpoint nApp t := match t with | tApp u _ => S (nApp u) @@ -445,18 +478,6 @@ Proof. - assumption. Qed. -Lemma decompose_app_rec_inv {t l' f l} : - decompose_app_rec t l' = (f, l) -> - mkApps t l' = mkApps f l. -Proof. - induction t in f, l', l |- *; try intros [= <- <-]; try reflexivity. - simpl. apply/IHt1. -Qed. - -Lemma decompose_app_inv {t f l} : - decompose_app t = (f, l) -> t = mkApps f l. -Proof. by apply/decompose_app_rec_inv. Qed. - Lemma mkApps_Fix_spec mfix idx args t : mkApps (tFix mfix idx) args = t -> match decompose_app t with | (tFix mfix idx, args') => args' = args @@ -777,6 +798,19 @@ Proof. reflexivity. Qed. +Lemma destArity_mkApps_None ctx t l : + destArity ctx t = None -> destArity ctx (mkApps t l) = None. +Proof. + induction l in t |- *. trivial. + intros H. cbn. apply IHl. reflexivity. +Qed. + +Lemma destArity_mkApps_Ind ctx ind u l : + destArity ctx (mkApps (tInd ind u) l) = None. +Proof. + apply destArity_mkApps_None. reflexivity. +Qed. + (* Helper for nested recursive functions on well-typed terms *) Section MapInP. @@ -788,7 +822,6 @@ Section MapInP. map_InP nil _ := nil; map_InP (cons x xs) H := cons (f x (H x (or_introl eq_refl))) (map_InP xs (fun x inx => H x _)). End MapInP. -Global Transparent map_InP. Lemma map_InP_spec {A B : Type} {P : A -> Type} (f : A -> B) (l : list A) (H : forall x, In x l -> P x) : map_InP (fun (x : A) (_ : P x) => f x) l H = List.map f l. @@ -815,4 +848,62 @@ Lemma map_InP_length {A B : Type} {P : A -> Type} (f : forall x : A, P x -> B) ( Proof. induction l; simpl; auto. Qed. -Hint Rewrite @map_InP_length : len. \ No newline at end of file +Hint Rewrite @map_InP_length : len. + +(** Views *) + +Definition isSort T := + match T with + | tSort u => true + | _ => false + end. + +Inductive view_sort : term -> Type := +| view_sort_sort s : view_sort (tSort s) +| view_sort_other t : ~ isSort t -> view_sort t. + +Equations view_sortc (t : term) : view_sort t := + view_sortc (tSort s) := view_sort_sort s; + view_sortc t := view_sort_other t _. + +Fixpoint isProd t := + match t with + | tProd na A B => true + | _ => false + end. + +Inductive view_prod : term -> Type := +| view_prod_prod na A b : view_prod (tProd na A b) +| view_prod_other t : ~ isProd t -> view_prod t. + +Equations view_prodc (t : term) : view_prod t := + view_prodc (tProd na A b) := view_prod_prod na A b; + view_prodc t := view_prod_other t _. + +Definition isInd (t : term) : bool := + match t with + | tInd _ _ => true + | _ => false + end. + +Inductive view_ind : term -> Type := +| view_ind_tInd ind u : view_ind (tInd ind u) +| view_ind_other t : negb (isInd t) -> view_ind t. + +Equations view_indc (t : term) : view_ind t := + view_indc (tInd ind u) => view_ind_tInd ind u; + view_indc t => view_ind_other t _. + +Inductive view_prod_sort : term -> Type := +| view_prod_sort_prod na A B : view_prod_sort (tProd na A B) +| view_prod_sort_sort u : view_prod_sort (tSort u) +| view_prod_sort_other t : + ~isProd t -> + ~isSort t -> + view_prod_sort t. + +Equations view_prod_sortc (t : term) : view_prod_sort t := { + | tProd na A B => view_prod_sort_prod na A B; + | tSort u => view_prod_sort_sort u; + | t => view_prod_sort_other t _ _ + }. diff --git a/pcuic/theories/PCUICCSubst.v b/pcuic/theories/PCUICCSubst.v index 46e5e19ff..42f4781f6 100644 --- a/pcuic/theories/PCUICCSubst.v +++ b/pcuic/theories/PCUICCSubst.v @@ -11,7 +11,6 @@ From Equations Require Import Equations. no lifting involved and one term at a time. *) - Local Ltac inv H := inversion H; subst. Fixpoint csubst t k u := diff --git a/pcuic/theories/PCUICCanonicity.v b/pcuic/theories/PCUICCanonicity.v index 96dffdc1b..86594c38b 100644 --- a/pcuic/theories/PCUICCanonicity.v +++ b/pcuic/theories/PCUICCanonicity.v @@ -18,7 +18,7 @@ Module P := PCUICWcbvEval. Local Existing Instance config.extraction_checker_flags. Require Import Equations.Prop.DepElim. -Require Import ssreflect. +Require Import ssreflect ssrbool. Lemma negb_False (p : bool) : negb p -> p -> False. Proof. @@ -285,8 +285,8 @@ Section Spines. * elimtype False; depelim ass. * simpl. rewrite !it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /= //. intros wf sp; depelim sp. rewrite nth_error_nil //. - pose proof (All_local_env_app _ _ _ wf) as [_ wfty]. - eapply All_local_env_app in wfty as [wfty _]. depelim wfty. + pose proof (wf_local_app_inv wf) as [_ wfty]. + eapply wf_local_rel_app_inv in wfty as [wfty _]. depelim wfty. eapply cumul_Prod_inv in c as [dom codom]. 2-3:pcuic. assert (Σ ;;; Γ |- hd : ty). { eapply type_Cumul'; pcuic. eapply conv_cumul. now symmetry. } @@ -410,8 +410,8 @@ Section Spines. * elimtype False; depelim ass. * simpl. rewrite !it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /= //. intros wf sp. - pose proof (All_local_env_app _ _ _ wf) as [_ wfty]. - eapply All_local_env_app in wfty as [wfty _]. depelim wfty. + pose proof (wf_local_app_inv wf) as [_ wfty]. + eapply All_local_env_app_inv in wfty as [wfty _]. depelim wfty. intros Hargs. eapply nth_error_None in Hargs. len in Hargs. len; simpl. depelim sp. @@ -606,188 +606,6 @@ Section Spines. End Spines. -(* -Section Normalization. - Context {cf:checker_flags} (Σ : global_env_ext). - Context {wfΣ : wf Σ}. - - Section reducible. - Lemma reducible Γ t : sum (∑ t', red1 Σ Γ t t') (forall t', red1 Σ Γ t t' -> False). - Proof. - Local Ltac lefte := left; eexists; econstructor; eauto. - Local Ltac leftes := left; eexists; econstructor; solve [eauto]. - Local Ltac righte := right; intros t' red; depelim red; solve_discr; eauto 2. - induction t in Γ |- * using term_forall_list_ind. - (*all:try solve [righte]. - - destruct (nth_error Γ n) eqn:hnth. - destruct c as [na [b|] ty]; [lefte|righte]. - * rewrite hnth; reflexivity. - * rewrite hnth /= // in e. - * righte. rewrite hnth /= // in e. - - admit. - - destruct (IHt1 Γ) as [[? ?]|]; [lefte|]. - destruct (IHt2 (Γ ,, vass n t1)) as [[? ?]|]; [|righte]. - leftes. - - destruct (IHt1 Γ) as [[? ?]|]; [lefte|]. - destruct (IHt2 (Γ ,, vass n t1)) as [[? ?]|]; [|righte]. - leftes. - - destruct (IHt1 Γ) as [[? ?]|]; [lefte|]. - destruct (IHt2 Γ) as [[? ?]|]; [lefte|]. - destruct (IHt3 (Γ ,, vdef n t1 t2)) as [[? ?]|]; [|]. - leftes. lefte. - - destruct (IHt1 Γ) as [[? ?]|]; [lefte|]. - destruct (IHt2 Γ) as [[? ?]|]; [leftes|]. - destruct (PCUICParallelReductionConfluence.view_lambda_fix_app t1 t2). - * rewrite [tApp _ _](mkApps_nested _ _ [a]). - destruct (unfold_fix mfix i) as [[rarg body]|] eqn:unf. - destruct (is_constructor rarg (l ++ [a])) eqn:isc; [leftes|]; eauto. - right => t' red; depelim red; solve_discr; eauto. - rewrite -mkApps_nested in H. noconf H. eauto. - rewrite -mkApps_nested in H. noconf H. eauto. - eapply (f_equal (@length _)) in H1. rewrite /= app_length /= // in H1; lia. - eapply (f_equal (@length _)) in H1. rewrite /= app_length /= // in H1; lia. - righte; try (rewrite -mkApps_nested in H; noconf H); eauto. - eapply (f_equal (@length _)) in H1. rewrite /= app_length /= // in H1; lia. - eapply (f_equal (@length _)) in H1. rewrite /= app_length /= // in H1; lia. - * admit. - * righte. destruct args using rev_case; solve_discr; noconf H. - rewrite H in i. eapply negb_False; eauto. - rewrite -mkApps_nested; eapply isFixLambda_app_mkApps' => //. - - admit. - - admit. - - admit. - - admit. - - admit.*) - - Admitted. - End reducible. - - Lemma reducible' Γ t : sum (∑ t', red1 Σ Γ t t') (normal Σ Γ t). - Proof. - Ltac lefte := left; eexists; econstructor; eauto. - Ltac leftes := left; eexists; econstructor; solve [eauto]. - Ltac righte := right; (solve [repeat (constructor; eauto)])||(repeat constructor). - induction t in Γ |- * using term_forall_list_ind. - all:try solve [righte]. - - destruct (nth_error Γ n) eqn:hnth. - destruct c as [na [b|] ty]; [lefte|]. - * rewrite hnth; reflexivity. - * right. do 2 constructor; rewrite hnth /= //. - * righte. rewrite hnth /= //. - - admit. - - destruct (IHt1 Γ) as [[? ?]|]; [lefte|]. - destruct (IHt2 (Γ ,, vass n t1)) as [[? ?]|]; [|]. - leftes. right; solve[constructor; eauto]. - - destruct (IHt1 Γ) as [[? ?]|]; [lefte|]. - destruct (IHt2 (Γ ,, vass n t1)) as [[? ?]|]; [leftes|leftes]. - - destruct (IHt1 Γ) as [[? ?]|]; [lefte|]. - destruct (IHt2 Γ) as [[? ?]|]; [leftes|]. - destruct (PCUICParallelReductionConfluence.view_lambda_fix_app t1 t2). - * rewrite [tApp _ _](mkApps_nested _ _ [a]). - destruct (unfold_fix mfix i) as [[rarg body]|] eqn:unf. - destruct (is_constructor rarg (l ++ [a])) eqn:isc; [leftes|]; eauto. - right; constructor. rewrite -mkApps_nested. constructor. admit. admit. admit. - * admit. - * admit. - - admit. - - admit. - - admit. - - admit. - - admit. - Admitted. - - Lemma normalizer {Γ t ty} : - Σ ;;; Γ |- t : ty -> - ∑ nf, (red Σ.1 Γ t nf) * normal Σ Γ nf. - Proof. - intros Hty. - unshelve epose proof (PCUICSN.normalisation Σ Γ t (iswelltyped _ _ _ ty Hty)). - clear ty Hty. - move: t H. eapply Fix_F. - intros x IH. - destruct (reducible' Γ x) as [[t' red]|nred]. - specialize (IH t'). forward IH by (constructor; auto). - destruct IH as [nf [rednf norm]]. - exists nf; split; auto. now transitivity t'. - exists x. split; [constructor|assumption]. - Qed. - - Derive Signature for neutral normal. - - Lemma typing_var {Γ n ty} : Σ ;;; Γ |- (tVar n) : ty -> False. - Proof. intros Hty; depind Hty; eauto. Qed. - - Lemma typing_evar {Γ n l ty} : Σ ;;; Γ |- (tEvar n l) : ty -> False. - Proof. intros Hty; depind Hty; eauto. Qed. - - Definition axiom_free Σ := - forall c decl, declared_constant Σ c decl -> cst_body decl <> None. - - Lemma neutral_empty t ty : axiom_free Σ -> Σ ;;; [] |- t : ty -> neutral Σ [] t -> False. - Proof. - intros axfree typed ne. - pose proof (PCUICClosed.subject_closed wfΣ typed) as cl. - depind ne. - - now simpl in cl. - - now eapply typing_var in typed. - - now eapply typing_evar in typed. - - eapply inversion_Const in typed as [decl [wfd [declc [cu cum]]]]; eauto. - specialize (axfree _ _ declc). specialize (H decl). - destruct (cst_body decl); try congruence. - now specialize (H t declc eq_refl). - - simpl in cl; move/andP: cl => [clf cla]. - eapply inversion_App in typed as [na [A [B [Hf _]]]]; eauto. - - simpl in cl; move/andP: cl => [/andP[_ clc] _]. - eapply inversion_Case in typed; firstorder eauto. - - eapply inversion_Proj in typed; firstorder auto. - Qed. - - Lemma ind_normal_constructor t i u args : - axiom_free Σ -> - Σ ;;; [] |- t : mkApps (tInd i u) args -> normal Σ [] t -> construct_cofix_discr (head t). - Proof. - intros axfree Ht capp. destruct capp. - - eapply neutral_empty in H; eauto. - - eapply inversion_Sort in Ht as (? & ? & ? & ? & ?); auto. - eapply invert_cumul_sort_l in c as (? & ? & ?). - eapply red_mkApps_tInd in r as (? & eq & ?); eauto; eauto. - solve_discr. - - eapply inversion_Prod in Ht as (? & ? & ? & ? & ?); auto. - eapply invert_cumul_sort_l in c as (? & ? & ?). - eapply red_mkApps_tInd in r as (? & eq & ?); eauto; eauto. - solve_discr. - - eapply inversion_Lambda in Ht as (? & ? & ? & ? & ?); auto. - eapply invert_cumul_prod_l in c as (? & ? & ? & (? & ?) & ?); auto. - eapply red_mkApps_tInd in r as (? & eq & ?); eauto; eauto. - solve_discr. - - now rewrite head_mkApps /= /head /=. - - eapply PCUICValidity.inversion_mkApps in Ht as (? & ? & ?); auto. - eapply inversion_Ind in t as (? & ? & ? & decli & ? & ?); auto. - eapply PCUICSpine.typing_spine_strengthen in t0; eauto. - pose proof (on_declared_inductive wfΣ decli) as [onind oib]. - rewrite oib.(ind_arity_eq) in t0. - rewrite !subst_instance_constr_it_mkProd_or_LetIn in t0. - eapply typing_spine_arity_mkApps_Ind in t0; eauto. - eexists; split; [sq|]; eauto. - now do 2 eapply isArity_it_mkProd_or_LetIn. - - admit. (* wf of fixpoints *) - - now rewrite /head /=. - Admitted. - - Lemma red_normal_constructor t i u args : - axiom_free Σ -> - Σ ;;; [] |- t : mkApps (tInd i u) args -> - ∑ hnf, (red Σ.1 [] t hnf) * construct_cofix_discr (head hnf). - Proof. - intros axfree Ht. destruct (normalizer Ht) as [nf [rednf capp]]. - exists nf; split; auto. - eapply subject_reduction in Ht; eauto. - now eapply ind_normal_constructor. - Qed. - -End Normalization. -*) - (** Evaluation is a subrelation of reduction: *) Tactic Notation "redt" uconstr(y) := eapply (CRelationClasses.transitivity (R:=red _ _) (y:=y)). @@ -802,7 +620,7 @@ Section WeakNormalization. Transparent construct_cofix_discr. - Lemma value_whnf t : closed t -> value Σ t -> wh_normal Σ [] t. + Lemma value_whnf t : closed t -> value t -> wh_normal Σ [] t. Proof. intros cl ev. induction ev in cl |- * using value_values_ind. @@ -814,10 +632,6 @@ Section WeakNormalization. now rewrite nth_error_nil. - eapply (whnf_cofixapp _ _ [] _ _ []). - unfold value_head in H. destruct t => //. - constructor; eapply whne_mkApps. - cbn in H; destruct lookup_env eqn:eq => //. - destruct g => //. destruct c => //. destruct cst_body => //. - eapply whne_const; eauto. - destruct f => //. cbn in H. destruct cunfold_fix as [[rarg body]|] eqn:unf => //. pose proof cl as cl'. @@ -837,69 +651,12 @@ Section WeakNormalization. now eapply value_whnf. Qed. - (* - Lemma reducible Γ t : sum (∑ t', red1 Σ Γ t t') (wh_normal Σ Γ t). - Proof. - Ltac lefte := left; eexists; econstructor; eauto. - Ltac leftes := left; eexists; econstructor; solve [eauto]. - Ltac righte := right; (solve [repeat (constructor; eauto)])||(repeat constructor). - induction t in Γ |- * using term_forall_list_ind. - all:try solve [righte]. - - destruct (nth_error Γ n) eqn:hnth. - destruct c as [na [b|] ty]; [lefte|]. - * rewrite hnth; reflexivity. - * right. do 2 constructor; rewrite hnth /= //. - * righte. rewrite hnth /= //. admit. - - destruct (IHt1 Γ) as [[? ?]|]; [lefte|]. - destruct (IHt2 Γ) as [[? ?]|]; [leftes|]. - destruct (IHt2 (Γ ,, vass n t1)) as [[? ?]|]; [|]. - leftes. leftes. - - destruct (IHt1 Γ) as [[? ?]|]; [lefte|]. - destruct (IHt2 Γ) as [[? ?]|]; [leftes|]. - destruct (PCUICParallelReductionConfluence.view_lambda_fix_app t1 t2). - * rewrite [tApp _ _](mkApps_nested _ _ [a]). - destruct (unfold_fix mfix i) as [[rarg body]|] eqn:unf. - + destruct (is_constructor rarg (l ++ [a])) eqn:isc; [leftes|]; eauto. - right. rewrite /is_constructor in isc. - destruct nth_error eqn:eq. - ++ constructor; eapply whne_fixapp; eauto. admit. - ++ eapply whnf_fixapp. rewrite unf //. - + right. eapply whnf_fixapp. rewrite unf //. - * left. induction l; simpl. eexists. constructor. - eexists. eapply app_red_l. eapply red1_mkApps_l. constructor. - * admit. - - admit. - - admit. - - admit. - - admit. - - admit. - Admitted. - - Lemma normalizer {Γ t ty} : - Σ ;;; Γ |- t : ty -> - ∑ nf, (red Σ.1 Γ t nf) * wh_normal Σ Γ nf. - Proof. - intros Hty. - unshelve epose proof (PCUICSN.normalisation Σ Γ t (iswelltyped _ _ _ ty Hty)). - clear ty Hty. - move: t H. eapply Fix_F. - intros x IH. - destruct (reducible Γ x) as [[t' red]|nred]. - specialize (IH t'). forward IH by (constructor; auto). - destruct IH as [nf [rednf norm]]. - exists nf; split; auto. now transitivity t'. - exists x. split; [constructor|assumption]. - Qed. *) - Lemma typing_var {Γ n ty} : Σ ;;; Γ |- (tVar n) : ty -> False. Proof. intros Hty; depind Hty; eauto. Qed. Lemma typing_evar {Γ n l ty} : Σ ;;; Γ |- (tEvar n l) : ty -> False. Proof. intros Hty; depind Hty; eauto. Qed. - Definition axiom_free Σ := - forall c decl, declared_constant Σ c decl -> cst_body decl <> None. - Lemma invert_cumul_prod_ind {Γ na dom codom ind u args} : Σ ;;; Γ |- tProd na dom codom <= mkApps (tInd ind u) args -> False. Proof. @@ -1014,39 +771,92 @@ Section WeakNormalization. eapply PCUICConfluence.red_mkApps_tInd in r as [? [eq _]]; auto. solve_discr. Qed. - - Lemma wh_neutral_empty_gen t Γ : axiom_free Σ -> wh_neutral Σ Γ t -> forall ty, Σ ;;; Γ |- t : ty -> Γ = [] -> False - with wh_normal_empty_gen t Γ i u args : axiom_free Σ -> wh_normal Σ Γ t -> Σ ;;; Γ |- t : mkApps (tInd i u) args -> - Γ = [] -> construct_cofix_discr (head t). + + Fixpoint axiom_free_value Σ args t := + match t with + | tApp hd arg => axiom_free_value Σ (axiom_free_value Σ [] arg :: args) hd + | tConst kn _ => match lookup_env Σ kn with + | Some (ConstantDecl {| cst_body := None |}) => false + | _ => true + end + | tCase _ _ discr _ => axiom_free_value Σ [] discr + | tProj _ t => axiom_free_value Σ [] t + | tFix defs i => + match nth_error defs i with + | Some def => nth (rarg def) args true + | None => true + end + | _ => true + end. + + Lemma axiom_free_value_mkApps Σ' args hd args' : + axiom_free_value Σ' args (mkApps hd args') = + axiom_free_value Σ' (map (axiom_free_value Σ' []) args' ++ args) hd. Proof. - intros axfree ne ty typed. - 2:intros axfree ne typed. - all:pose proof (subject_closed wfΣ typed) as cl; - destruct ne; intros eqΓ; simpl in *; try discriminate. - - rewrite eqΓ in cl => //. + revert hd args. + induction args' using MCList.rev_ind; intros hd args; cbn; auto. + rewrite -mkApps_nested /=. + rewrite IHargs'. + now rewrite map_app /= -app_assoc. + Qed. + + Lemma wh_neutral_empty_gen Γ free t ty : + axiom_free_value Σ free t -> + wh_neutral Σ [] t -> + Σ ;;; Γ |- t : ty -> + Γ = [] -> + False. + Proof. + intros axfree ne typed. + induction ne in free, t, ty, axfree, ne, typed |- *; intros ->; simpl in *; try discriminate. + - apply inversion_Rel in typed as (?&?&?&?); auto. + rewrite nth_error_nil in e0; discriminate. - now eapply typing_var in typed. - now eapply typing_evar in typed. - - clear wh_neutral_empty_gen wh_normal_empty_gen. subst. - apply inversion_Const in typed as [decl' [wfd [declc [cu cum]]]]; eauto. - specialize (axfree _ _ declc). - red in declc. rewrite declc in e. noconf e. congruence. - - simpl in cl; move/andP: cl => [clf cla]. - eapply inversion_App in typed as [na [A [B [Hf _]]]]; eauto. - - specialize (wh_neutral_empty_gen _ _ axfree ne). subst. - simpl in cl. - eapply inversion_mkApps in typed as (? & ? & ?); eauto. + - apply inversion_Const in typed as [decl' [wfd [declc [cu cum]]]]; eauto. + red in declc. + rewrite declc in e, axfree. + noconf e. + destruct decl; cbn in *. + rewrite e0 in axfree; congruence. + - eapply inversion_App in typed as [na [A [B [Hf _]]]]; eauto. + - eapply inversion_mkApps in typed as (? & ? & ?); eauto. eapply inversion_Fix in t as (? & ? & ? & ? & ? & ? & ?); auto. eapply typing_spine_strengthen in t0; eauto. eapply nth_error_all in a; eauto. simpl in a. rewrite /unfold_fix in e. rewrite e1 in e. noconf e. eapply (wf_fixpoint_spine wfΣ) in t0; eauto. rewrite e0 in t0. destruct t0 as [ind [u [indargs [tyarg ckind]]]]. - clear wh_normal_empty_gen. - now specialize (wh_neutral_empty_gen _ tyarg eq_refl). - - move/andP: cl => [/andP[_ clc] _]. - eapply inversion_Case in typed; firstorder eauto. - - eapply inversion_Proj in typed; firstorder auto. - - eapply wh_neutral_empty_gen in w; eauto. + rewrite axiom_free_value_mkApps in axfree. + cbn in axfree. + rewrite e1 nth_nth_error nth_error_app1 in axfree. + 1: { rewrite map_length. + apply nth_error_Some_length in e0; auto. } + rewrite nth_error_map e0 in axfree. + cbn in axfree. + eauto. + - eapply inversion_Case in typed as + (? & ? & ? & ? & ? & ? & ? & ? & ? & ? & ? & ? & ? & ? & ? & ? & ?); eauto. + - eapply inversion_Proj in typed as (? & ? & ? & ? & ? & ? & ? & ? & ?); eauto. + Qed. + + Lemma wh_neutral_empty t ty : + axiom_free_value Σ [] t -> + wh_neutral Σ [] t -> + Σ ;;; [] |- t : ty -> + False. + Proof. eauto using wh_neutral_empty_gen. Qed. + + Lemma wh_normal_ind_discr t i u args : + axiom_free_value Σ [] t -> + wh_normal Σ [] t -> + Σ ;;; [] |- t : mkApps (tInd i u) args -> + construct_cofix_discr (head t). + Proof. + intros axfree ne typed. + pose proof (subject_closed wfΣ typed) as cl. + depelim ne; simpl in *. + - eauto using wh_neutral_empty. - eapply inversion_Sort in typed as (? & ? & ?); auto. eapply invert_cumul_sort_l in c as (? & ? & ?); auto. eapply red_mkApps_tInd in r as (? & eq & ?); eauto; eauto. @@ -1063,25 +873,14 @@ Section WeakNormalization. - exfalso; eapply invert_ind_ind; eauto. - exfalso; eapply invert_fix_ind; eauto. - now rewrite head_mkApps /head /=. + - now eapply inversion_Prim in typed. Qed. - Lemma wh_neutral_empty t ty : axiom_free Σ -> - Σ ;;; [] |- t : ty -> - wh_neutral Σ [] t -> - False. - Proof. intros; eapply wh_neutral_empty_gen; eauto. Qed. - - Lemma wh_normal_ind_discr t i u args : axiom_free Σ -> - Σ ;;; [] |- t : mkApps (tInd i u) args -> - wh_normal Σ [] t -> - construct_cofix_discr (head t). - Proof. intros. eapply wh_normal_empty_gen; eauto. Qed. - Lemma whnf_ind_finite t ind u indargs : - axiom_free Σ -> + axiom_free_value Σ [] t -> Σ ;;; [] |- t : mkApps (tInd ind u) indargs -> wh_normal Σ [] t -> - check_recursivity_kind Σ (inductive_mind ind) Finite -> + ~check_recursivity_kind Σ (inductive_mind ind) CoFinite -> isConstruct_app t. Proof. intros axfree typed whnf ck. @@ -1092,19 +891,18 @@ Section WeakNormalization. destruct hd eqn:eqh => //. subst hd. eapply decompose_app_inv in da. subst. eapply typing_cofix_coind in typed. - now move: (check_recursivity_kind_inj typed ck). + auto. Qed. Lemma fix_app_is_constructor {mfix idx args ty narg fn} : - axiom_free Σ -> Σ;;; [] |- mkApps (tFix mfix idx) args : ty -> unfold_fix mfix idx = Some (narg, fn) -> match nth_error args narg return Type with - | Some a => wh_normal Σ [] a -> isConstruct_app a + | Some a => axiom_free_value Σ [] a -> wh_normal Σ [] a -> isConstruct_app a | None => ∑ na dom codom, Σ ;;; [] |- tProd na dom codom <= ty end. Proof. - intros axfree typed unf. + intros typed unf. eapply inversion_mkApps in typed as (? & ? & ?); eauto. eapply inversion_Fix in t as (? & ? & ? & ? & ? & ? & ?); auto. eapply typing_spine_strengthen in t0; eauto. @@ -1112,11 +910,35 @@ Section WeakNormalization. rewrite /unfold_fix in unf. rewrite e in unf. noconf unf. eapply (wf_fixpoint_spine wfΣ) in t0; eauto. - rewrite /is_constructor. destruct (nth_error args (rarg x0)) eqn:hnth. + rewrite /is_constructor. destruct (nth_error args (rarg x0)) eqn:hnth; [|assumption]. destruct_sigma t0. destruct t0. - intros norm. + intros axfree norm. eapply whnf_ind_finite in t0; eauto. - assumption. + intros chk. + pose proof (check_recursivity_kind_inj chk i1). + discriminate. + Qed. + + Lemma value_axiom_free Σ' t : + value t -> + axiom_free_value Σ' [] t. + Proof. + intros val. + induction val using value_values_ind. + - destruct t; try discriminate; auto. + cbn. + destruct ?; auto. + destruct ?; auto. + - rewrite axiom_free_value_mkApps. + destruct t; auto. + - rewrite axiom_free_value_mkApps. + destruct f; try discriminate. + cbn. + unfold isStuckFix, cunfold_fix in H. + destruct nth_error; auto. + rewrite nth_overflow; auto. + rewrite app_nil_r map_length; auto. + toProp; auto. Qed. (** Evaluation on well-typed terms corresponds to reduction. @@ -1126,11 +948,10 @@ Section WeakNormalization. have a constructor at their recursive argument as it is ensured by typing. *) Lemma wcbeval_red t ty u : - axiom_free Σ -> Σ ;;; [] |- t : ty -> eval Σ t u -> red Σ [] t u. Proof. - intros axfree Hc He. + intros Hc He. revert ty Hc. induction He; simpl; move=> ty Ht; try solve[econstructor; eauto]. @@ -1216,8 +1037,8 @@ Section WeakNormalization. rewrite closedn_mkApps in Ht2. now move/andP: Ht2 => [clf _]. eapply red_fix; eauto. assert (Σ ;;; [] |- mkApps (tFix mfix idx) (argsv ++ [av]) : B {0 := av}). - { rewrite -mkApps_nested /=. eapply type_App; eauto. } - epose proof (fix_app_is_constructor axfree X0 e); eauto. + { rewrite -mkApps_nested /=. eapply type_App'; eauto. } + epose proof (fix_app_is_constructor X0 e); eauto. rewrite /is_constructor. destruct nth_error eqn:hnth => //. assert (All (closedn 0) (argsv ++ [av])). @@ -1225,7 +1046,7 @@ Section WeakNormalization. rewrite closedn_mkApps in X0. move/andP: X0 => [clfix clargs]. now eapply forallb_All in clargs. } - assert (All (value Σ) (argsv ++ [av])). + assert (All value (argsv ++ [av])). { eapply All_app_inv; [|constructor; [|constructor]]. eapply eval_to_value in He1. eapply value_mkApps_inv in He1 as [[[-> Hat]|[vh vargs]]|[hstuck vargs]] => //. @@ -1233,6 +1054,7 @@ Section WeakNormalization. solve_all. eapply nth_error_all in X3; eauto. simpl in X3. destruct X3 as [cl val]. eapply X1, value_whnf; auto. + eapply value_axiom_free; auto. eapply nth_error_None in hnth; len in hnth; simpl in *. lia. } redt _; eauto. @@ -1270,14 +1092,15 @@ Section WeakNormalization. Qed. Lemma eval_ind_canonical t i u args : - axiom_free Σ -> Σ ;;; [] |- t : mkApps (tInd i u) args -> forall t', eval Σ t t' -> construct_cofix_discr (head t'). Proof. - intros axfree Ht t' eval. + intros Ht t' eval. pose proof (subject_closed wfΣ Ht). eapply subject_reduction in Ht. 3:eapply wcbeval_red; eauto. 2:auto. + eapply eval_to_value in eval as axfree. + eapply value_axiom_free in axfree. eapply eval_whne in eval; auto. eapply wh_normal_ind_discr; eauto. Qed. diff --git a/pcuic/theories/PCUICClosed.v b/pcuic/theories/PCUICClosed.v index 7152f08a0..35a50d762 100644 --- a/pcuic/theories/PCUICClosed.v +++ b/pcuic/theories/PCUICClosed.v @@ -331,8 +331,6 @@ Proof. now rewrite closedn_subst_instance_constr IHl. Qed. -Arguments skipn : simpl never. - Lemma destArity_spec ctx T : match destArity ctx T with | Some (ctx', s) => it_mkProd_or_LetIn ctx T = it_mkProd_or_LetIn ctx' (tSort s) @@ -715,7 +713,7 @@ Proof. - intuition auto. generalize (closedn_subst [u] #|Γ| 0 B). rewrite Nat.add_0_r. - move=> Hs. apply: Hs => /=. rewrite H0 => //. + move=> Hs. apply: Hs => /=. rewrite H1 => //. rewrite Nat.add_1_r. auto. - rewrite closedn_subst_instance_constr. @@ -940,8 +938,7 @@ Proof. now move/andP: oib => []. - pose proof (onConstructors oib). red in X. eapply All_forallb. eapply All2_All_left; eauto. - firstorder auto. - eapply on_ctype in X0. + intros cdecl cs X0; eapply on_ctype in X0. now move/andP: X0 => []. - eapply All_forallb. pose proof (onProjections oib). @@ -997,6 +994,9 @@ Proof. now move: c => [_ /andP [_ ct]]. Qed. +Lemma isType_closed {cf:checker_flags} {Σ Γ T} : wf Σ.1 -> isType Σ Γ T -> closedn #|Γ| T. +Proof. intros wfΣ [s Hs]. now eapply subject_closed in Hs. Qed. + Lemma closed_wf_local `{checker_flags} {Σ Γ} : wf Σ.1 -> wf_local Σ Γ -> @@ -1127,8 +1127,8 @@ Proof. erewrite hidecl in clbodies. simpl in clbodies. unfold closed_inductive_body in clbodies. move/andP: clbodies => [/andP [_ cl] _]. - eapply forallb_All in cl. apply (All_impl cl). - intros [[? ?] ?]; simpl; firstorder. + eapply forallb_All in cl. apply (All_impl cl). + now intros [[? ?] ?]; simpl. Qed. Lemma declared_minductive_closed_inds {cf:checker_flags} : @@ -1212,6 +1212,7 @@ Lemma term_closedn_list_ind : (forall k (s : projection) (t : term), P k t -> P k (tProj s t)) -> (forall k (m : mfixpoint term) (n : nat), tFixProp (P k) (P (#|fix_context m| + k)) m -> P k (tFix m n)) -> (forall k (m : mfixpoint term) (n : nat), tFixProp (P k) (P (#|fix_context m| + k)) m -> P k (tCoFix m n)) -> + (forall k p, P k (tPrim p)) -> forall k (t : term), closedn k t -> P k t. Proof. intros until t. revert k t. @@ -1261,3 +1262,79 @@ Proof. simpl in clt. move/andP: clt. intuition auto. move/andP: clt => [cd cmfix]. apply auxm; auto. Defined. + +Lemma term_noccur_between_list_ind : + forall (P : nat -> nat -> term -> Type), + (forall k n (i : nat), i < k \/ k + n <= i -> P k n (tRel i)) -> + (forall k n (i : ident), P k n (tVar i)) -> + (forall k n (id : nat) (l : list term), All (P k n) l -> P k n (tEvar id l)) -> + (forall k n s, P k n (tSort s)) -> + (forall k n (na : aname) (t : term), P k n t -> forall t0 : term, P (S k) n t0 -> P k n (tProd na t t0)) -> + (forall k n (na : aname) (t : term), P k n t -> forall t0 : term, P (S k) n t0 -> P k n (tLambda na t t0)) -> + (forall k n (na : aname) (t : term), + P k n t -> forall t0 : term, P k n t0 -> forall t1 : term, P (S k) n t1 -> P k n (tLetIn na t t0 t1)) -> + (forall k n (t u : term), P k n t -> P k n u -> P k n (tApp t u)) -> + (forall k n s (u : list Level.t), P k n (tConst s u)) -> + (forall k n (i : inductive) (u : list Level.t), P k n (tInd i u)) -> + (forall k n (i : inductive) (c : nat) (u : list Level.t), P k n (tConstruct i c u)) -> + (forall k n (p : inductive * nat) (t : term), + P k n t -> forall t0 : term, P k n t0 -> forall l : list (nat * term), + tCaseBrsProp (P k n) l -> P k n (tCase p t t0 l)) -> + (forall k n (s : projection) (t : term), P k n t -> P k n (tProj s t)) -> + (forall k n (m : mfixpoint term) (i : nat), tFixProp (P k n) (P (#|fix_context m| + k) n) m -> P k n (tFix m i)) -> + (forall k n (m : mfixpoint term) (i : nat), tFixProp (P k n) (P (#|fix_context m| + k) n) m -> P k n (tCoFix m i)) -> + (forall k n p, P k n (tPrim p)) -> + forall k n (t : term), noccur_between k n t -> P k n t. +Proof. + intros until t. revert k n t. + fix auxt 3. + move auxt at top. + intros k n t. + destruct t; intros clt; match goal with + H : _ |- _ => apply H + end; auto; simpl in clt; + try move/andP: clt => [cl1 cl2]; + try move/andP: cl1 => [cl1 cl1']; + try solve[apply auxt; auto]; + simpl in *. + + - move/orP: clt => [cl|cl]. + now apply Nat.ltb_lt in cl; eauto. + now apply Nat.leb_le in cl; eauto. + + - revert l clt. + fix auxl' 1. + destruct l; constructor; [|apply auxl']. + apply auxt. simpl in clt. now move/andP: clt => [clt cll]. + now move/andP: clt => [clt cll]. + + - red. + revert brs cl2. + fix auxl' 1. + destruct brs; constructor; [|apply auxl']. + simpl in cl2. move/andP: cl2 => [clt cll]. + apply auxt, clt. move/andP: cl2 => [clt cll]. + apply cll. + + - red. + rewrite fix_context_length. + revert clt. + generalize (#|mfix|). + revert mfix. + fix auxm 1. + destruct mfix; intros; constructor. + simpl in clt. move/andP: clt => [clt cll]. + simpl in clt. move/andP: clt. intuition auto. + move/andP: clt => [cd cmfix]. apply auxm; auto. + + - red. + rewrite fix_context_length. + revert clt. + generalize (#|mfix|). + revert mfix. + fix auxm 1. + destruct mfix; intros; constructor. + simpl in clt. move/andP: clt => [clt cll]. + simpl in clt. move/andP: clt. intuition auto. + move/andP: clt => [cd cmfix]. apply auxm; auto. +Defined. diff --git a/pcuic/theories/PCUICConfluence.v b/pcuic/theories/PCUICConfluence.v index 2c8e50de4..34b950da2 100644 --- a/pcuic/theories/PCUICConfluence.v +++ b/pcuic/theories/PCUICConfluence.v @@ -1,6 +1,6 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Template Require Import config utils. -From MetaCoq.PCUIC Require Import PCUICRelations PCUICAst PCUICLiftSubst PCUICTyping +From MetaCoq.PCUIC Require Import PCUICAst PCUICLiftSubst PCUICTyping PCUICReduction PCUICWeakening PCUICEquality PCUICUnivSubstitution PCUICParallelReduction PCUICParallelReductionConfluence. @@ -2047,15 +2047,15 @@ Section RedConfluence. intros x y H. induction H; firstorder. - red in p. - induction p. repeat constructor. firstorder. + induction p. repeat constructor. pcuicfo. constructor 2. econstructor 3 with (Γ ,, vass na y); auto. - subst. - induction a. repeat constructor. firstorder. + induction a. repeat constructor. pcuicfo. constructor 2. econstructor 3 with (Γ ,, vdef na y t'); auto. - subst. - induction a. constructor. constructor. red. right. firstorder. + induction a. constructor. constructor. red. right. pcuicfo. constructor 2. econstructor 3 with (Γ ,, vdef na b' y); auto. - clear H. induction IHOnOne2_local_env. constructor. now constructor 3. @@ -2262,7 +2262,7 @@ Section RedConfluence. Lemma clos_refl_trans_out Γ x y : clos_refl_trans (red1 Σ Γ) x y -> clos_refl_trans red1_rel (Γ, x) (Γ, y). Proof. - induction 1. constructor. red. left. firstorder. + induction 1. constructor. red. left. pcuicfo. constructor 2. econstructor 3; eauto. Qed. @@ -2322,8 +2322,8 @@ Section RedConfluence. clear - Hctx. induction (fix_context mfix0). + assumption. + simpl. destruct a as [na [b|] ty]. - * constructor ; firstorder (hnf ; auto). - * constructor ; firstorder (hnf ; auto). + * constructor ; pcuicfo (hnf ; auto). + * constructor ; pcuicfo (hnf ; auto). - eapply red_cofix_one_ty. eapply OnOne2_impl ; eauto. intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. @@ -2338,8 +2338,8 @@ Section RedConfluence. clear - Hctx. induction (fix_context mfix0). + assumption. + simpl. destruct a as [na [b|] ty]. - * constructor ; firstorder (hnf ; auto). - * constructor ; firstorder (hnf ; auto). + * constructor ; pcuicfo (hnf ; auto). + * constructor ; pcuicfo (hnf ; auto). - auto. - eapply red_trans; eauto. Qed. @@ -2637,7 +2637,7 @@ Section RedConfluence. Lemma clos_rt_red1_red1_rel_alpha Γ x y : clos_refl_trans (red1 Σ Γ) x y -> clos_refl_trans red1_rel_alpha (Γ, x) (Γ, y). Proof. - induction 1. constructor. red. left. firstorder. + induction 1. constructor. red. left. pcuicfo. constructor 2. econstructor 3; eauto. Qed. diff --git a/pcuic/theories/PCUICContextConversion.v b/pcuic/theories/PCUICContextConversion.v index 741669314..c226a00bd 100644 --- a/pcuic/theories/PCUICContextConversion.v +++ b/pcuic/theories/PCUICContextConversion.v @@ -890,6 +890,16 @@ Qed. Hint Extern 4 (eq_term_upto_univ _ _ _ _ _) => reflexivity : pcuic. +Axiom fix_guard_context_cumulativity : forall {cf:checker_flags} Σ Γ Γ' mfix, + cumul_context Σ Γ' Γ -> + fix_guard Σ Γ mfix -> + fix_guard Σ Γ' mfix. + +Axiom cofix_guard_context_cumulativity : forall {cf:checker_flags} Σ Γ Γ' mfix, + cumul_context Σ Γ' Γ -> + cofix_guard Σ Γ mfix -> + cofix_guard Σ Γ' mfix. + Lemma context_cumulativity_prop {cf:checker_flags} : env_prop (fun Σ Γ t T => @@ -951,19 +961,8 @@ Proof. - econstructor; pcuic. intuition auto. eapply isdecl. eapply isdecl. eauto. solve_all. destruct b0 as [? [? ?]]; eauto. - - econstructor; pcuic. - eapply (All_impl X0). - intros x [s [Hs IH]]. - exists s; eauto. - eapply (All_impl X1). - intros x [[Hs Hl] IH]. split; auto. - eapply IH. - now apply cumul_context_app_same. - eapply (All_mfix_wf); auto. - apply (All_impl X0); simpl. - intros x' [s [Hs' IH']]. exists s. - eapply IH'; auto. - - econstructor; pcuic. + - econstructor. eapply fix_guard_context_cumulativity; eauto. + all:pcuic. eapply (All_impl X0). intros x [s [Hs IH]]. exists s; eauto. @@ -975,6 +974,20 @@ Proof. apply (All_impl X0); simpl. intros x' [s [Hs' IH']]. exists s. eapply IH'; auto. + - econstructor. + eapply cofix_guard_context_cumulativity; eauto. + all:pcuic. + + eapply (All_impl X0). + intros x [s [Hs IH]]. + exists s; eauto. + + eapply (All_impl X1). + intros x [Hs IH]. + eapply IH. + now apply cumul_context_app_same. + eapply (All_mfix_wf); auto. + apply (All_impl X0); simpl. + intros x' [s [Hs' IH']]. exists s. + eapply IH'; auto. - econstructor; eauto. eapply cumul_cumul_ctx; eauto. diff --git a/pcuic/theories/PCUICContexts.v b/pcuic/theories/PCUICContexts.v index 39b491063..1a5a648e9 100644 --- a/pcuic/theories/PCUICContexts.v +++ b/pcuic/theories/PCUICContexts.v @@ -514,19 +514,19 @@ Proof. intros wfΣ. induction Γ in Δ |- *; simpl; auto. intros wfΓ wfΔ. depelim wfΓ; simpl. - - apply IHΓ; auto. eapply All_local_env_app_inv. split; auto. + - apply IHΓ; auto. eapply All_local_env_app. split; auto. repeat constructor; auto. eapply All_local_env_impl; eauto. simpl; intros. unfold lift_typing in X |- *. now rewrite app_context_assoc. - apply IHΓ. auto. eapply All_local_env_subst; eauto. simpl; intros. - destruct T; unfold lift_typing in X |- *; simpl in *; firstorder auto. + destruct T; unfold lift_typing in X |- *; simpl in *; pcuicfo auto. rewrite Nat.add_0_r. eapply (substitution _ (Δ' ,,, Γ) [vdef na b t] [b] Γ0) in X; eauto. rewrite -{1}(subst_empty 0 b). repeat constructor. now rewrite !subst_empty. - exists x. + destruct X as [s Hs]; exists s. rewrite Nat.add_0_r. - eapply (substitution _ _ [vdef na b t] [b] Γ0) in p; eauto. + eapply (substitution _ _ [vdef na b t] [b] Γ0) in Hs; eauto. rewrite -{1}(subst_empty 0 b). repeat constructor. now rewrite !subst_empty. Qed. @@ -543,17 +543,17 @@ Lemma wf_local_smash_end {cf:checker_flags} Σ Γ Δ : wf_local Σ (Γ ,,, Δ) -> wf_local Σ (Γ ,,, smash_context [] Δ). Proof. intros wfΣ wf. - apply All_local_env_app_inv. split. - now apply All_local_env_app in wf. + apply All_local_env_app. split. + now apply All_local_env_app_inv in wf. eapply wf_local_rel_smash_context; auto. Qed. Lemma wf_local_rel_empty {cf:checker_flags} Σ Γ : wf_local_rel Σ [] Γ <~> wf_local Σ Γ. Proof. split. - - intros h. eapply (All_local_env_impl _ _ _ h). firstorder. + - intros h. eapply (All_local_env_impl _ _ _ h). pcuicfo. red in X |- *. now rewrite app_context_nil_l in X. - - intros h. eapply (All_local_env_impl _ _ _ h). firstorder. + - intros h. eapply (All_local_env_impl _ _ _ h). pcuicfo. red in X |- *. now rewrite app_context_nil_l. Qed. diff --git a/pcuic/theories/PCUICConvCumInversion.v b/pcuic/theories/PCUICConvCumInversion.v index b90b3755b..1710f1d55 100644 --- a/pcuic/theories/PCUICConvCumInversion.v +++ b/pcuic/theories/PCUICConvCumInversion.v @@ -108,7 +108,8 @@ Section fixed. all: depelim e. all: depelim w0. all: apply All2_length in a. - all: constructor; constructor; rewrite a; auto. + all: try (constructor; constructor; rewrite a; auto). + all: destruct leq; cbn; repeat constructor. - clear -a1 a a0. induction a1 in args, args', x2, a, x3, a0, a1 |- *; depelim a; depelim a0; [now constructor|]. constructor. diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index f1c1d56cf..527ce3503 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -58,7 +58,7 @@ Proof. eapply cumul_alt in X0 as [w [w' [[redl' redr'] eq']]]. destruct (red_confluence wfΣ redr redl') as [nf [nfl nfr]]. eapply cumul_alt. - eapply red_eq_term_upto_univ_r in eq; tc;eauto with pcuic. + eapply red_eq_term_upto_univ_r in eq. all:tc;eauto with pcuic. destruct eq as [v'0 [red'0 eq2]]. eapply red_eq_term_upto_univ_l in eq'; tc;eauto with pcuic. destruct eq' as [v'1 [red'1 eq1]]. @@ -2595,12 +2595,12 @@ Proof. induction Δ as [|d Δ] in cl, wf0 |- *. - constructor. - simpl. - rewrite closedn_ctx_cons in cl. apply andP in cl as [clctx cld]. + rewrite closedn_ctx_cons in cl. apply andb_and in cl as [clctx cld]. simpl in wf0. destruct d as [na [b|] ty] => /=. * depelim wf0; simpl in *. simpl in cld. unfold closed_decl in cld. simpl in cld. simpl. - apply andP in cld as [clb clty]. + apply andb_and in cld as [clb clty]. constructor; auto. constructor; [reflexivity|..]. ** apply weaken_conv; auto; autorewrite with len. 1:now rewrite closedn_subst_instance_context. @@ -3047,8 +3047,8 @@ Proof. - simpl. constructor. - rewrite /= closed_ctx_decl in wf. rewrite /= closed_ctx_decl in wf'. - move/andP: wf => [wfd wf]. - move/andP: wf' => [wfd' wf']. + move/andb_and: wf => [wfd wf]. + move/andb_and: wf' => [wfd' wf']. constructor; auto. + now eapply IHX. + depelim p. constructor; auto. @@ -3058,12 +3058,12 @@ Proof. now autorewrite with len in wfd'. - rewrite /= closed_ctx_decl in wf. rewrite /= closed_ctx_decl in wf'. - move/andP: wf => [wfd wf]. - move/andP: wf' => [wfd' wf']. + move/andb_and: wf => [wfd wf]. + move/andb_and: wf' => [wfd' wf']. constructor; auto. + now eapply IHX. - + move/andP: wfd => /= [wfb wft]. - move/andP: wfd' => /= [wfb' wft']. + + move/andb_and: wfd => /= [wfb wft]. + move/andb_and: wfd' => /= [wfb' wft']. autorewrite with len in *. rewrite <- (context_relation_length X) in *. depelim p; constructor; auto. diff --git a/pcuic/theories/PCUICCumulProp.v b/pcuic/theories/PCUICCumulProp.v index 174ba913c..5de299dce 100644 --- a/pcuic/theories/PCUICCumulProp.v +++ b/pcuic/theories/PCUICCumulProp.v @@ -904,14 +904,14 @@ Proof. 14:{ specialize (X1 _ _ H X5 _ X6). now eapply cumul_prop_cum_l. } - 6:{ eapply inversion_App in X4 as (na' & A' & B' & hf & ha & cum); auto. - specialize (X1 _ _ H hf _ X5_1). - specialize (X3 _ _ H ha _ (eq_term_upto_univ_napp_leq X5_2)). + 6:{ eapply inversion_App in X6 as (na' & A' & B' & hf & ha & cum); auto. + specialize (X3 _ _ H hf _ X7_1). + specialize (X5 _ _ H ha _ (eq_term_upto_univ_napp_leq X7_2)). eapply cumul_cumul_prop in cum; auto. transitivity (B' {0 := u0}) => //. - eapply cumul_prop_prod_inv in X1 => //. + eapply cumul_prop_prod_inv in X3 => //. transitivity (B' {0 := u}). - now eapply substitution1_untyped_cumul_prop in X1. + now eapply substitution1_untyped_cumul_prop in X3. constructor. eapply eq_term_eq_term_prop_impl => //. eapply PCUICEquality.eq_term_upto_univ_substs. diff --git a/pcuic/theories/PCUICCumulativity.v b/pcuic/theories/PCUICCumulativity.v index 30de075ac..346361e62 100644 --- a/pcuic/theories/PCUICCumulativity.v +++ b/pcuic/theories/PCUICCumulativity.v @@ -1,6 +1,8 @@ (* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import config utils. -From MetaCoq.PCUIC Require Import PCUICRelations PCUICAst PCUICAstUtils +From Coq Require Import CRelationClasses. +From Equations.Type Require Import Relation Relation_Properties. +From MetaCoq.Template Require Import config utils BasicAst. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICLiftSubst PCUICEquality PCUICUnivSubst PCUICReduction. Set Default Goal Selector "!". @@ -53,6 +55,7 @@ Section ConvCumulDefs. End ConvCumulDefs. +(* todo mode typing notation *) Reserved Notation " Σ ;;; Γ |- t : T " (at level 50, Γ, t, T at next level). Reserved Notation " Σ ;;; Γ |- t <= u " (at level 50, Γ, t, u at next level). Reserved Notation " Σ ;;; Γ |- t = u " (at level 50, Γ, t, u at next level). @@ -286,10 +289,6 @@ Proof. eapply red_conv_conv_inv; eauto. now constructor. Qed. -Inductive conv_pb := -| Conv -| Cumul. - Definition conv_cum {cf:checker_flags} leq Σ Γ u v := match leq with | Conv => ∥ Σ ;;; Γ |- u = v ∥ diff --git a/pcuic/theories/PCUICElimination.v b/pcuic/theories/PCUICElimination.v index 3ce747b18..73b33e6df 100644 --- a/pcuic/theories/PCUICElimination.v +++ b/pcuic/theories/PCUICElimination.v @@ -19,7 +19,7 @@ Definition SingletonProp `{cf : checker_flags} (Σ : global_env_ext) (ind : indu declared_inductive (fst Σ) mdecl ind idecl -> forall Γ args u n (Σ' : global_env_ext), wf Σ' -> - PCUICWeakeningEnv.extends Σ Σ' -> + extends Σ Σ' -> welltyped Σ' Γ (mkApps (tConstruct ind n u) args) -> ∥Is_proof Σ' Γ (mkApps (tConstruct ind n u) args)∥ /\ #|ind_ctors idecl| <= 1 /\ @@ -30,7 +30,7 @@ Definition Computational `{cf : checker_flags} (Σ : global_env_ext) (ind : indu declared_inductive (fst Σ) mdecl ind idecl -> forall Γ args u n (Σ' : global_env_ext), wf Σ' -> - PCUICWeakeningEnv.extends Σ Σ' -> + extends Σ Σ' -> welltyped Σ' Γ (mkApps (tConstruct ind n u) args) -> Is_proof Σ' Γ (mkApps (tConstruct ind n u) args) -> False. @@ -39,7 +39,7 @@ Definition Informative `{cf : checker_flags} (Σ : global_env_ext) (ind : induct declared_inductive (fst Σ) mdecl ind idecl -> forall Γ args u n (Σ' : global_env_ext), wf_ext Σ' -> - PCUICWeakeningEnv.extends Σ Σ' -> + extends Σ Σ' -> Is_proof Σ' Γ (mkApps (tConstruct ind n u) args) -> #|ind_ctors idecl| <= 1 /\ squash (All (Is_proof Σ' Γ) (skipn (ind_npars mdecl) args)). @@ -292,12 +292,12 @@ Proof. eapply All_local_assum_subst; eauto. simpl. intros. destruct X as [s [Ht2 isp]]. - exists s; firstorder. + exists s; pcuicfo. rewrite Nat.add_0_r. eapply (substitution _ Γ [vdef na b ty] [b] Γ1 _ (tSort s)); auto. rewrite -{1}(subst_empty 0 b). repeat (constructor; auto). rewrite !subst_empty. eapply typing_wf_local in Ht2. - rewrite app_context_assoc in Ht2. eapply All_local_env_app in Ht2 as [Ht2 _]. + rewrite app_context_assoc in Ht2. eapply All_local_env_app_inv in Ht2 as [Ht2 _]. depelim Ht2. apply l0. now rewrite app_context_assoc in Ht2. * intros mdecl idec decli oib. @@ -342,7 +342,7 @@ Proof. eapply All_local_assum_subst; eauto. simpl. intros. destruct X as [s [Ht2 isp]]. - exists s; firstorder auto. + exists s; pcuicfo auto. rewrite Nat.add_0_r. eapply (substitution _ Γ [vass na ty] [t] Γ1 _ (tSort s)); auto. repeat (constructor; auto). now rewrite subst_empty. now rewrite app_context_assoc in Ht2. } @@ -477,7 +477,7 @@ Proof. unshelve eapply PCUICInductiveInversion.on_constructor_inst in oi; eauto. destruct oi as [oi _]. rewrite !subst_instance_context_app in oi. - now eapply wf_local_app in oi. } + now eapply wf_local_app_l in oi. } apply s. rewrite subst_app_context in X0. @@ -619,7 +619,7 @@ Proof. eapply build_branches_type_lookup in e2. eauto. 2:eauto. 3:destruct declc; eauto. - 2:{ eapply (All2_impl a); firstorder eauto. } + 2:{ eapply (All2_impl a); pcuicfo eauto. } destruct e2 as [nargs [br [brty [[[Hnth Hnth'] brtyped]]]]]. epose proof (All2_nth_error _ _ _ a H). specialize (X0 Hnth'). @@ -773,7 +773,7 @@ Proof. now apply is_prop_gt in lt. - rewrite app_context_assoc in eq. pose proof eq as eq'. - eapply All_local_env_app in eq' as [wfΓ wf']. depelim wfΓ; + eapply All_local_env_app_inv in eq' as [wfΓ wf']. depelim wfΓ; rewrite it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /= in X2 |- *. + eapply invert_cumul_prod_l in X2 as (na' & A & B' & (red & conv) & cum). eapply subject_reduction in X1. 3:eassumption. all:auto. @@ -837,7 +837,7 @@ Proof. apply is_prop_gt in lt; auto. - rewrite app_context_assoc in eq. pose proof eq as eq'. - eapply All_local_env_app in eq' as [wfΓ wf']. depelim wfΓ; + eapply All_local_env_app_inv in eq' as [wfΓ wf']. depelim wfΓ; rewrite it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /= in X2 |- *. + eapply invert_cumul_prod_r in X2 as (na' & A' & B' & (red & conv) & cum). eapply subject_reduction in X1. 3:eassumption. all:auto. @@ -885,7 +885,7 @@ Lemma cumul_prop1 (Σ : global_env_ext) Γ A B u : Proof. intros. destruct X0 as [s Hs]. - eapply cumul_prop_inv in H. 4:eauto. firstorder. auto. + eapply cumul_prop_inv in H. 4:eauto. pcuicfo. auto. right; eauto. Qed. @@ -899,7 +899,7 @@ Lemma cumul_prop2 (Σ : global_env_ext) Γ A B u : Proof. intros. destruct X0 as [s Hs]. - eapply cumul_prop_inv in H. 4:eauto. firstorder. auto. + eapply cumul_prop_inv in H. 4:eauto. pcuicfo. auto. left; eauto. Qed. @@ -913,7 +913,7 @@ Lemma cumul_sprop1 (Σ : global_env_ext) Γ A B u : Proof. intros. destruct X0 as [s Hs]. - eapply cumul_sprop_inv in H. 4:eauto. firstorder. auto. + eapply cumul_sprop_inv in H. 4:eauto. pcuicfo. auto. right; eauto. Qed. @@ -927,7 +927,7 @@ Lemma cumul_sprop2 (Σ : global_env_ext) Γ A B u : Proof. intros. destruct X0 as [s Hs]. - eapply cumul_sprop_inv in H. 4:eauto. firstorder. auto. + eapply cumul_sprop_inv in H. 4:eauto. pcuicfo. auto. left; eauto. Qed. diff --git a/pcuic/theories/PCUICEquality.v b/pcuic/theories/PCUICEquality.v index ab95796d1..9f6ce36e5 100644 --- a/pcuic/theories/PCUICEquality.v +++ b/pcuic/theories/PCUICEquality.v @@ -8,7 +8,6 @@ Require Import ssreflect. From Equations.Prop Require Import DepElim. Set Equations With UIP. - Definition R_universe_instance R := fun u u' => Forall2 R (List.map Universe.make u) (List.map Universe.make u'). @@ -201,7 +200,9 @@ Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) (x.(rarg) = y.(rarg)) * eq_binder_annot x.(dname) y.(dname) ) mfix mfix' -> - eq_term_upto_univ_napp Σ Re Rle napp (tCoFix mfix idx) (tCoFix mfix' idx). + eq_term_upto_univ_napp Σ Re Rle napp (tCoFix mfix idx) (tCoFix mfix' idx) + +| eq_Prim i : eq_term_upto_univ_napp Σ Re Rle napp (tPrim i) (tPrim i). Notation eq_term_upto_univ Σ Re Rle := (eq_term_upto_univ_napp Σ Re Rle 0). @@ -890,6 +891,8 @@ Fixpoint eqb_term_upto_univ_napp Σ (equ lequ : Universe.t -> Universe.t -> bool eqb_annot x.(dname) y.(dname) ) mfix mfix' + | tPrim p, tPrim p' => eqb p p' + | _, _ => false end. @@ -1030,6 +1033,7 @@ Proof. cbn -[eqb]. intros x X0 y. eqspec; [|rewrite andb_false_r; discriminate]. intro. rtoProp. split; tas. split;tas. split; eapply X0; tea. now apply eqb_annot_spec. + - eqspec; [|discriminate]. constructor. Qed. Lemma reflect_eq_term_upto_univ Σ equ lequ (Re Rle : Universe.t -> Universe.t -> Prop) napp : @@ -1222,6 +1226,7 @@ Proof. apply X2. -- constructor. intro bot. apply f. inversion bot. subst. inversion X0. subst. apply X2. + - cbn - [eqb]. eqspecs. do 2 constructor. Qed. Lemma compare_global_instance_refl : @@ -1239,6 +1244,15 @@ Proof. eapply forallb2_map, forallb2_refl; intro; apply eqb_refl. Qed. +Lemma eq_dec_to_bool_refl {A} {ea : Classes.EqDec A} (x : A) : + eq_dec_to_bool x x. +Proof. + unfold eq_dec_to_bool. + destruct (Classes.eq_dec x x). + constructor. + congruence. +Qed. + Lemma eqb_term_upto_univ_refl : forall Σ (eqb leqb : Universe.t -> Universe.t -> bool) napp t, (forall u, eqb u u) -> @@ -1281,6 +1295,7 @@ Proof. destruct p as [e1 e2]. rewrite -> e1 by assumption. rewrite -> e2 by assumption. rewrite eqb_annot_refl; assumption. + - apply eq_dec_to_bool_refl. Qed. (** ** Behavior on mkApps and it_mkLambda_or_LetIn ** *) diff --git a/pcuic/theories/PCUICGeneration.v b/pcuic/theories/PCUICGeneration.v index ecd8976f6..40755651e 100644 --- a/pcuic/theories/PCUICGeneration.v +++ b/pcuic/theories/PCUICGeneration.v @@ -38,9 +38,9 @@ Section Generation. intros. specialize (IHHsp (tApp t0 hd)). apply IHHsp. - eapply type_App. + destruct i as [s Hs]. + eapply type_App; eauto. eapply type_Cumul; eauto. - eapply i.π2. eauto. Qed. Lemma type_it_mkLambda_or_LetIn : diff --git a/pcuic/theories/PCUICGlobalEnv.v b/pcuic/theories/PCUICGlobalEnv.v new file mode 100644 index 000000000..bc59d5ee4 --- /dev/null +++ b/pcuic/theories/PCUICGlobalEnv.v @@ -0,0 +1,98 @@ +(* Distributed under the terms of the MIT license. *) +From Coq Require Import ProofIrrelevance. +From MetaCoq.Template Require Import config utils uGraph. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils + PCUICReflect PCUICTyping. + +Definition wf_global_uctx_invariants {cf:checker_flags} Σ : + wf Σ -> + global_uctx_invariants (global_uctx Σ). +Proof. + intros HΣ. split. + - cbn. eapply LevelSet.mem_spec, global_levels_Set. + - unfold global_uctx. + simpl. intros [[l ct] l'] Hctr. simpl in *. + induction Σ in HΣ, l, ct, l', Hctr |- *. + + apply ConstraintSetFact.empty_iff in Hctr; contradiction. + + simpl in *. apply ConstraintSet.union_spec in Hctr. + destruct Hctr as [Hctr|Hctr]. + * split. + -- inversion HΣ; subst. + destruct H2 as [HH1 [HH HH3]]. + subst udecl. destruct d as [decl|decl]; simpl in *. + ++ destruct decl; simpl in *. + destruct cst_universes ; [ + eapply (HH (l, ct, l') Hctr) + | apply ConstraintSetFact.empty_iff in Hctr ; contradiction + ]. + ++ destruct decl. simpl in *. + destruct ind_universes ; [ + eapply (HH (l, ct, l') Hctr) + | apply ConstraintSetFact.empty_iff in Hctr; contradiction + ]. + -- inversion HΣ. subst. + destruct H2 as [HH1 [HH HH3]]. + subst udecl. destruct d as [decl|decl]. + all: simpl in *. + ++ destruct decl. simpl in *. + destruct cst_universes ; [ + eapply (HH (l, ct, l') Hctr) + | apply ConstraintSetFact.empty_iff in Hctr; contradiction + ]. + ++ destruct decl. simpl in *. + destruct ind_universes; [ + eapply (HH (l, ct, l') Hctr) + | apply ConstraintSetFact.empty_iff in Hctr; contradiction + ]. + * inversion HΣ. subst. + split. + all: apply LevelSet.union_spec. + all: right. + all: unshelve eapply (IHΣ _ _ _ _ Hctr). + all: try eassumption. +Qed. + +Definition wf_ext_global_uctx_invariants {cf:checker_flags} Σ : + wf_ext Σ -> + global_uctx_invariants (global_ext_uctx Σ). +Proof. + intros HΣ. split. + - apply LevelSet.union_spec. right. apply LevelSet.mem_spec, global_levels_Set. + - destruct Σ as [Σ φ]. destruct HΣ as [HΣ Hφ]. + destruct (wf_global_uctx_invariants _ HΣ) as [_ XX]. + unfold global_ext_uctx, global_ext_levels, global_ext_constraints. + simpl. intros [[l ct] l'] Hctr. simpl in *. apply ConstraintSet.union_spec in Hctr. + destruct Hctr as [Hctr|Hctr]. + + destruct Hφ as [_ [HH _]]. apply (HH _ Hctr). + + specialize (XX _ Hctr). + split; apply LevelSet.union_spec; right; apply XX. +Qed. + +Lemma wf_consistent {cf:checker_flags} Σ : wf Σ -> consistent (global_constraints Σ). +Proof. + destruct Σ. + - exists {| valuation_mono := fun _ => 1%positive; valuation_poly := fun _ => 0 |}. + intros x Hx; now apply ConstraintSetFact.empty_iff in Hx. + - inversion 1; subst. subst udecl. clear -H2. + destruct H2 as [_ [_ [_ [v Hv]]]]. + exists v. intros ct Hc. apply Hv. simpl in *. + apply ConstraintSet.union_spec in Hc. destruct Hc. + apply ConstraintSet.union_spec; simpl. + + left. destruct d. + destruct c, cst_universes. assumption. + apply ConstraintSetFact.empty_iff in H; contradiction. + destruct m, ind_universes. assumption. + apply ConstraintSetFact.empty_iff in H; contradiction. + + apply ConstraintSet.union_spec; simpl. + now right. +Qed. + +Definition global_ext_uctx_consistent {cf:checker_flags} Σ + : wf_ext Σ -> consistent (global_ext_uctx Σ).2. +Proof. + intros HΣ. cbn. unfold global_ext_constraints. + unfold wf_ext, on_global_env_ext in HΣ. + destruct HΣ as [_ [_ [_ HH]]]. apply HH. +Qed. + + diff --git a/pcuic/theories/PCUICInduction.v b/pcuic/theories/PCUICInduction.v index d0b2a0d99..22194a6b9 100644 --- a/pcuic/theories/PCUICInduction.v +++ b/pcuic/theories/PCUICInduction.v @@ -12,6 +12,7 @@ Set Asymmetric Patterns. Allows to get the right induction principle on lists of terms appearing in the term syntax (in evar, applications, branches of cases and (co-)fixpoints. *) +Notation prim_ind P p := (P (tPrim p)). (** Custom induction principle on syntax, dealing with the various lists appearing in terms. *) @@ -35,6 +36,7 @@ Lemma term_forall_list_ind : (forall (s : projection) (t : term), P t -> P (tProj s t)) -> (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tCoFix m n)) -> + (forall p, prim_ind P p) -> forall t : term, P t. Proof. intros until t. revert t. @@ -231,9 +233,11 @@ Lemma term_forall_mkApps_ind : (forall (s : projection) (t : term), P t -> P (tProj s t)) -> (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tCoFix m n)) -> + (forall i, prim_ind P i) -> forall t : term, P t. Proof. intros until t. + rename X14 into Pprim. assert (Acc (MR lt size) t) by eapply measure_wf, Wf_nat.lt_wf. induction H. rename X14 into auxt. clear H. rename x into t. move auxt at top. diff --git a/pcuic/theories/PCUICInductiveInversion.v b/pcuic/theories/PCUICInductiveInversion.v index 4c87eb12a..8a886151c 100644 --- a/pcuic/theories/PCUICInductiveInversion.v +++ b/pcuic/theories/PCUICInductiveInversion.v @@ -329,7 +329,7 @@ Proof. intros wfΣ H. depind H. - unfold unfold_fix. rewrite e. specialize (nth_error_all e a0) as [s Hs]. - specialize (nth_error_all e a1) as [Hty Hlam]. + specialize (nth_error_all e a1) as Hty. simpl. destruct decl as [name ty body rarg]; simpl in *. clear e. @@ -359,7 +359,7 @@ Proof. apply IHAll. --- intros. simpl in Hi. specialize (Hi (S i)). apply Hi. lia. --- lia. - ** clear IHAll. destruct p. + ** clear IHAll. simpl in Hlen. assert ((Nat.pred #|mfix| - (#|mfix| - S #|l|)) = #|l|) by lia. rewrite H0. rewrite simpl_subst_k. --- clear. induction l; simpl; auto with arith. @@ -379,7 +379,7 @@ Qed. Lemma subslet_cofix {cf:checker_flags} (Σ : global_env_ext) Γ mfix : wf_local Σ Γ -> - cofix_guard mfix -> + cofix_guard Σ Γ mfix -> All (fun d : def term => ∑ s : Universe.t, Σ;;; Γ |- dtype d : tSort s) mfix -> All (fun d : def term => @@ -456,7 +456,7 @@ Proof. intros eqr. exists ind. destruct (map_option_out (map check_one_cofix mfix)) eqn:eqfixes. - move/andP: eqr => [eqinds rk]. + move/andb_and: eqr => [eqinds rk]. split; auto. constructor. - move: hcof. unfold check_one_cofix. @@ -482,7 +482,7 @@ Proof. noconf eqfixes. specialize (IHmfix _ eq_refl). simpl in eqinds. - move/andP: eqinds => [eqk eql0]. + move/andb_and: eqinds => [eqk eql0]. constructor; auto. clear IHmfix hcofs d. destruct a as [dname dbody dtype rarg] => /=. unfold check_one_cofix in hcof. @@ -588,7 +588,7 @@ Proof. apply weaken_wf_local => //. rewrite [_ ,,, _]app_context_assoc in wfΓ. - eapply All_local_env_app in wfΓ as [? ?]. + eapply All_local_env_app_inv in wfΓ as [? ?]. apply on_minductive_wf_params_indices => //. pcuic. Qed. @@ -716,8 +716,8 @@ Qed. Lemma test_decl_impl (f g : term -> bool) x : (forall x, f x -> g x) -> test_decl f x -> test_decl g x. Proof. intros Hf; rewrite /test_decl. - move/andP=> [Hd Hb]. - apply/andP; split; eauto. + move/andb_and=> [Hd Hb]. + apply/andb_and; split; eauto. eapply option_all_impl; eauto. Qed. @@ -733,7 +733,7 @@ Lemma closedn_ctx_upwards k k' Γ : closedn_ctx k' Γ. Proof. induction Γ; auto. rewrite !closed_ctx_decl /=. - move/andP => [cla clΓ] le. + move/andb_and => [cla clΓ] le. rewrite (IHΓ clΓ le). rewrite (closed_decl_upwards _ _ cla) //. lia. Qed. @@ -769,9 +769,9 @@ Proof. induction Γ as [|[na [b|] ty] Γ] using ctx_length_rev_ind; intros k k' t; rewrite ?closedn_ctx_app /= /id /=; simpl; auto. - now rewrite /expand_lets /expand_lets_k /= subst_empty lift0_id. - - rewrite andb_true_r /id; move/andP=> [cld clΓ]. len. + - rewrite andb_true_r /id; move/andb_and=> [cld clΓ]. len. rewrite !expand_lets_k_vdef. - simpl in cld |- *. move/andP: cld => /= [clb _]. + simpl in cld |- *. move/andb_and: cld => /= [clb _]. rewrite Nat.add_0_r in clb. specialize (H (subst_context [b] 0 Γ) ltac:(len; lia)). autorewrite with len in H. rewrite H /=. @@ -779,7 +779,7 @@ Proof. now rewrite Nat.add_0_r. now rewrite /= clb. rewrite -Nat.add_assoc -closedn_subst_eq. simpl. now rewrite clb. simpl; lia_f_equal. - - len. rewrite andb_true_r /=; move/andP => [clty clΓ]. + - len. rewrite andb_true_r /=; move/andb_and => [clty clΓ]. rewrite !expand_lets_k_vass. simpl. specialize (H Γ ltac:(len; lia) (S k)). rewrite Nat.add_assoc !Nat.add_succ_r !Nat.add_0_r. apply H. @@ -925,13 +925,13 @@ Proof. induction Γ in k |- *; intros cl; simpl; auto. destruct a as [na [b|] ty] => /=. len. - rewrite closed_ctx_decl in cl. move/andP: cl => [cld clΓ]. + rewrite closed_ctx_decl in cl. move/andb_and: cl => [cld clΓ]. simpl. f_equal. rewrite distr_subst. len. simpl in cld. rewrite IHΓ //. f_equal. rewrite simpl_subst_k. len. rewrite context_assumptions_smash_context /= //. - rewrite lift_closed //. now move/andP: cld => /= //. + rewrite lift_closed //. now move/andb_and: cld => /= //. rewrite IHΓ //. simpl map. @@ -950,7 +950,7 @@ Proof. rewrite lift_extended_subst. rewrite map_map_compose. rewrite map_subst_lift1. - rewrite closed_ctx_decl in cl. move/andP: cl => [cld clΓ]. + rewrite closed_ctx_decl in cl. move/andb_and: cl => [cld clΓ]. now rewrite IHΓ // Nat.add_1_r. Qed. @@ -1003,13 +1003,13 @@ Lemma subst_map_lift_lift_context (Γ : context) k s : Proof. induction Γ as [|[? [] ?] ?] in k |- *; intros cl; auto; rewrite lift_context_snoc !subst_context_snoc /= /subst_decl /map_decl /=; - rewrite closed_ctx_decl in cl; move/andP: cl => [cld clΓ]. + rewrite closed_ctx_decl in cl; move/andb_and: cl => [cld clΓ]. - rewrite IHΓ //. f_equal. f_equal. f_equal; len. - rewrite closed_subst_map_lift //. now move/andP: cld => /=. + rewrite closed_subst_map_lift //. now move/andb_and: cld => /=. lia_f_equal. len. - rewrite closed_subst_map_lift //. now move/andP: cld => /=. + rewrite closed_subst_map_lift //. now move/andb_and: cld => /=. lia_f_equal. - f_equal. apply IHΓ => //. f_equal; len. rewrite closed_subst_map_lift //. @@ -1076,7 +1076,7 @@ Proof. { unshelve epose proof (on_minductive_wf_params_indices _ _ _ _ wfΣ _ oib). pcuic. eapply closed_wf_local in X => //. rewrite closedn_ctx_app in X; simpl; eauto. - move/andP: X; intuition auto; + move/andb_and: X; intuition auto; now rewrite closedn_subst_instance_context. } assert (closedn_ctx (#|ind_params mdecl| + #|cshape_args cshape|) (subst_instance_context u (ind_indices oib))) as clinds'. @@ -1701,7 +1701,7 @@ Proof. specialize (X Δ ltac:(len; lia) _ _ H). simpl; autorewrite with len in X |- *. destruct X; split; auto. simpl. - eapply All_local_env_app_inv; split. + eapply All_local_env_app; split. constructor; auto. eapply (All_local_env_impl _ _ _ a). intros; auto. now rewrite app_context_assoc. simpl. @@ -2066,7 +2066,7 @@ Proof. autorewrite with len in IHpos. eapply IHpos; eauto. simpl; constructor; auto. simpl in isty. eapply isType_tProd in isty as [Hty Ht]; eauto. - rewrite closedn_ctx_cons /=. apply/andP; split; auto. simpl. + rewrite closedn_ctx_cons /=. apply/andb_and; split; auto. simpl. eapply closed_upwards; eauto; lia. simpl in isty. eapply isType_tProd in isty as [Hty Ht]; eauto. Qed. @@ -2144,7 +2144,7 @@ Proof. now rewrite subst_instance_context_app app_context_assoc. * eapply closed_wf_local in wf; eauto. rewrite closedn_subst_instance_context app_assoc in wf. - now rewrite closedn_ctx_app /= in wf; move/andP: wf => [wfars wf]. + now rewrite closedn_ctx_app /= in wf; move/andb_and: wf => [wfars wf]. * now rewrite subst_instance_context_app app_context_assoc. - elimtype False; now depelim ass. - elimtype False; now depelim ass. @@ -2190,7 +2190,7 @@ Proof. rewrite -it_mkProd_or_LetIn_app in cpos. eapply positive_cstr_it_mkProd_or_LetIn in cpos as [hpars _]. rewrite smash_context_app_expand in hpars. - eapply All_local_env_app in hpars as [_hpars hargs]. + eapply All_local_env_app_inv in hpars as [_hpars hargs]. rewrite expand_lets_smash_context /= expand_lets_k_ctx_nil /= in hargs. eapply positive_cstr_closed_args_subst_arities in hargs; eauto. split. @@ -2523,14 +2523,14 @@ Proof. specialize (clcstra _ H0). simpl in *. destruct x as [[l c] r]; rewrite /subst_instance_cstr; simpl. - move/andP: clcstra => [cll clr]. + move/andb_and: clcstra => [cll clr]. rewrite !closedu_subst_instance_level_app //. subst c; exists x; split; auto. specialize (clcstra _ H0). simpl in *. destruct x as [[l c] r]; rewrite /subst_instance_cstr; simpl. - move/andP: clcstra => [cll clr]. + move/andb_and: clcstra => [cll clr]. rewrite !closedu_subst_instance_level_app //. Qed. @@ -2586,17 +2586,17 @@ Proof. specialize (clcstra _ inc'). simpl in *. destruct c' as [[l c] r]; rewrite /subst_instance_cstr; simpl. - move/andP: clcstra => [cll clr]. + move/andb_and: clcstra => [cll clr]. rewrite !closedu_subst_instance_level_lift //. - subst c. exists (lift_constraint #|u| x). rewrite -> In_lift_constraints. - firstorder eauto. + pcuicfo eauto. specialize (clcstra _ H0). simpl in *. destruct x as [[l c] r]; rewrite /subst_instance_cstr; simpl. - move/andP: clcstra => [cll clr]. + move/andb_and: clcstra => [cll clr]. rewrite !closedu_subst_instance_level_lift //. Qed. @@ -2615,8 +2615,9 @@ Proof. subst. exists x; intuition eauto. now rewrite ConstraintSetFact.add_iff. - firstorder eauto. - subst. exists x0; firstorder auto with *. + destruct H0 as [c' [-> ?]]. + eexists c'; split; firstorder eauto. + now rewrite ConstraintSetFact.add_iff. Qed. Lemma subst_instance_variance_cstrs l u i i' : @@ -3046,13 +3047,13 @@ Proof. eapply (weaken_cumul_ctx _ Γ) in args => //. 4:eapply spu. 2:{ eapply closed_wf_local; eauto. } - 2:{ rewrite closedn_ctx_app; apply /andP. split => //. simpl. + 2:{ rewrite closedn_ctx_app; apply /andb_and. split => //. simpl. len. simpl. eapply closedn_smash_context => //. len; simpl. pose proof (on_minductive_wf_params_indices_inst _ _ _ _ _ wfΣ (proj1 decli) oib cu') as wf'. eapply closed_wf_local in wf'; eauto. rewrite subst_instance_context_app in wf'. - rewrite closedn_ctx_app in wf'. move/andP: wf'=> [_ clargs]. + rewrite closedn_ctx_app in wf'. move/andb_and: wf'=> [_ clargs]. simpl in clargs; autorewrite with len in clargs. eapply closedn_smash_context => //. rewrite closedn_subst_instance_context. @@ -3181,14 +3182,14 @@ Proof. eapply (weaken_cumul_ctx _ Γ) in args => //. 4:eapply spu. 2:{ eapply closed_wf_local; eauto. } - 2:{ rewrite closedn_ctx_app; apply /andP. split => //. simpl. + 2:{ rewrite closedn_ctx_app; apply /andb_and. split => //. simpl. len. simpl. eapply closedn_smash_context. rewrite -(Nat.add_0_l (context_assumptions _)). eapply closedn_ctx_subst. len; simpl. 2:{ eapply declared_minductive_closed_inds; eauto. } epose proof (on_constructor_wf_args _ _ _ _ _ _ wfΣ decli onind oib onc) as wf'; eauto. eapply closed_wf_local in wf'; eauto. - rewrite closedn_ctx_app in wf'. move/andP: wf'=> [_ clargs]. + rewrite closedn_ctx_app in wf'. move/andb_and: wf'=> [_ clargs]. simpl in clargs; autorewrite with len in clargs. rewrite closedn_subst_instance_context. rewrite Nat.add_comm. @@ -3290,7 +3291,7 @@ Proof. eapply closedn_ctx_subst; cbn. rewrite /ind_subst; len. epose proof (on_constructor_inst u wfΣ decli onind _ onc cu) as [wfarpars _]; auto. move/closed_wf_local: wfarpars. rewrite !subst_instance_context_app closedn_ctx_app. - now move/andP; len. + now move/andb_and; len. rewrite /ind_subst. eapply declared_minductive_closed_inds; eauto. } { len. simpl. rewrite expand_lets_app in clx. rewrite -(subst_instance_context_assumptions u' (ind_params _)). @@ -3309,7 +3310,7 @@ Proof. eapply closedn_ctx_subst; cbn. rewrite /ind_subst; len. epose proof (on_constructor_inst _ wfΣ decli onind _ onc cu') as [wfarpars _]; auto. move/closed_wf_local: wfarpars. rewrite !subst_instance_context_app closedn_ctx_app. - now move/andP; len. + now move/andb_and; len. rewrite /ind_subst. eapply declared_minductive_closed_inds; eauto. now rewrite /argctx /argctx'; len. } rewrite smash_context_app smash_context_acc in cxy. @@ -3409,7 +3410,7 @@ Proof. eapply cumul_ctx_subst_instance => //. rewrite -app_context_assoc. eapply weaken_wf_local; eauto. rewrite !subst_instance_context_app in wfargs. - now eapply All_local_env_app in wfargs as [wfargs _]. + now eapply All_local_env_app_inv in wfargs as [wfargs _]. } { rewrite /pargctx. rewrite (smash_context_subst []). @@ -3876,7 +3877,7 @@ Proof. fold parsubst. move: (context_subst_length sppars); len => <-. epose proof (on_minductive_wf_params_indices_inst _ _ _ _ _ wfΣ (proj1 isdecl) oib cu). eapply PCUICClosed.closed_wf_local in X; auto. move: X. - now rewrite subst_instance_context_app PCUICClosed.closedn_ctx_app /=; len => /andP [_ H]. + now rewrite subst_instance_context_app PCUICClosed.closedn_ctx_app /=; len => /andb_and [_ H]. rewrite lift_context_subst_context //. + eapply isType_weakening; eauto. eapply spargs. @@ -3957,8 +3958,8 @@ Proof. { eapply validity in Hp; auto. eapply PCUICWfUniverses.isType_wf_universes in Hp. rewrite PCUICWfUniverses.wf_universes_it_mkProd_or_LetIn in Hp. - move/andP: Hp => [_ Hp]. - now apply (PCUICWfUniverses.reflect_bP (PCUICWfUniverses.wf_universe_reflect _ _)) in Hp. auto. } + move/andb_and: Hp => [_ Hp]. + now apply (ssrbool.elimT PCUICWfUniverses.wf_universe_reflect) in Hp. auto. } rewrite !subst_instance_context_app in wf. assert (sorts_local_ctx (lift_typing typing) Σ Γ (subst_context parsubst 0 @@ -3975,7 +3976,7 @@ Proof. eapply (subst_sorts_local_ctx) in X; simpl in *; eauto. 3:{ eapply subslet_inds; eauto. } 2:{ rewrite app_context_nil_l. - now eapply All_local_env_app in wf as [? ?]. } + now eapply All_local_env_app_inv in wf as [? ?]. } simpl in X. len in X. eapply weaken_sorts_local_ctx in X. 2:eauto. 2:eapply typing_wf_local; eauto. rewrite app_context_nil_l in X. @@ -4063,7 +4064,7 @@ Proof. fold parsubst. move: (context_subst_length sppars); len => <-. epose proof (on_minductive_wf_params_indices_inst _ _ _ _ _ wfΣ (proj1 decli) oib cu). eapply PCUICClosed.closed_wf_local in X0; auto. move: X0. - now rewrite subst_instance_context_app PCUICClosed.closedn_ctx_app /=; len => /andP [_ ?]. + now rewrite subst_instance_context_app PCUICClosed.closedn_ctx_app /=; len => /andb_and [_ ?]. rewrite lift_context_subst_context //. } split. { eapply isType_it_mkProd_or_LetIn; eauto. @@ -4137,7 +4138,7 @@ Proof. eapply (closedn_ctx_subst 0). len. simpl. 2:eapply declared_minductive_closed_inds; eauto. eapply closed_wf_local in wf; eauto. move: wf. - now rewrite !closedn_ctx_app /=; len => /andP [_ ?]. + now rewrite !closedn_ctx_app /=; len => /andb_and [_ ?]. now rewrite lift_context_subst_context. } len. rewrite (subst_cstr_concl_head ind u mdecl (cshape_args cs) _ _). diff --git a/pcuic/theories/PCUICInductives.v b/pcuic/theories/PCUICInductives.v index 0761f8d14..8cdfe1dd0 100644 --- a/pcuic/theories/PCUICInductives.v +++ b/pcuic/theories/PCUICInductives.v @@ -589,7 +589,7 @@ Proof. (arities_context (ind_bodies mdecl))). { eapply wf_arities_context; eauto. } eapply PCUICClosed.sorts_local_ctx_All_local_env in wfargs. - 2:{ eapply All_local_env_app_inv. split; auto. + 2:{ eapply All_local_env_app. split; auto. red in onpars. eapply (All_local_env_impl _ _ _ onpars). intros. destruct T; simpl in *. - eapply weaken_ctx; auto. @@ -676,7 +676,7 @@ Proof. eapply declared_minductive_closed_inds. 2:destruct isdecl as [ [] ?]; eauto. eauto. } rewrite -app_assoc in wfl. - apply All_local_env_app in wfl as [wfctx wfsargs]. + apply All_local_env_app_inv in wfl as [wfctx wfsargs]. rewrite smash_context_app in Heq'. rewrite smash_context_acc in Heq'. rewrite nth_error_app_lt in Heq'. @@ -717,8 +717,8 @@ Proof. { assert(wf_local (Σ.1, ind_universes mdecl) (arities_context (ind_bodies mdecl) ,,, ind_params mdecl ,,, smash_context [] (cshape_args c))). - apply All_local_env_app_inv; split; auto. - now apply All_local_env_app in wfargs as [wfindpars wfargs]. + apply All_local_env_app; split; auto. + now apply All_local_env_app_inv in wfargs as [wfindpars wfargs]. apply wf_local_rel_smash_context; auto. eapply closed_wf_local in X0; auto. eapply (closedn_ctx_decl (n:=idx)) in X0; eauto. @@ -1190,21 +1190,21 @@ Lemma invert_type_mkApps_ind {cf:checker_flags} Σ Γ ind u args T mdecl idecl : Proof. intros wfΣ decli. intros H; dependent induction H; solve_discr. - - destruct args using rev_case; solve_discr. noconf H1. - rewrite -PCUICAstUtils.mkApps_nested in H1. simpl in H1. - noconf H1. clear IHtyping2. - specialize (IHtyping1 _ _ _ _ _ _ _ wfΣ decli eq_refl) as [IH cu]; + - destruct args using rev_case; solve_discr. noconf H2. + rewrite -PCUICAstUtils.mkApps_nested in H2. simpl in H2. + noconf H2. clear IHtyping1 IHtyping3. + specialize (IHtyping2 _ _ _ _ _ _ _ wfΣ decli eq_refl) as [IH cu]; split; auto. destruct (on_declared_inductive wfΣ decli) as [onmind oib]. eapply typing_spine_app; eauto. - - invs H0. destruct (declared_inductive_inj isdecl decli) as [-> ->]. + - invs H0. destruct (declared_inductive_inj d decli) as [-> ->]. clear decli. split; auto. constructor; [|reflexivity]. - destruct (on_declared_inductive wfΣ isdecl) as [onmind oib]. + destruct (on_declared_inductive wfΣ d) as [onmind oib]. pose proof (oib.(onArity)) as ar. eapply isType_weaken; eauto. eapply (isType_subst_instance_decl _ []); eauto. - destruct isdecl; eauto. + destruct d; eauto. eapply oib.(onArity). auto. - specialize (IHtyping1 _ _ wfΣ decli) as [IH cu]; split; auto. eapply typing_spine_weaken_concl; pcuic. @@ -1242,7 +1242,7 @@ Proof. rewrite subst_instance_context_length in subp, suba. subst parctx argctx. repeat split; eauto; rewrite ?subst_instance_context_length => //. - rewrite app_context_assoc in wfcodom. now apply All_local_env_app in wfcodom as [? ?]. + rewrite app_context_assoc in wfcodom. now apply All_local_env_app_inv in wfcodom as [? ?]. simpl. eapply substitution_wf_local; eauto. now rewrite app_context_assoc in wfcodom. unshelve eapply on_inductive_inst in declm; pcuic. @@ -1300,7 +1300,10 @@ Lemma invert_red1_letin {cf:checker_flags} (Σ : global_env_ext) Γ C na d ty b red1 Σ.1 (Γ ,, vdef na d ty) b b') + (C = subst10 d b)%type. Proof. - intros H; depelim H; try solve_discr; firstorder auto. + intros H; depelim H; try solve_discr; pcuicfo auto. + - left. left. left. eexists; eauto. + - left. left. right. eexists; eauto. + - left. right. eexists; eauto. Qed. Lemma decompose_prod_assum_it_mkProd_or_LetIn' ctx Γ t : @@ -1360,7 +1363,7 @@ Lemma destInd_head_subst s k t f : destInd (head (subst s k t)) = Some f -> (destInd (head t) = Some f) + (∑ n u, (head t = tRel n) /\ k <= n /\ (nth_error s (n - k) = Some u /\ destInd (head (lift0 k u)) = Some f)). Proof. - induction t in s, k, f |- *; simpl; try solve [firstorder auto]. + induction t in s, k, f |- *; simpl; try solve [pcuicfo auto]. elim: leb_spec_Set => le; intros destI. right. destruct nth_error eqn:Heq. diff --git a/pcuic/theories/PCUICInversion.v b/pcuic/theories/PCUICInversion.v index a9a128af3..089f47d55 100644 --- a/pcuic/theories/PCUICInversion.v +++ b/pcuic/theories/PCUICInversion.v @@ -209,13 +209,11 @@ Section Inversion. Σ ;;; Γ |- tFix mfix n : T -> ∑ decl, let types := fix_context mfix in - fix_guard mfix × + fix_guard Σ Γ mfix × nth_error mfix n = Some decl × All (fun d => isType Σ Γ (dtype d)) mfix × All (fun d => - Σ ;;; Γ ,,, types |- dbody d : (lift0 #|types|) (dtype d) × - isLambda (dbody d) = true - ) mfix × + Σ ;;; Γ ,,, types |- dbody d : (lift0 #|types|) (dtype d)) mfix × wf_fixpoint Σ mfix × Σ ;;; Γ |- dtype decl <= T. Proof. @@ -226,7 +224,7 @@ Section Inversion. forall {Γ mfix idx T}, Σ ;;; Γ |- tCoFix mfix idx : T -> ∑ decl, - cofix_guard mfix × + cofix_guard Σ Γ mfix × let types := fix_context mfix in nth_error mfix idx = Some decl × All (fun d => isType Σ Γ (dtype d)) mfix × @@ -239,6 +237,14 @@ Section Inversion. intros Γ mfix idx T h. invtac h. Qed. + (** At this stage we don't typecheck primitive values *) + Lemma inversion_Prim : + forall {Γ i T}, + Σ ;;; Γ |- tPrim i : T -> False. + Proof. + intros Γ i T h. now depind h. + Qed. + Lemma inversion_it_mkLambda_or_LetIn : forall {Γ Δ t T}, Σ ;;; Γ |- it_mkLambda_or_LetIn Δ t : T -> diff --git a/pcuic/theories/PCUICLiftSubst.v b/pcuic/theories/PCUICLiftSubst.v index e6ff1f1bc..cdf0c1567 100644 --- a/pcuic/theories/PCUICLiftSubst.v +++ b/pcuic/theories/PCUICLiftSubst.v @@ -189,7 +189,7 @@ Notation closed t := (closedn 0 t). Fixpoint noccur_between k n (t : term) : bool := match t with - | tRel i => Nat.leb k i && Nat.ltb i (k + n) + | tRel i => Nat.ltb i k || Nat.leb (k + n) i | tEvar ev args => List.forallb (noccur_between k n) args | tLambda _ T M | tProd _ T M => noccur_between k n T && noccur_between (S k) n M | tApp u v => noccur_between k n u && noccur_between k n v @@ -1793,6 +1793,7 @@ Hint Rewrite @subst_inst : sigma. Hint Rewrite shiftk_shift_l shiftk_shift : sigma. (* Hint Rewrite Upn_eq : sigma. *) +(** Can't move to PCUICInduction because of fix_context definition *) Lemma term_forall_ctx_list_ind : forall P : context -> term -> Type, (forall Γ (n : nat), P Γ (tRel n)) -> @@ -1813,6 +1814,7 @@ Lemma term_forall_ctx_list_ind : (forall Γ (s : projection) (t : term), P Γ t -> P Γ (tProj s t)) -> (forall Γ (m : mfixpoint term) (n : nat), tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tFix m n)) -> (forall Γ (m : mfixpoint term) (n : nat), tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tCoFix m n)) -> + (forall Γ p, P Γ (tPrim p)) -> forall Γ (t : term), P Γ t. Proof. intros. revert Γ t0. @@ -1841,6 +1843,7 @@ Proof. split. apply auxt. apply auxt. apply auxm. Defined. + Fixpoint subst_app (t : term) (us : list term) : term := match t, us with | tLambda _ A t, u :: us => subst_app (t {0 := u}) us diff --git a/pcuic/theories/PCUICNameless.v b/pcuic/theories/PCUICNameless.v index af703f73e..198c9d82d 100644 --- a/pcuic/theories/PCUICNameless.v +++ b/pcuic/theories/PCUICNameless.v @@ -47,6 +47,7 @@ Fixpoint nameless (t : term) : bool := | tCoFix mfix idx => forallb (fun d => banon d.(dname)) mfix && forallb (test_def nameless nameless) mfix + | tPrim _ => true end. Definition anonymize (b : binder_annot name) : binder_annot name := @@ -76,6 +77,7 @@ Fixpoint nl (t : term) : term := | tProj p c => tProj p (nl c) | tFix mfix idx => tFix (map (map_def_anon nl nl) mfix) idx | tCoFix mfix idx => tCoFix (map (map_def_anon nl nl) mfix) idx + | tPrim p => tPrim p end. Definition map_decl_anon f (d : context_decl) := {| @@ -122,7 +124,7 @@ Definition nlg (Σ : global_env_ext) : global_env_ext := Ltac destruct_one_andb := lazymatch goal with | h : is_true (_ && _) |- _ => - apply andP in h ; destruct h as [? ?] + apply andb_and in h ; destruct h as [? ?] end. Ltac destruct_andb := @@ -1338,10 +1340,11 @@ Proof. = nlctx (Γ ,,, fix_context mfix)) by now rewrite <- nl_fix_context, <- nlctx_app_context. constructor. - + eapply fix_guard_eq_term with (idx:=n). 1: eassumption. + + todo "fix_guard spec". + (*eapply fix_guard_eq_term with (idx:=n). 1: eassumption. constructor. clear. induction mfix. 1: constructor. simpl. constructor; tas. cbn. - repeat split; now apply eq_term_upto_univ_tm_nl. + repeat split; now apply eq_term_upto_univ_tm_nl.*) + now rewrite nth_error_map H0. + auto. + clear -X0. @@ -1353,7 +1356,6 @@ Proof. * rewrite fix_context_length map_length. rewrite fix_context_length in Hs. now rewrite -> XX, <- nl_lift. - * destruct dbody; simpl in *; congruence. + now rewrite <-nl_wf_fixpoint. - replace (nl (dtype decl)) with (dtype (map_def_anon nl nl decl)); [|destruct decl; reflexivity]. @@ -1361,10 +1363,11 @@ Proof. = nlctx (Γ ,,, fix_context mfix)) by now rewrite <- nl_fix_context, <- nlctx_app_context. constructor; auto. - + eapply cofix_guard_eq_term with (idx:=n). 1: eassumption. + + todo "cofix_guard eq_term". + (* eapply cofix_guard_eq_term with (idx:=n). 1: eassumption. constructor. clear. induction mfix. 1: constructor. simpl. constructor; tas. cbn. - repeat split; now apply eq_term_upto_univ_tm_nl. + repeat split; now apply eq_term_upto_univ_tm_nl.*) + now rewrite nth_error_map H0. + clear -X0. apply All_map. eapply All_impl; tea. diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index c4dfa2dc2..679a65f37 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -52,6 +52,7 @@ Section Normal. end -> whnf Γ (mkApps (tFix mfix idx) v) | whnf_cofixapp mfix idx v : whnf Γ (mkApps (tCoFix mfix idx) v) + | whnf_prim p : whnf Γ (tPrim p) with whne (Γ : context) : term -> Type := | whne_rel i : @@ -155,7 +156,7 @@ Section Normal. right. exists x0, x1, x2, x3, x4. repeat split; eauto. now eapply nth_error_app_left. + rewrite <- mkApps_snoc in x. eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. + rewrite !decompose_app_mkApps in x; cbn in *; try intuition congruence. inversion x. subst. right. exists mfix, idx, rarg, body, arg. repeat split; eauto. Qed. @@ -173,7 +174,7 @@ Local Ltac inv H := inversion H; subst; clear H. Ltac help' := try repeat match goal with | [ H0 : _ = mkApps _ _ |- _ ] => eapply (f_equal decompose_app) in H0; - rewrite !decompose_app_mkApps in H0; cbn in *; firstorder congruence + rewrite !decompose_app_mkApps in H0; cbn in *; intuition congruence | [ H1 : tApp _ _ = mkApps _ ?args |- _ ] => destruct args using rev_ind; cbn in *; [ inv H1 | rewrite <- mkApps_nested in H1; cbn in *; inv H1 @@ -198,7 +199,7 @@ Defined. Lemma negb_is_true b : ~ is_true b -> is_true (negb b). Proof. - destruct b; firstorder. + destruct b; pcuicfo. Qed. Hint Resolve negb_is_true : core. @@ -249,6 +250,8 @@ Proof. lia. - destruct (mkApps_elim t l). apply mkApps_eq_inj in eq as (<-&<-); auto. + - destruct l using MCList.rev_ind; [|now rewrite <- mkApps_nested in eq]. + cbn in *; subst; auto. Qed. Lemma whnf_fixapp' {flags} Σ Γ mfix idx narg body v : @@ -307,6 +310,8 @@ Proof. - inv whn; solve_discr; easy. Qed. +Set Firstorder Solver auto with core. + Definition whnf_whne_dec flags Σ Γ t : ({∥whnf flags Σ Γ t∥} + {~∥whnf flags Σ Γ t∥}) * ({∥whne flags Σ Γ t∥} + {~∥whne flags Σ Γ t∥}). @@ -335,18 +340,18 @@ Proof with eauto using sq with pcuic. right. intros [w]. depelim w. depelim w. all:help. clear IHt. eapply whne_mkApps_inv in w as []... -- depelim w. help. - -- firstorder congruence. + -- destruct s0 as [? [? [? [? [? [? ?]]]]]]. congruence. * destruct v as [ | ? v]... right. intros [w]. depelim w. depelim w. all:help. clear IHt. eapply whne_mkApps_inv in w as []... -- depelim w. help. - -- firstorder congruence. + -- destruct s0 as (?&?&?&?&?&?&?); congruence. * destruct v as [ | ? v]... destruct (RedFlags.beta flags) eqn:?. -- right. intros [w]. depelim w. depelim w. all:help. clear IHl. eapply whne_mkApps_inv in w as []... ++ depelim w. congruence. help. - ++ firstorder congruence. + ++ destruct s0 as (?&?&?&?&?&?&?); congruence. -- left; constructor. apply whnf_mkApps. constructor. @@ -360,23 +365,29 @@ Proof with eauto using sq with pcuic. eapply whne_mkApps_inv in w as []; eauto. --- depelim w; [|congruence]. help. eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. + rewrite !decompose_app_mkApps in x; cbn in *; try intuition congruence. inv x. destruct rarg; inv e0. --- destruct s0 as (? & ? & ? & ? & ? & ? & ? & ? & ?). inv e. rewrite E1 in e0. inv e0. eapply (nth_error_app_left v [x0]) in e1. rewrite E2 in e1. inv e1... - --- help. - eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; - cbn in *; try firstorder congruence. - inv x. rewrite E1 in y. congruence. + --- solve_discr. + rewrite E1 in e. injection e as -> ->. + rewrite E2 in e0. injection e0 as ->. + apply n; sq; auto. + --- solve_discr. + rewrite E1 in y. congruence. ** left. constructor. apply whnf_mkApps. constructor. assumption. -- left. constructor. eapply whnf_fixapp. rewrite E1. eauto. + * destruct v as [ | ? v]... + right. intros [w]. depelim w. depelim w. all:help. clear IHt. + eapply whne_mkApps_inv in w as []... + -- depelim w. help. + -- destruct s0 as [? [? [? [? [? [? ?]]]]]]. congruence. + right. intros [w]. eapply n. constructor. now eapply whnf_mkApps_inv. - destruct (IHt Γ) as [_ []]. + left. destruct s as [w]. constructor. now eapply whne_mkApps. @@ -392,7 +403,10 @@ Proof with eauto using sq with pcuic. destruct s as (? & ? & ? & ? & ? & ? & ? & ? & ?). inv e. rewrite E1 in e0. inv e0. eapply (nth_error_app_left v [x0]) in e1. - rewrite E2 in e1. inv e1... + rewrite E2 in e1. inv e1... solve_discr. + rewrite E1 in e. injection e as -> ->. + rewrite E2 in e0. injection e0 as ->. + apply n0; sq; auto. --- left. constructor. apply whne_mkApps. constructor. assumption. ++ right. intros [[ | (mfix' & idx' & narg' & body' & a' & [=] & ? & ? & ?) ] % whne_mkApps_inv]; subst; cbn... congruence. @@ -914,6 +928,7 @@ Proof. destruct s as [->|(?&?)]; [easy|]. now inv e. - eapply red1_mkApps_tCoFix_inv in r as [[(?&->&?)|(?&->&?)]|(?&->&?)]; eauto. + - depelim r. solve_discr. Qed. Lemma whnf_pres Σ Γ t t' : @@ -977,7 +992,8 @@ Inductive whnf_red Σ Γ : term -> term -> Type := red Σ Γ (dtype d) (dtype d') × red Σ (Γ,,, fix_context mfix) (dbody d) (dbody d')) mfix mfix' -> - whnf_red Σ Γ (tCoFix mfix idx) (tCoFix mfix' idx). + whnf_red Σ Γ (tCoFix mfix idx) (tCoFix mfix' idx) +| whnf_red_tPrim i : whnf_red Σ Γ (tPrim i) (tPrim i). Derive Signature for whnf_red. @@ -1263,6 +1279,7 @@ Proof. cbn. intros ? ? (?&[= -> -> ->]). auto. + - depelim r; solve_discr. Qed. Lemma whnf_red_inv Σ Γ t t' : @@ -1352,6 +1369,7 @@ Proof. - apply eq_term_upto_univ_napp_mkApps_l_inv in eq as (?&?&(?&?)&->). depelim e. apply whnf_cofixapp. + - depelim eq; auto. Qed. Lemma whnf_eq_term {cf:checker_flags} f Σ φ Γ t t' : diff --git a/pcuic/theories/PCUICParallelReduction.v b/pcuic/theories/PCUICParallelReduction.v index b77089b53..4f74ca916 100644 --- a/pcuic/theories/PCUICParallelReduction.v +++ b/pcuic/theories/PCUICParallelReduction.v @@ -96,6 +96,7 @@ Lemma term_forall_ctx_list_ind : (forall Γ (m : mfixpoint term) (n : nat), All_local_env (on_local_decl (fun Γ' t => P (Γ ,,, Γ') t)) (fix_context m) -> tFixProp (P Γ) (P (Γ ,,, fix_context m)) m -> P Γ (tCoFix m n)) -> + (forall Γ p, P Γ (tPrim p)) -> forall Γ (t : term), P Γ t. Proof. intros. @@ -148,6 +149,7 @@ Proof. eapply X13; try (apply aux; red; simpl; lia). apply auxl'. simpl. lia. red. apply All_pair. split; apply auxl; simpl; auto. + Defined. (** All2 lemmas *) @@ -385,7 +387,8 @@ Section ParallelReduction. | tVar _ | tSort _ | tInd _ _ - | tConstruct _ _ _ => true + | tConstruct _ _ _ + | tPrim _ => true | _ => false end. diff --git a/pcuic/theories/PCUICParallelReductionConfluence.v b/pcuic/theories/PCUICParallelReductionConfluence.v index 2de5cad52..65637381e 100644 --- a/pcuic/theories/PCUICParallelReductionConfluence.v +++ b/pcuic/theories/PCUICParallelReductionConfluence.v @@ -244,6 +244,7 @@ Lemma term_ind_size_app : tFixProp P P m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), tFixProp (P) P m -> P (tCoFix m n)) -> + (forall p, P (tPrim p)) -> forall (t : term), P t. Proof. intros. @@ -336,32 +337,13 @@ Ltac finish_discr := | [ H : pred_atom (mkApps _ _) |- _ ] => apply pred_atom_mkApps in H; intuition subst end. -Definition application_atom t := - match t with - | tVar _ - | tSort _ - | tInd _ _ - | tConstruct _ _ _ - | tLambda _ _ _ => true - | _ => false - end. - -Lemma application_atom_mkApps {t l} : application_atom (mkApps t l) -> application_atom t /\ l = []. -Proof. - induction l in t |- *; simpl; auto. - intros. destruct (IHl _ H). discriminate. -Qed. - Ltac solve_discr ::= (try (progress (prepare_discr; finish_discr; cbn [mkApps] in * ))); (try (match goal with - | [ H : is_true (application_atom _) |- _ ] => discriminate | [ H : is_true (atom _) |- _ ] => discriminate | [ H : is_true (atom (mkApps _ _)) |- _ ] => destruct (atom_mkApps H); subst; try discriminate | [ H : is_true (pred_atom _) |- _ ] => discriminate | [ H : is_true (pred_atom (mkApps _ _)) |- _ ] => destruct (pred_atom_mkApps H); subst; try discriminate - | [ H : is_true (application_atom (mkApps _ _)) |- _ ] => - destruct (application_atom_mkApps H); subst; try discriminate end)). Lemma is_constructor_app_ge n l l' : is_constructor n l -> is_constructor n (l ++ l'). @@ -1521,7 +1503,7 @@ Section Rho. rename r (rho Γ t) = rho Δ (rename r t). Proof. revert t Γ Δ r. - refine (term_ind_size_app _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); + refine (term_ind_size_app _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); intros; try subst Γ; try rename Γ0 into Γ(* ; rename_all_hyps *). all:auto 2. diff --git a/pcuic/theories/PCUICPretty.v b/pcuic/theories/PCUICPretty.v index eb834ed63..408b3f8ab 100644 --- a/pcuic/theories/PCUICPretty.v +++ b/pcuic/theories/PCUICPretty.v @@ -190,6 +190,7 @@ Section print_term. | tCoFix l n => parens top ("let cofix " ^ print_defs print_term Γ l ^ nl ^ " in " ^ List.nth_default (string_of_nat n) (map (string_of_aname ∘ dname) l) n) + | tPrim i => parens top (string_of_prim (print_term Γ true false) i) end. End print_term. diff --git a/pcuic/theories/PCUICPrimitive.v b/pcuic/theories/PCUICPrimitive.v new file mode 100644 index 000000000..92d60ada7 --- /dev/null +++ b/pcuic/theories/PCUICPrimitive.v @@ -0,0 +1,105 @@ +(* Distributed under the terms of the MIT license. *) +From MetaCoq.Template Require Import utils Universes BasicAst Reflect + Environment EnvironmentTyping. +From Equations Require Import Equations. +From Coq Require Import ssreflect. + +Variant prim_tag := + | primInt + | primFloat. + (* | primArray. *) +Derive NoConfusion EqDec for prim_tag. + +(** We don't enforce the type of the array here*) +Record array_model (term : Type) := + { array_instance : Instance.t; + array_type : term; + array_default : term; + array_value : list term }. +Derive NoConfusion for array_model. +Instance array_model_eqdec {term} (e : EqDec term) : EqDec (array_model term). +Proof. eqdec_proof. Qed. + +Inductive prim_model (term : Type) : prim_tag -> Type := +| primIntModel (i : uint63_model) : prim_model term primInt +| primFloatModel (f : float64_model) : prim_model term primFloat. +(* | primArrayModel (a : array_model term) : prim_model term primArray. *) +Arguments primIntModel {term}. +Arguments primFloatModel {term}. +(* Arguments primArrayModel {term}. *) + +Derive Signature NoConfusion for prim_model. + +Definition prim_model_of (term : Type) (p : prim_tag) : Type := + match p with + | primInt => uint63_model + | primFloat => float64_model + (* | primArray => array_model term *) + end. + +Definition prim_val term := ∑ t : prim_tag, prim_model term t. +Definition prim_val_tag {term} (s : prim_val term) := s.π1. +Definition prim_val_model {term} (s : prim_val term) : prim_model term (prim_val_tag s) := s.π2. + +Definition prim_model_val {term} (p : prim_val term) : prim_model_of term (prim_val_tag p) := + match prim_val_model p in prim_model _ t return prim_model_of term t with + | primIntModel i => i + | primFloatModel f => f + (* | primArrayModel a => a *) + end. + +Lemma exist_irrel_eq {A} (P : A -> bool) (x y : sig P) : proj1_sig x = proj1_sig y -> x = y. +Proof. + destruct x as [x p], y as [y q]; simpl; intros ->. + now destruct (uip p q). +Qed. + +Instance reflect_eq_Z : ReflectEq Z := EqDec_ReflectEq _. + +Local Obligation Tactic := idtac. +#[program] +Instance reflect_eq_uint63 : ReflectEq uint63_model := + { eqb x y := eqb (proj1_sig x) (proj1_sig y) }. +Next Obligation. + cbn -[eqb]. + intros x y. + elim: eqb_spec. constructor. + now apply exist_irrel_eq. + intros neq; constructor => H'; apply neq; now subst x. +Qed. + +Instance reflect_eq_spec_float : ReflectEq SpecFloat.spec_float := EqDec_ReflectEq _. + +#[program] +Instance reflect_eq_float64 : ReflectEq float64_model := + { eqb x y := eqb (proj1_sig x) (proj1_sig y) }. +Next Obligation. + cbn -[eqb]. + intros x y. + elim: eqb_spec. constructor. + now apply exist_irrel_eq. + intros neq; constructor => H'; apply neq; now subst x. +Qed. + +(** Propositional UIP is needed below *) +Set Equations With UIP. + +Instance prim_model_eqdec {term} (*e : EqDec term*) : forall p : prim_tag, EqDec (prim_model term p). +Proof. eqdec_proof. Qed. + +Instance prim_tag_model_eqdec term : EqDec (prim_val term). +Proof. eqdec_proof. Defined. + +Instance prim_val_reflect_eq term : ReflectEq (prim_val term) := EqDec_ReflectEq _. + +(** Printing *) + +Definition string_of_float64_model (f : float64_model) := + "<>". + +Definition string_of_prim {term} (soft : term -> string) (p : prim_val term) : string := + match p.π2 return string with + | primIntModel f => "(int: " ^ string_of_Z (proj1_sig f) ^ ")" + | primFloatModel f => "(float: " ^ string_of_float64_model f ^ ")" + (* | primArrayModel a => "(array:" ^ ")" *) + end. diff --git a/pcuic/theories/PCUICPrincipality.v b/pcuic/theories/PCUICPrincipality.v index 214be6a56..90b81cdb5 100644 --- a/pcuic/theories/PCUICPrincipality.v +++ b/pcuic/theories/PCUICPrincipality.v @@ -178,10 +178,11 @@ Section Principality. eapply cumul_conv_ctx; eauto. constructor. apply conv_ctx_refl. constructor. assumption. eapply conv_trans; eauto. now apply conv_sym. - econstructor. eapply type_reduction; eauto. + eapply type_App'. + eapply type_reduction; eauto. eapply type_Cumul'; eauto. 2:transitivity dom; auto; now apply conv_cumul. - eapply type_reduction in t0. 3:eapply redA. 2:auto. + eapply type_reduction in t0. 2:eapply redA. eapply validity in t0; auto. eapply isType_tProd in t0 as [? ?]; eauto. eapply typing_wf_local; eauto. @@ -278,7 +279,7 @@ Section Principality. solve_discr. etransitivity; eauto. assert (consistent_instance_ext Σ (ind_universes x0) u'). - { eapply type_reduction in t1. 3:eapply redr. all:pcuic. + { eapply type_reduction in t1. 2:eapply redr. eapply validity in t1; eauto. destruct t1 as [s Hs]. eapply invert_type_mkApps_ind in Hs. intuition eauto. all:auto. eapply d. } @@ -296,7 +297,7 @@ Section Principality. (projection_context x0 x1 ind x2) []); auto. eapply (projection_subslet _ _ _ _ _ _ (ind, k, pars)); eauto. simpl. eapply type_reduction; eauto. simpl. - eapply type_reduction in t0. 3:eapply redr. eapply validity; eauto. auto. + eapply type_reduction in t0. 2:eapply redr. eapply validity; eauto. eapply (projection_subslet _ _ _ _ _ _ (ind, k, pars)); eauto. simpl. eapply validity; eauto. constructor; auto. now apply All2_rev. @@ -309,7 +310,7 @@ Section Principality. eapply (wf_projection_context _ (p:=(ind, k, pars))); pcuic. eapply (projection_subslet _ _ _ _ _ _ (ind, k, pars)); eauto. simpl. eapply type_reduction; eauto. simpl. - eapply type_reduction in t0. 3:eapply redr. all:eauto. + eapply type_reduction in t0. 2:eapply redr. eapply validity; eauto. rewrite e0 in redu'. unshelve epose proof (projection_cumulative_indices wfΣ d _ H H0 redu'). @@ -343,6 +344,7 @@ Section Principality. rewrite nthe' in nthe; noconf nthe. repeat split; eauto. eapply type_CoFix; eauto. + - now apply inversion_Prim in hA. Qed. (** A weaker version that is often convenient to use. *) @@ -521,22 +523,22 @@ Proof. econstructor; eauto. eapply cum_LetIn; pcuic. - - eapply inversion_App in X4 as (na' & A' & B' & hf & ha & cum); auto. + - eapply inversion_App in X6 as (na' & A' & B' & hf & ha & cum); auto. unfold leq_term in X1. - eapply eq_term_upto_univ_empty_impl in X5_1. - specialize (X1 onu _ _ hf X5_1). all:try typeclasses eauto. - specialize (X3 onu _ _ ha (eq_term_empty_leq_term X5_2)). - eapply leq_term_empty_leq_term in X5_1. - eapply eq_term_empty_eq_term in X5_2. + eapply eq_term_upto_univ_empty_impl in X7_1. + specialize (X3 onu _ _ hf X7_1). all:try typeclasses eauto. + specialize (X5 onu _ _ ha (eq_term_empty_leq_term X7_2)). + eapply leq_term_empty_leq_term in X7_1. + eapply eq_term_empty_eq_term in X7_2. eapply type_Cumul'. - econstructor; [eapply X1|eapply X3]. + eapply type_App'; [eapply X3|eapply X5]. eapply PCUICValidity.validity; pcuic. eapply type_App; eauto. eapply conv_cumul. eapply (subst_conv Γ [vass na A] [vass na A] []); pcuic. repeat constructor. now rewrite subst_empty. repeat constructor. now rewrite subst_empty. - eapply PCUICValidity.validity in X0; auto. - apply PCUICArities.isType_tProd in X0 as [tyA]; auto. + eapply PCUICValidity.validity in X2; auto. + apply PCUICArities.isType_tProd in X2 as [tyA]; auto. constructor; auto. - eapply inversion_Const in X1 as [decl' [wf [declc [cu cum]]]]; auto. @@ -594,7 +596,8 @@ Proof. eapply PCUICValidity.validity; eauto. eapply (type_Case _ _ (ind, npar)). eapply isdecl. all:eauto. - eapply (All2_impl X5); firstorder. + eapply (All2_impl X5); pcuicfo. + destruct b1 as [s [? ?]]. now exists s. eapply conv_cumul. eapply mkApps_conv_args; pcuic. eapply All2_app. simpl in *. @@ -637,8 +640,9 @@ Proof. econstructor; eauto. eapply PCUICValidity.validity; eauto. econstructor. 2:eapply H0. all:eauto. - eapply (All_impl X0); firstorder. - eapply (All_impl X1); firstorder. + eapply (All_impl X0); pcuicfo. + destruct X2 as [s [Hs ?]]; now exists s. + eapply (All_impl X1); pcuicfo. eapply All2_nth_error in a; eauto. destruct a as [[[eqty _] _] _]. constructor. eapply eq_term_empty_leq_term in eqty. @@ -649,12 +653,13 @@ Proof. econstructor; eauto. eapply PCUICValidity.validity; eauto. eapply type_CoFix. 2:eapply H0. all:eauto. - eapply (All_impl X0); firstorder. - eapply (All_impl X1); firstorder. + eapply (All_impl X0); pcuicfo. + destruct X2 as [s [? ?]]; now exists s. + eapply (All_impl X1); pcuicfo. eapply All2_nth_error in a; eauto. destruct a as [[[eqty _] _] _]. constructor. apply eq_term_empty_leq_term in eqty. - now eapply leq_term_empty_leq_term. Show Existentials. + now eapply leq_term_empty_leq_term. Qed. Lemma typing_eq_term {cf:checker_flags} (Σ : global_env_ext) Γ t t' T T' : diff --git a/pcuic/theories/PCUICReduction.v b/pcuic/theories/PCUICReduction.v index 7fa6b58f5..a63e073c1 100644 --- a/pcuic/theories/PCUICReduction.v +++ b/pcuic/theories/PCUICReduction.v @@ -1,13 +1,13 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Template Require Import config utils. -From MetaCoq.PCUIC Require Import PCUICRelations PCUICAst PCUICAstUtils +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICLiftSubst PCUICEquality PCUICUnivSubst PCUICInduction. Require Import ssreflect. Require Import Equations.Prop.DepElim. +From Equations.Type Require Import Relation Relation_Properties. From Equations Require Import Equations. - Set Default Goal Selector "!". Definition tDummy := tVar String.EmptyString. @@ -15,7 +15,6 @@ Definition tDummy := tVar String.EmptyString. Definition iota_red npar c args brs := (mkApps (snd (List.nth c brs (0, tDummy))) (List.skipn npar args)). - (** ** Reduction *) (** *** Helper functions for reduction *) @@ -320,7 +319,6 @@ Hint Constructors red1 : pcuic. Definition red Σ Γ := clos_refl_trans (red1 Σ Γ). - Lemma refl_red Σ Γ t : red Σ Γ t t. Proof. reflexivity. @@ -713,7 +711,7 @@ Section ReductionCongruence. red Σ Γ M M' -> red Σ (Γ ,, vass na M') N N' -> red Σ Γ (tLambda na M N) (tLambda na M' N'). Proof. - intros. eapply (transitivity (y := tLambda na M' N)). + intros. transitivity (tLambda na M' N). - now apply (red_ctx (tCtxLambda_l _ tCtxHole _)). - now eapply (red_ctx (tCtxLambda_r _ _ tCtxHole)). Qed. @@ -730,7 +728,7 @@ Section ReductionCongruence. red Σ Γ N0 N1 -> red Σ Γ (tApp M0 N0) (tApp M1 N1). Proof. - intros; eapply (transitivity (y := tApp M1 N0)). + intros; transitivity (tApp M1 N0). - now apply (red_ctx (tCtxApp_l tCtxHole _)). - now eapply (red_ctx (tCtxApp_r _ tCtxHole)). Qed. @@ -786,9 +784,9 @@ Section ReductionCongruence. red Σ Γ d0 d1 -> red Σ Γ t0 t1 -> red Σ (Γ ,, vdef na d1 t1) b0 b1 -> red Σ Γ (tLetIn na d0 t0 b0) (tLetIn na d1 t1 b1). Proof. - intros; eapply (transitivity (y := tLetIn na d1 t0 b0)). + intros; transitivity (tLetIn na d1 t0 b0). - now apply (red_ctx (tCtxLetIn_l _ tCtxHole _ _)). - - eapply (transitivity (y := tLetIn na d1 t1 b0)). + - transitivity (tLetIn na d1 t1 b0). + now eapply (red_ctx (tCtxLetIn_b _ _ tCtxHole _)). + now eapply (red_ctx (tCtxLetIn_r _ _ _ tCtxHole)). Qed. diff --git a/pcuic/theories/PCUICReflect.v b/pcuic/theories/PCUICReflect.v index 39e5b23de..b430b53f7 100644 --- a/pcuic/theories/PCUICReflect.v +++ b/pcuic/theories/PCUICReflect.v @@ -34,6 +34,7 @@ Local Ltac term_dec_tac term_dec := | i : string, i' : kername |- _ => fcase (string_dec i i') | n : name, n' : name |- _ => fcase (eq_dec n n') | n : aname, n' : aname |- _ => fcase (eq_dec n n') + | i : prim_val, j : prim_val |- _ => fcase (eq_dec i j) | i : inductive, i' : inductive |- _ => fcase (eq_dec i i') | x : inductive * nat, y : inductive * nat |- _ => fcase (eq_dec x y) diff --git a/pcuic/theories/PCUICRelations.v b/pcuic/theories/PCUICRelations.v deleted file mode 100644 index 029b0d8af..000000000 --- a/pcuic/theories/PCUICRelations.v +++ /dev/null @@ -1,165 +0,0 @@ -(* Distributed under the terms of the MIT license. *) -Require Import ssreflect. -Require Export CRelationClasses. -Require Export Equations.Type.Relation Equations.Type.Relation_Properties. - - -Section Flip. - Local Set Universe Polymorphism. - Context {A : Type} (R : crelation A). - - Lemma flip_Reflexive : Reflexive R -> Reflexive (flip R). - Proof. - intros HR x. unfold flip. apply reflexivity. - Qed. - - Lemma flip_Symmetric : Symmetric R -> Symmetric (flip R). - Proof. - intros HR x y. unfold flip. apply symmetry. - Qed. - - Lemma flip_Transitive : Transitive R -> Transitive (flip R). - Proof. - intros HR x y z xy yz. - unfold flip in *. eapply HR; eassumption. - Qed. - -End Flip. - - - - -Definition commutes {A} (R S : relation A) := - forall x y z, R x y -> S x z -> { w & S y w * R z w}%type. - - -Lemma clos_t_rt {A} {R : A -> A -> Type} x y : trans_clos R x y -> clos_refl_trans R x y. -Proof. - induction 1; try solve [econstructor; eauto]. -Qed. - - -Arguments rt_step {A} {R} {x y}. -Polymorphic Hint Resolve rt_refl rt_step : core. - - -Definition clos_rt_monotone {A} (R S : relation A) : - inclusion R S -> inclusion (clos_refl_trans R) (clos_refl_trans S). -Proof. - move => incls x y. - induction 1; solve [econstructor; eauto]. -Qed. - -Lemma relation_equivalence_inclusion {A} (R S : relation A) : - inclusion R S -> inclusion S R -> relation_equivalence R S. -Proof. firstorder. Qed. - -Lemma clos_rt_disjunction_left {A} (R S : relation A) : - inclusion (clos_refl_trans R) - (clos_refl_trans (relation_disjunction R S)). -Proof. - apply clos_rt_monotone. - intros x y H; left; exact H. -Qed. - -Lemma clos_rt_disjunction_right {A} (R S : relation A) : - inclusion (clos_refl_trans S) - (clos_refl_trans (relation_disjunction R S)). -Proof. - apply clos_rt_monotone. - intros x y H; right; exact H. -Qed. - -Global Instance clos_rt_trans A R : Transitive (@clos_refl_trans A R). -Proof. - intros x y z H H'. econstructor 3; eauto. -Qed. - -Global Instance clos_rt_refl A R : Reflexive (@clos_refl_trans A R). -Proof. intros x. constructor 2. Qed. - -Lemma clos_refl_trans_prod_l {A B} (R : relation A) (S : relation (A * B)) : - (forall x y b, R x y -> S (x, b) (y, b)) -> - forall (x y : A) b, - clos_refl_trans R x y -> - clos_refl_trans S (x, b) (y, b). -Proof. - intros. induction X0; try solve [econstructor; eauto]. -Qed. - -Lemma clos_refl_trans_prod_r {A B} (R : relation B) (S : relation (A * B)) a : - (forall x y, R x y -> S (a, x) (a, y)) -> - forall (x y : B), - clos_refl_trans R x y -> - clos_refl_trans S (a, x) (a, y). -Proof. - intros. induction X0; try solve [econstructor; eauto]. -Qed. - -Lemma clos_rt_t_incl {A} {R : relation A} `{Reflexive A R} : - inclusion (clos_refl_trans R) (trans_clos R). -Proof. - intros x y. induction 1; try solve [econstructor; eauto]. -Qed. - -Lemma clos_t_rt_incl {A} {R : relation A} `{Reflexive A R} : - inclusion (trans_clos R) (clos_refl_trans R). -Proof. - intros x y. induction 1; try solve [econstructor; eauto]. -Qed. - -Lemma clos_t_rt_equiv {A} {R} `{Reflexive A R} : - relation_equivalence (trans_clos R) (clos_refl_trans R). -Proof. - apply relation_equivalence_inclusion. - apply clos_t_rt_incl. - apply clos_rt_t_incl. -Qed. - -Global Instance relation_disjunction_refl_l {A} {R S : relation A} : - Reflexive R -> Reflexive (relation_disjunction R S). -Proof. - intros HR x. left; auto. -Qed. - -Global Instance relation_disjunction_refl_r {A} {R S : relation A} : - Reflexive S -> Reflexive (relation_disjunction R S). -Proof. - intros HR x. right; auto. -Qed. - -Global Instance clos_rt_trans_Symmetric A R : - Symmetric R -> Symmetric (@clos_refl_trans A R). -Proof. - intros X x y H. induction H; eauto. - eapply clos_rt_trans; eassumption. -Qed. - -Definition clos_sym {A} (R : relation A) := relation_disjunction R (flip R). - -Global Instance clos_sym_Symmetric A R : - Symmetric (@clos_sym A R). -Proof. - intros x y []; [right|left]; assumption. -Qed. - -Global Instance clos_refl_sym_trans_Symmetric A R : - Symmetric (@clos_refl_sym_trans A R) - := rst_sym R. - -Global Instance clos_refl_sym_trans_Reflexive A R : - Reflexive (@clos_refl_sym_trans A R) - := rst_refl R. - -Global Instance clos_refl_sym_trans_Transitive A R : - Transitive (@clos_refl_sym_trans A R) - := rst_trans R. - -Global Instance relation_disjunction_Symmetric {A} {R S : relation A} : - Symmetric R -> Symmetric S -> Symmetric (relation_disjunction R S). -Proof. - intros HR HS x y [X|X]; [left|right]; eauto. -Qed. - -Ltac rst_induction h := - induction h; [constructor|reflexivity|etransitivity]. diff --git a/pcuic/theories/PCUICSR.v b/pcuic/theories/PCUICSR.v index 9efbcf3f0..07fc68e7c 100644 --- a/pcuic/theories/PCUICSR.v +++ b/pcuic/theories/PCUICSR.v @@ -12,10 +12,10 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICUtils Require Import ssreflect. From Equations Require Import Equations. Require Import Equations.Prop.DepElim. - - Local Set SimplIsCbn. +Implicit Types (cf : checker_flags) (Σ : global_env_ext). + Derive Signature for OnOne2_local_env. Ltac rename_hyp h ht ::= my_rename_hyp h ht. @@ -29,21 +29,20 @@ Hint Rewrite projs_length : len. (** The subject reduction property of the system: *) -Definition SR_red1 {cf:checker_flags} (Σ : global_env_ext) Γ t T := +Definition SR_red1 {cf} Σ Γ t T := forall u (Hu : red1 Σ Γ t u), Σ ;;; Γ |- u : T. (* Preservation of wf_*fixpoint *) -Lemma wf_fixpoint_red1_type {cf:checker_flags} (Σ : global_env_ext) Γ mfix mfix1 : - wf Σ.1 -> - wf_fixpoint Σ.1 mfix -> +Lemma wf_fixpoint_red1_type {cf Σ} {wfΣ : wf Σ} Γ mfix mfix1 : + wf_fixpoint Σ mfix -> OnOne2 (fun x y : def term => red1 Σ Γ (dtype x) (dtype y) × (dname x, dbody x, rarg x) = (dname y, dbody y, rarg y)) mfix mfix1 -> - wf_fixpoint Σ.1 mfix1. + wf_fixpoint Σ mfix1. Proof. - intros wfΣ wffix o. + intros wffix o. move: wffix; unfold wf_fixpoint. enough (forall inds, map_option_out (map check_one_fix mfix) = Some inds -> map_option_out (map check_one_fix mfix1) = Some inds) => //. @@ -75,16 +74,15 @@ Proof. discriminate. Qed. -Lemma wf_fixpoint_red1_body {cf:checker_flags} (Σ : global_env_ext) Γ mfix mfix1 : - wf Σ.1 -> - wf_fixpoint Σ.1 mfix -> +Lemma wf_fixpoint_red1_body {cf Σ} {wfΣ : wf Σ} Γ mfix mfix1 : + wf_fixpoint Σ mfix -> OnOne2 (fun x y : def term => red1 Σ Γ (dbody x) (dbody y) × (dname x, dtype x, rarg x) = (dname y, dtype y, rarg y)) mfix mfix1 -> - wf_fixpoint Σ.1 mfix1. + wf_fixpoint Σ mfix1. Proof. - intros wfΣ wffix o. + intros wffix o. move: wffix; unfold wf_fixpoint. enough (map check_one_fix mfix = map check_one_fix mfix1) as -> => //. induction o. @@ -275,10 +273,9 @@ Proof. specialize (validity _ wf _ _ _ appty) as [_ vT']. eapply type_tFix_inv in appty as [T [arg [fn' [[[Hnth wffix] Hty]]]]]; auto. rewrite e in Hnth. noconf Hnth. - eapply type_App. + eapply type_App; eauto. eapply type_mkApps. eapply type_Cumul'; eauto. eapply spty. - eauto. - + - (* Congruence *) eapply type_Cumul'; [eapply type_App| |]; eauto with wf. eapply validity. eauto. eauto. @@ -354,8 +351,8 @@ Proof. assert (wfps : wf_universe Σ ps). { eapply validity in typep; auto. eapply PCUICWfUniverses.isType_wf_universes in typep. rewrite PCUICWfUniverses.wf_universes_it_mkProd_or_LetIn in typep. - move/andP: typep => /= [_ /andP[_ typep]]. - now apply (PCUICWfUniverses.reflect_bP (PCUICWfUniverses.wf_universe_reflect _ _)) in typep. auto. } + move/andb_and: typep => /= [_ /andb_and[_ typep]]. + now apply (ssrbool.elimT PCUICWfUniverses.wf_universe_reflect) in typep. auto. } eapply wf_arity_spine_typing_spine => //. split. { (* Predicate instantiation is well typed *) @@ -412,7 +409,7 @@ Proof. eapply closed_wf_local in wfcl; auto. rewrite !subst_instance_context_app in wfcl. rewrite closedn_ctx_app in wfcl. - move/andP: wfcl => []. autorewrite with len. now auto. } + move/andb_and: wfcl => []. autorewrite with len. now auto. } eapply arity_spine_it_mkProd_or_LetIn; eauto. { unfold to_extended_list, to_extended_list_k. rewrite /argctxu in X0. simpl. rewrite -H in X0. eapply X0. } @@ -467,7 +464,7 @@ Proof. epose proof (on_minductive_wf_params_indices_inst _ _ u _ _ _ (proj1 decli) oib cu). rewrite subst_instance_context_app in X1. eapply closed_wf_local in X1; eauto. rewrite closedn_ctx_app in X1. autorewrite with len in X1. - now move/andP: X1 => []. + now move/andb_and: X1 => []. } simpl. eapply conv_cumul; apply mkApps_conv_args; auto. @@ -700,8 +697,8 @@ Proof. assert (wfps : wf_universe Σ ps). { eapply validity in typep; auto. eapply PCUICWfUniverses.isType_wf_universes in typep. rewrite PCUICWfUniverses.wf_universes_it_mkProd_or_LetIn in typep. - move/andP: typep => /= [_ /andP[_ typep]]. - now apply (PCUICWfUniverses.reflect_bP (PCUICWfUniverses.wf_universe_reflect _ _)) in typep. auto. } + move/andb_and: typep => /= [_ /andb_and[_ typep]]. + now apply (ssrbool.elimT PCUICWfUniverses.wf_universe_reflect) in typep. auto. } pose proof (context_subst_fun cparsubst spars). subst parsubst'. clear cparsubst. eapply type_mkApps. eauto. eapply wf_arity_spine_typing_spine; eauto. @@ -718,21 +715,20 @@ Proof. - (* Case congruence on discriminee *) eapply type_Cumul. eapply type_Case; eauto. - * solve_all. - firstorder auto. + * solve_all. destruct b0 as [s Hs]; exists s; pcuic. * pose proof typec as typec'. eapply (env_prop_typing _ _ validity) in typec' as wat; auto. unshelve eapply isType_mkApps_Ind in wat as [parsubst [argsubst wat]]; eauto. set (oib := on_declared_inductive wf isdecl) in *. clearbody oib. destruct oib as [onind oib]. - destruct wat as [[spars sargs] cu]. + destruct wat as [[spars sargs] cu]. unshelve eapply (build_case_predicate_type_spec (Σ.1, _)) in heq_build_case_predicate_type as [parsubst' [cparsubst Hpty]]; eauto. rewrite {}Hpty in typep. assert (wfps : wf_universe Σ ps). { eapply validity in typep; auto. eapply PCUICWfUniverses.isType_wf_universes in typep. rewrite PCUICWfUniverses.wf_universes_it_mkProd_or_LetIn in typep. - move/andP: typep => /= [_ /andP[_ typep]]. - now apply (PCUICWfUniverses.reflect_bP (PCUICWfUniverses.wf_universe_reflect _ _)) in typep. auto. } + move/andb_and: typep => /= [_ /andb_and[_ typep]]. + now apply (ssrbool.elimT PCUICWfUniverses.wf_universe_reflect) in typep. auto. } subst npar. pose proof (context_subst_fun cparsubst spars). subst parsubst'. clear cparsubst. eapply type_mkApps. eauto. @@ -1102,7 +1098,7 @@ Proof. clear projsubsl. eapply closed_wf_local in wfdecl. rewrite closedn_ctx_app in wfdecl. - move/andP: wfdecl => [_ wfdecl]. + move/andb_and: wfdecl => [_ wfdecl]. autorewrite with len in wfdecl. simpl in wfdecl. eapply closedn_ctx_decl in wfdecl; eauto. @@ -1161,12 +1157,11 @@ Proof. constructor; eauto. * eapply (OnOne2_All_mix_left X0) in o. apply (OnOne2_All_All o X1). - + intros x [[Hb Hlam] IH]. - split; auto. + + intros x [Hb IH]. eapply context_conversion; eauto. now rewrite -fixl. - + move=> [na ty b rarg] [na' ty' b' rarg'] /= [[red eq] [s [Hs IH]]] [[Hb Hlam] IH']. - noconf eq. split; auto. + + move=> [na ty b rarg] [na' ty' b' rarg'] /= [[red eq] [s [Hs IH]]] [Hb IH']. + noconf eq. eapply context_conversion; eauto. rewrite -fixl. eapply type_Cumul'. eapply Hb. @@ -1209,15 +1204,13 @@ Proof. apply fix_red_body; eauto. * eapply (OnOne2_All_mix_left X0) in o. apply (OnOne2_All_All o X1). - + intros x [[Hb Hlam] IH]. - split; auto. + + intros x [Hb IH]. eapply context_conversion; eauto. now rewrite -fixl. rewrite convctx. apply conv_ctx_refl. - + move=> [na ty b rarg] [na' ty' b' rarg'] /= [[red eq] [s [Hs IH]]] [[Hb Hlam] IH']. + + move=> [na ty b rarg] [na' ty' b' rarg'] /= [[red eq] [s [Hs IH]]] [Hb IH']. noconf eq. - rewrite -convctx. split; auto. - now eapply isLambda_red1. + now rewrite -convctx. * eapply wf_fixpoint_red1_body; eauto. * eapply All_nth_error in X2; eauto. * apply conv_cumul, conv_sym, red_conv. destruct disj as [<-|[_ eq]]. @@ -1314,8 +1307,6 @@ Proof. eapply type_Cumul; eauto. Qed. -Print Assumptions sr_red1. - Definition sr_stmt {cf:checker_flags} (Σ : global_env_ext) Γ t T := forall u, red Σ Γ t u -> Σ ;;; Γ |- u : T. @@ -1327,8 +1318,8 @@ Proof. eapply sr_red1 in Hty; eauto with wf. Qed. -Lemma subject_reduction1 {cf:checker_flags} {Σ Γ t u T} - : wf Σ.1 -> Σ ;;; Γ |- t : T -> red1 Σ.1 Γ t u -> Σ ;;; Γ |- u : T. +Lemma subject_reduction1 {cf Σ} {wfΣ : wf Σ} {Γ t u T} + : Σ ;;; Γ |- t : T -> red1 Σ Γ t u -> Σ ;;; Γ |- u : T. Proof. intros. eapply subject_reduction; try eassumption. now apply red1_red. @@ -1337,24 +1328,11 @@ Defined. Section SRContext. Context {cf:checker_flags}. - (* todo: rename wf_local_app *) - Definition wf_local_app' {Σ Γ1 Γ2} : - wf_local Σ Γ1 -> wf_local_rel Σ Γ1 Γ2 - -> wf_local Σ (Γ1 ,,, Γ2). - Proof. - intros H1 H2. apply wf_local_local_rel. - apply wf_local_rel_local in H1. - apply wf_local_rel_app_inv; tas. - now rewrite app_context_nil_l. - Qed. - - Definition cumul_red_l' `{checker_flags} : - forall Σ Γ t u, - wf Σ.1 -> - red (fst Σ) Γ t u -> + Definition cumul_red_l' {Σ} {wfΣ : wf Σ} Γ t u : + red Σ Γ t u -> Σ ;;; Γ |- t <= u. Proof. - intros Σ Γ t u hΣ h. + intros h. induction h using red_rect'. - eapply cumul_refl'. - eapply PCUICConversion.cumul_trans ; try eassumption. @@ -1366,8 +1344,7 @@ Section SRContext. Hint Constructors OnOne2_local_env : aa. Hint Unfold red1_ctx : aa. - - Lemma red1_ctx_app Σ Γ Γ' Δ : + Lemma red1_ctx_app (Σ : global_env) Γ Γ' Δ : red1_ctx Σ Γ Γ' -> red1_ctx Σ (Γ ,,, Δ) (Γ' ,,, Δ). Proof. @@ -1375,7 +1352,7 @@ Section SRContext. intro H. simpl. constructor. now apply IHΔ. Qed. - Lemma red1_red_ctx Σ Γ Γ' : + Lemma red1_red_ctx (Σ : global_env) Γ Γ' : red1_ctx Σ Γ Γ' -> red_ctx Σ Γ Γ'. Proof. @@ -1388,7 +1365,7 @@ Section SRContext. apply refl_red. Qed. - Lemma nth_error_red1_ctx Σ Γ Γ' n decl : + Lemma nth_error_red1_ctx (Σ : global_env) Γ Γ' n decl : wf Σ -> nth_error Γ n = Some decl -> red1_ctx Σ Γ Γ' -> @@ -1422,13 +1399,12 @@ Section SRContext. eapply (weakening_red_0 wfΣ _ [_]); tas; cbnr. Qed. - Lemma wf_local_isType_nth Σ Γ n decl : - wf Σ.1 -> + Lemma wf_local_isType_nth {Σ} {wfΣ : wf Σ} Γ n decl : wf_local Σ Γ -> nth_error Γ n = Some decl -> ∑ s, Σ ;;; Γ |- lift0 (S n) (decl_type decl) : tSort s. Proof. - induction n in Γ, decl |- *; intros hΣ hΓ e; destruct Γ; + induction n in Γ, decl |- *; intros hΓ e; destruct Γ; cbn; inversion e; inversion hΓ; subst. all: try (destruct X0 as [s Ht]; exists s; eapply (weakening _ _ [_] _ (tSort s)); tas). @@ -1443,9 +1419,8 @@ Section SRContext. Ltac invs H := inversion H; subst. Ltac invc H := inversion H; subst; clear H. - Lemma subject_reduction_ctx Σ Γ Γ' t T : - wf Σ.1 -> - red1_ctx Σ.1 Γ Γ' -> + Lemma subject_reduction_ctx {Σ} {wfΣ : wf Σ} Γ Γ' t T : + red1_ctx Σ Γ Γ' -> Σ ;;; Γ |- t : T -> Σ ;;; Γ' |- t : T. Proof. assert(OnOne2_local_env @@ -1463,7 +1438,7 @@ Section SRContext. constructor. pcuic. constructor; pcuics. now apply red_conv, red1_red. - destruct d as [na [b|] ?]; constructor; auto; constructor; auto. all:pcuic. } - intros wfΣ r H. + intros r H. specialize (X r). assert(wf_local Σ Γ'). apply typing_wf_local in H. @@ -1495,11 +1470,10 @@ Section SRContext. - eapply context_conversion; eauto. Qed. - Lemma wf_local_red1 {Σ Γ Γ'} : - wf Σ.1 -> - red1_ctx Σ.1 Γ Γ' -> wf_local Σ Γ -> wf_local Σ Γ'. + Lemma wf_local_red1 {Σ} {wfΣ : wf Σ} {Γ Γ'} : + red1_ctx Σ Γ Γ' -> wf_local Σ Γ -> wf_local Σ Γ'. Proof. - intro hΣ. induction 1; cbn in *. + induction 1; cbn in *. - intro e. inversion e; subst; cbn in *. constructor; tas. destruct X0 as [s Ht]. exists s. eapply subject_reduction1; tea. @@ -1526,17 +1500,16 @@ Section SRContext. all:constructor; cbnr; eauto. Qed. - Lemma wf_local_red {Σ Γ Γ'} : - wf Σ.1 -> - red_ctx Σ.1 Γ Γ' -> wf_local Σ Γ -> wf_local Σ Γ'. + Lemma wf_local_red {Σ} {wfΣ : wf Σ} {Γ Γ'} : + red_ctx Σ Γ Γ' -> wf_local Σ Γ -> wf_local Σ Γ'. Proof. - intros hΣ h. apply red_ctx_clos_rt_red1_ctx in h. + intros h. apply red_ctx_clos_rt_red1_ctx in h. induction h; eauto using wf_local_red1. apply eq_context_upto_names_upto_names in e. eauto using wf_local_alpha. Qed. - Lemma wf_local_subst1 Σ (wfΣ : wf Σ.1) Γ na b t Γ' : + Lemma wf_local_subst1 {Σ} {wfΣ : wf Σ} Γ na b t Γ' : wf_local Σ (Γ ,,, [],, vdef na b t ,,, Γ') -> wf_local Σ (Γ ,,, subst_context [b] 0 Γ'). Proof. @@ -1549,7 +1522,7 @@ Section SRContext. assert (subslet Σ Γ [b] [vdef na b t]). { pose proof (cons_let_def Σ Γ [] [] na b t) as XX. rewrite !subst_empty in XX. apply XX. constructor. - apply wf_local_app in X. inversion X; subst; cbn in *; assumption. + apply wf_local_app_l in X. inversion X; subst; cbn in *; assumption. } constructor; cbn; auto. 1: exists s. 1: unfold PCUICTerm.tSort. @@ -1561,7 +1534,7 @@ Section SRContext. assert (subslet Σ Γ [b] [vdef na b t]). { pose proof (cons_let_def Σ Γ [] [] na b t) as XX. rewrite !subst_empty in XX. apply XX. constructor. - apply wf_local_app in X. inversion X; subst; cbn in *; assumption. } + apply wf_local_app_l in X. inversion X; subst; cbn in *; assumption. } constructor; cbn; auto. exists s. unfold PCUICTerm.tSort. change (tSort s) with (subst [b] #|Γ'| (tSort s)). @@ -1576,23 +1549,21 @@ Section SRContext. intro H; simpl; constructor; cbn; eauto; now apply IHΔ. Qed. - Lemma isType_red1 {Σ Γ A B} : - wf Σ.1 -> - red1 (fst Σ) Γ A B -> + Lemma isType_red1 {Σ : global_env_ext} {wfΣ : wf Σ} {Γ A B} : isType Σ Γ A -> + red1 Σ Γ A B -> isType Σ Γ B. Proof. - intros wf red [s Hs]. + intros [s Hs] red. exists s. eapply subject_reduction1; eauto. Qed. - Lemma isWfArity_red1 {Σ Γ A B} : - wf Σ.1 -> - red1 (fst Σ) Γ A B -> + Lemma isWfArity_red1 {Σ} {wfΣ : wf Σ} {Γ A B} : isWfArity Σ Γ A -> + red1 Σ Γ A B -> isWfArity Σ Γ B. Proof. - intros wfΣ re [isty H]. + intros [isty H] re. split. eapply isType_red1; eauto. clear isty; revert H. induction re using red1_ind_all. @@ -1629,44 +1600,38 @@ Section SRContext. eexists _, s''. cbn. rewrite destArity_app ee'. reflexivity. Qed. - Lemma isWfArity_red {Σ Γ A B} : - wf Σ.1 -> - red (fst Σ) Γ A B -> + Lemma isWfArity_red {Σ} {wfΣ : wf Σ} {Γ A B} : isWfArity Σ Γ A -> + red Σ Γ A B -> isWfArity Σ Γ B. Proof. induction 2 using red_rect'. - easy. - - intro. now eapply isWfArity_red1. + - now eapply isWfArity_red1. Qed. - Lemma isType_red {Σ Γ A B} : - wf Σ.1 -> - red (fst Σ) Γ A B -> - isType Σ Γ A -> - isType Σ Γ B. + Lemma isType_red {Σ} {wfΣ : wf Σ} {Γ T U} : + isType Σ Γ T -> red Σ Γ T U -> isType Σ Γ U. Proof. - intros ? ? [? ?]. - eexists. eapply subject_reduction; tea. + intros [s Hs] red; exists s. + eapply subject_reduction; eauto. Qed. - Lemma type_reduction {Σ Γ t A B} - : wf Σ.1 -> Σ ;;; Γ |- t : A -> red (fst Σ) Γ A B -> Σ ;;; Γ |- t : B. + Lemma type_reduction {Σ} {wfΣ : wf Σ} {Γ t A B} : + Σ ;;; Γ |- t : A -> red Σ Γ A B -> Σ ;;; Γ |- t : B. Proof. - intros HΣ' Ht Hr. + intros Ht Hr. eapply type_Cumul'. eassumption. 2: now eapply cumul_red_l'. - destruct (validity_term HΣ' Ht) as [s HA]. + destruct (validity_term wfΣ Ht) as [s HA]. exists s; eapply subject_reduction; eassumption. Defined. End SRContext. -Lemma isWAT_tLetIn {cf:checker_flags} {Σ : global_env_ext} (HΣ' : wf Σ) - {Γ} (HΓ : wf_local Σ Γ) {na t A B} +Lemma isType_tLetIn {cf} {Σ} {HΣ' : wf Σ} {Γ} {na t A B} : isType Σ Γ (tLetIn na t A B) - <~> (isType Σ Γ A × (Σ ;;; Γ |- t : A) - × isType Σ (Γ,, vdef na t A) B). + <~> (isType Σ Γ A × (Σ ;;; Γ |- t : A) × isType Σ (Γ,, vdef na t A) B). Proof. split; intro HH. - destruct HH as [s H]. @@ -1690,3 +1655,6 @@ Proof. * apply red1_red. apply red_zeta with (b':=tSort sB). Defined. + +(** Keep at the end to not disturb asynchronous proof processing *) +Print Assumptions sr_red1. diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index dc26dc084..0cc13c891 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -274,8 +274,30 @@ Section Lemmata. Arguments iswelltyped {Σ Γ t A} h. + Definition isType_welltyped {Γ T} + : isType Σ Γ T -> welltyped Σ Γ T. + Proof. + intros []. now econstructor. + Qed. + Hint Resolve isType_welltyped : pcuic. + Context (hΣ : ∥ wf Σ ∥). + Lemma validity_wf {Γ t T} : + ∥ Σ ;;; Γ |- t : T ∥ -> welltyped Σ Γ T. + Proof. + destruct hΣ as [wΣ]. intros [X]. + intros. eapply validity_term in X; try assumption. + destruct X. now exists (tSort x). + Defined. + + Lemma wat_welltyped {Γ T} : + ∥ isType Σ Γ T ∥ -> welltyped Σ Γ T. + Proof. + destruct hΣ as [wΣ]. intros [X]. + now apply isType_welltyped. + Defined. + Lemma welltyped_alpha Γ u v : welltyped Σ Γ u -> eq_term_upto_univ [] eq eq u v -> @@ -871,12 +893,6 @@ Section Lemmata. | _ => false end. - Definition isProd t := - match t with - | tProd na A B => true - | _ => false - end. - Lemma isAppProd_mkApps : forall t l, isAppProd (mkApps t l) = isAppProd t. Proof. @@ -909,7 +925,7 @@ Section Lemmata. induction l ; intros t h. - reflexivity. - cbn in h. specialize IHl with (1 := h). subst. - cbn in h. exfalso. assumption. + cbn in h. exfalso. discriminate. Qed. Lemma isAppProd_isProd : @@ -1364,3 +1380,5 @@ Section Lemmata. Qed. End Lemmata. + +Hint Resolve isType_welltyped : pcuic. diff --git a/pcuic/theories/PCUICSigmaCalculus.v b/pcuic/theories/PCUICSigmaCalculus.v index dc7b9767d..e3b3cbc59 100644 --- a/pcuic/theories/PCUICSigmaCalculus.v +++ b/pcuic/theories/PCUICSigmaCalculus.v @@ -5,14 +5,13 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICInduction PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICClosed PCUICEquality. -Require Import ssreflect. +Require Import ssreflect ssrbool. From Equations Require Import Equations. Require Import Equations.Prop.DepElim. Set Equations With UIP. (** * Type preservation for σ-calculus *) - Open Scope sigma_scope. Set Keyed Unification. @@ -359,21 +358,22 @@ Proof. try (try rewrite H; try rewrite H0 ; try rewrite H1 ; easy); try solve [f_equal; solve_all]. - apply Hs. now eapply Nat.ltb_lt. - - move/andP: clt => []. intros. f_equal; eauto. + - move/andb_and: clt => []. intros. f_equal; eauto. eapply H0; eauto. intros. eapply up_ext_closed; eauto. - - move/andP: clt => []. intros. f_equal; eauto. now eapply H0, up_ext_closed. - - move/andP: clt => [] /andP[] ?. intros. f_equal; eauto. + - move/andb_and: clt => []. intros. f_equal; eauto. now eapply H0, up_ext_closed. + - move/andb_and: clt => [] /andb_and[] ?. intros. f_equal; eauto. now eapply H1, up_ext_closed. - - move/andP: clt => [] ? ?. f_equal; eauto. - - move/andP: clt => [] /andP[] ? ? b1. + - move/andb_and: clt => [] ? ?. f_equal; eauto. + - move/andb_and: clt => [] /andb_and[] ? ? b1. red in X. solve_all. f_equal; eauto. - eapply All_map_eq. eapply (All_impl b1). firstorder auto with *. + eapply All_map_eq. eapply (All_impl b1). pcuicfo. + unfold on_snd. f_equal. now eapply a0. - f_equal; eauto. red in X. solve_all. - move/andP: b => []. eauto. intros. + move/andb_and: b => []. eauto. intros. apply map_def_eq_spec; eauto. eapply b0; eauto. now apply up_ext_closed. - f_equal; eauto. red in X. solve_all. - move/andP: b => []. eauto. intros. + move/andb_and: b => []. eauto. intros. apply map_def_eq_spec; eauto. eapply b0; eauto. now apply up_ext_closed. Qed. @@ -749,15 +749,15 @@ Proof. Qed. Lemma rename_shiftn : - forall f t, - rename (shiftn 1 f) (lift0 1 t) = lift0 1 (rename f t). + forall f k t, + rename (shiftn k f) (lift0 k t) = lift0 k (rename f t). Proof. - intros f t. + intros f k t. autorewrite with sigma. eapply inst_ext. intro i. unfold ren, lift_renaming, shiftn, subst_compose. simpl. - replace (i - 0) with i by lia. - reflexivity. + destruct (Nat.ltb_spec (k + i) k); try lia. + unfold shiftk. lia_f_equal. Qed. Lemma urenaming_vass : @@ -1071,6 +1071,15 @@ Proof. intros Σ Γ t A B h []. assumption. Qed. +Lemma meta_conv_term : + forall Σ Γ t t' A, + Σ ;;; Γ |- t : A -> + t = t' -> + Σ ;;; Γ |- t' : A. +Proof. + intros Σ Γ t A B h []. assumption. +Qed. + (* Could be more precise *) Lemma instantiate_params_subst_length : forall params pars s t s' t', @@ -1151,7 +1160,7 @@ Lemma inst_decl_closed : Proof. intros σ k d. case: d => na [body|] ty. all: rewrite /closed_decl /inst_decl /map_decl /=. - - move /andP => [cb cty]. rewrite !inst_closed //. + - move /andb_and => [cb cty]. rewrite !inst_closed //. - move => cty. rewrite !inst_closed //. Qed. @@ -1166,7 +1175,7 @@ Proof. induction ctx using rev_ind; try easy. move => n. rewrite /closedn_ctx !rev_app_distr /id /=. - move /andP => [closedx Hctx]. + move /andb_and => [closedx Hctx]. rewrite inst_decl_closed //. f_equal. now rewrite IHctx. Qed. @@ -1604,6 +1613,20 @@ Proof. + eapply red1_rename. all: try eassumption. Qed. +Lemma fix_guard_rename Σ Γ Δ mfix f : + renaming Σ Γ Δ f -> + let mfix' := map (map_def (rename f) (rename (shiftn (List.length mfix) f))) mfix in + fix_guard Σ Δ mfix -> + fix_guard Σ Γ mfix'. +Admitted. + +Lemma cofix_guard_rename Σ Γ Δ mfix f : + renaming Σ Γ Δ f -> + let mfix' := map (map_def (rename f) (rename (shiftn (List.length mfix) f))) mfix in + cofix_guard Σ Δ mfix -> + cofix_guard Σ Γ mfix'. +Admitted. + Lemma typing_rename_prop : env_prop (fun Σ Γ t A => forall Δ f, @@ -1650,9 +1673,10 @@ Proof. * destruct hf. assumption. * simpl. eexists. eapply ihB. assumption. * simpl. eapply ihb. assumption. - - intros Σ wfΣ Γ wfΓ t na A B u X ht iht hu ihu Δ f hf. + - intros Σ wfΣ Γ wfΓ t na A B s u X hty ihty ht iht hu ihu Δ f hf. simpl. eapply meta_conv. - + econstructor. + + eapply type_App. + * simpl in ihty. eapply ihty; eassumption. * simpl in iht. eapply iht. assumption. * eapply ihu. assumption. + autorewrite with sigma. rewrite !subst1_inst. sigma. @@ -1722,24 +1746,22 @@ Proof. simpl. eapply meta_conv. + eapply type_Fix. - * eapply fix_guard_rename. assumption. + * eapply fix_guard_rename; eauto. * rewrite nth_error_map. rewrite hdecl. simpl. reflexivity. * apply hf. * apply All_map, (All_impl ihmfixt). intros x [s [Hs IHs]]. exists s. now apply IHs. * apply All_map, (All_impl ihmfixb). - intros x [[Hb Hlam] IHb]. + intros x [Hb IHb]. destruct x as [na ty bo rarg]. simpl in *. - split. - -- rewrite rename_fix_context. - eapply meta_conv. - ++ apply (IHb (Δ ,,, rename_context f types) (shiftn #|mfix| f)). - split; auto. subst types. rewrite -(fix_context_length mfix). - apply urenaming_context; auto. apply hf. - ++ autorewrite with sigma. subst types. rewrite fix_context_length. - now rewrite -ren_shiftn up_Upn shiftn_consn_idsn. - -- eapply isLambda_rename. assumption. + rewrite rename_fix_context. + eapply meta_conv. + ++ apply (IHb (Δ ,,, rename_context f types) (shiftn #|mfix| f)). + split; auto. subst types. rewrite -(fix_context_length mfix). + apply urenaming_context; auto. apply hf. + ++ autorewrite with sigma. subst types. rewrite fix_context_length. + now rewrite -ren_shiftn up_Upn shiftn_consn_idsn. * admit (* wf_fixpoint renaming *). + reflexivity. @@ -1777,7 +1799,7 @@ Proof. + eapply ihB. assumption. + eapply cumul_rename. all: try eassumption. apply hf. -Abort. +Admitted. Lemma typing_rename : forall Σ Γ Δ f t A, @@ -1786,7 +1808,7 @@ Lemma typing_rename : Σ ;;; Γ |- t : A -> Σ ;;; Δ |- rename f t : rename f A. Proof. -Abort. +Admitted. (* intros Σ Γ Δ f t A hΣ hf h. revert Σ hΣ Γ t A h Δ f hf. apply typing_rename_prop. @@ -2218,9 +2240,12 @@ Proof. constructor; auto. ** exists s1. apply ihB; auto. ** apply ihb; auto. - - intros Σ wfΣ Γ wfΓ t na A B u X ht iht hu ihu Δ σ hΔ hσ. + - intros Σ wfΣ Γ wfΓ t na A B s u X hty ihty ht iht hu ihu Δ σ hΔ hσ. autorewrite with sigma. econstructor. + * specialize (ihty _ _ hΔ hσ). + simpl in ihty. eapply meta_conv_term; [eapply ihty|]. + now rewrite up_Up. * specialize (iht _ _ hΔ hσ). simpl in iht. eapply meta_conv; [eapply iht|]. now rewrite up_Up. diff --git a/pcuic/theories/PCUICSpine.v b/pcuic/theories/PCUICSpine.v index 3bdb05f71..b116d3b3a 100644 --- a/pcuic/theories/PCUICSpine.v +++ b/pcuic/theories/PCUICSpine.v @@ -130,8 +130,8 @@ Proof. rewrite H [firstn _ [u]]firstn_0 // app_nil_r in Hr |- *. specialize (IHctx Hr). rewrite app_assoc. now econstructor. - * destruct a as [na' [b'|] ty']; simpl in *; noconf H. - rewrite skipn_S in Hl, Hr, H. subst b. + * rewrite skipn_S in Hl, Hr, H. + destruct a as [na' [b'|] ty']; simpl in *; noconf H. pose proof (context_subst_length Hl). rewrite subst_context_length in H. rewrite {3}H -subst_app_simpl [firstn #|ctx| _ ++ _]firstn_skipn. constructor. apply IHctx => //. @@ -298,7 +298,7 @@ Lemma arity_typing_spine {cf:checker_flags} Σ Γ Γ' s inst s' : ∑ instsubst, spine_subst Σ Γ inst instsubst Γ'. Proof. intros wfΣ wfΓ'; revert s inst s'. - assert (wf_local Σ Γ). now apply wf_local_app in wfΓ'. move X after wfΓ'. + assert (wf_local Σ Γ). now apply wf_local_app_l in wfΓ'. move X after wfΓ'. rename X into wfΓ. generalize (le_n #|Γ'|). generalize (#|Γ'|) at 2. @@ -325,12 +325,12 @@ Proof. specialize (IHn (subst_context [b] 0 l)). forward IHn. { rewrite app_context_assoc in wfΓ'. - apply All_local_env_app in wfΓ' as [wfb wfa]. + apply All_local_env_app_inv in wfΓ' as [wfb wfa]. depelim wfb. eapply substitution_wf_local. eauto. epose proof (cons_let_def Σ Γ [] [] na b ty ltac:(constructor)). rewrite !subst_empty in X. eapply X. auto. - eapply All_local_env_app_inv; split. + eapply All_local_env_app; split. constructor; auto. apply wfa. } forward IHn by rewrite subst_context_length; lia. specialize (IHn s inst s' Hsp). @@ -349,7 +349,7 @@ Proof. * apply subslet_app. now rewrite subst_empty. repeat constructor. rewrite app_context_assoc in wfΓ'. simpl in wfΓ'. - apply wf_local_app in wfΓ'. depelim wfΓ'; now rewrite !subst_empty. + apply wf_local_app_l in wfΓ'. depelim wfΓ'; now rewrite !subst_empty. + rewrite PCUICCtxShape.context_assumptions_app /=. depelim Hsp. now eapply cumul_Prod_Sort_inv in c. @@ -359,14 +359,14 @@ Proof. specialize (IHn (subst_context [hd] 0 l)). forward IHn. { rewrite app_context_assoc in wfΓ'. - apply All_local_env_app in wfΓ' as [wfb wfa]; eauto. + apply All_local_env_app_inv in wfΓ' as [wfb wfa]; eauto. depelim wfb. eapply substitution_wf_local. auto. constructor. constructor. rewrite subst_empty. eapply type_Cumul'. eapply t. eapply l0. eapply conv_cumul; auto. now symmetry. - eapply All_local_env_app_inv; eauto; split. + eapply All_local_env_app; eauto; split. constructor; eauto. eapply isType_tProd in i; intuition eauto. } forward IHn by rewrite subst_context_length; lia. specialize (IHn s tl s'). @@ -388,7 +388,7 @@ Proof. * apply subslet_app => //. repeat constructor. rewrite app_context_assoc in wfΓ'. simpl in wfΓ'. - apply wf_local_app in wfΓ'. depelim wfΓ'. + apply wf_local_app_l in wfΓ'. depelim wfΓ'. rewrite !subst_empty. red in l0. eapply type_Cumul'; eauto. eapply conv_cumul. now symmetry. Qed. @@ -650,7 +650,7 @@ Proof. rewrite firstn_0 in csr => //. rewrite app_nil_r in csr. eapply subslet_app_inv in subs as [sl sr]. split; split; auto. rewrite app_context_assoc in wfcodom. - now eapply All_local_env_app in wfcodom as [? ?]. + now eapply All_local_env_app_inv in wfcodom as [? ?]. eapply substitution_wf_local; eauto. now rewrite app_context_assoc in wfcodom. Qed. @@ -1430,7 +1430,7 @@ Proof. pose proof wf as wf'. rewrite -eql in wf'. rewrite app_context_assoc in wf'. - apply wf_local_app in wf'. depelim wf'. + apply wf_local_app_l in wf'. depelim wf'. eapply (weakening_typing); auto. * rewrite app_length /= Nat.add_1_r in IHc. constructor; auto. @@ -1438,7 +1438,7 @@ Proof. pose proof wf as wf'. rewrite -eql in wf'. rewrite app_context_assoc in wf'. - apply wf_local_app in wf'. depelim wf'. + apply wf_local_app_l in wf'. depelim wf'. rewrite Nat.add_0_r. eapply type_Cumul'. @@ -1721,8 +1721,8 @@ Proof. * eapply subslet_app => //. eapply sps. rewrite -{1}(subst_empty 0 b). repeat constructor. rewrite !subst_empty. - eapply All_local_env_app in wfΓΔ as [_ wf]. - eapply All_local_env_app in wf as [wfd _]. + eapply All_local_env_app_inv in wfΓΔ as [_ wf]. + eapply All_local_env_app_inv in wf as [wfd _]. depelim wfd. apply l0. * rewrite subst_app_simpl. move: (context_subst_length sps). @@ -1747,7 +1747,7 @@ Proof. 2:{ eauto. } forward X. { pose proof wfΓΔ as wfΓΔ'. - rewrite app_context_assoc in wfΓΔ'. eapply All_local_env_app in wfΓΔ' as [wfΓΔ' _]. + rewrite app_context_assoc in wfΓΔ'. eapply All_local_env_app_inv in wfΓΔ' as [wfΓΔ' _]. eapply (isType_subst wfΣ wfΓΔ') in wat''; eauto. 2:{ repeat constructor. now rewrite subst_empty. } now rewrite subst_it_mkProd_or_LetIn Nat.add_0_r in wat''. } @@ -1783,7 +1783,7 @@ Lemma typing_spine_app {cf:checker_flags} Σ Γ ty args na A B arg : Proof. intros wfΣ H; revert arg. dependent induction H. - - intros arg Harg. simpl. econstructor; eauto. + - intros arg Harg. simpl. econstructor; eauto. constructor. 2:reflexivity. eapply isType_tProd in i as [watd wat]. eapply (isType_subst wfΣ (Δ:=[vass na A])); eauto. @@ -1889,7 +1889,7 @@ Proof. simpl. rewrite smash_context_acc. simpl. auto. auto. } split; auto. - - eapply All_local_env_app_inv; split; auto. + - eapply All_local_env_app; split; auto. eapply wf_local_rel_smash_context; auto. - induction inst_subslet0 in inst, inst_ctx_subst0, spine_codom_wf0 |- *. depelim inst_ctx_subst0. diff --git a/pcuic/theories/PCUICSubstitution.v b/pcuic/theories/PCUICSubstitution.v index 4ab91eead..46164e095 100644 --- a/pcuic/theories/PCUICSubstitution.v +++ b/pcuic/theories/PCUICSubstitution.v @@ -100,7 +100,7 @@ Qed. Lemma subst_decl_closed n k d : closed_decl k d -> subst_decl n k d = d. Proof. case: d => na [body|] ty; rewrite /closed_decl /subst_decl /map_decl /=. - - move/andP => [cb cty]. rewrite !subst_closedn //. + - move/andb_and => [cb cty]. rewrite !subst_closedn //. - move=> cty; now rewrite !subst_closedn //. Qed. @@ -122,7 +122,7 @@ Proof. induction ctx using rev_ind; try easy. move=> n0. rewrite /closedn_ctx !rev_app_distr /id /=. - move/andP => [closedx Hctx]. + move/andb_and => [closedx Hctx]. rewrite subst_decl_closed //. - rewrite (closed_decl_upwards n0) //; lia. - f_equal. now rewrite IHctx. @@ -353,7 +353,7 @@ Proof. move: Hdecl. rewrite /closed_inductive_decl /lift_mutual_inductive_body. destruct decl; simpl. - move/andP => [clpar clbodies]. f_equal. + move/andb_and => [clpar clbodies]. f_equal. - now rewrite [fold_context _ _]closed_ctx_subst. - eapply forallb_All in clbodies. eapply Alli_mapi_id. @@ -361,7 +361,7 @@ Proof. intros. eapply H. * simpl; intros. move: H. rewrite /closed_inductive_body. - destruct x; simpl. move=> /andP[/andP [ci ct] cp]. + destruct x; simpl. move=> /andb_and[/andb_and [ci ct] cp]. f_equal. + rewrite subst_closedn; eauto. eapply closed_upwards; eauto; lia. @@ -372,7 +372,7 @@ Proof. eapply closed_upwards; eauto; lia. + simpl in X. rewrite -X in cp. eapply forallb_All in cp. eapply All_map_id; eauto. - eapply (All_impl cp); firstorder auto. + eapply (All_impl cp); pcuicfo. destruct x; unfold on_snd; simpl; f_equal. apply subst_closedn. rewrite context_assumptions_fold. eapply closed_upwards; eauto; lia. @@ -1004,14 +1004,14 @@ Lemma subst_decompose_prod_assum_rec ctx t s k : (decompose_prod_assum (subst_context s k ctx) (subst s (length ctx + k) t) = (subst_context s k ctx' ,,, ctx'', t'')). Proof. - induction t in ctx, k |- *; simpl; try solve [ eexists _, _; firstorder eauto ]. + induction t in ctx, k |- *; simpl; try solve [ eexists _, _; pcuicfo ]. - elim: leb_spec_Set => comp. + destruct (nth_error s (n - (#|ctx| + k))) eqn:Heq. * destruct decompose_prod_assum eqn:Heq'. eexists _, _; intuition eauto. now rewrite decompose_prod_assum_ctx Heq'. - * eexists _,_; firstorder eauto. - + eexists _,_; firstorder eauto. + * eexists _,_; pcuicfo. + + eexists _,_; pcuicfo. - destruct decompose_prod_assum eqn:Heq. rewrite decompose_prod_assum_ctx in Heq. destruct (decompose_prod_assum [] t2) eqn:Heq'. @@ -2328,7 +2328,7 @@ Proof. eapply on_declared_inductive in isdecl as [on_mind on_ind]; auto. apply onArity in on_ind as [[s' Hindty] _]. apply typecheck_closed in Hindty as [_ Hindty]; eauto. symmetry. - move/andP/proj1: Hindty. rewrite -(closedn_subst_instance_constr _ _ u) => Hty. + move/andb_and/proj1: Hindty. rewrite -(closedn_subst_instance_constr _ _ u) => Hty. apply: (subst_closedn s #|Δ|); auto with wf. eapply closed_upwards. eauto. simpl; lia. @@ -2371,20 +2371,20 @@ Proof. - assert (wf_local Σ (Γ ,,, subst_context s 0 Δ ,,, subst_context s #|Δ| (fix_context mfix))). subst types. - apply All_local_env_app in X as [X Hfixc]. - apply All_local_env_app_inv. intuition. + apply All_local_env_app_inv in X as [X Hfixc]. + apply All_local_env_app. intuition. revert Hfixc. clear X0 X H. induction 1; simpl; auto. + destruct t0 as [u [Ht IHt]]. specialize (IHt Γ Γ' (Δ ,,, Γ0) s sub). forward IHt. now rewrite app_context_assoc. rewrite app_context_length subst_context_app app_context_assoc Nat.add_0_r in IHt. unfold snoc; rewrite subst_context_snoc; econstructor; auto. exists u. - apply IHt; apply All_local_env_app_inv; intuition. + apply IHt; apply All_local_env_app; intuition. + destruct t0 as [Ht IHt]. specialize (IHt Γ Γ' (Δ ,,, Γ0) s sub). forward IHt. now rewrite app_context_assoc. rewrite app_context_length subst_context_app app_context_assoc Nat.add_0_r in IHt. unfold snoc; rewrite subst_context_snoc; econstructor; auto; - apply IHt; apply All_local_env_app_inv; intuition. + apply IHt; apply All_local_env_app; intuition. + erewrite map_dtype. eapply type_Fix. * rewrite nth_error_map H. reflexivity. * now rewrite subst_fix_context. @@ -2405,20 +2405,20 @@ Proof. - assert (wf_local Σ (Γ ,,, subst_context s 0 Δ ,,, subst_context s #|Δ| (fix_context mfix))). subst types. - apply All_local_env_app in X as [X Hfixc]. - apply All_local_env_app_inv. intuition. + apply All_local_env_app_inv in X as [X Hfixc]. + apply All_local_env_app. intuition. revert Hfixc. clear X0 X H. induction 1; simpl; auto. + destruct t0 as [u [Ht IHt]]. specialize (IHt Γ Γ' (Δ ,,, Γ0) s sub). forward IHt. now rewrite app_context_assoc. rewrite app_context_length subst_context_app app_context_assoc Nat.add_0_r in IHt. unfold snoc; rewrite subst_context_snoc; econstructor; auto. exists u. - apply IHt; apply All_local_env_app_inv; intuition. + apply IHt; apply All_local_env_app; intuition. + destruct t0 as [Ht IHt]. specialize (IHt Γ Γ' (Δ ,,, Γ0) s sub). forward IHt. now rewrite app_context_assoc. rewrite app_context_length subst_context_app app_context_assoc Nat.add_0_r in IHt. unfold snoc; rewrite subst_context_snoc; econstructor; auto; - apply IHt; apply All_local_env_app_inv; intuition. + apply IHt; apply All_local_env_app; intuition. + erewrite map_dtype. eapply type_CoFix. * rewrite nth_error_map H. reflexivity. * now rewrite subst_fix_context. @@ -2444,7 +2444,7 @@ Proof. pose proof (subst_destArity [] B s #|Δ|). rewrite Hd in H. rewrite H. clear H. split; auto. - apply All_local_env_app_inv; intuition auto. + apply All_local_env_app; intuition auto. clear -sub wfsubs a. induction ctx; try constructor; depelim a. -- rewrite subst_context_snoc. @@ -2455,7 +2455,7 @@ Proof. specialize (t0 _ _ (Δ ,,, ctx) _ sub). forward t0. now rewrite app_context_assoc. simpl in t0. forward t0. rewrite subst_context_app app_context_assoc Nat.add_0_r. - apply All_local_env_app_inv. split; auto. + apply All_local_env_app. split; auto. eapply IHctx. eapply a. now rewrite subst_context_app Nat.add_0_r app_context_assoc app_length in t0. -- rewrite subst_context_snoc. @@ -2466,7 +2466,7 @@ Proof. specialize (t0 _ _ (Δ ,,, ctx) _ sub). forward t0. now rewrite app_context_assoc. simpl in t0. forward t0. rewrite subst_context_app app_context_assoc Nat.add_0_r. - apply All_local_env_app_inv. split; auto. + apply All_local_env_app. split; auto. eapply IHctx. eapply a. now rewrite subst_context_app Nat.add_0_r app_context_assoc app_length in t0. + right; exists u; intuition eauto. @@ -2490,7 +2490,7 @@ Proof. - induction Δ. - + clear X. simpl in *. now apply All_local_env_app in wfΓ0. + + clear X. simpl in *. now apply All_local_env_app_inv in wfΓ0. + rewrite subst_context_snoc. depelim X; simpl; econstructor; eauto. @@ -2567,7 +2567,7 @@ Proof. rewrite subst_context_snoc0 in X5. econstructor; eauto. - - specialize (X1 Γ Γ' Δ s sub eq_refl). + - specialize (X3 _ _ _ s0 sub eq_refl). eapply refine_type. 1: econstructor; eauto. unfold subst1. rewrite -> distr_subst. simpl. reflexivity. @@ -2578,7 +2578,7 @@ Proof. eapply on_declared_inductive in isdecl as [on_mind on_ind]; auto. apply onArity in on_ind as [s' Hindty]. apply typecheck_closed in Hindty as [_ Hindty]; eauto. symmetry. - move/andP/proj1: Hindty. rewrite -(closedn_subst_instance_constr _ _ u) => Hty. + move/andb_and/proj1: Hindty. rewrite -(closedn_subst_instance_constr _ _ u) => Hty. apply: (subst_closedn s #|Δ|); auto with wf. eapply closed_upwards. 1: eauto. simpl; lia. @@ -2602,7 +2602,7 @@ Proof. destruct isdecl as [_ oind]. clear -oind wfΣ. apply onArity in oind; destruct oind as [s HH]; cbn in *. apply typecheck_closed in HH; eauto. - apply snd in HH. apply andP in HH; apply HH. } + apply snd in HH. apply andb_and in HH; apply HH. } simpl. econstructor. all: eauto. + eapply subst_build_case_predicate_type in H0; tea. * simpl in *. subst params. rewrite firstn_map. @@ -2644,16 +2644,14 @@ Proof. now specialize (Hs' _ _ _ _ sub eq_refl). * eapply All_map. eapply (All_impl X1); simpl. - intros x [[Hb Hlam] IH]. + intros x [Hb IH]. rewrite subst_fix_context. specialize (IH Γ Γ' (Δ ,,, (fix_context mfix)) _ sub). rewrite app_context_assoc in IH. specialize (IH eq_refl). - split; auto. - + rewrite subst_context_app Nat.add_0_r app_context_assoc in IH. - rewrite app_context_length fix_context_length in IH. - rewrite subst_context_length fix_context_length. - rewrite commut_lift_subst_rec; try lia. now rewrite (Nat.add_comm #|Δ|). - + now rewrite isLambda_subst. + rewrite subst_context_app Nat.add_0_r app_context_assoc in IH. + rewrite app_context_length fix_context_length in IH. + rewrite subst_context_length fix_context_length. + rewrite commut_lift_subst_rec; try lia. now rewrite (Nat.add_comm #|Δ|). * move: H1. rewrite /wf_fixpoint. pose proof (substitution_check_one_fix s #|Δ| mfix). @@ -2663,7 +2661,7 @@ Proof. - rewrite -> (map_dtype _ (subst s (#|mfix| + #|Δ|))). eapply type_CoFix; auto. - * eapply cofix_guard_subst; auto. + * eapply cofix_guard_subst; eauto. * now rewrite -> nth_error_map, H0. * eapply All_map. eapply (All_impl X0); simpl. diff --git a/pcuic/theories/PCUICTemplateEquiv.v b/pcuic/theories/PCUICTemplateEquiv.v new file mode 100644 index 000000000..60b3b0edd --- /dev/null +++ b/pcuic/theories/PCUICTemplateEquiv.v @@ -0,0 +1,34 @@ + +Lemma T_cumul_trans {cf} (Σ : PCUICAst.global_env_ext) Γ T U V : + STy.wf (trans_global Σ) -> + PCUICTyping.wf (TemplateToPCUIC.trans_global (trans_global Σ)).1 -> + wf T -> wf U -> wf V -> + TT.cumul (trans_global Σ) (trans_local Γ) T U -> + TT.cumul (trans_global Σ) (trans_local Γ) U V -> + TT.cumul (trans_global Σ) (trans_local Γ) T V. +Proof. + intros wftΣ wftΣ' wfT wfU wfV. + assert (wfΓ : Forall wf_decl (trans_local Γ)). admit. + intros TU UV. + pose proof (TemplateToPCUICCorrectness.trans_cumul (trans_global Σ) (trans_local Γ) _ _ wftΣ wfΓ wfT wfU TU). + pose proof (TemplateToPCUICCorrectness.trans_cumul (trans_global Σ) (trans_local Γ) _ _ wftΣ wfΓ wfU wfV UV). + pose proof (PCUICConversion.cumul_trans _ _ _ _ _ _ X X0). + apply trans_cumul in X1. +Admitted. + +Lemma T_cumul_trans' {cf} (Σ : global_env_ext) Γ T U V : + STy.wf Σ -> + PCUICTyping.wf (TemplateToPCUIC.trans_global_decls Σ) -> + wf T -> wf U -> wf V -> + TT.cumul Σ Γ T U -> + TT.cumul Σ Γ U V -> + TT.cumul Σ Γ T V. +Proof. + intros wftΣ wftransb wfT wfU wfV. + assert (wfΓ : Forall wf_decl Γ). admit. + intros TU UV. + pose proof (TemplateToPCUICCorrectness.trans_cumul Σ Γ _ _ wftΣ wfΓ wfT wfU TU). + pose proof (TemplateToPCUICCorrectness.trans_cumul Σ Γ _ _ wftΣ wfΓ wfU wfV UV). + unshelve epose proof (PCUICConversion.cumul_trans _ _ _ _ _ _ X X0). + apply trans_cumul in X1. +Admitted. diff --git a/pcuic/theories/PCUICToTemplate.v b/pcuic/theories/PCUICToTemplate.v index 7883b7a58..add6facdd 100644 --- a/pcuic/theories/PCUICToTemplate.v +++ b/pcuic/theories/PCUICToTemplate.v @@ -1,6 +1,21 @@ (* Distributed under the terms of the MIT license. *) +From Coq Require Import Int63 FloatOps FloatAxioms. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils. +Set Warnings "-notation-overridden". From MetaCoq.Template Require Import config utils AstUtils BasicAst Ast. +Set Warnings "+notation-overridden". + +Definition uint63_from_model (i : uint63_model) : Int63.int := + Int63.of_Z (proj1_sig i). + +Definition float64_from_model (f : float64_model) : PrimFloat.float := + FloatOps.SF2Prim (proj1_sig f). + +Definition trans_prim (t : prim_val) : Ast.term := + match t.π2 with + | primIntModel i => Ast.tInt (uint63_from_model i) + | primFloatModel f => Ast.tFloat (float64_from_model f) + end. Fixpoint trans (t : PCUICAst.term) : Ast.term := match t with @@ -25,6 +40,7 @@ Fixpoint trans (t : PCUICAst.term) : Ast.term := | PCUICAst.tCoFix mfix idx => let mfix' := List.map (map_def trans trans) mfix in tCoFix mfix' idx + | PCUICAst.tPrim i => trans_prim i end. Definition trans_decl (d : PCUICAst.context_decl) := diff --git a/pcuic/theories/PCUICToTemplateCorrectness.v b/pcuic/theories/PCUICToTemplateCorrectness.v index c3889746a..c537f943b 100644 --- a/pcuic/theories/PCUICToTemplateCorrectness.v +++ b/pcuic/theories/PCUICToTemplateCorrectness.v @@ -1,12 +1,20 @@ (* Distributed under the terms of the MIT license. *) +From Coq Require Import ssreflect ssrbool CRelationClasses. +From Equations.Type Require Import Relation Relation_Properties. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICInduction - PCUICLiftSubst PCUICEquality PCUICUnivSubst PCUICTyping PCUICGeneration. + PCUICLiftSubst PCUICEquality PCUICReduction PCUICUnivSubst PCUICTyping PCUICGeneration + PCUICConversion. (* Needs transitivity of cumulativity *) +Set Warnings "-notation-overridden". From MetaCoq.Template Require Import config utils Ast TypingWf WfInv UnivSubst - LiftSubst. + TermEquality LiftSubst Reduction. + +Set Warnings "+notation_overridden". Require Import PCUICToTemplate. +Implicit Types cf : checker_flags. (* Use {cf} to parameterize by checker_flags where needed *) + (* from Checker Generation to avoid dependencies *) Derive Signature for Template.Typing.typing. @@ -15,48 +23,41 @@ Derive NoConfusion for term. From Equations Require Import Equations. Require Import Equations.Prop.DepElim. -Module P := PCUICAst. -Module PT := PCUICTyping. +(** Source = PCUIC, Target = Coq *) +Module S := PCUICAst. +Module ST := PCUICTyping. Module T := Template.Ast. Module TT := Template.Typing. -Local Existing Instance default_checker_flags. - -Module PL := PCUICLiftSubst. +Module SL := PCUICLiftSubst. Module TL := Template.LiftSubst. Ltac solve_list := - rewrite !map_map_compose, ?compose_on_snd, ?compose_map_def; + rewrite !map_map_compose ?compose_on_snd ?compose_map_def; try rewrite !map_length; try solve_all; try typeclasses eauto with core. -Lemma mkAppMkApps s t: - mkApp s t = mkApps s [t]. -Proof. - cbn. unfold mkApp. - reflexivity. -Qed. - Lemma mapOne {X Y} (f:X->Y) x: map f [x] = [f x]. Proof. reflexivity. Qed. -Lemma trans_lift (t:P.term) n k: - trans(PL.lift n k t) = TL.lift n k (trans t). +Lemma trans_lift (t : S.term) n k: + trans (SL.lift n k t) = TL.lift n k (trans t). Proof. revert k. induction t using PCUICInduction.term_forall_list_ind; simpl; intros; try congruence. - f_equal. rewrite !map_map_compose. solve_all. - - rewrite IHt1, IHt2. - rewrite mkAppMkApps. + - rewrite IHt1 IHt2. + rewrite AstUtils.mkAppMkApps. rewrite <- mapOne. rewrite <- TL.lift_mkApps. f_equal. - f_equal; auto. red in X. solve_list. - f_equal; auto; red in X; solve_list. - f_equal; auto; red in X; solve_list. + - destruct p as [? []]; eauto. Qed. Definition on_fst {A B C} (f:A->C) (p:A×B) := (f p.1, p.2). @@ -72,22 +73,22 @@ Proof. Defined. Lemma trans_global_ext_levels Σ: -PT.global_ext_levels Σ = TT.global_ext_levels (trans_global Σ). +ST.global_ext_levels Σ = TT.global_ext_levels (trans_global Σ). Proof. - unfold TT.global_ext_levels, PT.global_ext_levels. + unfold TT.global_ext_levels, ST.global_ext_levels. destruct Σ. cbn [trans_global fst snd]. f_equal. induction g. - reflexivity. - - unfold PT.global_levels in IHg. + - unfold ST.global_levels in IHg. cbn. rewrite IHg. f_equal. destruct a. cbn. unfold TT.monomorphic_levels_decl, TT.monomorphic_udecl_decl, TT.on_udecl_decl. - unfold PT.monomorphic_levels_decl, PT.monomorphic_udecl_decl, PT.on_udecl_decl. + unfold ST.monomorphic_levels_decl, ST.monomorphic_udecl_decl, ST.on_udecl_decl. destruct g0. + cbn. destruct c. @@ -98,16 +99,20 @@ Proof. Qed. Lemma trans_global_ext_constraints Σ : - PT.global_ext_constraints Σ = TT.global_ext_constraints (trans_global Σ). + ST.global_ext_constraints Σ = TT.global_ext_constraints (trans_global Σ). Proof. destruct Σ. - unfold PT.global_ext_constraints, trans_global. - cbn [fst snd]. -Admitted. + unfold ST.global_ext_constraints, TT.global_ext_constraints. simpl. + f_equal. clear u. + induction g. + - reflexivity. + - simpl. rewrite IHg. f_equal. clear. + destruct a as [? []]; reflexivity. +Qed. Lemma trans_mem_level_set l Σ: -LevelSet.mem l (PT.global_ext_levels Σ) -> -LevelSet.mem l (TT.global_ext_levels (trans_global Σ)). + LevelSet.mem l (ST.global_ext_levels Σ) -> + LevelSet.mem l (TT.global_ext_levels (trans_global Σ)). Proof. intros. rewrite <- trans_global_ext_levels. @@ -115,60 +120,52 @@ Proof. Qed. Lemma trans_in_level_set l Σ: -LevelSet.In l (PT.global_ext_levels Σ) -> -LevelSet.In l (TT.global_ext_levels (trans_global Σ)). + LevelSet.In l (ST.global_ext_levels Σ) -> + LevelSet.In l (TT.global_ext_levels (trans_global Σ)). Proof. intros. rewrite <- trans_global_ext_levels. apply H. Qed. - -Lemma trans_lookup Σ cst decl: -P.lookup_env Σ.1 cst = Some decl -> -Ast.lookup_env (trans_global Σ).1 cst = -Some (trans_global_decl decl). +Lemma trans_lookup Σ cst : + Ast.lookup_env (trans_global_decls Σ) cst = option_map trans_global_decl (S.lookup_env Σ cst). Proof. - intros H. - destruct Σ. cbn in *. - induction g;cbn in H. - - discriminate H. + induction Σ. + - reflexivity. - cbn. unfold eq_kername in *; destruct kername_eq_dec; subst. - + destruct a. - cbn in *. - injection H as ->. - reflexivity. - + apply IHg, H. + + destruct a; auto. + + now rewrite IHΣ. Qed. Lemma trans_declared_constant Σ cst decl: -PT.declared_constant Σ.1 cst decl -> -TT.declared_constant (trans_global Σ).1 cst (trans_constant_body decl). + ST.declared_constant Σ cst decl -> + TT.declared_constant (trans_global_decls Σ) cst (trans_constant_body decl). Proof. - intros H. unfold TT.declared_constant. - unfold PT.declared_constant in H. - apply trans_lookup in H as ->. + rewrite trans_lookup. + unfold ST.declared_constant. + intros ->. reflexivity. Qed. Lemma trans_constraintSet_in x Σ: -ConstraintSet.In x (PT.global_ext_constraints Σ) -> -ConstraintSet.In x (TT.global_ext_constraints (trans_global Σ)). + ConstraintSet.In x (ST.global_ext_constraints Σ) -> + ConstraintSet.In x (TT.global_ext_constraints (trans_global Σ)). Proof. rewrite trans_global_ext_constraints. trivial. Qed. -Lemma trans_consistent_instance_ext Σ decl u: -PT.consistent_instance_ext Σ decl u -> -TT.consistent_instance_ext (trans_global Σ) decl u. +Lemma trans_consistent_instance_ext {cf} Σ decl u: + ST.consistent_instance_ext Σ decl u -> + TT.consistent_instance_ext (trans_global Σ) decl u. Proof. intros H. - unfold consistent_instance_ext, PT.consistent_instance_ext in *. - unfold consistent_instance, PT.consistent_instance in *. + unfold consistent_instance_ext, ST.consistent_instance_ext in *. + unfold consistent_instance, ST.consistent_instance in *. destruct decl;trivial. destruct H as (?&?&?). repeat split;trivial. @@ -190,20 +187,19 @@ Proof. Qed. Lemma trans_declared_inductive Σ mdecl ind idecl: -PT.declared_inductive Σ.1 mdecl ind idecl -> -TT.declared_inductive (trans_global Σ).1 (trans_minductive_body mdecl) ind (trans_one_ind_body idecl). + ST.declared_inductive Σ mdecl ind idecl -> + TT.declared_inductive (trans_global_decls Σ) (trans_minductive_body mdecl) ind (trans_one_ind_body idecl). Proof. intros []. split. - - unfold TT.declared_minductive, PT.declared_minductive in *. - apply trans_lookup in H as ->. - reflexivity. + - unfold TT.declared_minductive, ST.declared_minductive in *. + now rewrite trans_lookup H. - now apply map_nth_error. Qed. Lemma trans_mkApps t args: -trans (PCUICAst.mkApps t args) = -mkApps (trans t) (map trans args). + trans (PCUICAst.mkApps t args) = + mkApps (trans t) (map trans args). Proof. induction args in t |- *. - reflexivity. @@ -214,30 +210,16 @@ Proof. reflexivity. Qed. - Lemma trans_decl_type decl: -trans (PCUICAst.decl_type decl) = -decl_type (trans_decl decl). + trans (PCUICAst.decl_type decl) = + decl_type (trans_decl decl). Proof. destruct decl. reflexivity. Qed. -Lemma All_forall {X Y} (f:X->Y->Prop) xs: - All (fun a => forall b, f a b) xs -> - (forall b, All (fun a => f a b) xs). -Proof. - intros. - induction X0. - - constructor. - - constructor. - + apply p. - + apply IHX0. -Qed. - - Lemma trans_subst xs k t: - trans (PL.subst xs k t) = + trans (SL.subst xs k t) = TL.subst (map trans xs) k (trans t). Proof. induction t in k |- * using PCUICInduction.term_forall_list_ind. @@ -253,7 +235,7 @@ Proof. apply All_map_eq. eapply All_forall in X. apply X. - - rewrite IHt1, IHt2. + - rewrite IHt1 IHt2. destruct (trans t1);cbn;trivial. rewrite AstUtils.mkApp_mkApps. do 2 f_equal. @@ -280,7 +262,7 @@ Proof. destruct x. unfold map_def. cbn in *. - now rewrite e, e0. + now rewrite e e0. + apply IHX. - f_equal. rewrite map_length. @@ -292,31 +274,32 @@ Proof. destruct x. unfold map_def. cbn in *. - now rewrite e, e0. + now rewrite e e0. + apply IHX. + - destruct p as [? []]; eauto. Qed. Lemma trans_subst10 u B: -trans (PL.subst1 u 0 B)= -TL.subst10 (trans u) (trans B). + trans (SL.subst1 u 0 B) = + TL.subst10 (trans u) (trans B). Proof. - unfold PL.subst1. + unfold SL.subst1. rewrite trans_subst. reflexivity. Qed. Lemma trans_subst_instance_constr u t: -trans (PCUICUnivSubst.subst_instance_constr u t) = -Template.UnivSubst.subst_instance_constr u (trans t). + trans (PCUICUnivSubst.subst_instance_constr u t) = + Template.UnivSubst.subst_instance_constr u (trans t). Proof. induction t using PCUICInduction.term_forall_list_ind;cbn;auto;try congruence. - do 2 rewrite map_map. f_equal. apply All_map_eq. apply X. - - rewrite IHt1, IHt2. - do 2 rewrite mkAppMkApps. + - rewrite IHt1 IHt2. + do 2 rewrite AstUtils.mkAppMkApps. rewrite subst_instance_constr_mkApps. cbn. reflexivity. @@ -342,7 +325,7 @@ Proof. + destruct x;cbn in *. unfold map_def;cbn. destruct p. - now rewrite e, e0. + now rewrite e e0. + apply IHX. - f_equal. unfold tFixProp in X. @@ -351,38 +334,39 @@ Proof. + destruct x;cbn in *. unfold map_def;cbn. destruct p. - now rewrite e, e0. + now rewrite e e0. + apply IHX. + - destruct p as [? []]; eauto. Qed. Lemma trans_cst_type decl: -trans (PCUICAst.cst_type decl) = -cst_type (trans_constant_body decl). + trans (PCUICAst.cst_type decl) = + cst_type (trans_constant_body decl). Proof. reflexivity. Qed. Lemma trans_ind_type idecl: -trans (PCUICAst.ind_type idecl) = -ind_type (trans_one_ind_body idecl). + trans (PCUICAst.ind_type idecl) = + ind_type (trans_one_ind_body idecl). Proof. reflexivity. Qed. Lemma trans_type_of_constructor mdecl cdecl ind i u: -trans (PT.type_of_constructor mdecl cdecl (ind, i) u) = -TT.type_of_constructor - (trans_minductive_body mdecl) - (trans_ctor cdecl) - (ind,i) - u. -Proof. - unfold PT.type_of_constructor. + trans (ST.type_of_constructor mdecl cdecl (ind, i) u) = + TT.type_of_constructor + (trans_minductive_body mdecl) + (trans_ctor cdecl) + (ind,i) + u. +Proof. + unfold ST.type_of_constructor. rewrite trans_subst. unfold TT.type_of_constructor. f_equal. - cbn [fst]. - rewrite PT.inds_spec. + rewrite ST.inds_spec. rewrite TT.inds_spec. rewrite map_rev. rewrite map_mapi. @@ -402,24 +386,24 @@ Proof. Qed. Lemma trans_dtype decl: -trans (dtype decl) = -dtype (trans_def decl). + trans (dtype decl) = + dtype (trans_def decl). Proof. destruct decl. reflexivity. Qed. Lemma trans_dbody decl: -trans (dbody decl) = -dbody (trans_def decl). + trans (dbody decl) = + dbody (trans_def decl). Proof. destruct decl. reflexivity. Qed. Lemma trans_declared_projection Σ mdecl idecl p pdecl : -PT.declared_projection Σ.1 mdecl idecl p pdecl -> -TT.declared_projection (trans_global Σ).1 (trans_minductive_body mdecl) (trans_one_ind_body idecl) p (on_snd trans pdecl). + ST.declared_projection Σ.1 mdecl idecl p pdecl -> + TT.declared_projection (trans_global Σ).1 (trans_minductive_body mdecl) (trans_one_ind_body idecl) p (on_snd trans pdecl). Proof. intros (?&?&?). split;[|split]. @@ -447,14 +431,13 @@ Proof. repeat split;eassumption. Qed. - Lemma trans_instantiate_params params pars ty: option_map trans - (PT.instantiate_params + (ST.instantiate_params params pars ty) = TT.instantiate_params (trans_local params) (map trans pars) (trans ty). Proof. - rewrite TT.instantiate_params_, PT.instantiate_params_. + rewrite TT.instantiate_params_ ST.instantiate_params_. rewrite option_map_two. match goal with |- option_map _ (_ _ _ ?A _) = option_map _ (_ _ _ ?B _) => change B with (map trans A) @@ -473,28 +456,28 @@ Proof. 2: destruct (trans ty1);cbn;trivial. cbn. rewrite IHparams. cbn. now rewrite trans_subst. + destruct prim as [? []]; eauto. + destruct ty;cbn;trivial. 2: destruct (trans ty1);cbn;trivial. cbn. destruct pars;trivial. cbn. now rewrite IHparams. + destruct prim as [? []]; eauto. Qed. - -Lemma trans_instantiate_params' u mdecl idecl args npar x: -PT.instantiate_params +Lemma trans_instantiate_params' u mdecl args npar x ty : + ST.instantiate_params (PCUICUnivSubst.subst_instance_context u (PCUICAst.ind_params mdecl)) (firstn npar args) - (PCUICUnivSubst.subst_instance_constr u (PCUICAst.ind_type idecl)) = + ty = Some x -> TT.instantiate_params (subst_instance_context u (trans_local (PCUICAst.ind_params mdecl))) (firstn npar (map trans args)) - (subst_instance_constr u (trans (PCUICAst.ind_type idecl))) = + (trans ty) = Some (trans x). Proof. intros H. - rewrite <- trans_subst_instance_constr. rewrite firstn_map. match goal with |- TT.instantiate_params ?A _ _ = _ => @@ -524,8 +507,8 @@ Proof. Qed. Lemma trans_destr_arity x: -TT.destArity [] (trans x) -= option_map (fun '(xs,u) => (map trans_decl xs,u)) (PCUICAstUtils.destArity [] x). + AstUtils.destArity [] (trans x) = + option_map (fun '(xs,u) => (map trans_decl xs,u)) (PCUICAstUtils.destArity [] x). Proof. remember (@nil PCUICAst.context_decl) as xs. replace (@nil context_decl) with (map trans_decl xs) by (now subst). @@ -536,6 +519,7 @@ Proof. - rewrite <- IHx3. reflexivity. - destruct (trans x1);cbn;trivial. + - destruct prim as [? []]; eauto. Qed. Lemma trans_mkProd_or_LetIn a t: @@ -556,8 +540,278 @@ Proof. Qed. +Lemma trans_nth n l x : trans (nth n l x) = nth n (List.map trans l) (trans x). +Proof. + induction l in n |- *; destruct n; trivial. + simpl in *. congruence. +Qed. + +Lemma trans_iota_red pars c args brs : + trans (iota_red pars c args brs) = + TT.iota_red pars c (List.map trans args) (List.map (on_snd trans) brs). +Proof. + unfold iota_red, TT.iota_red. + rewrite trans_mkApps. + f_equal. induction brs in c |- *; simpl; destruct c; trivial. + now rewrite map_skipn. +Qed. + +Lemma trans_isLambda t : + T.isLambda (trans t) = S.isLambda t. +Proof. + destruct t; cbnr. + generalize (trans t1) (trans t2); clear. + induction t; intros; cbnr. + destruct prim as [? []]; cbnr. +Qed. + +Lemma trans_unfold_fix mfix idx narg fn : + unfold_fix mfix idx = Some (narg, fn) -> + TT.unfold_fix (map (map_def trans trans) mfix) idx = Some (narg, trans fn). +Proof. + unfold TT.unfold_fix, unfold_fix. + rewrite nth_error_map. destruct (nth_error mfix idx) eqn:Hdef => //. + cbn. + intros [= <- <-]. simpl. + repeat f_equal. + rewrite trans_subst. + f_equal. clear Hdef. simpl. + unfold fix_subst, TT.fix_subst. rewrite map_length. + generalize mfix at 1 3. + induction mfix; trivial. + simpl; intros mfix'. f_equal. + eapply IHmfix. +Qed. + +Lemma trans_unfold_cofix mfix idx narg fn : + unfold_cofix mfix idx = Some (narg, fn) -> + TT.unfold_cofix (map (map_def trans trans) mfix) idx = Some (narg, trans fn). +Proof. + unfold TT.unfold_cofix, unfold_cofix. + rewrite nth_error_map. destruct (nth_error mfix idx) eqn:Hdef => //. + intros [= <- <-]. simpl. repeat f_equal. + rewrite trans_subst. + f_equal. clear Hdef. + unfold cofix_subst, TT.cofix_subst. rewrite map_length. + generalize mfix at 1 3. + induction mfix; trivial. + simpl; intros mfix'. f_equal. + eapply IHmfix. +Qed. + +Lemma trans_is_constructor: + forall (args : list S.term) (narg : nat), + is_constructor narg args = true -> TT.is_constructor narg (map trans args) = true. +Proof. + intros args narg. + unfold TT.is_constructor, is_constructor. + rewrite nth_error_map //. destruct nth_error => //. simpl. intros. + unfold AstUtils.isConstruct_app, isConstruct_app, decompose_app in *. + destruct (decompose_app_rec t []) eqn:da. simpl in H. + destruct t0 => //. + apply decompose_app_rec_inv in da. simpl in da. subst t. + rewrite trans_mkApps /=. + rewrite decompose_app_mkApps //. +Qed. + +Lemma refine_red1_r Σ Γ t u u' : u = u' -> TT.red1 Σ Γ t u -> TT.red1 Σ Γ t u'. +Proof. + intros ->. trivial. +Qed. + +Lemma refine_red1_Γ Σ Γ Γ' t u : Γ = Γ' -> TT.red1 Σ Γ t u -> TT.red1 Σ Γ' t u. +Proof. + intros ->. trivial. +Qed. +Ltac wf_inv H := try apply wf_inv in H; simpl in H; rdest. + +Lemma trans_wf t : wf (trans t). +Proof. + induction t using term_forall_list_ind; simpl; try constructor; auto; + solve_all. + now eapply wf_mkApp. + destruct p as [? []]; constructor. +Qed. +Hint Resolve trans_wf : wf. + +Lemma trans_fix_context mfix: + map trans_decl (PCUICLiftSubst.fix_context mfix) = + TT.fix_context (map (map_def trans trans) mfix). +Proof. + unfold trans_local. + destruct mfix;trivial. + unfold TT.fix_context, PCUICLiftSubst.fix_context. + rewrite map_rev map_mapi. + cbn. f_equal. + 2: { + destruct d. cbn. + rewrite lift0_p. + rewrite SL.lift0_p. + unfold vass. + reflexivity. + } + f_equal. + unfold vass. + remember 1 as k. + induction mfix in k |- *;trivial. + cbn;f_equal. + - now rewrite trans_lift. + - apply IHmfix. +Qed. + +Lemma trans_red1 Σ Γ T U : + red1 Σ Γ T U -> + TT.red1 (trans_global_decls Σ) (trans_local Γ) (trans T) (trans U). +Proof. + induction 1 using red1_ind_all; simpl in *; + try solve [econstructor; eauto]. + + - simpl. + rewrite trans_subst; auto. constructor. + + - rewrite trans_subst; eauto. repeat constructor. + - rewrite trans_lift; eauto. repeat constructor. + rewrite nth_error_map. + destruct nth_error eqn:Heq => //. simpl in H. noconf H. + simpl. destruct c; noconf H => //. + + - rewrite trans_mkApps; eauto with wf; simpl. + erewrite trans_iota_red; eauto. repeat constructor. + + - simpl. rewrite !trans_mkApps /=. + unfold is_constructor in H0. + destruct nth_error eqn:hnth. + pose proof (nth_error_Some_length hnth). + destruct args. simpl. elimtype False; depelim H1. + cbn -[mkApps]. + eapply TT.red_fix. + apply trans_unfold_fix; eauto. + eapply (trans_is_constructor (t0 :: args)). + now rewrite /is_constructor hnth. + discriminate. + + - rewrite trans_mkApps. + rewrite !trans_mkApps; eauto with wf. + apply trans_unfold_cofix in H; eauto with wf. + eapply TT.red_cofix_case; eauto. + + - rewrite !trans_mkApps. + apply trans_unfold_cofix in H; eauto with wf. + eapply TT.red_cofix_proj; eauto. + + - rewrite trans_subst_instance_constr. econstructor. + apply (trans_declared_constant _ c decl H). + destruct decl. now simpl in *; subst cst_body. + + - rewrite trans_mkApps; eauto with wf. + simpl. constructor. now rewrite nth_error_map H. + + - constructor. solve_all. + apply OnOne2_map. + solve_all. red. simpl in *. + intuition eauto. + + - eapply (red1_mkApps_l _ _ _ _ [_]); auto with wf. + - eapply (red1_mkApps_r _ _ _ [_] [_]); auto with wf. + + - constructor. eapply OnOne2_map. solve_all. + + - constructor. eapply OnOne2_map. solve_all. + red. intuition auto. destruct x, y; simpl in *. noconf b. + reflexivity. + + - apply TT.fix_red_body. eapply OnOne2_map; solve_all. + red. destruct x, y; simpl in *. noconf b. + intuition auto. + rewrite /trans_local map_app in b0. + now rewrite -trans_fix_context. + + - constructor. eapply OnOne2_map. solve_all. + red. intuition auto. destruct x, y; simpl in *. noconf b. + reflexivity. + + - apply TT.cofix_red_body. eapply OnOne2_map; solve_all. + red. destruct x, y; simpl in *. noconf b. + intuition auto. + rewrite /trans_local map_app in b0. + now rewrite -trans_fix_context. +Qed. + +Lemma context_assumptions_map ctx : Ast.context_assumptions (map trans_decl ctx) = PCUICAst.context_assumptions ctx. +Proof. + induction ctx as [|[na [b|] ty] ctx]; simpl; auto. +Qed. + +Lemma trans_R_global_instance Σ Re Rle gref napp u u' : + PCUICEquality.R_global_instance Σ Re Rle gref napp u u' -> + R_global_instance (trans_global_decls Σ) Re Rle gref napp u u'. +Proof. + unfold PCUICEquality.R_global_instance, PCUICEquality.global_variance. + unfold R_global_instance, global_variance. + destruct gref; simpl; auto. + - unfold PCUICEquality.lookup_inductive, PCUICEquality.lookup_minductive. + unfold lookup_inductive, lookup_minductive. + rewrite trans_lookup. destruct PCUICAst.lookup_env => //; simpl. + destruct g => /= //. rewrite nth_error_map. + destruct nth_error => /= //. + rewrite trans_destr_arity. + destruct PCUICAstUtils.destArity as [[ctx ps]|] => /= //. + now rewrite context_assumptions_map. + - unfold PCUICEquality.lookup_constructor, PCUICEquality.lookup_inductive, PCUICEquality.lookup_minductive. + unfold lookup_constructor, lookup_inductive, lookup_minductive. + rewrite trans_lookup. destruct PCUICAst.lookup_env => //; simpl. + destruct g => /= //. rewrite nth_error_map. + destruct nth_error => /= //. + rewrite nth_error_map. + destruct nth_error => /= //. + destruct p as [[id t] nargs]; simpl. + destruct Nat.leb => //. +Qed. + +Lemma trans_eq_term_upto_univ {cf} : + forall Σ Re Rle t u napp, + PCUICEquality.eq_term_upto_univ_napp Σ Re Rle napp t u -> + eq_term_upto_univ_napp (trans_global_decls Σ) Re Rle napp (trans t) (trans u). +Proof. + intros Σ Re Rle t u napp e. + induction t using term_forall_list_ind in Rle, napp, u, e |- *. + all: invs e; cbn. + all: try solve [ constructor ; auto ]. + all: try solve [ + repeat constructor ; + match goal with + | ih : forall Rle u (napp : nat), _ |- _ => + eapply ih ; auto + end + ]. + all: try solve [ constructor; solve_all ]. + all: try solve [ constructor; now eapply trans_R_global_instance ]. + - eapply (eq_term_upto_univ_mkApps _ _ _ _ _ [_] _ [_]); simpl; eauto. + - destruct p as [? []]; constructor. +Qed. + +Lemma trans_leq_term {cf} Σ ϕ T U : + PCUICEquality.leq_term Σ ϕ T U -> + leq_term (trans_global_decls Σ) ϕ (trans T) (trans U). +Proof. + eapply trans_eq_term_upto_univ ; eauto. +Qed. + +Lemma trans_cumul {cf} Σ Γ T U : + cumul Σ Γ T U -> + TT.cumul (trans_global Σ) (trans_local Γ) (trans T) (trans U). +Proof. + induction 1. constructor; auto. + eapply trans_leq_term in l. + now rewrite -trans_global_ext_constraints. + apply trans_red1 in r; auto. econstructor 2; eauto. + econstructor 3. + apply IHX; auto. + apply trans_red1 in r; auto. +Qed. + Lemma trans_build_case_predicate_type ind mdecl idecl npar args u ps pty: -PT.build_case_predicate_type ind mdecl idecl (firstn npar args) u ps = Some pty -> +ST.build_case_predicate_type ind mdecl idecl (firstn npar args) u ps = Some pty -> TT.build_case_predicate_type ind (trans_minductive_body mdecl) (trans_one_ind_body idecl) (firstn npar (map trans args)) u ps = Some (trans pty). @@ -566,13 +820,13 @@ Proof. apply inv_opt_monad in H as (?&?&?&?&[=]). unfold TT.build_case_predicate_type. simpl in *. - apply trans_instantiate_params' in H as ->. + apply trans_instantiate_params' in H. + rewrite trans_subst_instance_constr in H. rewrite H. simpl. - rewrite trans_destr_arity, H0. + rewrite trans_destr_arity H0. simpl. f_equal. rewrite <- H2. - rewrite trans_it_mkProd_or_LetIn, - trans_mkProd_or_LetIn. + rewrite trans_it_mkProd_or_LetIn trans_mkProd_or_LetIn. f_equal. - now destruct x0. - f_equal. cbn. @@ -586,7 +840,7 @@ Proof. f_equal. + rewrite firstn_map. induction (firstn npar args);cbn;trivial. - now rewrite IHl0, map_length, trans_lift. + now rewrite IHl0 map_length trans_lift. + remember 0 as n;clear Heqn. remember (@nil PCUICAst.term) as xs. replace (@nil term) with (map trans xs) by (now subst). @@ -595,42 +849,16 @@ Proof. now destruct a as [? [] ?];rewrite <- IHl;cbn. Qed. +Definition TTwf_local {cf} Σ Γ := TT.All_local_env (TT.lift_typing TT.typing Σ) Γ. -Lemma trans_fix_context mfix: -map trans_decl (PCUICLiftSubst.fix_context mfix) = -TT.fix_context (map (map_def trans trans) mfix). -Proof. - unfold trans_local. - destruct mfix;trivial. - unfold TT.fix_context, PCUICLiftSubst.fix_context. - rewrite map_rev, map_mapi. - cbn. f_equal. - 2: { - destruct d. cbn. - rewrite lift0_p. - rewrite PL.lift0_p. - unfold vass. - reflexivity. - } - f_equal. - unfold vass. - remember 1 as k. - induction mfix in k |- *;trivial. - cbn;f_equal. - - now rewrite trans_lift. - - apply IHmfix. -Qed. - - - -Lemma trans_wf_local': - forall (Σ : PCUICAst.global_env_ext) Γ (wfΓ : PT.wf_local Σ Γ), +Lemma trans_wf_local' {cf} : + forall (Σ : PCUICAst.global_env_ext) Γ (wfΓ : wf_local Σ Γ), let P := (fun Σ0 Γ0 _ (t T : PCUICAst.term) _ => TT.typing (trans_global Σ0) (trans_local Γ0) (trans t) (trans T)) in - PT.All_local_env_over PT.typing P Σ Γ wfΓ -> - TT.wf_local (trans_global Σ) (trans_local Γ). + ST.All_local_env_over ST.typing P Σ Γ wfΓ -> + TTwf_local (trans_global Σ) (trans_local Γ). Proof. intros. induction X. @@ -641,13 +869,13 @@ Proof. - simpl. constructor; auto. red. destruct tu. exists x. auto. Qed. -Lemma trans_wf_local_env Σ Γ : - PT.All_local_env - (PT.lift_typing +Lemma trans_wf_local_env {cf} Σ Γ : + ST.All_local_env + (ST.lift_typing (fun (Σ : PCUICAst.global_env_ext) Γ b ty => - PT.typing Σ Γ b ty × TT.typing (trans_global Σ) (trans_local Γ) (trans b) (trans ty)) Σ) + ST.typing Σ Γ b ty × TT.typing (trans_global Σ) (trans_local Γ) (trans b) (trans ty)) Σ) Γ -> - TT.wf_local (trans_global Σ) (trans_local Γ). + TTwf_local (trans_global Σ) (trans_local Γ). Proof. intros. induction X. @@ -659,19 +887,19 @@ Proof. red. red in t1. destruct t1. eapply t2. Qed. -Lemma trans_branches Σ Γ brs btys ps: -All2 - (fun br bty : nat × PCUICAst.term => - (((br.1 = bty.1 × PT.typing Σ Γ br.2 bty.2) - × TT.typing (trans_global Σ) (trans_local Γ) (trans br.2) (trans bty.2)) - × PT.typing Σ Γ bty.2 (PCUICAst.tSort ps)) - × TT.typing (trans_global Σ) (trans_local Γ) (trans bty.2) - (trans (PCUICAst.tSort ps))) brs btys -> -All2 - (fun br bty : nat × term => - (br.1 = bty.1 × TT.typing (trans_global Σ) (trans_local Γ) br.2 bty.2) - × TT.typing (trans_global Σ) (trans_local Γ) bty.2 (tSort ps)) - (map (on_snd trans) brs) (map (fun '(x, y) => (x, trans y)) btys). +Lemma trans_branches {cf} Σ Γ brs btys ps: + All2 + (fun br bty : nat × PCUICAst.term => + (((br.1 = bty.1 × ST.typing Σ Γ br.2 bty.2) + × TT.typing (trans_global Σ) (trans_local Γ) (trans br.2) (trans bty.2)) + × ST.typing Σ Γ bty.2 (PCUICAst.tSort ps)) + × TT.typing (trans_global Σ) (trans_local Γ) (trans bty.2) + (trans (PCUICAst.tSort ps))) brs btys -> + All2 + (fun br bty : nat × term => + (br.1 = bty.1 × TT.typing (trans_global Σ) (trans_local Γ) br.2 bty.2) + × TT.typing (trans_global Σ) (trans_local Γ) bty.2 (tSort ps)) + (map (on_snd trans) brs) (map (fun '(x, y) => (x, trans y)) btys). Proof. induction 1;cbn. - constructor. @@ -683,23 +911,21 @@ Proof. + apply IHX. Qed. - - -Lemma trans_mfix_All Σ Γ mfix: -PT.All_local_env - (PT.lift_typing - (fun (Σ : PCUICEnvironment.global_env_ext) - (Γ : PCUICEnvironment.context) (b ty : PCUICAst.term) => - PT.typing Σ Γ b ty - × TT.typing (trans_global Σ) (trans_local Γ) (trans b) (trans ty)) Σ) - (PCUICAst.app_context Γ (PCUICLiftSubst.fix_context mfix)) -> -TT.wf_local (trans_global Σ) - (trans_local Γ ,,, TT.fix_context (map (map_def trans trans) mfix)). +Lemma trans_mfix_All {cf} Σ Γ mfix: + ST.All_local_env + (ST.lift_typing + (fun (Σ : PCUICEnvironment.global_env_ext) + (Γ : PCUICEnvironment.context) (b ty : PCUICAst.term) => + ST.typing Σ Γ b ty + × TT.typing (trans_global Σ) (trans_local Γ) (trans b) (trans ty)) Σ) + (PCUICAst.app_context Γ (PCUICLiftSubst.fix_context mfix)) -> + TTwf_local (trans_global Σ) + (trans_local Γ ,,, TT.fix_context (map (map_def trans trans) mfix)). Proof. intros. rewrite <- trans_fix_context. match goal with - |- TT.wf_local _ ?A => + |- TTwf_local _ ?A => replace A with (trans_local (PCUICAst.app_context Γ (PCUICLiftSubst.fix_context mfix))) end. 2: { @@ -717,53 +943,52 @@ Proof. assumption. Qed. -Lemma trans_mfix_All2 Σ Γ mfix xfix: -All - (fun d : def PCUICAst.term => - (PT.typing Σ (PCUICAst.app_context Γ (PCUICLiftSubst.fix_context xfix)) - (dbody d) - (PCUICLiftSubst.lift0 #|PCUICLiftSubst.fix_context xfix| (dtype d)) - × PCUICAst.isLambda (dbody d) = true) - × TT.typing (trans_global Σ) - (trans_local (PCUICAst.app_context Γ (PCUICLiftSubst.fix_context xfix))) - (trans (dbody d)) - (trans - (PCUICLiftSubst.lift0 #|PCUICLiftSubst.fix_context xfix| - (dtype d)))) mfix -> -All - (fun d : def term => - TT.typing (trans_global Σ) - (trans_local Γ ,,, TT.fix_context (map (map_def trans trans) xfix)) - (dbody d) (TL.lift0 #|TT.fix_context (map (map_def trans trans) xfix)| (dtype d)) - × isLambda (dbody d) = true) (map (map_def trans trans) mfix). +Lemma trans_mfix_All2 {cf} Σ Γ mfix xfix: + All + (fun d : def PCUICAst.term => + (ST.typing Σ (PCUICAst.app_context Γ (PCUICLiftSubst.fix_context xfix)) + (dbody d) + (PCUICLiftSubst.lift0 #|PCUICLiftSubst.fix_context xfix| (dtype d))) + × TT.typing (trans_global Σ) + (trans_local (PCUICAst.app_context Γ (PCUICLiftSubst.fix_context xfix))) + (trans (dbody d)) + (trans + (PCUICLiftSubst.lift0 #|PCUICLiftSubst.fix_context xfix| + (dtype d)))) mfix -> + All + (fun d : def term => + TT.typing (trans_global Σ) + (trans_local Γ ,,, TT.fix_context (map (map_def trans trans) xfix)) + (dbody d) (TL.lift0 #|TT.fix_context (map (map_def trans trans) xfix)| (dtype d))) + (map (map_def trans trans) mfix). Proof. induction 1. - constructor. - - constructor. - + destruct p as [[]]. + - simpl; constructor. + + destruct p as []. unfold app_context, PCUICAst.app_context in *. - split. - * unfold trans_local in t0. - rewrite map_app, trans_fix_context in t0. - rewrite trans_dbody, trans_lift, trans_dtype in t0. - replace(#|PCUICLiftSubst.fix_context xfix|) with - (#|TT.fix_context (map(map_def trans trans) xfix)|) in t0. - 2: admit. -Admitted. - - -Lemma All_over_All Σ Γ wfΓ : -PT.All_local_env_over PT.typing - (fun (Σ : PCUICAst.global_env_ext) (Γ : PCUICAst.context) - (_ : PCUICTyping.wf_local Σ Γ) (t T : PCUICAst.term) - (_ : PT.typing Σ Γ t T) => - TT.typing (trans_global Σ) (trans_local Γ) (trans t) (trans T)) Σ Γ wfΓ -> -PT.All_local_env - (PT.lift_typing - (fun (Σ0 : PCUICAst.global_env_ext) (Γ0 : PCUICEnvironment.context) - (b ty : PCUICAst.term) => - PT.typing Σ0 Γ0 b ty - × TT.typing (trans_global Σ0) (trans_local Γ0) (trans b) (trans ty)) Σ) Γ. + unfold trans_local in t0. + rewrite map_app trans_fix_context in t0. + rewrite trans_dbody trans_lift trans_dtype in t0. + replace(#|PCUICLiftSubst.fix_context xfix|) with + (#|TT.fix_context (map(map_def trans trans) xfix)|) in t0. + 2:now rewrite TT.fix_context_length map_length fix_context_length. + now destruct x. + + auto. +Qed. + +Lemma All_over_All {cf} Σ Γ wfΓ : + ST.All_local_env_over ST.typing + (fun (Σ : PCUICAst.global_env_ext) (Γ : PCUICAst.context) + (_ : wf_local Σ Γ) (t T : PCUICAst.term) + (_ : ST.typing Σ Γ t T) => + TT.typing (trans_global Σ) (trans_local Γ) (trans t) (trans T)) Σ Γ wfΓ -> + ST.All_local_env + (ST.lift_typing + (fun (Σ0 : PCUICAst.global_env_ext) (Γ0 : PCUICEnvironment.context) + (b ty : PCUICAst.term) => + ST.typing Σ0 Γ0 b ty + × TT.typing (trans_global Σ0) (trans_local Γ0) (trans b) (trans ty)) Σ) Γ. Proof. intros H. induction H;cbn. @@ -782,128 +1007,551 @@ Proof. split;eassumption. Qed. +Lemma trans_decompose_prod_assum : + forall Γ t, + let '(Δ, c) := decompose_prod_assum Γ t in + AstUtils.decompose_prod_assum (trans_local Γ) (trans t) = (trans_local Δ, trans c). +Proof. + intros Γ t. + destruct decompose_prod_assum as [Δ c] eqn:e. + induction t in Γ, Δ, c, e |- *. + all: simpl in *. + all: try solve [ inversion e ; subst ; reflexivity ]. + - eapply IHt2 in e as e'. apply e'. + - eapply IHt3 in e as e'. assumption. + - noconf e. simpl. + now destruct (mkApp_ex (trans t1) (trans t2)) as [f [args ->]]. + - noconf e. now destruct prim as [? []] => /=. +Qed. -Lemma invert_type_App `{checker_flags} Σ Γ f u T : - TT.typing Σ Γ (tApp f u) T -> - { T' : term & { U' & ((TT.typing Σ Γ f T') * TT.typing_spine Σ Γ T' u U' * - (isApp f = false) * (u <> []) * - (TT.cumul Σ Γ U' T))%type } }. +Lemma trans_subst_context s k Γ : + trans_local (SL.subst_context s k Γ) = TL.subst_context (map trans s) k (trans_local Γ). Proof. - intros Hty. - dependent induction Hty. - - exists t_ty, t'. intuition. - - edestruct IHHty as [T' [U' [H' H'']]]. - exists T', U'. split; auto. - eapply TT.cumul_trans; eauto. + induction Γ as [|[na [b|] ty] Γ]; simpl; auto. + - rewrite SL.subst_context_snoc /=. rewrite [TL.subst_context _ _ _ ]subst_context_snoc. + f_equal; auto. rewrite IHΓ /snoc /subst_decl /map_decl /=; f_equal. + now rewrite !trans_subst map_length. + - rewrite SL.subst_context_snoc /=. rewrite [TL.subst_context _ _ _ ]subst_context_snoc. + f_equal; auto. rewrite IHΓ /snoc /subst_decl /map_decl /=; f_equal. + now rewrite !trans_subst map_length. Qed. -Lemma type_mkApp `{checker_flags} Σ Γ f u T T' : - TT.typing Σ Γ f T -> - TT.typing_spine Σ Γ T [u] T' -> - TT.typing Σ Γ (mkApp f u) T'. -Proof. - intros Hf Hu. inv Hu. - simpl. destruct (isApp f) eqn:Happ. - destruct f; try discriminate. simpl. - eapply invert_type_App in Hf. - destruct Hf as (T'' & U' & (((Hf & HU) & Happf) & Hunil) & Hcumul). - eapply TT.type_App; eauto. intro. destruct args; discriminate. - inv X2. clear Happ Hf Hunil. - induction HU. simpl. econstructor; eauto. - eapply TT.cumul_trans; eauto. - econstructor. econstructor. eapply t. eauto. eauto. - apply IHHU; eauto. - rewrite -> AstUtils.mkApp_tApp; eauto. - econstructor; eauto. congruence. - econstructor; eauto. +Lemma trans_smash_context Γ Δ : trans_local (smash_context Γ Δ) = TT.smash_context (trans_local Γ) (trans_local Δ). +Proof. + induction Δ in Γ |- *; simpl; auto. + destruct a as [na [b|] ty] => /=. + rewrite IHΔ. f_equal. + now rewrite (trans_subst_context [_]). + rewrite IHΔ. f_equal. rewrite /trans_local map_app //. Qed. +Lemma trans_isApp t : PCUICAst.isApp t = false -> Ast.isApp (trans t) = false. +Proof. + destruct t => //. + now destruct prim as [? []]. +Qed. -(* from Checker Substitution *) +Lemma trans_nisApp t : ~~ PCUICAst.isApp t -> ~~ Ast.isApp (trans t). +Proof. + destruct t => //. + now destruct prim as [? []]. +Qed. +Lemma trans_destInd t : ST.destInd t = TT.destInd (trans t). +Proof. + destruct t => //. simpl. + now destruct (mkApp_ex (trans t1) (trans t2)) as [f [u ->]]. + now destruct prim as [? []]. +Qed. +Lemma trans_decompose_app t : + let '(hd, args) := decompose_app t in + AstUtils.decompose_app (trans t) = (trans hd, map trans args). +Proof. + destruct (decompose_app t) eqn:da. + pose proof (decompose_app_notApp _ _ _ da). + eapply decompose_app_rec_inv in da. simpl in da. subst t. + rewrite trans_mkApps. + apply AstUtils.decompose_app_mkApps. + now rewrite trans_isApp. +Qed. -Lemma eq_term_upto_univ_refl Re Rle : - RelationClasses.Reflexive Re -> - RelationClasses.Reflexive Rle -> - forall t, TT.eq_term_upto_univ Re Rle t t. +Lemma All2_map_option_out_mapi_Some_spec : + forall {A B B'} (f : nat -> A -> option B) (g' : B -> B') + (g : nat -> A -> option B') l t, + (forall i x t, f i x = Some t -> g i x = Some (g' t)) -> + map_option_out (mapi f l) = Some t -> + map_option_out (mapi g l) = Some (map g' t). Proof. - intros hRe hRle. - induction t in Rle, hRle |- * using Induction.term_forall_list_rect; simpl; - try constructor; try apply Forall_Forall2; try apply All_All2 ; try easy; - try now (try apply Forall_All ; apply Forall_True). - - eapply All_All2. 1: eassumption. - intros. easy. - - eapply All_All2. 1: eassumption. - intros. easy. - - destruct p. constructor; try easy. - eapply All_All2. 1: eassumption. - intros. split ; auto. - - eapply All_All2. 1: eassumption. - intros x [? ?]. repeat split ; auto. - - eapply All_All2. 1: eassumption. - intros x [? ?]. repeat split ; auto. + intros A B B' f g' g l t. + unfold mapi. generalize 0. + intros n h e. + induction l in n, e, t |- *. + - simpl in *. apply some_inj in e. subst. reflexivity. + - simpl in *. + destruct (f n a) eqn:e'. 2: discriminate. + eapply h in e' as h'. + rewrite h'. + match type of e with + | context [ map_option_out ?t ] => + destruct (map_option_out t) eqn:e'' + end. 2: discriminate. + eapply IHl in e''. rewrite e''. now noconf e. Qed. -Lemma leq_term_refl `{checker_flags} φ t : TT.leq_term φ t t. +Lemma trans_inds ind u mdecl : + map trans (ST.inds (inductive_mind ind) u (PCUICAst.ind_bodies mdecl)) = + TT.inds (inductive_mind ind) u (ind_bodies (trans_minductive_body mdecl)). Proof. - apply eq_term_upto_univ_refl. - - intro; apply eq_universe_refl. - - intro; apply leq_universe_refl. + rewrite ST.inds_spec TT.inds_spec. + rewrite map_rev map_mapi. simpl. + f_equal. rewrite mapi_map. apply mapi_ext => n //. Qed. +Lemma trans_reln l p Γ : map trans (PCUICAst.reln l p Γ) = + reln (map trans l) p (trans_local Γ). +Proof. + induction Γ as [|[na [b|] ty] Γ] in l, p |- *; simpl; auto. + now rewrite IHΓ. +Qed. +Lemma trans_to_extended_list Γ n : map trans (PCUICAst.to_extended_list_k Γ n) = to_extended_list_k (trans_local Γ) n. +Proof. + now rewrite /to_extended_list_k trans_reln. +Qed. +Lemma trans_build_branches_type ind mdecl idecl npar args u p brtys : + map_option_out (ST.build_branches_type ind mdecl idecl (firstn npar args) u p) + = Some brtys -> + map_option_out + (TT.build_branches_type ind (trans_minductive_body mdecl) + (trans_one_ind_body idecl) (firstn npar (map trans args)) u (trans p)) + = Some (map (on_snd trans) brtys). +Proof. + unfold ST.build_branches_type. + unfold TT.build_branches_type. + rewrite mapi_map. + apply All2_map_option_out_mapi_Some_spec. + intros i [[id ty] argn] [nargs br]. + simpl. + destruct ST.instantiate_params eqn:ipars => //. + apply trans_instantiate_params' in ipars. + rewrite trans_subst in ipars. + rewrite -trans_inds. + rewrite -trans_subst_instance_constr ipars. + move: (trans_decompose_prod_assum [] t). + destruct decompose_prod_assum => -> /=. + move: (trans_decompose_app t0). + destruct decompose_app => /= // ->. + simpl. destruct chop as [params argsl] eqn:ceq. + rewrite (chop_map _ _ _ _ _ ceq). + intros [= ->]. f_equal. unfold on_snd. simpl. f_equal. subst br. + rewrite trans_it_mkProd_or_LetIn trans_mkApps map_app /= trans_mkApps /= map_app. + now rewrite trans_to_extended_list trans_lift map_length. +Qed. +Lemma strip_casts_trans t : AstUtils.strip_casts (trans t) = trans t. +Proof. + induction t using term_forall_list_ind; simpl; auto; + rewrite ?map_map_compose; + f_equal; solve_all. + + - rewrite strip_casts_mkApp_wf; auto with wf. + now rewrite IHt1 IHt2. + + - now destruct p as [? []]. +Qed. +Lemma trans_check_one_fix mfix : + TT.check_one_fix (map_def trans trans mfix) = ST.check_one_fix mfix. +Proof. + unfold ST.check_one_fix, TT.check_one_fix. + destruct mfix as [na ty arg rarg] => /=. + move: (trans_decompose_prod_assum [] ty). + destruct decompose_prod_assum as [ctx p] => /= ->. + rewrite -(trans_smash_context []) /trans_local. + rewrite -List.map_rev nth_error_map. + destruct nth_error => /= //. + move: (trans_decompose_app (PCUICAst.decl_type c)). + destruct decompose_app => /=. + simpl. destruct c. simpl. intros ->. + now rewrite -trans_destInd. +Qed. -Theorem template_to_pcuic (Σ : P.global_env_ext) Γ t T : - PT.wf Σ -> - PT.typing Σ Γ t T -> - TT.typing (trans_global Σ) (trans_local Γ) (trans t) (trans T). +Lemma map_option_out_check_one_fix mfix : + map (fun x => TT.check_one_fix (map_def trans trans x)) mfix = + map ST.check_one_fix mfix. Proof. + eapply map_ext => x. apply trans_check_one_fix. +Qed. +Lemma trans_check_one_cofix mfix : + TT.check_one_cofix (map_def trans trans mfix) = ST.check_one_cofix mfix. +Proof. + unfold ST.check_one_cofix, TT.check_one_cofix. + destruct mfix as [na ty arg rarg] => /=. + move: (trans_decompose_prod_assum [] ty). + destruct decompose_prod_assum as [ctx p] => -> /=. + move: (trans_decompose_app p). + destruct decompose_app => /= ->. + now rewrite -trans_destInd. +Qed. + +Lemma map_option_out_check_one_cofix mfix : + map (fun x => TT.check_one_cofix (map_def trans trans x)) mfix = + map ST.check_one_cofix mfix. +Proof. + eapply map_ext => x. apply trans_check_one_cofix. +Qed. + +Lemma trans_check_rec_kind Σ k f : + ST.check_recursivity_kind Σ k f = TT.check_recursivity_kind (trans_global_decls Σ) k f. +Proof. + unfold ST.check_recursivity_kind, TT.check_recursivity_kind. + rewrite trans_lookup. + destruct S.lookup_env as [[]|] => //. +Qed. + +Lemma trans_wf_fixpoint Σ mfix : + TT.wf_fixpoint (trans_global_decls Σ) (map (map_def trans trans) mfix) = + ST.wf_fixpoint Σ mfix. +Proof. + unfold ST.wf_fixpoint, TT.wf_fixpoint. + rewrite map_map_compose. + rewrite map_option_out_check_one_fix. + destruct map_option_out as [[]|] => //. + now rewrite trans_check_rec_kind. +Qed. + +Lemma trans_wf_cofixpoint Σ mfix : + TT.wf_cofixpoint (trans_global_decls Σ) (map (map_def trans trans) mfix) = + ST.wf_cofixpoint Σ mfix. +Proof. + unfold ST.wf_cofixpoint, TT.wf_cofixpoint. + rewrite map_map_compose. + rewrite map_option_out_check_one_cofix. + destruct map_option_out as [[]|] => //. + now rewrite trans_check_rec_kind. +Qed. + +Lemma type_mkApps_napp `{checker_flags} Σ Γ f u T T' : + ~~ isApp f -> + TT.typing Σ Γ f T -> + TT.typing_spine Σ Γ T u T' -> + TT.typing Σ Γ (mkApps f u) T'. +Proof. + intros hf hty hsp. + depelim hsp. simpl; auto. + rewrite mkApps_tApp. + destruct f => //. congruence. + econstructor; eauto. + destruct f => //. congruence. + econstructor; eauto. +Qed. + +Axiom fix_guard_trans : + forall Σ Γ mfix, + ST.fix_guard Σ Γ mfix -> + TT.fix_guard (trans_global Σ) (trans_local Γ) (map (map_def trans trans) mfix). + +Axiom cofix_guard_trans : + forall Σ Γ mfix, + ST.cofix_guard Σ Γ mfix -> + TT.cofix_guard (trans_global Σ) (trans_local Γ) (map (map_def trans trans) mfix). + +(* This version of typing spine allows to not require [isType] assumptions at the end of the + application chain, which is crucial to be able to build a spine that both encompasses PCUIC's + applications and mimicks the typing rule of Template. Otherwise we would only model: + + |- tProd na A B : s + |- f : tProd na A B + |- u : A + *** |- B [0 := u] : s *** <- this last hypothesis is problematic, it's not present in PCUIC + ------------------------ derivations for example. + |- tApp f u : B [0 := u] + + The typing_spine_nil constructor allows to stack cumulativity steps using transitivity of cumulativity, + following from confluence. + + In the end a typing derivation for an application that can mix application and cumulativity steps at any + point: + + tApp (tApp f u) v : T' can be translated to + + typing_spine fty [u; v] T' +*) + +Import PCUICAst PCUICLiftSubst PCUICTyping. +Inductive typing_spine {cf} (Σ : global_env_ext) (Γ : context) : term -> list term -> term -> Type := +| type_spine_eq ty : typing_spine Σ Γ ty [] ty + +| type_spine_nil ty ty' : + isType Σ Γ ty' -> + Σ ;;; Γ |- ty <= ty' -> + typing_spine Σ Γ ty [] ty' + +| type_spine_cons hd tl na A B T B' : + isType Σ Γ (tProd na A B) -> + Σ ;;; Γ |- T <= tProd na A B -> + Σ ;;; Γ |- hd : A -> + typing_spine Σ Γ (subst10 hd B) tl B' -> + typing_spine Σ Γ T (hd :: tl) B'. +Derive Signature for typing_spine. + +(* The size of a spine is the maximal depth of any typing derivation appearing in it *) +Section Typing_Spine_size. + Context `{checker_flags}. + Context {Σ : global_env_ext} {Γ : context}. + + Fixpoint typing_spine_size {t T U} (s : typing_spine Σ Γ t T U) : size := + match s with + | type_spine_eq ty => 0 + | type_spine_nil ty ty' ist cum => typing_size ist.π2 + | type_spine_cons hd tl na A B T B' typrod cumul ty s' => + (max (typing_size typrod.π2) (max (typing_size ty) (typing_spine_size s')))%nat + end. +End Typing_Spine_size. + +(* We can weaken a spine by cumulativity, using the type_spine_nil *) +Lemma typing_spine_weaken_concl {cf:checker_flags} {Σ Γ T args S S'} : + wf Σ.1 -> + typing_spine Σ Γ T args S -> + Σ ;;; Γ |- S <= S' -> + isType Σ Γ S' -> + typing_spine Σ Γ T args S'. +Proof. + intros wfΣ sp. + induction sp in S' => cum. constructor; auto. constructor; auto. + now transitivity ty'. + intros isType. econstructor. eauto. eauto. eauto. + eauto. +Defined. + +(* The size includes the new [isType] hypothesis *) +Lemma typing_spine_weaken_concl_size {cf:checker_flags} {Σ Γ T args S S'} + (wf : wf Σ.1) + (sp :typing_spine Σ Γ T args S) + (Hs : Σ ;;; Γ |- S <= S') + (ist : isType Σ Γ S') : + typing_spine_size (typing_spine_weaken_concl wf sp Hs ist) <= + max (typing_spine_size sp) (typing_size ist.π2). +Proof. + induction sp; simpl; auto. lia. + rewrite - !Nat.max_assoc. specialize (IHsp Hs). lia. +Qed. + +Ltac sig := unshelve eexists. + +(** We can also append an argument to a spine *without* requiring a typing for the new conclusion type. + This would be a derivation for a substitution of [B], hence incomparable in depth to the arguments. + The invariant is witnessed by the spine size being lower than the arguments. +*) +Lemma typing_spine_app {cf:checker_flags} Σ Γ ty args na A B arg + (wf : wf Σ.1) + (isty : isType Σ Γ (tProd na A B)) + (sp : typing_spine Σ Γ ty args (tProd na A B)) + (argd : Σ ;;; Γ |- arg : A) : + ∑ sp' : typing_spine Σ Γ ty (args ++ [arg]) (B {0 := arg}), + typing_spine_size sp' <= max (typing_size isty.π2) (max (typing_spine_size sp) (typing_size argd)). +Proof. + revert arg argd. + depind sp. + - intros arg Harg. simpl. + sig. + * econstructor; eauto. apply cumul_refl'. + constructor. + * simpl. lia. + - intros arg Harg. simpl. + sig. + * econstructor; eauto. constructor. + * simpl. lia. + - intros arg Harg. simpl. + specialize (IHsp wf isty _ Harg) as [sp' Hsp']. + sig. + * econstructor. apply i. auto. auto. exact sp'. + * simpl. lia. +Qed. + +(** Likewise, in Template-Coq, we can append without re-typing the result *) +Lemma TT_typing_spine_app {cf:checker_flags} Σ Γ ty T' args na A B arg s : + TT.typing Σ Γ (T.tProd na A B) (T.tSort s) -> + TT.typing_spine Σ Γ ty args T' -> + TT.cumul Σ Γ T' (T.tProd na A B) -> + TT.typing Σ Γ arg A -> + TT.typing_spine Σ Γ ty (args ++ [arg]) (TL.subst1 arg 0 B). +Proof. + intros isty H; revert arg. + remember (T.tProd na A B) as prod. + revert na A B Heqprod. + induction H. + - intros na A B eq arg Harg. simpl. econstructor; eauto. now rewrite <-eq. rewrite <-eq. + apply Harg. + constructor. + - intros na' A' B'' eq arg Harg. simpl. + econstructor. apply t. auto. auto. + eapply IHtyping_spine. now rewrite eq. exact eq. apply Harg. apply X. +Defined. + +(** This is the central lemma for this translation: + any PCUIC typing of an application `t` can be decomposed as + a typing of the head of `t` followed by a typing spine for + its arguments `l`. *) + +Lemma type_app {cf} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : + wf Σ.1 -> + if isApp t then + let (f, l) := decompose_app t in + ∑ fty (d' : Σ ;;; Γ |- f : fty), + ((typing_size d' <= typing_size d) * + ∑ (sp : typing_spine Σ Γ fty l T), + (typing_spine_size sp <= typing_size d))%type + else True. +Proof. + intros wfΣ. + induction d; simpl; auto. + - destruct (isApp t) eqn:isapp. + destruct (decompose_app (tApp t u)) eqn:da. + unfold decompose_app in da. + simpl in da. + apply decompose_app_rec_inv' in da as [n [napp [equ eqt]]]. + subst t. clear IHd1 IHd3. + rewrite PCUICAstUtils.decompose_app_mkApps // in IHd2. + destruct IHd2 as [fty [d' [sp hd']]]. + exists fty, d'. + split. lia. + destruct hd' as [sp' Hsp]. + destruct (typing_spine_app _ _ _ _ _ _ _ _ wfΣ (s; d1) sp' d3) as [appsp Happ]. + simpl in Happ. + move: appsp Happ. rewrite equ firstn_skipn. + intros app happ. exists app. lia. + + unfold decompose_app. simpl. + rewrite PCUICAstUtils.atom_decompose_app. destruct t => /= //. + eexists _, d2. split. lia. + unshelve eexists. + econstructor. exists s; eauto. reflexivity. assumption. constructor. + simpl. lia. + + - destruct (isApp t) eqn:isapp => //. + destruct (decompose_app t) eqn:da. + assert (l <> []). + { eapply decompose_app_nonnil in da => //. } + destruct IHd1 as [fty [d' [? ?]]]. + exists fty, d'. split. lia. + destruct s0 as [sp Hsp]. + unshelve eexists. eapply typing_spine_weaken_concl; eauto. now exists s. + epose proof (typing_spine_weaken_concl_size wfΣ sp c (s; d2)). simpl in H0. lia. +Qed. + +(** Finally, for each typing spine built above, assuming we can apply the induction hypothesis + of the translation to any of the typing derivations in the spine, then we can produce a typing + spine in the n-ary application template-coq spec. + + We have two cases to consider at the "end" of the spine: either we have the same translation of the + PCUIC conclusion type, or there is exactly one cumulativity step to get to this type. *) +Lemma make_typing_spine {cf} {Σ Γ fty l T} + (sp : typing_spine Σ Γ fty l T) n + (IH : forall t' T' (Ht' : Σ ;;; Γ |- t' : T'), + typing_size Ht' <= n -> + TT.typing (trans_global Σ) (trans_local Γ) (trans t') (trans T')) : + typing_spine_size sp <= n -> + ∑ T', TT.typing_spine (trans_global Σ) (trans_local Γ) (trans fty) (map trans l) T' * + ((T' = trans T) + + (TT.cumul (trans_global Σ) (trans_local Γ) T' (trans T) * + ∑ s, TT.typing (trans_global Σ) (trans_local Γ) (trans T) (T.tSort s)))%type. +Proof. + induction sp; simpl; intros. + - exists (trans ty); split. + * constructor. + * left; auto. + - exists (trans ty). + split. constructor. right. + split. + now eapply trans_cumul in c. + specialize (IH _ _ i.π2 H). now exists i.π1. + + - simpl; intros. + forward IHsp by lia. + specialize (IH _ _ t) as IHt. + forward IHt by lia. + apply trans_cumul in c. + specialize (IH _ _ i.π2) as IHi. + forward IHi by lia. + simpl in IHi. + destruct IHsp as [T' [Hsp eq]]. + destruct eq. subst T'. + exists (trans B'). split; auto. + econstructor. + eauto. + apply c. + apply IHt. + now rewrite trans_subst in Hsp. + exists T'. split; auto. + econstructor. + eauto. apply c. apply IHt. + now rewrite trans_subst in Hsp. +Qed. + +Theorem pcuic_to_template {cf} (Σ : S.global_env_ext) Γ t T : + ST.wf Σ -> + ST.typing Σ Γ t T -> + TT.typing (trans_global Σ) (trans_local Γ) (trans t) (trans T). +Proof. intros X X0. revert Σ X Γ t T X0. - apply (PT.typing_ind_env (fun Σ Γ t T => + (** We use an induction principle allowing to apply induction to any subderivation of + functions in applications. *) + apply (typing_ind_env_app_size (fun Σ Γ t T => TT.typing (trans_global Σ) (trans_local Γ) (trans t) (trans T) )%type (fun Σ Γ wfΓ => TT.All_local_env (TT.lift_typing TT.typing (trans_global Σ)) (trans_local Γ)) );intros. - now eapply trans_wf_local_env, All_over_All. -(* -typing_ind_env is like induction but with -additional assumptions for the nested typing -in All (for wf_local assumptions) -*) - rewrite trans_lift. rewrite trans_decl_type. eapply TT.type_Rel; eauto. + now apply map_nth_error. - - admit. (* TODO adapt Template Coq to new sort typing rule *) - (* eapply TT.type_Sort; eauto. - now apply trans_in_level_set. *) + - econstructor; eauto. + destruct u; auto. simpl in H |- *. + intros l inl. now rewrite <-trans_global_ext_levels. - eapply TT.type_Prod;assumption. - eapply TT.type_Lambda;eassumption. - eapply TT.type_LetIn;eassumption. - - eapply type_mkApp. - + eassumption. - + econstructor. - 3: eassumption. - 3: { - rewrite trans_subst10. - constructor. - } - 2: { - cbn. constructor. - apply leq_term_refl. - } - admit. - (* - tProd na (trans A) (trans B) has a type - obtain with inversion from IHX0_1 - *) + - simpl. rewrite trans_subst10. + destruct (isApp t) eqn:isapp. + move: (type_app Ht wfΣ). rewrite isapp. + destruct (decompose_app t) eqn:da. + intros [fty [d' [hd' [sp hsp]]]]. + assert (IHt := X3 _ _ d' hd'). + epose proof (make_typing_spine sp (typing_size Ht) X3 hsp) as [T' [IH []]]. + * subst T'. rewrite (PCUICAstUtils.decompose_app_inv da). + rewrite trans_mkApps mkApp_mkApps AstUtils.mkApps_nested. + pose proof (AstUtils.mkApps_nested (trans t0) (map trans l) [trans u]). + apply decompose_app_notApp in da. + apply trans_isApp in da. + eapply type_mkApps_napp. rewrite da //. + eassumption. + eapply TT_typing_spine_app. simpl in X1. eapply X1. + apply IH. apply TT.cumul_refl'. + apply X5. + * destruct p as [hcum _]. + rewrite (PCUICAstUtils.decompose_app_inv da). + rewrite trans_mkApps mkApp_mkApps AstUtils.mkApps_nested. + pose proof (AstUtils.mkApps_nested (trans t0) (map trans l) [trans u]). + apply decompose_app_notApp in da. + apply trans_isApp in da. + eapply type_mkApps_napp. rewrite da //. + eassumption. + eapply TT_typing_spine_app. simpl in X1. eapply X1. + apply IH. apply hcum. + apply X5. + * rewrite mkApp_mkApps. + eapply type_mkApps_napp. + apply trans_isApp in isapp. rewrite isapp //. + now simpl in X2. + econstructor. simpl in X1. eapply X1. + apply TT.cumul_refl'. assumption. constructor. - rewrite trans_subst_instance_constr. rewrite trans_cst_type. eapply TT.type_Const; eauto. @@ -921,8 +1569,8 @@ in All (for wf_local assumptions) * now apply trans_declared_inductive. * now apply map_nth_error. + now apply trans_consistent_instance_ext. - - replace (trans (PCUICAst.mkApps p (skipn npar args ++ [c]))) - with (mkApps (trans p) (skipn npar (map trans args) ++ [trans c])). + - replace (trans (mkApps p (skipn npar args ++ [c]))) + with (Ast.mkApps (trans p) (skipn npar (map trans args) ++ [trans c])). 2: { rewrite trans_mkApps. rewrite map_app. @@ -930,6 +1578,7 @@ in All (for wf_local assumptions) repeat f_equal. now rewrite map_skipn. } + simpl. eapply TT.type_Case. + now apply trans_declared_inductive. + cbn. assumption. @@ -938,12 +1587,11 @@ in All (for wf_local assumptions) + eassumption. + rewrite <- trans_global_ext_constraints. eassumption. - + rewrite trans_mkApps in X4. - cbn in X4. - apply X4. - + admit. (* map_option_out build branche type *) - (* this should be similar to trans_build_case_predicate_type *) - + admit. (* now apply trans_branches.*) + + auto. + + now rewrite trans_mkApps in X4. + + now apply trans_build_branches_type in H3. + + eapply All2_map. solve_all. + destruct b0 as [s [Hs Hy]]. exists s; eauto. - rewrite trans_subst. rewrite trans_subst_instance_constr. cbn. @@ -959,31 +1607,37 @@ in All (for wf_local assumptions) reflexivity. - rewrite trans_dtype. simpl. eapply TT.type_Fix; auto. - + admit. (* fix guard *) + + now rewrite fix_guard_trans. + erewrite map_nth_error. 2: apply H0. destruct decl. unfold map_def. reflexivity. + eapply All_map, (All_impl X0). - firstorder. + intuition auto. destruct X2 as [s [? ?]]. + exists s; intuition auto. + fold trans. subst types. - eapply trans_mfix_All2;eassumption. - - rewrite trans_dtype. + eapply trans_mfix_All2; eassumption. + + now rewrite trans_wf_fixpoint. + - rewrite trans_dtype. simpl. eapply TT.type_CoFix; auto. + + now eapply cofix_guard_trans. + erewrite map_nth_error. 2: eassumption. destruct decl. unfold map_def. reflexivity. + fold trans. - eapply All_map, (All_impl X0); firstorder. + eapply All_map, (All_impl X0). + intros x [s ?]; exists s; intuition auto. + fold trans;subst types. - (* like trans_mfix_All2 without isLambda *) - admit. + now apply trans_mfix_All2. + + now rewrite trans_wf_cofixpoint. - eapply TT.type_Conv. + eassumption. - + admit. (* trans isWfArity *) - + admit. (* trans_cumul *) -Admitted. + + eassumption. + + now eapply trans_cumul. +Qed. + +Print Assumptions pcuic_to_template. diff --git a/pcuic/theories/PCUICTyping.v b/pcuic/theories/PCUICTyping.v index dd7fa91cb..e3bde6629 100644 --- a/pcuic/theories/PCUICTyping.v +++ b/pcuic/theories/PCUICTyping.v @@ -13,6 +13,8 @@ Require Import Equations.Type.Relation. From Equations Require Import Equations. Set Equations With UIP. +Implicit Types (cf : checker_flags) (Σ : global_env_ext). + (** * Typing derivations Inductive relations for reduction, conversion and typing of PCUIC terms. @@ -27,13 +29,6 @@ Hint Rewrite subst_context_length subst_instance_context_length map_length app_length lift_context_length @mapi_length @mapi_rec_length List.rev_length Nat.add_0_r : len. - -Definition isSort T := - match T with - | tSort u => True - | _ => False - end. - Fixpoint isArity T := match T with | tSort u => True @@ -71,93 +66,94 @@ Definition type_of_constructor mdecl (cdecl : ident * term * nat) (c : inductive let mind := inductive_mind (fst c) in subst0 (inds mind u mdecl.(ind_bodies)) (subst_instance_constr u (snd (fst cdecl))). - +Definition extends (Σ Σ' : global_env) := + { Σ'' & Σ' = Σ'' ++ Σ }. + (** ** Typing relation *) Module PCUICEnvTyping := EnvTyping PCUICTerm PCUICEnvironment. Include PCUICEnvTyping. -Derive Signature NoConfusion for All_local_env. +Derive NoConfusion for All_local_env. Derive NoConfusion for context_decl. Derive NoConfusion for list. (* AXIOM GUARD CONDITION *) -Axiom fix_guard : mfixpoint term -> bool. - -Axiom fix_guard_red1 : - forall Σ Γ mfix mfix' idx, - fix_guard mfix -> - red1 Σ Γ (tFix mfix idx) (tFix mfix' idx) -> - fix_guard mfix'. - -Axiom fix_guard_eq_term : - forall mfix mfix' idx, - fix_guard mfix -> - upto_names (tFix mfix idx) (tFix mfix' idx) -> - fix_guard mfix'. - -Axiom fix_guard_rename : - forall mfix f, - let mfix' := - map (map_def (rename f) (rename (shiftn (List.length mfix) f))) mfix - in - fix_guard mfix -> - fix_guard mfix'. - -Axiom fix_guard_lift : - forall mfix n k, - let k' := (#|mfix| + k)%nat in - let mfix' := map (map_def (lift n k) (lift n k')) mfix in - fix_guard mfix -> - fix_guard mfix'. -Axiom fix_guard_subst : - forall mfix s k, +Class GuardChecker := +{ (* Structural recursion check *) + fix_guard : global_env_ext -> context -> mfixpoint term -> bool ; + (* Guarded by destructors check *) + cofix_guard : global_env_ext -> context -> mfixpoint term -> bool ; + + fix_guard_red1 Σ Γ mfix mfix' idx : + fix_guard Σ Γ mfix -> + red1 Σ Γ (tFix mfix idx) (tFix mfix' idx) -> + fix_guard Σ Γ mfix' ; + + fix_guard_eq_term Σ Γ mfix mfix' idx : + fix_guard Σ Γ mfix -> + upto_names (tFix mfix idx) (tFix mfix' idx) -> + fix_guard Σ Γ mfix' ; + + fix_guard_lift Σ Γ Γ' Γ'' mfix : + let k' := (#|mfix| + #|Γ'|)%nat in + let mfix' := map (map_def (lift #|Γ''| #|Γ'|) (lift #|Γ''| k')) mfix in + fix_guard Σ (Γ ,,, Γ') mfix -> + fix_guard Σ (Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ') mfix' ; + + fix_guard_subst Σ Γ Γ' Δ mfix s k : let k' := (#|mfix| + k)%nat in let mfix' := map (map_def (subst s k) (subst s k')) mfix in - fix_guard mfix -> - fix_guard mfix'. - -(* AXIOM Cofix Guard Condition (guarded by constructors) *) - -Axiom cofix_guard : mfixpoint term -> bool. - -Axiom cofix_guard_red1 : - forall Σ Γ mfix mfix' idx, - cofix_guard mfix -> + fix_guard Σ (Γ ,,, Γ' ,,, Δ) mfix -> + fix_guard Σ (Γ ,,, subst_context s 0 Δ) mfix' ; + + fix_guard_subst_instance {cf:checker_flags} Σ Γ mfix u univs : + consistent_instance_ext (Σ.1, univs) Σ.2 u -> + fix_guard Σ Γ mfix -> + fix_guard (Σ.1, univs) (subst_instance_context u Γ) (map (map_def (subst_instance_constr u) (subst_instance_constr u)) + mfix) ; + + fix_guard_extends Σ Γ mfix Σ' : + fix_guard Σ Γ mfix -> + extends Σ.1 Σ' -> + fix_guard Σ' Γ mfix ; + + cofix_guard_red1 Σ Γ mfix mfix' idx : + cofix_guard Σ Γ mfix -> red1 Σ Γ (tCoFix mfix idx) (tCoFix mfix' idx) -> - cofix_guard mfix'. + cofix_guard Σ Γ mfix' ; -Axiom cofix_guard_eq_term : - forall mfix mfix' idx, - cofix_guard mfix -> + cofix_guard_eq_term Σ Γ mfix mfix' idx : + cofix_guard Σ Γ mfix -> upto_names (tCoFix mfix idx) (tCoFix mfix' idx) -> - cofix_guard mfix'. - -Axiom cofix_guard_rename : - forall mfix f, - let mfix' := - map (map_def (rename f) (rename (shiftn (List.length mfix) f))) mfix - in - cofix_guard mfix -> - cofix_guard mfix'. - -Axiom cofix_guard_lift : - forall mfix n k, - let k' := (#|mfix| + k)%nat in - let mfix' := map (map_def (lift n k) (lift n k')) mfix in - cofix_guard mfix -> - cofix_guard mfix'. + cofix_guard Σ Γ mfix' ; -Axiom cofix_guard_subst : - forall mfix s k, + cofix_guard_lift Σ Γ Γ' Γ'' mfix : + let k' := (#|mfix| + #|Γ'|)%nat in + let mfix' := map (map_def (lift #|Γ''| #|Γ'|) (lift #|Γ''| k')) mfix in + cofix_guard Σ (Γ ,,, Γ') mfix -> + cofix_guard Σ (Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ') mfix' ; + + cofix_guard_subst Σ Γ Γ' Δ mfix s k : let k' := (#|mfix| + k)%nat in let mfix' := map (map_def (subst s k) (subst s k')) mfix in - cofix_guard mfix -> - cofix_guard mfix'. + cofix_guard Σ (Γ ,,, Γ' ,,, Δ) mfix -> + cofix_guard Σ (Γ ,,, subst_context s 0 Δ) mfix' ; + + cofix_guard_subst_instance {cf:checker_flags} Σ Γ mfix u univs : + consistent_instance_ext (Σ.1, univs) Σ.2 u -> + cofix_guard Σ Γ mfix -> + cofix_guard (Σ.1, univs) (subst_instance_context u Γ) (map (map_def (subst_instance_constr u) (subst_instance_constr u)) + mfix) ; + + cofix_guard_extends Σ Γ mfix Σ' : + cofix_guard Σ Γ mfix -> + extends Σ.1 Σ' -> + cofix_guard Σ' Γ mfix }. -(* AXIOM INDUCTIVE GUARD CONDITION *) -Axiom ind_guard : mutual_inductive_body -> bool. +Axiom guard_checking : GuardChecker. +Existing Instance guard_checking. (** Compute the type of a case from the predicate [p], actual parameters [pars] and an inductive declaration. *) @@ -267,7 +263,7 @@ Definition isCoFinite (r : recursivity_kind) := | _ => false end. -Definition check_recursivity_kind Σ ind r := +Definition check_recursivity_kind (Σ : global_env) ind r := match lookup_env Σ ind with | Some (InductiveDecl mib) => Reflect.eqb mib.(ind_finite) r | _ => false @@ -331,14 +327,16 @@ Definition wf_universe Σ s := forall l, UnivExprSet.In l u -> LevelSet.In (UnivExpr.get_level l) (global_ext_levels Σ) end. +Reserved Notation "'wf_local' Σ Γ " (at level 9, Σ, Γ at next level). + Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> term -> Type := | type_Rel n decl : - All_local_env (lift_typing typing Σ) Γ -> + wf_local Σ Γ -> nth_error Γ n = Some decl -> Σ ;;; Γ |- tRel n : lift0 (S n) decl.(decl_type) | type_Sort s : - All_local_env (lift_typing typing Σ) Γ -> + wf_local Σ Γ -> wf_universe Σ s -> Σ ;;; Γ |- tSort s : tSort (Universe.super s) @@ -350,7 +348,7 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> | type_Lambda na A t s1 B : Σ ;;; Γ |- A : tSort s1 -> Σ ;;; Γ ,, vass na A |- t : B -> - Σ ;;; Γ |- (tLambda na A t) : tProd na A B + Σ ;;; Γ |- tLambda na A t : tProd na A B | type_LetIn na b B t s1 A : Σ ;;; Γ |- B : tSort s1 -> @@ -358,33 +356,40 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ ,, vdef na b B |- t : A -> Σ ;;; Γ |- tLetIn na b B t : tLetIn na b B A -| type_App t na A B u : +| type_App t na A B s u : + (* Paranoid assumption, allows to show equivalence with template-coq, + but eventually unnecessary thanks to validity. *) + Σ ;;; Γ |- tProd na A B : tSort s -> Σ ;;; Γ |- t : tProd na A B -> Σ ;;; Γ |- u : A -> Σ ;;; Γ |- tApp t u : B{0 := u} | type_Const cst u : - All_local_env (lift_typing typing Σ) Γ -> - forall decl (isdecl : declared_constant Σ.1 cst decl), + wf_local Σ Γ -> + forall decl, + declared_constant Σ.1 cst decl -> consistent_instance_ext Σ decl.(cst_universes) u -> Σ ;;; Γ |- (tConst cst u) : subst_instance_constr u decl.(cst_type) | type_Ind ind u : - All_local_env (lift_typing typing Σ) Γ -> - forall mdecl idecl (isdecl : declared_inductive Σ.1 mdecl ind idecl), + wf_local Σ Γ -> + forall mdecl idecl, + declared_inductive Σ.1 mdecl ind idecl -> consistent_instance_ext Σ mdecl.(ind_universes) u -> Σ ;;; Γ |- (tInd ind u) : subst_instance_constr u idecl.(ind_type) | type_Construct ind i u : - All_local_env (lift_typing typing Σ) Γ -> - forall mdecl idecl cdecl (isdecl : declared_constructor Σ.1 mdecl idecl (ind, i) cdecl), + wf_local Σ Γ -> + forall mdecl idecl cdecl, + declared_constructor Σ.1 mdecl idecl (ind, i) cdecl -> consistent_instance_ext Σ mdecl.(ind_universes) u -> Σ ;;; Γ |- (tConstruct ind i u) : type_of_constructor mdecl cdecl (ind, i) u | type_Case indnpar u p c brs args : let ind := indnpar.1 in let npar := indnpar.2 in - forall mdecl idecl (isdecl : declared_inductive Σ.1 mdecl ind idecl), + forall mdecl idecl, + declared_inductive Σ.1 mdecl ind idecl -> mdecl.(ind_npars) = npar -> let params := List.firstn npar args in forall ps pty, build_case_predicate_type ind mdecl idecl params u ps = Some pty -> @@ -393,30 +398,33 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ |- c : mkApps (tInd ind u) args -> isCoFinite mdecl.(ind_finite) = false -> forall btys, map_option_out (build_branches_type ind mdecl idecl params u p) = Some btys -> - All2 (fun br bty => (br.1 = bty.1) * (Σ ;;; Γ |- br.2 : bty.2) * (∑ s, Σ ;;; Γ |- bty.2 : tSort s)) brs btys -> + All2 (fun br bty => (br.1 = bty.1) * (Σ ;;; Γ |- br.2 : bty.2) * + (* This is a paranoid assumption *) + (∑ s, Σ ;;; Γ |- bty.2 : tSort s)) brs btys -> Σ ;;; Γ |- tCase indnpar p c brs : mkApps p (skipn npar args ++ [c]) | type_Proj p c u : - forall mdecl idecl pdecl (isdecl : declared_projection Σ.1 mdecl idecl p pdecl) args, + forall mdecl idecl pdecl, + declared_projection Σ.1 mdecl idecl p pdecl -> + forall args, Σ ;;; Γ |- c : mkApps (tInd (fst (fst p)) u) args -> #|args| = ind_npars mdecl -> let ty := snd pdecl in Σ ;;; Γ |- tProj p c : subst0 (c :: List.rev args) (subst_instance_constr u ty) | type_Fix mfix n decl : - fix_guard mfix -> + fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All_local_env (lift_typing typing Σ) Γ -> + wf_local Σ Γ -> All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> - All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype)) - * (isLambda d.(dbody) = true)%type) mfix -> + All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype))) mfix -> wf_fixpoint Σ.1 mfix -> Σ ;;; Γ |- tFix mfix n : decl.(dtype) | type_CoFix mfix n decl : - cofix_guard mfix -> + cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All_local_env (lift_typing typing Σ) Γ -> + wf_local Σ Γ -> All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> All (fun d => Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype)) mfix -> wf_cofixpoint Σ.1 mfix -> @@ -427,9 +435,8 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ |- B : tSort s -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- t : B -where " Σ ;;; Γ |- t : T " := (typing Σ Γ t T) : type_scope. - -Notation wf_local Σ Γ := (All_local_env (lift_typing typing Σ) Γ). +where " Σ ;;; Γ |- t : T " := (typing Σ Γ t T) +and "'wf_local' Σ Γ " := (All_local_env (lift_typing typing Σ) Γ). Lemma meta_conv {cf : checker_flags} Σ Γ t A B : Σ ;;; Γ |- t : A -> @@ -451,7 +458,6 @@ Definition unlift_opt_pred (P : global_env_ext -> context -> option term -> term Module PCUICTypingDef <: Typing PCUICTerm PCUICEnvironment PCUICEnvTyping. - Definition ind_guard := ind_guard. Definition typing := @typing. Definition wf_universe := @wf_universe. Definition conv := @conv. @@ -508,9 +514,9 @@ Proof. - exact (S (S (wf_local_size _ typing_size _ a))). - exact (S (S (wf_local_size _ typing_size _ a))). - exact (S (S (wf_local_size _ typing_size _ a))). - - exact (S (Nat.max d1 (Nat.max d2 + - exact (S (Nat.max d2 (Nat.max d3 (all2_size _ (fun x y p => Nat.max (typing_size Σ Γ (snd x) (snd y) (snd (fst p))) (typing_size _ _ _ _ (snd p).π2)) a)))). - - exact (S (Nat.max (Nat.max (wf_local_size _ typing_size _ a) (all_size _ (fun x p => typing_size Σ _ _ _ p.π2) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ (fst p)) a1))). + - exact (S (Nat.max (Nat.max (wf_local_size _ typing_size _ a) (all_size _ (fun x p => typing_size Σ _ _ _ p.π2) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - exact (S (Nat.max (Nat.max (wf_local_size _ typing_size _ a) (all_size _ (fun x p => typing_size Σ _ _ _ p.π2) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). Defined. @@ -553,17 +559,17 @@ Lemma wf_ext_consistent {cf:checker_flags} Σ : Proof. intros [? [? [? [? ?]]]]; assumption. Qed. Hint Resolve wf_ext_consistent : core. -Lemma wf_local_app `{checker_flags} Σ (Γ Γ' : context) : wf_local Σ (Γ ,,, Γ') -> wf_local Σ Γ. +Lemma wf_local_app_l `{checker_flags} Σ (Γ Γ' : context) : wf_local Σ (Γ ,,, Γ') -> wf_local Σ Γ. Proof. induction Γ'. auto. simpl. intros H'; inv H'; eauto. Defined. -Hint Resolve wf_local_app : wf. +Hint Resolve wf_local_app_l : wf. Lemma typing_wf_local `{checker_flags} {Σ} {Γ t T} : Σ ;;; Γ |- t : T -> wf_local Σ Γ. Proof. - induction 1; eauto using wf_local_app. + induction 1; eauto using wf_local_app_l. Defined. Hint Extern 4 (wf_local _ ?Γ) => @@ -602,7 +608,7 @@ Lemma type_Prop `{checker_flags} Σ : Σ ;;; [] |- tSort Universe.lProp : tSort Defined. Lemma env_prop_sigma `{checker_flags} P PΓ : env_prop P PΓ -> - forall Σ (wfΣ : wf Σ), Forall_decls_typing P Σ. + forall (Σ : global_env) (wfΣ : wf Σ), Forall_decls_typing P Σ. Proof. intros. red in X. eapply (X (empty_ext Σ)). apply wfΣ. @@ -620,7 +626,7 @@ Proof. Qed. Lemma size_wf_local_app `{checker_flags} {Σ} (Γ Γ' : context) (Hwf : wf_local Σ (Γ ,,, Γ')) : - wf_local_size Σ (@typing_size _) _ (wf_local_app _ _ _ Hwf) <= + wf_local_size Σ (@typing_size _) _ (wf_local_app_l _ _ _ Hwf) <= wf_local_size Σ (@typing_size _) _ Hwf. Proof. induction Γ' in Γ, Hwf |- *; try lia. simpl. lia. @@ -680,10 +686,439 @@ Qed. (** *** An induction principle ensuring the Σ declarations enjoy the same properties. Also theads the well-formedness of the local context and the induction principle for it, - and gives the right induction hypothesis - on typing judgments in application spines, fix and cofix blocks. + and gives the right induction hypothesis on typing judgments in application spines, + fix and cofix blocks. This general version allows to get the induction hypothesis on + any subderivation of the head of applications. + + The specialized version `typing_ind_env` below is the one used in general, with + no special case for applications. *) +Lemma typing_ind_env_app_size `{cf : checker_flags} : + forall (P : global_env_ext -> context -> term -> term -> Type) + (Pdecl := fun Σ Γ wfΓ t T tyT => P Σ Γ t T) + (PΓ : forall Σ Γ, wf_local Σ Γ -> Type), + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ), + All_local_env_over typing Pdecl Σ Γ wfΓ -> PΓ Σ Γ wfΓ) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : nat) decl, + nth_error Γ n = Some decl -> + PΓ Σ Γ wfΓ -> + P Σ Γ (tRel n) (lift0 (S n) decl.(decl_type))) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (u : Universe.t), + PΓ Σ Γ wfΓ -> + wf_universe Σ u -> + P Σ Γ (tSort u) (tSort (Universe.super u))) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) (s1 s2 : Universe.t), + PΓ Σ Γ wfΓ -> + Σ ;;; Γ |- t : tSort s1 -> + P Σ Γ t (tSort s1) -> + Σ ;;; Γ,, vass n t |- b : tSort s2 -> + P Σ (Γ,, vass n t) b (tSort s2) -> P Σ Γ (tProd n t b) (tSort (Universe.sort_of_product s1 s2))) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (t b : term) + (s1 : Universe.t) (bty : term), + PΓ Σ Γ wfΓ -> + Σ ;;; Γ |- t : tSort s1 -> + P Σ Γ t (tSort s1) -> + Σ ;;; Γ,, vass n t |- b : bty -> P Σ (Γ,, vass n t) b bty -> P Σ Γ (tLambda n t b) (tProd n t bty)) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (n : aname) (b b_ty b' : term) + (s1 : Universe.t) (b'_ty : term), + PΓ Σ Γ wfΓ -> + Σ ;;; Γ |- b_ty : tSort s1 -> + P Σ Γ b_ty (tSort s1) -> + Σ ;;; Γ |- b : b_ty -> + P Σ Γ b b_ty -> + Σ ;;; Γ,, vdef n b b_ty |- b' : b'_ty -> + P Σ (Γ,, vdef n b b_ty) b' b'_ty -> P Σ Γ (tLetIn n b b_ty b') (tLetIn n b b_ty b'_ty)) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t : term) na A B u s, + PΓ Σ Γ wfΓ -> + + + Σ ;;; Γ |- tProd na A B : tSort s -> P Σ Γ (tProd na A B) (tSort s) -> + forall (Ht : Σ ;;; Γ |- t : tProd na A B), P Σ Γ t (tProd na A B) -> + + (* Give a stronger induction hypothesis allowing to crawl under applications *) + (forall t' T' (Ht' : Σ ;;; Γ |- t' : T'), typing_size Ht' <= typing_size Ht -> P Σ Γ t' T') -> + + Σ ;;; Γ |- u : A -> P Σ Γ u A -> + P Σ Γ (tApp t u) (B{0 := u})) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) cst u (decl : constant_body), + Forall_decls_typing P Σ.1 -> + PΓ Σ Γ wfΓ -> + declared_constant Σ.1 cst decl -> + consistent_instance_ext Σ decl.(cst_universes) u -> + P Σ Γ (tConst cst u) (subst_instance_constr u (cst_type decl))) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) u + mdecl idecl (isdecl : declared_inductive Σ.1 mdecl ind idecl), + Forall_decls_typing P Σ.1 -> + PΓ Σ Γ wfΓ -> + consistent_instance_ext Σ mdecl.(ind_universes) u -> + P Σ Γ (tInd ind u) (subst_instance_constr u (ind_type idecl))) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) (i : nat) u + mdecl idecl cdecl (isdecl : declared_constructor Σ.1 mdecl idecl (ind, i) cdecl), + Forall_decls_typing P Σ.1 -> + PΓ Σ Γ wfΓ -> + consistent_instance_ext Σ mdecl.(ind_universes) u -> + P Σ Γ (tConstruct ind i u) (type_of_constructor mdecl cdecl (ind, i) u)) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (ind : inductive) u (npar : nat) + (p c : term) (brs : list (nat * term)) + (args : list term) (mdecl : mutual_inductive_body) (idecl : one_inductive_body) + (isdecl : declared_inductive (fst Σ) mdecl ind idecl), + Forall_decls_typing P Σ.1 -> PΓ Σ Γ wfΓ -> + ind_npars mdecl = npar -> + let params := firstn npar args in + forall ps pty, build_case_predicate_type ind mdecl idecl params u ps = Some pty -> + Σ ;;; Γ |- p : pty -> + P Σ Γ p pty -> + is_allowed_elimination (global_ext_constraints Σ) ps idecl.(ind_kelim) -> + Σ ;;; Γ |- c : mkApps (tInd ind u) args -> + isCoFinite mdecl.(ind_finite) = false -> + P Σ Γ c (mkApps (tInd ind u) args) -> + forall btys, map_option_out (build_branches_type ind mdecl idecl params u p) = Some btys -> + All2 (fun br bty => (br.1 = bty.1) * + (Σ ;;; Γ |- br.2 : bty.2) * P Σ Γ br.2 bty.2 * + ∑ s, (Σ ;;; Γ |- bty.2 : tSort s) * P Σ Γ bty.2 (tSort s)) + brs btys -> + P Σ Γ (tCase (ind, npar) p c brs) (mkApps p (skipn npar args ++ [c]))) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (p : projection) (c : term) u + mdecl idecl pdecl (isdecl : declared_projection Σ.1 mdecl idecl p pdecl) args, + Forall_decls_typing P Σ.1 -> PΓ Σ Γ wfΓ -> + Σ ;;; Γ |- c : mkApps (tInd (fst (fst p)) u) args -> + P Σ Γ c (mkApps (tInd (fst (fst p)) u) args) -> + #|args| = ind_npars mdecl -> + let ty := snd pdecl in P Σ Γ (tProj p c) (subst0 (c :: List.rev args) (subst_instance_constr u ty))) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (mfix : list (def term)) (n : nat) decl, + let types := fix_context mfix in + fix_guard Σ Γ mfix -> + nth_error mfix n = Some decl -> + PΓ Σ Γ wfΓ -> + All (fun d => {s & (Σ ;;; Γ |- d.(dtype) : tSort s)%type * P Σ Γ d.(dtype) (tSort s)})%type mfix -> + All (fun d => (Σ ;;; Γ ,,, types |- d.(dbody) : lift0 #|types| d.(dtype))%type * + P Σ (Γ ,,, types) d.(dbody) (lift0 #|types| d.(dtype)))%type mfix -> + wf_fixpoint Σ.1 mfix -> + P Σ Γ (tFix mfix n) decl.(dtype)) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (mfix : list (def term)) (n : nat) decl, + let types := fix_context mfix in + cofix_guard Σ Γ mfix -> + nth_error mfix n = Some decl -> + PΓ Σ Γ wfΓ -> + All (fun d => {s & (Σ ;;; Γ |- d.(dtype) : tSort s)%type * P Σ Γ d.(dtype) (tSort s)})%type mfix -> + All (fun d => (Σ ;;; Γ ,,, types |- d.(dbody) : lift0 #|types| d.(dtype))%type * + P Σ (Γ ,,, types) d.(dbody) (lift0 #|types| d.(dtype)))%type mfix -> + wf_cofixpoint Σ.1 mfix -> + P Σ Γ (tCoFix mfix n) decl.(dtype)) -> + + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t A B : term) s, + PΓ Σ Γ wfΓ -> + Σ ;;; Γ |- t : A -> + P Σ Γ t A -> + Σ ;;; Γ |- B : tSort s -> + P Σ Γ B (tSort s) -> + Σ ;;; Γ |- A <= B -> + P Σ Γ t B) -> + + env_prop P PΓ. +Proof. + intros P Pdecl PΓ; unfold env_prop. + intros XΓ X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 X12 Σ wfΣ Γ t T H. + (* NOTE (Danil): while porting to 8.9, I had to split original "pose" into 2 pieces, + otherwise it takes forever to execure the "pose", for some reason *) + pose (@Fix_F ({ Σ : _ & { wfΣ : wf Σ.1 & { Γ : context & + { t : term & { T : term & Σ ;;; Γ |- t : T }}}}})) as p0. + specialize (p0 (PCUICUtils.dlexprod (precompose lt (fun Σ => globenv_size (fst Σ))) + (fun Σ => precompose lt (fun x => typing_size (projT2 (projT2 (projT2 (projT2 x)))))))) as p. + set(foo := existT _ Σ (existT _ wfΣ (existT _ Γ (existT _ t (existT _ _ H)))) : { Σ : _ & { wfΣ : wf Σ.1 & { Γ : context & { t : term & { T : term & Σ ;;; Γ |- t : T }}}}}). + change Σ with (projT1 foo). + change Γ with (projT1 (projT2 (projT2 foo))). + change t with (projT1 (projT2 (projT2 (projT2 foo)))). + change T with (projT1 (projT2 (projT2 (projT2 (projT2 foo))))). + change H with (projT2 (projT2 (projT2 (projT2 (projT2 foo))))). + revert foo. + match goal with + |- let foo := _ in @?P foo => specialize (p (fun x => P x)) + end. + forward p; [ | apply p; apply PCUICUtils.wf_dlexprod; intros; apply wf_precompose; apply lt_wf]. + clear p. + clear Σ wfΣ Γ t T H. + intros (Σ & wfΣ & Γ & t & t0 & H). simpl. + intros IH. simpl in IH. + split. split. + destruct Σ as [Σ φ]. destruct Σ. + constructor. + cbn in wfΣ; inversion_clear wfΣ. auto. + inv wfΣ. + rename X14 into Xg. + constructor; auto. unfold Forall_decls_typing in IH. + - simple refine (let IH' := IH ((Σ, udecl); (X13; []; (tSort Universe.lProp); _; _)) in _). + shelve. simpl. apply type_Prop. + forward IH'. constructor 1; cbn. lia. + apply IH'; auto. + - simpl. simpl in *. + destruct d; simpl. + + destruct c; simpl in *. + destruct cst_body; simpl in *. + simpl. + intros. red in Xg. simpl in Xg. + specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ [] (existT _ _ (existT _ _ Xg)))))). + simpl in IH. + forward IH. constructor 1. simpl; lia. + apply IH. + red. simpl. red in Xg; simpl in Xg. + destruct Xg as [s Hs]. red. simpl. + specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ [] (existT _ _ (existT _ _ Hs)))))). + simpl in IH. + forward IH. constructor 1. simpl; lia. exists s. eapply IH. + + red in Xg. + destruct Xg as [onI onP onnp]; constructor; eauto. + eapply Alli_impl; eauto. clear onI onP onnp; intros n x Xg. + refine {| ind_indices := Xg.(ind_indices); + ind_arity_eq := Xg.(ind_arity_eq); + ind_cshapes := Xg.(ind_cshapes) |}. + + ++ apply onArity in Xg. destruct Xg as [s Hs]. exists s; auto. + specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ [] (existT _ _ (existT _ _ Hs)))))). + simpl in IH. simpl. apply IH; constructor 1; simpl; lia. + ++ pose proof Xg.(onConstructors) as Xg'. + eapply All2_impl; eauto. intros. + destruct X14 as [cass chead tyeq onctyp oncargs oncind]. + unshelve econstructor; eauto. + destruct onctyp as [s Hs]. + simpl in Hs. + specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ Hs)))))). + simpl in IH. simpl. exists s. simpl. apply IH; constructor 1; simpl; auto with arith. + eapply sorts_local_ctx_impl; eauto. simpl. intros. red in X14. + destruct T. + specialize (IH ((Σ, udecl); (X13; _; _; _; X14))). + apply IH. simpl. constructor 1. simpl. auto with arith. + destruct X14 as [u Hu]. exists u. + specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ Hu)))))). + apply IH. simpl. constructor 1. simpl. auto with arith. + clear -X13 IH oncind. + revert oncind. + generalize (List.rev (lift_context #|cshape_args y| 0 (ind_indices Xg))). + generalize (cshape_indices y). induction 1; constructor; auto. + red in p0 |- *. + specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ p0)))))). + apply IH. simpl. constructor 1. simpl. auto with arith. + ++ intros Hprojs; pose proof (onProjections Xg Hprojs); auto. + ++ destruct Xg. simpl. unfold check_ind_sorts in *. + destruct Universe.is_prop; auto. + destruct Universe.is_sprop; auto. + split. apply ind_sorts0. destruct indices_matter; auto. + eapply type_local_ctx_impl. eapply ind_sorts0. + intros. red in X14. + destruct T. + specialize (IH ((Σ, udecl); (X13; _; _; _; X14))). + apply IH. simpl. constructor 1. simpl. auto with arith. + destruct X14 as [u Hu]. exists u. + specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ Hu)))))). + apply IH. simpl. constructor 1. simpl. auto with arith. + ++ apply onIndices. + ++ red in onP |- *. + eapply All_local_env_impl; eauto. + intros. destruct T; simpl in X14. + specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ X14)))))). + simpl in IH. apply IH. constructor 1. simpl. lia. + destruct X14 as [u Hu]. + specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ Hu)))))). + simpl in IH. simpl. exists u. apply IH. constructor 1. simpl. lia. + + - assert (forall Γ t T (Hty : Σ ;;; Γ |- t : T), + typing_size Hty < typing_size H -> + Forall_decls_typing P Σ.1 * P Σ Γ t T). + intros. + specialize (IH (existT _ Σ (existT _ wfΣ (existT _ _ (existT _ _ (existT _ _ Hty)))))). + simpl in IH. + forward IH. + constructor 2. simpl. apply H0. + split; apply IH. clear IH. + rename X13 into X14. + + assert (All_local_env_over typing Pdecl Σ Γ (typing_wf_local H)). + { clear -Pdecl wfΣ X14. + pose proof (typing_wf_local_size H). + set (foo := typing_wf_local H) in *. + clearbody foo. + revert foo H0. generalize Γ at 1 2 4. + induction foo; simpl in *; try destruct t2 as [u Hu]; simpl in *; constructor. + - simpl in *. apply IHfoo. lia. + - red. eapply (X14 _ _ _ Hu). lia. + - simpl in *. apply IHfoo. lia. + - red. apply (X14 _ _ _ t3). lia. + - red. apply (X14 _ _ _ Hu). lia. } + eapply XΓ; eauto. + + - assert (forall Γ t T (Hty : Σ ;;; Γ |- t : T), + typing_size Hty < typing_size H -> + Forall_decls_typing P Σ.1 * P Σ Γ t T). + intros. + specialize (IH (existT _ Σ (existT _ wfΣ (existT _ _ (existT _ _ (existT _ _ Hty)))))). + simpl in IH. + forward IH. + constructor 2. simpl. apply H0. + split; apply IH. clear IH. + rename X13 into X14. + + assert (Hdecls: typing_size H > 1 -> Forall_decls_typing P Σ.1). + { specialize (X14 _ _ _ (type_Prop _)). + simpl in X14. intros Hle. apply X14. lia. } + + assert (All_local_env_over typing Pdecl Σ Γ (typing_wf_local H)). + { clear -Pdecl wfΣ X14. + pose proof (typing_wf_local_size H). + set (foo := typing_wf_local H) in *. + clearbody foo. + revert foo H0. generalize Γ at 1 2 4. + induction foo; simpl in *; try destruct t2 as [u Hu]; simpl in *; constructor. + - simpl in *. apply IHfoo. lia. + - red. eapply (X14 _ _ _ Hu). lia. + - simpl in *. apply IHfoo. lia. + - red. apply (X14 _ _ _ t3). lia. + - red. apply (X14 _ _ _ Hu). lia. } + apply XΓ in X13. all:auto. + + destruct H; + try solve [ match reverse goal with + H : _ |- _ => eapply H + end; eauto; + unshelve eapply X14; simpl; auto with arith]. + + -- match reverse goal with + H : _ |- _ => eapply H + end; eauto; + unshelve eapply X14; simpl; eauto with arith wf. + + -- match reverse goal with + H : _ |- _ => eapply H + end; eauto. all:try unshelve eapply X14; simpl; auto; try lia. + Unshelve. 2:exact H0. + simpl. intros. + eapply X14. instantiate (1 := Ht'). + simpl. lia. + + -- match reverse goal with + H : _ |- _ => eapply H + end; eauto. + simpl in Hdecls. apply Hdecls; lia. + + -- eapply X6; eauto. + apply Hdecls; simpl; lia. + + -- eapply X7; eauto. apply Hdecls; simpl; lia. + + -- destruct indnpar as [ind' npar']; + cbn in ind; cbn in npar; subst ind; subst npar. + eapply X8; eauto. + ++ eapply (X14 _ _ _ H); eauto. simpl; auto with arith. + ++ eapply (X14 _ _ _ H); eauto. simpl; auto with arith. + ++ simpl in *. + eapply (X14 _ _ _ H0); eauto. clear. lia. + ++ clear X13 Hdecls. revert a X14. simpl. clear. intros. + induction a; simpl in *. + ** constructor. + ** destruct r as [[? ?] ?]. constructor. + --- intuition eauto. + +++ eapply (X14 _ _ _ t); eauto. simpl; auto with arith. + lia. + +++ destruct s as [s Hs]. exists s; split; [auto|]. + eapply (X14 _ _ _ Hs); eauto. simpl; auto with arith. + lia. + --- apply IHa. auto. intros. + eapply (X14 _ _ _ Hty). lia. + + -- eapply X9; eauto. apply Hdecls; simpl. + pose proof (typing_size_pos H). lia. + eapply (X14 _ _ _ H). simpl. lia. + + -- clear X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X11 X12. + eapply X10; eauto; clear X10. simpl in *. + * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), + typing_size Hty < + S (all_size (fun x : def term => + ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) + (fun (x : def term) + (p : ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) => + typing_size p.π2) a0) -> + Forall_decls_typing P Σ.1 * P Σ Γ t T). + intros; eauto. eapply (X14 _ _ _ Hty); eauto. lia. + clear X13 X14 a Hdecls. + clear -a0 X. + induction a0; constructor. + destruct p as [s Hs]. exists s; split; auto. + apply (X (dtype x) (tSort s) Hs). simpl. lia. + apply IHa0. intros. eapply (X _ _ Hty); eauto. + simpl. lia. + * simpl in X14. + assert(forall Γ0 : context, + wf_local Σ Γ0 -> + forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), + typing_size Hty < + S + (all_size _ (fun (x : def term) p => typing_size p) a1) -> + Forall_decls_typing P Σ.1 * P Σ Γ0 t T). + {intros. eapply (X14 _ _ _ Hty); eauto. lia. } + clear X14 X13. + clear e decl i a0 Hdecls i0. + remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. + + induction a1; econstructor; eauto. + ++ split; auto. + eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. + ++ eapply IHa1. intros. + eapply (X _ X0 _ _ Hty). simpl; lia. + + -- clear X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X12. + eapply X11; eauto; clear X11. simpl in *. + * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), + typing_size Hty < + S (all_size (fun x : def term => + ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) + (fun (x : def term) + (p : ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) => + typing_size p.π2) a0) -> + Forall_decls_typing P Σ.1 * P Σ Γ t T). + intros; eauto. eapply (X14 _ _ _ Hty); eauto. lia. + clear X13 X14 a Hdecls. + clear -a0 X. + induction a0; constructor. + destruct p as [s Hs]. exists s; split; auto. + apply (X (dtype x) (tSort s) Hs). simpl. lia. + apply IHa0. intros. eapply (X _ _ Hty); eauto. + simpl. lia. + * simpl in X14. + assert(forall Γ0 : context, + wf_local Σ Γ0 -> + forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), + typing_size Hty < + S + (all_size (fun x : def term => (Σ;;; Γ ,,, fix_context mfix |- dbody x : lift0 #|fix_context mfix| (dtype x))%type) + (fun (x : def term) p => typing_size p) a1) -> + Forall_decls_typing P Σ.1 * P Σ Γ0 t T). + {intros. eapply (X14 _ _ _ Hty); eauto. lia. } + clear X14 X13. + clear e decl i a0 Hdecls i0. + remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. + + induction a1; econstructor; eauto. + ++ split; auto. + eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. + ++ eapply IHa1. intros. + eapply (X _ X0 _ _ Hty). simpl; lia. +Qed. + Lemma typing_ind_env `{cf : checker_flags} : forall (P : global_env_ext -> context -> term -> term -> Type) (Pdecl := fun Σ Γ wfΓ t T tyT => P Σ Γ t T) @@ -696,6 +1131,7 @@ Lemma typing_ind_env `{cf : checker_flags} : nth_error Γ n = Some decl -> PΓ Σ Γ wfΓ -> P Σ Γ (tRel n) (lift0 (S n) decl.(decl_type))) -> + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (u : Universe.t), PΓ Σ Γ wfΓ -> wf_universe Σ u -> @@ -725,8 +1161,9 @@ Lemma typing_ind_env `{cf : checker_flags} : Σ ;;; Γ,, vdef n b b_ty |- b' : b'_ty -> P Σ (Γ,, vdef n b b_ty) b' b'_ty -> P Σ Γ (tLetIn n b b_ty b') (tLetIn n b b_ty b'_ty)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t : term) na A B u, + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t : term) na A B u s, PΓ Σ Γ wfΓ -> + Σ ;;; Γ |- tProd na A B : tSort s -> P Σ Γ (tProd na A B) (tSort s) -> Σ ;;; Γ |- t : tProd na A B -> P Σ Γ t (tProd na A B) -> Σ ;;; Γ |- u : A -> P Σ Γ u A -> P Σ Γ (tApp t u) (B{0 := u})) -> @@ -783,19 +1220,18 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (mfix : list (def term)) (n : nat) decl, let types := fix_context mfix in - fix_guard mfix -> + fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ Γ wfΓ -> All (fun d => {s & (Σ ;;; Γ |- d.(dtype) : tSort s)%type * P Σ Γ d.(dtype) (tSort s)})%type mfix -> All (fun d => (Σ ;;; Γ ,,, types |- d.(dbody) : lift0 #|types| d.(dtype))%type * - (isLambda d.(dbody) = true)%type * P Σ (Γ ,,, types) d.(dbody) (lift0 #|types| d.(dtype)))%type mfix -> wf_fixpoint Σ.1 mfix -> P Σ Γ (tFix mfix n) decl.(dtype)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (mfix : list (def term)) (n : nat) decl, let types := fix_context mfix in - cofix_guard mfix -> + cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> PΓ Σ Γ wfΓ -> All (fun d => {s & (Σ ;;; Γ |- d.(dtype) : tSort s)%type * P Σ Γ d.(dtype) (tSort s)})%type mfix -> @@ -817,283 +1253,8 @@ Lemma typing_ind_env `{cf : checker_flags} : Proof. intros P Pdecl PΓ; unfold env_prop. intros XΓ X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X11 X12 Σ wfΣ Γ t T H. - (* NOTE (Danil): while porting to 8.9, I had to split original "pose" into 2 pieces, - otherwise it takes forever to execure the "pose", for some reason *) - pose (@Fix_F ({ Σ : _ & { wfΣ : wf Σ.1 & { Γ : context & - { t : term & { T : term & Σ ;;; Γ |- t : T }}}}})) as p0. - specialize (p0 (dlexprod (precompose lt (fun Σ => globenv_size (fst Σ))) - (fun Σ => precompose lt (fun x => typing_size (projT2 (projT2 (projT2 (projT2 x)))))))) as p. - set(foo := existT _ Σ (existT _ wfΣ (existT _ Γ (existT _ t (existT _ _ H)))) : { Σ : _ & { wfΣ : wf Σ.1 & { Γ : context & { t : term & { T : term & Σ ;;; Γ |- t : T }}}}}). - change Σ with (projT1 foo). - change Γ with (projT1 (projT2 (projT2 foo))). - change t with (projT1 (projT2 (projT2 (projT2 foo)))). - change T with (projT1 (projT2 (projT2 (projT2 (projT2 foo))))). - change H with (projT2 (projT2 (projT2 (projT2 (projT2 foo))))). - revert foo. - match goal with - |- let foo := _ in @?P foo => specialize (p (fun x => P x)) - end. - forward p; [ | apply p; apply wf_dlexprod; intros; apply wf_precompose; apply lt_wf]. - clear p. - clear Σ wfΣ Γ t T H. - intros (Σ & wfΣ & Γ & t & t0 & H). simpl. - intros IH. simpl in IH. - split. split. - destruct Σ as [Σ φ]. destruct Σ. - constructor. - cbn in wfΣ; inversion_clear wfΣ. auto. - inv wfΣ. - rename X14 into Xg. - constructor; auto. unfold Forall_decls_typing in IH. - - simple refine (let IH' := IH ((Σ, udecl); (X13; []; (tSort Universe.lProp); _; _)) in _). - shelve. simpl. apply type_Prop. - forward IH'. constructor 1; cbn. lia. - apply IH'; auto. - - simpl. simpl in *. - destruct d; simpl. - + destruct c; simpl in *. - destruct cst_body; simpl in *. - simpl. - intros. red in Xg. simpl in Xg. - specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ [] (existT _ _ (existT _ _ Xg)))))). - simpl in IH. - forward IH. constructor 1. simpl; lia. - apply IH. - red. simpl. red in Xg; simpl in Xg. - destruct Xg as [s Hs]. red. simpl. - specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ [] (existT _ _ (existT _ _ Hs)))))). - simpl in IH. - forward IH. constructor 1. simpl; lia. exists s. eapply IH. - + red in Xg. - destruct Xg as [onI onP onnp]; constructor; eauto. - eapply Alli_impl; eauto. clear onI onP onnp; intros n x Xg. - refine {| ind_indices := Xg.(ind_indices); - ind_arity_eq := Xg.(ind_arity_eq); - ind_cshapes := Xg.(ind_cshapes) |}. - - ++ apply onArity in Xg. destruct Xg as [s Hs]. exists s; auto. - specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ [] (existT _ _ (existT _ _ Hs)))))). - simpl in IH. simpl. apply IH; constructor 1; simpl; lia. - ++ pose proof Xg.(onConstructors) as Xg'. - eapply All2_impl; eauto. intros. - destruct X14 as [cass chead tyeq onctyp oncargs oncind]. - unshelve econstructor; eauto. - destruct onctyp as [s Hs]. - simpl in Hs. - specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ Hs)))))). - simpl in IH. simpl. exists s. simpl. apply IH; constructor 1; simpl; auto with arith. - eapply sorts_local_ctx_impl; eauto. simpl. intros. red in X14. - destruct T. - specialize (IH ((Σ, udecl); (X13; _; _; _; X14))). - apply IH. simpl. constructor 1. simpl. auto with arith. - destruct X14 as [u Hu]. exists u. - specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ Hu)))))). - apply IH. simpl. constructor 1. simpl. auto with arith. - clear -X13 IH oncind. - revert oncind. - generalize (List.rev (lift_context #|cshape_args y| 0 (ind_indices Xg))). - generalize (cshape_indices y). induction 1; constructor; auto. - red in p0 |- *. - specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ p0)))))). - apply IH. simpl. constructor 1. simpl. auto with arith. - ++ intros Hprojs; pose proof (onProjections Xg Hprojs); auto. - ++ destruct Xg. simpl. unfold check_ind_sorts in *. - destruct Universe.is_prop; auto. - destruct Universe.is_sprop; auto. - split. apply ind_sorts0. destruct indices_matter; auto. - eapply type_local_ctx_impl. eapply ind_sorts0. - intros. red in X14. - destruct T. - specialize (IH ((Σ, udecl); (X13; _; _; _; X14))). - apply IH. simpl. constructor 1. simpl. auto with arith. - destruct X14 as [u Hu]. exists u. - specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ Hu)))))). - apply IH. simpl. constructor 1. simpl. auto with arith. - ++ apply onIndices. - ++ red in onP |- *. - eapply All_local_env_impl; eauto. - intros. destruct T; simpl in X14. - specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ X14)))))). - simpl in IH. apply IH. constructor 1. simpl. lia. - destruct X14 as [u Hu]. - specialize (IH (existT _ (Σ, udecl) (existT _ X13 (existT _ _ (existT _ _ (existT _ _ Hu)))))). - simpl in IH. simpl. exists u. apply IH. constructor 1. simpl. lia. - - - assert (forall Γ t T (Hty : Σ ;;; Γ |- t : T), - typing_size Hty < typing_size H -> - Forall_decls_typing P Σ.1 * P Σ Γ t T). - intros. - specialize (IH (existT _ Σ (existT _ wfΣ (existT _ _ (existT _ _ (existT _ _ Hty)))))). - simpl in IH. - forward IH. - constructor 2. simpl. apply H0. - split; apply IH. clear IH. - rename X13 into X14. - - assert (All_local_env_over typing Pdecl Σ Γ (typing_wf_local H)). - { clear -Pdecl wfΣ X14. - pose proof (typing_wf_local_size H). - set (foo := typing_wf_local H) in *. - clearbody foo. - revert foo H0. generalize Γ at 1 2 4. - induction foo; simpl in *; try destruct t2 as [u Hu]; simpl in *; constructor. - - simpl in *. apply IHfoo. lia. - - red. eapply (X14 _ _ _ Hu). lia. - - simpl in *. apply IHfoo. lia. - - red. apply (X14 _ _ _ t3). lia. - - red. apply (X14 _ _ _ Hu). lia. } - eapply XΓ; eauto. - - - assert (forall Γ t T (Hty : Σ ;;; Γ |- t : T), - typing_size Hty < typing_size H -> - Forall_decls_typing P Σ.1 * P Σ Γ t T). - intros. - specialize (IH (existT _ Σ (existT _ wfΣ (existT _ _ (existT _ _ (existT _ _ Hty)))))). - simpl in IH. - forward IH. - constructor 2. simpl. apply H0. - split; apply IH. clear IH. - rename X13 into X14. - - assert (Hdecls: typing_size H > 1 -> Forall_decls_typing P Σ.1). - { specialize (X14 _ _ _ (type_Prop _)). - simpl in X14. intros Hle. apply X14. lia. } - - assert (All_local_env_over typing Pdecl Σ Γ (typing_wf_local H)). - { clear -Pdecl wfΣ X14. - pose proof (typing_wf_local_size H). - set (foo := typing_wf_local H) in *. - clearbody foo. - revert foo H0. generalize Γ at 1 2 4. - induction foo; simpl in *; try destruct t2 as [u Hu]; simpl in *; constructor. - - simpl in *. apply IHfoo. lia. - - red. eapply (X14 _ _ _ Hu). lia. - - simpl in *. apply IHfoo. lia. - - red. apply (X14 _ _ _ t3). lia. - - red. apply (X14 _ _ _ Hu). lia. } - apply XΓ in X13. all:auto. - - destruct H; - try solve [ match reverse goal with - H : _ |- _ => eapply H - end; eauto; - unshelve eapply X14; simpl; auto with arith]. - - -- match reverse goal with - H : _ |- _ => eapply H - end; eauto; - unshelve eapply X14; simpl; eauto with arith wf. - - -- match reverse goal with - H : _ |- _ => eapply H - end; eauto. - simpl in Hdecls. apply Hdecls; lia. - - -- eapply X6; eauto. - apply Hdecls; simpl; lia. - - -- eapply X7; eauto. apply Hdecls; simpl; lia. - - -- destruct indnpar as [ind' npar']; - cbn in ind; cbn in npar; subst ind; subst npar. - eapply X8; eauto. - ++ eapply (X14 _ _ _ H); eauto. simpl; auto with arith. - ++ eapply (X14 _ _ _ H); eauto. simpl; auto with arith. - ++ simpl in *. - eapply (X14 _ _ _ H0); eauto. clear. lia. - ++ clear X13 Hdecls. revert a X14. simpl. clear. intros. - induction a; simpl in *. - ** constructor. - ** destruct r as [[? ?] ?]. constructor. - --- intuition eauto. - +++ eapply (X14 _ _ _ t); eauto. simpl; auto with arith. - lia. - +++ destruct s as [s Hs]. exists s; split; [auto|]. - eapply (X14 _ _ _ Hs); eauto. simpl; auto with arith. - lia. - --- apply IHa. auto. intros. - eapply (X14 _ _ _ Hty). lia. - - -- eapply X9; eauto. apply Hdecls; simpl. - pose proof (typing_size_pos H). lia. - eapply (X14 _ _ _ H). simpl. lia. - - -- clear X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X11 X12. - eapply X10; eauto; clear X10. simpl in *. - * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < - S (all_size (fun x : def term => - ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) - (fun (x : def term) - (p : ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) => - typing_size p.π2) a0) -> - Forall_decls_typing P Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. lia. - clear X13 X14 a Hdecls. - clear -a0 X. - induction a0; constructor. - destruct p as [s Hs]. exists s; split; auto. - apply (X (dtype x) (tSort s) Hs). simpl. lia. - apply IHa0. intros. eapply (X _ _ Hty); eauto. - simpl. lia. - * simpl in X14. - assert(forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size _ (fun (x : def term) p => typing_size (fst p)) a1) -> - Forall_decls_typing P Σ.1 * P Σ Γ0 t T). - {intros. eapply (X14 _ _ _ Hty); eauto. lia. } - clear X14 X13. - clear e decl i a0 Hdecls i0. - remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. - - induction a1; econstructor; eauto. - ++ split; auto. - eapply (X _ (typing_wf_local (fst p)) _ _ (fst p)). simpl. lia. - ++ eapply IHa1. intros. - eapply (X _ X0 _ _ Hty). simpl; lia. - - -- clear X X0 X1 X2 X3 X4 X5 X6 X7 X8 X9 X10 X12. - eapply X11; eauto; clear X11. simpl in *. - * assert(forall (t T : term) (Hty : Σ;;; Γ |- t : T), - typing_size Hty < - S (all_size (fun x : def term => - ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) - (fun (x : def term) - (p : ∑ s : Universe.t, Σ;;; Γ |- dtype x : tSort s) => - typing_size p.π2) a0) -> - Forall_decls_typing P Σ.1 * P Σ Γ t T). - intros; eauto. eapply (X14 _ _ _ Hty); eauto. lia. - clear X13 X14 a Hdecls. - clear -a0 X. - induction a0; constructor. - destruct p as [s Hs]. exists s; split; auto. - apply (X (dtype x) (tSort s) Hs). simpl. lia. - apply IHa0. intros. eapply (X _ _ Hty); eauto. - simpl. lia. - * simpl in X14. - assert(forall Γ0 : context, - wf_local Σ Γ0 -> - forall (t T : term) (Hty : Σ;;; Γ0 |- t : T), - typing_size Hty < - S - (all_size (fun x : def term => (Σ;;; Γ ,,, fix_context mfix |- dbody x : lift0 #|fix_context mfix| (dtype x))%type) - (fun (x : def term) p => typing_size p) a1) -> - Forall_decls_typing P Σ.1 * P Σ Γ0 t T). - {intros. eapply (X14 _ _ _ Hty); eauto. lia. } - clear X14 X13. - clear e decl i a0 Hdecls i0. - remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. - - induction a1; econstructor; eauto. - ++ split; auto. - eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. - ++ eapply IHa1. intros. - eapply (X _ X0 _ _ Hty). simpl; lia. + apply typing_ind_env_app_size; eauto. Qed. -Print Assumptions typing_ind_env. Ltac my_rename_hyp h th := match th with @@ -1128,7 +1289,7 @@ Section All_local_env. apply IHX. simpl in *. lia. Qed. - Lemma lookup_on_global_env P Σ c decl : + Lemma lookup_on_global_env P (Σ : global_env) c decl : on_global_env P Σ -> lookup_env Σ c = Some decl -> { Σ' & { wfΣ' : on_global_env P Σ'.1 & on_global_decl P Σ' c decl } }. @@ -1141,6 +1302,15 @@ Section All_local_env. Qed. Lemma All_local_env_app (P : context -> term -> option term -> Type) l l' : + All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l' -> + All_local_env P (l ,,, l'). + Proof. + induction l'; simpl; auto. intuition. + intuition. destruct a. destruct decl_body. + inv b. econstructor; eauto. inv b; econstructor; eauto. + Qed. + + Lemma All_local_env_app_inv (P : context -> term -> option term -> Type) l l' : All_local_env P (l ,,, l') -> All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l'. Proof. @@ -1159,12 +1329,12 @@ Section All_local_env. * apply X1. * apply X2. Qed. - - Definition wf_local_rel_app {Σ Γ1 Γ2 Γ3} : + + Definition wf_local_rel_app_inv {Σ Γ1 Γ2 Γ3} : wf_local_rel Σ Γ1 (Γ2 ,,, Γ3) -> wf_local_rel Σ Γ1 Γ2 * wf_local_rel Σ (Γ1 ,,, Γ2) Γ3. Proof. - intros h. apply All_local_env_app in h as [h1 h2]. + intros h. apply All_local_env_app_inv in h as [h1 h2]. split. - exact h1. - eapply All_local_env_impl. 1: exact h2. @@ -1175,7 +1345,6 @@ Section All_local_env. all: auto. Defined. - Lemma All_local_env_lookup {P Γ n} {decl} : All_local_env P Γ -> nth_error Γ n = Some decl -> @@ -1188,20 +1357,11 @@ Section All_local_env. eapply IHX. Qed. - Lemma All_local_env_app_inv (P : context -> term -> option term -> Type) l l' : - All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l' -> - All_local_env P (l ,,, l'). - Proof. - induction l'; simpl; auto. intuition. - intuition. destruct a. destruct decl_body. - inv b. econstructor; eauto. inv b; econstructor; eauto. Qed. - - - Definition wf_local_rel_app_inv {Σ Γ1 Γ2 Γ3} : + Definition wf_local_rel_app {Σ Γ1 Γ2 Γ3} : wf_local_rel Σ Γ1 Γ2 -> wf_local_rel Σ (Γ1 ,,, Γ2) Γ3 -> wf_local_rel Σ Γ1 (Γ2 ,,, Γ3). Proof. - intros h1 h2. eapply All_local_env_app_inv. + intros h1 h2. eapply All_local_env_app. split. - assumption. - eapply All_local_env_impl. @@ -1210,7 +1370,28 @@ Section All_local_env. intros Γ t []; cbn; now rewrite app_context_assoc. Defined. + + Definition wf_local_app {Σ Γ1 Γ2} : + wf_local Σ Γ1 -> + wf_local_rel Σ Γ1 Γ2 -> + wf_local Σ (Γ1 ,,, Γ2). + Proof. + intros H1 H2. apply wf_local_local_rel. + apply wf_local_rel_local in H1. + apply wf_local_rel_app; tas. + now rewrite app_context_nil_l. + Qed. + Definition wf_local_app_inv {Σ Γ1 Γ2} : + wf_local Σ (Γ1 ,,, Γ2) -> + wf_local Σ Γ1 * wf_local_rel Σ Γ1 Γ2. + Proof. + intros H. + apply wf_local_rel_local in H. + apply wf_local_rel_app_inv in H as [H H']; tas. + rewrite app_context_nil_l in H'. + now split; [eapply wf_local_local_rel|]. + Qed. Lemma All_local_env_map (P : context -> term -> option term -> Type) f l : (forall u, f (tSort u) = tSort u) -> @@ -1222,7 +1403,7 @@ Section All_local_env. Definition property := forall (Σ : global_env_ext) (Γ : context), - All_local_env (lift_typing typing Σ) Γ -> forall t T : term, typing Σ Γ t T -> Type. + wf_local Σ Γ -> forall t T : term, typing Σ Γ t T -> Type. Definition lookup_wf_local {Γ P} (wfΓ : All_local_env P Γ) (n : nat) (isdecl : n < #|Γ|) : @@ -1247,7 +1428,7 @@ Section All_local_env. rewrite -skipn_app. now apply All_local_env_skipn. replace (n - #|Γ'|) with 0 by lia. now rewrite skipn_0. rewrite List.skipn_all2. lia. - now eapply wf_local_app in wf. + now eapply wf_local_app_l in wf. Qed. Definition on_local_decl_glob (P : term -> option term -> Type) d := @@ -1281,7 +1462,7 @@ Section All_local_env. | {| decl_name := na; decl_body := None; decl_type := ty |} => fun H => P Σ Γ wfΓ _ _ (projT2 H) end H. - Lemma nth_error_All_local_env_over {P Σ Γ n decl} (eq : nth_error Γ n = Some decl) {wfΓ : All_local_env (lift_typing typing Σ) Γ} : + Lemma nth_error_All_local_env_over {P Σ Γ n decl} (eq : nth_error Γ n = Some decl) {wfΓ : wf_local Σ Γ} : All_local_env_over typing P Σ Γ wfΓ -> let Γ' := skipn (S n) Γ in let p := lookup_wf_local_decl wfΓ n eq in diff --git a/pcuic/theories/PCUICUnivSubst.v b/pcuic/theories/PCUICUnivSubst.v index 77d68a0c9..857c69f5e 100644 --- a/pcuic/theories/PCUICUnivSubst.v +++ b/pcuic/theories/PCUICUnivSubst.v @@ -11,7 +11,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICInduction PCUICLiftSubst. Instance subst_instance_constr : UnivSubst term := fix subst_instance_constr u c {struct c} : term := match c with - | tRel _ | tVar _ => c + | tRel _ | tVar _ => c | tEvar ev args => tEvar ev (List.map (subst_instance_constr u) args) | tSort s => tSort (subst_instance_univ u s) | tConst c u' => tConst c (subst_instance_instance u u') @@ -32,6 +32,7 @@ Instance subst_instance_constr : UnivSubst term := | tCoFix mfix idx => let mfix' := List.map (map_def (subst_instance_constr u) (subst_instance_constr u)) mfix in tCoFix mfix' idx + | tPrim _ => c end. Instance subst_instance_decl : UnivSubst context_decl diff --git a/pcuic/theories/PCUICUnivSubstitution.v b/pcuic/theories/PCUICUnivSubstitution.v index 85b7082c4..a3e6e7679 100644 --- a/pcuic/theories/PCUICUnivSubstitution.v +++ b/pcuic/theories/PCUICUnivSubstitution.v @@ -10,14 +10,10 @@ From Equations Require Import Equations. (** * Universe Substitution lemmas for typing derivations. *) - Local Set Keyed Unification. Set Default Goal Selector "!". -Module CS := ConstraintSet. -Module LS := LevelSet. - Create HintDb univ_subst. Local Ltac aa := rdest; eauto with univ_subst. @@ -351,7 +347,7 @@ Proof. /\ In c' (CS.elements ctrs)). - generalize (CS.elements ctrs), CS.empty. induction l; cbn. - + firstorder. + + pcuicfo. now destruct H0 as [? ?]. + intros t. etransitivity. 1: eapply IHl. split; intros [HH|HH]. * destruct a as [[l1 a] l2]. apply CS.add_spec in HH. @@ -442,8 +438,8 @@ Qed. Context {cf : checker_flags}. Lemma consistent_instance_declared lvs φ uctx u : - consistent_instance lvs φ uctx u - -> forallb (fun l => LS.mem l lvs) u. + consistent_instance lvs φ uctx u -> + forallb (fun l => LS.mem l lvs) u. Proof. unfold consistent_instance. destruct uctx as [ctx|ctx]. 1: destruct u; [reflexivity|discriminate]. @@ -1089,9 +1085,11 @@ Proof. econstructor. now rewrite nth_error_map, H. - cbn. econstructor; eauto. eapply OnOne2_map. eapply OnOne2_impl. 1: eassumption. - firstorder. + (* Used to be pcuicfo *) + simpl in *; intuition; simpl in *. unfold on_Trel. + simpl. split; auto. - cbn; econstructor; - eapply OnOne2_map; eapply OnOne2_impl; [ eassumption | firstorder]. + eapply OnOne2_map; eapply OnOne2_impl; [ eassumption | pcuicfo]. - cbn; econstructor; eapply OnOne2_map; eapply OnOne2_impl; [ eassumption | ]. intros. destruct X1. destruct p. inversion e. destruct x, y; cbn in *; subst. @@ -1454,20 +1452,6 @@ Proof. destruct destInd as [[i u']|]; simpl; auto. Qed. -Axiom fix_guard_subst_instance : - forall mfix u, - fix_guard mfix -> - fix_guard (map (map_def (subst_instance_constr u) (subst_instance_constr u)) - mfix). - - -Axiom cofix_guard_subst_instance : - forall mfix u, - cofix_guard mfix -> - cofix_guard (map (map_def (subst_instance_constr u) (subst_instance_constr u)) - mfix). - - Lemma All_local_env_over_subst_instance Σ Γ (wfΓ : wf_local Σ Γ) : All_local_env_over typing (fun Σ0 Γ0 (_ : wf_local Σ0 Γ0) t T (_ : Σ0;;; Γ0 |- t : T) => @@ -1596,10 +1580,11 @@ Proof. + eapply X3; aa. - intros n b b_ty b' s1 b'_ty X X0 X1 X2 X3 X4 X5 u univs wfΣ' HSub H. econstructor; eauto. eapply X5; aa. - - intros t0 na A B u X X0 X1 X2 X3 u0 univs wfΣ' HSub H. + - intros t0 na A B s u X X0 X1 X2 X3 X4 X5 u0 univs wfΣ' HSub H. rewrite <- subst_subst_instance_constr. cbn. econstructor. + eapply X1; eauto. + eapply X3; eauto. + + eapply X5; eauto. - intros. rewrite subst_instance_constr_two. econstructor; [aa|aa|]. clear X X0; cbn in *. eapply consistent_ext_trans; eauto. @@ -1664,23 +1649,22 @@ Proof. - intros mfix n decl H H0 H1 X X0 wffix u univs wfΣ' HSub. erewrite map_dtype. econstructor. - + now apply fix_guard_subst_instance. + + now eapply fix_guard_subst_instance. + rewrite nth_error_map, H0. reflexivity. + eapply H1; eauto. + apply All_map, (All_impl X); simpl; intuition auto. destruct X1 as [s Hs]. exists (subst_instance_univ u s). now apply Hs. + eapply All_map, All_impl; tea. - intros x [[X1 X2] X3]. split. - * specialize (X3 u univs wfΣ' HSub H2). erewrite map_dbody in X3. - rewrite <- lift_subst_instance_constr in X3. - rewrite fix_context_length, map_length in *. - erewrite map_dtype with (d := x) in X3. - unfold subst_instance_context, map_context in *. - rewrite map_app in *. - rewrite <- (fix_context_subst_instance u mfix). - eapply X3. - * destruct x as [? ? []]; cbn in *; tea. + intros x [X1 X3]. + specialize (X3 u univs wfΣ' HSub H2). erewrite map_dbody in X3. + rewrite <- lift_subst_instance_constr in X3. + rewrite fix_context_length, map_length in *. + erewrite map_dtype with (d := x) in X3. + unfold subst_instance_context, map_context in *. + rewrite map_app in *. + rewrite <- (fix_context_subst_instance u mfix). + eapply X3. + red; rewrite <- wffix. unfold wf_fixpoint. rewrite map_map_compose. @@ -1688,7 +1672,7 @@ Proof. - intros mfix n decl guard H X X0 X1 wfcofix u univs wfΣ' HSub H1. erewrite map_dtype. econstructor; tas. - + now apply cofix_guard_subst_instance. + + now eapply cofix_guard_subst_instance. + rewrite nth_error_map, H. reflexivity. + apply X; eauto. + apply All_map, (All_impl X0); simpl; intuition auto. @@ -1889,74 +1873,6 @@ Section SubstIdentity. now rewrite mapi_nth. Qed. - Fixpoint unfold {A} (n : nat) (f : nat -> A) : list A := - match n with - | 0 => [] - | S n => unfold n f ++ [f n] - end. - - Lemma mapi_irrel_list {A B} (f : nat -> A) (l l' : list B) : - #|l| = #|l'| -> - mapi (fun i (x : B) => f i) l = mapi (fun i x => f i) l'. - Proof. - induction l in f, l' |- *; destruct l' => //; simpl; auto. - intros [= eq]. f_equal. - rewrite !mapi_rec_Sk. - now rewrite [mapi_rec _ _ _](IHl (fun x => (f (S x))) l'). - Qed. - - Lemma mapi_unfold {A B} (f : nat -> B) l : mapi (fun i (x : A) => f i) l = unfold #|l| f. - Proof. - induction l in f |- *; simpl; auto. - rewrite mapi_rec_Sk. - rewrite -IHl. rewrite -(mapi_rec_Sk (fun i x => f i) l 0). - change [f #|l|] with (mapi_rec (fun i x => f i) [a] #|l|). - rewrite -(Nat.add_0_r #|l|). rewrite -mapi_rec_app. - change (f 0 :: _) with (mapi (fun i x => f i) (a :: l)). - apply mapi_irrel_list. simpl. rewrite app_length /=; lia. - Qed. - - Lemma forallb_mapi {A B} (p : B -> bool) (f : nat -> B) l : - (forall i, i < #|l| -> p (f i)) -> - forallb p (mapi (fun i (x : A) => f i) l). - Proof. - intros Hp. rewrite (mapi_unfold f). - induction #|l| in *; simpl; auto. - rewrite forallb_app. simpl. now rewrite Hp // !andb_true_r. - Qed. - - Lemma In_unfold n i : In (Level.Var i) (unfold n Level.Var) -> i < n. - Proof. - induction n; simpl => //. - intros H; apply in_app_or in H. - destruct H. - - specialize (IHn H). lia. - - simpl in H. destruct H; [injection H|]. - * intros ->. auto. - * destruct H. - Qed. - - Lemma In_fold_right_add x l : - In x l <-> LevelSet.In x (fold_right LevelSet.add LevelSet.empty l). - Proof. - split. - - induction l; simpl => //. - intros [<-|H]. - * eapply LevelSet.add_spec; left; auto. - * eapply LevelSet.add_spec; right; auto. - - induction l; simpl => //. - * now rewrite LevelSetFact.empty_iff. - * rewrite LevelSet.add_spec. intuition auto. - Qed. - - Lemma CS_For_all_union f cst cst' : ConstraintSet.For_all f (ConstraintSet.union cst cst') -> - ConstraintSet.For_all f cst. - Proof. - unfold CS.For_all. - intros IH x inx. apply (IH x). - now eapply CS.union_spec; left. - Qed. - Lemma declared_inductive_wf_ext_wk Σ mdecl mind : wf Σ -> declared_minductive Σ mind mdecl -> @@ -1981,50 +1897,6 @@ Section SubstIdentity. Hint Resolve declared_inductive_wf_ext_wk declared_inductive_wf_global_ext : pcuic. - Instance For_all_proper P : Morphisms.Proper (CS.Equal ==> iff)%signature (ConstraintSet.For_all P). - Proof. - intros s s' eqs. - unfold CS.For_all. split; intros IH x inxs; apply (IH x); - now apply eqs. - Qed. - - Lemma unfold_length {A} (f : nat -> A) m : #|unfold m f| = m. - Proof. - induction m; simpl; rewrite ?app_length /=; auto. lia. - Qed. - - Lemma nth_error_unfold {A} (f : nat -> A) m n : n < m <-> nth_error (unfold m f) n = Some (f n). - Proof. - induction m in n |- *; split; intros Hn; try lia. - - simpl in Hn. rewrite nth_error_nil in Hn. discriminate. - - destruct (eq_dec n m); [subst|]. - * simpl. rewrite nth_error_app_ge unfold_length // Nat.sub_diag /= //. - * simpl. rewrite nth_error_app_lt ?unfold_length //; try lia. - apply IHm; lia. - - simpl in Hn. eapply nth_error_Some_length in Hn. - rewrite app_length /= unfold_length in Hn. lia. - Qed. - - Lemma nth_error_unfold_inv {A} (f : nat -> A) m n t : nth_error (unfold m f) n = Some t -> t = (f n). - Proof. - induction m in n |- *; intros Hn; try lia. - - simpl in Hn. rewrite nth_error_nil in Hn. discriminate. - - simpl in Hn. - pose proof (nth_error_Some_length Hn). - rewrite app_length /= unfold_length in H. - destruct (eq_dec n m); [subst|]. - * simpl. revert Hn. rewrite nth_error_app_ge unfold_length // Nat.sub_diag /= //; congruence. - * simpl. revert Hn. rewrite nth_error_app_lt ?unfold_length //; try lia. auto. - Qed. - - Lemma CS_For_all_add P x s : CS.For_all P (CS.add x s) -> P x /\ CS.For_all P s. - Proof. - intros. - split. - * apply (H x), CS.add_spec; left => //. - * intros y iny. apply (H y), CS.add_spec; right => //. - Qed. - Lemma subst_instance_level_abs l n Σ : wf Σ -> LevelSet.In l (LevelSet.union @@ -2035,9 +1907,9 @@ Section SubstIdentity. intros wfΣ lin. eapply LevelSet.union_spec in lin. destruct lin. - - apply In_fold_right_add in H. + - apply LevelSet_In_fold_right_add in H. destruct l; simpl; auto. - eapply In_unfold in H. + eapply In_unfold_inj in H; [|congruence]. pose proof (proj1 (nth_error_unfold Level.Var n n0) H). now rewrite (nth_error_nth _ _ _ H0). - eapply not_var_global_levels in wfΣ. @@ -2061,7 +1933,7 @@ Section SubstIdentity. apply LevelSet.mem_spec, LevelSet.union_spec. left. unfold levels_of_udecl. simpl. rewrite (mapi_unfold Level.Var). - eapply In_fold_right_add. + eapply LevelSet_In_fold_right_add. induction #|univs| in i, Hi |- *; try lia. simpl. eapply in_or_app. destruct (eq_dec i n). * subst. right; simpl; auto. @@ -2102,9 +1974,6 @@ Section SubstIdentity. eapply CS_For_all_union. Qed. - Lemma isType_closed {Σ Γ T} : wf Σ.1 -> isType Σ Γ T -> closedn #|Γ| T. - Proof. intros wfΣ [s Hs]. now eapply subject_closed in Hs. Qed. - Lemma udecl_prop_in_var_poly {Σ n} : on_udecl_prop Σ.1 Σ.2 -> LevelSet.In (Level.Var n) (levels_of_udecl Σ.2) -> ∑ ctx, Σ.2 = Polymorphic_ctx ctx. Proof. @@ -2128,7 +1997,7 @@ Section SubstIdentity. - destruct cu as [decl' [sizeu vc]]. clear sizeu vc. induction u; simpl; auto. - move/andP: decl' => [ina au]. specialize (IHu au). + move/andb_and: decl' => [ina au]. specialize (IHu au). rewrite [map _ u]IHu. f_equal. clear au. destruct a; simpl; auto. eapply LevelSet.mem_spec in ina. @@ -2136,8 +2005,8 @@ Section SubstIdentity. destruct (udecl_prop_in_var_poly onu ina) as [[univs csts] eq]. rewrite eq in IHu, ina |- *. simpl in *. rewrite mapi_unfold in IHu, ina |- *. - eapply In_fold_right_add in ina. - eapply In_unfold in ina. + eapply LevelSet_In_fold_right_add in ina. + eapply In_unfold_inj in ina; try congruence. eapply (nth_error_unfold Level.Var) in ina. now rewrite (nth_error_nth _ _ _ ina). Qed. @@ -2154,12 +2023,12 @@ Section SubstIdentity. eapply udecl_prop_in_var_poly in onu as [[ctx cstrs] eq]; eauto. rewrite eq. simpl. rewrite eq in cu. simpl in cu. - apply In_fold_right_add in cu. + apply LevelSet_In_fold_right_add in cu. unfold AUContext.repr in *. rewrite (mapi_unfold Level.Var) in cu |- *. destruct nth_error eqn:hnth. * apply nth_error_unfold_inv in hnth. subst; auto. * apply nth_error_None in hnth. rewrite unfold_length in hnth. - apply In_unfold in cu. lia. + apply In_unfold_inj in cu; try lia. congruence. Qed. Lemma consistent_instance_ext_subst_abs_univ Σ u : diff --git a/pcuic/theories/PCUICValidity.v b/pcuic/theories/PCUICValidity.v index 859eea3ac..dd01730c4 100644 --- a/pcuic/theories/PCUICValidity.v +++ b/pcuic/theories/PCUICValidity.v @@ -160,7 +160,7 @@ Section Validity. reflexivity. - (* Application *) - destruct X1 as [u' Hu']. exists u'. + destruct X3 as [u' Hu']. exists u'. move: (typing_wf_universe wf Hu') => wfu'. eapply (substitution0 _ _ na _ _ _ (tSort u')); eauto. apply inversion_Prod in Hu' as [na' [s1 [s2 Hs]]]; tas. intuition. @@ -267,12 +267,11 @@ Section Validity. assumption. - (* Fix *) - eapply nth_error_all in X0; eauto. - firstorder auto. + eapply nth_error_all in X0 as [s Hs]; eauto. + pcuic. - (* CoFix *) - eapply nth_error_all in X0; eauto. - firstorder auto. + eapply nth_error_all in X0 as [s Hs]; pcuic. - (* Conv *) now exists s. @@ -315,3 +314,13 @@ Proof. clear -H'' HA''' wfΣ. depind H''. econstructor; eauto. eapply cumul_trans; eauto. Qed. + +(** "Economical" typing rule for applications, not requiring to check the product type *) +Lemma type_App' {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ t na A B u} : + Σ;;; Γ |- t : tProd na A B -> + Σ;;; Γ |- u : A -> Σ;;; Γ |- tApp t u : B {0 := u}. +Proof. + intros Ht Hu. + have [s Hs] := validity_term wfΣ Ht. + eapply type_App; eauto. +Qed. \ No newline at end of file diff --git a/pcuic/theories/PCUICWcbvEval.v b/pcuic/theories/PCUICWcbvEval.v index 2fcb83e2f..772cbd61a 100644 --- a/pcuic/theories/PCUICWcbvEval.v +++ b/pcuic/theories/PCUICWcbvEval.v @@ -93,26 +93,6 @@ Definition isConstruct t := | _ => false end. -Definition isAssRel Γ x := - match x with - | tRel i => - match option_map decl_body (nth_error Γ i) with - | Some None => true - | _ => false - end - | _ => false - end. - -Definition isAxiom Σ x := - match x with - | tConst c u => - match lookup_env Σ c with - | Some (ConstantDecl {| cst_body := None |}) => true - | _ => false - end - | _ => false - end. - Definition substl defs body : term := fold_left (fun bod term => csubst term 0 bod) defs body. @@ -165,29 +145,12 @@ Section Wcbv. eval (csubst b0' 0 b1) res -> eval (tLetIn na b0 t b1) res - (** - (** Local variables: defined or undefined *) - | eval_rel_def i body res : - option_map decl_body (nth_error Γ i) = Some (Some body) -> - eval (lift0 (S i) body) res -> - eval (tRel i) res - - | eval_rel_undef i : - option_map decl_body (nth_error Γ i) = Some None -> - eval (tRel i) (tRel i) - *) - (** Constant unfolding *) | eval_delta c decl body (isdecl : declared_constant Σ c decl) u res : decl.(cst_body) = Some body -> eval (subst_instance_constr u body) res -> eval (tConst c u) res - (** Axiom *) - | eval_axiom c decl (isdecl : declared_constant Σ c decl) u : - decl.(cst_body) = None -> - eval (tConst c u) (tConst c u) - (** Case *) | eval_iota ind pars discr c u args p brs res : eval discr (mkApps (tConstruct ind c u) args) -> @@ -257,14 +220,13 @@ Section Wcbv. *) Definition value_head x := - isInd x || isConstruct x || isCoFix x || isAxiom Σ x. + isInd x || isConstruct x || isCoFix x. (* Lemma value_head_atom x : value_head x -> atom x. *) (* Proof. destruct x; auto. Qed. *) Inductive value : term -> Type := | value_atom t : atom t -> value t - (* | value_evar n l l' : Forall value l -> Forall value l' -> value (mkApps (tEvar n l) l') *) | value_app t l : value_head t -> All value l -> value (mkApps t l) | value_stuck_fix f args : All value args -> @@ -274,8 +236,6 @@ Section Wcbv. Lemma value_values_ind : forall P : term -> Type, (forall t, atom t -> P t) -> - (* (forall n l l', Forall value l -> Forall P l -> Forall value l' -> Forall P l' -> *) - (* P (mkApps (tEvar n l) l')) -> *) (forall t l, value_head t -> All value l -> All P l -> P (mkApps t l)) -> (forall f args, All value args -> @@ -335,11 +295,6 @@ Section Wcbv. Lemma eval_to_value e e' : eval e e' -> value e'. Proof. induction 1; simpl; auto using value. - - eapply (value_app (tConst c u) []). - red in isdecl. - rewrite /value_head /= isdecl. - destruct decl as [? [b|] ?]; try discriminate. - constructor. constructor. - change (tApp ?h ?a) with (mkApps h [a]). rewrite mkApps_nested. apply value_mkApps_inv in IHX1; [|easy]. @@ -443,15 +398,6 @@ Section Wcbv. move/(_ H) => Ht. induction l using rev_ind. simpl. destruct t; try discriminate. - (* * constructor. - unfold value_head in H. simpl in H. destruct option_map as [[o|]|] => //. *) - * unfold value_head in H. simpl in H. - destruct lookup_env eqn:Heq => //. - destruct g eqn:Heq' => //. - destruct c as [? [b|] ?] eqn:Heq'' => //. subst. - eapply eval_axiom. red. - rewrite Heq. reflexivity. - easy. * now eapply eval_atom. * now eapply eval_atom. * now eapply eval_atom. @@ -644,7 +590,7 @@ Section Wcbv. (* intros H; depind H; try solve_discr. *) (* - depelim H. *) (* - depelim H. *) - (* - eexists _, _; firstorder eauto. *) + (* - eexists _, _; pcuicfo eauto. *) (* - now depelim H. *) (* - discriminate. *) (* Qed. *) @@ -723,10 +669,6 @@ Section Wcbv. assert (e0 = e) as -> by now apply uip. assert (isdecl0 = isdecl) as -> by now apply uip. now specialize (IHev _ ev'); noconf IHev. - - depelim ev'; try go. - pose proof (PCUICWeakeningEnv.declared_constant_inj _ _ isdecl isdecl0) as <-. - assert (isdecl0 = isdecl) as -> by now apply uip. - now assert (e0 = e) as -> by now apply uip. - depelim ev'; try go. + specialize (IHev1 _ ev'1); noconf IHev1. apply (f_equal pr1) in IHev1 as apps_eq; cbn in *. @@ -861,13 +803,10 @@ Section Wcbv. ∑ decl, declared_constant Σ c decl * match cst_body decl with | Some body => eval (subst_instance_constr u body) v - | None => v = tConst c u + | None => False end. Proof. intros H; depelim H. - - exists decl. - split; [easy|]. - now rewrite e. - exists decl. split; [easy|]. now rewrite e. diff --git a/pcuic/theories/PCUICWeakening.v b/pcuic/theories/PCUICWeakening.v index c7c0021f8..e9f27ea0d 100644 --- a/pcuic/theories/PCUICWeakening.v +++ b/pcuic/theories/PCUICWeakening.v @@ -118,8 +118,6 @@ Proof. * destruct H2. rewrite H2. simpl. now rewrite Nat.sub_0_r. Qed. -Ltac lia_f_equal := repeat (lia || f_equal). - Lemma smash_context_lift Δ k n Γ : smash_context (lift_context n (k + #|Γ|) Δ) (lift_context n k Γ) = lift_context n k (smash_context Δ Γ). @@ -274,7 +272,7 @@ Proof. move: Hdecl. rewrite /closed_inductive_decl /lift_mutual_inductive_body. destruct decl; simpl. - move/andP => [clpar clbodies]. f_equal. + move/andb_and => [clpar clbodies]. f_equal. - now rewrite [fold_context _ _]closed_ctx_lift. - eapply forallb_All in clbodies. eapply Alli_mapi_id. @@ -282,7 +280,7 @@ Proof. intros. eapply H0. * simpl; intros. move: H0. rewrite /closed_inductive_body. - destruct x; simpl. move=> /andP[/andP [ci ct] cp]. + destruct x; simpl. move=> /andb_and[/andb_and [ci ct] cp]. f_equal. + rewrite lift_closed; eauto. eapply closed_upwards; eauto; lia. @@ -293,7 +291,7 @@ Proof. eapply closed_upwards; eauto; lia. + simpl in X. rewrite -X in cp. eapply forallb_All in cp. eapply All_map_id; eauto. - eapply (All_impl cp); firstorder auto. + eapply (All_impl cp); intuition auto. destruct x; unfold on_snd; simpl; f_equal. apply lift_closed. rewrite context_assumptions_fold. eapply closed_upwards; eauto; lia. @@ -475,7 +473,7 @@ Proof. 1: move=> //. move=> n0. rewrite /closedn_ctx !rev_app_distr /id /=. - move/andP => [closedx Hctx]. + move/andb_and => [closedx Hctx]. rewrite lift_decl_closed. - rewrite (@closed_decl_upwards n0) //; try lia. - f_equal. now rewrite IHctx. @@ -887,15 +885,6 @@ Proof. destruct t; simpl; try congruence. Qed. -Lemma nth_error_rev_inv {A} (l : list A) i : - i < #|l| -> - nth_error (List.rev l) i = nth_error l (#|l| - S i). -Proof. - intros Hi. - rewrite nth_error_rev; autorewrite with len; auto. - now rewrite List.rev_involutive. -Qed. - Lemma weakening_check_one_fix (Γ' Γ'' : context) mfix : map (fun x : def term => @@ -1066,16 +1055,14 @@ Proof. now specialize (Hs' _ _ _ wf eq_refl). * eapply All_map. eapply (All_impl X1); simpl. - intros x [[Hb Hlam] IH]. + intros x [Hb IH]. rewrite lift_fix_context. specialize (IH Γ (Γ' ,,, (fix_context mfix)) Γ'' wf). rewrite app_context_assoc in IH. specialize (IH eq_refl). - split; auto. - + rewrite lift_context_app Nat.add_0_r app_context_assoc in IH. - rewrite app_context_length fix_context_length in IH. - rewrite lift_context_length fix_context_length. - rewrite permute_lift; try lia. now rewrite (Nat.add_comm #|Γ'|). - + now rewrite isLambda_lift. + rewrite lift_context_app Nat.add_0_r app_context_assoc in IH. + rewrite app_context_length fix_context_length in IH. + rewrite lift_context_length fix_context_length. + rewrite permute_lift; try lia. now rewrite (Nat.add_comm #|Γ'|). * red; rewrite <-H1. unfold wf_fixpoint. rewrite map_map_compose. now rewrite weakening_check_one_fix. @@ -1187,7 +1174,7 @@ Proof. pose proof (closed_wf_local wfΣ (typing_wf_local ty)). specialize (X _ ty). eapply PCUICClosed.typecheck_closed in ty as [_ closed]; auto. - move/andP: closed => [ct cT]. + move/andb_and: closed => [ct cT]. rewrite !lift_closed // in X. now rewrite closed_ctx_lift in X. Qed. @@ -1212,11 +1199,11 @@ Proof. generalize (@nil context_decl) as Δ. rewrite /fix_context_gen. intros Δ wfΔ. - eapply All_local_env_app_inv. split; auto. + eapply All_local_env_app. split; auto. induction a in Δ, wfΔ |- *; simpl; auto. + constructor. + simpl. - eapply All_local_env_app_inv. split; auto. + eapply All_local_env_app. split; auto. * repeat constructor. simpl. destruct p as [s Hs]. diff --git a/pcuic/theories/PCUICWeakeningEnv.v b/pcuic/theories/PCUICWeakeningEnv.v index aa29c3b49..f10496b07 100644 --- a/pcuic/theories/PCUICWeakeningEnv.v +++ b/pcuic/theories/PCUICWeakeningEnv.v @@ -4,8 +4,6 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICEquality PCUICTyping. Require Import ssreflect. -Derive Signature for Alli. - Set Default Goal Selector "!". Lemma global_ext_constraints_app Σ Σ' φ @@ -64,9 +62,6 @@ Qed. Generalizable Variables Σ Γ t T. -Definition extends (Σ Σ' : global_env) := - { Σ'' & Σ' = Σ'' ++ Σ }. - Definition weaken_env_prop_full {cf:checker_flags} (P : global_env_ext -> context -> term -> term -> Type) := forall (Σ : global_env_ext) (Σ' : global_env), wf Σ' -> extends Σ.1 Σ' -> @@ -218,7 +213,7 @@ Proof. eapply extends_lookup in look; eauto. rewrite look //. - destruct (lookup_env Σ (inductive_mind i)) eqn:look => //. eapply extends_lookup in look; eauto. rewrite look //. -Qed. +Qed. (** The definition of [R_global_instance] is defined so that it is weakenable. *) Lemma R_global_instance_weaken_env {cf:checker_flags} Σ Σ' Re Re' Rle Rle' gr napp : @@ -417,7 +412,7 @@ Proof. intros ext wfΣ'. unfold wf_fixpoint. destruct map_option_out as [[|ind inds]|]; auto. - move/andP => [->] /=. + move/andb_and => [->] /=. now apply extends_check_recursivity_kind. Qed. @@ -427,7 +422,7 @@ Proof. intros ext wfΣ'. unfold wf_cofixpoint. destruct map_option_out as [[|ind inds]|]; auto. - move/andP => [->] /=. + move/andb_and => [->] /=. now apply extends_check_recursivity_kind. Qed. @@ -497,10 +492,12 @@ Proof. close_Forall. intros; intuition eauto with extends. destruct b as [s [Hs IH]]; eauto. - econstructor; eauto with extends. + + eapply fix_guard_extends; eauto. + eapply (All_impl X0); simpl; intuition eauto with extends. destruct X as [s Hs]; exists s. intuition eauto with extends. + eapply All_impl; eauto; simpl; intuition eauto with extends. - econstructor; eauto with extends. + + eapply cofix_guard_extends; eauto. + eapply (All_impl X0); simpl; intuition eauto with extends. destruct X as [s Hs]; exists s. intuition eauto with extends. + eapply All_impl; eauto; simpl; intuition eauto with extends. diff --git a/pcuic/theories/PCUICWfUniverses.v b/pcuic/theories/PCUICWfUniverses.v index d052f5bcd..2bf2fb4cc 100644 --- a/pcuic/theories/PCUICWfUniverses.v +++ b/pcuic/theories/PCUICWfUniverses.v @@ -13,16 +13,6 @@ Require Import ssreflect ssrbool. From MetaCoq.PCUIC Require Import PCUICInduction. -Lemma forallbP {A} (P : A -> Prop) (p : A -> bool) l : (forall x, reflect (P x) (p x)) -> reflect (Forall P l) (forallb p l). -Proof. - intros Hp. - apply: (iffP idP). - - induction l; rewrite /= //. move/andP => [pa pl]. - constructor; auto. now apply/(Hp _). - - induction 1; rewrite /= // IHForall // andb_true_r. - now apply/(Hp _). -Qed. - Section CheckerFlags. Context {cf:checker_flags}. @@ -269,7 +259,7 @@ Section CheckerFlags. rewrite mapi_unfold in wfx. eapply (proj1 (LevelSetProp.of_list_1 _ _)) in wfx. apply SetoidList.InA_alt in wfx as [? [<- wfx]]. simpl in wfx. - eapply In_unfold in wfx. + eapply In_unfold_inj in wfx; [|congruence]. destruct (nth_in_or_default n u (Level.lSet)). red in cu. eapply Forall_In in cu; eauto. rewrite e. red. eapply LS.union_spec. right. eapply global_levels_Set. @@ -284,7 +274,7 @@ Section CheckerFlags. | _ => true end. - Lemma wf_universe_reflect (u : Universe.t) : + Lemma wf_universe_reflect {u : Universe.t} : reflect (wf_universe Σ u) (wf_universeb u). Proof. destruct u; simpl; try now constructor. @@ -297,12 +287,6 @@ Section CheckerFlags. now eapply LS.mem_spec in H. Qed. - Lemma reflect_bP {P b} (r : reflect P b) : b -> P. - Proof. destruct r; auto. discriminate. Qed. - - Lemma reflect_Pb {P b} (r : reflect P b) : P -> b. - Proof. destruct r; auto. Qed. - Fixpoint wf_universes t := match t with | tSort s => wf_universeb s @@ -366,6 +350,7 @@ Section CheckerFlags. Qed. End WfUniverses. + Arguments wf_universe_reflect {Σ u}. Ltac to_prop := repeat match goal with @@ -375,8 +360,8 @@ Section CheckerFlags. Ltac to_wfu := repeat match goal with - | [ H: is_true (wf_universeb _ ?x) |- _ ] => apply (reflect_bP (wf_universe_reflect _ x)) in H - | [ |- is_true (wf_universeb _ ?x) ] => apply (reflect_Pb (wf_universe_reflect _ x)) + | [ H: is_true (wf_universeb _ ?x) |- _ ] => apply (elimT (@wf_universe_reflect _ x)) in H + | [ |- is_true (wf_universeb _ ?x) ] => apply (introT (@wf_universe_reflect _ x)) end. Lemma wf_universes_inst {Σ : global_env_ext} univs t u : @@ -580,7 +565,7 @@ Section CheckerFlags. rewrite wf_universes_subst. clear -s. generalize 0. induction args as [|[na [b|] ty] args] in sorts, s |- *; simpl in *; auto. - - destruct s as [? [[s' wf] [? ?]%MCProd.andP]]. + - destruct s as [? [[s' wf] [? ?]%andb_and]]. constructor; eauto. rewrite wf_universes_subst. eapply IHargs; eauto. now rewrite wf_universes_lift. @@ -692,7 +677,7 @@ Section CheckerFlags. - simpl in *; to_wfu; eauto with pcuic. - rewrite wf_universes_subst. constructor. to_wfu; auto. constructor. - now move/andP: H3 => []. + now move/andP: H4 => []. - apply/andP; split. { apply/wf_universe_instanceP. @@ -803,7 +788,7 @@ Section CheckerFlags. eapply wf_sorts_local_ctx_smash in s. eapply wf_sorts_local_ctx_nth_error in s as [? [? H]]; eauto. red in H. destruct x0. now move/andP: H => []. - now destruct H as [s [Hs _]%MCProd.andP]. + now destruct H as [s [Hs _]%andb_and]. - apply/andP; split; auto. solve_all. move:a => [s [Hty /andP[wfty wfs]]]. @@ -831,17 +816,18 @@ Section CheckerFlags. Σ ;;; Γ |- t : tSort s -> wf_universe Σ s. Proof. intros wfΣ Hty. - apply typing_wf_universes in Hty as [_ wfs]%MCProd.andP; auto. + apply typing_wf_universes in Hty as [_ wfs]%andb_and; auto. simpl in wfs. now to_wfu. Qed. Lemma isType_wf_universes {Σ Γ T} : wf Σ.1 -> isType Σ Γ T -> wf_universes Σ T. Proof. - intros wfΣ [s Hs]. now eapply typing_wf_universes in Hs as [HT _]%MCProd.andP. + intros wfΣ [s Hs]. now eapply typing_wf_universes in Hs as [HT _]%andb_and. Qed. End CheckerFlags. +Arguments wf_universe_reflect {Σ u}. Hint Resolve @wf_universe_type1 @wf_universe_super @wf_universe_sup @wf_universe_product : pcuic. Hint Extern 4 (wf_universe _ ?u) => diff --git a/pcuic/theories/TemplateToPCUIC.v b/pcuic/theories/TemplateToPCUIC.v index 1214cb324..a707b82c4 100644 --- a/pcuic/theories/TemplateToPCUIC.v +++ b/pcuic/theories/TemplateToPCUIC.v @@ -1,7 +1,20 @@ (* Distributed under the terms of the MIT license. *) +From Coq Require Import Int63 FloatOps FloatAxioms. From MetaCoq.Template Require Import config utils AstUtils. From MetaCoq.PCUIC Require Import PCUICAst. +Lemma to_Z_bounded_bool (i : Int63.int) : + ((0 <=? φ (i)%int63) && (φ (i)%int63 %Z.leb_le ->%Z.ltb_lt]. +Qed. + +Definition uint63_to_model (i : Int63.int) : uint63_model := + exist _ (Int63.to_Z i) (to_Z_bounded_bool i). + +Definition float64_to_model (f : PrimFloat.float) : float64_model := + exist _ (FloatOps.Prim2SF f) (FloatAxioms.Prim2SF_valid f). Fixpoint trans (t : Ast.term) : term := match t with @@ -27,6 +40,8 @@ Fixpoint trans (t : Ast.term) : term := | Ast.tCoFix mfix idx => let mfix' := List.map (map_def trans trans) mfix in tCoFix mfix' idx + | Ast.tInt n => tPrim (primInt; primIntModel (uint63_to_model n)) + | Ast.tFloat n => tPrim (primFloat; primFloatModel (float64_to_model n)) end. Definition trans_decl (d : Ast.context_decl) := diff --git a/pcuic/theories/TemplateToPCUICCorrectness.v b/pcuic/theories/TemplateToPCUICCorrectness.v index 87fd86d9c..18b010637 100644 --- a/pcuic/theories/TemplateToPCUICCorrectness.v +++ b/pcuic/theories/TemplateToPCUICCorrectness.v @@ -1,20 +1,20 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Template Require Import config utils Ast TypingWf WfInv. +From MetaCoq.Template Require TermEquality. +Set Warnings "-notation-overridden". From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICCumulativity PCUICLiftSubst PCUICEquality PCUICUnivSubst PCUICTyping TemplateToPCUIC - PCUICSubstitution PCUICGeneration. + PCUICWeakening PCUICSubstitution PCUICGeneration PCUICValidity. +Set Warnings "+notation-overridden". From Equations.Prop Require Import DepElim. +Implicit Types cf : checker_flags. - -Module T := Template.Ast. -Module TTy := Template.Typing. -Module TEnv := Template.Ast.TemplateEnvironment. - -Local Existing Instance default_checker_flags. - -Module TL := Template.LiftSubst. -Derive Signature for All. +(* Source = Template, Target (unqualified) = Coq *) +Module S := Template.Ast. +Module SEq := Template.TermEquality. +Module ST := Template.Typing. +Module SL := Template.LiftSubst. Lemma mkApps_morphism (f : term -> term) u v : (forall x y, f (tApp x y) = tApp (f x) (f y)) -> @@ -32,7 +32,7 @@ Ltac solve_list := try solve_all; try typeclasses eauto with core. Lemma trans_lift n k t : - trans (TL.lift n k t) = lift n k (trans t). + trans (SL.lift n k t) = lift n k (trans t). Proof. revert k. induction t using Template.Induction.term_forall_list_ind; simpl; intros; try congruence. - f_equal. rewrite !map_map_compose. solve_all. @@ -49,7 +49,7 @@ Proof. revert f l'; induction l; simpl; trivial. Qed. -Lemma trans_mkApp u a : trans (T.mkApp u a) = tApp (trans u) (trans a). +Lemma trans_mkApp u a : trans (S.mkApp u a) = tApp (trans u) (trans a). Proof. induction u; simpl; try reflexivity. rewrite map_app. @@ -58,8 +58,8 @@ Proof. rewrite mkApps_app. reflexivity. Qed. -Lemma trans_mkApps u v : T.wf u -> List.Forall T.wf v -> - trans (T.mkApps u v) = mkApps (trans u) (List.map trans v). +Lemma trans_mkApps u v : + trans (S.mkApps u v) = mkApps (trans u) (List.map trans v). Proof. revert u; induction v. simpl; trivial. @@ -67,16 +67,12 @@ Proof. rewrite <- Template.LiftSubst.mkApps_mkApp; auto. rewrite IHv. simpl. f_equal. apply trans_mkApp. - - apply Template.LiftSubst.wf_mkApp; auto. inversion_clear H0. auto. - inversion_clear H0. auto. Qed. -Lemma trans_subst t k u : All T.wf t -> T.wf u -> - trans (TL.subst t k u) = subst (map trans t) k (trans u). +Lemma trans_subst t k u : + trans (SL.subst t k u) = subst (map trans t) k (trans u). Proof. - intros wft wfu. - revert k. induction wfu using Template.Induction.term_wf_forall_list_ind; simpl; intros; try congruence. + revert k. induction u using Template.Induction.term_forall_list_ind; simpl; intros; try congruence. - repeat nth_leb_simpl; auto. rewrite trans_lift. @@ -85,11 +81,9 @@ Proof. - f_equal; solve_list. - - rewrite subst_mkApps. rewrite <- IHwfu. + - rewrite subst_mkApps. rewrite <- IHu. rewrite trans_mkApps. f_equal. solve_list. - apply Template.LiftSubst.wf_subst; auto. - solve_all. solve_all. apply Template.LiftSubst.wf_subst; auto. solve_all. - f_equal; auto; solve_list. - f_equal; auto; solve_list. @@ -114,10 +108,10 @@ Qed. Require Import ssreflect. Lemma forall_decls_declared_constant Σ cst decl : - TTy.declared_constant Σ cst decl -> + ST.declared_constant Σ cst decl -> declared_constant (trans_global_decls Σ) cst (trans_constant_body decl). Proof. - unfold declared_constant, TTy.declared_constant. + unfold declared_constant, ST.declared_constant. induction Σ => //; try discriminate. case: a => // /= k b. unfold eq_kername; destruct kername_eq_dec; subst; auto. @@ -125,10 +119,10 @@ Proof. Qed. Lemma forall_decls_declared_minductive Σ cst decl : - TTy.declared_minductive Σ cst decl -> + ST.declared_minductive Σ cst decl -> declared_minductive (trans_global_decls Σ) cst (trans_minductive_body decl). Proof. - unfold declared_minductive, TTy.declared_minductive. + unfold declared_minductive, ST.declared_minductive. induction Σ => //; try discriminate. case: a => // /= k b. unfold eq_kername; destruct kername_eq_dec; subst; auto. @@ -136,10 +130,10 @@ Proof. Qed. Lemma forall_decls_declared_inductive Σ mdecl ind decl : - TTy.declared_inductive Σ mdecl ind decl -> + ST.declared_inductive Σ mdecl ind decl -> declared_inductive (trans_global_decls Σ) (trans_minductive_body mdecl) ind (trans_one_ind_body decl). Proof. - unfold declared_inductive, TTy.declared_inductive. + unfold declared_inductive, ST.declared_inductive. move=> [decl' Hnth]. pose proof (forall_decls_declared_minductive _ _ _ decl'). eexists. eauto. destruct decl'; simpl in *. @@ -148,11 +142,11 @@ Proof. Qed. Lemma forall_decls_declared_constructor Σ cst mdecl idecl decl : - TTy.declared_constructor Σ mdecl idecl cst decl -> + ST.declared_constructor Σ mdecl idecl cst decl -> declared_constructor (trans_global_decls Σ) (trans_minductive_body mdecl) (trans_one_ind_body idecl) cst ((fun '(x, y, z) => (x, trans y, z)) decl). Proof. - unfold declared_constructor, TTy.declared_constructor. + unfold declared_constructor, ST.declared_constructor. move=> [decl' Hnth]. pose proof (forall_decls_declared_inductive _ _ _ _ decl'). split; auto. destruct idecl; simpl. @@ -160,11 +154,11 @@ Proof. Qed. Lemma forall_decls_declared_projection Σ cst mdecl idecl decl : - TTy.declared_projection Σ mdecl idecl cst decl -> + ST.declared_projection Σ mdecl idecl cst decl -> declared_projection (trans_global_decls Σ) (trans_minductive_body mdecl) (trans_one_ind_body idecl) cst ((fun '(x, y) => (x, trans y)) decl). Proof. - unfold declared_constructor, TTy.declared_constructor. + unfold declared_constructor, ST.declared_constructor. move=> [decl' [Hnth Hnpar]]. pose proof (forall_decls_declared_inductive _ _ _ _ decl'). split; auto. destruct idecl; simpl. @@ -179,18 +173,20 @@ Proof. Qed. Lemma trans_destArity ctx t : - T.wf t -> - match TTy.destArity ctx t with + S.wf t -> + match AstUtils.destArity ctx t with | Some (args, s) => destArity (trans_local ctx) (trans t) = Some (trans_local args, s) - | None => True + | None => destArity (trans_local ctx) (trans t) = None end. Proof. intros wf; revert ctx. induction wf using Template.Induction.term_wf_forall_list_ind; intros ctx; simpl; trivial. - apply (IHwf0 (T.vass n t :: ctx)). - apply (IHwf1 (T.vdef n t t0 :: ctx)). + apply (IHwf0 (S.vass n t :: ctx)). + apply (IHwf1 (S.vdef n t t0 :: ctx)). + destruct l. congruence. + now apply destArity_mkApps. Qed. (* TODO Duplicate? *) @@ -241,9 +237,9 @@ Qed. Definition on_pair {A B C D} (f : A -> B) (g : C -> D) (x : A * C) := (f (fst x), g (snd x)). -Lemma trans_inds kn u bodies : map trans (TTy.inds kn u bodies) = inds kn u (map trans_one_ind_body bodies). +Lemma trans_inds kn u bodies : map trans (ST.inds kn u bodies) = inds kn u (map trans_one_ind_body bodies). Proof. - unfold inds, TTy.inds. rewrite map_length. + unfold inds, ST.inds. rewrite map_length. induction bodies. simpl. reflexivity. simpl; f_equal. auto. Qed. @@ -251,7 +247,7 @@ Lemma trans_instantiate_params_subst params args s t : All TypingWf.wf_decl params -> All Ast.wf s -> All Ast.wf args -> forall s' t', - TTy.instantiate_params_subst params args s t = Some (s', t') -> + ST.instantiate_params_subst params args s t = Some (s', t') -> instantiate_params_subst (map trans_decl params) (map trans args) (map trans s) (trans t) = Some (map trans s', trans t'). @@ -265,7 +261,6 @@ Proof. erewrite <- IHparams. f_equal. 5:eauto. simpl. rewrite trans_subst; auto. inv Hparams. red in H. simpl in H. intuition auto. - now (inv Hparams). constructor; auto. inv Hparams; red in H; simpl in H; intuition auto. apply Template.LiftSubst.wf_subst; auto. solve_all. @@ -279,18 +274,18 @@ Proof. Qed. Lemma trans_instantiate_params params args t : - T.wf t -> + S.wf t -> All TypingWf.wf_decl params -> All Ast.wf args -> forall t', - TTy.instantiate_params params args t = Some t' -> + ST.instantiate_params params args t = Some t' -> instantiate_params (map trans_decl params) (map trans args) (trans t) = Some (trans t'). Proof. intros wft wfpars wfargs t' eq. revert eq. - unfold TTy.instantiate_params. - case_eq (TTy.instantiate_params_subst (List.rev params) args [] t). + unfold ST.instantiate_params. + case_eq (ST.instantiate_params_subst (List.rev params) args [] t). all: try discriminate. intros [ctx u] e eq. inversion eq. subst. clear eq. assert (wfargs' : Forall Ast.wf args) by (apply All_Forall ; assumption). @@ -305,8 +300,6 @@ Proof. rewrite map_rev in e. rewrite e. f_equal. symmetry. apply trans_subst. - - apply Forall_All. assumption. - - assumption. Qed. Lemma trans_it_mkProd_or_LetIn ctx t : @@ -337,58 +330,309 @@ Proof. reflexivity. Qed. -Lemma trans_decompose_prod_assum : +Lemma decompose_prod_assum_mkApps_nonnil ctx f args : + args <> [] -> + decompose_prod_assum ctx (mkApps f args) = (ctx, mkApps f args). +Proof. + induction args => //. + intros. + rewrite mkApps_nonempty //. +Qed. + +(* This is if Template's decompose_prod_assum handled casts. This does not work unless the specs + use a stronger decompose_prod_assum in PCUIC, as decompose_prod_assum is used to check fixpoints types. + Would be fixed having contexts instead of lambda abstractions in the Ast for fixpoints, where casts + cannot be introduced. +*) +(*Lemma trans_decompose_prod_assum : forall Γ t, - let '(Δ, c) := AstUtils.decompose_prod_assum Γ t in - decompose_prod_assum (trans_local Γ) (trans t) = (trans_local Δ, trans c). + let '(Δ, c) := AstUtils.decompose_prod_assum (S.map_context AstUtils.strip_casts Γ) (AstUtils.strip_casts t) in + decompose_prod_assum (trans_local (S.map_context AstUtils.strip_casts Γ)) (trans (AstUtils.strip_casts t)) = + (trans_local Δ, trans c). Proof. intros Γ t. destruct AstUtils.decompose_prod_assum as [Δ c] eqn:e. induction t in Γ, Δ, c, e |- *. (* all: simpl in *. *) all: try solve [ inversion e ; subst ; reflexivity ]. - - eapply IHt1 in e as e'. give_up. - - eapply IHt2 in e as e'. assumption. - - eapply IHt3 in e as e'. assumption. - - simpl in *. give_up. -Abort. - -Lemma trans_build_branches_type : - forall ind mdecl idecl args u p brtys, - map_option_out (TTy.build_branches_type ind mdecl idecl args u p) - = Some brtys -> - map_option_out - (build_branches_type ind (trans_minductive_body mdecl) - (trans_one_ind_body idecl) (map trans args) u - (trans p)) - = Some (map (on_snd trans) brtys). + - eapply IHt1 in e as e'. now simpl. + - simpl. simpl in e. eapply (IHt2 (Γ ,, Ast.vass na t1)) in e as e'. assumption. + - simpl. simpl in e. eapply (IHt3 (Γ ,, Ast.vdef na t1 t2)) in e as e'. assumption. + - simpl in *. rewrite trans_mkApps. + destruct args. + * simpl. now apply IHt. + * move: e. cbn [map]. + destruct (SL.mkApps_ex (AstUtils.strip_casts t) (AstUtils.strip_casts t0) (map AstUtils.strip_casts args)) as [f [? eq]]. + rewrite eq. intros [= <- <-]. rewrite -eq. + rewrite decompose_prod_assum_mkApps_nonnil //. + f_equal. + now rewrite trans_mkApps. +Qed. +*) + +Lemma trans_decompose_prod_assum : + forall Γ t, + S.wf t -> + let '(Δ, c) := AstUtils.decompose_prod_assum Γ t in + decompose_prod_assum (trans_local Γ) (trans t) = (trans_local Δ, trans c). +Proof. + intros Γ t wf. + destruct AstUtils.decompose_prod_assum as [Δ c] eqn:e. + revert Γ Δ c e. + induction wf using Template.Induction.term_wf_forall_list_ind; intros. + all: try solve [ inversion e ; subst ; reflexivity ]. + - simpl in e. eapply IHwf0 in e as e'. now simpl. + - simpl. simpl in e. eapply (IHwf1 (Γ ,, Ast.vdef n t t0)) in e as e'. assumption. + - simpl in *. noconf e. simpl. + destruct l => //. + cbn [map]. + rewrite decompose_prod_assum_mkApps_nonnil //. +Qed. + +(* Lemma trans_isApp t : Ast.isApp t = false -> PCUICAst.isApp (trans t) = false. +Proof. + destruct t => //. simpl. + now destruct prim as [? []]. +Qed. *) +(* +Lemma trans_decompose_app_args t f l : + Ast.wf t -> + AstUtils.decompose_app t = (f, l) -> + (decompose_app (trans t)).2 = map trans l. +Proof. + induction 1 using Template.Induction.term_wf_forall_list_ind; simpl; + intros [= <- <-]; auto. admit. + rewrite decompose_app_mkApps. + destruct (AstUtils.decompose_app t). +Admitted. *) + +Lemma trans_ind_params mdecl : trans_local (Ast.ind_params mdecl) = ind_params (trans_minductive_body mdecl). +Proof. reflexivity. Qed. + +Lemma trans_ind_bodies mdecl : map trans_one_ind_body (Ast.ind_bodies mdecl) = ind_bodies (trans_minductive_body mdecl). +Proof. reflexivity. Qed. +From MetaCoq.PCUIC Require Import PCUICClosed PCUICInductiveInversion. + +Lemma instantiate_params_spec params paramsi concl ty : + instantiate_params params paramsi (it_mkProd_or_LetIn params concl) = Some ty -> + ty = (subst (List.rev paramsi) 0 (expand_lets params concl)). +Proof. + intros h. eapply instantiate_params_make_context_subst in h as [ctx'' [ms [s' [da [dp ->]]]]]. + eapply make_context_subst_spec in dp. + rewrite List.rev_involutive in dp. + pose proof (PCUICContexts.context_subst_length2 dp). + rewrite decompose_prod_n_assum_it_mkProd app_nil_r in da. noconf da. + apply context_subst_subst_extended_subst in dp as ->. + rewrite map_subst_expand_lets ?List.rev_length //. +Qed. + +Lemma trans_ind_bodies_length mdecl : #|TemplateEnvironment.ind_bodies mdecl| = #|ind_bodies (trans_minductive_body mdecl)|. +Proof. simpl. now rewrite map_length. Qed. + +Lemma trans_ind_params_length mdecl : #|TemplateEnvironment.ind_params mdecl| = #|ind_params (trans_minductive_body mdecl)|. +Proof. simpl. now rewrite map_length. Qed. + +Lemma trans_ind_npars mdecl : Ast.ind_npars mdecl = ind_npars (trans_minductive_body mdecl). +Proof. simpl. reflexivity. Qed. + +Lemma trans_reln l p Γ : map trans (Ast.reln l p Γ) = + reln (map trans l) p (trans_local Γ). +Proof. + induction Γ as [|[na [b|] ty] Γ] in l, p |- *; simpl; auto. + now rewrite IHΓ. +Qed. + +Lemma trans_to_extended_list Γ n : map trans (Ast.to_extended_list_k Γ n) = to_extended_list_k (trans_local Γ) n. +Proof. + now rewrite /to_extended_list_k trans_reln. +Qed. + +Definition trans_constructor_shape (t : ST.constructor_shape) : constructor_shape := + {| cshape_args := trans_local t.(ST.cshape_args); + cshape_indices := map trans t.(ST.cshape_indices); + cshape_sorts := t.(ST.cshape_sorts) |}. + +Lemma trans_cshape_args cs : trans_local (ST.cshape_args cs) = cshape_args (trans_constructor_shape cs). +Proof. reflexivity. Qed. + +Lemma trans_cshape_args_length cs : #|ST.cshape_args cs| = #|cshape_args (trans_constructor_shape cs)|. +Proof. now rewrite -trans_cshape_args map_length. Qed. + +Lemma context_assumptions_map ctx : context_assumptions (map trans_decl ctx) = Ast.context_assumptions ctx. +Proof. + induction ctx as [|[na [b|] ty] ctx]; simpl; auto. +Qed. + +(* Lemma decompose_prod_assum_mkApps ctx t f l : + decompose_prod_assum ctx t = (ctx', l) -> + it_mkProd_or_LetIn ctx t = it_mkProd_or_LetIn ctx' (mkApps t' -> + () *) +(* +Lemma decompose_app_trans_inv t ind u l : + Ast.wf t -> + decompose_app (trans t) = (tInd ind u, l) -> + ∑ l', l = map trans l' /\ AstUtils.decompose_app t = (Ast.tInd ind u, l'). +Proof. + destruct t => //. simpl. + intros wf. + destruct decompose_app eqn:da. + pose proof (decompose_app_notApp _ _ _ da). + rewrite decompose_app_mkApps in da. depelim wf. destruct t => //. + destruct t0 => //. + intros [= -> -> ->]. + exists args. noconf da. split; auto. f_equal. rewrite -H0 in H. + depelim wf. + destruct t => //. simpl in H2. noconf H2. reflexivity. + intros. simpl in H0. noconf H0. + exists []. split; auto. +Qed. *) +Lemma trans_build_branches_type ind mdecl idecl args u p brtys : + (* (ST.on_inductive (ST.lift_typing ST.typing) Σ (inductive_mind ind) mdecl) -> + (ST.on_ind_body (ST.lift_typing ST.typing) Σ (inductive_mind ind) mdecl (inductive_ind ind) idecl) -> + inductive_ind ind < #|ind_bodies (trans_minductive_body mdecl)| -> *) + map_option_out (ST.build_branches_type ind mdecl idecl args u p) + = Some brtys -> + map_option_out + (build_branches_type ind (trans_minductive_body mdecl) + (trans_one_ind_body idecl) (map trans args) u + (trans p)) + = Some (map (on_snd trans) brtys). +Proof. Admitted. +(* Will likely change with the change of representation of cases *) +(* + intros onmind onind indlt. + pose proof (ST.onConstructors onind) as onc. + unfold ST.build_branches_type. + unfold build_branches_type. + rewrite mapi_map. + eapply All2_map_option_out_mapi_Some_spec. eapply onc. + intros i [[id ty] argn] [nargs br] s oncs. + simpl. + destruct ST.instantiate_params eqn:ipars => //. + pose proof ipars as ipars'. + apply trans_instantiate_params in ipars. + rewrite trans_subst trans_inds trans_subst_instance_constr in ipars. + rewrite [map _ _]trans_local_subst_instance_context trans_ind_params + trans_ind_bodies in ipars. + + rewrite trans_ind_params trans_ind_bodies. + rewrite ipars. + assert (wft : S.wf t). eapply wf_instantiate_params. 4:eapply ipars'. 1-3:admit. + pose proof (trans_decompose_prod_assum [] t wft). + destruct (AstUtils.decompose_prod_assum [] t) eqn:da'. + rewrite H. + destruct chop eqn:eqchop. + intros [= -> <-]. + destruct (chop _ (decompose_app (trans t0)).2) eqn:c'. + f_equal. f_equal. + rewrite [ty]oncs.(ST.cstr_eq) in ipars. + rewrite trans_it_mkProd_or_LetIn subst_instance_constr_it_mkProd_or_LetIn + subst_it_mkProd_or_LetIn in ipars. len in ipars. + rewrite closed_ctx_subst in ipars. + rewrite [map _ _]trans_ind_params. admit. + apply instantiate_params_spec in ipars. + rewrite trans_it_mkProd_or_LetIn. f_equal. + rewrite trans_mkApps map_length trans_lift. f_equal. + rewrite map_app /= trans_mkApps /= map_app trans_to_extended_list. + assert (l1 = map trans l /\ l2 = map trans l0) as [-> ->]. + { simpl in H. + move: ipars. + rewrite trans_it_mkProd_or_LetIn subst_instance_constr_it_mkProd_or_LetIn subst_it_mkProd_or_LetIn. + rewrite /expand_lets expand_lets_it_mkProd_or_LetIn. len. + rewrite subst_it_mkProd_or_LetIn. + rewrite trans_mkApps. + rewrite /ST.cstr_concl_head. cbn -[subst]. + rewrite trans_ind_bodies. len. + rewrite trans_ind_bodies_length trans_ind_params_length. + rewrite map_app trans_to_extended_list trans_ind_params trans_cshape_args_length. + rewrite (subst_cstr_concl_head ind u (trans_minductive_body mdecl)) //. + rewrite expand_lets_k_mkApps subst_mkApps /=. + intros eqit. + rewrite eqit in H. + rewrite decompose_prod_assum_it_mkProd // in H. + apply is_ind_app_head_mkApps. + injection H. clear H. rewrite app_nil_r. + intros eqind _. + apply (f_equal decompose_app) in eqind. + destruct (decompose_app (trans t0)) eqn:dt0. + simpl in *. + rewrite decompose_app_mkApps // in eqind. noconf eqind. + + noconf H. + move: H. + + rewrite !map_app map_map_compose chop_n_app. len. + rewrite context_assumptions_map. now rewrite onmind.(ST.onNpars). + destruct AstUtils.decompose_prod_assum eqn:da'. + rewrite trans_it_mkProd_or_LetIn trans_mkApps. len. + simpl in H. rewrite H. + + + + reflexivity. + + + + + + f_equal. f_equal. rewrite map_app. + rewrite trans_to_extended_list. f_equal. + + + + len. + + + apply instantiate_params_make_context_subst in ipars as [ctx' [ty'' [s' [dass [makes eq]]]]]. + rewrite eq. + eapply make_context_subst_spec in makes. + rewrite List.rev_involutive in makes. + + + + move: (trans_decompose_prod_assum [] t wft). + destruct oncs. simpl in cstr_eq. + rewrite trans_inds in ipars. + rewrite trans_subst_instance_constr in ipars. + intros [= -> <-]. + + destruct AstUtils.decompose_prod_assum eqn:dp => -> /=. + destruct (AstUtils.decompose_app t0) eqn:da'. simpl. + destruct chop as [params argsl] eqn:ceq. + + rewrite (chop_map _ _ _ _ _ ceq). + intros [= ->]. f_equal. unfold on_snd. simpl. f_equal. subst br. + rewrite trans_it_mkProd_or_LetIn trans_mkApps map_app /= trans_mkApps /= map_app. + now rewrite trans_to_extended_list trans_lift map_length. +Qed. +*) Lemma trans_build_case_predicate_type ind mdecl idecl params u ps : build_case_predicate_type ind (trans_minductive_body mdecl) (trans_one_ind_body idecl) (map trans params) u ps - = option_map trans (TTy.build_case_predicate_type ind mdecl idecl params u ps). + = option_map trans (ST.build_case_predicate_type ind mdecl idecl params u ps). Admitted. -(* Lemma trans_types_of_case (Σ : T.global_env) ind mdecl idecl args p u pty indctx pctx ps btys : *) -(* T.wf p -> T.wf pty -> T.wf (T.ind_type idecl) -> *) -(* All T.wf args -> *) -(* TTy.wf Σ -> *) -(* TTy.declared_inductive Σ mdecl ind idecl -> *) -(* TTy.types_of_case ind mdecl idecl args u p pty = Some (indctx, pctx, ps, btys) -> *) +(* Lemma trans_types_of_case (Σ : S.global_env) ind mdecl idecl args p u pty indctx pctx ps btys : *) +(* S.wf p -> S.wf pty -> S.wf (S.ind_type idecl) -> *) +(* All S.wf args -> *) +(* ST.wf Σ -> *) +(* ST.declared_inductive Σ mdecl ind idecl -> *) +(* ST.types_of_case ind mdecl idecl args u p pty = Some (indctx, pctx, ps, btys) -> *) (* types_of_case ind (trans_minductive_body mdecl) (trans_one_ind_body idecl) *) (* (map trans args) u (trans p) (trans pty) = *) (* Some (trans_local indctx, trans_local pctx, ps, map (on_snd trans) btys). *) (* Proof. *) (* intros wfp wfpty wfdecl wfargs wfΣ Hidecl. *) (* pose proof (on_declared_inductive wfΣ Hidecl) as [onmind onind]. *) -(* apply TTy.onParams in onmind as Hparams. *) +(* apply ST.onParams in onmind as Hparams. *) (* (* Maybe have a lemma for this we do it all the time *) *) -(* assert (wc : Forall wf_decl (UnivSubst.subst_instance_context u (T.ind_params mdecl))). *) -(* { assert (h : Forall wf_decl (T.ind_params mdecl)). *) +(* assert (wc : Forall wf_decl (UnivSubst.subst_instance_context u (S.ind_params mdecl))). *) +(* { assert (h : Forall wf_decl (S.ind_params mdecl)). *) (* { eapply typing_all_wf_decl ; revgoals. *) -(* - apply TTy.onParams in onmind. *) -(* unfold TTy.on_context in onmind. *) +(* - apply ST.onParams in onmind. *) +(* unfold ST.on_context in onmind. *) (* eassumption. *) (* - simpl. assumption. *) (* } *) @@ -402,11 +646,11 @@ Admitted. (* { eapply closed_wf_local ; eauto. eauto. } *) (* assert (wfparams : All wf_decl (Ast.ind_params mdecl)). *) (* { apply Forall_All. eapply typing_all_wf_decl ; eauto. simpl. eauto. } *) -(* unfold TTy.types_of_case, types_of_case. simpl. *) +(* unfold ST.types_of_case, types_of_case. simpl. *) (* match goal with *) -(* | |- context [ TTy.instantiate_params ?p ?a ?t ] => *) +(* | |- context [ ST.instantiate_params ?p ?a ?t ] => *) (* pose proof (trans_instantiate_params p a t) as ht ; *) -(* case_eq (TTy.instantiate_params p a t) *) +(* case_eq (ST.instantiate_params p a t) *) (* end. *) (* 2: discriminate. *) (* intros ity e. *) @@ -419,16 +663,16 @@ Admitted. (* apply wf_instantiate_params in e as wfity. *) (* all: auto with wf. *) (* pose proof (trans_destArity [] ity wfity) as e'. *) -(* destruct TTy.destArity as [[ctx s] | ]. 2: discriminate. *) +(* destruct ST.destArity as [[ctx s] | ]. 2: discriminate. *) (* rewrite e'. *) (* pose proof (trans_destArity [] pty wfpty) as e''. *) -(* destruct TTy.destArity as [[ctx' s'] | ]. 2: discriminate. *) +(* destruct ST.destArity as [[ctx' s'] | ]. 2: discriminate. *) (* simpl in e''. rewrite e''. *) -(* pose proof (TTy.onConstructors onind) as onc. *) +(* pose proof (ST.onConstructors onind) as onc. *) (* assert ( *) (* hb : *) (* forall brtys, *) -(* map_option_out (TTy.build_branches_type ind mdecl idecl args u p) *) +(* map_option_out (ST.build_branches_type ind mdecl idecl args u p) *) (* = Some brtys -> *) (* map_option_out *) (* (build_branches_type ind (trans_minductive_body mdecl) *) @@ -437,18 +681,18 @@ Admitted. (* = Some (map (on_snd trans) brtys) *) (* ). *) (* { intro brtys. *) -(* unfold TTy.build_branches_type, build_branches_type. *) +(* unfold ST.build_branches_type, build_branches_type. *) (* unfold trans_one_ind_body. simpl. rewrite -> mapi_map. *) -(* unfold TTy.on_constructors in onc. *) +(* unfold ST.on_constructors in onc. *) (* eapply All2_map_option_out_mapi_Some_spec. *) (* - eapply onc. *) (* - intros i [[id t] n] [t0 ar] z. *) (* unfold on_snd. simpl. *) (* intros [ont [cs ?]]. simpl in *. *) -(* unfold TTy.on_type in ont. simpl in ont. *) +(* unfold ST.on_type in ont. simpl in ont. *) (* destruct ont. *) -(* unfold TTy.instantiate_params, instantiate_params. *) -(* destruct TTy.instantiate_params_subst eqn:he. 2: discriminate. *) +(* unfold ST.instantiate_params, instantiate_params. *) +(* destruct ST.instantiate_params_subst eqn:he. 2: discriminate. *) (* destruct p0 as [s0 ty]. *) (* apply instantiate_params_subst_make_context_subst in he as he'. *) (* destruct he' as [ctx'' hctx'']. *) @@ -514,11 +758,11 @@ Admitted. (* reflexivity. *) (* Admitted. *) -Hint Constructors T.wf : wf. +Hint Constructors S.wf : wf. Hint Resolve Template.TypingWf.typing_wf : wf. -Lemma mkApps_trans_wf U l : T.wf (T.tApp U l) -> exists U' V', trans (T.tApp U l) = tApp U' V'. +Lemma mkApps_trans_wf U l : S.wf (S.tApp U l) -> exists U' V', trans (S.tApp U l) = tApp U' V'. Proof. simpl. induction l using rev_ind. intros. inv H. congruence. intros. rewrite map_app. simpl. exists (mkApps (trans U) (map trans l)), (trans x). @@ -527,9 +771,9 @@ Proof. rewrite mkApps_app. simpl. reflexivity. Qed. -Derive Signature for TTy.eq_term_upto_univ. +Derive Signature for SEq.eq_term_upto_univ_napp. -Lemma leq_term_mkApps Σ ϕ t u t' u' : +Lemma leq_term_mkApps {cf} Σ ϕ t u t' u' : eq_term Σ ϕ t t' -> All2 (eq_term Σ ϕ) u u' -> leq_term Σ ϕ (mkApps t u) (mkApps t' u'). Proof. @@ -572,82 +816,175 @@ Proof. * assumption. Qed. +Lemma trans_lookup Σ cst : + lookup_env (trans_global_decls Σ) cst = option_map trans_global_decl (Ast.lookup_env Σ cst). +Proof. + cbn in *. + induction Σ. + - reflexivity. + - cbn. + unfold eq_kername in *; destruct kername_eq_dec; subst. + + destruct a; auto. + + now rewrite IHΣ. +Qed. + +Lemma on_global_env_impl `{checker_flags} Σ P Q : + (forall Σ Γ t T, ST.on_global_env P Σ.1 -> P Σ Γ t T -> Q Σ Γ t T) -> + ST.on_global_env P Σ -> ST.on_global_env Q Σ. +Proof. + intros X X0. + simpl in *. induction X0; constructor; auto. + clear IHX0. destruct d; simpl. + - destruct c; simpl. destruct cst_body; simpl in *. + red in o |- *. simpl in *. now eapply X. + red in o |- *. simpl in *. now eapply X. + - red in o. simpl in *. + destruct o0 as [onI onP onNP]. + constructor; auto. + -- eapply Alli_impl. exact onI. eauto. intros. + refine {| ST.ind_indices := X1.(ST.ind_indices); + ST.ind_arity_eq := X1.(ST.ind_arity_eq); + ST.ind_cshapes := X1.(ST.ind_cshapes) |}. + --- apply ST.onArity in X1. unfold on_type in *; simpl in *. + now eapply X. + --- pose proof X1.(ST.onConstructors) as X11. red in X11. + eapply All2_impl; eauto. + simpl. intros. destruct X2 as [? ? ? ?]; unshelve econstructor; eauto. + * apply X; eauto. + * clear -X0 X on_cargs. revert on_cargs. + generalize (ST.cshape_args y), (ST.cshape_sorts y). + induction c; destruct l; simpl; auto; + destruct a as [na [b|] ty]; simpl in *; auto; + split; intuition eauto. + * clear -X0 X on_cindices. + revert on_cindices. + generalize (List.rev (SL.lift_context #|ST.cshape_args y| 0 (ST.ind_indices X1))). + generalize (ST.cshape_indices y). + induction 1; simpl; constructor; auto. + --- simpl; intros. pose (ST.onProjections X1 H0). simpl in *; auto. + --- destruct X1. simpl. unfold ST.check_ind_sorts in *. + destruct Universe.is_prop, Universe.is_sprop; auto. + split. + * apply ind_sorts. + * destruct indices_matter; auto. + eapply ST.type_local_ctx_impl. eapply ind_sorts. auto. + --- eapply X1.(ST.onIndices). + -- red in onP. red. + eapply ST.All_local_env_impl. eauto. + intros. now apply X. +Qed. + +Lemma typing_wf_wf {cf}: + forall (Σ : Template.Ast.global_env), + ST.wf Σ -> + ST.Forall_decls_typing + (fun (_ : Template.Ast.global_env_ext) (_ : Tcontext) (t T : Tterm) => Ast.wf t /\ Ast.wf T) Σ. +Proof. + intros Σ. + eapply on_global_env_impl. clear. + intros Σ Γ t S. + red. unfold ST.lift_typing. + intros ong. destruct S. + * intros ty. now eapply typing_wf. + * intros [s ty]. exists s. + now eapply typing_wf in ty. +Qed. + +Lemma trans_R_global_instance {cf} Σ Re Rle gref napp u u' : + ST.wf Σ -> + SEq.R_global_instance Σ Re Rle gref napp u u' -> + R_global_instance (trans_global_decls Σ) Re Rle gref napp u u'. +Proof. + intros wfe. + unfold PCUICEquality.R_global_instance, PCUICEquality.global_variance. + unfold SEq.R_global_instance, SEq.global_variance. + destruct gref; simpl; auto. + - unfold PCUICEquality.lookup_inductive, PCUICEquality.lookup_minductive. + unfold SEq.lookup_inductive, SEq.lookup_minductive. + rewrite trans_lookup. destruct Ast.lookup_env eqn:look => //; simpl. + destruct g => /= //. + rewrite nth_error_map. + destruct nth_error eqn:hnth => /= //. + assert (wfty : Ast.wf (Ast.ind_type o)). + { eapply declared_inductive_wf; eauto. eapply typing_wf_wf; eauto. split; eauto. } + move: (trans_destArity [] (Ast.ind_type o) wfty). + destruct AstUtils.destArity as [[ctx ps]|] eqn:eq => /= // -> //. + now rewrite context_assumptions_map. + - unfold PCUICEquality.lookup_constructor, PCUICEquality.lookup_inductive, PCUICEquality.lookup_minductive. + unfold SEq.lookup_constructor, SEq.lookup_inductive, SEq.lookup_minductive. + rewrite trans_lookup. destruct Ast.lookup_env => //; simpl. + destruct g => /= //. rewrite nth_error_map. + destruct nth_error => /= //. + rewrite nth_error_map. + destruct nth_error => /= //. + destruct p as [[id t] nargs]; simpl. + destruct Nat.leb => //. +Qed. + +Ltac tc := typeclasses eauto. + (* TODO update Template Coq's eq_term to reflect PCUIC's cumulativity *) -Lemma trans_eq_term_upto_univ : +Lemma trans_eq_term_upto_univ {cf} : forall Σ Re Rle t u napp, - T.wf t -> - T.wf u -> - TTy.eq_term_upto_univ Re Rle t u -> - eq_term_upto_univ_napp Σ Re Rle napp (trans t) (trans u). + RelationClasses.subrelation Re Rle -> + ST.wf Σ -> + S.wf t -> + S.wf u -> + SEq.eq_term_upto_univ_napp Σ Re Rle napp t u -> + eq_term_upto_univ_napp (trans_global_decls Σ) Re Rle napp (trans t) (trans u). Proof. -Admitted. -(* intros Σ Re Rle t u napp wt wu e. - induction t using Induction.term_forall_list_rect in Rle, napp, wt, u, wu, e |- *. + intros Σ Re Rle t u napp sub wfΣ wt wu e. + induction t using Induction.term_forall_list_rect in sub, Rle, napp, wt, u, wu, e |- *. all: invs e; cbn. all: try solve [ constructor ; auto ]. + all: repeat (match goal with + | H : S.wf (_ _) |- _ => apply wf_inv in H; simpl in H + | H : _ /\ _ |- _ => destruct H + end). all: try solve [ - repeat constructor ; + repeat constructor ; auto ; match goal with | ih : forall Rle (u : Tterm) (napp : nat), _ |- _ => - eapply ih ; [ - inversion wt ; assumption - | inversion wu ; assumption - | assumption - ] + now eapply ih end ]. - constructor. - assert (w1 : All T.wf l). - { eapply Forall_All. inversion wt. assumption. } - assert (w2 : All T.wf args'). - { eapply Forall_All. inversion wu. assumption. } + apply Forall_All in wt. + apply Forall_All in wu. pose proof (All2_All_mix_left X X0) as h1. simpl in h1. - pose proof (All2_All_mix_left w1 h1) as h2. - pose proof (All2_All_mix_right w2 h2) as h3. + pose proof (All2_All_mix_left wt h1) as h2. + pose proof (All2_All_mix_right wu h2) as h3. simpl in h3. eapply All2_map. eapply All2_impl. 1: exact h3. simpl. intros ? ? [[? [ih ?]] ?]. simpl in *. - eapply ih. all: auto. - -s apply eq_term_eq_term_napp. - + eapply IHt. - * inversion wt. assumption. - * inversion wu. assumption. - * assumption. + eapply ih. all: auto. tc. + - eapply eq_term_upto_univ_napp_mkApps. + + rewrite map_length. now eapply IHt. + pose proof (All2_All_mix_left X X1) as h. simpl in h. - assert (wl : Forall T.wf l). - { inversion wt. assumption. } - assert (wargs' : Forall T.wf args'). - { inversion wu. assumption. } - apply Forall_All in wl. - apply Forall_All in wargs'. - pose proof (All2_All_mix_left wl h) as h1. - pose proof (All2_All_mix_right wargs' h1) as h2. + apply Forall_All in H6. + apply Forall_All in H2. + pose proof (All2_All_mix_left H6 h) as h1. + pose proof (All2_All_mix_right H2 h1) as h2. simpl in h2. eapply All2_map. eapply All2_impl. 1: exact h2. simpl. intros u v [[? [ih ?]] ?]. - eapply ih. all: auto. - - constructor. todounivs. - - constructor. todounivs. + eapply ih. all: auto; tc. + - constructor. apply trans_R_global_instance; auto. + - constructor. apply trans_R_global_instance; auto. - constructor. all: try solve [ match goal with | ih : forall Rle u, _ |- _ => - eapply ih ; [ - inversion wt ; assumption - | inversion wu ; assumption - | assumption - ] + now eapply ih end ]. - assert (wl : All (T.wf ∘ snd) l). - { eapply Forall_All. inversion wt. assumption. } - assert (wbrs' : All (T.wf ∘ snd) brs'). - { eapply Forall_All. inversion wu. assumption. } + assert (wl : All (S.wf ∘ snd) l) by solve_all. + assert (wbrs' : All (S.wf ∘ snd) brs') by solve_all. pose proof (All2_All_mix_left X X2) as h1. simpl in h1. pose proof (All2_All_mix_left wl h1) as h2. pose proof (All2_All_mix_right wbrs' h2) as h3. @@ -656,26 +993,23 @@ Admitted. eapply All2_impl. 1: exact h3. simpl. intros [n u] [m v] [[? [ih [? ?]]] ?]. simpl in *. - intuition eauto. + intuition eauto. now apply ih. - constructor. assert ( w1 : All (fun def => - T.wf (dtype def) /\ - T.wf (dbody def) /\ - T.isLambda (dbody def) = true + S.wf (dtype def) /\ + S.wf (dbody def) ) m ). - { eapply Forall_All. inversion wt. assumption. } + { solve_all. } assert ( w2 : All (fun def => - T.wf (dtype def) /\ - T.wf (dbody def) /\ - T.isLambda (dbody def) = true - ) mfix' + S.wf (dtype def) /\ + S.wf (dbody def)) mfix' ). - { eapply Forall_All. inversion wu. assumption. } + { solve_all. } pose proof (All2_All_mix_left X X0) as h1. simpl in h1. pose proof (All2_All_mix_left w1 h1) as h2. pose proof (All2_All_mix_right w2 h2) as h3. @@ -683,20 +1017,20 @@ Admitted. eapply All2_map. eapply All2_impl. 1: exact h3. simpl. - intros [? ? ? ?] [? ? ? ?] [[[? [? ?]] [[ih1 ih2] [? [? ?]]]] [? [? ?]]]. + intros [? ? ? ?] [? ? ? ?] [[[? ?] [[ih1 ih2] [? ?]]] [? ?]]. simpl in *. - intuition eauto. + intuition eauto. now eapply ih1. now eapply ih2. - constructor. assert ( w1 : - All (fun def => T.wf (dtype def) /\ T.wf (dbody def)) m + All (fun def => S.wf (dtype def) /\ S.wf (dbody def)) m ). - { eapply Forall_All. inversion wt. assumption. } + { by eapply Forall_All. } assert ( w2 : - All (fun def => T.wf (dtype def) /\ T.wf (dbody def)) mfix' + All (fun def => S.wf (dtype def) /\ S.wf (dbody def)) mfix' ). - { eapply Forall_All. inversion wu. assumption. } + { by eapply Forall_All. } pose proof (All2_All_mix_left X X0) as h1. simpl in h1. pose proof (All2_All_mix_left w1 h1) as h2. pose proof (All2_All_mix_right w2 h2) as h3. @@ -704,35 +1038,38 @@ Admitted. eapply All2_map. eapply All2_impl. 1: exact h3. simpl. - intros [? ? ? ?] [? ? ? ?] [[[? ?] [[ih1 ih2] [? [? ?]]]] [? ?]]. + intros [? ? ? ?] [? ? ? ?] [[[? ?] [[ih1 ih2] [? ?]]] [? ?]]. simpl in *. - intuition eauto. -Qed. *) + intuition eauto. now eapply ih1. now eapply ih2. +Qed. -Lemma trans_leq_term Σ ϕ T U : - T.wf T -> T.wf U -> TTy.leq_term ϕ T U -> - leq_term Σ ϕ (trans T) (trans U). +Lemma trans_leq_term {cf} Σ ϕ T U : + ST.wf Σ -> + S.wf T -> S.wf U -> SEq.leq_term Σ ϕ T U -> + leq_term (trans_global_decls Σ) ϕ (trans T) (trans U). Proof. intros HT HU H. - eapply trans_eq_term_upto_univ ; eauto. + eapply trans_eq_term_upto_univ ; eauto. tc. Qed. -Lemma trans_eq_term Σ φ t u : - T.wf t -> T.wf u -> TTy.eq_term φ t u -> - eq_term Σ φ (trans t) (trans u). +Lemma trans_eq_term {cf} Σ φ t u : + ST.wf Σ -> + S.wf t -> S.wf u -> SEq.eq_term Σ φ t u -> + eq_term (trans_global_decls Σ) φ (trans t) (trans u). Proof. intros HT HU H. - eapply trans_eq_term_upto_univ ; eauto. + eapply trans_eq_term_upto_univ ; eauto. tc. Qed. -Lemma trans_eq_term_list : +Lemma trans_eq_term_list {cf}: forall Σ φ l l', - List.Forall T.wf l -> - List.Forall T.wf l' -> - All2 (TTy.eq_term φ) l l' -> - All2 (eq_term Σ φ) (List.map trans l) (List.map trans l'). + ST.wf Σ -> + List.Forall S.wf l -> + List.Forall S.wf l' -> + All2 (SEq.eq_term Σ φ) l l' -> + All2 (eq_term (trans_global_decls Σ) φ) (List.map trans l) (List.map trans l'). Proof. - intros Σ φ l l' w w' h. + intros Σ φ l l' wfΣ w w' h. eapply All2_map. apply Forall_All in w. apply Forall_All in w'. pose proof (All2_All_mix_left w h) as h1. @@ -742,12 +1079,12 @@ Proof. intuition auto using trans_eq_term. Qed. -(* Lemma wf_mkApps t u : T.wf (T.mkApps t u) -> List.Forall T.wf u. *) +(* Lemma wf_mkApps t u : S.wf (S.mkApps t u) -> List.Forall S.wf u. *) (* Proof. *) (* induction u in t |- *; simpl. *) (* - intuition. *) (* - intros H; destruct t; try solve [inv H; intuition auto]. *) -(* specialize (IHu (T.tApp t (l ++ [a]))). *) +(* specialize (IHu (S.tApp t (l ++ [a]))). *) (* forward IHu. *) (* induction u; trivial. *) (* simpl. rewrite <- app_assoc. simpl. apply H. *) @@ -762,25 +1099,18 @@ Proof. simpl in *. congruence. Qed. -Lemma trans_iota_red pars ind c u args brs : - T.wf (Template.Ast.mkApps (Template.Ast.tConstruct ind c u) args) -> - List.Forall (T.wf ∘ snd) brs -> - trans (TTy.iota_red pars c args brs) = +Lemma trans_iota_red pars c args brs : + trans (ST.iota_red pars c args brs) = iota_red pars c (List.map trans args) (List.map (on_snd trans) brs). Proof. - unfold TTy.iota_red, iota_red. intros wfapp wfbrs. + unfold ST.iota_red, iota_red. rewrite trans_mkApps. - - - induction wfbrs in c |- *. - destruct c; simpl; constructor. - destruct c; simpl; try constructor; auto with wf. - - apply wf_mkApps_napp in wfapp. solve_all. apply All_skipn. solve_all. constructor. - - f_equal. induction brs in c |- *; simpl; destruct c; trivial. + f_equal. induction brs in c |- *; simpl; destruct c; trivial. now rewrite map_skipn. Qed. Lemma trans_isLambda t : - T.wf t -> isLambda (trans t) = Ast.isLambda t. + S.wf t -> isLambda (trans t) = Ast.isLambda t. Proof. destruct 1; cbnr. destruct u; [contradiction|]. cbn. @@ -789,23 +1119,18 @@ Proof. Qed. Lemma trans_unfold_fix mfix idx narg fn : - List.Forall (fun def : def Tterm => T.wf (dtype def) /\ T.wf (dbody def) /\ - T.isLambda (dbody def) = true) mfix -> - TTy.unfold_fix mfix idx = Some (narg, fn) -> + List.Forall (fun def : def Tterm => S.wf (dtype def) /\ S.wf (dbody def)) mfix -> + ST.unfold_fix mfix idx = Some (narg, fn) -> unfold_fix (map (map_def trans trans) mfix) idx = Some (narg, trans fn). Proof. - unfold TTy.unfold_fix, unfold_fix. intros wffix. + unfold ST.unfold_fix, unfold_fix. intros wffix. rewrite nth_error_map. destruct (nth_error mfix idx) eqn:Hdef => //. cbn. - destruct (Ast.isLambda (dbody d)); [|discriminate]. intros [= <- <-]. simpl. repeat f_equal. - rewrite trans_subst. clear Hdef. - unfold TTy.fix_subst. generalize mfix at 2. - induction mfix0. constructor. simpl. repeat (constructor; auto). - apply (nth_error_forall Hdef) in wffix. simpl in wffix; intuition. + rewrite trans_subst. f_equal. clear Hdef. - unfold fix_subst, TTy.fix_subst. rewrite map_length. + unfold fix_subst, ST.fix_subst. rewrite map_length. generalize mfix at 1 3. induction wffix; trivial. simpl; intros mfix. f_equal. @@ -813,19 +1138,16 @@ Proof. Qed. Lemma trans_unfold_cofix mfix idx narg fn : - List.Forall (fun def : def Tterm => T.wf (dtype def) /\ T.wf (dbody def)) mfix -> - TTy.unfold_cofix mfix idx = Some (narg, fn) -> + List.Forall (fun def : def Tterm => S.wf (dtype def) /\ S.wf (dbody def)) mfix -> + ST.unfold_cofix mfix idx = Some (narg, fn) -> unfold_cofix (map (map_def trans trans) mfix) idx = Some (narg, trans fn). Proof. - unfold TTy.unfold_cofix, unfold_cofix. intros wffix. + unfold ST.unfold_cofix, unfold_cofix. intros wffix. rewrite nth_error_map. destruct (nth_error mfix idx) eqn:Hdef. intros [= <- <-]. simpl. repeat f_equal. - rewrite trans_subst. clear Hdef. - unfold TTy.cofix_subst. generalize mfix at 2. - induction mfix0. constructor. simpl. repeat (constructor; auto). - apply (nth_error_forall Hdef) in wffix. simpl in wffix; intuition. + rewrite trans_subst. f_equal. clear Hdef. - unfold cofix_subst, TTy.cofix_subst. rewrite map_length. + unfold cofix_subst, ST.cofix_subst. rewrite map_length. generalize mfix at 1 3. induction wffix; trivial. simpl; intros mfix. f_equal. @@ -836,10 +1158,10 @@ Definition isApp t := match t with tApp _ _ => true | _ => false end. Lemma trans_is_constructor: forall (args : list Tterm) (narg : nat), - TTy.is_constructor narg args = true -> is_constructor narg (map trans args) = true. + ST.is_constructor narg args = true -> is_constructor narg (map trans args) = true. Proof. intros args narg. - unfold TTy.is_constructor, is_constructor. + unfold ST.is_constructor, is_constructor. rewrite nth_error_map. destruct nth_error. simpl. intros. destruct t; try discriminate || reflexivity. simpl in H. destruct t; try discriminate || reflexivity. @@ -859,20 +1181,30 @@ Proof. Qed. Ltac wf_inv H := try apply wf_inv in H; simpl in H; rdest. -Lemma trans_red1 Σ Γ T U : - TTy.on_global_env (fun Σ => wf_decl_pred) Σ -> +Lemma wf_wf_decl_pred {cf} Σ : ST.wf Σ -> Typing.on_global_env (fun _ => wf_decl_pred) Σ. +Proof. + move/typing_wf_wf. + apply on_global_env_impl. + intros. + destruct T; simpl in *; auto. + destruct X0 as [s [Ht Hs]]. + red. split; auto. +Qed. +Hint Resolve wf_wf_decl_pred : wf. + +Lemma trans_red1 {cf} Σ Γ T U : + ST.wf Σ -> List.Forall wf_decl Γ -> - T.wf T -> - TTy.red1 Σ Γ T U -> + S.wf T -> + ST.red1 Σ Γ T U -> red1 (map (on_snd trans_global_decl) Σ) (trans_local Γ) (trans T) (trans U). Proof. intros wfΣ wfΓ Hwf. - induction 1 using TTy.red1_ind_all; wf_inv Hwf; simpl in *; + induction 1 using ST.red1_ind_all; wf_inv Hwf; simpl in *; try solve [econstructor; eauto]. - simpl. wf_inv H1. apply Forall_All in H2. inv H2. - rewrite trans_mkApps; auto. apply Template.LiftSubst.wf_subst; auto with wf; solve_all. - apply All_Forall. auto. + rewrite trans_mkApps; auto. rewrite trans_subst; auto. apply red1_mkApps_l. constructor. - rewrite trans_subst; eauto. repeat constructor. @@ -886,8 +1218,9 @@ Proof. - rewrite trans_mkApps; eauto with wf; simpl. erewrite trans_iota_red; eauto. repeat constructor. - - simpl. eapply red_fix. wf_inv H3. - now apply trans_unfold_fix; eauto. + - simpl. rewrite trans_mkApps. + eapply red_fix. wf_inv H3. + apply trans_unfold_fix; eauto. now apply trans_is_constructor. - apply wf_mkApps_napp in H1; auto. @@ -920,7 +1253,7 @@ Proof. intuition eauto. apply b2. all: solve_all. - - rewrite !trans_mkApps; auto with wf. eapply wf_red1 in X; auto. + - rewrite !trans_mkApps; auto with wf. eapply wf_red1 in X; auto with wf. apply red1_mkApps_l. auto. - apply Forall_All in H2. clear H H0 H1. revert M1. induction X. @@ -928,7 +1261,7 @@ Proof. apply red1_mkApps_l. apply app_red_r. auto. inv H2. specialize (IHX X0). simpl. intros. - eapply (IHX (T.tApp M1 [hd])). + eapply (IHX (S.tApp M1 [hd])). - constructor. apply IHX. constructor; hnf; simpl; auto. auto. - constructor. induction X; simpl; repeat constructor. apply p; auto. now inv Hwf. @@ -954,7 +1287,7 @@ Proof. red. rewrite <- !map_dtype. rewrite <- !map_dbody. intuition eauto. + unfold Template.Ast.app_context, trans_local in b. simpl in a. rewrite -> map_app in b. - unfold app_context. unfold TTy.fix_context in b. + unfold app_context. unfold ST.fix_context in b. rewrite map_rev map_mapi in b. simpl in b. unfold fix_context. rewrite mapi_map. simpl. forward b. @@ -989,7 +1322,7 @@ Proof. red. rewrite <- !map_dtype. rewrite <- !map_dbody. intuition eauto. + unfold Template.Ast.app_context, trans_local in b. simpl in a. rewrite -> map_app in b. - unfold app_context. unfold TTy.fix_context in b. + unfold app_context. unfold ST.fix_context in b. rewrite map_rev map_mapi in b. simpl in b. unfold fix_context. rewrite mapi_map. simpl. forward b. @@ -1008,23 +1341,28 @@ Proof. + simpl. inversion b0. reflexivity. Qed. +Lemma global_levels_trans Σ + : global_levels (trans_global_decls Σ) = ST.global_levels Σ. +Proof. + induction Σ; simpl; auto. + rewrite IHΣ. f_equal. clear. + destruct a as [? []]; reflexivity. +Qed. + Lemma global_ext_levels_trans Σ - : global_ext_levels (trans_global Σ) = TTy.global_ext_levels Σ. + : global_ext_levels (trans_global Σ) = ST.global_ext_levels Σ. Proof. destruct Σ. - unfold trans_global, TTy.global_ext_levels, global_ext_levels; simpl. + unfold trans_global, ST.global_ext_levels, global_ext_levels; simpl. f_equal. clear u. - induction g. - - reflexivity. - - simpl. rewrite IHg. f_equal. clear. - destruct a as [? []]; reflexivity. + now rewrite global_levels_trans. Qed. Lemma global_ext_constraints_trans Σ - : global_ext_constraints (trans_global Σ) = TTy.global_ext_constraints Σ. + : global_ext_constraints (trans_global Σ) = ST.global_ext_constraints Σ. Proof. destruct Σ. - unfold trans_global, TTy.global_ext_constraints, global_ext_constraints; simpl. + unfold trans_global, ST.global_ext_constraints, global_ext_constraints; simpl. f_equal. clear u. induction g. - reflexivity. @@ -1032,53 +1370,59 @@ Proof. destruct a as [? []]; reflexivity. Qed. -Lemma trans_cumul (Σ : Ast.global_env_ext) Γ T U : - TTy.on_global_env (fun Σ => wf_decl_pred) Σ -> +Lemma trans_cumul {cf} (Σ : Ast.global_env_ext) Γ T U : + ST.wf Σ -> List.Forall wf_decl Γ -> - T.wf T -> T.wf U -> TTy.cumul Σ Γ T U -> + S.wf T -> S.wf U -> ST.cumul Σ Γ T U -> trans_global Σ ;;; trans_local Γ |- trans T <= trans U. Proof. - intros wfΣ wfΓ. - induction 3. constructor; auto. - eapply trans_leq_term in l; eauto. - now rewrite global_ext_constraints_trans. - pose proof r as H3. apply wf_red1 in H3; auto. - apply trans_red1 in r; auto. econstructor 2; eauto. - econstructor 3. - apply IHX; auto. apply wf_red1 in r; auto. - apply trans_red1 in r; auto. + intros wfΣ. + induction 4. + - constructor; auto. + eapply trans_leq_term in l; eauto. + now rewrite global_ext_constraints_trans. + - pose proof r as H3. apply wf_red1 in H3; auto with wf. + apply trans_red1 in r; auto. econstructor 2; eauto with wf. + - econstructor 3. + apply IHX; auto. apply wf_red1 in r; auto with wf. + apply trans_red1 in r; auto. Qed. Definition Tlift_typing (P : Template.Ast.global_env_ext -> Tcontext -> Tterm -> Tterm -> Type) := fun Σ Γ t T => match T with | Some T => P Σ Γ t T - | None => { s : Universe.t & P Σ Γ t (T.tSort s) } + | None => { s : Universe.t & P Σ Γ t (S.tSort s) } end. -Lemma trans_wf_local: - forall (Σ : Template.Ast.global_env_ext) (Γ : Tcontext) (wfΓ : TTy.wf_local Σ Γ), +Definition TTy_wf_local {cf : checker_flags} Σ Γ := ST.All_local_env (ST.lift_typing ST.typing Σ) Γ. + +Lemma trans_wf_local {cf}: + forall (Σ : Template.Ast.global_env_ext) (Γ : Tcontext) (wfΓ : TTy_wf_local Σ Γ), let P := (fun Σ0 (Γ0 : Tcontext) _ (t T : Tterm) _ => + wf (trans_global Σ0).1 -> trans_global Σ0;;; trans_local Γ0 |- trans t : trans T) in - TTy.All_local_env_over TTy.typing P Σ Γ wfΓ -> + wf (trans_global Σ).1 -> + ST.All_local_env_over ST.typing P Σ Γ wfΓ -> wf_local (trans_global Σ) (trans_local Γ). Proof. intros. - induction X. + induction X0. - simpl. constructor. - simpl. econstructor. - + eapply IHX. - + simpl. destruct tu. exists x. eapply p. - - simpl. constructor; auto. red. destruct tu. exists x. auto. + + eapply IHX0. + + simpl. destruct tu. exists x. now eapply p. + - simpl. constructor; auto. red. destruct tu. exists x; auto. + simpl. eauto. Qed. -Lemma trans_wf_local_env Σ Γ : - TTy.All_local_env - (TTy.lift_typing +Lemma trans_wf_local_env {cf} Σ Γ : + ST.All_local_env + (ST.lift_typing (fun (Σ : Ast.global_env_ext) (Γ : Tcontext) (b ty : Tterm) => - TTy.typing Σ Γ b ty × trans_global Σ;;; trans_local Γ |- trans b : trans ty) Σ) + ST.typing Σ Γ b ty × trans_global Σ;;; trans_local Γ |- trans b : trans ty) Σ) Γ -> wf_local (trans_global Σ) (trans_local Γ). Proof. @@ -1092,31 +1436,7 @@ Proof. red. red in t1. destruct t1. eapply t2. Qed. -Lemma typing_wf_wf: - forall (Σ : Template.Ast.global_env_ext), - TTy.wf Σ -> - TTy.Forall_decls_typing - (fun (_ : Template.Ast.global_env_ext) (_ : Tcontext) (t T : Tterm) => Ast.wf t /\ Ast.wf T) Σ. -Proof. - intros Σ wf. - red. unfold TTy.lift_typing. - induction wf. 1: constructor. - constructor. all: auto. - red. red in o0. - destruct d. - - hnf. destruct c as [ty [d|] ?] => /=. - + red in o0. simpl in *. eapply typing_wf. 2: eassumption. - simpl. auto. - + red. simpl in *. destruct o0 as [s ?]. - exists s. simpl in *. eapply typing_wf. 2: eassumption. auto. - - simpl in *. destruct o0. constructor. all: auto. - (* + eapply Alli_impl. 1: eassumption. - intros ? [? ? ? ? ?] []. simpl in *. - econstructor. all: eauto. *) - (* General implication lemma would be nicer: Simon *) -Admitted. - -(* eapply TTy.on_global_decls_impl. 2:eapply wf. eauto. intros * ongl ont. destruct t. *) +(* eapply ST.on_global_decls_impl. 2:eapply wf. eauto. intros * ongl ont. destruct t. *) (* red in ont. *) (* eapply typing_wf; eauto. destruct ont. exists x; eapply typing_wf; intuition eauto. *) (* Qed. *) @@ -1124,35 +1444,212 @@ Admitted. Hint Resolve trans_wf_local : trans. Axiom fix_guard_trans : - forall mfix, - TTy.fix_guard mfix -> - fix_guard (map (map_def trans trans) mfix). - -Lemma isWFArity_wf (Σ : Ast.global_env_ext) Γ T : Typing.wf Σ -> TTy.isWfArity TTy.typing Σ Γ T -> T.wf T. -Proof. - intros wfΣ []. - destruct s as [s [eq wf]]. - generalize (destArity_spec [] T). rewrite eq. - simpl. move => ->. - apply (it_mkProd_or_LetIn_wf Γ). - rewrite -TEnv.it_mkProd_or_LetIn_app. - eapply wf_it_mkProd_or_LetIn. instantiate (1:=wf). - induction wf; constructor; auto. - destruct t0. eapply typing_wf; eauto. - eapply typing_wf; eauto. simpl. - destruct t0. - eapply typing_wf; eauto. constructor. -Qed. - -Theorem template_to_pcuic (Σ : T.global_env_ext) Γ t T : - TTy.wf Σ -> - TTy.typing Σ Γ t T -> - typing (trans_global Σ) (trans_local Γ) (trans t) (trans T). + forall Σ Γ mfix, + ST.fix_guard Σ Γ mfix -> + fix_guard (trans_global Σ) (trans_local Γ) (map (map_def trans trans) mfix). + +Axiom cofix_guard_trans : + forall Σ Γ mfix, + ST.cofix_guard Σ Γ mfix -> + cofix_guard (trans_global Σ) (trans_local Γ) (map (map_def trans trans) mfix). + +Notation Swf_fix def := (S.wf (dtype def) /\ S.wf (dbody def)). + +Lemma trans_subst_context s k Γ : + trans_local (SL.subst_context s k Γ) = subst_context (map trans s) k (trans_local Γ). Proof. - intros X X0. - pose proof (TTy.typing_wf_local X0). - revert Σ X Γ X1 t T X0. - apply (TTy.typing_ind_env (fun Σ Γ t T => + induction Γ as [|[na [b|] ty] Γ]; simpl; auto. + - rewrite SL.subst_context_snoc /=. rewrite [subst_context _ _ _ ]subst_context_snoc. + f_equal; auto. rewrite IHΓ /snoc /subst_decl /map_decl /=; f_equal. + now rewrite !trans_subst map_length. + - rewrite SL.subst_context_snoc /=. rewrite [subst_context _ _ _ ]subst_context_snoc. + f_equal; auto. rewrite IHΓ /snoc /subst_decl /map_decl /=; f_equal. + now rewrite !trans_subst map_length. +Qed. + +Lemma trans_smash_context Γ Δ : trans_local (ST.smash_context Γ Δ) = smash_context (trans_local Γ) (trans_local Δ). +Proof. + induction Δ in Γ |- *; simpl; auto. + destruct a as [na [b|] ty] => /=. + rewrite IHΔ. f_equal. + now rewrite (trans_subst_context [_]). + rewrite IHΔ. f_equal. rewrite /trans_local map_app //. +Qed. + +Lemma trans_decompose_app {t ind u l} : + S.wf t -> + AstUtils.decompose_app t = (Ast.tInd ind u, l) -> + ∑ l', decompose_app (trans t) = (tInd ind u, l'). +Proof. + intros wft. destruct (AstUtils.decompose_app t) eqn:da. + pose proof (TypingWf.decompose_app_inv _ [] _ _ wft da) as [n [napp [eqsk ->]]]. + rewrite trans_mkApps. + intros [= -> ->]. + rewrite decompose_app_mkApps //. eexists _; eauto. +Qed. + +Lemma decompose_prod_assum_it_mkProd_or_LetIn ctx t ctx' t' : + AstUtils.decompose_prod_assum ctx t = (ctx', t') -> + Ast.it_mkProd_or_LetIn ctx t = Ast.it_mkProd_or_LetIn ctx' t'. +Proof. + induction t in ctx, ctx', t' |- *; simpl; try intros [= -> <-]; auto. + - intros H. now apply IHt2 in H. + - intros H. now apply IHt3 in H. +Qed. + +Lemma it_mkProd_or_LetIn_wf Γ t + : Ast.wf (Ast.it_mkProd_or_LetIn Γ t) <-> Forall wf_decl Γ /\ Ast.wf t . +Proof. + revert t. induction Γ. + - intros t. simpl. split => //; try split => //; trivial. now intros []. + - intros t. + destruct a as [na [b|] ty]; simpl in *. rewrite /Ast.mkProd_or_LetIn /=. + * rewrite IHΓ /=. split; intros []. + depelim H0. intuition constructor; auto. split; auto. + depelim H. red in H. simpl in H. split; auto with wf. constructor; intuition auto. + * rewrite IHΓ /=. split; intros []. + depelim H0. simpl in H0_. split; [constructor|];auto. + split; simpl; auto. + depelim H. destruct H as [_ wfty]. simpl in wfty. + split; auto. constructor; simpl; auto. +Qed. + +Lemma trans_check_one_fix mfix ind : + Swf_fix mfix -> + ST.check_one_fix mfix = Some ind -> + check_one_fix (map_def trans trans mfix) = Some ind. +Proof. + unfold ST.check_one_fix, check_one_fix. + case: mfix => [na ty arg rarg] /= [wfty wfarg]. + move: (trans_decompose_prod_assum [] ty wfty) => /=. + destruct AstUtils.decompose_prod_assum as [ctx p] eqn:dp => /= // ->. + rewrite -(trans_smash_context []) /trans_local. + rewrite -List.map_rev nth_error_map. + destruct nth_error eqn:hnth => /= //. + destruct AstUtils.decompose_app eqn:da. + destruct t => //. + have [| l' eq] := (trans_decompose_app _ da). + { eapply decompose_prod_assum_it_mkProd_or_LetIn in dp. + simpl in dp; subst ty. + eapply it_mkProd_or_LetIn_wf in wfty as [wfctx _]. + eapply (nth_error_forall (P:=wf_decl)) in hnth; cycle 1. + eapply rev_Forall. + pose proof (nth_error_Some_length hnth). + rewrite nth_error_rev // List.rev_involutive in hnth. + now apply wf_smash_context. apply hnth. } + destruct c => /=. rewrite eq /= //. +Qed. + +Lemma Forall_map_option_out_map_Some_spec {A B} {f g : A -> option B} {P : A -> Prop} {l} : + (forall x t, P x -> f x = Some t -> g x = Some t) -> + Forall P l -> + forall t, + map_option_out (map f l) = Some t -> + map_option_out (map g l) = Some t. +Proof. + move => Hfg Hl. move: Hl. + induction 1; try constructor; auto. + simpl. move=> //. + case E: (f x) => [b|] //. + rewrite (Hfg _ _ H E). + case E' : map_option_out => [b'|] //. + move=> t [= <-]. now rewrite (IHHl _ E'). +Qed. + +Lemma map_option_out_check_one_fix {mfix} : + Forall (fun def => (S.wf (dtype def) /\ S.wf (dbody def))) mfix -> + forall l, + map_option_out (map (fun x => ST.check_one_fix x) mfix) = Some l -> + map_option_out (map (fun x => check_one_fix (map_def trans trans x)) mfix) = Some l. +Proof. + eapply Forall_map_option_out_map_Some_spec => x kn. + apply trans_check_one_fix. +Qed. + +Lemma trans_check_one_cofix mfix ind : + Swf_fix mfix -> + ST.check_one_cofix mfix = Some ind -> + check_one_cofix (map_def trans trans mfix) = Some ind. +Proof. + unfold ST.check_one_cofix, check_one_cofix. + case: mfix => [na ty arg rarg] /= [wfty wfarg]. + move: (trans_decompose_prod_assum [] ty wfty) => /=. + destruct AstUtils.decompose_prod_assum as [ctx p] eqn:dp => /= // ->. + destruct AstUtils.decompose_app eqn:da. + destruct t => //. + have [| l' ->] := (trans_decompose_app _ da) => //. + { eapply decompose_prod_assum_it_mkProd_or_LetIn in dp. + simpl in dp; subst ty. + now eapply it_mkProd_or_LetIn_wf in wfty as [_ wfp]. } +Qed. + +Lemma map_option_out_check_one_cofix {mfix} : + Forall (fun def => (S.wf (dtype def) /\ S.wf (dbody def))) mfix -> + forall l, + map_option_out (map (fun x => ST.check_one_cofix x) mfix) = Some l -> + map_option_out (map (fun x => check_one_cofix (map_def trans trans x)) mfix) = Some l. +Proof. + apply Forall_map_option_out_map_Some_spec => x kn. + apply trans_check_one_cofix. +Qed. + +Lemma trans_check_rec_kind {Σ k f} : + ST.check_recursivity_kind Σ k f = check_recursivity_kind (trans_global_decls Σ) k f. +Proof. + unfold ST.check_recursivity_kind, check_recursivity_kind. + rewrite trans_lookup. + destruct S.lookup_env as [[]|] => //. +Qed. + +Lemma trans_wf_fixpoint Σ mfix : + Forall (fun def => (S.wf (dtype def) /\ S.wf (dbody def))) mfix -> + ST.wf_fixpoint Σ mfix -> + wf_fixpoint (trans_global_decls Σ) (map (map_def trans trans) mfix). +Proof. + unfold ST.wf_fixpoint, wf_fixpoint. + rewrite map_map_compose. + intros wf. + destruct map_option_out eqn:ma => //. + rewrite (map_option_out_check_one_fix wf _ ma). + destruct l; auto. now rewrite -trans_check_rec_kind. +Qed. + +Lemma trans_wf_cofixpoint Σ mfix : + Forall (fun def => (S.wf (dtype def) /\ S.wf (dbody def))) mfix -> + ST.wf_cofixpoint Σ mfix -> + wf_cofixpoint (trans_global_decls Σ) (map (map_def trans trans) mfix). +Proof. + unfold ST.wf_cofixpoint, wf_cofixpoint. + rewrite map_map_compose. + intros wf. + destruct map_option_out eqn:ma => //. + rewrite (map_option_out_check_one_cofix wf _ ma). + destruct l; auto. now rewrite -trans_check_rec_kind. +Qed. + +Lemma trans_global_decl_universes d : + ST.universes_decl_of_decl d = + universes_decl_of_decl (trans_global_decl d). +Proof. + destruct d; reflexivity. +Qed. + +Lemma trans_consistent_instance_ext {cf:checker_flags} Σ d u : + ST.consistent_instance_ext Σ (ST.universes_decl_of_decl d) u -> + consistent_instance_ext (trans_global Σ) (universes_decl_of_decl (trans_global_decl d)) u. +Proof. + unfold ST.consistent_instance_ext, consistent_instance_ext. + rewrite global_ext_levels_trans global_ext_constraints_trans trans_global_decl_universes. + trivial. +Qed. + +Theorem template_to_pcuic {cf} : + ST.env_prop (fun Σ Γ t T => + wf (trans_global Σ).1 -> + typing (trans_global Σ) (trans_local Γ) (trans t) (trans T)). +Proof. + apply (ST.typing_ind_env (fun Σ Γ t T => + wf (trans_global Σ).1 -> typing (trans_global Σ) (trans_local Γ) (trans t) (trans T) )%type). all: simpl. @@ -1169,77 +1666,90 @@ Proof. - (* Sorts *) constructor; eauto. eapply trans_wf_local; eauto. - admit. (* todo *) - (* now rewrite global_ext_levels_trans. *) + destruct u; simpl in *; eauto. + intros. + now rewrite global_ext_levels_trans. - (* Casts *) - eapply refine_type. eapply type_App with _ (trans t). - eapply type_Lambda; eauto. eapply type_Rel. econstructor; auto. - eapply typing_wf_local. eauto. eauto. simpl. exists s; auto. reflexivity. eauto. + eapply refine_type. eapply type_App with _ (trans t) _. + 2:{ eapply type_Lambda; eauto. eapply type_Rel. econstructor; auto. + eapply typing_wf_local. eauto. eauto. simpl. exists s; auto. reflexivity. } + eapply type_Prod. eauto. + instantiate (1 := s). simpl. + eapply (weakening _ _ [_] _ (tSort _)); eauto. + constructor; eauto. eapply typing_wf_local; eauto. red. + exists s; eauto. auto. simpl. unfold subst1. rewrite simpl_subst; auto. now rewrite lift0_p. - (* The interesting application case *) eapply type_mkApps; eauto. + specialize (X0 X2). eapply typing_wf in X; eauto. destruct X. - clear H1 X0 H H0. revert H2. - induction X1. econstructor; eauto. - (* Need updated typing_spine in template-coq *) admit. reflexivity. - simpl in p. - destruct (TypingWf.typing_wf _ wfΣ _ _ _ typrod) as [wfAB _]. - intros wfT. - econstructor; eauto. exists s; eauto. - change (tProd na (trans A) (trans B)) with (trans (T.tProd na A B)). - apply trans_cumul; auto with trans. - apply TypingWf.typing_wf_sigma; auto. - eapply Forall_impl. eapply TypingWf.typing_all_wf_decl; eauto. - intros. auto. - eapply typing_wf in ty; eauto. destruct ty as [wfhd _]. - rewrite trans_subst in IHX1; eauto with wf. now inv wfAB. - eapply IHX1. apply Template.LiftSubst.wf_subst; try constructor; auto. now inv wfAB. + eapply PCUICValidity.validity in X0. + clear H1 H H0. revert H2. + induction X1. + * econstructor; eauto. reflexivity. + * simpl in p. + destruct (TypingWf.typing_wf _ wfΣ _ _ _ typrod) as [wfAB _]. + intros wfS. + econstructor; eauto. + + exists s; eauto. eapply p; eauto. + + change (tProd na (trans A) (trans B)) with (trans (S.tProd na A B)). + apply trans_cumul; auto with trans. + eapply Forall_impl. eapply TypingWf.typing_all_wf_decl; eauto. + intros. auto. + + eapply typing_wf in ty; eauto. destruct ty as [wfhd _]. + rewrite trans_subst in IHX1; eauto with wf. + eapply IHX1; cycle 1. + apply Template.LiftSubst.wf_subst; try constructor; auto. now inv wfAB. + specialize (p X2). specialize (p0 X2). + eapply PCUICInversion.inversion_Prod in p as [s1 [s2 [HA [HB Hs]]]]. + eapply PCUICArities.isType_subst with [vass na (trans A)]; eauto. + constructor; eauto with pcuic. + constructor; eauto with pcuic. now rewrite subst_empty. + now exists s2. + apply X2. + * apply X2. - rewrite trans_subst_instance_constr. pose proof (forall_decls_declared_constant _ _ _ H). replace (trans (Template.Ast.cst_type decl)) with (cst_type (trans_constant_body decl)) by (destruct decl; reflexivity). - constructor; eauto with trans. admit. + constructor; eauto with trans. + now apply (trans_consistent_instance_ext _ (S.ConstantDecl decl)). - rewrite trans_subst_instance_constr. pose proof (forall_decls_declared_inductive _ _ _ _ isdecl). replace (trans (Template.Ast.ind_type idecl)) with (ind_type (trans_one_ind_body idecl)) by (destruct idecl; reflexivity). - eapply type_Ind; eauto. eauto with trans. admit. + eapply type_Ind; eauto. eauto with trans. + now apply (trans_consistent_instance_ext _ (S.InductiveDecl mdecl)). - pose proof (forall_decls_declared_constructor _ _ _ _ _ isdecl). - unfold TTy.type_of_constructor in *. + unfold ST.type_of_constructor in *. rewrite trans_subst. - -- apply Forall_All. apply wf_inds. - -- apply wf_subst_instance_constr. - eapply declared_constructor_wf in isdecl; eauto with wf. - eauto with wf trans. - apply typing_wf_wf; auto. - -- rewrite trans_subst_instance_constr. - rewrite trans_inds. simpl. - eapply refine_type. econstructor; eauto with trans. admit. - unfold type_of_constructor. simpl. f_equal. f_equal. - destruct cdecl as [[id t] p]; simpl; auto. + rewrite trans_subst_instance_constr. + rewrite trans_inds. simpl. + eapply refine_type. econstructor; eauto with trans. + now apply (trans_consistent_instance_ext _ (S.InductiveDecl mdecl)). + unfold type_of_constructor. simpl. f_equal. f_equal. + destruct cdecl as [[id t] p]; simpl; auto. - rewrite trans_mkApps; auto with wf trans. apply typing_wf in X1; intuition auto. solve_all. apply typing_wf in X3; auto. intuition auto. - apply wf_mkApps_inv in H4. - apply All_app_inv. apply All_skipn. solve_all. constructor; auto. rewrite map_app map_skipn. eapply type_Case. apply forall_decls_declared_inductive; eauto. destruct mdecl; auto. -- simpl. rewrite firstn_map. rewrite trans_build_case_predicate_type. erewrite H0. reflexivity. - -- eapply X2. + -- eapply X1. -- rewrite global_ext_constraints_trans; exact H1. - -- rewrite trans_mkApps in X4; auto with wf. - eapply typing_wf in X3; auto. intuition. eapply wf_mkApps_inv in H4; auto. - -- admit. - -- apply trans_build_branches_type in H2. - rewrite firstn_map. exact H2. + -- rewrite trans_mkApps in X2; auto with wf. + -- apply H2. + -- eapply trans_build_branches_type in H3; eauto. + rewrite firstn_map. exact H3. -- eapply All2_map. solve_all. + destruct b0 as [s [Hs IHs]]; eauto. - destruct pdecl as [arity ty]; simpl in *. pose proof (TypingWf.declared_projection_wf _ _ _ _ _ isdecl). @@ -1247,18 +1757,16 @@ Proof. eapply forall_decls_declared_projection in isdecl. destruct (typing_wf _ wfΣ _ _ _ X1) as [wfc wfind]. eapply wf_mkApps_inv in wfind; auto. - rewrite trans_subst; auto with wf. constructor; solve_all. - apply wf_subst_instance_constr, H0. - eapply typing_wf_wf; auto. + rewrite trans_subst; auto with wf. simpl. rewrite map_rev. rewrite trans_subst_instance_constr. eapply (type_Proj _ _ _ _ _ _ _ (arity, trans ty)). eauto. - rewrite trans_mkApps in X2; auto. constructor. rewrite map_length. + rewrite trans_mkApps in X2; auto. rewrite map_length. destruct mdecl; auto. - eapply refine_type. assert (fix_context (map (map_def trans trans) mfix) = - trans_local (TTy.fix_context mfix)). - { unfold trans_local, TTy.fix_context. + trans_local (ST.fix_context mfix)). + { unfold trans_local, ST.fix_context. rewrite map_rev map_mapi /fix_context mapi_map. f_equal. apply mapi_ext; intros i x. simpl. rewrite trans_lift. reflexivity. } @@ -1271,55 +1779,235 @@ Proof. now exists s. -- apply All_map. eapply All_impl; eauto. intuition eauto 3 with wf; cbn. - rewrite H1. rewrite /trans_local map_length. - rewrite /trans_local map_app in X2. - rewrite <- trans_lift. apply X2. - rdest. destruct (dbody x); simpl in *; congruence. - -- admit. + rewrite H2. rewrite /trans_local map_length. + rewrite /trans_local map_app in X3. + rewrite <- trans_lift. apply X3; auto. + -- eapply trans_wf_fixpoint => //. + solve_all. destruct a as [s [Hs IH]]. + now eapply TypingWf.typing_wf in Hs. + now eapply TypingWf.typing_wf in a0. -- destruct decl; reflexivity. - eapply refine_type. assert (fix_context (map (map_def trans trans) mfix) = - trans_local (TTy.fix_context mfix)). - { unfold trans_local, TTy.fix_context. + trans_local (ST.fix_context mfix)). + { unfold trans_local, ST.fix_context. rewrite map_rev map_mapi /fix_context mapi_map. f_equal. apply mapi_ext => i x. simpl. rewrite trans_lift. reflexivity. } econstructor; eauto. - -- admit. - -- now rewrite nth_error_map H. + -- now eapply cofix_guard_trans. + -- now rewrite nth_error_map H0. -- eapply trans_wf_local; eauto. -- eapply All_map, (All_impl X0). intros x [s [Hs Hts]]. now exists s. -- apply All_map. eapply All_impl; eauto. intuition eauto 3 with wf. - rewrite H0. rewrite /trans_local map_length. - unfold Template.Ast.app_context in X2. - rewrite /trans_local map_app in X2. - cbn. rewrite <- trans_lift. apply X2. - -- admit. + rewrite H2. rewrite /trans_local map_length. + unfold Template.Ast.app_context in X3. + rewrite /trans_local map_app in X3. + cbn. rewrite <- trans_lift. now apply X3. + -- eapply trans_wf_cofixpoint => //. + solve_all. destruct a as [s [Hs IH]]. + now eapply TypingWf.typing_wf in Hs. + now eapply TypingWf.typing_wf in a0. -- destruct decl; reflexivity. - - assert (T.wf B). - { destruct X2. destruct i as [wfa allwfa]. - now apply (isWFArity_wf _ _ _ wfΣ wfa). - destruct s as [s [H ?]]. - eapply typing_wf in H; intuition eauto. } - destruct X2. red in i. - destruct i as [wfa allwfa]. - destruct wfa as [ctx [s ?]]. - * admit. - (* exists (trans_local ctx), s. destruct p. - generalize (trans_destArity [] B); rewrite e. - intros. split; auto. - eapply trans_wf_local in allwfa. simpl in allwfa. - rewrite /trans_local in allwfa. - now rewrite map_app in allwfa. - right. *) - * destruct s as [s [? ?]]; eauto. - eapply type_Cumul; eauto. - eapply trans_cumul; eauto with trans. - eapply typing_wf_sigma; auto. - clear X. apply typing_all_wf_decl in wfΓ; auto. - eapply typing_wf in X0; eauto. destruct X0. auto. + - assert (S.wf B). + { now apply typing_wf in X2. } + eapply type_Cumul; eauto. + eapply trans_cumul; eauto with trans. + clear X. apply typing_all_wf_decl in wfΓ; auto. + eapply typing_wf in X0; eauto. destruct X0. auto. +Qed. + +Lemma Alli_map {A B} (P : nat -> B -> Type) n (f : A -> B) l : + Alli (fun n x => P n (f x)) n l -> + Alli P n (map f l). +Proof. + induction 1; constructor; auto. +Qed. + +Lemma trans_arities_context m : trans_local (Ast.arities_context (Ast.ind_bodies m)) = + arities_context (map trans_one_ind_body (Ast.ind_bodies m)). +Proof. + rewrite /trans_local /Ast.arities_context rev_map_spec map_rev map_map_compose + /PCUICEnvironment.arities_context rev_map_spec map_map_compose /= //. +Qed. + +Lemma trans_lift_context n k Γ : lift_context n k (trans_local Γ) = + trans_local (LiftSubst.lift_context n k Γ). +Proof. + rewrite /lift_context /LiftSubst.lift_context /fold_context /Ast.fold_context. + rewrite /trans_local map_rev map_mapi -List.map_rev mapi_map. + f_equal. apply mapi_ext. intros ? [na [b|] ty]; simpl; rewrite /trans_decl /= /map_decl; simpl; + f_equal; now rewrite trans_lift. +Qed. + +Lemma trans_subst_telescope s k Γ : + map trans_decl (LiftSubst.subst_telescope s k Γ) = + subst_telescope (map trans s) k (map trans_decl Γ). +Proof. + rewrite /subst_telescope /LiftSubst.subst_telescope. + rewrite map_mapi mapi_map. + apply mapi_ext => n [na [b|] ty] /=; + rewrite /map_decl /=; f_equal; + now rewrite trans_subst. +Qed. + +Lemma trans_on_global_env `{checker_flags} Σ P Q : + (forall Σ Γ t T, ST.on_global_env P Σ.1 -> P Σ Γ t T -> + on_global_env Q (trans_global_decls Σ.1) -> + Q (trans_global Σ) (trans_local Γ) (trans t) (option_map trans T)) -> + ST.on_global_env P Σ -> on_global_env Q (trans_global_decls Σ). +Proof. + intros X X0. + simpl in *. induction X0; simpl; constructor; auto. + - red in f |- *. apply Forall_map. + solve_all. + - simpl. subst udecl. + clear -o. destruct d; simpl in *; + destruct o as [? [? [? ?]]]. + * split; rewrite global_levels_trans //. + repeat split => //. + red in H3 |- *. + rewrite -global_ext_constraints_trans in H3. + apply H3. + * split; rewrite global_levels_trans //. + repeat split => //. + red in H3 |- *. + rewrite -global_ext_constraints_trans in H3. + apply H3. + - simpl. + destruct d; simpl in *. + * destruct c; simpl. destruct cst_body; simpl in *. + red in o |- *. simpl in *. + apply (X (Σ, cst_universes) [] t (Some cst_type)); auto. + red in o0 |- *. simpl in *. + now apply (X (Σ, cst_universes) [] cst_type None). + * destruct o0 as [onI onP onNP]. + constructor; auto. + -- eapply Alli_map. eapply Alli_impl. exact onI. eauto. intros. + unshelve refine {| ind_indices := trans_local X1.(ST.ind_indices); + ind_sort := X1.(ST.ind_sort); + ind_cshapes := map trans_constructor_shape X1.(ST.ind_cshapes) |}. + --- simpl; rewrite X1.(ST.ind_arity_eq). + now rewrite !trans_it_mkProd_or_LetIn. + --- apply ST.onArity in X1. unfold on_type in *; simpl in *. + now apply (X (Σ, Ast.ind_universes m) [] (Ast.ind_type x) None). + --- pose proof X1.(ST.onConstructors) as X11. red in X11. + red. eapply All2_map. + eapply All2_impl; eauto. + simpl. intros [[? ?] ?] cs onc. destruct onc; unshelve econstructor; eauto. + + simpl. unfold trans_local. rewrite context_assumptions_map. + now rewrite cstr_args_length. + + simpl; unfold cdecl_type, ST.cdecl_type in cstr_eq |- *; simpl in *. + rewrite cstr_eq. rewrite !trans_it_mkProd_or_LetIn. + autorewrite with len. f_equal. f_equal. + rewrite !trans_mkApps //. + f_equal; auto. simpl. + now rewrite /trans_local !map_length. + rewrite map_app /=. f_equal. + rewrite /trans_local !map_length. + unfold TemplateEnvironment.to_extended_list_k. + now rewrite trans_reln. + + unfold cdecl_type, ST.cdecl_type in on_ctype |- *; simpl in *. + red. + move: (X (Σ, Ast.ind_universes m) (Ast.arities_context (Ast.ind_bodies m)) t None). + now rewrite trans_arities_context. + + clear -X0 IHX0 X on_cargs. revert on_cargs. simpl. + generalize (ST.cshape_args cs), (ST.cshape_sorts cs). + have foo := (X (Σ, udecl) _ _ _ X0). + rewrite -trans_arities_context. + induction c; destruct l; simpl; auto; + destruct a as [na [b|] ty]; simpl in *; auto; + split; intuition eauto; + specialize (foo + (Ast.app_context (Ast.app_context (Ast.arities_context (Ast.ind_bodies m)) (Ast.ind_params m)) c)); + rewrite /trans_local !map_app in foo. + now apply (foo ty None). + now apply (foo b (Some ty)). + now apply (foo ty None). + now apply (foo b (Some ty)). + now apply (foo ty (Some (Ast.tSort _))). + + clear -X0 IHX0 X on_cindices. + revert on_cindices. + rewrite trans_lift_context /trans_local -map_rev. + simpl. rewrite {2}/trans_local map_length. + generalize (List.rev (LiftSubst.lift_context #|ST.cshape_args cs| 0 (ST.ind_indices X1))). + generalize (ST.cshape_indices cs). + rewrite -trans_arities_context. + induction 1; simpl; constructor; auto; + have foo := (X (Σ, Ast.ind_universes m) _ _ _ X0); + specialize (foo (Ast.app_context (Ast.app_context (Ast.arities_context (Ast.ind_bodies m)) + (Ast.ind_params m)) (ST.cshape_args cs))). + rewrite /trans_local !map_app in foo. + now apply (foo i (Some t)). + now rewrite (trans_subst_telescope [i] 0) in IHon_cindices. + now rewrite (trans_subst_telescope [b] 0) in IHon_cindices. + + clear -IHX0 on_ctype_positive. + unfold ST.cdecl_type in *. unfold cdecl_type. simpl in *. + change [] with (map trans_decl []). revert on_ctype_positive. + generalize (@nil Ast.context_decl). induction 1; simpl. + rewrite trans_mkApps. simpl. + subst headrel. + assert (#|PCUICEnvironment.ind_bodies (trans_minductive_body m)| = #|TemplateEnvironment.ind_bodies m|) as <-. + now rewrite /trans_minductive_body /= map_length. + assert (#|ctx| = #|map trans_decl ctx|) as ->. now rewrite map_length. + eapply positive_cstr_concl. + rewrite map_length. + eapply All_map. eapply All_impl; eauto. + simpl; intros. admit. + constructor. now rewrite -(trans_subst [b]). + constructor. admit. + apply IHon_ctype_positive. + + clear -on_ctype_variance. + intros v indv. admit. + --- simpl; intros. have onp := X1.(ST.onProjections). + admit. + --- have inds := X1.(ST.ind_sorts). + admit. + --- have inds := X1.(ST.onIndices). + admit. + -- red in onP. red. + admit. + -- admit. Admitted. + +Lemma template_to_pcuic_env {cf} Σ : Template.Typing.wf Σ -> wf (trans_global_decls Σ). +Proof. + intros Hu. + eapply trans_on_global_env; eauto. simpl; intros. + epose proof (ST.env_prop_typing _ template_to_pcuic _ X Γ). + forward X2. + red in X0. destruct T. + now eapply ST.typing_wf_local. + destruct X0 as [s Hs]. eapply ST.typing_wf_local; eauto. + destruct T; simpl in *. + - eapply X2; eauto. + - destruct X0 as [s Hs]. exists s. + eapply (X2 _ (Ast.tSort s)); eauto. +Qed. + +Lemma template_to_pcuic_env_ext {cf} Σ : Template.Typing.wf_ext Σ -> wf_ext (trans_global Σ). +Proof. + intros [u Hu]. + split. now apply template_to_pcuic_env. + destruct Hu as [? [? [? ?]]]. + split; rewrite global_levels_trans //. + repeat split => //. + red in H2 |- *. + rewrite -global_ext_constraints_trans in H2. + apply H2. +Qed. + +Theorem template_to_pcuic_typing {cf} Σ Γ t T : + ST.wf Σ.1 -> + ST.typing Σ Γ t T -> + typing (trans_global Σ) (trans_local Γ) (trans t) (trans T). +Proof. + intros wf ty. + apply (ST.env_prop_typing _ template_to_pcuic _ wf); auto. + now eapply ST.typing_wf_local. + now apply template_to_pcuic_env. +Qed. diff --git a/safechecker/_CoqProject.in b/safechecker/_CoqProject.in index b9041b4c1..1b751f2da 100644 --- a/safechecker/_CoqProject.in +++ b/safechecker/_CoqProject.in @@ -1,10 +1,15 @@ -R theories MetaCoq.SafeChecker +theories/PCUICErrors.v theories/PCUICSafeReduce.v +theories/PCUICEqualityDec.v theories/PCUICSafeConversion.v theories/PCUICWfReduction.v +theories/PCUICWfEnv.v +theories/PCUICTypeChecker.v theories/PCUICSafeChecker.v theories/SafeTemplateChecker.v theories/PCUICSafeRetyping.v +theories/PCUICConsistency.v theories/Extraction.v diff --git a/safechecker/_PluginProject.in b/safechecker/_PluginProject.in index 1e90c5714..870755842 100644 --- a/safechecker/_PluginProject.in +++ b/safechecker/_PluginProject.in @@ -27,6 +27,8 @@ src/typing0.ml # From PCUIC +src/pCUICPrimitive.mli +src/pCUICPrimitive.ml src/pCUICAst.ml src/pCUICAst.mli src/pCUICAstUtils.ml @@ -83,10 +85,18 @@ src/pCUICPretty.mli src/pCUICPretty.ml # From SafeChecker +src/pCUICErrors.mli +src/pCUICErrors.ml src/pCUICSafeReduce.mli src/pCUICSafeReduce.ml +src/pCUICEqualityDec.mli +src/pCUICEqualityDec.ml src/pCUICSafeConversion.mli src/pCUICSafeConversion.ml +# src/pCUICWfEnv.mli +# src/pCUICWfEnv.ml +src/pCUICTypeChecker.mli +src/pCUICTypeChecker.ml src/pCUICSafeChecker.mli src/pCUICSafeChecker.ml src/safeTemplateChecker.mli diff --git a/safechecker/src/g_metacoq_safechecker.mlg b/safechecker/src/g_metacoq_safechecker.mlg index 6414c894a..871d8ac46 100644 --- a/safechecker/src/g_metacoq_safechecker.mlg +++ b/safechecker/src/g_metacoq_safechecker.mlg @@ -9,7 +9,18 @@ open PeanoNat.Nat open Datatypes open PCUICSafeChecker -let pr_char_list l = str (Scanf.unescaped (Tm_util.list_to_string l)) +let bytes_of_list l = + let bytes = Bytes.create (List.length l) in + let rec fill acc = function + | [] -> bytes + | c :: cs -> + Bytes.set bytes acc c; + fill (1 + acc) cs + in fill 0 l + +let pr_char_list l = + (* We allow utf8 encoding *) + str (Bytes.to_string (bytes_of_list l)) let time prefix f x = let start = Unix.gettimeofday () in diff --git a/safechecker/src/metacoq_safechecker_plugin.mlpack b/safechecker/src/metacoq_safechecker_plugin.mlpack index e4e3f623a..0868323b5 100644 --- a/safechecker/src/metacoq_safechecker_plugin.mlpack +++ b/safechecker/src/metacoq_safechecker_plugin.mlpack @@ -14,6 +14,7 @@ Classes0 Logic1 Relation Relation_Properties +PCUICPrimitive PCUICAst PCUICAstUtils PCUICInduction @@ -30,10 +31,13 @@ PCUICPosition PCUICPretty TemplateToPCUIC PCUICCumulativity - PCUICUnivSubst + +PCUICErrors PCUICSafeReduce +PCUICEqualityDec PCUICSafeConversion +PCUICTypeChecker PCUICSafeChecker SafeTemplateChecker diff --git a/safechecker/theories/Extraction.v b/safechecker/theories/Extraction.v index 8512dfc5b..9859f88f0 100644 --- a/safechecker/theories/Extraction.v +++ b/safechecker/theories/Extraction.v @@ -1,6 +1,6 @@ (* Distributed under the terms of the MIT license. *) -Require Import OrdersTac ExtrOcamlBasic ExtrOcamlString ExtrOcamlZInt. -Require Import MetaCoq.Template.utils. +From Coq Require Import OrdersTac ExtrOcamlBasic ExtrOcamlString ExtrOcamlZInt ExtrOCamlFloats. +From MetaCoq.Template Require Import utils MC_ExtrOCamlInt63 (*b/c nameclash with `comparion` *). From MetaCoq.SafeChecker Require Import PCUICSafeChecker PCUICSafeConversion SafeTemplateChecker. @@ -10,6 +10,7 @@ From MetaCoq.SafeChecker Require Import PCUICSafeChecker PCUICSafeConversion should use these same directives for consistency. *) +(** Here we could extract uint63_from/to_model to the identity *) (* Ignore [Decimal.int] before the extraction issue is solved: https://github.com/coq/coq/issues/7017. *) @@ -32,10 +33,12 @@ Extract Constant Equations.Init.pr1 => "fst". Extract Constant Equations.Init.pr2 => "snd". Extraction Inline Equations.Init.pr1 Equations.Init.pr2. Extraction Inline Equations.Prop.Logic.transport Equations.Prop.Logic.transport_r MCEquality.transport. +Extraction Inline Equations.Prop.Logic.True_rect_dep Equations.Prop.Logic.False_rect_dep. -Extract Constant PCUICTyping.fix_guard => "(fun x -> true)". -Extract Constant PCUICTyping.cofix_guard => "(fun x -> true)". -Extract Constant PCUICTyping.ind_guard => "(fun x -> true)". +(** This Inline is because of a problem of weak type variables (partial application?) *) +Extraction Inline PCUICPrimitive.prim_val_reflect_eq. + +Extract Constant PCUICTyping.guard_checking => "{ fix_guard = (fun _ _ _ -> true); cofix_guard = (fun _ _ _ -> true) }". Cd "src". diff --git a/safechecker/theories/PCUICConsistency.v b/safechecker/theories/PCUICConsistency.v new file mode 100644 index 000000000..297ca7f33 --- /dev/null +++ b/safechecker/theories/PCUICConsistency.v @@ -0,0 +1,220 @@ +(* The safe checker gives us a sound and complete wh-normalizer for + PCUIC, assuming strong normalization. Combined with canonicity, + this allows us to prove that PCUIC is consistent, i.e. there + is no axiom-free proof of [forall (P : Prop), P] in the empty + context. To do so we use weakening to add an empty inductive, the + provided term to build an inhabitant and then canonicity to show + that this is a contradiction. *) + +From Coq Require Import Ascii String. +From Equations Require Import Equations. +From MetaCoq.PCUIC Require Import PCUICAst. +From MetaCoq.PCUIC Require Import PCUICAstUtils. +From MetaCoq.PCUIC Require Import PCUICCanonicity. +From MetaCoq.PCUIC Require Import PCUICInductiveInversion. +From MetaCoq.PCUIC Require Import PCUICInversion. +From MetaCoq.PCUIC Require Import PCUICLiftSubst. +From MetaCoq.PCUIC Require Import PCUICSR. +From MetaCoq.PCUIC Require Import PCUICSafeLemmata. +From MetaCoq.PCUIC Require Import PCUICTyping. +From MetaCoq.PCUIC Require Import PCUICUnivSubst. +From MetaCoq.PCUIC Require Import PCUICValidity. +From MetaCoq.PCUIC Require Import PCUICWeakeningEnv. +From MetaCoq.Template Require Import config utils. +From MetaCoq.SafeChecker Require Import PCUICSafeReduce. + +Local Opaque hnf. + +Fixpoint string_repeat c (n : nat) : string := + match n with + | 0 => "" + | S n => String c (string_repeat c n) + end. + +Lemma string_repeat_length c n : + String.length (string_repeat c n) = n. +Proof. + induction n; cbn; auto. +Qed. + +Definition max_name_length (Σ : global_env) : nat := + fold_right max 0 (map (fun '(kn, _) => String.length (string_of_kername kn)) Σ). + +Lemma max_name_length_ge Σ : + Forall (fun '(kn, _) => String.length (string_of_kername kn) <= max_name_length Σ) Σ. +Proof. + induction Σ as [|(kn&decl) Σ IH]; cbn; constructor. + - lia. + - eapply Forall_impl; eauto. + intros (?&?); cbn; intros. + fold (max_name_length Σ). + lia. +Qed. + +Definition make_fresh_name (Σ : global_env) : kername := + (MPfile [], string_repeat "a"%char (S (max_name_length Σ))). + +Lemma make_fresh_name_fresh Σ : + fresh_global (make_fresh_name Σ) Σ. +Proof. + pose proof (max_name_length_ge Σ) as all. + eapply Forall_impl; eauto. + cbn. + intros (kn&decl) le. + cbn. + intros ->. + unfold make_fresh_name in le. + cbn in le. + rewrite string_repeat_length in le. + lia. +Qed. + +Definition Prop_univ := Universe.of_levels (inl PropLevel.lProp). + +Definition False_oib : one_inductive_body := + {| ind_name := ""; + ind_type := tSort Prop_univ; + ind_kelim := IntoAny; + ind_ctors := []; + ind_projs := []; + ind_relevance := Relevant |}. + +Definition False_mib : mutual_inductive_body := + {| ind_finite := BiFinite; + ind_npars := 0; + ind_params := []; + ind_bodies := [False_oib]; + ind_universes := Monomorphic_ctx ContextSet.empty; + ind_variance := None |}. + +Definition axiom_free Σ := + forall c decl, declared_constant Σ c decl -> cst_body decl <> None. + +Lemma axiom_free_axiom_free_value Σ t : + axiom_free Σ -> + axiom_free_value Σ [] t. +Proof. + intros axfree. + cut (Forall is_true []); [|constructor]. + generalize ([] : list bool). + induction t; intros axfree_args all_true; cbn; auto. + - destruct lookup_env eqn:find; auto. + destruct g; auto. + destruct c; auto. + apply axfree in find; cbn in *. + now destruct cst_body. + - destruct nth_error; auto. + rewrite nth_nth_error. + destruct nth_error eqn:nth; auto. + eapply nth_error_forall in nth; eauto. +Qed. + +Definition binder := {| binder_name := nNamed "P"; binder_relevance := Relevant |}. + +Theorem pcuic_consistent {cf:checker_flags} Σ t : + wf_ext Σ -> + axiom_free Σ -> + (* t : forall (P : Prop), P *) + Σ ;;; [] |- t : tProd binder (tSort Prop_univ) (tRel 0) -> + False. +Proof. + intros wfΣ axfree cons. + set (Σext := ((make_fresh_name Σ, InductiveDecl False_mib) :: Σ.1, Σ.2)). + assert (wf': wf_ext Σext). + { constructor; [constructor|]. + - destruct wfΣ; auto. + - apply make_fresh_name_fresh. + - red. + cbn. + split. + { now intros ? ?%LevelSet.empty_spec. } + split. + { now intros ? ?%ConstraintSet.empty_spec. } + split; auto. + destruct wfΣ as (?&(?&?&?&[val sat])). + exists val. + intros l isin. + apply sat; auto. + apply ConstraintSet.union_spec. + apply ConstraintSet.union_spec in isin as [?%ConstraintSet.empty_spec|]; auto. + - hnf. + constructor. + + constructor. + * econstructor; cbn. + -- instantiate (2 := []); reflexivity. + -- exists (Universe.super Prop_univ). + constructor; auto. + constructor. + -- instantiate (1 := []). + constructor. + -- intros. + congruence. + -- constructor. + -- intros ? [=]. + * constructor. + + constructor. + + reflexivity. + + reflexivity. + - destruct wfΣ; auto. } + eapply (env_prop_typing _ _ PCUICWeakeningEnv.weakening_env) in cons; auto. + 3: exists [(make_fresh_name Σ, InductiveDecl False_mib)]; reflexivity. + 2: now destruct wf'. + + set (Σ' := _ ++ _) in cons. + set (False_ty := tInd (mkInd (make_fresh_name Σ) 0) []). + assert (typ_false: (Σ', Σ.2);;; [] |- tApp t False_ty : False_ty). + { apply validity_term in cons as typ_prod; auto. + destruct typ_prod. + eapply type_App with (B := tRel 0) (u := False_ty); eauto. + eapply type_Ind with (u := []) (mdecl := False_mib) (idecl := False_oib); eauto. + - hnf. + cbn. + unfold declared_minductive. + cbn. + rewrite eq_kername_refl. + auto. + - cbn. + auto. } + assert (sqwf: ∥ wf (Σ', Σ.2).1 ∥) by now destruct wf'. + pose proof (iswelltyped _ _ _ _ typ_false) as wt. + pose proof (@hnf_sound _ _ sqwf _ _ wt) as [r]. + pose proof (@hnf_complete _ _ sqwf _ _ wt) as [w]. + eapply subject_reduction in typ_false; eauto. + eapply whnf_ind_finite with (indargs := []) in typ_false as ctor; auto. + - unfold isConstruct_app in ctor. + destruct decompose_app eqn:decomp. + apply decompose_app_inv in decomp. + rewrite decomp in typ_false. + destruct t0; try discriminate ctor. + apply inversion_mkApps in typ_false as H; auto. + destruct H as (?&typ_ctor&_). + apply inversion_Construct in typ_ctor as (?&?&?&?&?&?&?); auto. + unshelve eapply Construct_Ind_ind_eq with (args' := []) in typ_false. + 5: eassumption. + 1: eauto. + destruct on_declared_constructor. + destruct p. + destruct s. + destruct p. + destruct typ_false as ((((->&_)&_)&_)&_). + clear -d. + destruct d as ((?&?)&?). + cbn in *. + red in H. + cbn in *. + rewrite eq_kername_refl in H. + noconf H. + noconf H0. + cbn in H1. + rewrite nth_error_nil in H1. + discriminate. + - eapply axiom_free_axiom_free_value. + intros kn decl isdecl. + hnf in isdecl. + cbn in isdecl. + destruct eq_kername; [noconf isdecl|]. + eapply axfree; eauto. + - unfold check_recursivity_kind. + cbn. + rewrite eq_kername_refl; auto. +Qed. diff --git a/safechecker/theories/PCUICEqualityDec.v b/safechecker/theories/PCUICEqualityDec.v new file mode 100644 index 000000000..0c12c3d60 --- /dev/null +++ b/safechecker/theories/PCUICEqualityDec.v @@ -0,0 +1,175 @@ +(* Distributed under the terms of the MIT license. *) +From Coq Require Import ProofIrrelevance ssreflect ssrbool. +From MetaCoq.Template Require Import config utils uGraph. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils + PCUICReflect PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICGlobalEnv + PCUICCumulativity PCUICEquality. + +Require Import Equations.Prop.DepElim. +From Equations Require Import Equations. + +Local Set Keyed Unification. + +Set Default Goal Selector "!". + +(** Checking equality *) + +Section EqualityDec. + Context {cf : checker_flags}. + Context (Σ : global_env_ext). + Context (hΣ : ∥ wf Σ ∥) (Hφ : ∥ on_udecl Σ.1 Σ.2 ∥). + Context (G : universes_graph) (HG : is_graph_of_uctx G (global_ext_uctx Σ)). + + Local Definition hΣ' : ∥ wf_ext Σ ∥. + Proof. + destruct hΣ, Hφ; now constructor. + Defined. + + Definition conv_pb_relb pb := + match pb with + | Conv => check_eqb_universe G + | Cumul => check_leqb_universe G + end. + + Definition eqb_termp_napp pb napp := + eqb_term_upto_univ_napp Σ (check_eqb_universe G) (conv_pb_relb pb) napp. + + Lemma eqb_termp_napp_spec pb napp t u : + eqb_termp_napp pb napp t u -> + eq_termp_napp pb Σ napp t u. + Proof. + pose proof hΣ'. + apply eqb_term_upto_univ_impl. + - intros u1 u2. + eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). + + sq. now eapply wf_ext_global_uctx_invariants. + + sq; now eapply global_ext_uctx_consistent. + + assumption. + - intros u1 u2. + destruct pb. + + eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). + * sq; now eapply wf_ext_global_uctx_invariants. + * sq; now eapply global_ext_uctx_consistent. + * assumption. + + eapply (check_leqb_universe_spec' G (global_ext_uctx Σ)). + * sq; now eapply wf_ext_global_uctx_invariants. + * sq; now eapply global_ext_uctx_consistent. + * assumption. + Qed. + + Definition eqb_termp pb := (eqb_termp_napp pb 0). + Definition eqb_term := (eqb_termp Conv). + Definition leqb_term := (eqb_termp Cumul). + + Lemma leqb_term_spec t u : + leqb_term t u -> + leq_term Σ (global_ext_constraints Σ) t u. + Proof. + intros. + now apply eqb_termp_napp_spec in H. + Qed. + + Notation eq_term Σ t u := (eq_term Σ Σ t u). + + Lemma eqb_term_spec t u : + eqb_term t u -> + eq_term Σ t u. + Proof. + intros. + now apply eqb_termp_napp_spec in H. + Qed. + + Lemma eqb_term_refl : + forall t, eqb_term t t. + Proof. + intro t. eapply eqb_term_upto_univ_refl. + all: apply check_eqb_universe_refl. + Qed. + + Definition eqb_binder_annot {A} (b b' : binder_annot A) : bool := + eqb b.(binder_relevance) b'.(binder_relevance). + + Lemma eq_binder_annot_reflect {A} na na' : reflect (eq_binder_annot (A:=A) na na') (eqb_binder_annot na na'). + Proof. + unfold eq_binder_annot, eqb_binder_annot. + destruct (eqb_spec na.(binder_relevance) na'.(binder_relevance)); constructor; auto. + Qed. + + Fixpoint eqb_ctx (Γ Δ : context) : bool := + match Γ, Δ with + | [], [] => true + | {| decl_name := na1 ; decl_body := None ; decl_type := t1 |} :: Γ, + {| decl_name := na2 ; decl_body := None ; decl_type := t2 |} :: Δ => + eqb_binder_annot na1 na2 && eqb_term t1 t2 && eqb_ctx Γ Δ + | {| decl_name := na1 ; decl_body := Some b1 ; decl_type := t1 |} :: Γ, + {| decl_name := na2 ; decl_body := Some b2 ; decl_type := t2 |} :: Δ => + eqb_binder_annot na1 na2 && eqb_term b1 b2 && eqb_term t1 t2 && eqb_ctx Γ Δ + | _, _ => false + end. + + Lemma eqb_binder_annot_spec {A} na na' : eqb_binder_annot (A:=A) na na' -> eq_binder_annot (A:=A) na na'. + Proof. apply (elimT (eq_binder_annot_reflect _ _)). Qed. + + Lemma eqb_ctx_spec : + forall Γ Δ, + eqb_ctx Γ Δ -> + eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) Γ Δ. + Proof. + intros Γ Δ h. + induction Γ as [| [na [b|] A] Γ ih ] in Δ, h |- *. + all: destruct Δ as [| [na' [b'|] A'] Δ]. + all: try discriminate. + - constructor. + - simpl in h. apply andb_andI in h as [[[h1 h2]%andb_and h3]%andb_and h4]. + constructor. + + now apply eqb_binder_annot_spec in h1. + + eapply eqb_term_spec. assumption. + + eapply eqb_term_spec. assumption. + + eapply ih. assumption. + - simpl in h. apply andb_and in h as [[h1 h2]%andb_and h3]. + constructor. + + now apply eqb_binder_annot_spec. + + eapply eqb_term_spec. assumption. + + eapply ih. assumption. + Qed. + + Definition eqb_opt_term (t u : option term) := + match t, u with + | Some t, Some u => eqb_term t u + | None, None => true + | _, _ => false + end. + + Lemma eqb_opt_term_spec t u + : eqb_opt_term t u -> eq_opt_term false Σ (global_ext_constraints Σ) t u. + Proof. + destruct t, u; try discriminate; cbn => //. + apply eqb_term_spec; tea. + Qed. + + Definition eqb_decl (d d' : context_decl) := + eqb_binder_annot d.(decl_name) d'.(decl_name) && + eqb_opt_term d.(decl_body) d'.(decl_body) && eqb_term d.(decl_type) d'.(decl_type). + + Lemma eqb_decl_spec d d' + : eqb_decl d d' -> eq_decl false Σ (global_ext_constraints Σ) d d'. + Proof. + unfold eqb_decl, eq_decl. + intro H. rtoProp. apply eqb_opt_term_spec in H1. + eapply eqb_term_spec in H0; tea. + eapply eqb_binder_annot_spec in H. + repeat split; eauto. + Qed. + + Definition eqb_context (Γ Δ : context) := forallb2 eqb_decl Γ Δ. + + Lemma eqb_context_spec Γ Δ + : eqb_context Γ Δ -> eq_context false Σ (global_ext_constraints Σ) Γ Δ. + Proof. + unfold eqb_context, eq_context. + intro HH. apply forallb2_All2 in HH. + eapply All2_impl; try eassumption. + cbn. apply eqb_decl_spec. + Qed. + +End EqualityDec. \ No newline at end of file diff --git a/safechecker/theories/PCUICErrors.v b/safechecker/theories/PCUICErrors.v new file mode 100644 index 000000000..6136f62fb --- /dev/null +++ b/safechecker/theories/PCUICErrors.v @@ -0,0 +1,329 @@ + +(* Distributed under the terms of the MIT license. *) +From MetaCoq.Template Require Import config utils uGraph. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICPretty. +From Equations Require Import Equations. +Require Import ssreflect ssrbool. + +Local Set Keyed Unification. +Set Equations Transparent. + +Inductive ConversionError := +| NotFoundConstants (c1 c2 : kername) + +| NotFoundConstant (c : kername) + +| LambdaNotConvertibleTypes + (Γ1 : context) (na : aname) (A1 t1 : term) + (Γ2 : context) (na' : aname) (A2 t2 : term) + (e : ConversionError) + +| LambdaNotConvertibleAnn + (Γ1 : context) (na : aname) (A1 t1 : term) + (Γ2 : context) (na' : aname) (A2 t2 : term) + +| ProdNotConvertibleDomains + (Γ1 : context) (na : aname) (A1 B1 : term) + (Γ2 : context) (na' : aname) (A2 B2 : term) + (e : ConversionError) + +| ProdNotConvertibleAnn + (Γ1 : context) (na : aname) (A1 B1 : term) + (Γ2 : context) (na' : aname) (A2 B2 : term) + +| CaseOnDifferentInd + (Γ1 : context) + (ind : inductive) (par : nat) (p c : term) (brs : list (nat × term)) + (Γ2 : context) + (ind' : inductive) (par' : nat) (p' c' : term) (brs' : list (nat × term)) + +| CaseBranchNumMismatch + (ind : inductive) (par : nat) + (Γ : context) (p c : term) (brs1 : list (nat × term)) + (m : nat) (br : term) (brs2 : list (nat × term)) + (Γ' : context) (p' c' : term) (brs1' : list (nat × term)) + (m' : nat) (br' : term) (brs2' : list (nat × term)) + +| DistinctStuckProj + (Γ : context) (p : projection) (c : term) + (Γ' : context) (p' : projection) (c' : term) + +| CannotUnfoldFix + (Γ : context) (mfix : mfixpoint term) (idx : nat) + (Γ' : context) (mfix' : mfixpoint term) (idx' : nat) + +| FixRargMismatch (idx : nat) + (Γ : context) (u : def term) (mfix1 mfix2 : mfixpoint term) + (Γ' : context) (v : def term) (mfix1' mfix2' : mfixpoint term) + +| FixMfixMismatch (idx : nat) + (Γ : context) (mfix : mfixpoint term) + (Γ' : context) (mfix' : mfixpoint term) + +| DistinctCoFix + (Γ : context) (mfix : mfixpoint term) (idx : nat) + (Γ' : context) (mfix' : mfixpoint term) (idx' : nat) + +| CoFixRargMismatch (idx : nat) + (Γ : context) (u : def term) (mfix1 mfix2 : mfixpoint term) + (Γ' : context) (v : def term) (mfix1' mfix2' : mfixpoint term) + +| CoFixMfixMismatch (idx : nat) + (Γ : context) (mfix : mfixpoint term) + (Γ' : context) (mfix' : mfixpoint term) + +| StackHeadError + (leq : conv_pb) + (Γ1 : context) + (t1 : term) (args1 : list term) (u1 : term) (l1 : list term) + (Γ2 : context) + (t2 : term) (u2 : term) (l2 : list term) + (e : ConversionError) + +| StackTailError (leq : conv_pb) + (Γ1 : context) + (t1 : term) (args1 : list term) (u1 : term) (l1 : list term) + (Γ2 : context) + (t2 : term) (u2 : term) (l2 : list term) + (e : ConversionError) + +| StackMismatch + (Γ1 : context) (t1 : term) (args1 l1 : list term) + (Γ2 : context) (t2 : term) (l2 : list term) + +| HeadMismatch + (leq : conv_pb) + (Γ1 : context) (t1 : term) + (Γ2 : context) (t2 : term). + +Inductive type_error := +| UnboundRel (n : nat) +| UnboundVar (id : string) +| UnboundEvar (ev : nat) +| UndeclaredConstant (c : kername) +| UndeclaredInductive (c : inductive) +| UndeclaredConstructor (c : inductive) (i : nat) +| NotCumulSmaller (G : universes_graph) (Γ : context) (t u t' u' : term) (e : ConversionError) +| NotConvertible (G : universes_graph) (Γ : context) (t u : term) +| NotASort (t : term) +| NotAProduct (t t' : term) +| NotAnInductive (t : term) +| NotAnArity (t : term) +| IllFormedFix (m : mfixpoint term) (i : nat) +| UnsatisfiedConstraints (c : ConstraintSet.t) +| Msg (s : string). +Derive NoConfusion for type_error. + +Definition print_level := string_of_level. + +Definition string_of_Z z := + if (z <=? 0)%Z then "-" ^ string_of_nat (Z.to_nat (- z)) else string_of_nat (Z.to_nat z). + +Definition print_edge '(l1, n, l2) + := "(" ^ print_level l1 ^ ", " ^ string_of_Z n ^ ", " + ^ print_level l2 ^ ")". + +Definition print_universes_graph (G : universes_graph) := + let levels := LevelSet.elements G.1.1 in + let edges := wGraph.EdgeSet.elements G.1.2 in + string_of_list print_level levels + ^ "\n" ^ string_of_list print_edge edges. + +Definition string_of_conv_pb (c : conv_pb) : string := + match c with + | Conv => "conversion" + | Cumul => "cumulativity" + end. + +Definition print_term Σ Γ t := + print_term Σ Γ true false t. + +Fixpoint string_of_conv_error Σ (e : ConversionError) : string := + match e with + | NotFoundConstants c1 c2 => + "Both constants " ^ string_of_kername c1 ^ " and " ^ string_of_kername c2 ^ + " are not found in the environment." + | NotFoundConstant c => + "Constant " ^ string_of_kername c ^ + " common in both terms is not found in the environment." + | LambdaNotConvertibleAnn Γ1 na A1 t1 Γ2 na' A2 t2 => + "When comparing\n" ^ print_term Σ Γ1 (tLambda na A1 t1) ^ + "\nand\n" ^ print_term Σ Γ2 (tLambda na' A2 t2) ^ + "\nbinding annotations are not convertible\n" + | LambdaNotConvertibleTypes Γ1 na A1 t1 Γ2 na' A2 t2 e => + "When comparing\n" ^ print_term Σ Γ1 (tLambda na A1 t1) ^ + "\nand\n" ^ print_term Σ Γ2 (tLambda na' A2 t2) ^ + "\ntypes are not convertible:\n" ^ + string_of_conv_error Σ e + | ProdNotConvertibleAnn Γ1 na A1 B1 Γ2 na' A2 B2 => + "When comparing\n" ^ print_term Σ Γ1 (tProd na A1 B1) ^ + "\nand\n" ^ print_term Σ Γ2 (tProd na' A2 B2) ^ + "\nbinding annotations are not convertible:\n" + | ProdNotConvertibleDomains Γ1 na A1 B1 Γ2 na' A2 B2 e => + "When comparing\n" ^ print_term Σ Γ1 (tProd na A1 B1) ^ + "\nand\n" ^ print_term Σ Γ2 (tProd na' A2 B2) ^ + "\ndomains are not convertible:\n" ^ + string_of_conv_error Σ e + | CaseOnDifferentInd Γ ind par p c brs Γ' ind' par' p' c' brs' => + "The two stuck pattern-matching\n" ^ + print_term Σ Γ (tCase (ind, par) p c brs) ^ + "\nand\n" ^ print_term Σ Γ' (tCase (ind', par') p' c' brs') ^ + "\nare done on distinct inductive types." + | CaseBranchNumMismatch + ind par Γ p c brs1 m br brs2 Γ' p' c' brs1' m' br' brs2' => + "The two stuck pattern-matching\n" ^ + print_term Σ Γ (tCase (ind, par) p c (brs1 ++ (m,br) :: brs2)) ^ + "\nand\n" ^ + print_term Σ Γ' (tCase (ind, par) p' c' (brs1' ++ (m',br') :: brs2')) ^ + "\nhave a mistmatch in the branch number " ^ string_of_nat #|brs1| ^ + "\nthe number of parameters do not coincide\n" ^ + print_term Σ Γ br ^ + "\nhas " ^ string_of_nat m ^ " parameters while\n" ^ + print_term Σ Γ br' ^ + "\nhas " ^ string_of_nat m' ^ "." + | DistinctStuckProj Γ p c Γ' p' c' => + "The two stuck projections\n" ^ + print_term Σ Γ (tProj p c) ^ + "\nand\n" ^ + print_term Σ Γ' (tProj p' c') ^ + "\nare syntactically different." + | CannotUnfoldFix Γ mfix idx Γ' mfix' idx' => + "The two fixed-points\n" ^ + print_term Σ Γ (tFix mfix idx) ^ + "\nand\n" ^ print_term Σ Γ' (tFix mfix' idx') ^ + "\ncorrespond to syntactically distinct terms that can't be unfolded." + | FixRargMismatch idx Γ u mfix1 mfix2 Γ' v mfix1' mfix2' => + "The two fixed-points\n" ^ + print_term Σ Γ (tFix (mfix1 ++ u :: mfix2) idx) ^ + "\nand\n" ^ print_term Σ Γ' (tFix (mfix1' ++ v :: mfix2') idx) ^ + "\nhave a mismatch in the function number " ^ string_of_nat #|mfix1| ^ + ": arguments " ^ string_of_nat u.(rarg) ^ + " and " ^ string_of_nat v.(rarg) ^ "are different." + | FixMfixMismatch idx Γ mfix Γ' mfix' => + "The two fixed-points\n" ^ + print_term Σ Γ (tFix mfix idx) ^ + "\nand\n" ^ + print_term Σ Γ' (tFix mfix' idx) ^ + "\nhave a different number of mutually defined functions." + | DistinctCoFix Γ mfix idx Γ' mfix' idx' => + "The two cofixed-points\n" ^ + print_term Σ Γ (tCoFix mfix idx) ^ + "\nand\n" ^ print_term Σ Γ' (tCoFix mfix' idx') ^ + "\ncorrespond to syntactically distinct terms." + | CoFixRargMismatch idx Γ u mfix1 mfix2 Γ' v mfix1' mfix2' => + "The two co-fixed-points\n" ^ + print_term Σ Γ (tCoFix (mfix1 ++ u :: mfix2) idx) ^ + "\nand\n" ^ print_term Σ Γ' (tCoFix (mfix1' ++ v :: mfix2') idx) ^ + "\nhave a mismatch in the function number " ^ string_of_nat #|mfix1| ^ + ": arguments " ^ string_of_nat u.(rarg) ^ + " and " ^ string_of_nat v.(rarg) ^ "are different." + | CoFixMfixMismatch idx Γ mfix Γ' mfix' => + "The two co-fixed-points\n" ^ + print_term Σ Γ (tCoFix mfix idx) ^ + "\nand\n" ^ + print_term Σ Γ' (tCoFix mfix' idx) ^ + "\nhave a different number of mutually defined functions." + | StackHeadError leq Γ1 t1 args1 u1 l1 Γ2 t2 u2 l2 e => + "TODO stackheaderror\n" ^ + string_of_conv_error Σ e + | StackTailError leq Γ1 t1 args1 u1 l1 Γ2 t2 u2 l2 e => + "TODO stacktailerror\n" ^ + string_of_conv_error Σ e + | StackMismatch Γ1 t1 args1 l1 Γ2 t2 l2 => + "Convertible terms\n" ^ + print_term Σ Γ1 t1 ^ + "\nand\n" ^ print_term Σ Γ2 t2 ^ + "are convertible/convertible (TODO) but applied to a different number" ^ + " of arguments." + | HeadMismatch leq Γ1 t1 Γ2 t2 => + "Terms\n" ^ + print_term Σ Γ1 t1 ^ + "\nand\n" ^ print_term Σ Γ2 t2 ^ + "\ndo not have the same head when comparing for " ^ + string_of_conv_pb leq + end. + +Definition string_of_type_error Σ (e : type_error) : string := + match e with + | UnboundRel n => "Unbound rel " ^ string_of_nat n + | UnboundVar id => "Unbound var " ^ id + | UnboundEvar ev => "Unbound evar " ^ string_of_nat ev + | UndeclaredConstant c => "Undeclared constant " ^ string_of_kername c + | UndeclaredInductive c => "Undeclared inductive " ^ string_of_kername (inductive_mind c) + | UndeclaredConstructor c i => "Undeclared inductive " ^ string_of_kername (inductive_mind c) + | NotCumulSmaller G Γ t u t' u' e => "Terms are not <= for cumulativity:\n" ^ + print_term Σ Γ t ^ "\nand:\n" ^ print_term Σ Γ u ^ + "\nafter reduction:\n" ^ + print_term Σ Γ t' ^ "\nand:\n" ^ print_term Σ Γ u' ^ + "\nerror:\n" ^ string_of_conv_error Σ e ^ + "\nin universe graph:\n" ^ print_universes_graph G + | NotConvertible G Γ t u => "Terms are not convertible:\n" ^ + print_term Σ Γ t ^ "\nand:\n" ^ print_term Σ Γ u ^ + "\nin universe graph:\n" ^ print_universes_graph G + | NotASort t => "Not a sort" + | NotAProduct t t' => "Not a product" + | NotAnArity t => print_term Σ [] t ^ " is not an arity" + | NotAnInductive t => "Not an inductive" + | IllFormedFix m i => "Ill-formed recursive definition" + | UnsatisfiedConstraints c => "Unsatisfied constraints" + | Msg s => "Msg: " ^ s + end. + +Inductive typing_result (A : Type) := +| Checked (a : A) +| TypeError (t : type_error). +Global Arguments Checked {A} a. +Global Arguments TypeError {A} t. + +Instance typing_monad : Monad typing_result := + {| ret A a := Checked a ; + bind A B m f := + match m with + | Checked a => f a + | TypeError t => TypeError t + end + |}. + +Instance monad_exc : MonadExc type_error typing_result := + { raise A e := TypeError e; + catch A m f := + match m with + | Checked a => m + | TypeError t => f t + end + }. + +Inductive env_error := +| IllFormedDecl (e : string) (e : type_error) +| AlreadyDeclared (id : string). + +Inductive EnvCheck (A : Type) := +| CorrectDecl (a : A) +| EnvError (Σ : global_env_ext) (e : env_error). +Global Arguments EnvError {A} Σ e. +Global Arguments CorrectDecl {A} a. + +Global Instance envcheck_monad : Monad EnvCheck := + {| ret A a := CorrectDecl a ; + bind A B m f := + match m with + | CorrectDecl a => f a + | EnvError g e => EnvError g e + end + |}. + +Global Instance envcheck_monad_exc + : MonadExc (global_env_ext * env_error) EnvCheck := + { raise A '(g, e) := EnvError g e; + catch A m f := + match m with + | CorrectDecl a => m + | EnvError g t => f (g, t) + end + }. + +Definition wrap_error {A} Σ (id : string) (check : typing_result A) : EnvCheck A := + match check with + | Checked a => CorrectDecl a + | TypeError e => EnvError Σ (IllFormedDecl id e) + end. \ No newline at end of file diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index 13dedbd2c..330ad9ff0 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -6,12 +6,14 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICWeakening PCUICPosition PCUICCumulativity PCUICSafeLemmata PCUICSN PCUICPretty PCUICArities PCUICConfluence PCUICSize PCUICContextConversion PCUICConversion PCUICWfUniverses + PCUICGlobalEnv (* Used for support lemmas *) PCUICInductives PCUICWfUniverses PCUICContexts PCUICSubstitution PCUICSpine PCUICInductiveInversion PCUICClosed PCUICUnivSubstitution PCUICWeakeningEnv. -From MetaCoq.SafeChecker Require Import PCUICSafeReduce PCUICSafeConversion PCUICWfReduction. +From MetaCoq.SafeChecker Require Import PCUICSafeReduce PCUICErrors PCUICEqualityDec + PCUICSafeConversion PCUICWfReduction PCUICWfEnv PCUICTypeChecker. From Equations Require Import Equations. Require Import ssreflect ssrbool. @@ -28,1406 +30,8 @@ Ltac Coq.Program.Tactics.program_solve_wf ::= end end. -Lemma weakening_sq `{cf : checker_flags} {Σ Γ} Γ' {t T} : - ∥ wf Σ.1 ∥ -> ∥ wf_local Σ (Γ ,,, Γ') ∥ -> - ∥ Σ ;;; Γ |- t : T ∥ -> - ∥ Σ ;;; Γ ,,, Γ' |- lift0 #|Γ'| t : lift0 #|Γ'| T ∥. -Proof. - intros; sq; now eapply weakening. -Defined. - -Definition wf_local_rel_abs_sq `{cf : checker_flags} {Σ Γ Γ' A na} : - ∥ wf_local_rel Σ Γ Γ' ∥ -> {u & ∥ Σ ;;; Γ ,,, Γ' |- A : tSort u ∥ } - -> ∥ wf_local_rel Σ Γ (Γ',, vass na A) ∥. -Proof. - intros H [u Hu]; sq. now eapply wf_local_rel_abs. -Qed. - -Definition wf_local_rel_def_sq `{cf : checker_flags} {Σ Γ Γ' t A na} : - ∥ wf_local_rel Σ Γ Γ' ∥ -> ∥ isType Σ (Γ ,,, Γ') A ∥ -> ∥ Σ ;;; Γ ,,, Γ' |- t : A ∥ - -> ∥ wf_local_rel Σ Γ (Γ',, vdef na t A) ∥. -Proof. - intros; sq. now eapply wf_local_rel_def. -Qed. - -Definition wf_local_rel_app_inv_sq `{cf : checker_flags} {Σ Γ1 Γ2 Γ3} : - ∥ wf_local_rel Σ Γ1 Γ2 ∥ -> ∥ wf_local_rel Σ (Γ1 ,,, Γ2) Γ3 ∥ - -> ∥ wf_local_rel Σ Γ1 (Γ2 ,,, Γ3) ∥. -Proof. - intros; sq; now eapply wf_local_rel_app_inv. -Qed. - -Fixpoint monad_All {T} {M : Monad T} {A} {P} (f : forall x, T (P x)) l : T (@All A P l) - := match l with - | [] => ret All_nil - | a :: l => X <- f a ;; - Y <- monad_All f l ;; - ret (All_cons X Y) - end. - -Fixpoint monad_All2 {T E} {M : Monad T} {M' : MonadExc E T} wrong_sizes - {A B R} (f : forall x y, T (R x y)) l1 l2 - : T (@All2 A B R l1 l2) - := match l1, l2 with - | [], [] => ret All2_nil - | a :: l1, b :: l2 => X <- f a b ;; - Y <- monad_All2 wrong_sizes f l1 l2 ;; - ret (All2_cons X Y) - | _, _ => raise wrong_sizes - end. - - -Definition monad_prod {T} {M : Monad T} {A B} (x : T A) (y : T B): T (A * B) - := X <- x ;; Y <- y ;; ret (X, Y). - -Definition check_dec {T E} {M : Monad T} {M' : MonadExc E T} (e : E) {P} - (H : {P} + {~ P}) - : T P - := match H with - | left x => ret x - | right _ => raise e - end. - -Definition check_eq_true {T E} {M : Monad T} {M' : MonadExc E T} b (e : E) - : T (b = true) - := if b return T (b = true) then ret eq_refl else raise e. - -Program Definition check_eq_nat {T E} {M : Monad T} {M' : MonadExc E T} n m (e : E) - : T (n = m) - := match Nat.eq_dec n m with - | left p => ret p - | right p => raise e - end. - -Definition mkApps_decompose_app_rec t l : - mkApps t l = mkApps (fst (decompose_app_rec t l)) (snd (decompose_app_rec t l)). -Proof. - revert l; induction t; try reflexivity. - intro l; cbn in *. - transitivity (mkApps t1 ((t2 ::l))). reflexivity. - now rewrite IHt1. -Qed. - -Definition mkApps_decompose_app t : - t = mkApps (fst (decompose_app t)) (snd (decompose_app t)) - := mkApps_decompose_app_rec t []. - -Lemma isType_red {cf:checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} {Γ T U} : - isType Σ Γ T -> red Σ Γ T U -> isType Σ Γ U. -Proof. - intros [s Hs] red; exists s. - eapply subject_reduction; eauto. -Qed. - -Derive NoConfusion EqDec for allowed_eliminations. - -Inductive type_error := -| UnboundRel (n : nat) -| UnboundVar (id : string) -| UnboundEvar (ev : nat) -| UndeclaredConstant (c : kername) -| UndeclaredInductive (c : inductive) -| UndeclaredConstructor (c : inductive) (i : nat) -| NotCumulSmaller (G : universes_graph) (Γ : context) (t u t' u' : term) (e : ConversionError) -| NotConvertible (G : universes_graph) (Γ : context) (t u : term) -| NotASort (t : term) -| NotAProduct (t t' : term) -| NotAnInductive (t : term) -| NotAnArity (t : term) -| IllFormedFix (m : mfixpoint term) (i : nat) -| UnsatisfiedConstraints (c : ConstraintSet.t) -| Msg (s : string). - -Definition print_level := string_of_level. - -Definition string_of_Z z := - if (z <=? 0)%Z then "-" ^ string_of_nat (Z.to_nat (- z)) else string_of_nat (Z.to_nat z). - -Definition print_edge '(l1, n, l2) - := "(" ^ print_level l1 ^ ", " ^ string_of_Z n ^ ", " - ^ print_level l2 ^ ")". - -Definition print_universes_graph (G : universes_graph) := - let levels := LevelSet.elements G.1.1 in - let edges := wGraph.EdgeSet.elements G.1.2 in - string_of_list print_level levels - ^ "\n" ^ string_of_list print_edge edges. - -Definition string_of_conv_pb (c : conv_pb) : string := - match c with - | Conv => "conversion" - | Cumul => "cumulativity" - end. - -Definition print_term Σ Γ t := - print_term Σ Γ true false t. - -Fixpoint string_of_conv_error Σ (e : ConversionError) : string := - match e with - | NotFoundConstants c1 c2 => - "Both constants " ^ string_of_kername c1 ^ " and " ^ string_of_kername c2 ^ - " are not found in the environment." - | NotFoundConstant c => - "Constant " ^ string_of_kername c ^ - " common in both terms is not found in the environment." - | LambdaNotConvertibleAnn Γ1 na A1 t1 Γ2 na' A2 t2 => - "When comparing\n" ^ print_term Σ Γ1 (tLambda na A1 t1) ^ - "\nand\n" ^ print_term Σ Γ2 (tLambda na' A2 t2) ^ - "\nbinding annotations are not convertible\n" - | LambdaNotConvertibleTypes Γ1 na A1 t1 Γ2 na' A2 t2 e => - "When comparing\n" ^ print_term Σ Γ1 (tLambda na A1 t1) ^ - "\nand\n" ^ print_term Σ Γ2 (tLambda na' A2 t2) ^ - "\ntypes are not convertible:\n" ^ - string_of_conv_error Σ e - | ProdNotConvertibleAnn Γ1 na A1 B1 Γ2 na' A2 B2 => - "When comparing\n" ^ print_term Σ Γ1 (tProd na A1 B1) ^ - "\nand\n" ^ print_term Σ Γ2 (tProd na' A2 B2) ^ - "\nbinding annotations are not convertible:\n" - | ProdNotConvertibleDomains Γ1 na A1 B1 Γ2 na' A2 B2 e => - "When comparing\n" ^ print_term Σ Γ1 (tProd na A1 B1) ^ - "\nand\n" ^ print_term Σ Γ2 (tProd na' A2 B2) ^ - "\ndomains are not convertible:\n" ^ - string_of_conv_error Σ e - | CaseOnDifferentInd Γ ind par p c brs Γ' ind' par' p' c' brs' => - "The two stuck pattern-matching\n" ^ - print_term Σ Γ (tCase (ind, par) p c brs) ^ - "\nand\n" ^ print_term Σ Γ' (tCase (ind', par') p' c' brs') ^ - "\nare done on distinct inductive types." - | CaseBranchNumMismatch - ind par Γ p c brs1 m br brs2 Γ' p' c' brs1' m' br' brs2' => - "The two stuck pattern-matching\n" ^ - print_term Σ Γ (tCase (ind, par) p c (brs1 ++ (m,br) :: brs2)) ^ - "\nand\n" ^ - print_term Σ Γ' (tCase (ind, par) p' c' (brs1' ++ (m',br') :: brs2')) ^ - "\nhave a mistmatch in the branch number " ^ string_of_nat #|brs1| ^ - "\nthe number of parameters do not coincide\n" ^ - print_term Σ Γ br ^ - "\nhas " ^ string_of_nat m ^ " parameters while\n" ^ - print_term Σ Γ br' ^ - "\nhas " ^ string_of_nat m' ^ "." - | DistinctStuckProj Γ p c Γ' p' c' => - "The two stuck projections\n" ^ - print_term Σ Γ (tProj p c) ^ - "\nand\n" ^ - print_term Σ Γ' (tProj p' c') ^ - "\nare syntactically different." - | CannotUnfoldFix Γ mfix idx Γ' mfix' idx' => - "The two fixed-points\n" ^ - print_term Σ Γ (tFix mfix idx) ^ - "\nand\n" ^ print_term Σ Γ' (tFix mfix' idx') ^ - "\ncorrespond to syntactically distinct terms that can't be unfolded." - | FixRargMismatch idx Γ u mfix1 mfix2 Γ' v mfix1' mfix2' => - "The two fixed-points\n" ^ - print_term Σ Γ (tFix (mfix1 ++ u :: mfix2) idx) ^ - "\nand\n" ^ print_term Σ Γ' (tFix (mfix1' ++ v :: mfix2') idx) ^ - "\nhave a mismatch in the function number " ^ string_of_nat #|mfix1| ^ - ": arguments " ^ string_of_nat u.(rarg) ^ - " and " ^ string_of_nat v.(rarg) ^ "are different." - | FixMfixMismatch idx Γ mfix Γ' mfix' => - "The two fixed-points\n" ^ - print_term Σ Γ (tFix mfix idx) ^ - "\nand\n" ^ - print_term Σ Γ' (tFix mfix' idx) ^ - "\nhave a different number of mutually defined functions." - | DistinctCoFix Γ mfix idx Γ' mfix' idx' => - "The two cofixed-points\n" ^ - print_term Σ Γ (tCoFix mfix idx) ^ - "\nand\n" ^ print_term Σ Γ' (tCoFix mfix' idx') ^ - "\ncorrespond to syntactically distinct terms." - | CoFixRargMismatch idx Γ u mfix1 mfix2 Γ' v mfix1' mfix2' => - "The two co-fixed-points\n" ^ - print_term Σ Γ (tCoFix (mfix1 ++ u :: mfix2) idx) ^ - "\nand\n" ^ print_term Σ Γ' (tCoFix (mfix1' ++ v :: mfix2') idx) ^ - "\nhave a mismatch in the function number " ^ string_of_nat #|mfix1| ^ - ": arguments " ^ string_of_nat u.(rarg) ^ - " and " ^ string_of_nat v.(rarg) ^ "are different." - | CoFixMfixMismatch idx Γ mfix Γ' mfix' => - "The two co-fixed-points\n" ^ - print_term Σ Γ (tCoFix mfix idx) ^ - "\nand\n" ^ - print_term Σ Γ' (tCoFix mfix' idx) ^ - "\nhave a different number of mutually defined functions." - | StackHeadError leq Γ1 t1 args1 u1 l1 Γ2 t2 u2 l2 e => - "TODO stackheaderror\n" ^ - string_of_conv_error Σ e - | StackTailError leq Γ1 t1 args1 u1 l1 Γ2 t2 u2 l2 e => - "TODO stacktailerror\n" ^ - string_of_conv_error Σ e - | StackMismatch Γ1 t1 args1 l1 Γ2 t2 l2 => - "Convertible terms\n" ^ - print_term Σ Γ1 t1 ^ - "\nand\n" ^ print_term Σ Γ2 t2 ^ - "are convertible/convertible (TODO) but applied to a different number" ^ - " of arguments." - | HeadMismatch leq Γ1 t1 Γ2 t2 => - "Terms\n" ^ - print_term Σ Γ1 t1 ^ - "\nand\n" ^ print_term Σ Γ2 t2 ^ - "\ndo not have the same head when comparing for " ^ - string_of_conv_pb leq - end. - -Definition string_of_type_error Σ (e : type_error) : string := - match e with - | UnboundRel n => "Unbound rel " ^ string_of_nat n - | UnboundVar id => "Unbound var " ^ id - | UnboundEvar ev => "Unbound evar " ^ string_of_nat ev - | UndeclaredConstant c => "Undeclared constant " ^ string_of_kername c - | UndeclaredInductive c => "Undeclared inductive " ^ string_of_kername (inductive_mind c) - | UndeclaredConstructor c i => "Undeclared inductive " ^ string_of_kername (inductive_mind c) - | NotCumulSmaller G Γ t u t' u' e => "Terms are not <= for cumulativity:\n" ^ - print_term Σ Γ t ^ "\nand:\n" ^ print_term Σ Γ u ^ - "\nafter reduction:\n" ^ - print_term Σ Γ t' ^ "\nand:\n" ^ print_term Σ Γ u' ^ - "\nerror:\n" ^ string_of_conv_error Σ e ^ - "\nin universe graph:\n" ^ print_universes_graph G - | NotConvertible G Γ t u => "Terms are not convertible:\n" ^ - print_term Σ Γ t ^ "\nand:\n" ^ print_term Σ Γ u ^ - "\nin universe graph:\n" ^ print_universes_graph G - | NotASort t => "Not a sort" - | NotAProduct t t' => "Not a product" - | NotAnArity t => print_term Σ [] t ^ " is not an arity" - | NotAnInductive t => "Not an inductive" - | IllFormedFix m i => "Ill-formed recursive definition" - | UnsatisfiedConstraints c => "Unsatisfied constraints" - | Msg s => "Msg: " ^ s - end. - -Inductive typing_result (A : Type) := -| Checked (a : A) -| TypeError (t : type_error). -Global Arguments Checked {A} a. -Global Arguments TypeError {A} t. - -Instance typing_monad : Monad typing_result := - {| ret A a := Checked a ; - bind A B m f := - match m with - | Checked a => f a - | TypeError t => TypeError t - end - |}. - -Instance monad_exc : MonadExc type_error typing_result := - { raise A e := TypeError e; - catch A m f := - match m with - | Checked a => m - | TypeError t => f t - end - }. - - -Lemma wf_local_rel_inv `{checker_flags} {Σ Γ1 Γ2} (w : wf_local_rel Σ Γ1 Γ2) : - forall d Γ2', Γ2 = Γ2' ,, d -> - wf_local_rel Σ Γ1 Γ2' × (∑ u, Σ ;;; Γ1 ,,, Γ2' |- d.(decl_type) : tSort u) × - match d.(decl_body) with - | Some b => Σ ;;; Γ1 ,,, Γ2' |- b : d.(decl_type) - | None => True - end. -Proof. - intros d Γ. - destruct w; inversion 1; subst; cbn in *. - split; [assumption|]. split; [assumption| trivial]. - split; [assumption|]. split; [assumption| trivial]. -Defined. - -Definition wf_local_app_inv `{cf : checker_flags} {Σ Γ1 Γ2} : - wf_local Σ Γ1 -> wf_local_rel Σ Γ1 Γ2 - -> wf_local Σ (Γ1 ,,, Γ2). -Proof. - intros H1 H2. - apply wf_local_local_rel, wf_local_rel_app_inv. - now apply wf_local_rel_local. - now rewrite app_context_nil_l. -Defined. - -Definition fix_context_i i mfix := - List.rev (mapi_rec (fun i (d : BasicAst.def term) - => vass (dname d) ((lift0 i) (dtype d))) mfix i). - -Lemma lift_fix_context mfix : forall k k' j, j <= k -> - fix_context_i (k + k') mfix = lift_context k' j (fix_context_i k mfix). -Proof. - induction mfix. reflexivity. - intros k k' j Hj. - change ([vass (dname a) ((lift0 (k + k')) (dtype a))] ,,, fix_context_i (S (k + k')) mfix = lift_context k' j ([vass (dname a) ((lift0 k) (dtype a))] ,,, fix_context_i (S k) mfix)). - erewrite lift_context_app, (IHmfix (S k) k' (S j)); cbn. - unfold map_decl, vass; cbn. now rewrite simpl_lift. - now apply le_n_S. -Qed. - -Lemma wf_ext_gc_of_uctx {cf:checker_flags} {Σ : global_env_ext} (HΣ : ∥ wf_ext Σ ∥) - : ∑ uctx', gc_of_uctx (global_ext_uctx Σ) = Some uctx'. -Proof. - pose proof (global_ext_uctx_consistent _ HΣ) as HC. - destruct Σ as [Σ φ]. - simpl in HC. - unfold gc_of_uctx; simpl in *. - apply gc_consistent_iff in HC. - destruct (gc_of_constraints (global_ext_constraints (Σ, φ))). - eexists; reflexivity. - contradiction HC. -Defined. - -Lemma wf_ext_is_graph {cf:checker_flags} {Σ : global_env_ext} (HΣ : ∥ wf_ext Σ ∥) - : ∑ G, is_graph_of_uctx G (global_ext_uctx Σ). -Proof. - destruct (wf_ext_gc_of_uctx HΣ) as [uctx Huctx]. - exists (make_graph uctx). unfold is_graph_of_uctx. now rewrite Huctx. -Defined. - -Lemma map_squash {A B} (f : A -> B) : ∥ A ∥ -> ∥ B ∥. -Proof. - intros []; constructor; auto. -Qed. - -Lemma destArity_mkApps_None ctx t l : - destArity ctx t = None -> destArity ctx (mkApps t l) = None. -Proof. - induction l in t |- *. trivial. - intros H. cbn. apply IHl. reflexivity. -Qed. - -Lemma destArity_mkApps_Ind ctx ind u l : - destArity ctx (mkApps (tInd ind u) l) = None. -Proof. - apply destArity_mkApps_None. reflexivity. -Qed. - -Section Typecheck. - Context {cf : checker_flags} {Σ : global_env_ext} (HΣ : ∥ wf Σ ∥) - (Hφ : ∥ on_udecl Σ.1 Σ.2 ∥) - (G : universes_graph) (HG : is_graph_of_uctx G (global_ext_uctx Σ)). - - (* We get stack overflow on Qed after Equations definitions when this is transparent *) - Opaque reduce_stack_full. - - Local Definition HΣ' : ∥ wf_ext Σ ∥. - Proof. - destruct HΣ, Hφ; now constructor. - Defined. - - Definition isType_welltyped {Γ t} - : isType Σ Γ t -> welltyped Σ Γ t. - Proof. - intros []. now econstructor. - Qed. - - Lemma validity_wf {Γ t T} : - ∥ wf_local Σ Γ ∥ -> ∥ Σ ;;; Γ |- t : T ∥ -> welltyped Σ Γ T. - Proof. - destruct HΣ as [wΣ]. intros [wΓ] [X]. - intros. eapply validity_term in X; try assumption. - destruct X. now exists (tSort x). - Defined. - - Lemma wat_welltyped {Γ T} : - ∥ isType Σ Γ T ∥ -> welltyped Σ Γ T. - Proof. - destruct HΣ as [wΣ]. intros [X]. - now apply isType_welltyped. - Defined. - - Definition hnf := reduce_term RedFlags.default Σ HΣ. - - Theorem hnf_sound {Γ t h} : ∥ red (fst Σ) Γ t (hnf Γ t h) ∥. - Proof. - apply reduce_term_sound. - Defined. - - Theorem hnf_complete {Γ t h} : ∥whnf RedFlags.default Σ Γ (hnf Γ t h)∥. - Proof. - apply reduce_term_complete. - Qed. - - Inductive view_sort : term -> Type := - | view_sort_sort s : view_sort (tSort s) - | view_sort_other t : ~isSort t -> view_sort t. - - Equations view_sortc (t : term) : view_sort t := - view_sortc (tSort s) := view_sort_sort s; - view_sortc t := view_sort_other t _. - - Equations? reduce_to_sort (Γ : context) (t : term) (h : welltyped Σ Γ t) - : typing_result (∑ u, ∥ red (fst Σ) Γ t (tSort u) ∥) := - reduce_to_sort Γ t h with view_sortc t := { - | view_sort_sort s => ret (s; sq (refl_red _ _ _)); - | view_sort_other t _ with inspect (hnf Γ t h) := - | exist hnft eq with view_sortc hnft := { - | view_sort_sort s => ret (s; _); - | view_sort_other t _ => raise (NotASort t) - } - }. - Proof. - pose proof (hnf_sound (h:=h)). - now rewrite eq. - Qed. - - Lemma reduce_to_sort_complete {Γ t wt} e : - reduce_to_sort Γ t wt = TypeError e -> - (forall s, red Σ Γ t (tSort s) -> False). - Proof. - funelim (reduce_to_sort Γ t wt); try congruence. - intros _ s r. - pose proof (@hnf_complete Γ t0 h) as [wh]. - pose proof (@hnf_sound Γ t0 h) as [r']. - destruct HΣ. - eapply red_confluence in r as (?&r1&r2); eauto. - apply invert_red_sort in r2 as ->. - eapply whnf_red_inv in r1; eauto. - depelim r1. - clear Heq. - rewrite H in n0. - now cbn in n0. - Qed. - - Inductive view_prod : term -> Type := - | view_prod_prod na A b : view_prod (tProd na A b) - | view_prod_other t : ~isProd t -> view_prod t. - - Equations view_prodc (t : term) : view_prod t := - view_prodc (tProd na A b) := view_prod_prod na A b; - view_prodc t := view_prod_other t _. - - Equations? reduce_to_prod (Γ : context) (t : term) (h : welltyped Σ Γ t) - : typing_result (∑ na a b, ∥ red (fst Σ) Γ t (tProd na a b) ∥) := - reduce_to_prod Γ t h with view_prodc t := { - | view_prod_prod na a b => ret (na; a; b; sq (refl_red _ _ _)); - | view_prod_other t _ with inspect (hnf Γ t h) := - | exist hnft eq with view_prodc hnft := { - | view_prod_prod na a b => ret (na; a; b; _); - | view_prod_other t' _ => raise (NotAProduct t t') - } - }. - Proof. - pose proof (hnf_sound (h:=h)). - now rewrite eq. - Qed. - - Lemma reduce_to_prod_complete {Γ t wt} e : - reduce_to_prod Γ t wt = TypeError e -> - (forall na a b, red Σ Γ t (tProd na a b) -> False). - Proof. - funelim (reduce_to_prod Γ t wt); try congruence. - intros _ na a b r. - pose proof (@hnf_complete Γ t0 h) as [wh]. - pose proof (@hnf_sound Γ t0 h) as [r']. - destruct HΣ. - eapply red_confluence in r as (?&r1&r2); eauto. - apply invert_red_prod in r2 as (?&?&(->&?)&?); auto. - eapply whnf_red_inv in r1; auto. - depelim r1. - clear Heq. - rewrite H in n0. - now cbn in n0. - Qed. - - Definition isInd (t : term) : bool := - match t with - | tInd _ _ => true - | _ => false - end. - - Inductive view_ind : term -> Type := - | view_ind_tInd ind u : view_ind (tInd ind u) - | view_ind_other t : negb (isInd t) -> view_ind t. - - Equations view_indc (t : term) : view_ind t := - view_indc (tInd ind u) => view_ind_tInd ind u; - view_indc t => view_ind_other t _. - - Equations? reduce_to_ind (Γ : context) (t : term) (h : welltyped Σ Γ t) - : typing_result (∑ i u l, ∥ red (fst Σ) Γ t (mkApps (tInd i u) l) ∥) := - reduce_to_ind Γ t h with inspect (decompose_app t) := { - | exist (thd, args) eq_decomp with view_indc thd := { - | view_ind_tInd i u => ret (i; u; args; sq _); - | view_ind_other t _ with inspect (reduce_stack RedFlags.default Σ HΣ Γ t Empty h) := { - | exist (hnft, π) eq with view_indc hnft := { - | view_ind_tInd i u with inspect (decompose_stack π) := { - | exist (l, _) eq_decomp => ret (i; u; l; _) - }; - | view_ind_other _ _ => raise (NotAnInductive t) - } - } - } - }. - Proof. - - assert (X : mkApps (tInd i u) args = t); [|rewrite X; apply refl_red]. - etransitivity. 2: symmetry; eapply mkApps_decompose_app. - now rewrite <- eq_decomp. - - pose proof (reduce_stack_sound RedFlags.default Σ HΣ _ _ Empty h). - rewrite <- eq in H. - cbn in *. - assert (π = appstack l ε) as ->. - 2: { now rewrite zipc_appstack in H. } - unfold reduce_stack in eq. - destruct reduce_stack_full as (?&_&stack_val&?). - subst x. - unfold Pr in stack_val. - cbn in *. - assert (decomp: decompose_stack π = ((decompose_stack π).1, ε)). - { rewrite stack_val. - now destruct decompose_stack. } - apply decompose_stack_eq in decomp as ->. - now rewrite <- eq_decomp0. - Qed. - - Lemma reduce_to_ind_complete Γ ty wat e : - reduce_to_ind Γ ty wat = TypeError e -> - forall ind u args, - red Σ Γ ty (mkApps (tInd ind u) args) -> - False. - Proof. - funelim (reduce_to_ind Γ ty wat); try congruence. - intros _ ind u args r. - pose proof (reduce_stack_whnf RedFlags.default Σ HΣ Γ t ε h) as wh. - unfold reduce_stack in *. - destruct reduce_stack_full as ((hd&π)&r'&stack_valid&(notapp&_)). - destruct wh as [wh]. - apply Req_red in r' as [r']. - unfold Pr in stack_valid. - cbn in *. - destruct HΣ. - eapply red_confluence in r as (?&r1&r2); [|eassumption|exact r']. - assert (exists args, π = appstack args ε) as (?&->). - { exists ((decompose_stack π).1). - assert (decomp: decompose_stack π = ((decompose_stack π).1, ε)). - { now rewrite stack_valid; destruct decompose_stack. } - now apply decompose_stack_eq in decomp. } - - unfold zipp in wh. - rewrite stack_context_appstack decompose_stack_appstack in wh. - rewrite zipc_appstack in r1. - cbn in *. - rewrite app_nil_r in wh. - apply red_mkApps_tInd in r2 as (?&->&?); auto. - eapply whnf_red_inv in r1; eauto. - apply whnf_red_mkApps_inv in r1 as (?&?); auto. - depelim w. - noconf e0. - discriminate i0. - Qed. - - Definition isconv Γ := isconv_term Σ HΣ Hφ G HG Γ Conv. - Definition iscumul Γ := isconv_term Σ HΣ Hφ G HG Γ Cumul. - - Program Definition convert Γ t u - (ht : welltyped Σ Γ t) (hu : welltyped Σ Γ u) - : typing_result (∥ Σ ;;; Γ |- t = u ∥) := - match eqb_term Σ G t u with true => ret _ | false => - match isconv Γ t ht u hu with - | ConvSuccess => ret _ - | ConvError e => (* fallback *) (* nico *) - let t' := hnf Γ t ht in - let u' := hnf Γ u hu in - (* match leq_term (snd Σ) t' u' with true => ret _ | false => *) - raise (NotCumulSmaller G Γ t u t' u' e) - (* end *) - end end. - Next Obligation. - symmetry in Heq_anonymous; eapply eqb_term_spec in Heq_anonymous; tea. - constructor. now constructor. - Qed. - Next Obligation. - symmetry in Heq_anonymous; apply isconv_term_sound in Heq_anonymous. - assumption. - Qed. - - Program Definition convert_leq Γ t u - (ht : welltyped Σ Γ t) (hu : welltyped Σ Γ u) - : typing_result (∥ Σ ;;; Γ |- t <= u ∥) := - match leqb_term Σ G t u with true => ret _ | false => - match iscumul Γ t ht u hu with - | ConvSuccess => ret _ - | ConvError e => (* fallback *) (* nico *) - let t' := hnf Γ t ht in - let u' := hnf Γ u hu in - (* match leq_term (snd Σ) t' u' with true => ret _ | false => *) - raise (NotCumulSmaller G Γ t u t' u' e) - (* end *) - end end. - Next Obligation. - symmetry in Heq_anonymous; eapply leqb_term_spec in Heq_anonymous; tea. - constructor. now constructor. - Qed. - Next Obligation. - symmetry in Heq_anonymous; apply isconv_term_sound in Heq_anonymous. - assumption. - Qed. - - Section InferAux. - Variable (infer : forall Γ (HΓ : ∥ wf_local Σ Γ ∥) (t : term), - typing_result ({ A : term & ∥ Σ ;;; Γ |- t : A ∥ })). - - Program Definition infer_type Γ HΓ t - : typing_result ({u : Universe.t & ∥ Σ ;;; Γ |- t : tSort u ∥}) := - tx <- infer Γ HΓ t ;; - u <- reduce_to_sort Γ tx.π1 _ ;; - ret (u.π1; _). - Next Obligation. - eapply validity_wf; eassumption. - Defined. - Next Obligation. - destruct HΣ, HΓ, X, X0. - now constructor; eapply type_reduction. - Defined. - - Program Definition infer_cumul Γ HΓ t A (hA : ∥ isType Σ Γ A ∥) - : typing_result (∥ Σ ;;; Γ |- t : A ∥) := - A' <- infer Γ HΓ t ;; - X <- convert_leq Γ A'.π1 A _ _ ;; - ret _. - Next Obligation. now eapply validity_wf. Qed. - Next Obligation. destruct hA; now apply wat_welltyped. Qed. - Next Obligation. - destruct HΣ, HΓ, hA, X, X0. constructor. eapply type_Cumul'; eassumption. - Qed. - End InferAux. - - - Program Definition lookup_ind_decl ind - : typing_result - ({decl & {body & declared_inductive (fst Σ) decl ind body}}) := - match lookup_env (fst Σ) ind.(inductive_mind) with - | Some (InductiveDecl decl) => - match nth_error decl.(ind_bodies) ind.(inductive_ind) with - | Some body => ret (decl; body; _) - | None => raise (UndeclaredInductive ind) - end - | _ => raise (UndeclaredInductive ind) - end. - Next Obligation. - split. - - symmetry in Heq_anonymous0. - etransitivity. eassumption. reflexivity. - - now symmetry. - Defined. - - Lemma check_constraints_spec ctrs - : check_constraints G ctrs -> valid_constraints (global_ext_constraints Σ) ctrs. - Proof. - pose proof HΣ'. - intros HH. - refine (check_constraints_spec G (global_ext_uctx Σ) _ _ HG _ HH). - now apply wf_ext_global_uctx_invariants. - now apply global_ext_uctx_consistent. - Qed. - - Lemma is_graph_of_uctx_levels (l : Level.t) : - LevelSet.mem l (uGraph.wGraph.V G) -> - LevelSet.mem l (global_ext_levels Σ). - Proof. - unfold is_graph_of_uctx in HG. - case_eq (gc_of_uctx (global_ext_uctx Σ)); [intros [lvs cts] XX|intro XX]; - rewrite -> XX in *; simpl in *; [|contradiction]. - unfold gc_of_uctx in XX; simpl in XX. - destruct (gc_of_constraints Σ); [|discriminate]. - inversion XX; subst. generalize (global_ext_levels Σ); intros lvs; cbn. - clear. intro H. apply LevelSet.mem_spec. now apply LevelSet.mem_spec in H. - Qed. - - - Program Definition check_consistent_instance uctx u - : typing_result (consistent_instance_ext Σ uctx u) - := match uctx with - | Monomorphic_ctx _ => - check_eq_nat #|u| 0 (Msg "monomorphic instance should be of length 0") - | Polymorphic_ctx (inst, cstrs) => - let '(inst, cstrs) := AUContext.repr (inst, cstrs) in - check_eq_true (forallb (fun l => LevelSet.mem l (uGraph.wGraph.V G)) u) - (Msg "instance does not have the right length") ;; - (* check_eq_true (forallb (fun l => LevelSet.mem l lvs) u) ;; *) - X <- check_eq_nat #|u| #|inst| - (Msg "instance does not have the right length");; - match check_constraints G (subst_instance_cstrs u cstrs) with - | true => ret (conj _ _) - | false => raise (Msg "ctrs not satisfiable") - end - (* #|u| = #|inst| /\ valid_constraints φ (subst_instance_cstrs u cstrs) *) - end. - Next Obligation. - eapply forallb_All in H. eapply All_forallb'; tea. - clear -cf HG. intros x; simpl. apply is_graph_of_uctx_levels. - Qed. - Next Obligation. - repeat split. - - now rewrite mapi_length in X. - - eapply check_constraints_spec; eauto. - Qed. - - Definition eqb_opt_term (t u : option term) := - match t, u with - | Some t, Some u => eqb_term Σ G t u - | None, None => true - | _, _ => false - end. - - Lemma eqb_opt_term_spec t u - : eqb_opt_term t u -> eq_opt_term false Σ (global_ext_constraints Σ) t u. - Proof. - destruct t, u; try discriminate; cbn. - apply eqb_term_spec; tea. trivial. - Qed. - - Definition eqb_decl (d d' : context_decl) := - eqb_binder_annot d.(decl_name) d'.(decl_name) && - eqb_opt_term d.(decl_body) d'.(decl_body) && eqb_term Σ G d.(decl_type) d'.(decl_type). - - Lemma eqb_decl_spec d d' - : eqb_decl d d' -> eq_decl false Σ (global_ext_constraints Σ) d d'. - Proof. - unfold eqb_decl, eq_decl. - intro H. rtoProp. apply eqb_opt_term_spec in H1. - eapply eqb_term_spec in H0; tea. - eapply eqb_binder_annot_spec in H. - repeat split; eauto. - Qed. - - Definition eqb_context (Γ Δ : context) := forallb2 eqb_decl Γ Δ. - - Lemma eqb_context_spec Γ Δ - : eqb_context Γ Δ -> eq_context false Σ (global_ext_constraints Σ) Γ Δ. - Proof. - unfold eqb_context, eq_context. - intro HH. apply forallb2_All2 in HH. - eapply All2_impl; try eassumption. - cbn. apply eqb_decl_spec. - Qed. - - Obligation Tactic := Program.Tactics.program_simplify ; eauto 2. - - Program Definition check_is_allowed_elimination (u : Universe.t) (al : allowed_eliminations) : - typing_result (is_allowed_elimination Σ u al) := - let o := - match al return option (is_allowed_elimination Σ u al) with - | IntoSProp => - match Universe.is_sprop u with - | true => Some _ - | false => None - end - | IntoPropSProp => - match is_propositional u with - | true => Some _ - | false => None - end - | IntoSetPropSProp => - match is_propositional u || check_eqb_universe G u Universe.type0 with - | true => Some _ - | false => None - end - | IntoAny => Some _ - end in - match o with - | Some t => Checked t - | None => TypeError (Msg "Cannot eliminate over this sort") - end. - Next Obligation. - unfold is_allowed_elimination, is_allowed_elimination0. - destruct check_univs; auto. - intros val sat. - symmetry in Heq_anonymous. - apply is_sprop_val with (v := val) in Heq_anonymous. - now rewrite Heq_anonymous. - Qed. - Next Obligation. - unfold is_allowed_elimination, is_allowed_elimination0. - destruct check_univs; auto. - intros val sat. - unfold is_propositional in *. - destruct Universe.is_prop eqn:prop. - - apply is_prop_val with (v := val) in prop; rewrite prop; auto. - - destruct Universe.is_sprop eqn:sprop. - + apply is_sprop_val with (v := val) in sprop; rewrite sprop; auto. - + discriminate. - Qed. - Next Obligation. - unfold is_allowed_elimination, is_allowed_elimination0. - destruct check_univs eqn:cu; auto. - intros val sat. - unfold is_propositional in *. - destruct Universe.is_prop eqn:prop. - - apply is_prop_val with (v := val) in prop; rewrite prop; auto. - - destruct Universe.is_sprop eqn:sprop. - + apply is_sprop_val with (v := val) in sprop; rewrite sprop; auto. - + destruct check_eqb_universe eqn:check; [|discriminate]. - eapply check_eqb_universe_spec' in check; eauto. - * unfold eq_universe, eq_universe0 in check. - rewrite cu in check. - specialize (check val sat). - now rewrite check. - * destruct HΣ, Hφ. - now apply wf_ext_global_uctx_invariants. - * destruct HΣ, Hφ. - now apply global_ext_uctx_consistent. - Qed. - Next Obligation. - unfold is_allowed_elimination, is_allowed_elimination0. - destruct check_univs; auto. - Qed. - - Program Fixpoint infer (Γ : context) (HΓ : ∥ wf_local Σ Γ ∥) (t : term) {struct t} - : typing_result ({ A : term & ∥ Σ ;;; Γ |- t : A ∥ }) := - match t with - | tRel n => - match nth_error Γ n with - | Some c => ret ((lift0 (S n)) (decl_type c); _) - | None => raise (UnboundRel n) - end - - | tVar n => raise (UnboundVar n) - | tEvar ev args => raise (UnboundEvar ev) - - | tSort u => - check_eq_true (wf_universeb Σ u) - (Msg ("Sort contains an undeclared level " ^ string_of_sort u));; - ret (tSort (Universe.super u); _) - - | tProd na A B => - s1 <- infer_type infer Γ HΓ A ;; - s2 <- infer_type infer (Γ ,, vass na A) _ B ;; - ret (tSort (Universe.sort_of_product s1.π1 s2.π1); _) - - | tLambda na A t => - s1 <- infer_type infer Γ HΓ A ;; - B <- infer (Γ ,, vass na A) _ t ;; - ret (tProd na A B.π1; _) - - | tLetIn n b b_ty b' => - infer_type infer Γ HΓ b_ty ;; - infer_cumul infer Γ HΓ b b_ty _ ;; - b'_ty <- infer (Γ ,, vdef n b b_ty) _ b' ;; - ret (tLetIn n b b_ty b'_ty.π1; _) - - | tApp t u => - ty <- infer Γ HΓ t ;; - pi <- reduce_to_prod Γ ty.π1 _ ;; - infer_cumul infer Γ HΓ u pi.π2.π1 _ ;; - ret (subst10 u pi.π2.π2.π1; _) - - | tConst cst u => - match lookup_env (fst Σ) cst with - | Some (ConstantDecl d) => - check_consistent_instance d.(cst_universes) u ;; - let ty := subst_instance_constr u d.(cst_type) in - ret (ty; _) - | _ => raise (UndeclaredConstant cst) - end - - | tInd ind u => - d <- lookup_ind_decl ind ;; - check_consistent_instance d.π1.(ind_universes) u ;; - let ty := subst_instance_constr u d.π2.π1.(ind_type) in - ret (ty; _) - - | tConstruct ind k u => - d <- lookup_ind_decl ind ;; - match nth_error d.π2.π1.(ind_ctors) k with - | Some cdecl => - check_consistent_instance d.π1.(ind_universes) u ;; - ret (type_of_constructor d.π1 cdecl (ind, k) u; _) - | None => raise (UndeclaredConstructor ind k) - end - - | tCase (ind, par) p c brs => - cty <- infer Γ HΓ c ;; - I <- reduce_to_ind Γ cty.π1 _ ;; - let '(ind'; I') := I in let '(u; I'') := I' in let '(args; H) := I'' in - check_eq_true (eqb ind ind') - (* bad case info *) - (NotConvertible G Γ (tInd ind u) (tInd ind' u)) ;; - d <- lookup_ind_decl ind' ;; - let '(decl; d') := d in let '(body; HH) := d' in - check_coind <- check_eq_true (negb (isCoFinite (ind_finite decl))) - (Msg "Case on coinductives disallowed") ;; - check_eq_true (ind_npars decl =? par) - (Msg "not the right number of parameters") ;; - pty <- infer Γ HΓ p ;; - match destArity [] pty.π1 with - | None => raise (Msg "the type of the return predicate of a Case is not an arity") - | Some (pctx, ps) => - check_is_allowed_elimination ps (ind_kelim body);; - let params := firstn par args in - match build_case_predicate_type ind decl body params u ps with - | None => raise (Msg "failure in build_case_predicate_type") - | Some pty' => - (* We could avoid one useless sort comparison by only comparing *) - (* the contexts [pctx] and [indctx] (what is done in Coq). *) - match iscumul Γ pty.π1 _ pty' _ with - | ConvError e => raise (NotCumulSmaller G Γ pty.π1 pty' pty.π1 pty' e) - | ConvSuccess => - match map_option_out (build_branches_type ind decl body params u p) with - | None => raise (Msg "failure in build_branches_type") - | Some btys => - let btyswf : ∥ All (isType Σ Γ ∘ snd) btys ∥ := _ in - (fix check_branches (brs btys : list (nat * term)) - (HH : ∥ All (isType Σ Γ ∘ snd) btys ∥) {struct brs} - : typing_result - (All2 (fun br bty => br.1 = bty.1 /\ ∥ Σ ;;; Γ |- br.2 : bty.2 ∥) brs btys) - := match brs, btys with - | [], [] => ret All2_nil - | (n, t) :: brs , (m, A) :: btys => - W <- check_dec (Msg "not nat eq") - (EqDecInstances.nat_eqdec n m) ;; - Z <- infer_cumul infer Γ HΓ t A _ ;; - X <- check_branches brs btys _ ;; - ret (All2_cons (conj _ _) X) - | [], _ :: _ - | _ :: _, [] => raise (Msg "wrong number of branches") - end) brs btys btyswf ;; - ret (mkApps p (List.skipn par args ++ [c]); _) - end - end - end - end - - | tProj (ind, n, k) c => - d <- lookup_ind_decl ind ;; - match nth_error d.π2.π1.(ind_projs) k with - | Some pdecl => - c_ty <- infer Γ HΓ c ;; - I <- reduce_to_ind Γ c_ty.π1 _ ;; - let '(ind'; I') := I in let '(u; I'') := I' in let '(args; H) := I'' in - check_eq_true (eqb ind ind') - (NotConvertible G Γ (tInd ind u) (tInd ind' u)) ;; - check_eq_true (ind_npars d.π1 =? n) - (Msg "not the right number of parameters") ;; - let ty := snd pdecl in - ret (subst0 (c :: List.rev args) (subst_instance_constr u ty); - _) - | None => raise (Msg "projection not found") - end - - | tFix mfix n => - match nth_error mfix n with - | None => raise (IllFormedFix mfix n) - | Some decl => - XX <- (fix check_types (mfix : mfixpoint term) {struct mfix} - : typing_result (∥ All (fun x => isType Σ Γ (dtype x)) mfix ∥) - := match mfix with - | [] => ret (sq All_nil) - | def :: mfix => - (* probably not tail recursive but needed so that next line terminates *) - W <- infer_type infer Γ HΓ (dtype def) ;; - Z <- check_types mfix ;; - ret _ - end) - mfix ;; - YY <- (fix check_bodies (mfix' : mfixpoint term) - (XX : ∥ All (fun x => isType Σ Γ (dtype x)) mfix' ∥) - {struct mfix'} - : typing_result (All (fun d => - ∥ Σ ;;; Γ ,,, fix_context mfix |- dbody d : (lift0 #|fix_context mfix|) (dtype d) ∥ - /\ isLambda (dbody d) = true) mfix') - := match mfix' with - | [] => ret All_nil - | def :: mfix' => - W1 <- infer_cumul infer (Γ ,,, fix_context mfix) _ (dbody def) - (lift0 #|fix_context mfix| (dtype def)) _ ;; - W2 <- check_eq_true (isLambda (dbody def)) - (Msg "not a lambda") ;; - Z <- check_bodies mfix' _ ;; - ret (All_cons (conj W1 W2) Z) - end) mfix _ ;; - guarded <- check_eq_true (fix_guard mfix) (Msg "Unguarded fixpoint") ;; - wffix <- check_eq_true (wf_fixpoint Σ.1 mfix) (Msg "Ill-formed fixpoint: not defined on a mutually inductive family") ;; - ret (dtype decl; _) - end - - | tCoFix mfix n => - match nth_error mfix n with - | None => raise (IllFormedFix mfix n) - | Some decl => - XX <- (fix check_types (mfix : mfixpoint term) {struct mfix} - : typing_result (∥ All (fun x => isType Σ Γ (dtype x)) mfix ∥) - := match mfix with - | [] => ret (sq All_nil) - | def :: mfix => - (* probably not tail recursive but needed so that next line terminates *) - W <- infer_type infer Γ HΓ (dtype def) ;; - Z <- check_types mfix ;; - ret _ - end) - mfix ;; - YY <- (fix check_bodies (mfix' : mfixpoint term) - (XX' : ∥ All (fun x => isType Σ Γ (dtype x)) mfix' ∥) - {struct mfix'} - : typing_result (All (fun d => - ∥ Σ ;;; Γ ,,, fix_context mfix |- dbody d : (lift0 #|fix_context mfix|) (dtype d) ∥) mfix') - := match mfix' with - | [] => ret All_nil - | def :: mfix' => - W1 <- infer_cumul infer (Γ ,,, fix_context mfix) _ (dbody def) - (lift0 #|fix_context mfix| (dtype def)) _ ;; - Z <- check_bodies mfix' _ ;; - ret (All_cons W1 Z) - end) mfix _ ;; - guarded <- check_eq_true (cofix_guard mfix) (Msg "Unguarded cofixpoint") ;; - wfcofix <- check_eq_true (wf_cofixpoint Σ.1 mfix) (Msg "Ill-formed cofixpoint: not producing values in a mutually coinductive family") ;; - ret (dtype decl; _) - end - end. - - (* tRel *) - Next Obligation. intros; sq; now econstructor. Defined. - (* tSort *) - Next Obligation. - eapply (reflect_bP (wf_universe_reflect _ _)) in H. - sq; econstructor; tas. - Defined. - (* tProd *) - Next Obligation. - (* intros Γ HΓ t na A B Heq_t [s ?]; *) - sq; econstructor; cbn; easy. Defined. - Next Obligation. - (* intros Γ HΓ t na A B Heq_t [s1 ?] [s2 ?]; *) - sq; econstructor; eassumption. - Defined. - (* tLambda *) - Next Obligation. - (* intros Γ HΓ t0 na A t Heq_t [s ?]; *) - sq; econstructor; cbn; easy. - Defined. - Next Obligation. - (* intros Γ HΓ t0 na A t Heq_t [s ?] [B ?]; *) - sq; econstructor; eassumption. - Defined. - (* tLetIn *) - Next Obligation. - (* intros Γ HΓ t n b b_ty b' Heq_t [? ?]; *) - sq. econstructor; eauto. - Defined. - Next Obligation. - (* intros Γ HΓ t n b b_ty b' Heq_t [? ?] H0; *) - sq; econstructor; cbn; eauto. - Defined. - Next Obligation. - (* intros Γ HΓ t n b b_ty b' Heq_t [? ?] H0 [? ?]; *) - sq; econstructor; eassumption. - Defined. - - (* tApp *) - Next Obligation. simpl; eauto using validity_wf. Qed. - Next Obligation. - cbn in *; sq. - eapply type_reduction in X1 ; try eassumption. - eapply validity_term in X1 ; try assumption. destruct X1 as [s HH]. - eapply inversion_Prod in HH ; try assumption. - destruct HH as [s1 [_ [HH _]]]. - eexists. eassumption. - Defined. - Next Obligation. - cbn in *; sq; econstructor. - 2: eassumption. - eapply type_reduction; eassumption. - Defined. - - (* tConst *) - Next Obligation. - rename Heq_anonymous into HH. - sq; constructor; try assumption. - symmetry in HH. - etransitivity. eassumption. reflexivity. - Defined. - - (* tInd *) - Next Obligation. - sq; econstructor; eassumption. - Defined. - - (* tConstruct *) - Next Obligation. - sq; econstructor; tea. now split. - Defined. - - (* tCase *) - Next Obligation. simpl; eauto using validity_wf. Qed. - Next Obligation. simpl; eauto using validity_wf. Qed. - Next Obligation. - destruct X, X9. sq. - change (eqb ind I = true) in H0. - destruct (eqb_spec ind I) as [e|e]; [destruct e|discriminate]. - change (eqb (ind_npars d) par = true) in H1. - destruct (eqb_spec (ind_npars d) par) as [e|e]; [|discriminate]. - rename Heq_anonymous into HH. symmetry in HH. - simpl in *. - eapply type_reduction in t0; eauto. eapply validity in t0; eauto. - rewrite <- e in HH. - eapply PCUICInductiveInversion.WfArity_build_case_predicate_type in HH; eauto. - destruct HH as [[s Hs] ?]. eexists; eauto. - eapply validity in t; eauto. - generalize (PCUICClosed.destArity_spec [] pty). - rewrite -Heq_anonymous0 /=. intros ->. - eapply PCUICInductives.isType_it_mkProd_or_LetIn_inv in t; eauto. - eapply isType_wf_universes in t. simpl in t. - now exact (PCUICWfUniverses.reflect_bP (wf_universe_reflect _ _) t). auto. - Qed. - - Next Obligation. - symmetry in Heq_anonymous2. - unfold iscumul in Heq_anonymous2. simpl in Heq_anonymous2. - apply isconv_term_sound in Heq_anonymous2. - red in Heq_anonymous2. - noconf Heq_I''. - noconf Heq_I'. noconf Heq_I. - noconf Heq_d. noconf Heq_d'. - simpl in *; sq. - change (eqb ind I = true) in H0. - destruct (eqb_spec ind I) as [e|e]; [destruct e|discriminate H0]. - change (eqb (ind_npars d) par = true) in H1. - destruct (eqb_spec (ind_npars d) par) as [e|e]; [|discriminate]; subst. - assert (wfΣ : wf_ext Σ) by (split; auto). - eapply type_reduction in X9; eauto. - have val:= validity_term wfΣ X9. - eapply type_Cumul' in X; [| |eassumption]. - 2:{ eapply PCUICInductiveInversion.WfArity_build_case_predicate_type; eauto. - eapply validity in X; eauto. - generalize (PCUICClosed.destArity_spec [] pty). - rewrite -Heq_anonymous0 /=. intros ->. - eapply PCUICInductives.isType_it_mkProd_or_LetIn_inv in X; eauto. - eapply isType_wf_universes in X. simpl in X. - now exact (PCUICWfUniverses.reflect_bP (wf_universe_reflect _ _) X). auto. } - have [pctx' da] : (∑ pctx', destArity [] pty' = Some (pctx', ps)). - { symmetry in Heq_anonymous1. - unshelve eapply (PCUICInductives.build_case_predicate_type_spec (Σ.1, ind_universes d)) in Heq_anonymous1 as [parsubst [_ ->]]. - eauto. eapply (PCUICWeakeningEnv.on_declared_inductive wfΣ) in HH as [? ?]. eauto. - eexists. rewrite !destArity_it_mkProd_or_LetIn; simpl. reflexivity. } - eapply PCUICInductiveInversion.build_branches_type_wt. 6:eapply X. all:eauto. - Defined. - - Next Obligation. - rename Heq_anonymous2 into XX2. - symmetry in XX2. simpl in *. eapply isconv_sound in XX2. - change (eqb ind ind' = true) in H0. - destruct (eqb_spec ind ind') as [e|e]; [destruct e|discriminate H0]. - change (eqb (ind_npars decl) par = true) in H1. - destruct (eqb_spec (ind_npars decl) par) as [e|e]; [|discriminate]; subst. - depelim HH. - sq. auto. now depelim X10. - Defined. - Next Obligation. - sq. now depelim HH. - Defined. - - (* The obligation tactic removes a useful let here. *) - Obligation Tactic := idtac. - Next Obligation. - intros. clearbody btyswf. idtac; Program.Tactics.program_simplify. - symmetry in Heq_anonymous2. - unfold iscumul in Heq_anonymous2. simpl in Heq_anonymous2. - apply isconv_term_sound in Heq_anonymous2. - noconf Heq_I''. noconf Heq_I'. noconf Heq_I. - noconf Heq_d. noconf Heq_d'. simpl in *. - assert (∥ All2 (fun x y => ((fst x = fst y) * - (Σ;;; Γ |- snd x : snd y) * isType Σ Γ y.2)%type) brs btys ∥). { - solve_all. simpl in *. - destruct btyswf. eapply All2_sq. solve_all. simpl in *; intuition (sq; auto). } - clear H3. sq. - change (eqb ind I = true) in H0. - destruct (eqb_spec ind I) as [e|e]; [destruct e|discriminate H0]. - change (eqb (ind_npars d) par = true) in H1. - destruct (eqb_spec (ind_npars d) par) as [e|e]; [|discriminate]; subst. - assert (wfΣ : wf_ext Σ) by (split; auto). - eapply type_reduction in X9; eauto. - eapply type_Cumul' in X; eauto. - 2:{ eapply PCUICInductiveInversion.WfArity_build_case_predicate_type; eauto. - now eapply validity in X9. - eapply validity in X; eauto. - generalize (PCUICClosed.destArity_spec [] pty). - rewrite -Heq_anonymous0 /=. intros ->. - eapply PCUICInductives.isType_it_mkProd_or_LetIn_inv in X; eauto. - eapply isType_wf_universes in X. simpl in X. - now exact (PCUICWfUniverses.reflect_bP (wf_universe_reflect _ _) X). auto. } - have [pctx' da] : (∑ pctx', destArity [] pty' = Some (pctx', ps)). - { symmetry in Heq_anonymous1. - unshelve eapply (PCUICInductives.build_case_predicate_type_spec (Σ.1, ind_universes d)) in Heq_anonymous1 as [parsubst [_ ->]]. - eauto. eapply (PCUICWeakeningEnv.on_declared_inductive wfΣ) in HH as [? ?]. eauto. - eexists. rewrite !destArity_it_mkProd_or_LetIn; simpl. reflexivity. } - eapply type_Case with (pty0:=pty'); tea. - - reflexivity. - - symmetry; eassumption. - - destruct isCoFinite; auto. - - symmetry; eauto. - Defined. - - Obligation Tactic := Program.Tactics.program_simplify ; eauto 2. - - (* tProj *) - Next Obligation. simpl; eauto using validity_wf. Qed. - Next Obligation. - simpl in *; sq; eapply type_Proj with (pdecl := (i, t0)). - - split. eassumption. split. symmetry; eassumption. cbn in *. - now apply beq_nat_true. - - cbn. destruct (ssrbool.elimT (eqb_spec ind I)); [assumption|]. - eapply type_reduction; eassumption. - - eapply type_reduction in X5; eauto. - eapply validity_term in X5; eauto. - destruct (ssrbool.elimT (eqb_spec ind I)); auto. - unshelve eapply (PCUICInductives.isType_mkApps_Ind _ X7 _) in X5 as [parsubst [argsubst [[sp sp'] cu]]]; eauto. - pose proof (PCUICContexts.context_subst_length2 (PCUICSpine.inst_ctx_subst sp)). - pose proof (PCUICContexts.context_subst_length2 (PCUICSpine.inst_ctx_subst sp')). - autorewrite with len in H, H2. - destruct (PCUICWeakeningEnv.on_declared_inductive HΣ X7) eqn:ond. - rewrite -o.(onNpars) -H. - forward (o0.(onProjections)). - intros H'; rewrite H' nth_error_nil // in Heq_anonymous. - destruct ind_cshapes as [|cs []]; auto. - intros onps. - unshelve epose proof (onps.(on_projs_noidx _ _ _ _ _ _)). - rewrite ond /= in H2. - change (ind_indices o0) with (ind_indices o0) in *. - destruct (ind_indices o0) => //. - simpl in H2. - rewrite List.skipn_length in H2. - rewrite List.firstn_length. lia. - Qed. - - (* tFix *) - Next Obligation. sq. constructor; auto. exists W; auto. Defined. - Next Obligation. sq. now eapply PCUICWeakening.All_mfix_wf in XX0. Defined. - Next Obligation. - sq. cbn in *. depelim XX. - destruct i as [s HH]. - exists s. - change (tSort s) with (lift0 #|fix_context mfix| (tSort s)). - apply weakening; try assumption. - now apply All_mfix_wf. - Defined. - Next Obligation. - clear -XX HΣ. sq. - now depelim XX. - Qed. - Next Obligation. - assert (∥ All (fun d => ((Σ;;; Γ ,,, fix_context mfix |- dbody d : (lift0 #|fix_context mfix|) (dtype d)) * (isLambda (dbody d) = true))%type) mfix ∥). { - eapply All_sq, All_impl. exact YY. - cbn; intros ? []. sq; now constructor. } - sq; econstructor; try eassumption. - symmetry; eassumption. - Qed. - - (* tCoFix *) - Next Obligation. sq. constructor; auto. exists W; auto. Defined. - Next Obligation. sq. now eapply PCUICWeakening.All_mfix_wf in XX. Defined. - Next Obligation. - sq. cbn in *. depelim XX'. - destruct i as [s HH]. - exists s. - change (tSort s) with (lift0 #|fix_context mfix| (tSort s)). - apply weakening; try assumption. - now apply All_mfix_wf. - Defined. - Next Obligation. - clear -XX' HΣ. sq. - now depelim XX'. - Qed. - Next Obligation. - assert (∥ All (fun d => ((Σ;;; Γ ,,, fix_context mfix |- dbody d : (lift0 #|fix_context mfix|) (dtype d)))%type) mfix ∥). { - eapply All_sq, All_impl. exact YY. - now cbn; intros ? []. } - sq; econstructor; try eassumption. - symmetry; eassumption. - Qed. - - Lemma sq_wfl_nil : ∥ wf_local Σ [] ∥. - Proof. - repeat constructor. - Qed. - - Program Fixpoint check_context Γ : typing_result (∥ wf_local Σ Γ ∥) - := match Γ with - | [] => ret sq_wfl_nil - | {| decl_body := None; decl_type := A |} :: Γ => - HΓ <- check_context Γ ;; - XX <- infer_type infer Γ HΓ A ;; - ret _ - | {| decl_body := Some t; decl_type := A |} :: Γ => - HΓ <- check_context Γ ;; - XX <- infer_type infer Γ HΓ A ;; - XX <- infer_cumul infer Γ HΓ t A _ ;; - ret _ - end. - Next Obligation. - sq. econstructor; tas. econstructor; eauto. - Qed. - Next Obligation. - sq. econstructor; tea. - Qed. - Next Obligation. - sq. econstructor; tas. econstructor; eauto. - Qed. - -(* - Program Definition check_isWfArity Γ (HΓ : ∥ wf_local Σ Γ ∥) A - : typing_result (∥ isWfArity Σ Γ A ∥) := - match destArity [] A with - | None => raise (Msg (print_term Σ Γ A ^ " is not an arity")) - | Some (ctx, s) => XX <- check_context (Γ ,,, ctx) ;; - ret _ - end. - Next Obligation. - destruct XX. constructor. exists ctx, s. - split; auto. - Defined. *) - - Program Definition check_isType Γ (HΓ : ∥ wf_local Σ Γ ∥) A - : typing_result (∥ isType Σ Γ A ∥) := - s <- infer Γ HΓ A ;; - s' <- reduce_to_sort Γ s.π1 _ ;; - ret _. - Next Obligation. now eapply validity_wf. Defined. - Next Obligation. destruct X0. sq. eexists. eapply type_reduction; tea. Defined. - - Program Definition check Γ (HΓ : ∥ wf_local Σ Γ ∥) t A - : typing_result (∥ Σ;;; Γ |- t : A ∥) := - check_isType Γ HΓ A ;; - infer_cumul infer Γ HΓ t A _. - -End Typecheck. - - -Definition infer' {cf:checker_flags} {Σ} (HΣ : ∥ wf_ext Σ ∥) - := infer (map_squash fst HΣ) (map_squash snd HΣ). - -Definition make_graph_and_infer {cf:checker_flags} {Σ} (HΣ : ∥ wf_ext Σ ∥) - := let '(G; HG) := wf_ext_is_graph HΣ in infer' HΣ G HG. - Section CheckEnv. - Context {cf:checker_flags}. - - Inductive env_error := - | IllFormedDecl (e : string) (e : type_error) - | AlreadyDeclared (id : string). - - Inductive EnvCheck (A : Type) := - | CorrectDecl (a : A) - | EnvError (Σ : global_env_ext) (e : env_error). - Global Arguments EnvError {A} Σ e. - Global Arguments CorrectDecl {A} a. - - Global Instance envcheck_monad : Monad EnvCheck := - {| ret A a := CorrectDecl a ; - bind A B m f := - match m with - | CorrectDecl a => f a - | EnvError g e => EnvError g e - end - |}. - - Global Instance envcheck_monad_exc - : MonadExc (global_env_ext * env_error) EnvCheck := - { raise A '(g, e) := EnvError g e; - catch A m f := - match m with - | CorrectDecl a => m - | EnvError g t => f (g, t) - end - }. - - Definition wrap_error {A} Σ (id : string) (check : typing_result A) : EnvCheck A := - match check with - | Checked a => CorrectDecl a - | TypeError e => EnvError Σ (IllFormedDecl id e) - end. + Context {cf:checker_flags}. Definition check_wf_type kn Σ HΣ HΣ' G HG t : EnvCheck (∑ u, ∥ Σ;;; [] |- t : tSort u ∥) @@ -1471,79 +75,16 @@ Section CheckEnv. : ∥ Alli (on_ind_body (lift_typing typing) Σ ind mdecl) 0 (ind_bodies mdecl) ∥ -> ∥ wf_local Σ (ind_params mdecl) ∥ -> context_assumptions (ind_params mdecl) = ind_npars mdecl -> - ind_guard mdecl -> check_variance (ind_universes mdecl) (ind_variance mdecl) -> ∥ on_inductive (lift_typing typing) Σ ind mdecl ∥. Proof. - intros H H0 H1 H2 H3. sq. econstructor; try eassumption. - unfold check_variance in H3. unfold on_variance. + intros H H0 H1 H2. sq. econstructor; try eassumption. + unfold check_variance in H2. unfold on_variance. destruct (ind_universes mdecl) eqn:E; destruct (ind_variance mdecl) eqn:E'; try congruence. - 2:split. now eapply eqb_eq in H3. + 2:split. now eapply eqb_eq in H2. Defined. - Program Fixpoint monad_Alli {T} {M : Monad T} {A} {P} (f : forall n x, T (∥ P n x ∥)) l n - : T (∥ @Alli A P n l ∥) - := match l with - | [] => ret (sq Alli_nil) - | a :: l => X <- f n a ;; - Y <- monad_Alli f l (S n) ;; - ret _ - end. - Next Obligation. - sq. constructor; assumption. - Defined. - - Program Fixpoint monad_Alli_All {T} {M : Monad T} {A} {P} {Q} (f : forall n x, ∥ Q x ∥ -> T (∥ P n x ∥)) l n : - ∥ All Q l ∥ -> T (∥ @Alli A P n l ∥) - := match l with - | [] => fun _ => ret (sq Alli_nil) - | a :: l => fun allq => X <- f n a _ ;; - Y <- monad_Alli_All f l (S n) _ ;; - ret _ - end. - Next Obligation. sq. - now depelim allq. - Qed. - Next Obligation. - sq; now depelim allq. - Qed. - Next Obligation. - sq. constructor; assumption. - Defined. - - Section monad_Alli_nth. - Context {T} {M : Monad T} {A} {P : nat -> A -> Type}. - Program Fixpoint monad_Alli_nth_gen l k - (f : forall n x, nth_error l n = Some x -> T (∥ P (k + n) x ∥)) : - T (∥ @Alli A P k l ∥) - := match l with - | [] => ret (sq Alli_nil) - | a :: l => X <- f 0 a _ ;; - Y <- monad_Alli_nth_gen l (S k) (fun n x hnth => px <- f (S n) x hnth;; ret _) ;; - ret _ - end. - Next Obligation. - sq. now rewrite Nat.add_succ_r in px. - Qed. - Next Obligation. - sq. rewrite Nat.add_0_r in X. constructor; auto. - Qed. - - Definition monad_Alli_nth l (f : forall n x, nth_error l n = Some x -> T (∥ P n x ∥)) : T (∥ @Alli A P 0 l ∥) := - monad_Alli_nth_gen l 0 f. - - End monad_Alli_nth. - - (* Definition Build_on_ind_body Σ mind mdecl i idecl ind_indices ind_sort *) - (* : ind_type idecl = *) - (* it_mkProd_or_LetIn (ind_params mdecl) *) - (* (it_mkProd_or_LetIn ind_indices (tSort ind_sort)) -> *) - (* ∥ on_type (lift_typing typing) Σ [] (ind_type idecl) ∥ -> *) - (* forall onConstructors : on_constructors P Σ mind mdecl i idecl (ind_ctors idecl), *) - (* (ind_projs idecl <> [] -> *) - (* on_projections P Σ mind mdecl i idecl ind_indices (ind_projs idecl)) -> *) - (* check_ind_sorts P onConstructors ind_sort -> on_ind_body P Σ mind mdecl i idecl *) (* We pack up all the information required on the global environment and graph in a single record. *) @@ -1640,7 +181,7 @@ Section CheckEnv. [|intro XX; rewrite XX in Huctx; discriminate Huctx]. intros ctrs Hctrs. rewrite Hctrs in Huctx. simpl in *. eapply (is_consistent_spec (global_ext_uctx (Σ, udecl))). - { apply wf_global_uctx_invariants in HΣ. + { sq. apply wf_global_uctx_invariants in HΣ. split. + clear -HΣ. cbn. apply LevelSet.union_spec; right. apply HΣ. @@ -1747,34 +288,6 @@ Section CheckEnv. Definition wfnil {Σ : global_env_ext} : ∥ wf_local Σ [] ∥ := sq localenv_nil. Notation " ' pat <- m ;; f " := (bind m (fun pat => f)) (pat pattern, right associativity, at level 100, m at next level). - - Fixpoint split_at_aux {A} (n : nat) (acc : list A) (l : list A) : list A * list A := - match n with - | 0 => (List.rev acc, l) - | S n' => - match l with - | [] => (List.rev acc, []) - | hd :: l' => split_at_aux n' (hd :: acc) l' - end - end. - - Lemma split_at_aux_firstn_skipn {A} n acc (l : list A) : split_at_aux n acc l = (List.rev acc ++ firstn n l, skipn n l). - Proof. - induction n in acc, l |- *; destruct l; simpl; auto. - now rewrite app_nil_r skipn_0. - now rewrite app_nil_r skipn_0. - now rewrite app_nil_r skipn_nil. - rewrite IHn. simpl. - now rewrite -app_assoc skipn_S /=. - Qed. - - Definition split_at {A} (n : nat) (l : list A) : list A * list A := - split_at_aux n [] l. - - Lemma split_at_firstn_skipn {A} n (l : list A) : split_at n l = (firstn n l, skipn n l). - Proof. - now rewrite /split_at split_at_aux_firstn_skipn /= //. - Qed. Lemma inversion_it_mkProd_or_LetIn Σ {wfΣ : wf Σ.1}: forall {Γ Δ s A}, @@ -1895,8 +408,6 @@ Section CheckEnv. wf_local Σ Γ * wt_decl Σ Γ d. Proof. intros wfd; depelim wfd; split; simpl; pcuic. - destruct l as [s Hs]. now exists (tSort s). - destruct l as [s Hs]. now exists (tSort s). now exists t. Qed. @@ -2024,9 +535,9 @@ Section CheckEnv. Next Obligation. destruct Σ as [Σ wfΣ G wfG]; simpl in *. sq. destruct le; simpl. - - eapply leqb_term_spec in check0; sq; auto. + - eapply leqb_term_spec in check; sq; auto. eapply wfΣ. - - eapply eqb_term_spec in check0; sq; auto. + - eapply eqb_term_spec in check; sq; auto. apply wfΣ. Qed. @@ -2203,7 +714,7 @@ Section CheckEnv. unfold arities_context. induction 1; simpl; auto. rewrite rev_map_cons /=. - eapply All_local_env_app_inv; split. constructor; pcuic. + eapply All_local_env_app; split. constructor; pcuic. eapply All_local_env_impl; eauto. move=> Γ t [] /=. * move=> ty ht. eapply weaken_ctx; eauto. @@ -2430,8 +941,8 @@ Section CheckEnv. eapply cumul_Prod_Prod_inv in c as [eqann [eqdom codom]]; auto. rewrite List.rev_app_distr. constructor. - eapply All_local_env_app in wf as [wfΓ wfr]. - eapply All_local_env_app in wfr as [wfd wfΓ0]. + eapply All_local_env_app_inv in wf as [wfΓ wfr]. + eapply All_local_env_app_inv in wfr as [wfd wfΓ0]. depelim wfd. destruct l as [? Hs]. red. eapply type_Cumul'; pcuic. eapply conv_cumul. now symmetry. rewrite subst_telescope_subst_context. eapply X. now len. @@ -2439,8 +950,8 @@ Section CheckEnv. eapply substitution_wf_local; eauto. 2:rewrite app_context_assoc in wf; eapply wf. repeat constructor. rewrite subst_empty. - eapply All_local_env_app in wf as [wfΓ wfr]. - eapply All_local_env_app in wfr as [wfd wfΓ0]. + eapply All_local_env_app_inv in wf as [wfΓ wfr]. + eapply All_local_env_app_inv in wfr as [wfd wfΓ0]. depelim wfd. destruct l as [? Hs]. eapply type_Cumul'; pcuic. eapply conv_cumul. now symmetry. eapply typing_spine_strengthen; eauto. @@ -2547,7 +1058,7 @@ Section CheckEnv. forward X. rewrite closedn_ctx_app clΓ0 /=. red. rewrite -clΔ''. lia_f_equal. len in X. rewrite app_context_assoc in sp. now specialize (X sp). - rewrite app_context_assoc in wf. now eapply All_local_env_app in wf as [? ?]. + rewrite app_context_assoc in wf. now eapply All_local_env_app_inv in wf as [? ?]. - len in sp. rewrite lift_context_app /= it_mkProd_or_LetIn_app lift_context_app it_mkProd_or_LetIn_app /= -it_mkProd_or_LetIn_app in sp. @@ -2576,7 +1087,7 @@ Section CheckEnv. forward X. rewrite closedn_ctx_app clΓ0 /=. red. rewrite -clΔ''. lia_f_equal. len in X. rewrite app_context_assoc in sp. now specialize (X sp). - rewrite app_context_assoc in wf. now eapply All_local_env_app in wf as [? ?]. + rewrite app_context_assoc in wf. now eapply All_local_env_app_inv in wf as [? ?]. Qed. Lemma typing_spine_it_mkProd_or_LetIn_ext_list_inv {Σ : global_env_ext} (wfΣ : wf Σ) Γ Δ Δ' Δ'' s args s' : @@ -2807,39 +1318,17 @@ Section CheckEnv. Qed. End PositivityCheck. - Section MonadAllAll. - Context {T : Type -> Type} {M : Monad T} {A} {P : A -> Type} {Q} (f : forall x, ∥ Q x ∥ -> T (∥ P x ∥)). - Program Fixpoint monad_All_All l : ∥ All Q l ∥ -> T (∥ All P l ∥) := - match l return ∥ All Q l ∥ -> T (∥ All P l ∥) with - | [] => fun _ => ret (sq All_nil) - | a :: l => fun allq => - X <- f a _ ;; - Y <- monad_All_All l _ ;; - ret _ - end. - Next Obligation. sq. - now depelim allq. - Qed. - Next Obligation. - sq; now depelim allq. - Qed. - Next Obligation. - sq. constructor; assumption. - Defined. - End MonadAllAll. - - (* Does not work due to implicit arguments I guess - Coercion isType_welltyped : isType >-> welltyped. *) - - Lemma forallb_unfold {A} (f : A -> bool) (g : nat -> A) n : - (forall x, x < n -> f (g x)) -> - forallb f (PCUICUnivSubstitution.unfold n g). + Lemma In_unfold_inj {A} (f : nat -> A) n i : + (forall i j, f i = f j -> i = j) -> + In (f i) (unfold n f) <-> i < n. Proof. - intros fg. - induction n; simpl; auto. - rewrite forallb_app IHn //. - intros x lt; rewrite fg //. lia. - rewrite /= fg //. + intros hf. split. + now apply In_unfold_inj. + intros. + induction n in i, H |- *; simpl => //. lia. + eapply in_app_iff. + destruct (eq_dec i n). subst. right; left; auto. + left. eapply IHn. lia. Qed. Lemma In_Var_global_ext_poly {n Σ inst cstrs} : @@ -2849,11 +1338,10 @@ Section CheckEnv. intros Hn. unfold global_ext_levels; simpl. apply LevelSet.mem_spec; rewrite LevelSet.union_spec. left. - rewrite /AUContext.levels /= PCUICUnivSubstitution.mapi_unfold. + rewrite /AUContext.levels /= mapi_unfold. apply LevelSetProp.of_list_1, InA_In_eq. - eapply In_unfold_var. exists n. intuition auto. + eapply In_unfold_inj; try congruence. Qed. - Lemma on_udecl_poly_bounded Σ inst cstrs : wf Σ -> @@ -3024,7 +1512,7 @@ Section CheckEnv. rewrite -{1}(subst_empty 0 b). repeat constructor. rewrite !subst_empty. eapply typing_wf_local in h. - rewrite app_context_assoc in h. eapply All_local_env_app in h as [h _]. + rewrite app_context_assoc in h. eapply All_local_env_app_inv in h as [h _]. now depelim h. - intros Γ t T h. rewrite !smash_context_app_ass !expand_lets_vass. @@ -3078,6 +1566,68 @@ Section CheckEnv. eapply All2_refl. intros. simpl. eapply (eq_decl_upto_refl Σ). Qed. + Lemma wt_cstrs {Σ : wf_env_ext} {n mdecl cstrs cs} : + ∥ All2 + (fun (cstr : (ident × term) × nat) (cs0 : constructor_shape) => + check_constructor_spec Σ n mdecl cstr cs0) cstrs cs ∥ -> + ∥ All (fun cstr => welltyped Σ (arities_context mdecl.(ind_bodies)) cstr.1.2) cstrs ∥. + Proof. + intros; sq. + solve_all. simpl. + destruct X as [[[isTy _] _] _]. simpl in isTy. + now eapply isType_welltyped in isTy. + Qed. + + Lemma get_wt_indices {Σ : wf_env_ext} {mdecl cstrs cs} + (wfar : ∥ wf_ind_types Σ mdecl ∥) + (wfpars : ∥ wf_local Σ (ind_params mdecl) ∥) + (n : nat) (idecl : one_inductive_body) (indices : context) + (hnth : nth_error mdecl.(ind_bodies) n = Some idecl) + (heq : ∥ ∑ inds, idecl.(ind_type) = it_mkProd_or_LetIn (mdecl.(ind_params) ,,, indices) (tSort inds) ∥) : + ∥ All2 + (fun (cstr : (ident × term) × nat) (cs0 : constructor_shape) => + check_constructor_spec Σ (S n) mdecl cstr cs0) cstrs cs ∥ -> + ∥ All (fun cs => wt_indices Σ mdecl indices cs) cs ∥. + Proof. + destruct Σ as [Σ wfΣ G wfG]; simpl in *. destruct wfΣ. + intros; sq. + solve_all. simpl. + destruct X as [[[isTy eq] sorts] eq']. simpl in *. + assert(wf_local Σ (ind_params mdecl,,, indices)). + { eapply nth_error_all in wfar; eauto. simpl in wfar. + destruct heq as [s Hs]. rewrite Hs in wfar. + eapply isType_it_mkProd_or_LetIn_wf_local in wfar. + now rewrite app_context_nil_l in wfar. auto. } + red. rewrite eq in isTy. + eapply isType_it_mkProd_or_LetIn_inv in isTy; auto. + eapply isType_mkApps_inv in isTy as [fty [s [Hf Hsp]]]; auto. + eapply inversion_Rel in Hf as [decl [wfctx [Hnth cum]]]; auto. + rewrite nth_error_app_ge in Hnth. lia. + split. now rewrite app_context_assoc in wfctx. + replace (#|ind_params mdecl,,, cshape_args y| + (#|ind_bodies mdecl| - S n) - + #|ind_params mdecl,,, cshape_args y|) with (#|ind_bodies mdecl| - S n) in Hnth by lia. + pose proof (nth_error_Some_length hnth). + rewrite nth_error_rev in hnth => //. + eapply nth_error_arities_context in hnth. rewrite Hnth in hnth. + noconf hnth; simpl in *. clear Hnth. + eapply PCUICSpine.typing_spine_strengthen in Hsp; eauto. + clear cum. + destruct heq as [inds eqind]. + move: Hsp. rewrite eqind. + rewrite lift_it_mkProd_or_LetIn /= => Hs. + rewrite closed_ctx_lift in Hs. eapply PCUICClosed.closed_wf_local; eauto. + rewrite app_context_assoc in Hs wfctx. + eapply typing_spine_it_mkProd_or_LetIn_ext_list_inv in Hs; auto. + 2:{ eapply PCUICWeakening.weaken_wf_local; pcuic. + now eapply wf_ind_types_wf_arities. } + 2:{ eapply PCUICClosed.closed_wf_local; eauto. } + eapply typing_spine_it_mkProd_or_LetIn_inv in Hs => //. auto. + eapply weakening_wf_local; eauto. + rewrite -app_context_assoc. + eapply PCUICWeakening.weaken_wf_local; eauto. + now eapply wf_ind_types_wf_arities. + Qed. + Program Definition check_cstr_variance (Σ : wf_env) mdecl (id : kername) indices (mdeclvar : check_variance mdecl.(ind_universes) mdecl.(ind_variance) = true) cs (wfΣ : ∥ wf_ext (Σ, ind_universes mdecl) ∥) @@ -3142,67 +1692,27 @@ Section CheckEnv. destruct ind_universes as [|[]] => //. Qed. - Lemma wt_cstrs {Σ : wf_env_ext} {n mdecl cstrs cs} : - ∥ All2 - (fun (cstr : (ident × term) × nat) (cs0 : constructor_shape) => - check_constructor_spec Σ n mdecl cstr cs0) cstrs cs ∥ -> - ∥ All (fun cstr => welltyped Σ (arities_context mdecl.(ind_bodies)) cstr.1.2) cstrs ∥. - Proof. - intros; sq. - solve_all. simpl. - destruct X as [[[isTy _] _] _]. simpl in isTy. - apply (isType_welltyped isTy). - Qed. - - Lemma get_wt_indices {Σ : wf_env_ext} {mdecl cstrs cs} - (wfar : ∥ wf_ind_types Σ mdecl ∥) - (wfpars : ∥ wf_local Σ (ind_params mdecl) ∥) - (n : nat) (idecl : one_inductive_body) (indices : context) - (hnth : nth_error mdecl.(ind_bodies) n = Some idecl) - (heq : ∥ ∑ inds, idecl.(ind_type) = it_mkProd_or_LetIn (mdecl.(ind_params) ,,, indices) (tSort inds) ∥) : - ∥ All2 - (fun (cstr : (ident × term) × nat) (cs0 : constructor_shape) => - check_constructor_spec Σ (S n) mdecl cstr cs0) cstrs cs ∥ -> - ∥ All (fun cs => wt_indices Σ mdecl indices cs) cs ∥. - Proof. - destruct Σ as [Σ wfΣ G wfG]; simpl in *. destruct wfΣ. - intros; sq. - solve_all. simpl. - destruct X as [[[isTy eq] sorts] eq']. simpl in *. - assert(wf_local Σ (ind_params mdecl,,, indices)). - { eapply nth_error_all in wfar; eauto. simpl in wfar. - destruct heq as [s Hs]. rewrite Hs in wfar. - eapply isType_it_mkProd_or_LetIn_wf_local in wfar. - now rewrite app_context_nil_l in wfar. auto. } - red. rewrite eq in isTy. - eapply isType_it_mkProd_or_LetIn_inv in isTy; auto. - eapply isType_mkApps_inv in isTy as [fty [s [Hf Hsp]]]; auto. - eapply inversion_Rel in Hf as [decl [wfctx [Hnth cum]]]; auto. - rewrite nth_error_app_ge in Hnth. lia. - split. now rewrite app_context_assoc in wfctx. - replace (#|ind_params mdecl,,, cshape_args y| + (#|ind_bodies mdecl| - S n) - - #|ind_params mdecl,,, cshape_args y|) with (#|ind_bodies mdecl| - S n) in Hnth by lia. - pose proof (nth_error_Some_length hnth). - rewrite nth_error_rev in hnth => //. - eapply nth_error_arities_context in hnth. rewrite Hnth in hnth. - noconf hnth; simpl in *. clear Hnth. - eapply PCUICSpine.typing_spine_strengthen in Hsp; eauto. - clear cum. - destruct heq as [inds eqind]. - move: Hsp. rewrite eqind. - rewrite lift_it_mkProd_or_LetIn /= => Hs. - rewrite closed_ctx_lift in Hs. eapply PCUICClosed.closed_wf_local; eauto. - rewrite app_context_assoc in Hs wfctx. - eapply typing_spine_it_mkProd_or_LetIn_ext_list_inv in Hs; auto. - 2:{ eapply PCUICWeakening.weaken_wf_local; pcuic. - now eapply wf_ind_types_wf_arities. } - 2:{ eapply PCUICClosed.closed_wf_local; eauto. } - eapply typing_spine_it_mkProd_or_LetIn_inv in Hs => //. auto. - eapply weakening_wf_local; eauto. - rewrite -app_context_assoc. - eapply PCUICWeakening.weaken_wf_local; eauto. - now eapply wf_ind_types_wf_arities. - Qed. + (** Moving it causes a universe bug... *) + Section MonadAllAll. + Context {T : Type -> Type} {M : Monad T} {A} {P : A -> Type} {Q : A -> Type} (f : forall x, ∥ Q x ∥ -> T (∥ P x ∥)). + Program Fixpoint monad_All_All l : ∥ All Q l ∥ -> T (∥ All P l ∥) := + match l return ∥ All Q l ∥ -> T (∥ All P l ∥) with + | [] => fun _ => ret (sq All_nil) + | a :: l => fun allq => + X <- f a _ ;; + Y <- monad_All_All l _ ;; + ret _ + end. + Next Obligation. sq. + now depelim allq. + Qed. + Next Obligation. + sq; now depelim allq. + Qed. + Next Obligation. + sq. constructor; assumption. + Defined. + End MonadAllAll. Program Definition check_constructors (Σ0 : wf_env) (Σ : wf_env_ext) (id : kername) (mdecl : mutual_inductive_body) (HΣ : Σ.(wf_env_ext_env) = (Σ0, ind_universes mdecl)) @@ -3225,7 +1735,6 @@ Section CheckEnv. Next Obligation. now sq. Qed. Next Obligation. apply wf_env_ext_wf. Qed. - Next Obligation. epose proof (get_wt_indices wfar wfpars _ _ _ hnth heq Hcs). destruct Σ as [Σ wfΣ G wfG]; simpl in *. clear X. @@ -3332,18 +1841,6 @@ Section CheckEnv. Definition checkb_constructors_smaller (G : universes_graph) (cs : list constructor_shape) (ind_sort : Universe.t) := List.forallb (fun cs => List.forallb (fun argsort => check_leqb_universe G argsort ind_sort) cs.(cshape_sorts)) cs. - Lemma forallbP_cond {A} (P Q : A -> Prop) (p : A -> bool) l : - Forall Q l -> - (forall x, Q x -> reflect (P x) (p x)) -> reflect (Forall P l) (forallb p l). - Proof. - intros HQ Hp. - apply: (iffP idP). - - induction HQ; rewrite /= //. move/andP => [pa pl]. - constructor; auto. now apply/(Hp _). - - induction HQ; intros HP; depelim HP; rewrite /= // IHHQ // andb_true_r. - now apply/(Hp _). - Qed. - Lemma check_constructors_smallerP (Σ : wf_env_ext) cs ind_sort : Forall (fun cs => Forall (wf_universe Σ) cs.(cshape_sorts)) cs -> wf_universe Σ ind_sort -> ∥ reflect (check_constructors_smaller Σ cs ind_sort) (checkb_constructors_smaller Σ cs ind_sort) ∥. @@ -3354,7 +1851,7 @@ Section CheckEnv. eapply forallbP_cond; eauto. clear wfcs. simpl; intros c wfc. pose proof (check_leqb_universe_spec' G (global_ext_uctx Σ)). - forward H. apply wf_ext_global_uctx_invariants; auto. now sq. + forward H. apply wf_ext_global_uctx_invariants; auto. forward H. apply wfΣ. eapply forallbP_cond; eauto. simpl. intros x wfx. specialize (H wfG x ind_sort). simpl. @@ -3395,11 +1892,14 @@ Section CheckEnv. (Msg ("Incorrect inductive sort: The constructor arguments universes are not smaller than the declared inductive sort")) ;; match indices_matter with | true => + (* FIXME this is wrong, we should use sorts_local_ctx as the indices might be + in sorts which don't necessarily have a sup. *) tyloc <- check_type_local_ctx Σ params indices ind_sort wfparams ;; ret _ | false => ret _ end end. + Next Obligation. unfold check_ind_sorts. simpl. now rewrite H. @@ -3616,15 +2116,14 @@ Section CheckEnv. let wfΣ : wf_env_ext := {| wf_env_ext_env := Σ; wf_env_ext_wf := _; wf_env_ext_graph := G; wf_env_ext_graph_wf := HG |} in let id := string_of_kername kn in - X5 <- wrap_error Σ id (check_eq_true (check_variance mdecl.(ind_universes) mdecl.(ind_variance)) (Msg "variance"));; - X2 <- wrap_error Σ id (check_context HΣ HΣ' G HG (ind_params mdecl)) ;; - X3 <- wrap_error Σ id (check_eq_nat (context_assumptions (ind_params mdecl)) + check_var <- wrap_error Σ id (check_eq_true (check_variance mdecl.(ind_universes) mdecl.(ind_variance)) (Msg "variance"));; + check_pars <- wrap_error Σ id (check_context HΣ HΣ' G HG (ind_params mdecl)) ;; + check_npars <- wrap_error Σ id (check_eq_nat (context_assumptions (ind_params mdecl)) (ind_npars mdecl) (Msg "wrong number of parameters")) ;; onarities <- check_ind_types wfΣ mdecl ;; - X1 <- monad_Alli_nth mdecl.(ind_bodies) (fun i oib Hoib => check_one_ind_body Σ0 wfΣ kn mdecl eq X2 onarities X5 i oib Hoib);; - X4 <- wrap_error Σ id (check_eq_true (ind_guard mdecl) (Msg "guard"));; - ret (Build_on_inductive_sq X1 X2 X3 X4 X5) + check_bodies <- monad_Alli_nth mdecl.(ind_bodies) (fun i oib Hoib => check_one_ind_body Σ0 wfΣ kn mdecl eq check_pars onarities check_var i oib Hoib);; + ret (Build_on_inductive_sq check_bodies check_pars check_npars check_var) end. Next Obligation. sq. unfold on_constant_decl; rewrite <- Heq_anonymous; tas. @@ -3643,20 +2142,6 @@ Section CheckEnv. Obligation Tactic := Program.Tactics.program_simpl. - Lemma levelset_union_empty s : LevelSet.union LevelSet.empty s = s. - Proof. - apply LevelSet.eq_leibniz. - change LevelSet.eq with LevelSet.Equal. - intros x; rewrite LevelSet.union_spec. lsets. - Qed. - - Lemma constraintset_union_empty s : ConstraintSet.union ConstraintSet.empty s = s. - Proof. - apply ConstraintSet.eq_leibniz. - change ConstraintSet.eq with ConstraintSet.Equal. - intros x; rewrite ConstraintSet.union_spec. lsets. - Qed. - Program Fixpoint check_wf_env (Σ : global_env) : EnvCheck (∑ G, (is_graph_of_uctx G (global_uctx Σ) /\ ∥ wf Σ ∥)) := match Σ with @@ -3744,7 +2229,7 @@ Section CheckEnv. destruct g. destruct c, cst_universes; try discriminate; reflexivity. destruct m, ind_universes; try discriminate; reflexivity. } rewrite eq1; clear eq1. - now rewrite levelset_union_empty constraintset_union_empty. + now rewrite LevelSet_union_empty CS_union_empty. Qed. Obligation Tactic := idtac. @@ -3813,25 +2298,6 @@ Section CheckEnv. simpl; intros. exact (check_wf_env_bool_spec Σ Γ wfΓ t T eq). Qed. - Lemma wf_consistent Σ : wf Σ -> consistent (global_constraints Σ). - Proof. - destruct Σ. - - exists {| valuation_mono := fun _ => 1%positive; valuation_poly := fun _ => 0 |}. - intros x Hx; now apply ConstraintSetFact.empty_iff in Hx. - - inversion 1; subst. subst udecl. clear -H2. - destruct H2 as [_ [_ [_ [v Hv]]]]. - exists v. intros ct Hc. apply Hv. simpl in *. - apply ConstraintSet.union_spec in Hc. destruct Hc. - apply ConstraintSet.union_spec; simpl. - + left. destruct d. - destruct c, cst_universes. assumption. - apply ConstraintSetFact.empty_iff in H; contradiction. - destruct m, ind_universes. assumption. - apply ConstraintSetFact.empty_iff in H; contradiction. - + apply ConstraintSet.union_spec; simpl. - now right. - Qed. - Obligation Tactic := Program.Tactics.program_simpl. Program Definition typecheck_program (p : program) φ Hφ diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 85972cff6..96e14c866 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -2,12 +2,12 @@ From Coq Require Import ProofIrrelevance. From MetaCoq.Template Require Import config utils uGraph. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils - PCUICReflect PCUICLiftSubst PCUICUnivSubst PCUICTyping + PCUICReflect PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICGlobalEnv PCUICCumulativity PCUICEquality PCUICConversion PCUICSafeLemmata PCUICNormal PCUICInversion PCUICReduction PCUICPosition PCUICPrincipality PCUICContextConversion PCUICSN PCUICUtils PCUICWeakening PCUICConvCumInversion. -From MetaCoq.SafeChecker Require Import PCUICSafeReduce. +From MetaCoq.SafeChecker Require Import PCUICEqualityDec PCUICErrors PCUICSafeReduce. Require Import Equations.Prop.DepElim. From Equations Require Import Equations. @@ -25,79 +25,6 @@ Module PSR := PCUICSafeReduce. *) - -Definition wf_global_uctx_invariants {cf:checker_flags} Σ : - ∥ wf Σ ∥ -> - global_uctx_invariants (global_uctx Σ). -Proof. - intros [HΣ]. split. - - cbn. eapply LevelSet.mem_spec, global_levels_Set. - - unfold global_uctx. - simpl. intros [[l ct] l'] Hctr. simpl in *. - induction Σ in HΣ, l, ct, l', Hctr |- *. - + apply ConstraintSetFact.empty_iff in Hctr; contradiction. - + simpl in *. apply ConstraintSet.union_spec in Hctr. - destruct Hctr as [Hctr|Hctr]. - * split. - -- inversion HΣ; subst. - destruct H2 as [HH1 [HH HH3]]. - subst udecl. destruct d as [decl|decl]; simpl in *. - ++ destruct decl; simpl in *. - destruct cst_universes ; [ - eapply (HH (l, ct, l') Hctr) - | apply ConstraintSetFact.empty_iff in Hctr ; contradiction - ]. - ++ destruct decl. simpl in *. - destruct ind_universes ; [ - eapply (HH (l, ct, l') Hctr) - | apply ConstraintSetFact.empty_iff in Hctr; contradiction - ]. - -- inversion HΣ. subst. - destruct H2 as [HH1 [HH HH3]]. - subst udecl. destruct d as [decl|decl]. - all: simpl in *. - ++ destruct decl. simpl in *. - destruct cst_universes ; [ - eapply (HH (l, ct, l') Hctr) - | apply ConstraintSetFact.empty_iff in Hctr; contradiction - ]. - ++ destruct decl. simpl in *. - destruct ind_universes; [ - eapply (HH (l, ct, l') Hctr) - | apply ConstraintSetFact.empty_iff in Hctr; contradiction - ]. - * inversion HΣ. subst. - split. - all: apply LevelSet.union_spec. - all: right. - all: unshelve eapply (IHΣ _ _ _ _ Hctr). - all: try eassumption. -Qed. - -Definition wf_ext_global_uctx_invariants {cf:checker_flags} Σ : - ∥ wf_ext Σ ∥ -> - global_uctx_invariants (global_ext_uctx Σ). -Proof. - intros [HΣ]. split. - - apply LevelSet.union_spec. right. apply LevelSet.mem_spec, global_levels_Set. - - destruct Σ as [Σ φ]. destruct HΣ as [HΣ Hφ]. - destruct (wf_global_uctx_invariants _ (sq HΣ)) as [_ XX]. - unfold global_ext_uctx, global_ext_levels, global_ext_constraints. - simpl. intros [[l ct] l'] Hctr. simpl in *. apply ConstraintSet.union_spec in Hctr. - destruct Hctr as [Hctr|Hctr]. - + destruct Hφ as [_ [HH _]]. apply (HH _ Hctr). - + specialize (XX _ Hctr). - split; apply LevelSet.union_spec; right; apply XX. -Qed. - -Definition global_ext_uctx_consistent {cf:checker_flags} Σ - : ∥ wf_ext Σ ∥ -> consistent (global_ext_uctx Σ).2. - intros [HΣ]. cbn. unfold global_ext_constraints. - unfold wf_ext, on_global_env_ext in HΣ. - destruct HΣ as [_ [_ [_ HH]]]. apply HH. -Qed. - - Section Conversion. Context {cf : checker_flags}. @@ -492,112 +419,10 @@ Section Conversion. simpl in *. eapply R_aux_stateR. all: simpl. all: auto. Qed. - - Definition conv_pb_relb pb := - match pb with - | Conv => check_eqb_universe G - | Cumul => check_leqb_universe G - end. - - Definition eqb_termp_napp pb napp := - eqb_term_upto_univ_napp Σ (check_eqb_universe G) (conv_pb_relb pb) napp. - - Lemma eqb_termp_napp_spec pb napp t u : - eqb_termp_napp pb napp t u -> - eq_termp_napp pb Σ napp t u. - Proof. - pose proof hΣ'. - apply eqb_term_upto_univ_impl. - - intros u1 u2. - eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). - + now eapply wf_ext_global_uctx_invariants. - + now eapply global_ext_uctx_consistent. - + assumption. - - intros u1 u2. - destruct pb. - + eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). - * now eapply wf_ext_global_uctx_invariants. - * now eapply global_ext_uctx_consistent. - * assumption. - + eapply (check_leqb_universe_spec' G (global_ext_uctx Σ)). - * now eapply wf_ext_global_uctx_invariants. - * now eapply global_ext_uctx_consistent. - * assumption. - Qed. - - Definition eqb_termp pb := (eqb_termp_napp pb 0). - Definition eqb_term := (eqb_termp Conv). - Definition leqb_term := (eqb_termp Cumul). - - Lemma leqb_term_spec t u : - leqb_term t u -> - leq_term Σ (global_ext_constraints Σ) t u. - Proof. - intros. - now apply eqb_termp_napp_spec in H. - Qed. - - Lemma eqb_term_spec t u : - eqb_term t u -> - eq_term Σ t u. - Proof. - intros. - now apply eqb_termp_napp_spec in H. - Qed. - - Lemma eqb_term_refl : - forall t, eqb_term t t. - Proof. - intro t. eapply eqb_term_upto_univ_refl. - all: apply check_eqb_universe_refl. - Qed. - - Definition eqb_binder_annot {A} (b b' : binder_annot A) : bool := - eqb b.(binder_relevance) b'.(binder_relevance). - - Lemma eq_binder_annot_reflect {A} na na' : reflect (eq_binder_annot (A:=A) na na') (eqb_binder_annot na na'). - Proof. - unfold eq_binder_annot, eqb_binder_annot. - destruct (eqb_spec na.(binder_relevance) na'.(binder_relevance)); constructor; auto. - Qed. - - Fixpoint eqb_ctx (Γ Δ : context) : bool := - match Γ, Δ with - | [], [] => true - | {| decl_name := na1 ; decl_body := None ; decl_type := t1 |} :: Γ, - {| decl_name := na2 ; decl_body := None ; decl_type := t2 |} :: Δ => - eqb_binder_annot na1 na2 && eqb_term t1 t2 && eqb_ctx Γ Δ - | {| decl_name := na1 ; decl_body := Some b1 ; decl_type := t1 |} :: Γ, - {| decl_name := na2 ; decl_body := Some b2 ; decl_type := t2 |} :: Δ => - eqb_binder_annot na1 na2 && eqb_term b1 b2 && eqb_term t1 t2 && eqb_ctx Γ Δ - | _, _ => false - end. - Lemma eqb_binder_annot_spec {A} na na' : eqb_binder_annot (A:=A) na na' -> eq_binder_annot (A:=A) na na'. - Proof. apply (PCUICWfUniverses.reflect_bP (eq_binder_annot_reflect _ _)). Qed. - - Lemma eqb_ctx_spec : - forall Γ Δ, - eqb_ctx Γ Δ -> - eq_context_upto Σ (eq_universe Σ) (eq_universe Σ) Γ Δ. - Proof. - intros Γ Δ h. - induction Γ as [| [na [b|] A] Γ ih ] in Δ, h |- *. - all: destruct Δ as [| [na' [b'|] A'] Δ]. - all: try discriminate. - - constructor. - - simpl in h. apply andP in h as [[[h1 h2]%andP h3]%andP h4]. - constructor. - + now apply eqb_binder_annot_spec in h1. - + eapply eqb_term_spec. assumption. - + eapply eqb_term_spec. assumption. - + eapply ih. assumption. - - simpl in h. apply andP in h as [[h1 h2]%andP h3]. - constructor. - + now apply eqb_binder_annot_spec. - + eapply eqb_term_spec. assumption. - + eapply ih. assumption. - Qed. + Notation eqb_ctx := (eqb_ctx Σ G). + Notation eqb_term := (eqb_term Σ G). + Notation leqb_term := (leqb_term Σ G). Definition eqb_term_stack t1 π1 t2 π2 := eqb_ctx (stack_context π1) (stack_context π2) && @@ -612,12 +437,12 @@ Section Conversion. eq_term Σ (zipp t1 π1) (zipp t2 π2). Proof. intros Γ t1 π1 t2 π2 h. - apply andP in h as [h1 h2]. + apply andb_and in h as [h1 h2]. split. - eapply eq_context_upto_cat. + eapply eq_context_upto_refl; tc. - + eapply eqb_ctx_spec. assumption. - - eapply eqb_term_spec. assumption. + + eapply eqb_ctx_spec; eassumption. + - eapply eqb_term_spec; eassumption. Qed. Definition leqb_term_stack t1 π1 t2 π2 := @@ -633,12 +458,12 @@ Section Conversion. leq_term Σ Σ (zipp t1 π1) (zipp t2 π2). Proof. intros Γ t1 π1 t2 π2 h. - apply andP in h as [h1 h2]. + apply andb_and in h as [h1 h2]. split. - eapply eq_context_upto_cat. + eapply eq_context_upto_refl; tc. - + eapply eqb_ctx_spec. assumption. - - eapply leqb_term_spec. assumption. + + eapply eqb_ctx_spec; eassumption. + - eapply leqb_term_spec; eassumption. Qed. Notation conv_stack_ctx Γ π1 π2 := @@ -652,94 +477,6 @@ Section Conversion. (∥ Σ ;;; Γ ,,, stack_context π |- zipp t π = zipp t' π' ∥) (only parsing). - Inductive ConversionError := - | NotFoundConstants (c1 c2 : kername) - - | NotFoundConstant (c : kername) - - | LambdaNotConvertibleTypes - (Γ1 : context) (na : aname) (A1 t1 : term) - (Γ2 : context) (na' : aname) (A2 t2 : term) - (e : ConversionError) - - | LambdaNotConvertibleAnn - (Γ1 : context) (na : aname) (A1 t1 : term) - (Γ2 : context) (na' : aname) (A2 t2 : term) - - | ProdNotConvertibleDomains - (Γ1 : context) (na : aname) (A1 B1 : term) - (Γ2 : context) (na' : aname) (A2 B2 : term) - (e : ConversionError) - - | ProdNotConvertibleAnn - (Γ1 : context) (na : aname) (A1 B1 : term) - (Γ2 : context) (na' : aname) (A2 B2 : term) - - | CaseOnDifferentInd - (Γ1 : context) - (ind : inductive) (par : nat) (p c : term) (brs : list (nat × term)) - (Γ2 : context) - (ind' : inductive) (par' : nat) (p' c' : term) (brs' : list (nat × term)) - - | CaseBranchNumMismatch - (ind : inductive) (par : nat) - (Γ : context) (p c : term) (brs1 : list (nat × term)) - (m : nat) (br : term) (brs2 : list (nat × term)) - (Γ' : context) (p' c' : term) (brs1' : list (nat × term)) - (m' : nat) (br' : term) (brs2' : list (nat × term)) - - | DistinctStuckProj - (Γ : context) (p : projection) (c : term) - (Γ' : context) (p' : projection) (c' : term) - - | CannotUnfoldFix - (Γ : context) (mfix : mfixpoint term) (idx : nat) - (Γ' : context) (mfix' : mfixpoint term) (idx' : nat) - - | FixRargMismatch (idx : nat) - (Γ : context) (u : def term) (mfix1 mfix2 : mfixpoint term) - (Γ' : context) (v : def term) (mfix1' mfix2' : mfixpoint term) - - | FixMfixMismatch (idx : nat) - (Γ : context) (mfix : mfixpoint term) - (Γ' : context) (mfix' : mfixpoint term) - - | DistinctCoFix - (Γ : context) (mfix : mfixpoint term) (idx : nat) - (Γ' : context) (mfix' : mfixpoint term) (idx' : nat) - - | CoFixRargMismatch (idx : nat) - (Γ : context) (u : def term) (mfix1 mfix2 : mfixpoint term) - (Γ' : context) (v : def term) (mfix1' mfix2' : mfixpoint term) - - | CoFixMfixMismatch (idx : nat) - (Γ : context) (mfix : mfixpoint term) - (Γ' : context) (mfix' : mfixpoint term) - - | StackHeadError - (leq : conv_pb) - (Γ1 : context) - (t1 : term) (args1 : list term) (u1 : term) (l1 : list term) - (Γ2 : context) - (t2 : term) (u2 : term) (l2 : list term) - (e : ConversionError) - - | StackTailError (leq : conv_pb) - (Γ1 : context) - (t1 : term) (args1 : list term) (u1 : term) (l1 : list term) - (Γ2 : context) - (t2 : term) (u2 : term) (l2 : list term) - (e : ConversionError) - - | StackMismatch - (Γ1 : context) (t1 : term) (args1 l1 : list term) - (Γ2 : context) (t2 : term) (l2 : list term) - - | HeadMismatch - (leq : conv_pb) - (Γ1 : context) (t1 : term) - (Γ2 : context) (t2 : term). - Inductive ConversionResult (P : Prop) := | Success (h : P) | Error (e : ConversionError) (h : ~P). @@ -1508,14 +1245,15 @@ Section Conversion. eapply Forall2_impl. 1: eassumption. intros. eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). all: auto. - + eapply wf_ext_global_uctx_invariants. - eapply hΣ'. - + eapply global_ext_uctx_consistent. - eapply hΣ'. + + destruct hΣ'. sq. eapply wf_ext_global_uctx_invariants. + eapply X. + + destruct hΣ'. sq; eapply global_ext_uctx_consistent; assumption. Qed. Arguments LevelSet.mem : simpl never. + Notation conv_pb_relb := (conv_pb_relb G). + Lemma conv_pb_relb_complete leq u u' : wf_universe Σ u -> wf_universe Σ u' -> @@ -1525,11 +1263,11 @@ Section Conversion. intros all1 all2 conv. destruct leq; cbn. - eapply check_eqb_universe_complete; eauto. - + apply wf_ext_global_uctx_invariants, hΣ'. - + apply global_ext_uctx_consistent, hΣ'. + + destruct hΣ' as [hΣ']; apply wf_ext_global_uctx_invariants, hΣ'. + + destruct hΣ' as [hΣ']; apply global_ext_uctx_consistent, hΣ'. - eapply check_leqb_universe_complete; eauto. - + apply wf_ext_global_uctx_invariants, hΣ'. - + apply global_ext_uctx_consistent, hΣ'. + + destruct hΣ' as [hΣ']; apply wf_ext_global_uctx_invariants, hΣ'. + + destruct hΣ' as [hΣ']; apply global_ext_uctx_consistent, hΣ'. Qed. Lemma get_level_make l : @@ -2107,7 +1845,7 @@ Section Conversion. change (true = eqb u.(rarg) v.(rarg)) in eq1. destruct (eqb_spec u.(rarg) v.(rarg)). 2: discriminate. symmetry in eqann. - apply (PCUICWfUniverses.reflect_bP (eq_binder_annot_reflect _ _)) in eqann. + apply (ssrbool.elimT (eq_binder_annot_reflect _ _)) in eqann. split; auto. destruct fk; simpl in *. all: intuition eauto. @@ -2141,7 +1879,7 @@ Section Conversion. change (true = eqb u.(rarg) v.(rarg)) in eq1. destruct (eqb_spec u.(rarg) v.(rarg)). 2: discriminate. symmetry in eqann. - apply (PCUICWfUniverses.reflect_bP (eq_binder_annot_reflect _ _)) in eqann. + apply (ssrbool.elimT (eq_binder_annot_reflect _ _)) in eqann. clear eq1. destruct fk. all: intuition eauto. @@ -2160,8 +1898,9 @@ Section Conversion. Next Obligation. destruct H as [H]; inversion H; destruct X as [_ [eq_uv eqann]]. change (?ru =? ?rv) with (eqb ru rv) in eq1. - pose proof (PCUICWfUniverses.reflect_Pb (eq_binder_annot_reflect (dname u) (dname v))). - rewrite <- neqann in H0. specialize (H0 eqann). discriminate. + symmetry in neqann. + pose proof (eq_binder_annot_reflect (dname u) (dname v)) as r. + now apply (ssrbool.elimF r neqann). Qed. Next Obligation. destruct H as [H]; inversion H; destruct X as [_ [eq_uv eqann]]. @@ -2521,7 +2260,7 @@ Section Conversion. Proof. intros eq ir. destruct ir as (_&[wh]). - apply eqb_term_spec in eq. + apply eqb_term_spec in eq; tea. epose proof (reduce_term_complete _ _ _ _ _ _) as [wh']. eapply whnf_eq_term in eq; [|exact wh']. rewrite zipp_as_mkApps in wh. @@ -2573,7 +2312,7 @@ Section Conversion. Proof. intros eq%eq_sym ir. destruct ir as (_&[wh]). - apply eqb_term_spec in eq. + apply eqb_term_spec in eq; tea. epose proof (reduce_term_complete _ _ _ _ _ _) as [wh']. eapply whnf_eq_term in eq; [|exact wh']. rewrite zipp_as_mkApps in wh. @@ -3009,9 +2748,7 @@ Section Conversion. simpl_stacks. destruct hΣ. apply Lambda_conv_cum_inv in H as (?&?&?); auto. - apply (PCUICWfUniverses.reflect_Pb (eq_binder_annot_reflect _ _)) in H. - unfold eq_binder_annot in H. unfold eqb_binder_annot in H; simpl in H. - congruence. + eapply (ssrbool.elimF (eq_binder_annot_reflect _ _)); tea. Qed. Next Obligation. (* Contrapositive of previous obligation *) @@ -3087,9 +2824,8 @@ Section Conversion. apply mkApps_Prod_nil in h2; auto. rewrite h1, h2 in H. apply Prod_conv_cum_inv in H as (?&?&?); auto. - apply (PCUICWfUniverses.reflect_Pb (eq_binder_annot_reflect _ _)) in H. - unfold eq_binder_annot in H. unfold eqb_binder_annot in H; simpl in H. - congruence. + eapply (ssrbool.elimF (eq_binder_annot_reflect _ _)); tea. + now unfold eqb_binder_annot. Qed. Next Obligation. (* Domains are not convertible *) @@ -4320,7 +4056,25 @@ Section Conversion. - reflexivity. - reflexivity. Qed. - + + (* TODO move to PCUICNormal *) + Lemma whnf_mkApps_tPrim_inv : + forall (f : RedFlags.t) (Σ : global_env) (Γ : context) p (args : list term), + whnf f Σ Γ (mkApps (tPrim p) args) -> args = []. + Proof. + intros * wh. + inversion wh; solve_discr. + clear -X. + remember (mkApps (tPrim p) args) eqn:teq. + exfalso. + revert teq. + induction X in args |- *; intros; solve_discr. + destruct args as [|? ? _] using MCList.rev_ind; [easy|]. + rewrite <- mkApps_nested in teq. + cbn in teq. noconf teq. + eauto. + Qed. + Lemma reducible_head_None Γ t π h : isApp t = false -> whnf RedFlags.nodelta Σ (Γ,,, stack_context π) (mkApps t (decompose_stack π).1) -> @@ -4379,6 +4133,9 @@ Section Conversion. - constructor; eexists _, (decompose_stack π).1. split; [constructor; eauto with pcuic|]. eauto with pcuic. + - apply whnf_mkApps_tPrim_inv in wh as ->. + constructor; eexists _, []. + eauto using whnf_red with pcuic. - constructor; eexists _, (decompose_stack π).1. split; [econstructor|]; eauto. split; [eauto with pcuic|]. @@ -4440,7 +4197,7 @@ Section Conversion. isconv_prog leq t1 π1 rt2' (θ2' +++ θ2) aux } } ; - | @exist None nored2 with inspect (eqb_termp_napp leq #|(decompose_stack π1).1| t1 t2) := { + | @exist None nored2 with inspect (eqb_termp_napp Σ G leq #|(decompose_stack π1).1| t1 t2) := { | @exist true eq1 := isconv_args leq t1 π1 t2 π2 aux; | @exist false noteq := no ( @@ -4651,7 +4408,7 @@ Section Conversion. rewrite Nat.add_0_r. apply All2_length in a. rewrite a in eq1. - apply eqb_termp_napp_spec; eauto. + eapply eqb_termp_napp_spec; eauto. Qed. Next Obligation. apply h; clear h. @@ -4762,6 +4519,7 @@ Section Conversion. apply inversion_Sort in h2 as (_&h2&_); auto. apply inversion_Sort in h1 as (_&h1&_); auto. eapply conv_pb_relb_complete in H0; eauto. + - now rewrite eq_dec_to_bool_refl in noteq. Qed. Equations _isconv (s : state) (Γ : context) diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index 3a90d1f79..5f72d038f 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -6,6 +6,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICUnivSubst PCUICTyping PCUICPosition PCUICNormal PCUICInversion PCUICSafeLemmata PCUICSR PCUICSN PCUICUtils PCUICReduction PCUICValidity. +From MetaCoq.SafeChecker Require Import PCUICWfReduction. Require Import Equations.Prop.DepElim. From Equations Require Import Equations. @@ -1255,6 +1256,7 @@ Section Reduce. unfold is_true in typ. unfold PCUICAst.fst_ctx in *. congruence. + - now eapply inversion_Prim in typ. Qed. Definition isCoFix_app t := @@ -1286,6 +1288,7 @@ Section Reduce. - exfalso; eapply invert_fix_ind; eauto. - unfold isCoFix_app in cof. now rewrite decompose_app_mkApps in cof. + - now eapply inversion_Prim in typ. Qed. Lemma whnf_fix_arg_whne mfix idx body Γ t before args aftr ty : @@ -1434,6 +1437,9 @@ Section Reduce. apply inversion_App in h as (?&?&?&?&?); auto. apply inversion_Prod in t0 as (?&?&?&?&?); auto. eapply PCUICConversion.cumul_Sort_Prod_inv; eauto. + + exfalso. + eapply welltyped_context in h as [s Hs]; tas. + now eapply inversion_Prim in Hs. - unfold zipp. case_eq (decompose_stack π0). intros l ρ e. constructor. constructor. eapply whne_mkApps. eapply whne_rel_nozeta. assumption. @@ -1721,3 +1727,301 @@ Section Reduce. Qed. End Reduce. + +From MetaCoq.PCUIC Require Import PCUICConfluence PCUICConversion. +From MetaCoq.SafeChecker Require Import PCUICErrors. + +Section ReduceFns. + Context {cf : checker_flags} {Σ : global_env_ext} (HΣ : ∥ wf Σ ∥). + + (* We get stack overflow on Qed after Equations definitions when this is transparent *) + Opaque reduce_stack_full. + + Definition hnf := reduce_term RedFlags.default Σ HΣ. + + Theorem hnf_sound {Γ t h} : ∥ red (fst Σ) Γ t (hnf Γ t h) ∥. + Proof. + apply reduce_term_sound. + Defined. + + Theorem hnf_complete {Γ t h} : ∥whnf RedFlags.default Σ Γ (hnf Γ t h)∥. + Proof. + apply reduce_term_complete. + Qed. + + Equations? reduce_to_sort (Γ : context) (t : term) (h : welltyped Σ Γ t) + : typing_result (∑ u, ∥ red (fst Σ) Γ t (tSort u) ∥) := + reduce_to_sort Γ t h with view_sortc t := { + | view_sort_sort s => ret (s; sq (refl_red _ _ _)); + | view_sort_other t _ with inspect (hnf Γ t h) := + | exist hnft eq with view_sortc hnft := { + | view_sort_sort s => ret (s; _); + | view_sort_other t _ => raise (NotASort t) + } + }. + Proof. + pose proof (hnf_sound (h:=h)). + now rewrite eq. + Qed. + + Lemma reduce_to_sort_complete {Γ t wt} e : + reduce_to_sort Γ t wt = TypeError e -> + (forall s, red Σ Γ t (tSort s) -> False). + Proof. + funelim (reduce_to_sort Γ t wt); try congruence. + intros _ s r. + pose proof (@hnf_complete Γ t0 h) as [wh]. + pose proof (@hnf_sound Γ t0 h) as [r']. + destruct HΣ. + eapply red_confluence in r as (?&r1&r2); eauto. + apply invert_red_sort in r2 as ->. + eapply whnf_red_inv in r1; eauto. + depelim r1. + clear Heq. + rewrite H in n0. + now cbn in n0. + Qed. + + Equations? reduce_to_prod (Γ : context) (t : term) (h : welltyped Σ Γ t) + : typing_result (∑ na a b, ∥ red (fst Σ) Γ t (tProd na a b) ∥) := + reduce_to_prod Γ t h with view_prodc t := { + | view_prod_prod na a b => ret (na; a; b; sq (refl_red _ _ _)); + | view_prod_other t _ with inspect (hnf Γ t h) := + | exist hnft eq with view_prodc hnft := { + | view_prod_prod na a b => ret (na; a; b; _); + | view_prod_other t' _ => raise (NotAProduct t t') + } + }. + Proof. + pose proof (hnf_sound (h:=h)). + now rewrite eq. + Qed. + + Lemma reduce_to_prod_complete {Γ t wt} e : + reduce_to_prod Γ t wt = TypeError e -> + (forall na a b, red Σ Γ t (tProd na a b) -> False). + Proof. + funelim (reduce_to_prod Γ t wt); try congruence. + intros _ na a b r. + pose proof (@hnf_complete Γ t0 h) as [wh]. + pose proof (@hnf_sound Γ t0 h) as [r']. + destruct HΣ. + eapply red_confluence in r as (?&r1&r2); eauto. + apply invert_red_prod in r2 as (?&?&(->&?)&?); auto. + eapply whnf_red_inv in r1; auto. + depelim r1. + clear Heq. + rewrite H in n0. + now cbn in n0. + Qed. + + Equations? reduce_to_ind (Γ : context) (t : term) (h : welltyped Σ Γ t) + : typing_result (∑ i u l, ∥ red (fst Σ) Γ t (mkApps (tInd i u) l) ∥) := + reduce_to_ind Γ t h with inspect (decompose_app t) := { + | exist (thd, args) eq_decomp with view_indc thd := { + | view_ind_tInd i u => ret (i; u; args; sq _); + | view_ind_other t _ with inspect (reduce_stack RedFlags.default Σ HΣ Γ t Empty h) := { + | exist (hnft, π) eq with view_indc hnft := { + | view_ind_tInd i u with inspect (decompose_stack π) := { + | exist (l, _) eq_decomp => ret (i; u; l; _) + }; + | view_ind_other _ _ => raise (NotAnInductive t) + } + } + } + }. + Proof. + - assert (X : mkApps (tInd i u) args = t); [|rewrite X; apply refl_red]. + etransitivity. 2: symmetry; eapply mkApps_decompose_app. + now rewrite <- eq_decomp. + - pose proof (reduce_stack_sound RedFlags.default Σ HΣ _ _ Empty h). + rewrite <- eq in H. + cbn in *. + assert (π = appstack l ε) as ->. + 2: { now rewrite zipc_appstack in H. } + unfold reduce_stack in eq. + destruct reduce_stack_full as (?&_&stack_val&?). + subst x. + unfold Pr in stack_val. + cbn in *. + assert (decomp: decompose_stack π = ((decompose_stack π).1, ε)). + { rewrite stack_val. + now destruct decompose_stack. } + apply decompose_stack_eq in decomp as ->. + now rewrite <- eq_decomp0. + Qed. + + Lemma reduce_to_ind_complete Γ ty wat e : + reduce_to_ind Γ ty wat = TypeError e -> + forall ind u args, + red Σ Γ ty (mkApps (tInd ind u) args) -> + False. + Proof. + funelim (reduce_to_ind Γ ty wat); try congruence. + intros _ ind u args r. + pose proof (reduce_stack_whnf RedFlags.default Σ HΣ Γ t ε h) as wh. + unfold reduce_stack in *. + destruct reduce_stack_full as ((hd&π)&r'&stack_valid&(notapp&_)). + destruct wh as [wh]. + apply Req_red in r' as [r']. + unfold Pr in stack_valid. + cbn in *. + destruct HΣ. + eapply red_confluence in r as (?&r1&r2); [|eassumption|exact r']. + assert (exists args, π = appstack args ε) as (?&->). + { exists ((decompose_stack π).1). + assert (decomp: decompose_stack π = ((decompose_stack π).1, ε)). + { now rewrite stack_valid; destruct decompose_stack. } + now apply decompose_stack_eq in decomp. } + + unfold zipp in wh. + rewrite stack_context_appstack, decompose_stack_appstack in wh. + rewrite zipc_appstack in r1. + cbn in *. + rewrite app_nil_r in wh. + apply red_mkApps_tInd in r2 as (?&->&?); auto. + eapply whnf_red_inv in r1; eauto. + apply whnf_red_mkApps_inv in r1 as (?&?); auto. + depelim w. + noconf e0. + discriminate i0. + Qed. + + (* Definition of assumption-only arities (without lets) *) + Definition arity_ass := aname * term. + + Fixpoint mkAssumArity (l : list arity_ass) (s : Universe.t) : term := + match l with + | [] => tSort s + | (na, A) :: l => tProd na A (mkAssumArity l s) + end. + + Definition arity_ass_context := rev_map (fun '(na, A) => vass na A). + + Lemma assumption_context_arity_ass_context l : + assumption_context (arity_ass_context l). + Proof. + unfold arity_ass_context. + rewrite rev_map_spec. + induction l using MCList.rev_ind; cbn. + - constructor. + - rewrite map_app, rev_app_distr. + cbn. + destruct x. + constructor; auto. + Qed. + + Lemma mkAssumArity_it_mkProd_or_LetIn (l : list arity_ass) (s : Universe.t) : + mkAssumArity l s = it_mkProd_or_LetIn (arity_ass_context l) (tSort s). + Proof. + unfold arity_ass_context. + rewrite rev_map_spec. + induction l; auto. + cbn. + destruct a. + rewrite IHl, it_mkProd_or_LetIn_app; auto. + Qed. + + Lemma isArity_mkAssumArity l s : + isArity (mkAssumArity l s). + Proof. + induction l as [|(na & A) l IH]; cbn; auto. + Qed. + + Record conv_arity {Γ T} : Type := + build_conv_arity { + conv_ar_context : list arity_ass; + conv_ar_univ : Universe.t; + conv_ar_red : ∥red Σ Γ T (mkAssumArity conv_ar_context conv_ar_univ)∥ + }. + + Global Arguments conv_arity : clear implicits. + + Lemma conv_arity_Is_conv_to_Arity {Γ T} : + conv_arity Γ T -> + Is_conv_to_Arity Σ Γ T. + Proof. + intros [asses univ r]. + eexists. + split; [eassumption|]. + apply isArity_mkAssumArity. + Qed. + + Lemma isArity_red Γ u v : + isArity u -> + red Σ Γ u v -> + isArity v. + Proof. + intros arity_u r. + induction r using red_rect_n1; [easy|]. + eapply isArity_red1; eassumption. + Qed. + + Lemma Is_conv_to_Arity_red Γ T T' : + Is_conv_to_Arity Σ Γ T -> + red Σ Γ T T' -> + Is_conv_to_Arity Σ Γ T'. + Proof. + unfold Is_conv_to_Arity. + intros [T'' (redT'' & is_arity)] red. + sq. + destruct (red_confluence HΣ red redT'') as (a & reda' & reda''). + exists a. + split; [easy|]. + clear -is_arity reda''. + eapply isArity_red; eauto. + Qed. + + Local Instance wellfounded : WellFounded (@hnf_subterm_rel _ Σ) := + @wf_hnf_subterm _ _ HΣ. + + Equations? (noeqns) reduce_to_arity (Γ : context) (T : term) (wt : welltyped Σ Γ T) + : (conv_arity Γ T) + {~Is_conv_to_Arity Σ Γ T} + by wf ((Γ;T; wt) : (∑ Γ t, welltyped Σ Γ t)) hnf_subterm_rel := + reduce_to_arity Γ T wt with inspect (hnf Γ T wt) := + | exist Thnf eqhnf with view_prod_sortc Thnf := { + | view_prod_sort_prod na A B with reduce_to_arity (Γ,, vass na A) B _ := { + | inleft car => inleft {| conv_ar_context := (na, A) :: conv_ar_context car; + conv_ar_univ := conv_ar_univ car |}; + | inright nocar => inright _ + }; + | view_prod_sort_sort u => inleft {| conv_ar_context := []; + conv_ar_univ := u |}; + | view_prod_sort_other nonar notprod notsort => inright _ + }. + Proof. + all: pose proof (@hnf_sound Γ T wt) as [r]. + all: rewrite <- ?eqhnf in r. + all: destruct HΣ as [wf]. + - destruct wt as (?&typ). + eapply subject_reduction in r; eauto. + apply inversion_Prod in r as (?&?&?&?&?); auto. + econstructor; eauto. + - constructor. + eexists _; split; eauto. + unshelve eexists _; [constructor; constructor|]; auto. + - destruct car as [c_ar c_univ [c_red]]; cbn. + constructor. + etransitivity; eauto. + eapply red_prod; eauto. + - eapply Is_conv_to_Arity_red in H as (?&[r']&isar); eauto. + apply invert_red_prod in r' as (?&?&(->&?)&?); auto. + contradiction nocar. + eexists; eauto using sq. + - constructor; auto. + - pose proof (@hnf_complete Γ T wt) as [w]. + destruct HΣ. + apply Is_conv_to_Arity_inv in H as [(na&A&B&[r'])|(u&[r'])]; auto. + + eapply red_confluence in r' as (?&r1&r2); eauto. + apply invert_red_prod in r2 as (?&?&(->&?)&?); auto. + eapply whnf_red_inv in r1; eauto. + depelim r1. + rewrite H in notprod; auto. + + eapply red_confluence in r' as (?&r1&r2); eauto. + apply invert_red_sort in r2 as ->. + eapply whnf_red_inv in r1; eauto. + depelim r1. + rewrite H in notsort; cbn in *; auto. + Qed. + +End ReduceFns. diff --git a/safechecker/theories/PCUICSafeRetyping.v b/safechecker/theories/PCUICSafeRetyping.v index 5a32c8965..800dbc0e4 100644 --- a/safechecker/theories/PCUICSafeRetyping.v +++ b/safechecker/theories/PCUICSafeRetyping.v @@ -10,7 +10,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICArities PCUICInduc PCUICGeneration PCUICInversion PCUICValidity PCUICInductives PCUICInductiveInversion PCUICSpine PCUICSR PCUICCumulativity PCUICConversion PCUICConfluence PCUICArities PCUICWeakeningEnv PCUICContexts. -From MetaCoq.SafeChecker Require Import PCUICSafeReduce PCUICSafeChecker. +From MetaCoq.SafeChecker Require Import PCUICErrors PCUICSafeReduce. Local Open Scope string_scope. Set Asymmetric Patterns. Import monad_utils.MonadNotation. @@ -261,7 +261,9 @@ Section TypeOf. infer Γ (tCoFix mfix n) wt with inspect (nth_error mfix n) := { | ret (Some f) => ret f.(dtype); - | ret None => ! } + | ret None => ! }; + + infer Γ (tPrim p) wt := ! }. Proof. all:try clear infer. @@ -342,7 +344,7 @@ Section TypeOf. sq. split. * simpl. eapply type_reduction in Htty; eauto. - eapply type_App; eauto. + eapply type_App'; eauto. specialize (pbty _ t0). assert (Σ ;;; Γ |- tProd na' A' B' <= tProd x x0 x1). eapply cumul_red_l_inv; eauto. @@ -655,6 +657,8 @@ Section TypeOf. congruence. - now eapply inversion_CoFix in HT as [decl [fg [hnth [htys [hbods [wf cum]]]]]]; auto. + + - now eapply inversion_Prim in HT. Defined. Definition type_of Γ t wt : term := (infer Γ t wt). diff --git a/safechecker/theories/PCUICTypeChecker.v b/safechecker/theories/PCUICTypeChecker.v new file mode 100644 index 000000000..2ce20f41c --- /dev/null +++ b/safechecker/theories/PCUICTypeChecker.v @@ -0,0 +1,764 @@ +(* Distributed under the terms of the MIT license. *) +From MetaCoq.Template Require Import config utils uGraph. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils + PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICNormal PCUICSR + PCUICGeneration PCUICReflect PCUICEquality PCUICInversion PCUICValidity + PCUICWeakening PCUICPosition PCUICCumulativity PCUICSafeLemmata PCUICSN + PCUICPretty PCUICArities PCUICConfluence PCUICSize + PCUICContextConversion PCUICConversion PCUICWfUniverses + PCUICGlobalEnv + (* Used for support lemmas *) + PCUICInductives PCUICWfUniverses + PCUICContexts PCUICSubstitution PCUICSpine PCUICInductiveInversion + PCUICClosed PCUICUnivSubstitution PCUICWeakeningEnv. + +From MetaCoq.SafeChecker Require Import PCUICSafeReduce PCUICErrors PCUICEqualityDec + PCUICSafeConversion PCUICWfReduction PCUICWfEnv. + +From Equations Require Import Equations. +Require Import ssreflect ssrbool. + +Local Set Keyed Unification. +Set Equations Transparent. + +(** It otherwise tries [auto with *], very bad idea. *) +Ltac Coq.Program.Tactics.program_solve_wf ::= + match goal with + | |- @Wf.well_founded _ _ => auto with subterm wf + | |- ?T => match type of T with + | Prop => auto + end + end. + +Section Typecheck. + Context {cf : checker_flags} {Σ : global_env_ext} (HΣ : ∥ wf Σ ∥) + (Hφ : ∥ on_udecl Σ.1 Σ.2 ∥) + (G : universes_graph) (HG : is_graph_of_uctx G (global_ext_uctx Σ)). + + (* We get stack overflow on Qed after Equations definitions when this is transparent *) + Opaque reduce_stack_full. + + Local Definition HΣ' : ∥ wf_ext Σ ∥. + Proof. + destruct HΣ, Hφ; now constructor. + Defined. + + Notation hnf := (hnf HΣ). + + Definition isconv Γ := isconv_term Σ HΣ Hφ G HG Γ Conv. + Definition iscumul Γ := isconv_term Σ HΣ Hφ G HG Γ Cumul. + + Program Definition convert Γ t u + (ht : welltyped Σ Γ t) (hu : welltyped Σ Γ u) + : typing_result (∥ Σ ;;; Γ |- t = u ∥) := + match eqb_term Σ G t u with true => ret _ | false => + match isconv Γ t ht u hu with + | ConvSuccess => ret _ + | ConvError e => (* fallback *) (* nico *) + let t' := hnf Γ t ht in + let u' := hnf Γ u hu in + (* match leq_term (snd Σ) t' u' with true => ret _ | false => *) + raise (NotCumulSmaller G Γ t u t' u' e) + (* end *) + end end. + Next Obligation. + symmetry in Heq_anonymous; eapply eqb_term_spec in Heq_anonymous; tea. + constructor. now constructor. + Qed. + Next Obligation. + symmetry in Heq_anonymous; apply isconv_term_sound in Heq_anonymous. + assumption. + Qed. + + Program Definition convert_leq Γ t u + (ht : welltyped Σ Γ t) (hu : welltyped Σ Γ u) + : typing_result (∥ Σ ;;; Γ |- t <= u ∥) := + match leqb_term Σ G t u with true => ret _ | false => + match iscumul Γ t ht u hu with + | ConvSuccess => ret _ + | ConvError e => (* fallback *) (* nico *) + let t' := hnf Γ t ht in + let u' := hnf Γ u hu in + (* match leq_term (snd Σ) t' u' with true => ret _ | false => *) + raise (NotCumulSmaller G Γ t u t' u' e) + (* end *) + end end. + Next Obligation. + symmetry in Heq_anonymous; eapply leqb_term_spec in Heq_anonymous; tea. + constructor. now constructor. + Qed. + Next Obligation. + symmetry in Heq_anonymous; apply isconv_term_sound in Heq_anonymous. + assumption. + Qed. + + Section InferAux. + Variable (infer : forall Γ (HΓ : ∥ wf_local Σ Γ ∥) (t : term), + typing_result ({ A : term & ∥ Σ ;;; Γ |- t : A ∥ })). + + Program Definition infer_type Γ HΓ t + : typing_result ({u : Universe.t & ∥ Σ ;;; Γ |- t : tSort u ∥}) := + tx <- infer Γ HΓ t ;; + u <- reduce_to_sort HΣ Γ tx.π1 _ ;; + ret (u.π1; _). + Next Obligation. + eapply validity_wf; eassumption. + Defined. + Next Obligation. + destruct HΣ, HΓ, X, X0. + now constructor; eapply type_reduction. + Defined. + + Program Definition infer_cumul Γ HΓ t A (hA : ∥ isType Σ Γ A ∥) + : typing_result (∥ Σ ;;; Γ |- t : A ∥) := + A' <- infer Γ HΓ t ;; + X <- convert_leq Γ A'.π1 A _ _ ;; + ret _. + Next Obligation. now eapply validity_wf. Qed. + Next Obligation. destruct hA; now apply wat_welltyped. Qed. + Next Obligation. + destruct HΣ, HΓ, hA, X, X0. constructor. eapply type_Cumul'; eassumption. + Qed. + + Program Definition infer_scheme Γ HΓ t : + typing_result (∑ ctx u, ∥ Σ ;;; Γ |- t : mkAssumArity ctx u ∥) := + '(T; p) <- infer Γ HΓ t;; + match reduce_to_arity HΣ Γ T _ with + | inleft car => ret (conv_ar_context car; conv_ar_univ car; _) + | inright _ => TypeError (NotAnArity T) + end. + Next Obligation. + intros; subst. + cbn in *. + eapply validity_wf; eauto. + Qed. + Next Obligation. + destruct car as [? ? [r]]. + cbn. + clear Heq_anonymous. + sq. + eapply type_reduction; eauto. + Qed. + End InferAux. + + Program Definition lookup_ind_decl ind + : typing_result + ({decl & {body & declared_inductive (fst Σ) decl ind body}}) := + match lookup_env (fst Σ) ind.(inductive_mind) with + | Some (InductiveDecl decl) => + match nth_error decl.(ind_bodies) ind.(inductive_ind) with + | Some body => ret (decl; body; _) + | None => raise (UndeclaredInductive ind) + end + | _ => raise (UndeclaredInductive ind) + end. + Next Obligation. + split. + - symmetry in Heq_anonymous0. + etransitivity. eassumption. reflexivity. + - now symmetry. + Defined. + + Program Definition check_consistent_instance uctx u + : typing_result (consistent_instance_ext Σ uctx u) + := match uctx with + | Monomorphic_ctx _ => + check_eq_nat #|u| 0 (Msg "monomorphic instance should be of length 0") + | Polymorphic_ctx (inst, cstrs) => + let '(inst, cstrs) := AUContext.repr (inst, cstrs) in + check_eq_true (forallb (fun l => LevelSet.mem l (uGraph.wGraph.V G)) u) + (Msg "undeclared level in instance") ;; + X <- check_eq_nat #|u| #|inst| + (Msg "instance does not have the right length");; + match check_constraints G (subst_instance_cstrs u cstrs) with + | true => ret (conj _ _) + | false => raise (Msg "ctrs not satisfiable") + end + (* #|u| = #|inst| /\ valid_constraints φ (subst_instance_cstrs u cstrs) *) + end. + Next Obligation. + eapply forallb_All in H. eapply All_forallb'; tea. + clear -cf HG. intros x; simpl. now apply is_graph_of_uctx_levels. + Qed. + Next Obligation. + repeat split. + - now rewrite mapi_length in X. + - eapply check_constraints_spec; eauto. + Qed. + + Obligation Tactic := Program.Tactics.program_simplify ; eauto 2. + + Program Definition check_is_allowed_elimination (u : Universe.t) (al : allowed_eliminations) : + typing_result (is_allowed_elimination Σ u al) := + let o := + match al return option (is_allowed_elimination Σ u al) with + | IntoSProp => + match Universe.is_sprop u with + | true => Some _ + | false => None + end + | IntoPropSProp => + match is_propositional u with + | true => Some _ + | false => None + end + | IntoSetPropSProp => + match is_propositional u || check_eqb_universe G u Universe.type0 with + | true => Some _ + | false => None + end + | IntoAny => Some _ + end in + match o with + | Some t => Checked t + | None => TypeError (Msg "Cannot eliminate over this sort") + end. + Next Obligation. + unfold is_allowed_elimination, is_allowed_elimination0. + destruct check_univs; auto. + intros val sat. + symmetry in Heq_anonymous. + apply is_sprop_val with (v := val) in Heq_anonymous. + now rewrite Heq_anonymous. + Qed. + Next Obligation. + unfold is_allowed_elimination, is_allowed_elimination0. + destruct check_univs; auto. + intros val sat. + unfold is_propositional in *. + destruct Universe.is_prop eqn:prop. + - apply is_prop_val with (v := val) in prop; rewrite prop; auto. + - destruct Universe.is_sprop eqn:sprop. + + apply is_sprop_val with (v := val) in sprop; rewrite sprop; auto. + + discriminate. + Qed. + Next Obligation. + unfold is_allowed_elimination, is_allowed_elimination0. + destruct check_univs eqn:cu; auto. + intros val sat. + unfold is_propositional in *. + destruct Universe.is_prop eqn:prop. + - apply is_prop_val with (v := val) in prop; rewrite prop; auto. + - destruct Universe.is_sprop eqn:sprop. + + apply is_sprop_val with (v := val) in sprop; rewrite sprop; auto. + + destruct check_eqb_universe eqn:check; [|discriminate]. + eapply check_eqb_universe_spec' in check; eauto. + * unfold eq_universe, eq_universe0 in check. + rewrite cu in check. + specialize (check val sat). + now rewrite check. + * destruct HΣ, Hφ. + now apply wf_ext_global_uctx_invariants. + * destruct HΣ, Hφ. + now apply global_ext_uctx_consistent. + Qed. + Next Obligation. + unfold is_allowed_elimination, is_allowed_elimination0. + destruct check_univs; auto. + Qed. + + Program Fixpoint infer (Γ : context) (HΓ : ∥ wf_local Σ Γ ∥) (t : term) {struct t} + : typing_result ({ A : term & ∥ Σ ;;; Γ |- t : A ∥ }) := + match t with + | tRel n => + match nth_error Γ n with + | Some c => ret ((lift0 (S n)) (decl_type c); _) + | None => raise (UnboundRel n) + end + + | tVar n => raise (UnboundVar n) + | tEvar ev args => raise (UnboundEvar ev) + + | tSort u => + check_eq_true (wf_universeb Σ u) + (Msg ("Sort contains an undeclared level " ^ string_of_sort u));; + ret (tSort (Universe.super u); _) + + | tProd na A B => + s1 <- infer_type infer Γ HΓ A ;; + s2 <- infer_type infer (Γ ,, vass na A) _ B ;; + ret (tSort (Universe.sort_of_product s1.π1 s2.π1); _) + + | tLambda na A t => + s1 <- infer_type infer Γ HΓ A ;; + B <- infer (Γ ,, vass na A) _ t ;; + ret (tProd na A B.π1; _) + + | tLetIn n b b_ty b' => + infer_type infer Γ HΓ b_ty ;; + infer_cumul infer Γ HΓ b b_ty _ ;; + b'_ty <- infer (Γ ,, vdef n b b_ty) _ b' ;; + ret (tLetIn n b b_ty b'_ty.π1; _) + + | tApp t u => + ty <- infer Γ HΓ t ;; + pi <- reduce_to_prod HΣ Γ ty.π1 _ ;; + infer_cumul infer Γ HΓ u pi.π2.π1 _ ;; + ret (subst10 u pi.π2.π2.π1; _) + + | tConst cst u => + match lookup_env (fst Σ) cst with + | Some (ConstantDecl d) => + check_consistent_instance d.(cst_universes) u ;; + let ty := subst_instance_constr u d.(cst_type) in + ret (ty; _) + | _ => raise (UndeclaredConstant cst) + end + + | tInd ind u => + d <- lookup_ind_decl ind ;; + check_consistent_instance d.π1.(ind_universes) u ;; + let ty := subst_instance_constr u d.π2.π1.(ind_type) in + ret (ty; _) + + | tConstruct ind k u => + d <- lookup_ind_decl ind ;; + match nth_error d.π2.π1.(ind_ctors) k with + | Some cdecl => + check_consistent_instance d.π1.(ind_universes) u ;; + ret (type_of_constructor d.π1 cdecl (ind, k) u; _) + | None => raise (UndeclaredConstructor ind k) + end + + | tCase (ind, par) p c brs => + cty <- infer Γ HΓ c ;; + I <- reduce_to_ind HΣ Γ cty.π1 _ ;; + let '(ind'; I') := I in let '(u; I'') := I' in let '(args; H) := I'' in + check_eq_true (eqb ind ind') + (* bad case info *) + (NotConvertible G Γ (tInd ind u) (tInd ind' u)) ;; + d <- lookup_ind_decl ind' ;; + let '(decl; d') := d in let '(body; HH) := d' in + check_coind <- check_eq_true (negb (isCoFinite (ind_finite decl))) + (Msg "Case on coinductives disallowed") ;; + check_eq_true (ind_npars decl =? par) + (Msg "not the right number of parameters") ;; + IS <- infer_scheme infer Γ HΓ p ;; + let '(pctx; IS') := IS in let '(ps; typ_p) := IS' in + check_is_allowed_elimination ps (ind_kelim body);; + let pty := mkAssumArity pctx ps in + let params := firstn par args in + match build_case_predicate_type ind decl body params u ps with + | None => raise (Msg "failure in build_case_predicate_type") + | Some pty' => + (* We could avoid one useless sort comparison by only comparing *) + (* the contexts [pctx] and [indctx] (what is done in Coq). *) + match iscumul Γ pty _ pty' _ with + | ConvError e => raise (NotCumulSmaller G Γ pty pty' pty pty' e) + | ConvSuccess => + match map_option_out (build_branches_type ind decl body params u p) with + | None => raise (Msg "failure in build_branches_type") + | Some btys => + let btyswf : ∥ All (isType Σ Γ ∘ snd) btys ∥ := _ in + (fix check_branches (brs btys : list (nat * term)) + (HH : ∥ All (isType Σ Γ ∘ snd) btys ∥) {struct brs} + : typing_result + (All2 (fun br bty => br.1 = bty.1 /\ ∥ Σ ;;; Γ |- br.2 : bty.2 ∥) brs btys) + := match brs, btys with + | [], [] => ret All2_nil + | (n, t) :: brs , (m, A) :: btys => + W <- check_dec (Msg "not nat eq") + (EqDecInstances.nat_eqdec n m) ;; + Z <- infer_cumul infer Γ HΓ t A _ ;; + X <- check_branches brs btys _ ;; + ret (All2_cons (conj _ _) X) + | [], _ :: _ + | _ :: _, [] => raise (Msg "wrong number of branches") + end) brs btys btyswf ;; + ret (mkApps p (List.skipn par args ++ [c]); _) + end + end + end + + | tProj (ind, n, k) c => + d <- lookup_ind_decl ind ;; + match nth_error d.π2.π1.(ind_projs) k with + | Some pdecl => + c_ty <- infer Γ HΓ c ;; + I <- reduce_to_ind HΣ Γ c_ty.π1 _ ;; + let '(ind'; I') := I in let '(u; I'') := I' in let '(args; H) := I'' in + check_eq_true (eqb ind ind') + (NotConvertible G Γ (tInd ind u) (tInd ind' u)) ;; + check_eq_true (ind_npars d.π1 =? n) + (Msg "not the right number of parameters") ;; + let ty := snd pdecl in + ret (subst0 (c :: List.rev args) (subst_instance_constr u ty); + _) + | None => raise (Msg "projection not found") + end + + | tFix mfix n => + match nth_error mfix n with + | None => raise (IllFormedFix mfix n) + | Some decl => + XX <- (fix check_types (mfix : mfixpoint term) {struct mfix} + : typing_result (∥ All (fun x => isType Σ Γ (dtype x)) mfix ∥) + := match mfix with + | [] => ret (sq All_nil) + | def :: mfix => + (* probably not tail recursive but needed so that next line terminates *) + W <- infer_type infer Γ HΓ (dtype def) ;; + Z <- check_types mfix ;; + ret _ + end) + mfix ;; + YY <- (fix check_bodies (mfix' : mfixpoint term) + (XX : ∥ All (fun x => isType Σ Γ (dtype x)) mfix' ∥) + {struct mfix'} + : typing_result (All (fun d => + ∥ Σ ;;; Γ ,,, fix_context mfix |- dbody d : (lift0 #|fix_context mfix|) (dtype d) ∥) mfix') + := match mfix' with + | [] => ret All_nil + | def :: mfix' => + W1 <- infer_cumul infer (Γ ,,, fix_context mfix) _ (dbody def) + (lift0 #|fix_context mfix| (dtype def)) _ ;; + Z <- check_bodies mfix' _ ;; + ret (All_cons W1 Z) + end) mfix _ ;; + guarded <- check_eq_true (fix_guard Σ Γ mfix) (Msg "Unguarded fixpoint") ;; + wffix <- check_eq_true (wf_fixpoint Σ.1 mfix) (Msg "Ill-formed fixpoint: not defined on a mutually inductive family") ;; + ret (dtype decl; _) + end + + | tCoFix mfix n => + match nth_error mfix n with + | None => raise (IllFormedFix mfix n) + | Some decl => + XX <- (fix check_types (mfix : mfixpoint term) {struct mfix} + : typing_result (∥ All (fun x => isType Σ Γ (dtype x)) mfix ∥) + := match mfix with + | [] => ret (sq All_nil) + | def :: mfix => + (* probably not tail recursive but needed so that next line terminates *) + W <- infer_type infer Γ HΓ (dtype def) ;; + Z <- check_types mfix ;; + ret _ + end) + mfix ;; + YY <- (fix check_bodies (mfix' : mfixpoint term) + (XX' : ∥ All (fun x => isType Σ Γ (dtype x)) mfix' ∥) + {struct mfix'} + : typing_result (All (fun d => + ∥ Σ ;;; Γ ,,, fix_context mfix |- dbody d : (lift0 #|fix_context mfix|) (dtype d) ∥) mfix') + := match mfix' with + | [] => ret All_nil + | def :: mfix' => + W1 <- infer_cumul infer (Γ ,,, fix_context mfix) _ (dbody def) + (lift0 #|fix_context mfix| (dtype def)) _ ;; + Z <- check_bodies mfix' _ ;; + ret (All_cons W1 Z) + end) mfix _ ;; + guarded <- check_eq_true (cofix_guard Σ Γ mfix) (Msg "Unguarded cofixpoint") ;; + wfcofix <- check_eq_true (wf_cofixpoint Σ.1 mfix) (Msg "Ill-formed cofixpoint: not producing values in a mutually coinductive family") ;; + ret (dtype decl; _) + end + + | tPrim _ => raise (Msg "Primitive types are not supported") + end. + + (* tRel *) + Next Obligation. intros; sq; now econstructor. Defined. + (* tSort *) + Next Obligation. + eapply (elimT wf_universe_reflect) in H. + sq; econstructor; tas. + Defined. + (* tProd *) + Next Obligation. + (* intros Γ HΓ t na A B Heq_t [s ?]; *) + sq; econstructor; cbn; easy. Defined. + Next Obligation. + (* intros Γ HΓ t na A B Heq_t [s1 ?] [s2 ?]; *) + sq; econstructor; eassumption. + Defined. + (* tLambda *) + Next Obligation. + (* intros Γ HΓ t0 na A t Heq_t [s ?]; *) + sq; econstructor; cbn; easy. + Defined. + Next Obligation. + (* intros Γ HΓ t0 na A t Heq_t [s ?] [B ?]; *) + sq; econstructor; eassumption. + Defined. + (* tLetIn *) + Next Obligation. + (* intros Γ HΓ t n b b_ty b' Heq_t [? ?]; *) + sq. econstructor; eauto. + Defined. + Next Obligation. + (* intros Γ HΓ t n b b_ty b' Heq_t [? ?] H0; *) + sq; econstructor; cbn; eauto. + Defined. + Next Obligation. + (* intros Γ HΓ t n b b_ty b' Heq_t [? ?] H0 [? ?]; *) + sq; econstructor; eassumption. + Defined. + + (* tApp *) + Next Obligation. simpl; eauto using validity_wf. Qed. + Next Obligation. + cbn in *; sq. + eapply type_reduction in X1 ; try eassumption. + eapply validity_term in X1 ; try assumption. destruct X1 as [s HH]. + eapply inversion_Prod in HH ; try assumption. + destruct HH as [s1 [_ [HH _]]]. + eexists. eassumption. + Defined. + Next Obligation. + cbn in *; sq; eapply type_App'. + 2: eassumption. + eapply type_reduction; eassumption. + Defined. + + (* tConst *) + Next Obligation. + rename Heq_anonymous into HH. + sq; constructor; try assumption. + symmetry in HH. + etransitivity. eassumption. reflexivity. + Defined. + + (* tInd *) + Next Obligation. + sq; econstructor; eassumption. + Defined. + + (* tConstruct *) + Next Obligation. + sq; econstructor; tea. now split. + Defined. + + (* tCase *) + Next Obligation. simpl; eauto using validity_wf. Qed. + Next Obligation. simpl; eauto using validity_wf. Qed. + Next Obligation. + destruct X1, X11. sq. + change (eqb ind I = true) in H0. + destruct (eqb_spec ind I) as [e|e]; [destruct e|discriminate]. + change (eqb (ind_npars d) par = true) in H1. + destruct (eqb_spec (ind_npars d) par) as [e|e]; [|discriminate]. + rename Heq_anonymous into HH. symmetry in HH. + simpl in *. + rewrite <- e in HH. + eapply PCUICInductiveInversion.WfArity_build_case_predicate_type in HH; eauto. + destruct HH as [[s Hs] ?]. eexists; eauto. + eapply isType_red; [|eassumption]. + eapply validity; eauto. + rewrite mkAssumArity_it_mkProd_or_LetIn in X. + eapply validity in X; auto. + eapply PCUICInductives.isType_it_mkProd_or_LetIn_inv in X; eauto. + eapply isType_wf_universes in X; auto. + now exact (elimT wf_universe_reflect X). + Qed. + + Next Obligation. + symmetry in Heq_anonymous1. + unfold iscumul in Heq_anonymous1. simpl in Heq_anonymous1. + apply isconv_term_sound in Heq_anonymous1. + red in Heq_anonymous1. + noconf Heq_I''. + noconf Heq_I'. noconf Heq_I. + noconf Heq_d. noconf Heq_d'. + noconf Heq_IS. noconf Heq_IS'. + simpl in *; sq. + change (eqb ind I = true) in H0. + destruct (eqb_spec ind I) as [e|e]; [destruct e|discriminate H0]. + change (eqb (ind_npars d) par = true) in H1. + destruct (eqb_spec (ind_npars d) par) as [e|e]; [|discriminate]; subst. + assert (wfΣ : wf_ext Σ) by (split; auto). + eapply type_reduction in X11; eauto. + have val:= validity_term wfΣ X11. + eapply type_Cumul' in typ_p; [| |eassumption]. + 2:{ eapply PCUICInductiveInversion.WfArity_build_case_predicate_type; eauto. + eapply validity in typ_p; eauto. + rewrite mkAssumArity_it_mkProd_or_LetIn in typ_p. + eapply PCUICInductives.isType_it_mkProd_or_LetIn_inv in typ_p; eauto. + apply isType_wf_universes in typ_p; auto. + now exact (elimT wf_universe_reflect typ_p). } + have [pctx' da] : (∑ pctx', destArity [] pty' = Some (pctx', X0)). + { symmetry in Heq_anonymous0. + unshelve eapply (PCUICInductives.build_case_predicate_type_spec (Σ.1, ind_universes d)) in Heq_anonymous0 as [parsubst [_ ->]]. + eauto. eapply (PCUICWeakeningEnv.on_declared_inductive wfΣ) in HH as [? ?]. eauto. + eexists. rewrite !destArity_it_mkProd_or_LetIn; simpl. reflexivity. } + eapply PCUICInductiveInversion.build_branches_type_wt. 6:eapply typ_p. all:eauto. + Defined. + + Next Obligation. + sq. + depelim HH; auto. + Defined. + Next Obligation. + sq. + depelim HH; auto. + Defined. + + (* The obligation tactic removes a useful let here. *) + Obligation Tactic := idtac. + Next Obligation. + intros. clearbody btyswf. idtac; Program.Tactics.program_simplify. + symmetry in Heq_anonymous1. + unfold iscumul in Heq_anonymous1. simpl in Heq_anonymous1. + apply isconv_term_sound in Heq_anonymous1. + noconf Heq_I''. noconf Heq_I'. noconf Heq_I. + noconf Heq_d. noconf Heq_d'. + noconf Heq_IS. noconf Heq_IS'. + simpl in *. + assert (∥ All2 (fun x y => ((fst x = fst y) * + (Σ;;; Γ |- snd x : snd y) * isType Σ Γ y.2)%type) brs btys ∥). { + solve_all. simpl in *. + destruct btyswf. eapply All2_sq. solve_all. simpl in *; intuition (sq; auto). } + clear H3. sq. + change (eqb ind I = true) in H0. + destruct (eqb_spec ind I) as [e|e]; [destruct e|discriminate H0]. + change (eqb (ind_npars d) par = true) in H1. + destruct (eqb_spec (ind_npars d) par) as [e|e]; [|discriminate]; subst. + assert (wfΣ : wf_ext Σ) by (split; auto). + eapply type_reduction in X11; eauto. + eapply type_Cumul' in typ_p; eauto. + 2:{ eapply PCUICInductiveInversion.WfArity_build_case_predicate_type; eauto. + eapply validity in X11; eauto. + eapply validity in typ_p; auto. + rewrite mkAssumArity_it_mkProd_or_LetIn in typ_p. + eapply PCUICInductives.isType_it_mkProd_or_LetIn_inv in typ_p; eauto. + apply isType_wf_universes in typ_p; auto. + now exact (elimT wf_universe_reflect typ_p). } + eapply type_Case with (pty0:=pty'); tea. + - reflexivity. + - symmetry; eassumption. + - destruct isCoFinite; auto. + - symmetry; eauto. + Defined. + + Obligation Tactic := Program.Tactics.program_simplify ; eauto 2. + + (* tProj *) + Next Obligation. simpl; eauto using validity_wf. Qed. + Next Obligation. + simpl in *; sq; eapply type_Proj with (pdecl := (i, t0)). + - split. eassumption. split. symmetry; eassumption. cbn in *. + now apply beq_nat_true. + - cbn. destruct (ssrbool.elimT (eqb_spec ind I)); [assumption|]. + eapply type_reduction; eassumption. + - eapply type_reduction in X5; eauto. + eapply validity_term in X5; eauto. + destruct (ssrbool.elimT (eqb_spec ind I)); auto. + unshelve eapply (PCUICInductives.isType_mkApps_Ind _ X7 _) in X5 as [parsubst [argsubst [[sp sp'] cu]]]; eauto. + pose proof (PCUICContexts.context_subst_length2 (PCUICSpine.inst_ctx_subst sp)). + pose proof (PCUICContexts.context_subst_length2 (PCUICSpine.inst_ctx_subst sp')). + autorewrite with len in H, H2. + destruct (PCUICWeakeningEnv.on_declared_inductive HΣ X7) eqn:ond. + rewrite -o.(onNpars) -H. + forward (o0.(onProjections)). + intros H'; rewrite H' nth_error_nil // in Heq_anonymous. + destruct ind_cshapes as [|cs []]; auto. + intros onps. + unshelve epose proof (onps.(on_projs_noidx _ _ _ _ _ _)). + rewrite ond /= in H2. + change (ind_indices o0) with (ind_indices o0) in *. + destruct (ind_indices o0) => //. + simpl in H2. + rewrite List.skipn_length in H2. + rewrite List.firstn_length. lia. + Qed. + + (* tFix *) + Next Obligation. sq. constructor; auto. exists W; auto. Defined. + Next Obligation. sq. now eapply PCUICWeakening.All_mfix_wf in XX0. Defined. + Next Obligation. + sq. cbn in *. depelim XX. + destruct i as [s HH]. + exists s. + change (tSort s) with (lift0 #|fix_context mfix| (tSort s)). + apply weakening; try assumption. + now apply All_mfix_wf. + Defined. + Next Obligation. + clear -XX HΣ. sq. + now depelim XX. + Qed. + Next Obligation. + assert (∥ All (fun d => ((Σ;;; Γ ,,, fix_context mfix |- dbody d : (lift0 #|fix_context mfix|) (dtype d)))%type) mfix ∥). { + eapply All_sq, All_impl. exact YY. + cbn; intros ? ?. now sq. } + sq; econstructor; try eassumption. + symmetry; eassumption. + Qed. + + (* tCoFix *) + Next Obligation. sq. constructor; auto. exists W; auto. Defined. + Next Obligation. sq. now eapply PCUICWeakening.All_mfix_wf in XX. Defined. + Next Obligation. + sq. cbn in *. depelim XX'. + destruct i as [s HH]. + exists s. + change (tSort s) with (lift0 #|fix_context mfix| (tSort s)). + apply weakening; try assumption. + now apply All_mfix_wf. + Defined. + Next Obligation. + clear -XX' HΣ. sq. + now depelim XX'. + Qed. + Next Obligation. + assert (∥ All (fun d => ((Σ;;; Γ ,,, fix_context mfix |- dbody d : (lift0 #|fix_context mfix|) (dtype d)))%type) mfix ∥). { + eapply All_sq, All_impl. exact YY. + now cbn; intros ? []. } + sq; econstructor; try eassumption. + symmetry; eassumption. + Qed. + + Lemma sq_wfl_nil : ∥ wf_local Σ [] ∥. + Proof. + repeat constructor. + Qed. + + Program Fixpoint check_context Γ : typing_result (∥ wf_local Σ Γ ∥) + := match Γ with + | [] => ret sq_wfl_nil + | {| decl_body := None; decl_type := A |} :: Γ => + HΓ <- check_context Γ ;; + XX <- infer_type infer Γ HΓ A ;; + ret _ + | {| decl_body := Some t; decl_type := A |} :: Γ => + HΓ <- check_context Γ ;; + XX <- infer_type infer Γ HΓ A ;; + XX <- infer_cumul infer Γ HΓ t A _ ;; + ret _ + end. + Next Obligation. + sq. econstructor; tas. econstructor; eauto. + Qed. + Next Obligation. + sq. econstructor; tea. + Qed. + Next Obligation. + sq. econstructor; tas. econstructor; eauto. + Qed. + +(* + Program Definition check_isWfArity Γ (HΓ : ∥ wf_local Σ Γ ∥) A + : typing_result (∥ isWfArity Σ Γ A ∥) := + match destArity [] A with + | None => raise (Msg (print_term Σ Γ A ^ " is not an arity")) + | Some (ctx, s) => XX <- check_context (Γ ,,, ctx) ;; + ret _ + end. + Next Obligation. + destruct XX. constructor. exists ctx, s. + split; auto. + Defined. *) + + Program Definition check_isType Γ (HΓ : ∥ wf_local Σ Γ ∥) A + : typing_result (∥ isType Σ Γ A ∥) := + s <- infer Γ HΓ A ;; + s' <- reduce_to_sort HΣ Γ s.π1 _ ;; + ret _. + Next Obligation. now eapply validity_wf. Defined. + Next Obligation. destruct X0. sq. eexists. eapply type_reduction; tea. Defined. + + Program Definition check Γ (HΓ : ∥ wf_local Σ Γ ∥) t A + : typing_result (∥ Σ;;; Γ |- t : A ∥) := + check_isType Γ HΓ A ;; + infer_cumul infer Γ HΓ t A _. + +End Typecheck. diff --git a/safechecker/theories/PCUICWfEnv.v b/safechecker/theories/PCUICWfEnv.v new file mode 100644 index 000000000..3dacd3fc4 --- /dev/null +++ b/safechecker/theories/PCUICWfEnv.v @@ -0,0 +1,62 @@ + +(* Distributed under the terms of the MIT license. *) +From Coq Require Import ProofIrrelevance. +From MetaCoq.Template Require Import config utils uGraph. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils + PCUICReflect PCUICTyping PCUICGlobalEnv. + +Lemma wf_ext_gc_of_uctx {cf:checker_flags} {Σ : global_env_ext} (HΣ : ∥ wf_ext Σ ∥) + : ∑ uctx', gc_of_uctx (global_ext_uctx Σ) = Some uctx'. +Proof. + assert (consistent (global_ext_uctx Σ).2) as HC. + { sq; apply (global_ext_uctx_consistent _ HΣ). } + destruct Σ as [Σ φ]. + simpl in HC. + unfold gc_of_uctx; simpl in *. + apply gc_consistent_iff in HC. + destruct (gc_of_constraints (global_ext_constraints (Σ, φ))). + eexists; reflexivity. + contradiction HC. +Defined. + +Lemma graph_of_wf_ext {cf:checker_flags} {Σ : global_env_ext} (HΣ : ∥ wf_ext Σ ∥) + : ∑ G, is_graph_of_uctx G (global_ext_uctx Σ). +Proof. + destruct (wf_ext_gc_of_uctx HΣ) as [uctx Huctx]. + exists (make_graph uctx). unfold is_graph_of_uctx. now rewrite Huctx. +Defined. + +Section GraphSpec. + Context {cf:checker_flags} {Σ : global_env_ext} (HΣ : ∥ wf Σ ∥) + (Hφ : ∥ on_udecl Σ.1 Σ.2 ∥) + (G : universes_graph) (HG : is_graph_of_uctx G (global_ext_uctx Σ)). + + Local Definition HΣ' : ∥ wf_ext Σ ∥. + Proof. + destruct HΣ, Hφ; now constructor. + Defined. + + Lemma check_constraints_spec ctrs + : check_constraints G ctrs -> valid_constraints (global_ext_constraints Σ) ctrs. + Proof. + pose proof HΣ'. + intros HH. + refine (check_constraints_spec G (global_ext_uctx Σ) _ _ HG _ HH). + sq; now eapply wf_ext_global_uctx_invariants. + sq; now apply global_ext_uctx_consistent. + Qed. + + Lemma is_graph_of_uctx_levels (l : Level.t) : + LevelSet.mem l (uGraph.wGraph.V G) -> + LevelSet.mem l (global_ext_levels Σ). + Proof. + unfold is_graph_of_uctx in HG. + case_eq (gc_of_uctx (global_ext_uctx Σ)); [intros [lvs cts] XX|intro XX]; + rewrite -> XX in *; simpl in *; [|contradiction]. + unfold gc_of_uctx in XX; simpl in XX. + destruct (gc_of_constraints Σ); [|discriminate]. + inversion XX; subst. generalize (global_ext_levels Σ); intros lvs; cbn. + clear. intro H. apply LevelSet.mem_spec. now apply LevelSet.mem_spec in H. + Qed. + +End GraphSpec. \ No newline at end of file diff --git a/safechecker/theories/PCUICWfReduction.v b/safechecker/theories/PCUICWfReduction.v index 91ae67eb6..befb3a7a8 100644 --- a/safechecker/theories/PCUICWfReduction.v +++ b/safechecker/theories/PCUICWfReduction.v @@ -6,7 +6,6 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICWeakening PCUICPosition PCUICCumulativity PCUICSafeLemmata PCUICSN PCUICPretty PCUICArities PCUICConfluence PCUICSize PCUICContextConversion PCUICConversion PCUICWfUniverses. -From MetaCoq.SafeChecker Require Import PCUICSafeReduce PCUICSafeConversion. From Equations Require Import Equations. Require Import ssreflect ssrbool. @@ -268,4 +267,4 @@ Section fix_sigma. | H : ∥ _ ∥ |- _ => destruct H end; try eapply sq. -End fix_sigma. \ No newline at end of file +End fix_sigma. diff --git a/safechecker/theories/SafeTemplateChecker.v b/safechecker/theories/SafeTemplateChecker.v index e07ebaff0..737e3f9d6 100644 --- a/safechecker/theories/SafeTemplateChecker.v +++ b/safechecker/theories/SafeTemplateChecker.v @@ -4,7 +4,7 @@ From MetaCoq.Template Require Import config utils. From MetaCoq.Template Require AstUtils Typing. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping TemplateToPCUIC. -From MetaCoq.SafeChecker Require Import PCUICSafeReduce PCUICSafeChecker. +From MetaCoq.SafeChecker Require Import PCUICErrors PCUICSafeChecker. Program Definition infer_template_program {cf : checker_flags} (p : Ast.program) φ Hφ @@ -12,7 +12,7 @@ Program Definition infer_template_program {cf : checker_flags} (p : Ast.program) p <- typecheck_program (cf:=cf) (trans_global_decls p.1, trans p.2) φ Hφ ;; ret (p.π1 ; _). -(** In Coq until 8.10, programs can be ill-formed w.r.t. universes as they don't include +(** In Coq until 8.11 at least, programs can be ill-formed w.r.t. universes as they don't include all declarations of universes and constraints coming from section variable declarations. We hence write a program that computes the dangling universes in an Ast.program and registers them appropriately. *) @@ -107,7 +107,7 @@ Program Definition infer_and_print_template_program {cf : checker_flags} (p : As | EnvError Σ (AlreadyDeclared id) => inr ("Already declared: " ^ id) | EnvError Σ (IllFormedDecl id e) => - inr ("Type error: " ^ PCUICSafeChecker.string_of_type_error Σ e ^ ", while checking " ^ id) + inr ("Type error: " ^ string_of_type_error Σ e ^ ", while checking " ^ id) end. (* Program Definition check_template_program {cf : checker_flags} (p : Ast.program) (ty : Ast.term) *) diff --git a/template-coq/_CoqProject b/template-coq/_CoqProject index 91e041d1e..16ab96933 100644 --- a/template-coq/_CoqProject +++ b/template-coq/_CoqProject @@ -16,6 +16,7 @@ theories/utils/MCSquash.v theories/utils/MCString.v theories/utils/wGraph.v theories/utils/MCUtils.v +theories/utils/MC_ExtrOCamlInt63.v # common theories/common/uGraph.v @@ -36,8 +37,13 @@ theories/UnivSubst.v theories/Pretty.v theories/EnvironmentTyping.v theories/WfInv.v +theories/TermEquality.v theories/Typing.v +theories/Reduction.v theories/TypingWf.v +theories/Normal.v +theories/WcbvEval.v +theories/Checker.v # The Template monad theories/TemplateMonad.v diff --git a/template-coq/_PluginProject b/template-coq/_PluginProject index 1cdd82476..1b4377926 100644 --- a/template-coq/_PluginProject +++ b/template-coq/_PluginProject @@ -7,6 +7,8 @@ gen-src/signature.mli gen-src/signature.ml gen-src/classes0.mli gen-src/classes0.ml +gen-src/eqDec.mli +gen-src/eqDec.ml gen-src/all_Forall.ml gen-src/all_Forall.mli gen-src/ascii.ml @@ -56,12 +58,37 @@ gen-src/hexadecimal.ml gen-src/hexadecimal.mli #gen-src/induction.ml #gen-src/induction.mli +gen-src/int63.ml +gen-src/int63.mli + +gen-src/sumbool.mli +gen-src/sumbool.ml +gen-src/zeven.mli +gen-src/zeven.ml +gen-src/zArith_dec.mli +gen-src/zArith_dec.ml +gen-src/zbool.mli +gen-src/zbool.ml +gen-src/zpower.mli +gen-src/zpower.ml +gen-src/specFloat.mli +gen-src/specFloat.ml +gen-src/floatOps.mli +gen-src/floatOps.ml +# gen-src/float64.mli +# gen-src/float64.ml gen-src/liftSubst.ml gen-src/liftSubst.mli gen-src/list0.ml gen-src/list0.mli gen-src/logic0.ml gen-src/logic0.mli +gen-src/logic1.mli +gen-src/logic1.ml +gen-src/relation.mli +gen-src/relation.ml +gen-src/mCPrelude.mli +gen-src/mCPrelude.ml gen-src/mCCompare.ml gen-src/mCCompare.mli gen-src/mCList.ml @@ -72,6 +99,10 @@ gen-src/mCProd.ml gen-src/mCProd.mli gen-src/mCRelations.ml gen-src/mCRelations.mli +gen-src/primFloat.mli +gen-src/primFloat.ml +gen-src/decimalString.mli +gen-src/decimalString.ml gen-src/mCString.ml gen-src/mCString.mli gen-src/mCUtils.ml diff --git a/template-coq/gen-src/metacoq_template_plugin.mlpack b/template-coq/gen-src/metacoq_template_plugin.mlpack index 4dffdfac7..14c6b04c4 100644 --- a/template-coq/gen-src/metacoq_template_plugin.mlpack +++ b/template-coq/gen-src/metacoq_template_plugin.mlpack @@ -13,6 +13,7 @@ BinPosDef BinPos BinNat BinInt +Int63 Ascii String0 Orders @@ -26,17 +27,31 @@ MSetDecide MSetList MSetProperties BinNums +EqDec CRelationClasses Compare_dec +Classes0 +EqDecInstances +Logic1 +Relation +MCPrelude MCList MCRelations MCOption MCProd MCCompare +Sumbool +Zeven +ZArith_dec +Zbool +Zpower +SpecFloat +PrimFloat +FloatOps +DecimalString MCString MCUtils Signature -Classes0 All_Forall Config0 Universes0 @@ -47,7 +62,6 @@ AstUtils LiftSubst UnivSubst0 Induction -EqDecInstances Reflect Pretty Common0 diff --git a/template-coq/gen-src/specFloat.ml.rej b/template-coq/gen-src/specFloat.ml.rej new file mode 100644 index 000000000..294627074 --- /dev/null +++ b/template-coq/gen-src/specFloat.ml.rej @@ -0,0 +1,10 @@ +--- gen-src/specFloat.ml.orig 2020-12-04 21:48:20.000000000 +0100 ++++ gen-src/specFloat.ml 2020-12-04 21:48:33.000000000 +0100 +@@ -4,6 +4,7 @@ + open Datatypes + open Zbool + open Zpower ++open Float64 + + type spec_float = + | S754_zero of bool diff --git a/template-coq/specFloat.patch b/template-coq/specFloat.patch new file mode 100644 index 000000000..294627074 --- /dev/null +++ b/template-coq/specFloat.patch @@ -0,0 +1,10 @@ +--- gen-src/specFloat.ml.orig 2020-12-04 21:48:20.000000000 +0100 ++++ gen-src/specFloat.ml 2020-12-04 21:48:33.000000000 +0100 +@@ -4,6 +4,7 @@ + open Datatypes + open Zbool + open Zpower ++open Float64 + + type spec_float = + | S754_zero of bool diff --git a/template-coq/src/ast_denoter.ml b/template-coq/src/ast_denoter.ml index 6f4e000cb..bfb59c939 100644 --- a/template-coq/src/ast_denoter.ml +++ b/template-coq/src/ast_denoter.ml @@ -9,6 +9,8 @@ struct type t = Ast0.term type quoted_ident = char list type quoted_int = Datatypes.nat + type quoted_int63 = Uint63.t + type quoted_float64 = Float64.t type quoted_bool = bool type quoted_name = name type quoted_aname = name binder_annot @@ -68,6 +70,8 @@ struct let mkInd = mkInd let mkCase = mkCase let mkProj = mkProj + let mkInt = mkInt + let mkFloat = mkFloat let unquote_def (x: 't BasicAst.def) : ('t, name binder_annot, quoted_int) adef = { @@ -77,7 +81,9 @@ struct rarg=rarg x } - let inspect_term (tt: t):(t, quoted_int, quoted_ident, quoted_aname, quoted_sort, quoted_cast_kind, quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj) structure_of_term= + let inspect_term (tt: t):(t, quoted_int, quoted_ident, quoted_aname, quoted_sort, quoted_cast_kind, + quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj, + quoted_int63, quoted_float64) structure_of_term= match tt with | Coq_tRel n -> ACoq_tRel n | Coq_tVar v -> ACoq_tVar v @@ -95,7 +101,8 @@ struct | Coq_tProj (a,b) -> ACoq_tProj (a,b) | Coq_tFix (a,b) -> ACoq_tFix (List.map unquote_def a,b) | Coq_tCoFix (a,b) -> ACoq_tCoFix (List.map unquote_def a,b) - + | Coq_tInt i -> ACoq_tInt i + | Coq_tFloat f -> ACoq_tFloat f let unquote_ident (qi: quoted_ident) : Id.t = let s = list_to_string qi in @@ -119,8 +126,16 @@ struct match q with | Datatypes.O -> 0 | Datatypes.S x -> succ (unquote_int x) + + let unquote_evar env evm n l = + let id = Evar.unsafe_of_int (unquote_int n) in + evm, mkEvar (id, l) let unquote_bool (q : quoted_bool) : bool = q + + let unquote_int63 i = i + + let unquote_float64 i = i (* val unquote_sort : quoted_sort -> Sorts.t *) (* val unquote_sort_family : quoted_sort_family -> Sorts.family *) diff --git a/template-coq/src/ast_quoter.ml b/template-coq/src/ast_quoter.ml index 1739e4a17..176f49aa3 100644 --- a/template-coq/src/ast_quoter.ml +++ b/template-coq/src/ast_quoter.ml @@ -9,6 +9,8 @@ struct type t = Ast0.term type quoted_ident = char list type quoted_int = Datatypes.nat + type quoted_int63 = Uint63.t + type quoted_float64 = Float64.t type quoted_bool = bool type quoted_name = BasicAst.name type quoted_aname = BasicAst.aname @@ -55,7 +57,7 @@ struct let quote_relevance = function | Sorts.Relevant -> BasicAst.Relevant | Sorts.Irrelevant -> BasicAst.Irrelevant - + let quote_name = function | Anonymous -> Coq_nAnon | Name i -> Coq_nNamed (quote_ident i) @@ -72,6 +74,10 @@ struct let quote_bool x = x + let quote_int63 x = x + + let quote_float64 x = x + (* NOTE: fails if it hits Prop or SProp *) let quote_nonprop_level (l : Univ.Level.t) : Universes0.Level.t = if Univ.Level.is_prop l || Univ.Level.is_sprop l then @@ -150,7 +156,7 @@ struct let is_Eq = function | Univ.Eq -> true | _ -> false - + let quote_univ_constraint ((l, ct, l') : Univ.univ_constraint) : quoted_univ_constraint = ((quote_nonprop_level l, quote_constraint_type ct), quote_nonprop_level l') @@ -232,6 +238,8 @@ struct let mkInd i u = Coq_tInd (i, u) let mkConstruct (ind, i) u = Coq_tConstruct (ind, i, u) let mkLetIn na b t t' = Coq_tLetIn (na,b,t,t') + let mkInt i = Coq_tInt i + let mkFloat f = Coq_tFloat f let rec seq f t = if f < t then @@ -278,7 +286,7 @@ struct let mk_mutual_inductive_body finite npars params inds uctx variance = {ind_finite = finite; - ind_npars = npars; ind_params = params; ind_bodies = inds; + ind_npars = npars; ind_params = params; ind_bodies = inds; ind_universes = uctx; ind_variance = variance} let mk_constant_body ty tm uctx = @@ -315,7 +323,7 @@ struct mind_entry_cumulative = cumulative; mind_entry_private = None } - let quote_definition_entry ty body ctx = + let quote_definition_entry ty body ctx = { definition_entry_type = ty; definition_entry_body = body; definition_entry_universes = ctx; @@ -324,11 +332,11 @@ struct let quote_parameter_entry ty ctx = { parameter_entry_type = ty; parameter_entry_universes = ctx } - + let quote_constant_entry = function | Left ce -> DefinitionEntry ce | Right pe -> ParameterEntry pe -(* +(* let quote_entry e = match e with | Some (Left (ty, body, ctx)) -> @@ -337,9 +345,12 @@ struct Some (Right mind_entry) | None -> None *) - let inspectTerm (t : term) : (term, quoted_int, quoted_ident, quoted_name, quoted_sort, quoted_cast_kind, quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj) structure_of_term = + let inspectTerm (t : term) : (term, quoted_int, quoted_ident, quoted_name, quoted_sort, quoted_cast_kind, + quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj, + quoted_int63, quoted_float64) structure_of_term = match t with | Coq_tRel n -> ACoq_tRel n + (* todo ! *) (* so on *) | _ -> failwith "not yet implemented" @@ -364,7 +375,7 @@ struct let mkPolymorphic_entry names c = Universes0.Polymorphic_entry (names, c) let mkMonomorphic_entry c = Universes0.Monomorphic_entry c - + end module ExtractedASTQuoter = Quoter.Quoter(ExtractedASTBaseQuoter) diff --git a/template-coq/src/constr_denoter.ml b/template-coq/src/constr_denoter.ml index e0f95a4b5..e8388c1cd 100644 --- a/template-coq/src/constr_denoter.ml +++ b/template-coq/src/constr_denoter.ml @@ -73,6 +73,16 @@ struct false else not_supported_verb trm "from_bool" + let unquote_int63 trm = + match Constr.kind trm with + | Constr.Int i -> i + | _ -> not_supported_verb trm "unquote_int63" + + let unquote_float64 trm = + match Constr.kind trm with + | Constr.Float f -> f + | _ -> not_supported_verb trm "unquote_float64" + let unquote_char trm = let (h,args) = app_full trm [] in if constr_equall h tAscii then @@ -129,6 +139,16 @@ struct else not_supported_verb trm "unquote_name" + let unquote_evar env evm id args = + if constr_equall id tfresh_evar_id then + let evm, (tyev, s) = Evarutil.new_type_evar env evm Evd.univ_flexible_alg in + let evm, ev = Evarutil.new_evar env evm tyev in + evm, EConstr.Unsafe.to_constr ev + else + let id = unquote_nat id in + let ev = Evar.unsafe_of_int id in + evm, Constr.mkEvar (ev, args) + let unquote_relevance trm = if Constr.equal trm (Lazy.force tRelevant) then Sorts.Relevant @@ -326,7 +346,9 @@ struct let inspect_term (t:Constr.t) - : (Constr.t, quoted_int, quoted_ident, quoted_name, quoted_sort, quoted_cast_kind, quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj) structure_of_term = + : (Constr.t, quoted_int, quoted_ident, quoted_name, quoted_sort, quoted_cast_kind, quoted_kernel_name, + quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj, + quoted_int63, quoted_float64) structure_of_term = let (h,args) = app_full t [] in if constr_equall h tRel then match args with @@ -336,6 +358,10 @@ struct match args with x :: _ -> ACoq_tVar x | _ -> CErrors.user_err (print_term t ++ Pp.str ("has bad structure")) + else if constr_equall h tEvar then + match args with + | [x; y] -> ACoq_tEvar (x, unquote_list y) + | _ -> CErrors.user_err (print_term t ++ Pp.str ("has bad structure")) else if constr_equall h tSort then match args with x :: _ -> ACoq_tSort x @@ -414,7 +440,14 @@ struct match args with proj::t::_ -> ACoq_tProj (proj, t) | _ -> CErrors.user_err (print_term t ++ Pp.str ("has bad structure")) - + else if constr_equall h tInt then + match args with + t::_ -> ACoq_tInt t + | _ -> CErrors.user_err (print_term t ++ Pp.str ("has bad structure")) + else if constr_equall h tFloat then + match args with + t::_ -> ACoq_tFloat t + | _ -> CErrors.user_err (print_term t ++ Pp.str ("has bad structure")) else CErrors.user_err (str"inspect_term: cannot recognize " ++ print_term t ++ str" (maybe you forgot to reduce it?)") diff --git a/template-coq/src/constr_quoter.ml b/template-coq/src/constr_quoter.ml index 8a7f3e86a..78720158d 100644 --- a/template-coq/src/constr_quoter.ml +++ b/template-coq/src/constr_quoter.ml @@ -80,6 +80,9 @@ struct let mkProj kn t = constr_mkApp (tProj, [| kn; t |]) + let mkInt i = i + let mkFloat f = f + let quote_option ty = function | Some tm -> constr_mkApp (cSome, [|ty; tm|]) | None -> constr_mkApp (cNone, [|ty|]) @@ -109,6 +112,10 @@ struct let quote_bool b = if b then Lazy.force ttrue else Lazy.force tfalse + let quote_int63 i = constr_mkApp (tInt, [| Constr.mkInt i |]) + + let quote_float64 f = constr_mkApp (tFloat, [| Constr.mkFloat f |]) + let quote_char i = constr_mkApp (tAscii, Array.of_list (List.map (fun m -> quote_bool ((i land m) = m)) (List.rev [128;64;32;16;8;4;2;1]))) diff --git a/template-coq/src/constr_reification.ml b/template-coq/src/constr_reification.ml index a4c85e2b2..1ffa5ff60 100644 --- a/template-coq/src/constr_reification.ml +++ b/template-coq/src/constr_reification.ml @@ -15,6 +15,8 @@ struct type quoted_kernel_name = Constr.t (* of type Ast.kername *) type quoted_inductive = Constr.t (* of type Ast.inductive *) type quoted_proj = Constr.t (* of type Ast.projection *) + type quoted_int63 = Constr.t (* of type UInt63.t *) + type quoted_float64 = Constr.t (* of type Float64.t *) type quoted_global_reference = Constr.t (* of type Ast.global_reference *) type quoted_sort_family = Constr.t (* of type Ast.sort_family *) @@ -77,17 +79,17 @@ struct let cZ0 = resolve "metacoq.Z.zero" let cZpos = resolve "metacoq.Z.pos" let cZneg = resolve "metacoq.Z.neg" - + let tpos = resolve "metacoq.pos.type" let cposzero = resolve "metacoq.pos.xH" let cposI = resolve "metacoq.pos.xI" let cposO = resolve "metacoq.pos.xO" - + let cSome_instance = resolve "metacoq.option_instance.some" let cNone_instance = resolve "metacoq.option_instance.none" let unit_tt = resolve "metacoq.unit.intro" - + let tAscii = resolve "metacoq.ascii.intro" let tlist = resolve "metacoq.list.type" let c_nil = resolve "metacoq.list.nil" @@ -138,17 +140,18 @@ struct let tmkInd = ast "mkInd" let tmkdecl = ast "mkdecl" let (tTerm,tRel,tVar,tEvar,tSort,tCast,tProd, - tLambda,tLetIn,tApp,tCase,tFix,tConstructor,tConst,tInd,tCoFix,tProj) = + tLambda,tLetIn,tApp,tCase,tFix,tConstructor,tConst,tInd,tCoFix,tProj,tInt,tFloat) = (ast "term", ast "tRel", ast "tVar", ast "tEvar", ast "tSort", ast "tCast", ast "tProd", ast "tLambda", ast "tLetIn", ast "tApp", ast "tCase", ast "tFix", - ast "tConstruct", ast "tConst", ast "tInd", ast "tCoFix", ast "tProj") + ast "tConstruct", ast "tConst", ast "tInd", ast "tCoFix", ast "tProj", ast "tInt", ast "tFloat") let tkername = ast "kername" let tmodpath = ast "modpath" let tMPfile = ast "MPfile" let tMPbound = ast "MPbound" let tMPdot = ast "MPdot" - + let tfresh_evar_id = ast "fresh_evar_id" + let tproplevel = ast "level.prop_level_type" let tlevelSProp = ast "level.lsprop" let tlevelProp = ast "level.lprop" @@ -222,7 +225,7 @@ struct let cMonomorphic_entry = ast "Monomorphic_entry" let cPolymorphic_entry = ast "Polymorphic_entry" - let (tcbv, tcbn, thnf, tall, tlazy, tunfold) = + let (tcbv, tcbn, thnf, tall, tlazy, tunfold) = (template "cbv", template "cbn", template "hnf", template "all", template "lazy", template "unfold") diff --git a/template-coq/src/denoter.ml b/template-coq/src/denoter.ml index 938d2c45d..1c673c757 100644 --- a/template-coq/src/denoter.ml +++ b/template-coq/src/denoter.ml @@ -10,8 +10,11 @@ sig val unquote_name : quoted_name -> Name.t val unquote_aname : quoted_aname -> Name.t Context.binder_annot val unquote_relevance : quoted_relevance -> Sorts.relevance - val unquote_int : quoted_int -> int + val unquote_evar : Environ.env -> Evd.evar_map -> quoted_int -> Constr.t list -> Evd.evar_map * Constr.t + val unquote_int : quoted_int -> int val unquote_bool : quoted_bool -> bool + val unquote_int63 : quoted_int63 -> Uint63.t + val unquote_float64 : quoted_float64 -> Float64.t (* val unquote_sort : quoted_sort -> Sorts.t *) (* val unquote_sort_family : quoted_sort_family -> Sorts.family *) val unquote_cast_kind : quoted_cast_kind -> Constr.cast_kind @@ -22,7 +25,9 @@ sig val unquote_universe : Evd.evar_map -> quoted_sort -> Evd.evar_map * Univ.Universe.t val unquote_universe_instance: Evd.evar_map -> quoted_univ_instance -> Evd.evar_map * Univ.Instance.t (* val representsIndConstuctor : quoted_inductive -> Term.constr -> bool *) - val inspect_term : t -> (t, quoted_int, quoted_ident, quoted_aname, quoted_sort, quoted_cast_kind, quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj) structure_of_term + val inspect_term : t -> (t, quoted_int, quoted_ident, quoted_aname, quoted_sort, quoted_cast_kind, + quoted_kernel_name, quoted_inductive, quoted_relevance, quoted_univ_instance, quoted_proj, + quoted_int63, quoted_float64) structure_of_term end @@ -36,6 +41,9 @@ let map_evm (f : 'a -> 'b -> 'a * 'c) (evm : 'a) (l : 'b list) : 'a * ('c list) let evm, res = List.fold_left (fun (evm, l) b -> let evm, c = f evm b in evm, c :: l) (evm, []) l in evm, List.rev res +let fold_env_evm_right (f : 'a -> 'b -> 'c -> 'a * 'b * 'd) (env : 'a) (evm : 'b) (l : 'c list) : 'a * 'b * ('d list) = + List.fold_right (fun b (env, evm, l) -> let env, evm, c = f env evm b in env, evm, c :: l) l (env, evm, []) + module Denoter (D : BaseDenoter) = struct @@ -45,28 +53,42 @@ struct ACoq_tApp (f, xs) -> app_full_abs f (xs @ acc) | _ -> (trm, acc) - let denote_term (evm : Evd.evar_map) (trm: D.t) : Evd.evar_map * Constr.t = - let rec aux evm (trm: D.t) : _ * Constr.t = + let push_rec_types (lna,typarray) env = + let open Context.Rel.Declaration in + let ctxt = CArray.map2_i (fun i na t -> LocalAssum (na, Vars.lift i t)) lna typarray in + Array.fold_left (fun e assum -> Environ.push_rel assum e) env ctxt + + let denote_term (env : Environ.env) (evm : Evd.evar_map) (trm: D.t) : Evd.evar_map * Constr.t = + let open Context.Rel.Declaration in + let rec aux env evm (trm: D.t) : _ * Constr.t = (* debug (fun () -> Pp.(str "denote_term" ++ spc () ++ pr_constr trm)) ; *) match D.inspect_term trm with | ACoq_tRel x -> evm, Constr.mkRel (D.unquote_int x + 1) | ACoq_tVar x -> evm, Constr.mkVar (D.unquote_ident x) + | ACoq_tEvar (n, l) -> + let evm, l' = map_evm (aux env) evm l in + D.unquote_evar env evm n l' | ACoq_tSort x -> let evm, u = D.unquote_universe evm x in evm, Constr.mkType u - | ACoq_tCast (t,c,ty) -> let evm, t = aux evm t in - let evm, ty = aux evm ty in + | ACoq_tCast (t,c,ty) -> let evm, t = aux env evm t in + let evm, ty = aux env evm ty in evm, Constr.mkCast (t, D.unquote_cast_kind c, ty) - | ACoq_tProd (n,t,b) -> let evm, t = aux evm t in - let evm, b = aux evm b in - evm, Constr.mkProd (D.unquote_aname n, t, b) - | ACoq_tLambda (n,t,b) -> let evm, t = aux evm t in - let evm, b = aux evm b in - evm, Constr.mkLambda (D.unquote_aname n, t, b) - | ACoq_tLetIn (n,e,t,b) -> let evm, e = aux evm e in - let evm, t = aux evm t in - let evm, b = aux evm b in - evm, Constr.mkLetIn (D.unquote_aname n, e, t, b) - | ACoq_tApp (f,xs) -> let evm, f = aux evm f in - let evm, xs = map_evm aux evm xs in + | ACoq_tProd (n,t,b) -> let evm, t = aux env evm t in + let n = D.unquote_aname n in + let evm, b = aux (Environ.push_rel (LocalAssum (n, t)) env) evm b in + evm, Constr.mkProd (n, t, b) + | ACoq_tLambda (n,t,b) -> + let n = D.unquote_aname n in + let evm, t = aux env evm t in + let evm, b = aux (Environ.push_rel (LocalAssum (n, t)) env) evm b in + evm, Constr.mkLambda (n, t, b) + | ACoq_tLetIn (n,e,t,b) -> + let n = D.unquote_aname n in + let evm, e = aux env evm e in + let evm, t = aux env evm t in + let evm, b = aux (Environ.push_rel (LocalDef (n, e, t)) env) evm b in + evm, Constr.mkLetIn (n, e, t, b) + | ACoq_tApp (f,xs) -> let evm, f = aux env evm f in + let evm, xs = map_evm (aux env) evm xs in evm, Constr.mkApp (f, Array.of_list xs) | ACoq_tConst (s,u) -> let s = D.unquote_kn s in @@ -83,28 +105,32 @@ struct | ACoq_tCase (((i, _), r), ty, d, brs) -> let ind = D.unquote_inductive i in let relevance = D.unquote_relevance r in - let evm, ty = aux evm ty in - let evm, d = aux evm d in - let evm, brs = map_evm aux evm (List.map snd brs) in + let evm, ty = aux env evm ty in + let evm, d = aux env evm d in + let evm, brs = map_evm (aux env) evm (List.map snd brs) in (* todo: reify better case_info *) let ci = Inductiveops.make_case_info (Global.env ()) ind relevance Constr.RegularStyle in evm, Constr.mkCase (ci, ty, d, Array.of_list brs) | ACoq_tFix (lbd, i) -> let (names,types,bodies,rargs) = (List.map (fun p->p.adname) lbd, List.map (fun p->p.adtype) lbd, List.map (fun p->p.adbody) lbd, List.map (fun p->p.rarg) lbd) in - let evm, types = map_evm aux evm types in - let evm, bodies = map_evm aux evm bodies in + let evm, types = map_evm (aux env) evm types in let (names,rargs) = (List.map D.unquote_aname names, List.map D.unquote_int rargs) in let la = Array.of_list in - evm, Constr.mkFix ((la rargs, D.unquote_int i), (la names, la types, la bodies)) + let lnames = la names and ltypes = la types in + let env = push_rec_types (lnames, ltypes) env in + let evm, bodies = map_evm (aux env) evm bodies in + evm, Constr.mkFix ((la rargs, D.unquote_int i), (lnames, ltypes, la bodies)) | ACoq_tCoFix (lbd, i) -> let (names,types,bodies,rargs) = (List.map (fun p->p.adname) lbd, List.map (fun p->p.adtype) lbd, List.map (fun p->p.adbody) lbd, List.map (fun p->p.rarg) lbd) in - let evm, types = map_evm aux evm types in - let evm, bodies = map_evm aux evm bodies in + let evm, types = map_evm (aux env) evm types in let (names,rargs) = (List.map D.unquote_aname names, List.map D.unquote_int rargs) in let la = Array.of_list in - evm, Constr.mkCoFix (D.unquote_int i, (la names, la types, la bodies)) + let lnames = la names and ltypes = la types in + let env = push_rec_types (lnames, ltypes) env in + let evm, bodies = map_evm (aux env) evm bodies in + evm, Constr.mkCoFix (D.unquote_int i, (lnames, ltypes, la bodies)) | ACoq_tProj (proj,t) -> let (ind, npars, arg) = D.unquote_proj proj in @@ -115,19 +141,11 @@ struct | Some p -> Names.Constant.label p | None -> failwith "tproj case of denote_term") in let p' = Names.Projection.make (Projection.Repr.make ind' ~proj_npars ~proj_arg l) false in - let evm, t' = aux evm t in + let evm, t' = aux env evm t in evm, Constr.mkProj (p', t') - (* | _ -> not_supported_verb trm "big_case" - * - * | ACoq_tProj (proj,t) -> - * let (ind, _, narg) = D.unquote_proj proj in (\* todo: is narg the correct projection? *\) - * let ind' = D.unquote_inductive ind in - * let projs = Recordops.lookup_projections ind' in - * let evm, t = aux evm t in - * (match List.nth projs (D.unquote_int narg) with - * | Some p -> evm, Constr.mkProj (Names.Projection.make p false, t) - * | None -> (\*bad_term trm *\) ) *) - | _ -> failwith "big case of denote_term" - in aux evm trm + | ACoq_tInt x -> evm, Constr.mkInt (D.unquote_int63 x) + | ACoq_tFloat x -> evm, Constr.mkFloat (D.unquote_float64 x) + + in aux env evm trm end diff --git a/template-coq/src/g_template_coq.mlg b/template-coq/src/g_template_coq.mlg index 8ef6c1e7d..fcf1d0cea 100644 --- a/template-coq/src/g_template_coq.mlg +++ b/template-coq/src/g_template_coq.mlg @@ -181,8 +181,9 @@ END TACTIC EXTEND TemplateCoq_denote_term | [ "denote_term" constr(c) tactic(tac) ] -> { Proofview.Goal.enter (begin fun gl -> + let env = Proofview.Goal.env gl in let evm = Proofview.Goal.sigma gl in - let evm, c = Constr_denoter.denote_term evm (to_constr_evars evm c) in + let evm, c = Constr_denoter.denote_term env evm (to_constr_evars evm c) in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS evm) (ltac_apply tac (List.map to_ltac_val [EConstr.of_constr c])) end) } diff --git a/template-coq/src/quoter.ml b/template-coq/src/quoter.ml index fa924c280..eb610eb35 100644 --- a/template-coq/src/quoter.ml +++ b/template-coq/src/quoter.ml @@ -58,6 +58,8 @@ sig val mkProj : quoted_proj -> t -> t val mkFix : (quoted_int array * quoted_int) * (quoted_aname array * t array * t array) -> t val mkCoFix : quoted_int * (quoted_aname array * t array * t array) -> t + val mkInt : quoted_int63 -> t + val mkFloat : quoted_float64 -> t val mkBindAnn : quoted_name -> quoted_relevance -> quoted_aname val mkName : quoted_ident -> quoted_name @@ -76,6 +78,10 @@ sig val quote_inductive : quoted_kernel_name * quoted_int -> quoted_inductive val quote_proj : quoted_inductive -> quoted_int -> quoted_int -> quoted_proj + (* Primitive objects *) + val quote_int63 : Uint63.t -> quoted_int63 + val quote_float64 : Float64.t -> quoted_float64 + val quote_constraint_type : Univ.constraint_type -> quoted_constraint_type val quote_univ_constraint : Univ.univ_constraint -> quoted_univ_constraint val quote_univ_instance : Univ.Instance.t -> quoted_univ_instance @@ -227,7 +233,7 @@ struct | Constr.Const (c,pu) -> let kn = Constant.canonical c in - (Q.mkConst (Q.quote_kn kn) (Q.quote_univ_instance pu), add_constant kn acc) + (Q.mkConst (Q.quote_kn kn) (Q.quote_univ_instance pu), add_constant kn acc) | Constr.Construct ((mind,c),pu) -> (Q.mkConstruct (quote_inductive' mind, Q.quote_int (c - 1)) (Q.quote_univ_instance pu), @@ -261,9 +267,9 @@ struct let p' = Q.quote_proj ind pars arg in let t', acc = quote_term acc env c in (Q.mkProj p' t', add_inductive (Projection.inductive p) acc) + | Constr.Int i -> (Q.mkInt (Q.quote_int63 i), acc) + | Constr.Float f -> (Q.mkFloat (Q.quote_float64 f), acc) | Constr.Meta _ -> failwith "Meta not supported by TemplateCoq" - | Constr.Int _ -> failwith "Native integers not supported by TemplateCoq" - | Constr.Float _ -> failwith "Native floating point numbers not supported by TemplateCoq" in let in_prop, env' = env in if is_cast_prop () && not in_prop then diff --git a/template-coq/src/reification.ml b/template-coq/src/reification.ml index d6abf45e8..c38303c01 100644 --- a/template-coq/src/reification.ml +++ b/template-coq/src/reification.ml @@ -14,6 +14,9 @@ sig type quoted_kernel_name type quoted_inductive type quoted_proj + type quoted_int63 + type quoted_float64 + type quoted_global_reference type quoted_sort_family diff --git a/template-coq/src/run_extractable.ml b/template-coq/src/run_extractable.ml index 5f488eeb6..fa4460561 100644 --- a/template-coq/src/run_extractable.ml +++ b/template-coq/src/run_extractable.ml @@ -137,11 +137,13 @@ let of_cast_kind (ck: BasicAst.cast_kind) : Constr.cast_kind = | RevertCast -> Constr.REVERTcast (* todo(gmm): determine what of these already exist. *) -let rec to_constr_ev (evm : Evd.evar_map) (t : Ast0.term) : Evd.evar_map * Constr.t = - denote_term evm t +let rec to_constr_ev env (evm : Evd.evar_map) (t : Ast0.term) : Evd.evar_map * Constr.t = + denote_term env evm t let to_constr (t : Ast0.term) : Constr.t = - snd (to_constr_ev Evd.empty t) + let env = Global.env () in + let evm = Evd.from_env env in + snd (to_constr_ev env evm t) let tmOfConstr (t : Constr.t) : Ast0.term tm = Plugin_core.with_env_evm (fun env _ -> tmReturn (quote_term env t)) diff --git a/template-coq/src/run_template_monad.ml b/template-coq/src/run_template_monad.ml index dce8f79be..e5874195a 100644 --- a/template-coq/src/run_template_monad.ml +++ b/template-coq/src/run_template_monad.ml @@ -220,15 +220,15 @@ let denote_universes_entry evm trm (* of type universes_entry *) : _ * Entries.u | _ -> bad_term_verb trm "denote_universes_entry" -let unquote_one_inductive_entry evm trm (* of type one_inductive_entry *) : _ * Entries.one_inductive_entry = +let unquote_one_inductive_entry env evm trm (* of type one_inductive_entry *) : _ * Entries.one_inductive_entry = let (h,args) = app_full trm [] in if constr_equall h tBuild_one_inductive_entry then match args with | id::ar::cnames::ctms::[] -> let id = unquote_ident id in - let evm, ar = denote_term evm ar in + let evm, ar = denote_term env evm ar in let cnames = List.map unquote_ident (unquote_list cnames) in - let evm, ctms = map_evm denote_term evm (unquote_list ctms) in + let evm, ctms = map_evm (denote_term env) evm (unquote_list ctms) in evm, { mind_entry_typename = id ; mind_entry_arity = ar; mind_entry_consnames = cnames; @@ -242,32 +242,33 @@ let map_option f o = | Some x -> Some (f x) | None -> None -let denote_decl evm d = +let denote_decl env evm d = let (h, args) = app_full d [] in if constr_equall h tmkdecl then match args with | name :: body :: typ :: [] -> let name = unquote_aname name in - let evm, ty = denote_term evm typ in - (match unquote_option body with + let evm, ty = denote_term env evm typ in + let evm, decl = (match unquote_option body with | None -> evm, Context.Rel.Declaration.LocalAssum (name, ty) - | Some body -> let evm, body = denote_term evm body in + | Some body -> let evm, body = denote_term env evm body in evm, Context.Rel.Declaration.LocalDef (name, body, ty)) + in Environ.push_rel decl env, evm, decl | _ -> bad_term_verb d "denote_decl" else bad_term_verb d "denote_decl" -let denote_context evm ctx = - map_evm denote_decl evm (unquote_list ctx) +let denote_context env evm ctx = + fold_env_evm_right denote_decl env evm (unquote_list ctx) -let unquote_mutual_inductive_entry evm trm (* of type mutual_inductive_entry *) : _ * Entries.mutual_inductive_entry = +let unquote_mutual_inductive_entry env evm trm (* of type mutual_inductive_entry *) : _ * Entries.mutual_inductive_entry = let (h,args) = app_full trm [] in if constr_equall h tBuild_mutual_inductive_entry then match args with | record::finite::params::inds::univs::template::cumulative::priv::[] -> let record = map_option (map_option (fun x -> [|x|])) (unquote_map_option (unquote_map_option unquote_ident) record) in let finite = denote_mind_entry_finite finite in - let evm, params = denote_context evm params in - let evm, inds = map_evm unquote_one_inductive_entry evm (unquote_list inds) in + let envpars, evm, params = denote_context env evm params in + let evm, inds = map_evm (unquote_one_inductive_entry env) evm (unquote_list inds) in let evm, univs = denote_universes_entry evm univs in let template = unquote_bool template in let cumulative = unquote_bool cumulative in @@ -287,7 +288,7 @@ let unquote_mutual_inductive_entry evm trm (* of type mutual_inductive_entry *) let declare_inductive (env: Environ.env) (evm: Evd.evar_map) (mind: Constr.t) : unit = let mind = reduce_all env evm mind in - let evm, mind = unquote_mutual_inductive_entry evm mind in + let evm, mind = unquote_mutual_inductive_entry env evm mind in ignore (DeclareInd.declare_mutual_inductive_with_eliminations mind Names.Id.Map.empty []) let not_in_tactic s = @@ -334,19 +335,19 @@ let rec run_template_program_rec ~poly ?(intactic=false) (k : Environ.env * Evd. else let name = unquote_ident (reduce_all env evm name) in let opaque = unquote_bool (reduce_all env evm opaque) in - let evm,body = denote_term evm (reduce_all env evm body) in + let evm,body = denote_term env evm (reduce_all env evm body) in let evm,typ = match unquote_option typ with | None -> (evm, None) | Some t -> - let (evm, t) = denote_term evm (reduce_all env evm t) in + let (evm, t) = denote_term env evm (reduce_all env evm t) in (evm, Some t) in Plugin_core.run (Plugin_core.tmDefinition name ~poly ~opaque typ body) env evm (fun env evm res -> k (env, evm, quote_kn res)) | TmLemmaTerm (name, typ) -> let ident = unquote_ident (reduce_all env evm name) in - let evm,typ = denote_term evm (reduce_all env evm typ) in + let evm,typ = denote_term env evm (reduce_all env evm typ) in Plugin_core.run (Plugin_core.tmLemma ident ~poly typ) env evm (fun env evm kn -> k (env, evm, quote_kn kn)) @@ -366,7 +367,7 @@ let rec run_template_program_rec ~poly ?(intactic=false) (k : Environ.env * Evd. then not_in_tactic "tmAxiom" else let name = unquote_ident (reduce_all env evm name) in - let evm,typ = denote_term evm (reduce_all env evm typ) in + let evm,typ = denote_term env evm (reduce_all env evm typ) in Plugin_core.run (Plugin_core.tmAxiom name ~poly typ) env evm (fun a b c -> k (a,b,quote_kn c)) @@ -439,7 +440,7 @@ let rec run_template_program_rec ~poly ?(intactic=false) (k : Environ.env * Evd. in k (env, evm, trm) | TmEvalTerm (s,trm) -> let red = unquote_reduction_strategy env evm (reduce_all env evm s) in - let evm,trm = denote_term evm (reduce_all env evm trm) in + let evm,trm = denote_term env evm (reduce_all env evm trm) in Plugin_core.run (Plugin_core.tmEval red trm) env evm (fun env evm trm -> k (env, evm, quote_term env trm)) | TmMkInductive mind -> @@ -450,7 +451,7 @@ let rec run_template_program_rec ~poly ?(intactic=false) (k : Environ.env * Evd. begin try let t = reduce_all env evm t in - let evm, t' = denote_term evm t in + let evm, t' = denote_term env evm t in let typ = Retyping.get_type_of env evm (EConstr.of_constr t') in let evm, typ = Evarsolve.refresh_universes (Some false) env evm typ in let make_typed_term typ term evm = @@ -467,7 +468,7 @@ let rec run_template_program_rec ~poly ?(intactic=false) (k : Environ.env * Evd. with Reduction.NotArity -> CErrors.user_err (str "unquoting ill-typed term") end | TmUnquoteTyped (typ, t) -> - let evm, t' = denote_term evm (reduce_all env evm t) in + let evm, t' = denote_term env evm (reduce_all env evm t) in (* let t' = Typing.e_solve_evars env evdref (EConstr.of_constr t') in *) let evm = Typing.check env evm (EConstr.of_constr t') (EConstr.of_constr typ) in k (env, evm, t') @@ -496,7 +497,7 @@ let rec run_template_program_rec ~poly ?(intactic=false) (k : Environ.env * Evd. Not_found -> k (env, evm, constr_mkApp (cNone_instance, [|typ|])) end | TmInferInstanceTerm typ -> - let evm,typ = denote_term evm (reduce_all env evm typ) in + let evm,typ = denote_term env evm (reduce_all env evm typ) in Plugin_core.run (Plugin_core.tmInferInstance typ) env evm (fun env evm -> function None -> k (env, evm, constr_mkAppl (cNone, [| tTerm|])) @@ -504,6 +505,6 @@ let rec run_template_program_rec ~poly ?(intactic=false) (k : Environ.env * Evd. let qtrm = quote_term env trm in k (env, evm, constr_mkApp (cSome, [| Lazy.force tTerm; qtrm |]))) | TmPrintTerm trm -> - let evm,trm = denote_term evm (reduce_all env evm trm) in + let evm,trm = denote_term env evm (reduce_all env evm trm) in Plugin_core.run (Plugin_core.tmPrint trm) env evm (fun env evm _ -> k (env, evm, Lazy.force unit_tt)) diff --git a/template-coq/src/tm_util.ml b/template-coq/src/tm_util.ml index 676361a3f..b2cf92ef8 100644 --- a/template-coq/src/tm_util.ml +++ b/template-coq/src/tm_util.ml @@ -51,7 +51,7 @@ let not_supported trm = let not_supported_verb trm rs = let env = Global.env () in - CErrors.user_err (str "Not Supported raised at " ++ str rs ++ str ":" ++ spc () ++ + CErrors.user_err (str "Not Supported raised at " ++ str rs ++ str ":" ++ spc () ++ Printer.pr_constr_env env (Evd.from_env env) trm) let bad_term trm = @@ -68,7 +68,7 @@ type ('term, 'name, 'nat) adef = { adname : 'name; adtype : 'term; adbody : 'ter type ('term, 'name, 'nat) amfixpoint = ('term, 'name, 'nat) adef list -type ('term, 'nat, 'ident, 'name, 'quoted_sort, 'cast_kind, 'kername, 'inductive, 'relevance, 'universe_instance, 'projection) structure_of_term = +type ('term, 'nat, 'ident, 'name, 'quoted_sort, 'cast_kind, 'kername, 'inductive, 'relevance, 'universe_instance, 'projection, 'int63, 'float64) structure_of_term = | ACoq_tRel of 'nat | ACoq_tVar of 'ident | ACoq_tEvar of 'nat * 'term list @@ -85,4 +85,6 @@ type ('term, 'nat, 'ident, 'name, 'quoted_sort, 'cast_kind, 'kername, 'inductive | ACoq_tProj of 'projection * 'term | ACoq_tFix of ('term, 'name, 'nat) amfixpoint * 'nat | ACoq_tCoFix of ('term, 'name, 'nat) amfixpoint * 'nat + | ACoq_tInt of 'int63 + | ACoq_tFloat of 'float64 diff --git a/template-coq/theories/Ast.v b/template-coq/theories/Ast.v index d6fd10797..a54be64fa 100644 --- a/template-coq/theories/Ast.v +++ b/template-coq/theories/Ast.v @@ -1,7 +1,8 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Template Require Import utils Environment. From MetaCoq.Template Require Export Universes. - +(* For primitive integers and floats *) +From Coq Require Int63 Floats.PrimFloat. (** * AST of Coq kernel terms and kernel data structures @@ -50,7 +51,15 @@ Inductive term : Type := (discr:term) (branches : list (nat * term)) | tProj (proj : projection) (t : term) | tFix (mfix : mfixpoint term) (idx : nat) -| tCoFix (mfix : mfixpoint term) (idx : nat). +| tCoFix (mfix : mfixpoint term) (idx : nat) +| tInt (i : Int63.int) +| tFloat (f : PrimFloat.float). + +(** This can be used to represent holes, that, when unquoted, turn into fresh existential variables. + The fresh evar will depend on the whole context at this point in the term, despite the empty instance. + Denotation will call Coq's Typing.solve_evars to try and fill these holes using typing information. +*) +Definition hole := tEvar fresh_evar_id []. Definition mkApps t us := match us with @@ -61,7 +70,6 @@ Definition mkApps t us := end end. - Module TemplateTerm <: Term. Definition term := term. @@ -116,9 +124,12 @@ Inductive wf : term -> Prop := | wf_tConstruct i k u : wf (tConstruct i k u) | wf_tCase ci p c brs : wf p -> wf c -> Forall (wf ∘ snd) brs -> wf (tCase ci p c brs) | wf_tProj p t : wf t -> wf (tProj p t) -| wf_tFix mfix k : Forall (fun def => wf def.(dtype) /\ wf def.(dbody) /\ isLambda def.(dbody) = true) mfix -> +| wf_tFix mfix k : Forall (fun def => wf def.(dtype) /\ wf def.(dbody)) mfix -> wf (tFix mfix k) -| wf_tCoFix mfix k : Forall (fun def => wf def.(dtype) /\ wf def.(dbody)) mfix -> wf (tCoFix mfix k). +| wf_tCoFix mfix k : Forall (fun def => wf def.(dtype) /\ wf def.(dbody)) mfix -> wf (tCoFix mfix k) +| wf_tInt i : wf (tInt i) +| wf_tFloat f : wf (tFloat f). +Derive Signature for wf. (** ** Entries diff --git a/template-coq/theories/AstUtils.v b/template-coq/theories/AstUtils.v index 0608e2c3f..11b81d2e1 100644 --- a/template-coq/theories/AstUtils.v +++ b/template-coq/theories/AstUtils.v @@ -1,9 +1,19 @@ +(* For primitive integers and floats *) +From Coq Require Numbers.Cyclic.Int63.Int63 Floats.PrimFloat. (* Distributed under the terms of the MIT license. *) From MetaCoq.Template Require Import utils BasicAst Ast. Require Import ssreflect. +Require Import ZArith. (** Raw term printing *) +Definition string_of_prim_int (i:Int63.int) : string := + (* Better? DecimalString.NilZero.string_of_uint (BinNat.N.to_uint (BinInt.Z.to_N (Int63.to_Z i))). ? *) + string_of_Z (Numbers.Cyclic.Int63.Int63.to_Z i). + +Definition string_of_float (f : PrimFloat.float) := + "". + Fixpoint string_of_term (t : term) := match t with | tRel n => "Rel(" ^ string_of_nat n ^ ")" @@ -42,8 +52,18 @@ Fixpoint string_of_term (t : term) := ^ string_of_term c ^ ")" | tFix l n => "Fix(" ^ (string_of_list (string_of_def string_of_term) l) ^ "," ^ string_of_nat n ^ ")" | tCoFix l n => "CoFix(" ^ (string_of_list (string_of_def string_of_term) l) ^ "," ^ string_of_nat n ^ ")" + | tInt i => "Int(" ^ string_of_prim_int i ^ ")" + | tFloat f => "Float(" ^ string_of_float f ^ ")" end. - + +Fixpoint destArity Γ (t : term) := + match t with + | tProd na t b => destArity (Γ ,, vass na t) b + | tLetIn na b b_ty b' => destArity (Γ ,, vdef na b b_ty) b' + | tSort s => Some (Γ, s) + | _ => None + end. + Definition decompose_app (t : term) := match t with | tApp f l => (f, l) @@ -70,6 +90,10 @@ Proof. rewrite <- mkApps_nested. reflexivity. Qed. +Lemma mkAppMkApps s t: + mkApp s t = mkApps s [t]. +Proof. reflexivity. Qed. + Lemma mkApp_tApp f u : isApp f = false -> mkApp f u = tApp f [u]. Proof. intros. destruct f; (discriminate || reflexivity). Qed. @@ -141,11 +165,32 @@ apply (List.firstn decl.(ind_npars)) in types. (snd (fst x))) ind_ctors). Defined. +Fixpoint strip_casts t := + match t with + | tEvar ev args => tEvar ev (List.map strip_casts args) + | tLambda na T M => tLambda na (strip_casts T) (strip_casts M) + | tApp u v => mkApps (strip_casts u) (List.map (strip_casts) v) + | tProd na A B => tProd na (strip_casts A) (strip_casts B) + | tCast c kind t => strip_casts c + | tLetIn na b t b' => tLetIn na (strip_casts b) (strip_casts t) (strip_casts b') + | tCase ind p c brs => + let brs' := List.map (on_snd (strip_casts)) brs in + tCase ind (strip_casts p) (strip_casts c) brs' + | tProj p c => tProj p (strip_casts c) + | tFix mfix idx => + let mfix' := List.map (map_def strip_casts strip_casts) mfix in + tFix mfix' idx + | tCoFix mfix idx => + let mfix' := List.map (map_def strip_casts strip_casts) mfix in + tCoFix mfix' idx + | tRel _ | tVar _ | tSort _ | tConst _ _ | tInd _ _ | tConstruct _ _ _ + | tInt _ | tFloat _ => t + end. + Fixpoint decompose_prod_assum (Γ : context) (t : term) : context * term := match t with | tProd n A B => decompose_prod_assum (Γ ,, vass n A) B | tLetIn na b bty b' => decompose_prod_assum (Γ ,, vdef na b bty) b' - | tCast t _ _ => decompose_prod_assum Γ t | _ => (Γ, t) end. diff --git a/template-coq/theories/BasicAst.v b/template-coq/theories/BasicAst.v index e4993b2d6..e521dec0b 100644 --- a/template-coq/theories/BasicAst.v +++ b/template-coq/theories/BasicAst.v @@ -1,6 +1,6 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Template Require Import utils. - +From Coq Require Import Floats.SpecFloat. (** ** Reification of names ** *) @@ -28,6 +28,8 @@ From MetaCoq.Template Require Import utils. - constructor => [inductive * nat] - Projection.t => [projection] - GlobRef.t => global_reference + + Finally, we define the models of primitive types (uint63 and floats64). *) Definition ident := string. (* e.g. nat *) @@ -37,6 +39,8 @@ Definition qualid := string. (* e.g. Datatypes.nat *) order is reversed to improve sharing. E.g. A.B.C is ["C";"B";"A"] *) Definition dirpath := list ident. +Instance dirpath_eqdec : Classes.EqDec dirpath := _. + Definition string_of_dirpath : dirpath -> string := String.concat "." ∘ rev. @@ -49,6 +53,7 @@ Inductive modpath := | MPfile (dp : dirpath) | MPbound (dp : dirpath) (id : ident) (i : nat) | MPdot (mp : modpath) (id : ident). +Derive NoConfusion EqDec for modpath. Fixpoint string_of_modpath (mp : modpath) : string := match mp with @@ -59,17 +64,19 @@ Fixpoint string_of_modpath (mp : modpath) : string := (** The absolute names of objects seen by kernel *) Definition kername := modpath × ident. +Instance kername_eqdec : Classes.EqDec kername := _. Definition string_of_kername (kn : kername) := string_of_modpath kn.1 ^ "." ^ kn.2. - (** Identifiers that are allowed to be anonymous (i.e. "_" in concrete syntax). *) Inductive name : Set := | nAnon | nNamed (_ : ident). +Derive NoConfusion EqDec for name. Inductive relevance : Set := Relevant | Irrelevant. +Derive NoConfusion EqDec for relevance. (** Binders annotated with relevance *) Record binder_annot (A : Type) := mkBindAnn { binder_name : A; binder_relevance : relevance }. @@ -78,6 +85,11 @@ Arguments mkBindAnn {_}. Arguments binder_name {_}. Arguments binder_relevance {_}. +Derive NoConfusion for binder_annot. + +Instance eqdec_binder_annot (A : Type) (e : Classes.EqDec A) : Classes.EqDec (binder_annot A). +Proof. ltac:(Equations.Prop.Tactics.eqdec_proof). Qed. + Definition map_binder_annot {A B} (f : A -> B) (b : binder_annot A) : binder_annot B := {| binder_name := f b.(binder_name); binder_relevance := b.(binder_relevance) |}. @@ -86,6 +98,7 @@ Definition eq_binder_annot {A} (b b' : binder_annot A) : Prop := (** Type of annotated names *) Definition aname := binder_annot name. +Instance anqme_eqdec : Classes.EqDec aname := _. Definition string_of_name (na : name) := match na with @@ -104,6 +117,8 @@ Record inductive : Set := mkInd { inductive_mind : kername ; inductive_ind : nat }. Arguments mkInd _%string _%nat. +Derive NoConfusion EqDec for inductive. + Definition string_of_inductive (i : inductive) := string_of_kername (inductive_mind i) ^ "," ^ string_of_nat (inductive_ind i). @@ -116,6 +131,8 @@ Inductive global_reference := | IndRef : inductive -> global_reference | ConstructRef : inductive -> nat -> global_reference. +Derive NoConfusion EqDec for global_reference. + Definition string_of_gref gr : string := match gr with @@ -127,20 +144,10 @@ Definition string_of_gref gr : string := "Constructor " ^ string_of_kername s ^ " " ^ (string_of_nat n) ^ " " ^ (string_of_nat k) end. -Definition kername_eq_dec (k k0 : kername) : {k = k0} + {k <> k0}. -Proof. - repeat (decide equality; eauto with eq_dec). -Defined. +Definition kername_eq_dec (k k0 : kername) : {k = k0} + {k <> k0} := Classes.eq_dec k k0. Hint Resolve kername_eq_dec : eq_dec. -Definition gref_eq_dec (gr gr' : global_reference) : {gr = gr'} + {~ gr = gr'}. -Proof. - decide equality; eauto with eq_dec. - destruct i, i0. - decide equality; eauto with eq_dec. - destruct i, i0. - decide equality; eauto with eq_dec. -Defined. +Definition gref_eq_dec (gr gr' : global_reference) : {gr = gr'} + {~ gr = gr'} := Classes.eq_dec gr gr'. Definition ident_eq (x y : ident) := match string_compare x y with @@ -194,21 +201,29 @@ Proof. now rewrite eq_inductive_refl, !PeanoNat.Nat.eqb_refl. Qed. - - (** The kind of a cast *) Inductive cast_kind : Set := | VmCast | NativeCast | Cast | RevertCast. +Derive NoConfusion EqDec for cast_kind. Inductive recursivity_kind := | Finite (* = inductive *) | CoFinite (* = coinductive *) | BiFinite (* = non-recursive, like in "Record" definitions *). +Derive NoConfusion EqDec for recursivity_kind. +(* The kind of a conversion problem *) +Inductive conv_pb := + | Conv + | Cumul. +Derive NoConfusion EqDec for conv_pb. +(* This opaque natural number is a hack to inform unquoting to generate + a fresh evar when encountering it. *) +Definition fresh_evar_id : nat. exact 0. Qed. (* Parametrized by term because term is not yet defined *) Record def term := mkdef { @@ -222,6 +237,10 @@ Arguments dtype {term} _. Arguments dbody {term} _. Arguments rarg {term} _. +Derive NoConfusion for def. +Instance def_eq_dec {A} : Classes.EqDec A -> Classes.EqDec (def A). +Proof. ltac:(Equations.Prop.Tactics.eqdec_proof). Qed. + Definition string_of_def {A} (f : A -> string) (def : def A) := "(" ^ string_of_name (binder_name (dname def)) ^ "," ^ string_of_relevance (binder_relevance (dname def)) @@ -247,7 +266,6 @@ Proof. destruct d; reflexivity. Qed. Definition mfixpoint term := list (def term). - Definition test_def {A} (tyf bodyf : A -> bool) (d : def A) := tyf d.(dtype) && bodyf d.(dbody). @@ -367,3 +385,21 @@ Ltac close_All := | H : All2 _ _ _ |- All _ _ => (apply (All2_All_left H) || apply (All2_All_right H)); clear H; simpl end. + +(** Primitive types models (axiom free) *) + +(** Model of unsigned integers *) +Definition uint_size := 63. +Definition uint_wB := (2 ^ (Z.of_nat uint_size))%Z. +Definition uint63_model := { z : Z | ((0 <=? z) && (z ". diff --git a/checker/theories/Checker.v b/template-coq/theories/Checker.v similarity index 76% rename from checker/theories/Checker.v rename to template-coq/theories/Checker.v index 984cffa87..7de4536e3 100644 --- a/checker/theories/Checker.v +++ b/template-coq/theories/Checker.v @@ -6,9 +6,7 @@ From MetaCoq.Template Require Import config Ast AstUtils utils Implemets [typecheck_program] which returns an error and on success should guarantee that the term has the given type. - Currently uses fuel to implement reduction. - - Correctness and completeness proofs are a work-in-progress. + Currently uses fuel to implement reduction and is unverified. This file implements reduction with a stack machine [reduce_stack], conversion/cumulativity with the first-order fast-path heuristic [isconv] @@ -447,7 +445,8 @@ Inductive type_error := | IllFormedFix (m : mfixpoint term) (i : nat) | UnsatisfiedConstraints (c : ConstraintSet.t) | UnsatisfiableConstraints (c : ConstraintSet.t) -| NotEnoughFuel (n : nat). +| NotEnoughFuel (n : nat) +| NotSupported (s : string). Definition string_of_type_error (e : type_error) : string := match e with @@ -468,6 +467,7 @@ Definition string_of_type_error (e : type_error) : string := | UnsatisfiedConstraints c => "Unsatisfied constraints" | UnsatisfiableConstraints c => "Unsatisfiable constraints" | NotEnoughFuel n => "Not enough fuel" + | NotSupported s => s ^ " are not supported" end. Inductive typing_result (A : Type) := @@ -576,7 +576,7 @@ Section Typecheck. end. End Typecheck. -Section Typecheck2. +Section Typecheck. Context {cf : checker_flags} {F : Fuel}. Context (Σ : global_env) (G : universes_graph). @@ -769,6 +769,8 @@ Section Typecheck2. | Some f => ret f.(dtype) | None => raise (IllFormedFix mfix n) end + + | tInt _ | tFloat _ => raise (NotSupported "primitive types") end. Definition check (Γ : context) (t : term) (ty : term) : typing_result unit := @@ -782,258 +784,11 @@ Section Typecheck2. | TypeError _ => false end. -End Typecheck2. - -Lemma cumul_convert_leq : forall `{checker_flags} {F:Fuel} Σ G Γ t t', - is_graph_of_global_env_ext Σ G -> - Σ ;;; Γ |- t <= t' <~> convert_leq (fst Σ) G Γ t t' = Checked (). -Proof. intros. todo "Checker.cumul_convert_leq". Defined. - -Lemma cumul_reduce_to_sort : forall `{checker_flags} {F:Fuel} Σ G Γ t s', - is_graph_of_global_env_ext Σ G -> - Σ ;;; Γ |- t <= tSort s' <~> - { s'' & reduce_to_sort (fst Σ) Γ t = Checked s'' - /\ check_leqb_universe G s'' s' = true }. -Proof. intros. todo "Checker.cumul_reduce_to_sort". Defined. - -Lemma cumul_reduce_to_product : forall `{checker_flags} {F:Fuel} Σ Γ t na a b, - Σ ;;; Γ |- t <= tProd na a b -> - { a' & { b' & - ((reduce_to_prod (fst Σ) Γ t = Checked (a', b')) * - cumul Σ Γ (tProd na a' b') (tProd na a b))%type } }. -Proof. intros. todo "Checker.cumul_reduce_to_product". Defined. - -Lemma cumul_reduce_to_ind : forall `{checker_flags} {F:Fuel} Σ Γ t i u args, - Σ ;;; Γ |- t <= mkApps (tInd i u) args <~> - { args' & - ((reduce_to_ind (fst Σ) Γ t = Checked (i, u, args')) * - cumul Σ Γ (mkApps (tInd i u) args') (mkApps (tInd i u) args))%type }. -Proof. intros. todo "Checker.cumul_reduce_to_ind". Defined. - -Lemma lookup_constant_type_declared Σ cst u decl (isdecl : declared_constant Σ cst decl) : - lookup_constant_type_cstrs Σ cst u = - Checked (subst_instance_constr u decl.(cst_type), - subst_instance_cstrs u (polymorphic_constraints decl.(cst_universes))). -Proof. - unfold lookup_constant_type_cstrs, lookup_env. - red in isdecl. rewrite isdecl. destruct decl. reflexivity. -Qed. - -Lemma lookup_constant_type_is_declared Σ cst u T : - lookup_constant_type_cstrs Σ cst u = Checked T -> - { decl | declared_constant Σ cst decl /\ - subst_instance_constr u decl.(cst_type) = fst T }. -Proof. - unfold lookup_constant_type_cstrs, lookup_env, declared_constant. - destruct Ast.lookup_env eqn:Hlook; try discriminate. - destruct g eqn:Hg; intros; try discriminate. destruct c. - injection H as eq. subst T. simpl. - eexists. split; eauto. -Qed. - +End Typecheck. Arguments bind _ _ _ _ ! _. Open Scope monad. - -Section InferOk. - Context {cf : checker_flags} {F : Fuel}. - Context (Σ : global_env_ext) (G : universes_graph) - (HG : is_graph_of_global_env_ext Σ G). - - Ltac tc := eauto with typecheck. - -(* - Lemma infer_complete Γ t T : - Σ ;;; Γ |- t : T -> { T' & ((infer Γ t = Checked T') /\ squash (cumul Σ Γ T' T)) }. - Proof. - induction 1; unfold infer_type, infer_cumul in *; simpl; unfold infer_type, infer_cumul in *; - repeat match goal with - H : { T' & _ } |- _ => destruct H as [? [-> H]] - end; simpl; try (eexists; split; [ reflexivity | solve [ tc ] ]). - - - eexists. rewrite e. - split; [ reflexivity | constructor; tc ]. - - - eexists. split; [reflexivity | tc]. - constructor. simpl. unfold leq_universe. - admit. - - - eexists. - apply cumul_reduce_to_sort in IHX1 as [s'' [-> Hs'']]. - apply cumul_convert_leq in IHX2 as ->. - simpl. split; tc. - - - apply cumul_reduce_to_sort in IHX1 as [s'' [-> Hs'']]. - apply cumul_reduce_to_sort in IHX2 as [s''' [-> Hs''']]. - simpl. eexists; split; tc. - admit. - - - eexists. apply cumul_reduce_to_sort in IHX1 as [s'' [-> Hs'']]. - split. reflexivity. - apply congr_cumul_prod; tc. - - - apply cumul_convert_leq in IHX2 as ->; simpl. - eexists ; split; tc. - admit. - - - admit. - - - erewrite lookup_constant_type_declared; eauto. - eexists ; split; [ try reflexivity | tc ]. - simpl. unfold consistent_universe_context_instance in c. - destruct cst_universes. simpl. reflexivity. - simpl in *. destruct cst0. simpl in *. - destruct c. unfold check_consistent_constraints. rewrite H0. reflexivity. - admit. - - - admit. - - admit. - - - (* destruct indpar. *) - apply cumul_reduce_to_ind in IHX2 as [args' [-> Hcumul]]. - simpl in *. rewrite (proj2 (eq_ind_refl ind ind) eq_refl). - eexists ; split; [ reflexivity | tc ]. - admit. - - - admit. - - - eexists. rewrite e. - split; [ reflexivity | tc ]. - - - eexists. rewrite e. - split; [ reflexivity | tc ]. - - - eexists. - split; [ reflexivity | tc ]. - eapply cumul_trans; eauto. - Admitted. - *) - Ltac infers := - repeat - match goal with - | |- context [infer ?Γ ?t] => - destruct (infer Γ t) eqn:?; [ simpl | simpl; intro; discriminate ] - | |- context [infer_type ?Γ ?t] => - destruct (infer_type Γ t) eqn:?; [ simpl | simpl; intro; discriminate ] - | |- context [infer_cumul ?Γ ?t ?t2] => - destruct (infer_cumul Γ t t2) eqn:?; [ simpl | simpl; intro; discriminate ] - | |- context [convert_leq ?Γ ?t ?t2] => - destruct (convert_leq Γ t t2) eqn:?; [ simpl | simpl; intro; discriminate ] - end; try intros [= <-]. - - - Lemma infer_type_correct Γ t x : - (forall (Γ : context) (T : term), infer (fst Σ) G Γ t = Checked T -> Σ ;;; Γ |- t : T) -> - infer_type (fst Σ) (infer (fst Σ) G) Γ t = Checked x -> - Σ ;;; Γ |- t : tSort x. - Proof. - intros IH H. - unfold infer_type in H. - revert H; infers. - specialize (IH _ _ Heqt0). - intros. - eapply type_Conv. apply IH. - admit. - eapply cumul_reduce_to_sort. eassumption. - exists x. split; tc. - Admitted. - - - Lemma infer_cumul_correct Γ t u x x' : - (forall (Γ : context) (T : term), infer (fst Σ) G Γ u = Checked T -> Σ ;;; Γ |- u : T) -> - (forall (Γ : context) (T : term), infer (fst Σ) G Γ t = Checked T -> Σ ;;; Γ |- t : T) -> - infer_type (fst Σ) (infer (fst Σ) G) Γ u = Checked x' -> - infer_cumul (fst Σ) G (infer (fst Σ) G) Γ t u = Checked x -> - Σ ;;; Γ |- t : u. - Proof. - intros IH IH' H H'. - unfold infer_cumul in H'. - revert H'; infers. - specialize (IH' _ _ Heqt0). - intros. - eapply type_Conv. apply IH'. - right. apply infer_type_correct in H; eauto. - destruct a0. eapply cumul_convert_leq; eassumption. - Qed. - - Ltac infco := eauto using infer_cumul_correct, infer_type_correct. - - (* Axiom cheat : forall A, A. *) - (* Ltac admit := apply cheat. *) - - Lemma infer_correct Γ t T : wf_local Σ Γ -> - infer (fst Σ) G Γ t = Checked T -> Σ ;;; Γ |- t : T. - Proof. - (* induction t in Γ, T |- * ; simpl; intros; try discriminate; *) - (* revert H; infers; try solve [econstructor; infco]. *) - - (* - destruct nth_error eqn:Heq; try discriminate. *) - (* intros [= <-]. constructor; auto. *) - - (* - admit. *) - (* - admit. *) - - (* - admit. *) - (* intros. *) - (* destruct (lookup_constant_type) eqn:?. simpl in *. *) - (* apply (lookup_constant_type_is_declared k u) in Heqt. *) - (* destruct Heqt as [decl [Hdecl Heq]]. *) - (* destruct a eqn:eqa. simpl in *. *) - (* destruct check_consistent_constraints eqn:cons. *) - - (* simpl in *. injection H as ->. rewrite <- Heq. constructor. auto. *) - (* red in Hdecl. *) - (* unfold consistent_universe_context_instance. *) - (* unfold check_consistent_constraints in cons. *) - (* unfold check_constraints in cons. *) - (* destruct decl. simpl in *. *) - - (* destruct decl; simpl in *. destruct cst_universes; simpl in *. auto. *) - (* destruct cst. simpl. unfold check_consistent_constraints in cons. split; auto. *) - (* unfold lookup_constant_type in Heqt. *) - - (* pose (lookup_constant_type_is_declared k u). _ _ _ H) as [decl [Hdecl <-]]. *) - (* constructor. auto. *) - - (* - (* Ind *) admit. *) - - (* - (* Construct *) admit. *) - - (* - (* Case *) *) - (* (* match goal with H : inductive * nat |- _ => destruct H end. *) *) - (* (* infers. *) *) - (* (* destruct reduce_to_ind eqn:?; try discriminate. simpl. *) *) - (* (* destruct a0 as [[ind' u] args]. *) *) - (* (* destruct eq_ind eqn:?; try discriminate. *) *) - (* (* intros [= <-]. *) *) - (* admit. *) - (* (* eapply type_Case. simpl in *. *) *) - (* (* eapply type_Conv. eauto. *) *) - (* (* admit. *) *) - (* (* rewrite cumul_reduce_to_ind. *) *) - (* (* exists args. split; auto. *) *) - (* (* rewrite Heqt0. repeat f_equal. apply eq_ind_refl in Heqb. congruence. *) *) - (* (* tc. *) *) - - (* - (* Proj *) admit. *) - (* - admit. *) - (* - admit. *) - (* - admit. *) - (* - admit. *) - - (* - destruct nth_error eqn:?; intros [= <-]. *) - (* constructor; auto. admit. admit. *) - - (* - destruct nth_error eqn:?; intros [= <-]. *) - (* constructor; auto. admit. admit. *) - Admitted. - -End InferOk. - -Extract Constant infer_type_correct => "(fun f sigma ctx t x -> assert false)". -Extract Constant infer_correct => "(fun f sigma ctx t ty -> assert false)". - Definition default_fuel : Fuel := Nat.pow 2 14. Fixpoint fresh id (env : global_env) : bool := @@ -1154,3 +909,4 @@ Definition infer' `{checker_flags} `{Fuel} (Σ : global_env_ext) Γ t | None => raise (UnsatisfiableConstraints uctx.2) | Some uctx => infer (fst Σ) (make_graph uctx) Γ t end. + diff --git a/template-coq/theories/Constants.v b/template-coq/theories/Constants.v index def487a2c..989b102cf 100644 --- a/template-coq/theories/Constants.v +++ b/template-coq/theories/Constants.v @@ -87,6 +87,7 @@ Register MetaCoq.Template.BasicAst.VarRef as metacoq.ast.VarRef. Register MetaCoq.Template.BasicAst.ConstRef as metacoq.ast.ConstRef. Register MetaCoq.Template.BasicAst.IndRef as metacoq.ast.IndRef. Register MetaCoq.Template.BasicAst.ConstructRef as metacoq.ast.ConstructRef. +Register MetaCoq.Template.BasicAst.fresh_evar_id as metacoq.ast.fresh_evar_id. (* Universes *) @@ -166,6 +167,8 @@ Register MetaCoq.Template.Ast.tCase as metacoq.ast.tCase. Register MetaCoq.Template.Ast.tProj as metacoq.ast.tProj. Register MetaCoq.Template.Ast.tFix as metacoq.ast.tFix. Register MetaCoq.Template.Ast.tCoFix as metacoq.ast.tCoFix. +Register MetaCoq.Template.Ast.tInt as metacoq.ast.tInt. +Register MetaCoq.Template.Ast.tFloat as metacoq.ast.tFloat. (* Local and global declarations *) Register MetaCoq.Template.Ast.parameter_entry as metacoq.ast.parameter_entry. diff --git a/template-coq/theories/EnvironmentTyping.v b/template-coq/theories/EnvironmentTyping.v index 080e64778..ae3cc946f 100644 --- a/template-coq/theories/EnvironmentTyping.v +++ b/template-coq/theories/EnvironmentTyping.v @@ -123,7 +123,6 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T). Import T E. Section TypeLocal. - Context (typing : forall (Γ : context), term -> option term -> Type). Inductive All_local_env : context -> Type := @@ -140,7 +139,7 @@ Module EnvTyping (T : Term) (E : EnvironmentSig T). typing Γ t None -> typing Γ b (Some t) -> All_local_env (Γ ,, vdef na b t). - + Derive Signature for All_local_env. End TypeLocal. Arguments localenv_nil {_}. @@ -289,8 +288,6 @@ Module Type Typing (T : Term) (E : EnvironmentSig T) (ET : EnvTypingSig T E). Import T E ET. - Parameter (ind_guard : mutual_inductive_body -> bool). - Parameter (conv : forall `{checker_flags}, global_env_ext -> context -> term -> term -> Type). Parameter (cumul : forall `{checker_flags}, global_env_ext -> context -> term -> term -> Type). @@ -819,7 +816,6 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) onParams : on_context Σ mdecl.(ind_params); onNpars : context_assumptions mdecl.(ind_params) = mdecl.(ind_npars); onVariance : on_variance mdecl.(ind_universes) mdecl.(ind_variance); - onGuard : ind_guard mdecl }. (** *** Typing of constant declarations *) @@ -905,7 +901,6 @@ Module DeclarationTyping (T : Term) (E : EnvironmentSig T) Arguments onParams {_ P Σ mind mdecl}. Arguments onNpars {_ P Σ mind mdecl}. Arguments onVariance {_ P Σ mind mdecl}. - Arguments onGuard {_ P Σ mind mdecl}. Lemma All_local_env_impl (P Q : context -> term -> option term -> Type) l : All_local_env P l -> diff --git a/template-coq/theories/Extraction.v b/template-coq/theories/Extraction.v index d7696f898..3bc31f816 100644 --- a/template-coq/theories/Extraction.v +++ b/template-coq/theories/Extraction.v @@ -1,7 +1,7 @@ (* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import utils Ast Reflect Induction. -Require Import FSets ExtrOcamlBasic ExtrOcamlString ExtrOcamlZInt. - +From MetaCoq.Template Require Import utils Ast Reflect Induction MC_ExtrOCamlInt63 (*b/c nameclash with `comparion` *). +From Coq Require Import FSets ExtrOcamlBasic ExtrOcamlString ExtrOcamlZInt ExtrOCamlFloats. +From Coq Require Extraction. (** * Extraction setup for template-coq. Any extracted code planning to link with the plugin's OCaml reifier @@ -22,7 +22,7 @@ Extract Inductive Equations.Init.sigma => "( * )" ["(,)"]. Extract Constant Equations.Init.pr1 => "fst". Extract Constant Equations.Init.pr2 => "snd". Extraction Inline Equations.Init.pr1 Equations.Init.pr2. - + Extraction Blacklist Classes config uGraph Universes Ast String List Nat Int UnivSubst Typing Checker Retyping OrderedType Logic Common Equality UnivSubst. Set Warnings "-extraction-opaque-accessed". @@ -32,14 +32,29 @@ Cd "gen-src". From MetaCoq.Template Require Import TemplateMonad.Extractable config Induction LiftSubst UnivSubst Pretty. +Import Init.Nat. +Locate Nat. + +(* Floats *) +(* Extraction Library Zeven. +Extraction Library Zeven. +Extraction Library ZArith_dec. +Extraction Library Sumbool. +Extraction Library Zbool. +Extraction Library SpecFloat. *) +Separate Extraction FloatOps.Prim2SF. Recursive Extraction Library Extractable. +Extraction Library MCPrelude. Extraction Library MCOption. Extraction Library MCUtils. +Extraction Library MCList. Extraction Library EqDecInstances. Extraction Library Induction. Extraction Library LiftSubst. Extraction Library UnivSubst. +Extraction Library BasicAst. + Extraction Library Reflect. Extraction Library Pretty. Extraction Library config. diff --git a/template-coq/theories/Induction.v b/template-coq/theories/Induction.v index 6b9cecb42..7182c9f1e 100644 --- a/template-coq/theories/Induction.v +++ b/template-coq/theories/Induction.v @@ -30,6 +30,8 @@ Lemma term_forall_list_ind : (forall (s : projection) (t : term), P t -> P (tProj s t)) -> (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tCoFix m n)) -> + (forall i, P (tInt i)) -> + (forall f, P (tFloat f)) -> forall t : term, P t. Proof. intros until t. revert t. @@ -85,47 +87,48 @@ Lemma term_wf_forall_list_ind : P t -> forall t0 : term, P t0 -> forall l : list (nat * term), tCaseBrsProp P l -> P (tCase p t t0 l)) -> (forall (s : projection) (t : term), P t -> P (tProj s t)) -> - (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> Forall (fun def => isLambda (dbody def) = true) m -> P (tFix m n)) -> + (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), tFixProp P P m -> P (tCoFix m n)) -> + (forall i, P (tInt i)) -> + (forall f, P (tFloat f)) -> forall t : term, wf t -> P t. Proof. - pose proof I as H1. (* can go away, to avoid renaming everything... *) - intros until t. revert t. + pose proof I as H1. (* can go away, to avoid renaming everything... *) + intros until t. revert t. rename H16 into H16'. apply (term_forall_list_ind (fun t => wf t -> P t)); intros; try solve [match goal with H : _ |- _ => apply H end; auto]. - apply H2. inv H17. + apply H2. inv H18. auto using lift_to_wf_list. - - inv H18; auto. - - inv H18; auto. - - inv H18; auto. - inv H19; auto. - - inv H18; auto. + - inv H19; auto. + - inv H19; auto. + - inv H20; auto. + - inv H19; auto. apply H8; auto. auto using lift_to_wf_list. - - inv H18; apply H12; auto. + - inv H19; apply H12; auto. red. red in X. induction X. + constructor. - + constructor. inv H21; auto. apply IHX. inv H21; auto. + + constructor. inv H22; auto. apply IHX. inv H22; auto. - - inv H17; auto. + - inv H18; auto. - inv H16; auto. apply H14. red. red in X. induction X; constructor. - + split; inv H17; intuition. - + apply IHX. now inv H17. - + eapply Forall_impl; tea. clear; intros; cbn in *; intuition. + + split; inv H18; intuition. + + apply IHX. now inv H18. - inv H16; auto. apply H15. red. red in X. induction X; constructor. - + split; inv H17; intuition. - + apply IHX. now inv H17. + + split; inv H18; intuition. + + apply IHX. now inv H18. Qed. Definition tCaseBrsType {A} (P : A -> Type) (l : list (nat * A)) := @@ -155,6 +158,8 @@ Lemma term_forall_list_rect : (forall (s : projection) (t : term), P t -> P (tProj s t)) -> (forall (m : mfixpoint term) (n : nat), tFixType P P m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), tFixType P P m -> P (tCoFix m n)) -> + (forall i, P (tInt i)) -> + (forall f, P (tFloat f)) -> forall t : term, P t. Proof. intros until t. revert t. diff --git a/template-coq/theories/LiftSubst.v b/template-coq/theories/LiftSubst.v index 07025fb88..3633858fe 100644 --- a/template-coq/theories/LiftSubst.v +++ b/template-coq/theories/LiftSubst.v @@ -1,5 +1,7 @@ (* Distributed under the terms of the MIT license. *) -From MetaCoq Require Import utils Ast AstUtils Induction. +From MetaCoq.Template Require Import utils Ast AstUtils Induction. +From Coq Require Import ssreflect. +From Equations Require Import Equations. (** * Lifting and substitution for the AST @@ -41,6 +43,13 @@ Definition lift_decl n k d := (map_decl (lift n k) d). Definition lift_context n k (Γ : context) : context := fold_context (fun k' => lift n (k' + k)) Γ. +Lemma lift_context_alt n k Γ : + lift_context n k Γ = + mapi (fun k' d => lift_decl n (Nat.pred #|Γ| - k' + k) d) Γ. +Proof. + unfold lift_context. apply fold_context_alt. +Qed. + (** Parallel substitution: it assumes that all terms in the substitution live in the same context *) @@ -83,15 +92,37 @@ Notation "M { j := N }" := (subst1 N j M) (at level 10, right associativity). Definition subst_context s k (Γ : context) : context := fold_context (fun k' => subst s (k' + k)) Γ. +Definition subst_decl s k (d : context_decl) := map_decl (subst s k) d. + Lemma subst_context_length s n Γ : #|subst_context s n Γ| = #|Γ|. Proof. induction Γ as [|[na [body|] ty] tl] in Γ |- *; cbn; eauto. - - rewrite !List.rev_length, !mapi_rec_length, !app_length, !List.rev_length. simpl. + - rewrite !List.rev_length !mapi_rec_length !app_length !List.rev_length. simpl. lia. - - rewrite !List.rev_length, !mapi_rec_length, !app_length, !List.rev_length. simpl. + - rewrite !List.rev_length !mapi_rec_length !app_length !List.rev_length. simpl. lia. Qed. +Lemma subst_context_nil s n : subst_context s n [] = []. +Proof. reflexivity. Qed. + +Lemma subst_context_alt s k Γ : + subst_context s k Γ = + mapi (fun k' d => subst_decl s (Nat.pred #|Γ| - k' + k) d) Γ. +Proof. + unfold subst_context, fold_context. rewrite rev_mapi. rewrite List.rev_involutive. + apply mapi_ext. intros. f_equal. now rewrite List.rev_length. +Qed. + +Lemma subst_context_snoc s k Γ d : subst_context s k (d :: Γ) = subst_context s k Γ ,, subst_decl s (#|Γ| + k) d. +Proof. + unfold subst_context, fold_context. + rewrite !rev_mapi !rev_involutive /mapi mapi_rec_eqn /snoc. + f_equal. now rewrite Nat.sub_0_r List.rev_length. + rewrite mapi_rec_Sk. simpl. apply mapi_rec_ext. intros. + rewrite app_length !List.rev_length. simpl. f_equal. f_equal. lia. +Qed. + Definition subst_telescope s k (Γ : context) : context := mapi (fun k' decl => map_decl (subst s (k' + k)) decl) Γ. @@ -120,7 +151,7 @@ Notation closed t := (closedn 0 t). Fixpoint noccur_between k n (t : term) : bool := match t with - | tRel i => Nat.leb k i && Nat.ltb i (k + n) + | tRel i => Nat.ltb i k && Nat.leb (k + n) i | tEvar ev args => List.forallb (noccur_between k n) args | tLambda _ T M | tProd _ T M => noccur_between k n T && noccur_between (S k) n M | tApp u v => noccur_between k n u && List.forallb (noccur_between k n) v @@ -241,7 +272,7 @@ Lemma subst_rel_eq : subst u n (tRel p) = lift0 n t. Proof. intros; simpl in |- *. subst p. - elim (leb_spec n (n + i)). intros. assert (n + i - n = i) by lia. rewrite H1, H. + elim (leb_spec n (n + i)). intros. assert (n + i - n = i) by lia. rewrite H1 H. reflexivity. intros. lia. Qed. @@ -614,7 +645,7 @@ Proof. - apply wf_mkApps; auto. apply Forall_map. eapply Forall_impl; eauto. - apply Forall_map. apply All_Forall. eapply All_impl; tea. intros [] XX; cbn in *; apply XX. - - solve_all. induction dbody; try discriminate. reflexivity. + - solve_all. - apply Forall_map. eapply All_Forall, All_impl; eauto. intros [] XX; cbn in *; split; apply XX. Qed. @@ -644,7 +675,7 @@ Lemma lift_to_extended_list_k Γ k : forall k', to_extended_list_k Γ (k' + k) = map (lift0 k') (to_extended_list_k Γ k). Proof. unfold to_extended_list_k. - intros k'. rewrite !reln_alt_eq, !app_nil_r. + intros k'. rewrite !reln_alt_eq !app_nil_r. induction Γ in k, k' |- *; simpl; auto. destruct a as [na [body|] ty]. now rewrite <- Nat.add_assoc, (IHΓ (k + 1) k'). @@ -702,12 +733,61 @@ Lemma map_vass_map_def g l n k : (map (map_def (lift n k) g) l)) = (mapi (fun i d => map_decl (lift n (i + k)) d) (mapi (fun i (d : def term) => vass (dname d) (lift0 i (dtype d))) l)). Proof. - rewrite mapi_mapi, mapi_map. apply mapi_ext. + rewrite mapi_mapi mapi_map. apply mapi_ext. intros. unfold map_decl, vass; simpl; f_equal. - rewrite permute_lift. f_equal; lia. lia. + rewrite permute_lift. lia. f_equal; lia. Qed. (* Lemma noccur_between_subst k n t : noccur_between k n t -> closedn (n + k) t -> closedn k t. Proof. Admitted. *) + +Lemma strip_casts_lift n k t : + strip_casts (lift n k t) = lift n k (strip_casts t). +Proof. + induction t in k |- * using term_forall_list_ind; simpl; auto; + rewrite ?map_map_compose ?compose_on_snd ?compose_map_def ?map_length; + f_equal; solve_all; eauto. + + - rewrite lift_mkApps IHt map_map_compose. + f_equal; solve_all. +Qed. +Lemma mkApps_ex t u l : ∑ f args, Ast.mkApps t (u :: l) = Ast.tApp f args. +Proof. + induction t; simpl; eexists _, _; reflexivity. +Qed. +(* +Lemma mkApps_tApp' f l l' : mkApps (tApp f l) l' = mkApps f (l ++ l'). +Proof. + induction l'; simpl. rewrite app_nil_r. *) + +Lemma list_length_ind {A} (P : list A -> Type) (p0 : P []) + (pS : forall d Γ, (forall Γ', #|Γ'| <= #|Γ| -> P Γ') -> P (d :: Γ)) + Γ : P Γ. +Proof. + generalize (le_n #|Γ|). + generalize #|Γ| at 2. + induction n in Γ |- *. + destruct Γ; [|simpl; intros; elimtype False; lia]. + intros. apply p0. + intros. + destruct Γ; simpl in *. + apply p0. apply pS. intros. apply IHn. simpl. lia. +Qed. + +Lemma strip_casts_mkApps_tApp f l : + isApp f = false -> + strip_casts (mkApps f l) = strip_casts (tApp f l). +Proof. + induction l. simpl; auto. + intros. + rewrite mkApps_tApp //. +Qed. + +Lemma strip_casts_mkApps f l : + isApp f = false -> + strip_casts (mkApps f l) = mkApps (strip_casts f) (map strip_casts l). +Proof. + intros Hf. rewrite strip_casts_mkApps_tApp //. +Qed. \ No newline at end of file diff --git a/checker/theories/Normal.v b/template-coq/theories/Normal.v similarity index 100% rename from checker/theories/Normal.v rename to template-coq/theories/Normal.v diff --git a/template-coq/theories/Pretty.v b/template-coq/theories/Pretty.v index 5a50a483a..893d593cb 100644 --- a/template-coq/theories/Pretty.v +++ b/template-coq/theories/Pretty.v @@ -67,6 +67,7 @@ Section print_term. | Some body => substring 0 1 (body.(ind_name)) | None => "X" end + | tInt _ => "i" | _ => "U" end. @@ -201,6 +202,8 @@ Section print_term. | tCoFix l n => parens top ("let cofix " ^ print_defs print_term Γ l ^ nl ^ " in " ^ List.nth_default (string_of_nat n) (map (string_of_name ∘ binder_name ∘ dname) l) n) + | tInt i => "Int(" ^ string_of_prim_int i ^ ")" + | tFloat f => "Float(" ^ string_of_float f ^ ")" end. End print_term. diff --git a/template-coq/theories/Reduction.v b/template-coq/theories/Reduction.v new file mode 100644 index 000000000..c138034ff --- /dev/null +++ b/template-coq/theories/Reduction.v @@ -0,0 +1,77 @@ +(* Distributed under the terms of the MIT license. *) +From MetaCoq.Template Require Import config utils Ast AstUtils Induction LiftSubst + UnivSubst TermEquality Typing. + +From Equations Require Import Equations. +Require Import ssreflect. + +Lemma red1_tApp_mkApps_l Σ Γ M1 N1 M2 : +red1 Σ Γ M1 N1 -> red1 Σ Γ (tApp M1 M2) (mkApps N1 M2). +Proof. constructor. auto. Qed. + +Lemma red1_tApp_mkApp Σ Γ M1 N1 M2 : + red1 Σ Γ M1 N1 -> red1 Σ Γ (tApp M1 [M2]) (mkApp N1 M2). +Proof. + intros. + change (mkApp N1 M2) with (mkApps N1 [M2]). + apply app_red_l. auto. +Qed. + +Lemma red1_mkApp Σ Γ M1 N1 M2 : + Ast.wf M1 -> + red1 Σ Γ M1 N1 -> red1 Σ Γ (mkApp M1 M2) (mkApp N1 M2). +Proof. + intros wfM1 H. + destruct (isApp M1) eqn:Heq. + destruct M1; try discriminate. simpl. + revert wfM1. inv H. simpl. intros. + rewrite mkApp_mkApps. constructor. + + intros. + rewrite mkApp_mkApps. + econstructor; eauto. + inv wfM1. simpl. + clear -H1. + unfold is_constructor in *. + destruct (nth_error args narg) eqn:Heq. + eapply nth_error_app_left in Heq. now rewrite -> Heq. discriminate. + + intros. rewrite mkApp_mkApps. now constructor. + + intros. simpl. + constructor. clear -X. induction X; constructor; auto. + + rewrite mkApp_tApp; auto. + now apply red1_tApp_mkApp. +Qed. + +Lemma red1_mkApps_l Σ Γ M1 N1 M2 : + Ast.wf M1 -> All Ast.wf M2 -> + red1 Σ Γ M1 N1 -> red1 Σ Γ (mkApps M1 M2) (mkApps N1 M2). +Proof. + induction M2 in M1, N1 |- *. simpl; auto. + intros. specialize (IHM2 (mkApp M1 a) (mkApp N1 a)). + inv X. + forward IHM2. apply wf_mkApp; auto. + forward IHM2. auto. + rewrite <- !mkApps_mkApp; auto. + apply IHM2. + apply red1_mkApp; auto. +Qed. + +Lemma red1_mkApps_r Σ Γ M1 M2 N2 : + Ast.wf M1 -> All Ast.wf M2 -> + OnOne2 (red1 Σ Γ) M2 N2 -> red1 Σ Γ (mkApps M1 M2) (mkApps M1 N2). +Proof. + intros. induction X0 in M1, H, X |- *. + inv X. + destruct (isApp M1) eqn:Heq. destruct M1; try discriminate. + simpl. constructor. apply OnOne2_app. constructor. auto. + rewrite mkApps_tApp; try congruence. + rewrite mkApps_tApp; try congruence. + constructor. constructor. auto. + inv X. + specialize (IHX0 (mkApp M1 hd)). forward IHX0. + apply wf_mkApp; auto. forward IHX0; auto. + now rewrite !mkApps_mkApp in IHX0. +Qed. diff --git a/template-coq/theories/Reflect.v b/template-coq/theories/Reflect.v index ee7b639c3..23df31391 100644 --- a/template-coq/theories/Reflect.v +++ b/template-coq/theories/Reflect.v @@ -1,4 +1,6 @@ (* Distributed under the terms of the MIT license. *) +(* For primitive integers and floats *) +From Coq Require Numbers.Cyclic.Int63.Int63 Floats.PrimFloat Floats.FloatAxioms. From MetaCoq.Template Require Import utils AstUtils BasicAst Ast Induction. Require Import ssreflect. From Equations Require Import Equations. @@ -135,6 +137,30 @@ Instance reflect_nat : ReflectEq nat := { eqb_spec := Nat.eqb_spec }. +#[program] +Instance reflect_prim_int : ReflectEq Numbers.Cyclic.Int63.Int63.int := + { eqb := Numbers.Cyclic.Int63.Int63.eqb +}. +Next Obligation. + destruct (Int63.eqb x y) eqn:eq; constructor. + now apply (Numbers.Cyclic.Int63.Int63.eqb_spec x y) in eq. + now apply (Numbers.Cyclic.Int63.Int63.eqb_false_spec x y) in eq. +Qed. + +Derive NoConfusion EqDec for SpecFloat.spec_float. + +Local Obligation Tactic := idtac. +#[program] +Instance reflect_prim_float : ReflectEq PrimFloat.float := + { eqb x y := eqb (ReflectEq := EqDec_ReflectEq SpecFloat.spec_float) (FloatOps.Prim2SF x) (FloatOps.Prim2SF y) }. +Next Obligation. + intros. cbn -[eqb]. + destruct (eqb_spec (ReflectEq := EqDec_ReflectEq SpecFloat.spec_float) (FloatOps.Prim2SF x) (FloatOps.Prim2SF y)); constructor. + now apply FloatAxioms.Prim2SF_inj. + intros e; apply n. rewrite e. + reflexivity. +Qed. + Definition eq_level l1 l2 := match l1, l2 with | Level.lSet, Level.lSet => true @@ -496,6 +522,10 @@ Proof. subst. inversion e1. subst. destruct (eq_dec rarg rarg0) ; nodec. subst. left. reflexivity. + - destruct (Int63.eqs i i0) ; nodec. + subst. left. reflexivity. + - destruct (eq_dec f f0) ; nodec. + subst. left. reflexivity. Defined. Instance reflect_term : ReflectEq term := diff --git a/template-coq/theories/TermEquality.v b/template-coq/theories/TermEquality.v new file mode 100644 index 000000000..57dab3b9c --- /dev/null +++ b/template-coq/theories/TermEquality.v @@ -0,0 +1,502 @@ +(* Distributed under the terms of the MIT license. *) +From Coq Require Import CMorphisms. +From MetaCoq.Template Require Import config utils Reflect Ast AstUtils Induction LiftSubst Reflect. + +Require Import ssreflect. +From Equations.Prop Require Import DepElim. +Set Equations With UIP. + +Definition R_universe_instance R := + fun u u' => Forall2 R (List.map Universe.make u) (List.map Universe.make u'). + +(** Cumulative inductive types: + + To simplify the development, we allow the variance list to not exactly + match the instances, so as to keep syntactic equality an equivalence relation + even on ill-formed terms. It corresponds to the right notion on well-formed terms. +*) + +Definition R_universe_variance (Re Rle : Universe.t -> Universe.t -> Prop) v u u' := + match v with + | Variance.Irrelevant => True + | Variance.Covariant => Rle (Universe.make u) (Universe.make u') + | Variance.Invariant => Re (Universe.make u) (Universe.make u') + end. + +Fixpoint R_universe_instance_variance Re Rle v u u' := + match u, u' return Prop with + | u :: us, u' :: us' => + match v with + | [] => R_universe_instance_variance Re Rle v us us' + (* Missing variance stands for irrelevance, we still check that the instances have + the same length. *) + | v :: vs => R_universe_variance Re Rle v u u' /\ + R_universe_instance_variance Re Rle vs us us' + end + | [], [] => True + | _, _ => False + end. + +Definition lookup_minductive Σ mind := + match lookup_env Σ mind with + | Some (InductiveDecl decl) => Some decl + | _ => None + end. + +Definition lookup_inductive Σ ind := + match lookup_minductive Σ (inductive_mind ind) with + | Some mdecl => + match nth_error mdecl.(ind_bodies) (inductive_ind ind) with + | Some idecl => Some (mdecl, idecl) + | None => None + end + | None => None + end. + +Definition lookup_constructor Σ ind k := + match lookup_inductive Σ ind with + | Some (mdecl, idecl) => + match nth_error idecl.(ind_ctors) k with + | Some cdecl => Some (mdecl, idecl, cdecl) + | None => None + end + | _ => None + end. + +Definition global_variance Σ gr napp := + match gr with + | IndRef ind => + match lookup_inductive Σ ind with + | Some (mdecl, idecl) => + match destArity [] idecl.(ind_type) with + | Some (ctx, _) => if (context_assumptions ctx) <=? napp then mdecl.(ind_variance) + else None + | None => None + end + | None => None + end + | ConstructRef ind k => + match lookup_constructor Σ ind k with + | Some (mdecl, idecl, cdecl) => + if (cdecl.2 + mdecl.(ind_npars))%nat <=? napp then + (** Fully applied constructors are always compared at the same supertype, + which implies that no universe equality needs to be checked here. *) + Some [] + else None + | _ => None + end + | _ => None + end. + +Definition R_opt_variance Re Rle v := + match v with + | Some v => R_universe_instance_variance Re Rle v + | None => R_universe_instance Re + end. + +Definition R_global_instance Σ Re Rle gr napp := + R_opt_variance Re Rle (global_variance Σ gr napp). + +Lemma R_universe_instance_impl R R' : + RelationClasses.subrelation R R' -> + RelationClasses.subrelation (R_universe_instance R) (R_universe_instance R'). +Proof. + intros H x y xy. eapply Forall2_impl ; tea. +Qed. + +Lemma R_universe_instance_impl' R R' : + RelationClasses.subrelation R R' -> + forall u u', R_universe_instance R u u' -> R_universe_instance R' u u'. +Proof. + intros H x y xy. eapply Forall2_impl ; tea. +Qed. + +(** ** Syntactic equality up-to universes + We don't look at printing annotations *) + +(** Equality is indexed by a natural number that counts the number of applications + that surround the current term, used to implement cumulativity of inductive types + correctly (only fully applied constructors and inductives benefit from it). *) + +Inductive eq_term_upto_univ_napp Σ (Re Rle : Universe.t -> Universe.t -> Prop) (napp : nat) : term -> term -> Type := +| eq_Rel n : + eq_term_upto_univ_napp Σ Re Rle napp (tRel n) (tRel n) + +| eq_Evar e args args' : + All2 (eq_term_upto_univ_napp Σ Re Re 0) args args' -> + eq_term_upto_univ_napp Σ Re Rle napp (tEvar e args) (tEvar e args') + +| eq_Var id : + eq_term_upto_univ_napp Σ Re Rle napp (tVar id) (tVar id) + +| eq_Sort s s' : + Rle s s' -> + eq_term_upto_univ_napp Σ Re Rle napp (tSort s) (tSort s') + +| eq_App t t' u u' : + eq_term_upto_univ_napp Σ Re Rle (#|u| + napp) t t' -> + All2 (eq_term_upto_univ_napp Σ Re Re 0) u u' -> + eq_term_upto_univ_napp Σ Re Rle napp (tApp t u) (tApp t' u') + +| eq_Const c u u' : + R_universe_instance Re u u' -> + eq_term_upto_univ_napp Σ Re Rle napp (tConst c u) (tConst c u') + +| eq_Ind i u u' : + R_global_instance Σ Re Rle (IndRef i) napp u u' -> + eq_term_upto_univ_napp Σ Re Rle napp (tInd i u) (tInd i u') + +| eq_Construct i k u u' : + R_global_instance Σ Re Rle (ConstructRef i k) napp u u' -> + eq_term_upto_univ_napp Σ Re Rle napp (tConstruct i k u) (tConstruct i k u') + +| eq_Lambda na na' ty ty' t t' : + eq_binder_annot na na' -> + eq_term_upto_univ_napp Σ Re Re 0 ty ty' -> + eq_term_upto_univ_napp Σ Re Rle 0 t t' -> + eq_term_upto_univ_napp Σ Re Rle napp (tLambda na ty t) (tLambda na' ty' t') + +| eq_Prod na na' a a' b b' : + eq_binder_annot na na' -> + eq_term_upto_univ_napp Σ Re Re 0 a a' -> + eq_term_upto_univ_napp Σ Re Rle 0 b b' -> + eq_term_upto_univ_napp Σ Re Rle napp (tProd na a b) (tProd na' a' b') + +| eq_LetIn na na' t t' ty ty' u u' : + eq_binder_annot na na' -> + eq_term_upto_univ_napp Σ Re Re 0 t t' -> + eq_term_upto_univ_napp Σ Re Re 0 ty ty' -> + eq_term_upto_univ_napp Σ Re Rle 0 u u' -> + eq_term_upto_univ_napp Σ Re Rle napp (tLetIn na t ty u) (tLetIn na' t' ty' u') + +| eq_Case indn p p' c c' brs brs' : + eq_term_upto_univ_napp Σ Re Re 0 p p' -> + eq_term_upto_univ_napp Σ Re Re 0 c c' -> + All2 (fun x y => + (fst x = fst y) * + eq_term_upto_univ_napp Σ Re Re 0 (snd x) (snd y) + ) brs brs' -> + eq_term_upto_univ_napp Σ Re Rle napp (tCase indn p c brs) (tCase indn p' c' brs') + +| eq_Proj p c c' : + eq_term_upto_univ_napp Σ Re Re 0 c c' -> + eq_term_upto_univ_napp Σ Re Rle napp (tProj p c) (tProj p c') + +| eq_Fix mfix mfix' idx : + All2 (fun x y => + eq_term_upto_univ_napp Σ Re Re 0 x.(dtype) y.(dtype) * + eq_term_upto_univ_napp Σ Re Re 0 x.(dbody) y.(dbody) * + (x.(rarg) = y.(rarg)) * + eq_binder_annot x.(dname) y.(dname) + )%type mfix mfix' -> + eq_term_upto_univ_napp Σ Re Rle napp (tFix mfix idx) (tFix mfix' idx) + +| eq_CoFix mfix mfix' idx : + All2 (fun x y => + eq_term_upto_univ_napp Σ Re Re 0 x.(dtype) y.(dtype) * + eq_term_upto_univ_napp Σ Re Re 0 x.(dbody) y.(dbody) * + (x.(rarg) = y.(rarg)) * + eq_binder_annot x.(dname) y.(dname) + ) mfix mfix' -> + eq_term_upto_univ_napp Σ Re Rle napp (tCoFix mfix idx) (tCoFix mfix' idx) + +| eq_Cast t1 c t2 t1' c' t2' : + eq_term_upto_univ_napp Σ Re Re 0 t1 t1' -> + eq_cast_kind c c' -> + eq_term_upto_univ_napp Σ Re Re 0 t2 t2' -> + eq_term_upto_univ_napp Σ Re Rle napp (tCast t1 c t2) (tCast t1' c' t2') + +| eq_Int i : eq_term_upto_univ_napp Σ Re Rle napp (tInt i) (tInt i) +| eq_Float f : eq_term_upto_univ_napp Σ Re Rle napp (tFloat f) (tFloat f). + +Notation eq_term_upto_univ Σ Re Rle := (eq_term_upto_univ_napp Σ Re Rle 0). + +(* ** Syntactic conversion up-to universes *) + +Definition eq_term `{checker_flags} Σ φ := + eq_term_upto_univ Σ (eq_universe φ) (eq_universe φ). + +(* ** Syntactic cumulativity up-to universes *) + +Definition leq_term `{checker_flags} Σ φ := + eq_term_upto_univ Σ (eq_universe φ) (leq_universe φ). + + Lemma R_global_instance_refl Σ Re Rle gr napp u : + RelationClasses.Reflexive Re -> + RelationClasses.Reflexive Rle -> + R_global_instance Σ Re Rle gr napp u u. +Proof. + intros rRE rRle. + rewrite /R_global_instance. + destruct global_variance as [v|] eqn:lookup. + - induction u in v |- *; simpl; auto; + unfold R_opt_variance in IHu; destruct v; simpl; auto. + split; auto. + destruct t; simpl; auto. + - apply Forall2_same; eauto. +Qed. + +Instance eq_binder_annot_equiv {A} : RelationClasses.Equivalence (@eq_binder_annot A). +Proof. + split. + - red. reflexivity. + - red; now symmetry. + - intros x y z; unfold eq_binder_annot. + congruence. +Qed. + +Definition eq_binder_annot_refl {A} x : @eq_binder_annot A x x. +Proof. reflexivity. Qed. + +Hint Resolve @eq_binder_annot_refl : core. + +Lemma eq_term_upto_univ_refl Σ Re Rle : + RelationClasses.Reflexive Re -> + RelationClasses.Reflexive Rle -> + forall napp t, eq_term_upto_univ_napp Σ Re Rle napp t t. +Proof. + intros hRe hRle napp. + induction t in napp, Rle, hRle |- * using term_forall_list_rect; simpl; + try constructor; try apply Forall_Forall2; try apply All_All2 ; try easy; + try now (try apply Forall_All ; apply Forall_True). + - eapply All_All2. 1: eassumption. + intros. simpl in X0. easy. + - destruct c; constructor. + - eapply All_All2. 1: eassumption. + intros. easy. + - now apply R_global_instance_refl. + - now apply R_global_instance_refl. + - red in X. eapply All_All2. 1:eassumption. + intros; easy. + - eapply All_All2. 1: eassumption. + simpl. + intros x [? ?]. repeat split ; auto. + - eapply All_All2. 1: eassumption. + intros x [? ?]. repeat split ; auto. +Qed. + +Lemma eq_term_refl `{checker_flags} Σ φ t : eq_term Σ φ t t. +Proof. + apply eq_term_upto_univ_refl. + - intro; apply eq_universe_refl. + - intro; apply eq_universe_refl. +Qed. + + +Lemma leq_term_refl `{checker_flags} Σ φ t : leq_term Σ φ t t. +Proof. + apply eq_term_upto_univ_refl. + - intro; apply eq_universe_refl. + - intro; apply leq_universe_refl. +Qed. +(* +Lemma eq_term_leq_term `{checker_flags} Σ φ napp t u : + eq_term_upto_univ_napp Σ napp φ t u -> leq_term Σ φ t u. +Proof. + induction t in u |- * using term_forall_list_rect; simpl; inversion 1; + subst; constructor; try (now unfold eq_term, leq_term in * ); + try eapply Forall2_impl' ; try eapply All2_impl' ; try easy. + now apply eq_universe_leq_universe. + all: try (apply Forall_True, eq_universe_leq_universe). + apply IHt. +Qed. *) + + +Instance R_global_instance_impl_same_napp Σ Re Re' Rle Rle' gr napp : + RelationClasses.subrelation Re Re' -> + RelationClasses.subrelation Rle Rle' -> + subrelation (R_global_instance Σ Re Rle gr napp) (R_global_instance Σ Re' Rle' gr napp). +Proof. + intros he hle t t'. + rewrite /R_global_instance /R_opt_variance. + destruct global_variance as [v|] eqn:glob. + induction t in v, t' |- *; destruct v, t'; simpl; auto. + intros []; split; auto. + destruct t0; simpl; auto. + now eapply R_universe_instance_impl'. +Qed. + +Lemma eq_term_upto_univ_morphism0 Σ (Re Re' : _ -> _ -> Prop) + (Hre : forall t u, Re t u -> Re' t u) + : forall t u napp, eq_term_upto_univ_napp Σ Re Re napp t u -> eq_term_upto_univ_napp Σ Re' Re' napp t u. +Proof. + fix aux 4. + destruct 1; constructor; eauto. + all: unfold R_universe_instance in *. + all: try solve[ match goal with + | H : All2 _ _ _ |- _ => clear -H aux; induction H; constructor; eauto + | H : Forall2 _ _ _ |- _ => induction H; constructor; eauto + end]. + - eapply R_global_instance_impl_same_napp; eauto. + - eapply R_global_instance_impl_same_napp; eauto. + - induction a; constructor; auto. intuition auto. + - induction a; constructor; auto. intuition auto. + - induction a; constructor; auto. intuition auto. +Qed. + +Lemma eq_term_upto_univ_morphism Σ (Re Re' Rle Rle' : _ -> _ -> Prop) + (Hre : forall t u, Re t u -> Re' t u) + (Hrle : forall t u, Rle t u -> Rle' t u) + : forall t u napp, eq_term_upto_univ_napp Σ Re Rle napp t u -> eq_term_upto_univ_napp Σ Re' Rle' napp t u. +Proof. + fix aux 4. + destruct 1; constructor; eauto using eq_term_upto_univ_morphism0. + all: unfold R_universe_instance in *. + all: try solve [match goal with + | H : Forall2 _ _ _ |- _ => induction H; constructor; + eauto using eq_term_upto_univ_morphism0 + | H : All2 _ _ _ |- _ => induction H; constructor; + eauto using eq_term_upto_univ_morphism0 + end]. + - clear X. induction a; constructor; eauto using eq_term_upto_univ_morphism0. + - eapply R_global_instance_impl_same_napp; eauto. + - eapply R_global_instance_impl_same_napp; eauto. + - clear X1 X2. induction a; constructor; eauto using eq_term_upto_univ_morphism0. + destruct r. split; eauto using eq_term_upto_univ_morphism0. + - induction a; constructor; eauto using eq_term_upto_univ_morphism0. + destruct r as [[[? ?] ?] ?]. + repeat split; eauto using eq_term_upto_univ_morphism0. + - induction a; constructor; eauto using eq_term_upto_univ_morphism0. + destruct r as [[[? ?] ?] ?]. + repeat split; eauto using eq_term_upto_univ_morphism0. +Qed. + + +Lemma global_variance_napp_mon {Σ gr napp napp' v} : + napp <= napp' -> + global_variance Σ gr napp = Some v -> + global_variance Σ gr napp' = Some v. +Proof. + intros hnapp. + rewrite /global_variance. + destruct gr; try congruence. + - destruct lookup_inductive as [[mdecl idec]|] => //. + destruct destArity as [[ctx s]|] => //. + elim: Nat.leb_spec => // cass indv. + elim: Nat.leb_spec => //. lia. + - destruct lookup_constructor as [[[mdecl idecl] cdecl]|] => //. + elim: Nat.leb_spec => // cass indv. + elim: Nat.leb_spec => //. lia. +Qed. + +Instance R_global_instance_impl Σ Re Re' Rle Rle' gr napp napp' : + RelationClasses.subrelation Re Re' -> + RelationClasses.subrelation Re Rle' -> + RelationClasses.subrelation Rle Rle' -> + napp <= napp' -> + subrelation (R_global_instance Σ Re Rle gr napp) (R_global_instance Σ Re' Rle' gr napp'). +Proof. + intros he hle hele hnapp t t'. + rewrite /R_global_instance /R_opt_variance. + destruct global_variance as [v|] eqn:glob. + rewrite (global_variance_napp_mon hnapp glob). + induction t in v, t' |- *; destruct v, t'; simpl; auto. + intros []; split; auto. + destruct t0; simpl; auto. + destruct (global_variance _ _ napp') as [v|] eqn:glob'; eauto using R_universe_instance_impl'. + induction t in v, t' |- *; destruct v, t'; simpl; auto; intros H; inv H. + eauto. + split; auto. + destruct t0; simpl; auto. +Qed. + +Instance eq_term_upto_univ_impl Σ Re Re' Rle Rle' napp napp' : + RelationClasses.subrelation Re Re' -> + RelationClasses.subrelation Rle Rle' -> + RelationClasses.subrelation Re Rle' -> + napp <= napp' -> + subrelation (eq_term_upto_univ_napp Σ Re Rle napp) (eq_term_upto_univ_napp Σ Re' Rle' napp'). +Proof. + intros he hle hele hnapp t t'. + induction t in napp, napp', hnapp, t', Rle, Rle', hle, hele |- * using term_forall_list_rect; + try (inversion 1; subst; constructor; + eauto using R_universe_instance_impl'; fail). + - inversion 1; subst; constructor. + eapply All2_impl'; tea. + eapply All_impl; eauto. + - inversion 1; subst; constructor. + eapply IHt. 4:eauto. all:auto with arith. eauto. + solve_all. + - inversion 1; subst; constructor. + eapply R_global_instance_impl. 5:eauto. all:auto. + - inversion 1; subst; constructor. + eapply R_global_instance_impl. 5:eauto. all:eauto. + - inversion 1; subst; constructor; eauto. + eapply All2_impl'; tea. + eapply All_impl; eauto. + cbn. intros x ? y [? ?]. split; eauto. + - inversion 1; subst; constructor. + eapply All2_impl'; tea. + eapply All_impl; eauto. + cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. + - inversion 1; subst; constructor. + eapply All2_impl'; tea. + eapply All_impl; eauto. + cbn. intros x [? ?] y [[[? ?] ?] ?]. repeat split; eauto. +Qed. + + +Lemma eq_term_leq_term `{checker_flags} Σ φ t u : + eq_term Σ φ t u -> leq_term Σ φ t u. +Proof. + eapply eq_term_upto_univ_morphism. auto. + intros. + now apply eq_universe_leq_universe. +Qed. + +Lemma eq_term_upto_univ_App `{checker_flags} Σ Re Rle napp f f' : + eq_term_upto_univ_napp Σ Re Rle napp f f' -> + isApp f = isApp f'. +Proof. + inversion 1; reflexivity. +Qed. + +Lemma eq_term_App `{checker_flags} Σ φ f f' : + eq_term Σ φ f f' -> + isApp f = isApp f'. +Proof. + inversion 1; reflexivity. +Qed. + +Lemma eq_term_upto_univ_mkApps `{checker_flags} Σ Re Rle napp f l f' l' : + eq_term_upto_univ_napp Σ Re Rle (#|l| + napp) f f' -> + All2 (eq_term_upto_univ Σ Re Re) l l' -> + eq_term_upto_univ_napp Σ Re Rle napp (mkApps f l) (mkApps f' l'). +Proof. + induction l in f, f' |- *; intro e; inversion_clear 1. + - assumption. + - pose proof (eq_term_upto_univ_App _ _ _ _ _ _ e). + case_eq (isApp f). + + intro X; rewrite X in H0. + destruct f; try discriminate. + destruct f'; try discriminate. + cbn. inversion_clear e. constructor. + rewrite app_length /= -Nat.add_assoc //. + apply All2_app. assumption. + now constructor. + + intro X; rewrite X in H0. + rewrite !mkApps_tApp; eauto. + intro; discriminate. + intro; discriminate. + constructor. simpl. now simpl in e. + now constructor. +Qed. + +Lemma leq_term_mkApps `{checker_flags} Σ φ f l f' l' : + leq_term Σ φ f f' -> + All2 (eq_term Σ φ) l l' -> + leq_term Σ φ (mkApps f l) (mkApps f' l'). +Proof. + intros. + eapply eq_term_upto_univ_mkApps. + eapply eq_term_upto_univ_impl. 5:eauto. 4:auto with arith. + 1-3:typeclasses eauto. + apply X0. +Qed. + +Lemma leq_term_App `{checker_flags} Σ φ f f' : + leq_term Σ φ f f' -> + isApp f = isApp f'. +Proof. + inversion 1; reflexivity. +Qed. \ No newline at end of file diff --git a/template-coq/theories/Typing.v b/template-coq/theories/Typing.v index 456768237..ba8633dfb 100644 --- a/template-coq/theories/Typing.v +++ b/template-coq/theories/Typing.v @@ -1,7 +1,9 @@ (* Distributed under the terms of the MIT license. *) -From Coq Require Import ssreflect Wellfounded Relation_Operators. +(** This defines relation operators in Type *) +From Equations.Type Require Import Relation. +From Coq Require Import ssreflect Wellfounded Relation_Operators CRelationClasses. From MetaCoq.Template Require Import config utils Ast AstUtils LiftSubst UnivSubst - EnvironmentTyping. + EnvironmentTyping Reflect TermEquality. (** * Typing derivations @@ -9,7 +11,7 @@ From MetaCoq.Template Require Import config utils Ast AstUtils LiftSubst UnivSub *) - + Definition isSort T := match T with | tSort u => True @@ -80,13 +82,7 @@ Definition fix_subst (l : mfixpoint term) := Definition unfold_fix (mfix : mfixpoint term) (idx : nat) := match List.nth_error mfix idx with - | Some d => - if isLambda d.(dbody) then - Some (d.(rarg), subst0 (fix_subst mfix) d.(dbody)) - else None (* We don't unfold ill-formed fixpoints, which would - render confluence unprovable, creating an infinite - number of critical pairs between unfoldings of fixpoints. - e.g. [fix f := f] is not allowed. *) + | Some d => Some (d.(rarg), subst0 (fix_subst mfix) d.(dbody)) | None => None end. @@ -167,7 +163,7 @@ Inductive red1 (Σ : global_env) (Γ : context) : term -> term -> Type := | red_fix mfix idx args narg fn : unfold_fix mfix idx = Some (narg, fn) -> is_constructor narg args = true -> - red1 Σ Γ (tApp (tFix mfix idx) args) (tApp fn args) + red1 Σ Γ (tApp (tFix mfix idx) args) (mkApps fn args) (** CoFix-case unfolding *) | red_cofix_case ip p mfix idx args narg fn brs : @@ -263,7 +259,7 @@ Lemma red1_ind_all : (forall (Γ : context) (mfix : mfixpoint term) (idx : nat) (args : list term) (narg : nat) (fn : term), unfold_fix mfix idx = Some (narg, fn) -> - is_constructor narg args = true -> P Γ (tApp (tFix mfix idx) args) (tApp fn args)) -> + is_constructor narg args = true -> P Γ (tApp (tFix mfix idx) args) (mkApps fn args)) -> (forall (Γ : context) (ip : inductive * nat * relevance) (p : term) (mfix : mfixpoint term) (idx : nat) (args : list term) (narg : nat) (fn : term) (brs : list (nat * term)), unfold_cofix mfix idx = Some (narg, fn) -> @@ -423,250 +419,138 @@ Inductive red Σ Γ M : term -> Type := We hence implement first an equality which considers casts and do a stripping phase of casts before checking equality. *) -Definition R_universe_instance R := - fun u u' => Forall2 R (List.map Universe.make u) (List.map Universe.make u'). - -Inductive eq_term_upto_univ (Re Rle : Universe.t -> Universe.t -> Prop) : term -> term -> Type := -| eq_Rel n : - eq_term_upto_univ Re Rle (tRel n) (tRel n) - -| eq_Evar e args args' : - All2 (eq_term_upto_univ Re Re) args args' -> - eq_term_upto_univ Re Rle (tEvar e args) (tEvar e args') - -| eq_Var id : - eq_term_upto_univ Re Rle (tVar id) (tVar id) - -| eq_Sort s s' : - Rle s s' -> - eq_term_upto_univ Re Rle (tSort s) (tSort s') - -| eq_Cast f f' k T T' : - eq_term_upto_univ Re Re f f' -> - eq_term_upto_univ Re Re T T' -> - eq_term_upto_univ Re Rle (tCast f k T) (tCast f' k T') - -| eq_App t t' args args' : - eq_term_upto_univ Re Rle t t' -> - All2 (eq_term_upto_univ Re Re) args args' -> - eq_term_upto_univ Re Rle (tApp t args) (tApp t' args') - -| eq_Const c u u' : - R_universe_instance Re u u' -> - eq_term_upto_univ Re Rle (tConst c u) (tConst c u') - -| eq_Ind i u u' : - R_universe_instance Re u u' -> - eq_term_upto_univ Re Rle (tInd i u) (tInd i u') - -| eq_Construct i k u u' : - R_universe_instance Re u u' -> - eq_term_upto_univ Re Rle (tConstruct i k u) (tConstruct i k u') - -| eq_Lambda na na' ty ty' t t' : - eq_term_upto_univ Re Re ty ty' -> - eq_term_upto_univ Re Rle t t' -> - eq_term_upto_univ Re Rle (tLambda na ty t) (tLambda na' ty' t') - -| eq_Prod na na' a a' b b' : - eq_term_upto_univ Re Re a a' -> - eq_term_upto_univ Re Rle b b' -> - eq_term_upto_univ Re Rle (tProd na a b) (tProd na' a' b') - -| eq_LetIn na na' ty ty' t t' u u' : - eq_term_upto_univ Re Re ty ty' -> - eq_term_upto_univ Re Re t t' -> - eq_term_upto_univ Re Rle u u' -> - eq_term_upto_univ Re Rle (tLetIn na ty t u) (tLetIn na' ty' t' u') - -| eq_Case ind par p p' c c' brs brs' : - eq_term_upto_univ Re Re p p' -> - eq_term_upto_univ Re Re c c' -> - All2 (fun x y => - fst x = fst y × - eq_term_upto_univ Re Re (snd x) (snd y) - ) brs brs' -> - eq_term_upto_univ Re Rle (tCase (ind, par) p c brs) (tCase (ind, par) p' c' brs') - -| eq_Proj p c c' : - eq_term_upto_univ Re Re c c' -> - eq_term_upto_univ Re Rle (tProj p c) (tProj p c') - -| eq_Fix mfix mfix' idx : - All2 (fun x y => - eq_term_upto_univ Re Re x.(dtype) y.(dtype) × - eq_term_upto_univ Re Re x.(dbody) y.(dbody) × - x.(rarg) = y.(rarg) - ) mfix mfix' -> - eq_term_upto_univ Re Rle (tFix mfix idx) (tFix mfix' idx) - -| eq_CoFix mfix mfix' idx : - All2 (fun x y => - eq_term_upto_univ Re Re x.(dtype) y.(dtype) × - eq_term_upto_univ Re Re x.(dbody) y.(dbody) × - x.(rarg) = y.(rarg) - ) mfix mfix' -> - eq_term_upto_univ Re Rle (tCoFix mfix idx) (tCoFix mfix' idx). - -Definition eq_term `{checker_flags} φ := - eq_term_upto_univ (eq_universe φ) (eq_universe φ). - -(* ** Syntactic cumulativity up-to universes - - We shouldn't look at printing annotations *) - -Definition leq_term `{checker_flags} φ := - eq_term_upto_univ (eq_universe φ) (leq_universe φ). - -Fixpoint strip_casts t := - match t with - | tEvar ev args => tEvar ev (List.map strip_casts args) - | tLambda na T M => tLambda na (strip_casts T) (strip_casts M) - | tApp u v => mkApps (strip_casts u) (List.map (strip_casts) v) - | tProd na A B => tProd na (strip_casts A) (strip_casts B) - | tCast c kind t => strip_casts c - | tLetIn na b t b' => tLetIn na (strip_casts b) (strip_casts t) (strip_casts b') - | tCase ind p c brs => - let brs' := List.map (on_snd (strip_casts)) brs in - tCase ind (strip_casts p) (strip_casts c) brs' - | tProj p c => tProj p (strip_casts c) - | tFix mfix idx => - let mfix' := List.map (map_def strip_casts strip_casts) mfix in - tFix mfix' idx - | tCoFix mfix idx => - let mfix' := List.map (map_def strip_casts strip_casts) mfix in - tCoFix mfix' idx - | tRel _ | tVar _ | tSort _ | tConst _ _ | tInd _ _ | tConstruct _ _ _ => t - end. - -Definition eq_term_nocast `{checker_flags} (φ : ConstraintSet.t) (t u : term) := - eq_term φ (strip_casts t) (strip_casts u). +Definition eq_term_nocast `{checker_flags} (Σ : global_env) (φ : ConstraintSet.t) (t u : term) := + eq_term Σ φ (strip_casts t) (strip_casts u). -Definition leq_term_nocast `{checker_flags} (φ : ConstraintSet.t) (t u : term) := - leq_term φ (strip_casts t) (strip_casts u). +Definition leq_term_nocast `{checker_flags} (Σ : global_env) (φ : ConstraintSet.t) (t u : term) := + leq_term Σ φ (strip_casts t) (strip_casts u). (** ** Utilities for typing *) (** Decompose an arity into a context and a sort *) -Fixpoint destArity Γ (t : term) := - match t with - | tProd na t b => destArity (Γ ,, vass na t) b - | tLetIn na b b_ty b' => destArity (Γ ,, vdef na b b_ty) b' - | tSort s => Some (Γ, s) - | _ => None - end. - - Reserved Notation " Σ ;;; Γ |- t : T " (at level 50, Γ, t, T at next level). Reserved Notation " Σ ;;; Γ |- t <= u " (at level 50, Γ, t, u at next level). Reserved Notation " Σ ;;; Γ |- t = u " (at level 50, Γ, t, u at next level). -(** ** Cumulativity *) + +(** ** Cumulativity: + + Reduction to terms in the leq_term relation. +*) Inductive cumul `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> term -> Type := -| cumul_refl t u : leq_term (global_ext_constraints Σ) t u -> Σ ;;; Γ |- t <= u +| cumul_refl t u : leq_term Σ (global_ext_constraints Σ) t u -> Σ ;;; Γ |- t <= u | cumul_red_l t u v : red1 Σ.1 Γ t v -> Σ ;;; Γ |- v <= u -> Σ ;;; Γ |- t <= u | cumul_red_r t u v : Σ ;;; Γ |- t <= v -> red1 Σ.1 Γ u v -> Σ ;;; Γ |- t <= u - where " Σ ;;; Γ |- t <= u " := (cumul Σ Γ t u) : type_scope. (** *** Conversion - Defined as cumulativity in both directions. + Reduction to terms in the eq_term relation *) -Inductive conv `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> term -> Type := - | conv_refl t u : eq_term (global_ext_constraints Σ) t u -> Σ ;;; Γ |- t = u + Inductive conv `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> term -> Type := + | conv_refl t u : eq_term Σ (global_ext_constraints Σ) t u -> Σ ;;; Γ |- t = u | conv_red_l t u v : red1 Σ.1 Γ t v -> Σ ;;; Γ |- v = u -> Σ ;;; Γ |- t = u | conv_red_r t u v : Σ ;;; Γ |- t = v -> red1 Σ.1 Γ u v -> Σ ;;; Γ |- t = u where " Σ ;;; Γ |- t = u " := (@conv _ Σ Γ t u) : type_scope. Lemma conv_refl' `{checker_flags} : forall Σ Γ t, Σ ;;; Γ |- t = t. - intros. todo "conv_refl". -Defined. - -Lemma cumul_refl' `{checker_flags} : forall Σ Γ t, Σ ;;; Γ |- t <= t. (* easy *) - intros. todo "cumul_refl'". + intros. constructor. apply eq_term_refl. Defined. -Lemma cumul_trans `{checker_flags} : forall Σ Γ t u v, Σ ;;; Γ |- t <= u -> Σ ;;; Γ |- u <= v -> Σ ;;; Γ |- t <= v. - intros. todo "cumul_trans". +Lemma cumul_refl' `{checker_flags} : forall Σ Γ t, Σ ;;; Γ |- t <= t. + intros. constructor. apply leq_term_refl. Defined. -Hint Resolve conv_refl cumul_refl' : typecheck. +Hint Resolve conv_refl' cumul_refl' : typecheck. -Lemma congr_cumul_prod `{checker_flags} : forall Σ Γ na na' M1 M2 N1 N2, - cumul Σ Γ M1 N1 -> - cumul Σ (Γ ,, vass na M1) M2 N2 -> - cumul Σ Γ (tProd na M1 M2) (tProd na' N1 N2). -Proof. intros. todo "congr_cumul_prod". Defined. - -Definition eq_opt_term `{checker_flags} φ (t u : option term) := +Definition eq_opt_term `{checker_flags} Σ φ (t u : option term) := match t, u with - | Some t, Some u => eq_term φ t u + | Some t, Some u => eq_term Σ φ t u | None, None => True | _, _ => False end. -Definition eq_decl `{checker_flags} φ (d d' : context_decl) := - eq_opt_term φ d.(decl_body) d'.(decl_body) * eq_term φ d.(decl_type) d'.(decl_type). +Definition eq_decl `{checker_flags} Σ φ (d d' : context_decl) := + eq_opt_term Σ φ d.(decl_body) d'.(decl_body) * eq_term Σ φ d.(decl_type) d'.(decl_type). -Definition eq_context `{checker_flags} φ (Γ Δ : context) := - All2 (eq_decl φ) Γ Δ. +Definition eq_context `{checker_flags} Σ φ (Γ Δ : context) := + All2 (eq_decl Σ φ) Γ Δ. (** ** Typing relation *) Module TemplateEnvTyping := EnvTyping TemplateTerm TemplateEnvironment. Include TemplateEnvTyping. -Section WfArity. - Context (typing : forall (Σ : global_env_ext) (Γ : context), term -> term -> Type). - - Definition isWfArity Σ (Γ : context) T := - { ctx & { s & (destArity [] T = Some (ctx, s)) * All_local_env (lift_typing typing Σ) (Γ ,,, ctx) } }. - - Context (property : forall (Σ : global_env_ext) (Γ : context), - All_local_env (lift_typing typing Σ) Γ -> - forall (t T : term), typing Σ Γ t T -> Type). - - Definition isWfArity_prop Σ (Γ : context) T := - { wfa : isWfArity Σ Γ T & All_local_env_over typing property Σ _ (snd (projT2 (projT2 wfa))) }. -End WfArity. +Definition extends (Σ Σ' : global_env) := { Σ'' & Σ' = Σ'' ++ Σ }. -(* AXIOM GUARD CONDITION *) -Axiom fix_guard : mfixpoint term -> bool. +Class GuardChecker := +{ (* Structural recursion check *) + fix_guard : global_env_ext -> context -> mfixpoint term -> bool ; + (* Guarded by destructors check *) + cofix_guard : global_env_ext -> context -> mfixpoint term -> bool ; -Axiom fix_guard_red1 : - forall Σ Γ mfix mfix' idx, - fix_guard mfix -> - red1 Σ Γ (tFix mfix idx) (tFix mfix' idx) -> - fix_guard mfix'. + fix_guard_red1 Σ Γ mfix mfix' idx : + fix_guard Σ Γ mfix -> + red1 Σ Γ (tFix mfix idx) (tFix mfix' idx) -> + fix_guard Σ Γ mfix' ; + + fix_guard_lift Σ Γ Γ' Γ'' mfix : + let k' := (#|mfix| + #|Γ'|)%nat in + let mfix' := map (map_def (lift #|Γ''| #|Γ'|) (lift #|Γ''| k')) mfix in + fix_guard Σ (Γ ,,, Γ') mfix -> + fix_guard Σ (Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ') mfix' ; -Axiom fix_guard_lift : - forall mfix n k, + fix_guard_subst Σ Γ Γ' Δ mfix s k : let k' := (#|mfix| + k)%nat in - let mfix' := map (map_def (lift n k) (lift n k')) mfix in - fix_guard mfix -> - fix_guard mfix'. - -Axiom fix_guard_subst : - forall mfix s k, + let mfix' := map (map_def (subst s k) (subst s k')) mfix in + fix_guard Σ (Γ ,,, Γ' ,,, Δ) mfix -> + fix_guard Σ (Γ ,,, subst_context s 0 Δ) mfix' ; + + fix_guard_subst_instance {cf:checker_flags} Σ Γ mfix u univs : + consistent_instance_ext (Σ.1, univs) Σ.2 u -> + fix_guard Σ Γ mfix -> + fix_guard (Σ.1, univs) (subst_instance_context u Γ) (map (map_def (subst_instance_constr u) (subst_instance_constr u)) + mfix) ; + + fix_guard_extends Σ Γ mfix (Σ' : global_env_ext) : + fix_guard Σ Γ mfix -> + extends Σ.1 Σ' -> + fix_guard Σ' Γ mfix ; + + cofix_guard_red1 Σ Γ mfix mfix' idx : + cofix_guard Σ Γ mfix -> + red1 Σ Γ (tCoFix mfix idx) (tCoFix mfix' idx) -> + cofix_guard Σ Γ mfix' ; + + cofix_guard_lift Σ Γ Γ' Γ'' mfix : + let k' := (#|mfix| + #|Γ'|)%nat in + let mfix' := map (map_def (lift #|Γ''| #|Γ'|) (lift #|Γ''| k')) mfix in + cofix_guard Σ (Γ ,,, Γ') mfix -> + cofix_guard Σ (Γ ,,, Γ'' ,,, lift_context #|Γ''| 0 Γ') mfix' ; + + cofix_guard_subst Σ Γ Γ' Δ mfix s k : let k' := (#|mfix| + k)%nat in let mfix' := map (map_def (subst s k) (subst s k')) mfix in - fix_guard mfix -> - fix_guard mfix'. - -(* AXIOM INDUCTIVE GUARD CONDITION *) -Axiom ind_guard : mutual_inductive_body -> bool. - -Extract Constant fix_guard => "fun m -> assert false". -Extract Constant fix_guard_red1 => "fun s g m m' i -> assert false". -Extract Constant fix_guard_lift => "fun m n k -> assert false". -Extract Constant fix_guard_subst => "fun m s k -> assert false". - -Extract Constant ind_guard => "fun m -> assert false". + cofix_guard Σ (Γ ,,, Γ' ,,, Δ) mfix -> + cofix_guard Σ (Γ ,,, subst_context s 0 Δ) mfix' ; + + cofix_guard_subst_instance {cf:checker_flags} Σ Γ mfix u univs : + consistent_instance_ext (Σ.1, univs) Σ.2 u -> + cofix_guard Σ Γ mfix -> + cofix_guard (Σ.1, univs) (subst_instance_context u Γ) (map (map_def (subst_instance_constr u) (subst_instance_constr u)) + mfix) ; + + cofix_guard_extends Σ Γ mfix (Σ' : global_env_ext) : + cofix_guard Σ Γ mfix -> + extends Σ.1 Σ' -> + cofix_guard Σ' Γ mfix }. +(* AXIOM GUARD CONDITION *) +Axiom guard_checking : GuardChecker. +Existing Instance guard_checking. (** Compute the type of a case from the predicate [p], actual parameters [pars] and an inductive declaration. *) @@ -758,6 +642,80 @@ Definition build_case_predicate_type ind mdecl idecl params u ps : option term : decl_type := mkApps (tInd ind u) (map (lift0 #|X.1|) params ++ to_extended_list X.1) |} in ret (it_mkProd_or_LetIn (X.1 ,, inddecl) (tSort ps)). +Definition destInd (t : term) := + match t with + | tInd ind u => Some (ind, u) + | _ => None + end. + +Definition isFinite (r : recursivity_kind) := + match r with + | Finite => true + | _ => false + end. + +Definition isCoFinite (r : recursivity_kind) := + match r with + | CoFinite => true + | _ => false + end. + +Definition check_recursivity_kind (Σ : global_env) ind r := + match lookup_env Σ ind with + | Some (InductiveDecl mib) => Reflect.eqb mib.(ind_finite) r + | _ => false + end. + +Definition check_one_fix d := + let '{| dname := na; + dtype := ty; + dbody := b; + rarg := arg |} := d in + let '(ctx, ty) := decompose_prod_assum [] ty in + match nth_error (List.rev (smash_context [] ctx)) arg with + | Some argd => + let (hd, args) := decompose_app argd.(decl_type) in + match destInd hd with + | Some (mkInd mind _, u) => Some mind + | None => None (* Not recursive on an inductive type *) + end + | None => None (* Recursive argument not found *) + end. + +Definition wf_fixpoint (Σ : global_env) mfix := + let checks := map check_one_fix mfix in + match map_option_out checks with + | Some (ind :: inds) => + (* Check that mutually recursive fixpoints are all on the same mututal + inductive block *) + forallb (Reflect.eqb ind) inds && + check_recursivity_kind Σ ind Finite + | _ => false + end. + +Definition check_one_cofix d := + let '{| dname := na; + dtype := ty; + dbody := b; + rarg := arg |} := d in + let '(ctx, ty) := decompose_prod_assum [] ty in + let (hd, args) := decompose_app ty in + match destInd hd with + | Some (mkInd ind _, u) => Some ind + | None => None (* Not recursive on an inductive type *) + end. + +Definition wf_cofixpoint (Σ : global_env) mfix := + let checks := map check_one_cofix mfix in + match map_option_out checks with + | Some (ind :: inds) => + (* Check that mutually recursive cofixpoints are all producing + coinductives in the same mututal coinductive block *) + forallb (Reflect.eqb ind) inds && + check_recursivity_kind Σ ind CoFinite + | _ => false + end. + Definition wf_universe Σ s := match s with | Universe.lProp @@ -765,20 +723,19 @@ Definition wf_universe Σ s := | Universe.lType u => forall l, UnivExprSet.In l u -> LevelSet.In (UnivExpr.get_level l) (global_ext_levels Σ) end. - + +Reserved Notation "'wf_local' Σ Γ " (at level 9, Σ, Γ at next level). Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> term -> Type := | type_Rel n decl : - All_local_env (lift_typing typing Σ) Γ -> + wf_local Σ Γ -> nth_error Γ n = Some decl -> Σ ;;; Γ |- tRel n : lift0 (S n) decl.(decl_type) -| type_Sort l : - All_local_env (lift_typing typing Σ) Γ -> - (l = inl PropLevel.lSProp \/ - l = inl PropLevel.lProp \/ - (exists l', LevelSet.In l' (global_ext_levels Σ) /\ l = inr l')) -> - Σ ;;; Γ |- tSort (Universe.of_levels l) : tSort (Universe.super (Universe.of_levels l)) +| type_Sort s : + wf_local Σ Γ -> + wf_universe Σ s -> + Σ ;;; Γ |- tSort s : tSort (Universe.super s) | type_Cast c k t s : Σ ;;; Γ |- t : tSort s -> @@ -808,19 +765,19 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ |- tApp t l : t' | type_Const cst u : - All_local_env (lift_typing typing Σ) Γ -> + wf_local Σ Γ -> forall decl (isdecl : declared_constant Σ.1 cst decl), consistent_instance_ext Σ decl.(cst_universes) u -> Σ ;;; Γ |- (tConst cst u) : subst_instance_constr u decl.(cst_type) | type_Ind ind u : - All_local_env (lift_typing typing Σ) Γ -> + wf_local Σ Γ -> forall mdecl idecl (isdecl : declared_inductive Σ.1 mdecl ind idecl), consistent_instance_ext Σ mdecl.(ind_universes) u -> Σ ;;; Γ |- (tInd ind u) : subst_instance_constr u idecl.(ind_type) | type_Construct ind i u : - All_local_env (lift_typing typing Σ) Γ -> + wf_local Σ Γ -> forall mdecl idecl cdecl (isdecl : declared_constructor Σ.1 mdecl idecl (ind, i) cdecl), consistent_instance_ext Σ mdecl.(ind_universes) u -> Σ ;;; Γ |- (tConstruct ind i u) : type_of_constructor mdecl cdecl (ind, i) u @@ -834,9 +791,10 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> forall ps pty, build_case_predicate_type ind mdecl idecl params u ps = Some pty -> Σ ;;; Γ |- p : pty -> is_allowed_elimination Σ ps idecl.(ind_kelim) -> + isCoFinite mdecl.(ind_finite) = false -> Σ ;;; Γ |- c : mkApps (tInd ind u) args -> forall btys, map_option_out (build_branches_type ind mdecl idecl params u p) = Some btys -> - All2 (fun br bty => (br.1 = bty.1) * (Σ ;;; Γ |- br.2 : bty.2) * (Σ ;;; Γ |- bty.2 : tSort ps)) brs btys -> + All2 (fun br bty => (br.1 = bty.1) * (Σ ;;; Γ |- br.2 : bty.2) * (∑ s, Σ ;;; Γ |- bty.2 : tSort s)) brs btys -> Σ ;;; Γ |- tCase indnparrel p c brs : mkApps p (skipn npar args ++ [c]) | type_Proj p c u : @@ -846,27 +804,30 @@ Inductive typing `{checker_flags} (Σ : global_env_ext) (Γ : context) : term -> Σ ;;; Γ |- tProj p c : subst0 (c :: List.rev args) (subst_instance_constr u pdecl.2) | type_Fix mfix n decl : - fix_guard mfix -> - nth_error mfix n = Some decl -> - All_local_env (lift_typing typing Σ) Γ -> - All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> - All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype)) - * (isLambda d.(dbody) = true)%type) mfix -> - Σ ;;; Γ |- tFix mfix n : decl.(dtype) + fix_guard Σ Γ mfix -> + nth_error mfix n = Some decl -> + wf_local Σ Γ -> + All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> + All (fun d => (Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype))) mfix -> + wf_fixpoint Σ mfix -> + Σ ;;; Γ |- tFix mfix n : decl.(dtype) | type_CoFix mfix n decl : + cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> - All_local_env (lift_typing typing Σ) Γ -> + wf_local Σ Γ -> All (fun d => {s & Σ ;;; Γ |- d.(dtype) : tSort s}) mfix -> All (fun d => Σ ;;; Γ ,,, fix_context mfix |- d.(dbody) : lift0 #|fix_context mfix| d.(dtype)) mfix -> + wf_cofixpoint Σ mfix -> Σ ;;; Γ |- tCoFix mfix n : decl.(dtype) -| type_Conv t A B : +| type_Conv t A B s : Σ ;;; Γ |- t : A -> - (isWfArity typing Σ Γ B + {s & Σ ;;; Γ |- B : tSort s})%type -> + Σ ;;; Γ |- B : tSort s -> Σ ;;; Γ |- A <= B -> Σ ;;; Γ |- t : B where " Σ ;;; Γ |- t : T " := (typing Σ Γ t T) : type_scope + and "'wf_local' Σ Γ " := (All_local_env (lift_typing typing Σ) Γ) : type_scope (* Typing of "spines", currently just the arguments of applications *) @@ -879,8 +840,6 @@ with typing_spine `{checker_flags} (Σ : global_env_ext) (Γ : context) : term - typing_spine Σ Γ (subst10 hd B) tl B' -> typing_spine Σ Γ T (cons hd tl) B'. -Notation wf_local Σ Γ := (All_local_env (lift_typing typing Σ) Γ). - (** ** Typechecking of global environments *) Definition has_nparams npars ty := @@ -893,7 +852,6 @@ Definition unlift_opt_pred (P : global_env_ext -> context -> option term -> term Module TemplateTyping <: Typing TemplateTerm TemplateEnvironment TemplateEnvTyping. - Definition ind_guard := ind_guard. Definition typing := @typing. Definition wf_universe := @wf_universe. Definition conv := @conv. @@ -964,19 +922,14 @@ Proof. - exact (S (S (wf_local_size _ typing_size _ a))). - exact (S (S (wf_local_size _ typing_size _ a))). - exact (S (Nat.max d1 (Nat.max d2 - (all2_size _ (fun x y p => Nat.max (typing_size Σ Γ (snd x) (snd y) (snd (fst p))) (typing_size _ _ _ _ (snd p))) a)))). - - exact (S (Nat.max (Nat.max (wf_local_size _ typing_size _ a) (all_size _ (fun x p => typing_size Σ _ _ _ p.π2) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ (fst p)) a1))). + (all2_size _ (fun x y p => Nat.max (typing_size Σ Γ (snd x) (snd y) (snd (fst p))) (typing_size _ _ _ _ (snd p).π2)) a)))). + - exact (S (Nat.max (Nat.max (wf_local_size _ typing_size _ a) (all_size _ (fun x p => typing_size Σ _ _ _ p.π2) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - exact (S (Nat.max (Nat.max (wf_local_size _ typing_size _ a) (all_size _ (fun x p => typing_size Σ _ _ _ p.π2) a0)) (all_size _ (fun x p => typing_size Σ _ _ _ p) a1))). - - destruct s. - + red in i. - exact (S (Nat.max (wf_local_size _ typing_size _ (snd (projT2 (projT2 i)))) d)). - + destruct s as [u Hu]. apply typing_size in Hu. - exact (S (Nat.max d Hu)). Defined. Lemma typing_size_pos `{checker_flags} {Σ Γ t T} (d : Σ ;;; Γ |- t : T) : typing_size d > 0. Proof. - induction d; simpl; try lia. destruct s as [s | [u Hu]]; try lia. + induction d; simpl; try lia. Qed. Fixpoint globenv_size (Σ : global_env) : size := @@ -1007,7 +960,7 @@ Proof. intros. now apply X. Qed. Lemma type_Prop `{checker_flags} Σ : Σ ;;; [] |- tSort Universe.lProp : tSort Universe.type1. change ( Σ ;;; [] |- tSort (Universe.of_levels (inl PropLevel.lProp)) : tSort Universe.type1); - constructor;auto. constructor. + constructor;auto. constructor. constructor. Defined. Lemma env_prop_sigma `{checker_flags} P : env_prop P -> @@ -1018,18 +971,17 @@ Proof. apply type_Prop. Defined. - -Lemma wf_local_app `{checker_flags} Σ (Γ Γ' : context) : wf_local Σ (Γ ,,, Γ') -> wf_local Σ Γ. +Lemma wf_local_app_l `{checker_flags} Σ (Γ Γ' : context) : wf_local Σ (Γ ,,, Γ') -> wf_local Σ Γ. Proof. induction Γ'. auto. simpl. intros H'; inv H'; eauto. Defined. -Hint Resolve wf_local_app : wf. +Hint Resolve wf_local_app_l : wf. Lemma typing_wf_local `{checker_flags} {Σ} {Γ t T} : Σ ;;; Γ |- t : T -> wf_local Σ Γ. Proof. - induction 1; eauto using wf_local_app. + induction 1; eauto using wf_local_app_l. Defined. Hint Resolve typing_wf_local : wf. @@ -1041,7 +993,6 @@ Proof. change (fun (x : global_env_ext) (x0 : context) (x1 x2 : term) (x3 : x;;; x0 |- x1 : x2) => typing_size x3) with (@typing_size H); try lia. - destruct indnparrel as ((ind' & npar') & ?); cbn in *; subst ind npar. lia. - - destruct s as [s | [u Hu]]; try lia. Qed. Lemma wf_local_inv `{checker_flags} {Σ Γ'} (w : wf_local Σ Γ') : @@ -1100,7 +1051,6 @@ Inductive Forall_typing_spine `{checker_flags} Σ Γ (P : term -> term -> Type) Forall_typing_spine Σ Γ P T (hd :: tl) B' (type_spine_cons Σ Γ hd tl na A B s T B' typrod cumul ty tls). - Lemma typing_ind_env `{cf : checker_flags} : forall (P : global_env_ext -> context -> term -> term -> Type) (Pdecl := fun Σ Γ wfΓ t T tyT => P Σ Γ t T), @@ -1108,14 +1058,10 @@ Lemma typing_ind_env `{cf : checker_flags} : nth_error Γ n = Some decl -> All_local_env_over typing Pdecl Σ Γ wfΓ -> P Σ Γ (tRel n) (lift0 (S n) decl.(decl_type))) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) - (l : PropLevel.t + Level.t), - All_local_env_over typing Pdecl Σ Γ wfΓ -> - (l = inl PropLevel.lSProp \/ - l = inl PropLevel.lProp \/ - (exists l', LevelSet.In l' (global_ext_levels Σ) /\ l = inr l')) -> - - P Σ Γ (tSort (Universe.of_levels l)) (tSort (Universe.super (Universe.of_levels l)))) -> + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (u : Universe.t), + All_local_env_over typing Pdecl Σ Γ wfΓ -> + wf_universe Σ u -> + P Σ Γ (tSort u) (tSort (Universe.super u))) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (c : term) (k : cast_kind) (t : term) (s : Universe.t), @@ -1186,10 +1132,11 @@ Lemma typing_ind_env `{cf : checker_flags} : is_allowed_elimination (global_ext_constraints Σ) ps idecl.(ind_kelim) -> Σ ;;; Γ |- c : mkApps (tInd ind u) args -> P Σ Γ c (mkApps (tInd ind u) args) -> + isCoFinite mdecl.(ind_finite) = false -> forall btys, map_option_out (build_branches_type ind mdecl idecl params u p) = Some btys -> All2 (fun br bty => (br.1 = bty.1) * (Σ ;;; Γ |- br.2 : bty.2) * P Σ Γ br.2 bty.2 * - (Σ ;;; Γ |- bty.2 : tSort ps) * P Σ Γ bty.2 (tSort ps)) + ∑ s, (Σ ;;; Γ |- bty.2 : tSort s) * P Σ Γ bty.2 (tSort s)) brs btys -> P Σ Γ (tCase (ind, npar, r) p c brs) (mkApps p (skipn npar args ++ [c]))) -> @@ -1203,29 +1150,32 @@ Lemma typing_ind_env `{cf : checker_flags} : (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (mfix : list (def term)) (n : nat) decl, let types := fix_context mfix in - fix_guard mfix -> + fix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> All_local_env_over typing Pdecl Σ Γ wfΓ -> All (fun d => {s & (Σ ;;; Γ |- d.(dtype) : tSort s)%type * P Σ Γ d.(dtype) (tSort s)})%type mfix -> All (fun d => (Σ ;;; Γ ,,, types |- d.(dbody) : lift0 #|types| d.(dtype))%type * - (isLambda d.(dbody) = true)%type * P Σ (Γ ,,, types) d.(dbody) (lift0 #|types| d.(dtype)))%type mfix -> + wf_fixpoint Σ.1 mfix -> P Σ Γ (tFix mfix n) decl.(dtype)) -> (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (mfix : list (def term)) (n : nat) decl, let types := fix_context mfix in + cofix_guard Σ Γ mfix -> nth_error mfix n = Some decl -> All_local_env_over typing Pdecl Σ Γ wfΓ -> All (fun d => {s & (Σ ;;; Γ |- d.(dtype) : tSort s)%type * P Σ Γ d.(dtype) (tSort s)})%type mfix -> All (fun d => (Σ ;;; Γ ,,, types |- d.(dbody) : lift0 #|types| d.(dtype))%type * P Σ (Γ ,,, types) d.(dbody) (lift0 #|types| d.(dtype)))%type mfix -> + wf_cofixpoint Σ.1 mfix -> P Σ Γ (tCoFix mfix n) decl.(dtype)) -> - (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t A B : term), + (forall Σ (wfΣ : wf Σ.1) (Γ : context) (wfΓ : wf_local Σ Γ) (t A B : term) s, All_local_env_over typing Pdecl Σ Γ wfΓ -> Σ ;;; Γ |- t : A -> P Σ Γ t A -> - (isWfArity_prop typing Pdecl Σ Γ B + {s & (Σ ;;; Γ |- B : tSort s) * P Σ Γ B (tSort s)})%type -> + Σ ;;; Γ |- B : tSort s -> + P Σ Γ B (tSort s) -> Σ ;;; Γ |- A <= B -> P Σ Γ t B) -> @@ -1432,7 +1382,8 @@ Proof. --- intuition eauto. +++ eapply (X14 _ wfΓ _ _ t); eauto. simpl; auto with arith. lia. - +++ eapply (X14 _ wfΓ _ _ t0); eauto. simpl; auto with arith. + +++ destruct s as [s Hs]. exists s; split; [auto|]. + eapply (X14 _ wfΓ _ _ Hs); eauto. simpl; auto with arith. lia. --- apply IHa. auto. intros. eapply (X14 _ wfΓ0 _ _ Hty). lia. @@ -1467,17 +1418,17 @@ Proof. typing_size Hty < S (all_size (fun x : def term => (Σ;;; Γ ,,, fix_context mfix |- dbody x : lift0 #|fix_context mfix| (dtype x))%type - * (isLambda (dbody x) = true)%type)%type - (fun (x : def term) p => typing_size (fst p)) a1) -> + )%type + (fun (x : def term) p => typing_size p) a1) -> Forall_decls_typing P Σ.1 * P Σ Γ0 t T). {intros. eapply (X14 _ X _ _ Hty); eauto. lia. } clear X14 X13. - clear e decl i a0. + clear e decl i a0 i0. remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. induction a1; econstructor; eauto. - ++ split; auto. - eapply (X _ (typing_wf_local (fst p)) _ _ (fst p)). simpl. lia. + ++ split; auto. + eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. ++ eapply IHa1. intros. eapply (X _ X0 _ _ Hty). simpl; lia. @@ -1510,36 +1461,16 @@ Proof. Forall_decls_typing P Σ.1 * P Σ Γ0 t T). {intros. eapply (X14 _ X _ _ Hty); eauto. lia. } clear X14 X13. - clear e decl a0. + clear e decl a0 i i0. remember (fix_context mfix) as mfixcontext. clear Heqmfixcontext. - + induction a1; econstructor; eauto. - ++ split; auto. + ++ split; auto. eapply (X _ (typing_wf_local p) _ _ p). simpl. lia. ++ eapply IHa1. intros. eapply (X _ X0 _ _ Hty). simpl; lia. - - -- eapply X12; eauto. - ++ eapply (X14 _ wfΓ _ _ H); simpl. destruct s as [s | [u Hu]]; lia. - ++ simpl in X14, X13. - destruct s as [s | [u Hu]]. - ** left. - red. exists s. red in s. revert X14. - generalize (snd (projT2 (projT2 s))). - induction a; simpl in *; intros X14 *; constructor. - --- apply IHa. intros. eapply (X14 _ wfΓ0 _ _ Hty). lia. - --- red. eapply (X14 _ a _ _ (projT2 t1)). clear. - destruct t1. simpl. lia. - --- apply IHa. intros. eapply (X14 _ wfΓ0 _ _ Hty). lia. - --- red. destruct t1. unshelve eapply X14. all: eauto. - simpl. lia. - --- red. destruct t1. unshelve eapply X14. all: eauto. - simpl. lia. - ** right. - exists u. intuition eauto. unshelve eapply X14. all: eauto. lia. Qed. - (** * Lemmas about All_local_env *) Lemma nth_error_All_local_env {P Γ n} (isdecl : n < #|Γ|) : @@ -1567,7 +1498,7 @@ Proof. - apply IHX. Qed. -Lemma All_local_env_app +Lemma All_local_env_app_inv (P : context -> term -> option term -> Type) l l' : All_local_env P (l ,,, l') -> All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l'. @@ -1603,7 +1534,7 @@ Proof. + apply IHX. Qed. -Lemma All_local_env_app_inv `{checker_flags} (P : context -> term -> option term -> Type) l l' : +Lemma All_local_env_app `{checker_flags} (P : context -> term -> option term -> Type) l l' : All_local_env P l * All_local_env (fun Γ t T => P (l ,,, Γ) t T) l' -> All_local_env P (l ,,, l'). Proof. @@ -1620,7 +1551,7 @@ Qed. Definition property `{checker_flags} := forall (Σ : global_env_ext) (Γ : context), - All_local_env (lift_typing typing Σ) Γ -> forall t T : term, typing Σ Γ t T -> Type. + wf_local Σ Γ -> forall t T : term, typing Σ Γ t T -> Type. Definition lookup_wf_local {Γ P} (wfΓ : All_local_env P Γ) (n : nat) (isdecl : n < #|Γ|) : All_local_env P (skipn (S n) Γ). @@ -1660,7 +1591,7 @@ Definition on_wf_local_decl `{checker_flags} {Σ Γ} end H. -Lemma nth_error_All_local_env_over `{checker_flags} {P Σ Γ n decl} (eq : nth_error Γ n = Some decl) {wfΓ : All_local_env (lift_typing typing Σ) Γ} : +Lemma nth_error_All_local_env_over `{checker_flags} {P Σ Γ n decl} (eq : nth_error Γ n = Some decl) {wfΓ : wf_local Σ Γ} : All_local_env_over typing P Σ Γ wfΓ -> let Γ' := skipn (S n) Γ in let p := lookup_wf_local_decl wfΓ n eq in diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index 298e683d1..ade75bae3 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -1,7 +1,8 @@ (* Distributed under the terms of the MIT license. *) From MetaCoq.Template Require Import config utils Ast AstUtils Induction LiftSubst - UnivSubst Typing. -Require Import ssreflect. + UnivSubst Reduction WfInv Typing. +From Equations Require Import Equations. +Require Import ssreflect ssrbool. (** * Well-formedness of terms and types in typing derivations @@ -59,7 +60,7 @@ Proof. simpl. intros H; now inv H. Qed. -Lemma wf_mkApps_napp t u : isApp t = false -> Ast.wf (mkApps t u) -> Ast.wf t /\ List.Forall Ast.wf u. +Lemma wf_mkApps_napp t u : ~~ isApp t -> Ast.wf (mkApps t u) -> Ast.wf t /\ List.Forall Ast.wf u. Proof. induction u in t |- *; simpl. - intuition. @@ -85,25 +86,22 @@ Qed. Hint Resolve wf_mkApps_inv : wf. Hint Constructors Ast.wf : wf. -Lemma isLambda_isApp t : isLambda t = true -> isApp t = false. +Lemma isLambda_isApp t : isLambda t = true -> ~~ isApp t. Proof. destruct t; simpl; congruence. Qed. Lemma unfold_fix_wf: forall (mfix : mfixpoint term) (idx : nat) (narg : nat) (fn : term), unfold_fix mfix idx = Some (narg, fn) -> Ast.wf (tFix mfix idx) -> - Ast.wf fn /\ isApp fn = false. + Ast.wf fn. Proof. intros mfix idx narg fn Hf Hwf. unfold unfold_fix in Hf. inv Hwf. destruct nth_error eqn:eqnth; try congruence. pose proof (nth_error_forall eqnth H). simpl in H0. - destruct H0 as [ _ [ wfd islamd ] ]. - rewrite islamd in Hf. + destruct H0 as [ _ wfd]. injection Hf. intros <- <-. - apply (isLambda_subst (fix_subst mfix) 0) in islamd. - apply isLambda_isApp in islamd. split; try congruence. - apply wf_subst; auto. clear wfd islamd Hf eqnth. + apply wf_subst; auto. clear wfd Hf eqnth. assert(forall n, Ast.wf (tFix mfix n)). constructor; auto. unfold fix_subst. generalize #|mfix|; intros. induction n; auto. Qed. @@ -133,8 +131,6 @@ Proof. now apply Forall_map. - constructor; auto. solve_all. - constructor. solve_all. - destruct x; simpl in *. repeat split; tas. - destruct dbody; simpl in *; congruence. - constructor. solve_all. Qed. @@ -203,7 +199,7 @@ Proof. induction brs in c, H1 |- *; destruct c; simpl in *. constructor. constructor. inv H1; auto. inv H1; auto. induction H0 in pars |- *; destruct pars; try constructor; auto. simpl. auto. - - apply unfold_fix_wf in H; auto. constructor; intuition auto. + - apply unfold_fix_wf in H; auto. eapply wf_mkApps; auto. - constructor; auto. apply wf_mkApps_napp in H1 as [Hcof Hargs]; auto. apply unfold_cofix_wf in H; auto. apply wf_mkApps; intuition auto. @@ -234,7 +230,7 @@ Proof. - constructor; auto. solve_all. pose proof H as H'. revert H. apply (OnOne2_All_All X). clear X. - intros [na bo ty ra] [nb bb tb rb] [[r ih] e] [? [? ?]]. + intros [na bo ty ra] [nb bb tb rb] [[r ih] e] [? ?]. simpl in *. inversion e. subst. clear e. intuition eauto. @@ -245,7 +241,6 @@ Proof. cbn. unfold wf_decl. simpl. intros ? [? ? ? ?] ?. simpl in *. intuition eauto with wf. - + eapply red1_isLambda. all: eassumption. - constructor; auto. induction X; inv H; constructor; intuition auto; congruence. - constructor; auto. solve_all. @@ -298,8 +293,6 @@ Proof. - destruct t; try reflexivity. discriminate. - destruct l; simpl in *; congruence. - - destruct x; simpl in *; intuition eauto. - destruct dbody; simpl in *; try discriminate. destruct Nat.leb; auto. Qed. Lemma declared_inductive_wf {cf:checker_flags} : @@ -366,15 +359,6 @@ Proof. try red; simpl; intuition eauto. Qed. -Lemma subst_context_snoc s k Γ d : subst_context s k (d :: Γ) = subst_context s k Γ ,, map_decl (subst s (#|Γ| + k)) d. -Proof. - unfold subst_context, fold_context. - rewrite !rev_mapi !rev_involutive /mapi mapi_rec_eqn /snoc. - f_equal. now rewrite Nat.sub_0_r List.rev_length. - rewrite mapi_rec_Sk. simpl. apply mapi_rec_ext. intros. - rewrite app_length !List.rev_length. simpl. f_equal. f_equal. lia. -Qed. - Lemma wf_subst_context s k Γ : Forall wf_decl Γ -> Forall Ast.wf s -> Forall wf_decl (subst_context s k Γ). Proof. intros wfΓ. induction wfΓ in s |- *. @@ -523,7 +507,7 @@ Proof. apply wf_subst; auto with wf. apply wf_subst_instance_constr. eapply declared_constructor_wf; eauto. - split. wf. constructor; eauto. solve_all. - apply wf_mkApps. wf. solve_all. apply wf_mkApps_inv in H7. solve_all. + apply wf_mkApps. wf. solve_all. apply wf_mkApps_inv in H8. solve_all. apply All_app_inv; solve_all. now apply All_skipn. - split. wf. apply wf_subst. solve_all. constructor. wf. apply wf_mkApps_inv in H2. apply All_rev. solve_all. @@ -544,15 +528,6 @@ Proof. solve_all. destruct a. intuition. + eapply All_nth_error in X0; eauto. destruct X0 as [s ?]; intuition. - - - split. apply H. destruct X1 as [X1|[s X1]]; [|apply X1]. - destruct X1 as [[Γ' [s [X1 X1']]] XX]; cbn in *. - assert (HB : B = it_mkProd_or_LetIn Γ' (tSort s)). { - clear -X1. pose proof (destArity_spec [] B) as HH. - rewrite X1 in HH. assumption. } - rewrite HB. clear -XX. - eapply wf_it_mkProd_or_LetIn in XX. rewrite it_mkProd_or_LetIn_app in XX. - apply it_mkProd_or_LetIn_wf in XX. exact XX. constructor. Qed. Lemma typing_all_wf_decl {cf:checker_flags} Σ (wfΣ : wf Σ.1) Γ (wfΓ : wf_local Σ Γ) : @@ -568,49 +543,6 @@ Proof. Qed. Hint Resolve typing_all_wf_decl : wf. -(* TODO MOVE? *) -Definition sort_irrelevant - (P : global_env_ext -> context -> term -> option term -> Type) := - forall Σ Γ b s s', P Σ Γ (tSort s) b -> P Σ Γ (tSort s') b. - - -(* TODO MOVE? *) -Lemma on_global_env_mix `{checker_flags} {Σ P Q} : - sort_irrelevant Q -> - on_global_env P Σ -> - on_global_env Q Σ -> - on_global_env (fun Σ Γ t T => (P Σ Γ t T * Q Σ Γ t T)%type) Σ. -Proof. - intros iQ hP hQ. - induction hP in Q, iQ, hQ |- *. - 1: constructor. - invs hQ. constructor. - - eapply IHhP. all: eauto. - - assumption. - - assumption. - (* - destruct d. all: simpl in *. - + destruct c as [ty [bo|] un]. all: simpl in *. - all: unfold on_constant_decl in *. - all: simpl in *. - * intuition eauto. - * unfold on_type in *. intuition eauto. - + destruct o0 as [oi op onpars ong]. - destruct o2 as [oi' op' onpars' ong']. - constructor. all: auto. - * clear - oi oi'. revert oi oi'. - generalize (TemplateEnvironment.ind_bodies m). - intros l h1 h2. - induction h1 in h2 |- *. 1: constructor. - dependent destruction h2. - constructor. 2: auto. - destruct p, o. econstructor. - -- eassumption. - -- unfold on_type in *. intuition eauto. - -- admit. - -- *) - - todo "simplify onConstructors"%string. -Qed. - Lemma on_global_env_impl `{checker_flags} Σ P Q : (forall Σ Γ t T, on_global_env P Σ.1 -> P Σ Γ t T -> Q Σ Γ t T) -> on_global_env P Σ -> on_global_env Q Σ. @@ -662,13 +594,8 @@ Proof. intros. pose proof (env_prop_sigma _ typing_wf_gen _ wfΣ). red in X. unfold lift_typing in X. do 2 red in wfΣ. - unshelve eapply (on_global_env_mix _ wfΣ) in X. - red. intros. destruct b; intuition auto with wf. - destruct X0 as [u Hu]. exists u. intuition auto with wf. - clear wfΣ. - eapply on_global_env_impl; eauto; simpl; intros. clear X. - destruct X1 as [Hty Ht]. - destruct T. apply Ht. destruct Ht; wf. + eapply on_global_env_impl; eauto; simpl; intros. + destruct T. red. apply X1. red. destruct X1 as [x [a wfs]]. split; auto. Qed. Lemma typing_wf {cf:checker_flags} Σ (wfΣ : wf Σ.1) Γ t T : @@ -791,3 +718,155 @@ Proof. eapply All_Forall, Alli_All; eauto. simpl; intros. eapply declared_projection_wf in H0; eauto. Qed. + +Lemma mkApp_ex_wf t u : Ast.wf (mkApp t u) -> + exists f args, mkApp t u = tApp f args /\ ~~ isApp f. +Proof. + induction t; simpl; try solve [eexists _, _; split; reflexivity]. + intros wf. + eapply wf_inv in wf as [appt [_ [wft wfargs]]]. + eapply Forall_app in wfargs as [wfargs wfu]. depelim wfu. + forward IHt. eapply wf_mkApp; intuition auto. + destruct IHt as [f [ar [eqf isap]]]. + eexists _, _; split; auto. rewrite appt //. +Qed. + +Lemma decompose_app_mkApp f u : + (decompose_app (mkApp f u)).2 <> []. +Proof. + induction f; simpl; auto; try congruence. + destruct args; simpl; congruence. +Qed. + +Lemma mkApps_tApp' f u f' u' : + ~~ isApp f' -> + mkApp f u = tApp f' u' -> mkApps f [u] = mkApps f' u'. +Proof. + intros. + rewrite -(mkApp_mkApps f u []). + simpl. rewrite H0. + rewrite -(mkApps_tApp f') // ?H //. + destruct u' => //. + eapply (f_equal decompose_app) in H0. + simpl in H0. pose proof (decompose_app_mkApp f u). + rewrite H0 /= in H1. congruence. +Qed. + +Lemma eq_decompose_app x y : + Ast.wf x -> Ast.wf y -> + decompose_app x = decompose_app y -> x = y. +Proof. + intros wfx; revert y. + induction wfx using term_wf_forall_list_ind; intros [] wfy; + eapply wf_inv in wfy; simpl in wfy; simpl; + intros [= ?]; try intuition congruence. +Qed. + +Lemma mkApp_ex t u : ∑ f args, mkApp t u = tApp f args. +Proof. + induction t; simpl; try solve [eexists _, _; reflexivity]. +Qed. + +Lemma strip_casts_decompose_app t : + Ast.wf t -> + forall f l, decompose_app t = (f, l) -> + strip_casts t = mkApps (strip_casts f) (map strip_casts l). +Proof. + intros wf. + induction wf using term_wf_forall_list_ind; simpl; intros; auto; noconf H; + try noconf H0; + rewrite ?map_map_compose ?compose_on_snd ?compose_map_def ?map_length; + f_equal; solve_all; eauto. + + - now noconf H3. + - now noconf H3. +Qed. + +Lemma mkApps_tApp f args : + ~~ isApp f -> + ~~ is_empty args -> + tApp f args = mkApps f args. +Proof. + intros. + destruct args, f; try discriminate; auto. +Qed. + +Lemma strip_casts_mkApps_napp_wf f u : + ~~ isApp f -> Ast.wf f -> Forall Ast.wf u -> + strip_casts (mkApps f u) = mkApps (strip_casts f) (map strip_casts u). +Proof. + intros nisapp wf wf'. + destruct u. + simpl. auto. + rewrite -(mkApps_tApp f (t :: u)) //. +Qed. + +Lemma mkApp_mkApps f u : mkApp f u = mkApps f [u]. +Proof. reflexivity. Qed. + +Lemma decompose_app_inv f l hd args : + Ast.wf f -> + decompose_app (mkApps f l) = (hd, args) -> + ∑ n, ~~ isApp hd /\ l = skipn n args /\ f = mkApps hd (firstn n args). +Proof. + destruct (isApp f) eqn:Heq. + revert l args hd. + induction f; try discriminate. intros. + simpl in H. + move/wf_inv: H => /= [isAppf [Hargs [wff wfargs]]]. + rewrite mkApps_tApp ?isAppf in H0 => //. destruct args => //. + rewrite mkApps_nested in H0. + rewrite decompose_app_mkApps ?isAppf in H0; auto. noconf H0. + exists #|args|; split; auto. now rewrite isAppf. + rewrite skipn_all_app. + rewrite firstn_app. rewrite firstn_all2. lia. + rewrite Nat.sub_diag firstn_O app_nil_r. split; auto. + rewrite mkApps_tApp ?isAppf //. now destruct args. + + intros wff fl. + rewrite decompose_app_mkApps in fl; auto. now apply negbT. + inversion fl. subst; exists 0. + split; auto. now eapply negbT. +Qed. + +Lemma eq_tip_skipn {A} (x : A) n l : [x] = skipn n l -> + exists l', l = l' ++ [x] /\ n = #|l'|. +Proof. + induction l in n |- *. rewrite skipn_nil //. + destruct n. simpl. destruct l => //. + intros eq. noconf eq. exists []; split; auto. + rewrite skipn_S. intros Hx. + destruct (IHl _ Hx) as [l' [-> ->]]. + exists (a :: l'); split; reflexivity. +Qed. + +Lemma strip_casts_mkApp_wf f u : + Ast.wf f -> Ast.wf u -> + strip_casts (mkApp f u) = mkApp (strip_casts f) (strip_casts u). +Proof. + intros wf wf'. + assert (wfa : Ast.wf (mkApp f u)). now apply wf_mkApp. + destruct (mkApp_ex_wf f u wfa) as [f' [args [eq isapp]]]. + eapply (f_equal decompose_app) in eq. simpl in eq. + epose proof (strip_casts_decompose_app _ wfa _ _ eq). + rewrite H. + rewrite mkApp_mkApps in eq. + destruct (decompose_app_inv _ _ _ _ wf eq) as [n [ng [stripeq stripf]]]. + apply eq_tip_skipn in stripeq. destruct stripeq as [l' [eqargs eqn]]. + subst n args. rewrite (firstn_app_left _ 0) // /= app_nil_r in stripf. subst f. + eapply wf_mkApps_napp in wf as [wff' wfl] => //. + rewrite strip_casts_mkApps_napp_wf //. + now rewrite mkApp_mkApps mkApps_nested map_app. +Qed. + +Lemma strip_casts_mkApps_wf f u : + Ast.wf f -> Forall Ast.wf u -> + strip_casts (mkApps f u) = mkApps (strip_casts f) (map strip_casts u). +Proof. + intros wf wf'. induction wf' in f, wf |- *. + simpl. auto. + rewrite -mkApps_mkApp IHwf'. + apply wf_mkApp; auto with wf. + rewrite strip_casts_mkApp_wf //. + now rewrite mkApps_mkApp. +Qed. diff --git a/template-coq/theories/UnivSubst.v b/template-coq/theories/UnivSubst.v index b46ece465..19838d06e 100644 --- a/template-coq/theories/UnivSubst.v +++ b/template-coq/theories/UnivSubst.v @@ -10,7 +10,7 @@ From MetaCoq Require Import utils Ast AstUtils Induction LiftSubst. Instance subst_instance_constr : UnivSubst term := fix subst_instance_constr u c {struct c} : term := match c with - | tRel _ | tVar _ => c + | tRel _ | tVar _ | tInt _ | tFloat _ => c | tEvar ev args => tEvar ev (List.map (subst_instance_constr u) args) | tSort s => tSort (subst_instance_univ u s) | tConst c u' => tConst c (subst_instance_instance u u') diff --git a/template-coq/theories/Universes.v b/template-coq/theories/Universes.v index 2c6b89001..6c4bc0004 100644 --- a/template-coq/theories/Universes.v +++ b/template-coq/theories/Universes.v @@ -1,5 +1,7 @@ From Coq Require Import MSetList MSetFacts MSetProperties MSetDecide. From MetaCoq.Template Require Import utils BasicAst config. +From Equations Require Import Equations. +Require Import ssreflect. Local Open Scope nat_scope. Local Open Scope string_scope2. @@ -23,15 +25,26 @@ Inductive universes := | UProp | USProp | UType (i : nat). +Derive NoConfusion EqDec for universes. Class Evaluable (A : Type) := val : valuation -> A -> nat. +(** This inductive classifies which eliminations are allowed for inductive types + in various sorts. *) +Inductive allowed_eliminations : Set := +| IntoSProp +| IntoPropSProp +| IntoSetPropSProp +| IntoAny. +Derive NoConfusion EqDec for allowed_eliminations. + (** Levels are Set or Level or Var *) Module Level. Inductive t_ : Set := | lSet | Level (_ : string) | Var (_ : nat) (* these are debruijn indices *). + Derive NoConfusion EqDec for t_. Definition t := t_. @@ -146,6 +159,7 @@ Module LevelSet := MSetList.MakeWithLeibniz Level. Module LevelSetFact := WFactsOn Level LevelSet. Module LevelSetProp := WPropertiesOn Level LevelSet. Module LevelSetDecide := WDecide (LevelSet). +Module LS := LevelSet. Ltac lsets := LevelSetDecide.fsetdec. Definition LevelSet_pair x y @@ -166,12 +180,31 @@ Proof. apply orb_true_iff. Qed. +Lemma LevelSet_In_fold_right_add x l : + In x l <-> LevelSet.In x (fold_right LevelSet.add LevelSet.empty l). +Proof. + split. + - induction l; simpl => //. + intros [<-|H]. + * eapply LevelSet.add_spec; left; auto. + * eapply LevelSet.add_spec; right; auto. + - induction l; simpl => //. + * now rewrite LevelSetFact.empty_iff. + * rewrite LevelSet.add_spec. intuition auto. +Qed. + +Lemma LevelSet_union_empty s : LevelSet.union LevelSet.empty s = s. +Proof. + apply LevelSet.eq_leibniz. + change LevelSet.eq with LevelSet.Equal. + intros x; rewrite LevelSet.union_spec. lsets. +Qed. (* prop level is Prop or SProp *) Module PropLevel. Inductive t := lSProp | lProp. - + Derive NoConfusion EqDec for t. (* Global Instance PropLevel_Evaluable : Evaluable t := fun v l => match l with lSProp => -1 @@ -253,7 +286,7 @@ Module UnivExpr. Qed. Definition lt_compat : Proper (Logic.eq ==> Logic.eq ==> iff) lt. - intros x x' H1 y y' H2; now rewrite H1, H2. + intros x x' H1 y y' H2; now rewrite H1 H2. Qed. Definition compare (x y : t) : comparison := @@ -299,6 +332,16 @@ Module UnivExprSet := MSetList.MakeWithLeibniz UnivExpr. Module UnivExprSetFact := WFactsOn UnivExpr UnivExprSet. Module UnivExprSetProp := WPropertiesOn UnivExpr UnivExprSet. +(* We have decidable equality w.r.t leibniz equality for sets of levels. + This means universes also have a decidable equality. *) +Instance univexprset_eq_dec : Classes.EqDec UnivExprSet.t. +Proof. + intros p p'. + destruct (UnivExprSet.eq_dec p p'). + - now left; eapply UnivExprSet.eq_leibniz in e. + - right. intros ->. apply n. reflexivity. +Qed. + Module Universe. (** A universe is a list of universe expressions which is: - sorted @@ -307,9 +350,20 @@ Module Universe. Record t0 := { t_set : UnivExprSet.t ; t_ne : UnivExprSet.is_empty t_set = false }. + Derive NoConfusion for t0. + + (** This needs a propositional UIP proof to show that [is_empty = false] is a set *) + Set Equations With UIP. + Instance t0_eqdec : EqDec t0. + Proof. eqdec_proof. Qed. + Inductive t_ := lProp | lSProp | lType (_ : t0). + Derive NoConfusion for t_. + Instance t_eqdec : EqDec t_. + Proof. eqdec_proof. Qed. + Definition t := t_. Coercion t_set : t0 >-> UnivExprSet.t. @@ -1011,10 +1065,12 @@ Qed. Module ConstraintType. Inductive t_ : Set := Le (z : Z) | Eq. + Derive NoConfusion EqDec for t_. + Definition t := t_. Definition eq : t -> t -> Prop := eq. Definition eq_equiv : Equivalence eq := _. - + Definition Le0 := Le 0. Definition Lt := Le 1. @@ -1059,6 +1115,7 @@ End ConstraintType. Module UnivConstraint. Definition t : Set := Level.t * ConstraintType.t * Level.t. + Definition eq : t -> t -> Prop := eq. Definition eq_equiv : Equivalence eq := _. @@ -1116,7 +1173,47 @@ End UnivConstraint. Module ConstraintSet := MSetList.MakeWithLeibniz UnivConstraint. Module ConstraintSetFact := WFactsOn UnivConstraint ConstraintSet. Module ConstraintSetProp := WPropertiesOn UnivConstraint ConstraintSet. +Module CS := ConstraintSet. + +Lemma CS_union_empty s : ConstraintSet.union ConstraintSet.empty s = s. +Proof. + apply ConstraintSet.eq_leibniz. + change ConstraintSet.eq with ConstraintSet.Equal. + intros x; rewrite ConstraintSet.union_spec. lsets. +Qed. +Lemma CS_For_all_union f cst cst' : ConstraintSet.For_all f (ConstraintSet.union cst cst') -> + ConstraintSet.For_all f cst. +Proof. + unfold CS.For_all. + intros IH x inx. apply (IH x). + now eapply CS.union_spec; left. +Qed. + +Lemma CS_For_all_add P x s : CS.For_all P (CS.add x s) -> P x /\ CS.For_all P s. +Proof. + intros. + split. + * apply (H x), CS.add_spec; left => //. + * intros y iny. apply (H y), CS.add_spec; right => //. +Qed. +Instance CS_For_all_proper P : Morphisms.Proper (CS.Equal ==> iff)%signature (ConstraintSet.For_all P). +Proof. + intros s s' eqs. + unfold CS.For_all. split; intros IH x inxs; apply (IH x); + now apply eqs. +Qed. + +(* Being built up from sorted lists without duplicates, constraint sets have + decidable equality. This is however not used in the development. *) +Set Equations With UIP. +Remark ConstraintSet_EqDec : EqDec ConstraintSet.t. +Proof. + intros p p'. + destruct (ConstraintSet.eq_dec p p'). + - now left; eapply ConstraintSet.eq_leibniz in e. + - right. intros ->. apply n. reflexivity. +Qed. (** {6 Universe instances} *) @@ -1140,8 +1237,6 @@ Module Instance. forallb2 f i j. End Instance. - - Module UContext. Definition t := Instance.t × ConstraintSet.t. @@ -1187,6 +1282,7 @@ Module Variance. | Irrelevant : t | Covariant : t | Invariant : t. + Derive NoConfusion EqDec for t. (* val check_subtype : t -> t -> bool *) (* val sup : t -> t -> t *) @@ -1196,6 +1292,8 @@ Inductive universes_decl : Type := | Monomorphic_ctx (ctx : ContextSet.t) | Polymorphic_ctx (cst : AUContext.t). +Derive NoConfusion for universes_decl. + Definition monomorphic_udecl u := match u with | Monomorphic_ctx ctx => ctx @@ -1422,7 +1520,7 @@ Section Univ. Proof. intros H1 H2 v Hv. specialize (is_prop_and_is_sprop_val_false _ H1 H2 v) as [n Hzero]. - rewrite val_universe_sup, Hzero. destruct (Universe.univ_val v s2); simpl; lia. + rewrite val_universe_sup Hzero. destruct (Universe.univ_val v s2); simpl; lia. Qed. Lemma leq_universe0_sup_r φ s1 s2 : @@ -1431,7 +1529,7 @@ Section Univ. Proof. intros H1 H2 v Hv. specialize (is_prop_and_is_sprop_val_false _ H1 H2 v) as [n Hzero]. - rewrite val_universe_sup, Hzero. + rewrite val_universe_sup Hzero. destruct (Universe.univ_val v s1); simpl; lia. Qed. @@ -1530,12 +1628,6 @@ Section Univ. (** Elimination restriction *) - Inductive allowed_eliminations : Set := - | IntoSProp - | IntoPropSProp - | IntoSetPropSProp - | IntoAny. - Definition is_allowed_elimination0 φ (into : Universe.t) (allowed : allowed_eliminations) : Prop := forall v, @@ -1563,7 +1655,6 @@ Section Univ. End Univ. - (* This universe is a hack used in plugings to generate fresh universes *) Definition fresh_universe : Universe.t. exact Universe.type0. Qed. (* This level is a hack used in plugings to generate fresh levels *) @@ -1669,8 +1760,9 @@ Section UniverseClosedSubst. intro e; split; intro He. - apply Universe.map_spec in He. destruct He as [e' [He' X]]. rewrite closedu_subst_instance_level_expr in X. - now subst. apply UnivExprSet.for_all_spec in H; proper. + apply UnivExprSet.for_all_spec in H; proper. exact (H _ He'). + now subst. - apply Universe.map_spec. exists e; split; tas. symmetry; apply closedu_subst_instance_level_expr. apply UnivExprSet.for_all_spec in H; proper. now apply H. @@ -1750,7 +1842,7 @@ Definition string_of_level (l : Level.t) : string := end. Definition string_of_level_expr (e : UnivExpr.t) : string := - let '(l, b) := e in string_of_level l ^ (if b then "+1" else ""). + let '(l, n) := e in string_of_level l ^ (if n is 0 then "" else "+" ^ string_of_nat n). Definition string_of_sort (u : Universe.t) := match u with @@ -1765,6 +1857,7 @@ Definition string_of_universe_instance u := Inductive universes_entry := | Monomorphic_entry (ctx : ContextSet.t) | Polymorphic_entry (names : list name) (ctx : UContext.t). +Derive NoConfusion for universes_entry. Definition universes_entry_of_decl (u : universes_decl) : universes_entry := match u with @@ -1849,8 +1942,8 @@ Proof. - destruct v;discriminate. - apply is_prop_val with (v:=vv) in e. specialize (is_prop_and_is_sprop_val_false _ f f1 vv) as [n HH]. - rewrite e, HH. - rewrite e, HH in H2. simpl in H2. + rewrite e HH. + rewrite e HH in H2. simpl in H2. simpl. destruct (Universe.univ_val vv u') eqn:eq; simpl; auto. - destruct v';discriminate. - apply is_prop_val with (v:=vv) in f. @@ -1858,13 +1951,13 @@ Proof. rewrite f in H2. eapply univ_le_prop_inv in H2. cong. - apply is_sprop_val with (v:=vv) in e1. specialize (is_prop_and_is_sprop_val_false _ f f1 vv) as [n HH]. - rewrite HH, e1 in H2. now simpl in H2. + rewrite HH e1 in H2. now simpl in H2. - apply is_sprop_val with (v:=vv) in f1. specialize (is_prop_and_is_sprop_val_false _ e e1 vv) as [n HH]. rewrite f1 in H2. apply univ_le_sprop_inv in H2; cong. - specialize (is_prop_and_is_sprop_val_false _ e e1 vv) as [n HH]. specialize (is_prop_and_is_sprop_val_false _ f f1 vv) as [n' HH']. - rewrite HH, HH'. apply univ_val_max_mon; auto. + rewrite HH HH'. apply univ_val_max_mon; auto. now rewrite <- HH, <- HH'. Qed. @@ -1916,7 +2009,7 @@ Section UniverseLemmas. intros s1 s1' H1 s2 s2' H2. unfold_eq_universe. specialize (H1 v Hv). specialize (H2 v Hv). rewrite !val_universe_sup. - now rewrite H1, H2. + now rewrite H1 H2. Qed. Lemma sort_of_product_twice u s : @@ -1925,7 +2018,7 @@ Section UniverseLemmas. Proof. destruct u,s;auto. unfold Universe.sort_of_product;cbn. - now rewrite sup0_assoc, sup0_idem. + now rewrite sup0_assoc sup0_idem. Qed. End UniverseLemmas. @@ -2057,7 +2150,7 @@ Section no_prop_leq_type. Proof. red. destruct check_univs; [|trivial]. intros H1 H2 v Hv. eapply is_prop_val in H1; eapply is_prop_val in H2. - rewrite H1, H2. lled; lia. + rewrite -> H1, H2. lled; lia. Qed. Lemma leq_sprop_sprop {l l'} : @@ -2066,7 +2159,7 @@ Section no_prop_leq_type. Proof. red. destruct check_univs; [|trivial]. intros H1 H2 v Hv. eapply is_sprop_val in H1; eapply is_sprop_val in H2. - rewrite H1, H2. lled; lia. + rewrite -> H1, H2. lled; lia. Qed. Lemma leq_prop_is_prop_sprop {x s} : @@ -2084,7 +2177,7 @@ Section no_prop_leq_type. consistent ϕ -> leq_universe ϕ (Universe.super l) l' -> Universe.is_prop l' -> False. Proof. - intros Hcf [v Hv] H1 H2. unfold leq_universe in *; rewrite Hcf in *. + intros Hcf [v Hv] H1 H2. rewrite /leq_universe Hcf in H1. eapply is_prop_val with (v:=v) in H2. specialize (H1 _ Hv). rewrite H2 in H1. destruct l as [| |]; destruct l'; lled; cbn -[Z.add] in *; lia. Qed. @@ -2094,7 +2187,7 @@ Section no_prop_leq_type. consistent ϕ -> leq_universe ϕ (Universe.super l) l' -> Universe.is_sprop l' -> False. Proof. - intros Hcf [v Hv] H1 H2. unfold leq_universe in *; rewrite Hcf in *. + intros Hcf [v Hv] H1 H2. rewrite /leq_universe Hcf in H1. eapply is_sprop_val with (v:=v) in H2. specialize (H1 _ Hv). rewrite H2 in H1. destruct l as [| |]; destruct l'; lled; cbn -[Z.add] in *; lia. Qed. diff --git a/checker/theories/WcbvEval.v b/template-coq/theories/WcbvEval.v similarity index 93% rename from checker/theories/WcbvEval.v rename to template-coq/theories/WcbvEval.v index 6aa05b0b9..ad96f1edf 100644 --- a/checker/theories/WcbvEval.v +++ b/template-coq/theories/WcbvEval.v @@ -71,26 +71,6 @@ Definition isConstruct t := | _ => false end. -Definition isAssRel Γ x := - match x with - | tRel i => - match option_map decl_body (nth_error Γ i) with - | Some None => true - | _ => false - end - | _ => false - end. - -Definition isAxiom Σ x := - match x with - | tConst c u => - match lookup_env Σ c with - | Some (ConstantDecl {| cst_body := None |}) => true - | _ => false - end - | _ => false - end. - Definition isStuckFix t args := match t with | tFix mfix idx => @@ -128,27 +108,18 @@ Section Wcbv. eval (subst10 b0' b1) res -> eval (tLetIn na b0 t b1) res - (** Local variables: defined or undefined *) + (** Local variable with bodies *) | eval_rel_def i body res : option_map decl_body (nth_error Γ i) = Some (Some body) -> eval (lift0 (S i) body) res -> eval (tRel i) res - | eval_rel_undef i : - option_map decl_body (nth_error Γ i) = Some None -> - eval (tRel i) (tRel i) - (** Constant unfolding *) | eval_delta c decl body (isdecl : declared_constant Σ c decl) u res : decl.(cst_body) = Some body -> eval (subst_instance_constr u body) res -> eval (tConst c u) res - (** Axiom *) - | eval_axiom c decl (isdecl : declared_constant Σ c decl) u : - decl.(cst_body) = None -> - eval (tConst c u) (tConst c u) - (** Case *) | eval_iota ind pars r discr c u args p brs res : eval discr (mkApps (tConstruct ind c u) args) -> @@ -223,14 +194,11 @@ Section Wcbv. (forall (i : nat) (body res : term), option_map decl_body (nth_error Γ i) = Some (Some body) -> eval ((lift0 (S i)) body) res -> P ((lift0 (S i)) body) res -> P (tRel i) res) -> - (forall i : nat, option_map decl_body (nth_error Γ i) = Some None -> P (tRel i) (tRel i)) -> (forall c (decl : constant_body) (body : term), declared_constant Σ c decl -> forall (u : Instance.t) (res : term), cst_body decl = Some body -> eval (subst_instance_constr u body) res -> P (subst_instance_constr u body) res -> P (tConst c u) res) -> - (forall c (decl : constant_body), - declared_constant Σ c decl -> forall u : Instance.t, cst_body decl = None -> P (tConst c u) (tConst c u)) -> (forall (ind : inductive) (pars : nat) r (discr : term) (c : nat) (u : Instance.t) (args : list term) (p : term) (brs : list (nat × term)) (res : term), eval discr (mkApps (tConstruct ind c u) args) -> @@ -277,15 +245,13 @@ Section Wcbv. All2 eval a a' -> All2 P a a' -> P (tApp f a) (mkApps f' a')) -> (forall t : term, atom t -> P t t) -> forall t t0 : term, eval t t0 -> P t t0. Proof. - intros P Hbeta Hlet Hreldef Hrelvar Hcst Hax Hcase Hproj Hfix Hstuckfix Hcofixcase Hcofixproj Happcong Hatom. + intros P Hbeta Hlet Hreldef Hcst Hcase Hproj Hfix Hstuckfix Hcofixcase Hcofixproj Happcong Hatom. fix eval_evals_ind 3. destruct 1; try solve [match goal with [ H : _ |- _ ] => match type of H with forall t t0, eval t t0 -> _ => fail 1 | _ => eapply H end end; eauto]. - - eapply Hrelvar, e. - - eapply Hax; [eapply isdecl|eapply e]. - eapply Hfix; eauto. clear -a eval_evals_ind. revert args argsv a. @@ -308,7 +274,7 @@ Section Wcbv. *) Definition value_head x := - isConstruct x || isCoFix x || isAssRel Γ x || isAxiom Σ x. + isConstruct x || isCoFix x. (* Lemma value_head_atom x : value_head x -> atom x. *) (* Proof. destruct x; auto. Qed. *) @@ -390,12 +356,6 @@ Section Wcbv. Proof. intros eve. induction eve using eval_evals_ind; simpl; intros; auto using value. - - eapply (value_app (tRel i) []). now rewrite /value_head /= H. constructor. - - eapply (value_app (tConst c u) []). - red in H. - rewrite /value_head /= H. - destruct decl as [? [b|] ?]; try discriminate. - constructor. constructor. - eapply value_stuck_fix. + apply All_app_inv. * now eapply value_mkApps_values. @@ -433,17 +393,8 @@ Section Wcbv. eval t t. Proof. destruct t; intros vt nt; try discriminate. - * constructor. - unfold value_head in vt. simpl in vt. destruct option_map as [[o|]|] => //. - * unfold value_head in vt. simpl in vt. - destruct lookup_env eqn:Heq => //. - destruct g eqn:Heq' => //. - destruct c0 as [? [b|] ?] eqn:Heq'' => //. subst. - intros. eapply eval_axiom. red. - now rewrite Heq. - easy. - * now eapply eval_atom. - * now eapply eval_atom. + - now eapply eval_atom. + - now eapply eval_atom. Qed. Lemma value_final e : value e -> eval e e. @@ -584,12 +535,11 @@ Section Wcbv. ∑ decl, declared_constant Σ c decl * match cst_body decl with | Some body => eval (subst_instance_constr u body) v - | None => v = tConst c u + | None => False end. Proof. intros H; depind H; try solve_discr'; try now easy. - exists decl. intuition auto. now rewrite e. - - exists decl. intuition auto. now rewrite e. - symmetry in H1. eapply mkApps_nisApp in H1 => //; intuition subst; auto. depelim a. eapply eval_to_stuck_fix_inv in H; [|easy]. diff --git a/template-coq/theories/WfInv.v b/template-coq/theories/WfInv.v index 7b1921c7e..9fff63947 100644 --- a/template-coq/theories/WfInv.v +++ b/template-coq/theories/WfInv.v @@ -7,7 +7,7 @@ From MetaCoq.Template Require Import config utils Ast AstUtils. Fixpoint wf_Inv (t : term) := match t with - | tRel _ | tVar _ | tSort _ => True + | tRel _ | tVar _ | tSort _ | tInt _ | tFloat _ => True | tEvar n l => Forall wf l | tCast t k t' => wf t /\ wf t' | tProd na t b => wf t /\ wf b @@ -17,7 +17,7 @@ Fixpoint wf_Inv (t : term) := | tConst _ _ | tInd _ _ | tConstruct _ _ _ => True | tCase ci p c brs => wf p /\ wf c /\ Forall (wf ∘ snd) brs | tProj p t => wf t - | tFix mfix k => Forall (fun def => wf def.(dtype) /\ wf def.(dbody) /\ isLambda def.(dbody) = true) mfix + | tFix mfix k => Forall (fun def => wf def.(dtype) /\ wf def.(dbody)) mfix | tCoFix mfix k => Forall (fun def => wf def.(dtype) /\ wf def.(dbody)) mfix end. @@ -61,7 +61,7 @@ Qed. Fixpoint wf_term (t : term) : bool := match t with - | tRel _ | tVar _ => true + | tRel _ | tVar _ | tInt _ | tFloat _ => true | tEvar n l => forallb wf_term l | tSort u => true | tCast t k t' => wf_term t && wf_term t' diff --git a/template-coq/theories/monad_utils.v b/template-coq/theories/monad_utils.v index c7bed7d9c..dc52a1d9a 100644 --- a/template-coq/theories/monad_utils.v +++ b/template-coq/theories/monad_utils.v @@ -1,7 +1,12 @@ (* todo(gmm): This file should really be replaced by a real * monad library. *) -Require Import List. +Require Import Arith List. +From MetaCoq.Template Require Import All_Forall MCSquash. +From Equations Require Import Equations. +Coercion is_true : bool >-> Sortclass. + +Import ListNotations. Set Universe Polymorphism. @@ -106,3 +111,115 @@ End MonadOperations. Definition monad_iter {T : Type -> Type} {M A} (f : A -> T unit) (l : list A) : T unit := @monad_fold_left T M _ _ (fun _ => f) l tt. + +Fixpoint monad_All {T} {M : Monad T} {A} {P} (f : forall x, T (P x)) l : T (@All A P l) := match l with + | [] => ret All_nil + | a :: l => X <- f a ;; + Y <- monad_All f l ;; + ret (All_cons X Y) + end. + +Fixpoint monad_All2 {T E} {M : Monad T} {M' : MonadExc E T} wrong_sizes + {A B R} (f : forall x y, T (R x y)) l1 l2 : T (@All2 A B R l1 l2) := + match l1, l2 with + | [], [] => ret All2_nil + | a :: l1, b :: l2 => X <- f a b ;; + Y <- monad_All2 wrong_sizes f l1 l2 ;; + ret (All2_cons X Y) + | _, _ => raise wrong_sizes + end. + +Definition monad_prod {T} {M : Monad T} {A B} (x : T A) (y : T B): T (A * B)%type + := X <- x ;; Y <- y ;; ret (X, Y). + +(** monadic checks *) +Definition check_dec {T E} {M : Monad T} {M' : MonadExc E T} (e : E) {P} + (H : {P} + {~ P}) : T P + := match H with + | left x => ret x + | right _ => raise e + end. + +Definition check_eq_true {T E} {M : Monad T} {M' : MonadExc E T} (b : bool) (e : E) : T b := + if b return T b then ret eq_refl else raise e. + +Definition check_eq_nat {T E} {M : Monad T} {M' : MonadExc E T} n m (e : E) : T (n = m) := + match PeanoNat.Nat.eq_dec n m with + | left p => ret p + | right p => raise e + end. + +Program Fixpoint monad_Alli {T} {M : Monad T} {A} {P} (f : forall n x, T (∥ P n x ∥)) l n + : T (∥ @Alli A P n l ∥) + := match l with + | [] => ret (sq Alli_nil) + | a :: l => X <- f n a ;; + Y <- monad_Alli f l (S n) ;; + ret _ + end. +Next Obligation. + sq. constructor; assumption. +Defined. + +Program Fixpoint monad_Alli_All {T} {M : Monad T} {A} {P} {Q} (f : forall n x, ∥ Q x ∥ -> T (∥ P n x ∥)) l n : + ∥ All Q l ∥ -> T (∥ @Alli A P n l ∥) + := match l with + | [] => fun _ => ret (sq Alli_nil) + | a :: l => fun allq => X <- f n a _ ;; + Y <- monad_Alli_All f l (S n) _ ;; + ret _ + end. +Next Obligation. sq. + now depelim allq. +Qed. +Next Obligation. + sq; now depelim allq. +Qed. +Next Obligation. + sq. constructor; assumption. +Defined. + +Section monad_Alli_nth. + Context {T} {M : Monad T} {A} {P : nat -> A -> Type}. + Program Fixpoint monad_Alli_nth_gen l k + (f : forall n x, nth_error l n = Some x -> T (∥ P (k + n) x ∥)) : + T (∥ @Alli A P k l ∥) + := match l with + | [] => ret (sq Alli_nil) + | a :: l => X <- f 0 a _ ;; + Y <- monad_Alli_nth_gen l (S k) (fun n x hnth => px <- f (S n) x hnth;; ret _) ;; + ret _ + end. + Next Obligation. + sq. now rewrite Nat.add_succ_r in px. + Qed. + Next Obligation. + sq. rewrite Nat.add_0_r in X. constructor; auto. + Qed. + + Definition monad_Alli_nth l (f : forall n x, nth_error l n = Some x -> T (∥ P n x ∥)) : T (∥ @Alli A P 0 l ∥) := + monad_Alli_nth_gen l 0 f. + +End monad_Alli_nth. + +Section MonadAllAll. + Context {T} {M : Monad T} {A} {P : A -> Type} {Q} (f : forall x, ∥ Q x ∥ -> T (∥ P x ∥)). + Program Fixpoint monad_All_All l : ∥ All Q l ∥ -> T (∥ All P l ∥) := + match l return ∥ All Q l ∥ -> T (∥ All P l ∥) with + | [] => fun _ => ret (sq All_nil) + | a :: l => fun allq => + X <- f a _ ;; + Y <- monad_All_All l _ ;; + ret _ + end. + Next Obligation. sq. + now depelim allq. + Qed. + Next Obligation. + sq; now depelim allq. + Qed. + Next Obligation. + sq. constructor; assumption. + Defined. +End MonadAllAll. + diff --git a/template-coq/theories/utils/All_Forall.v b/template-coq/theories/utils/All_Forall.v index 35add2c54..f1ce69c67 100644 --- a/template-coq/theories/utils/All_Forall.v +++ b/template-coq/theories/utils/All_Forall.v @@ -3,10 +3,7 @@ From MetaCoq.Template Require Import MCPrelude MCList MCRelations MCProd MCOptio From Equations Require Import Equations. Import ListNotations. -Local Ltac inv H := inversion_clear H. -Local Coercion is_true : bool >-> Sortclass. - -Derive Signature for Forall. +Derive Signature for Forall Forall2. (** Combinators *) @@ -99,10 +96,35 @@ Lemma forallb_Forall {A} (p : A -> bool) l Proof. split. induction 1; rewrite /= // H IHForall //. - induction l; rewrite /= //. move/andP => [pa pl]. + induction l; rewrite /= //. rewrite andb_and. + intros [pa pl]. constructor; auto. Qed. +Lemma forallbP {A} (P : A -> Prop) (p : A -> bool) l : + (forall x, reflect (P x) (p x)) -> + reflect (Forall P l) (forallb p l). +Proof. + intros Hp. + apply iff_reflect; change (forallb p l = true) with (forallb p l : Prop); split. + - induction 1; rewrite /= // IHForall // andb_true_r. + now destruct (Hp x). + - induction l; rewrite /= //. rewrite andb_and. + intros [pa pl]. + constructor; auto. now destruct (Hp a). +Qed. + +Lemma forallbP_cond {A} (P Q : A -> Prop) (p : A -> bool) l : + Forall Q l -> + (forall x, Q x -> reflect (P x) (p x)) -> reflect (Forall P l) (forallb p l). +Proof. + intros HQ Hp. + apply iff_reflect; split. + - induction HQ; intros HP; depelim HP; rewrite /= // IHHQ // andb_true_r. + now destruct (Hp x H). + - induction HQ; rewrite /= //. move/andb_and => [pa pl]. + constructor; auto. now destruct (Hp _ H). +Qed. Lemma map_eq_inj {A B} (f g : A -> B) l: map f l = map g l -> All (fun x => f x = g x) l. @@ -162,7 +184,7 @@ Lemma forallb2_app {A} (p : A -> A -> bool) l l' q q' : -> is_true (forallb2 p (l ++ q) (l' ++ q')). Proof. induction l in l' |- *; destruct l'; simpl; try congruence. - move=> /andP[/andP[pa pl] pq]. now rewrite pa IHl // pl pq. + move=> /andb_and[/andb_and[pa pl] pq]. now rewrite pa IHl // pl pq. Qed. Lemma All2_map {A B C D} (R : C -> D -> Type) (f : A -> C) (g : B -> D) l l' : @@ -440,18 +462,18 @@ Qed. Lemma Alli_rev {A} {P : nat -> A -> Type} k l : Alli P k l -> - Alli (fun k' => P (pred #|l| - k' + k)) 0 (List.rev l). + Alli (fun k' => P (Nat.pred #|l| - k' + k)) 0 (List.rev l). Proof. revert k. induction l using rev_ind; simpl; intros; try constructor. eapply Alli_app in X. intuition. rewrite rev_app_distr. rewrite app_length. simpl. constructor. - replace (pred (#|l| + 1) - 0) with #|l| by lia. + replace (Nat.pred (#|l| + 1) - 0) with #|l| by lia. inversion b. eauto. specialize (IHl _ a). eapply Alli_shift. eapply Alli_impl. eauto. simpl; intros. - now replace (pred (#|l| + 1) - S n) with (pred #|l| - n) by lia. + now replace (Nat.pred (#|l| + 1) - S n) with (Nat.pred #|l| - n) by lia. Qed. @@ -1654,8 +1676,8 @@ Lemma forallb_nth {A} (l : list A) (n : nat) P d : forallb P l -> n < #|l| -> exists x, (nth n l d = x) /\ P x. Proof. induction l in n |- *; destruct n; simpl; auto; try easy. - move/andP => [pa pl] pn. exists a; easy. - move/andP => [pa pl] pn. specialize (IHl n pl). + move/andb_and => [pa pl] pn. exists a; easy. + move/andb_and => [pa pl] pn. specialize (IHl n pl). apply IHl; lia. Qed. @@ -1888,7 +1910,7 @@ Proof. - destruct l'. 2: discriminate. constructor. - destruct l'. 1: discriminate. - simpl in h. apply andP in h as [? ?]. + simpl in h. move/andb_and: h => [? ?]. constructor. all: auto. Qed. @@ -1929,3 +1951,15 @@ Proof. now apply H; simpl. apply IHl. intros x xin; apply H; simpl; auto. Qed. + +Lemma All_forall {X Y} (f:X->Y->Prop) xs: + All (fun a => forall b, f a b) xs -> + (forall b, All (fun a => f a b) xs). +Proof. + intros. + induction X0. + - constructor. + - constructor. + + apply p. + + apply IHX0. +Qed. \ No newline at end of file diff --git a/template-coq/theories/utils/MCList.v b/template-coq/theories/utils/MCList.v index 626f3c2d6..79094596f 100644 --- a/template-coq/theories/utils/MCList.v +++ b/template-coq/theories/utils/MCList.v @@ -236,6 +236,16 @@ Proof. destruct l; simpl in *; congruence. Qed. +Lemma mapi_rec_compose {A B C} (g : nat -> B -> C) (f : nat -> A -> B) k l : + mapi_rec g (mapi_rec f l k) k = mapi_rec (fun k x => g k (f k x)) l k. +Proof. + induction l in k |- *; simpl; auto. now rewrite IHl. +Qed. + +Lemma mapi_compose {A B C} (g : nat -> B -> C) (f : nat -> A -> B) l : + mapi g (mapi f l) = mapi (fun k x => g k (f k x)) l. +Proof. apply mapi_rec_compose. Qed. + Lemma map_ext {A B : Type} (f g : A -> B) (l : list A) : (forall x, f x = g x) -> map f l = map g l. @@ -488,8 +498,6 @@ Proof. eapply IHl1 in H2 as []. split; congruence. Qed. -Arguments skipn : simpl nomatch. - Lemma skipn_all2 : forall {A n} (l : list A), #|l| <= n -> @@ -612,6 +620,34 @@ Proof. * simpl. apply IHn. Qed. +Fixpoint split_at_aux {A} (n : nat) (acc : list A) (l : list A) : list A * list A := + match n with + | 0 => (List.rev acc, l) + | S n' => + match l with + | [] => (List.rev acc, []) + | hd :: l' => split_at_aux n' (hd :: acc) l' + end + end. + +Definition split_at {A} (n : nat) (l : list A) : list A * list A := + split_at_aux n [] l. + +Lemma split_at_aux_firstn_skipn {A} n acc (l : list A) : split_at_aux n acc l = (List.rev acc ++ firstn n l, skipn n l). +Proof. + induction n in acc, l |- *; destruct l; simpl; auto. + now rewrite app_nil_r. + now rewrite app_nil_r. + now rewrite app_nil_r. + rewrite IHn. simpl. + now rewrite -app_assoc /=. +Qed. + +Lemma split_at_firstn_skipn {A} n (l : list A) : split_at n l = (firstn n l, skipn n l). +Proof. + now rewrite /split_at split_at_aux_firstn_skipn /= //. +Qed. + Lemma rev_app : forall A (l l' : list A), (rev (l ++ l') = rev l' ++ rev l)%list. @@ -744,6 +780,15 @@ Proof. reflexivity. Qed. +Lemma nth_error_rev_inv {A} (l : list A) i : + i < #|l| -> + nth_error (List.rev l) i = nth_error l (#|l| - S i). +Proof. + intros Hi. + rewrite nth_error_rev ?List.rev_length; auto. + now rewrite List.rev_involutive. +Qed. + Lemma nth_error_snoc {A} (l : list A) (a : A) (l' : list A) i : i = #|l| -> nth_error (l ++ a :: l') i = Some a. @@ -879,6 +924,97 @@ Proof. now rewrite andb_comm andb_true_r. Qed. +Fixpoint unfold {A} (n : nat) (f : nat -> A) : list A := + match n with + | 0 => [] + | S n => unfold n f ++ [f n] + end. + +Lemma mapi_irrel_list {A B} (f : nat -> A) (l l' : list B) : + #|l| = #|l'| -> + mapi (fun i (x : B) => f i) l = mapi (fun i x => f i) l'. +Proof. + induction l in f, l' |- *; destruct l' => //; simpl; auto. + intros [= eq]. unfold mapi. simpl. f_equal. + rewrite !mapi_rec_Sk. + now rewrite [mapi_rec _ _ _](IHl (fun x => (f (S x))) l'). +Qed. + +Lemma mapi_unfold {A B} (f : nat -> B) l : mapi (fun i (x : A) => f i) l = unfold #|l| f. +Proof. + unfold mapi. + induction l in f |- *; simpl; auto. + rewrite mapi_rec_Sk. + rewrite -IHl. rewrite -(mapi_rec_Sk (fun i x => f i) l 0). + change [f #|l|] with (mapi_rec (fun i x => f i) [a] #|l|). + rewrite -(Nat.add_0_r #|l|). rewrite -mapi_rec_app. + change (f 0 :: _) with (mapi (fun i x => f i) (a :: l)). + apply mapi_irrel_list. simpl. rewrite app_length /=; lia. +Qed. + +Lemma unfold_length {A} (f : nat -> A) m : #|unfold m f| = m. +Proof. + induction m; simpl; rewrite ?app_length /=; auto. lia. +Qed. + +Lemma nth_error_unfold {A} (f : nat -> A) m n : n < m <-> nth_error (unfold m f) n = Some (f n). +Proof. + induction m in n |- *; split; intros Hn; try lia. + - simpl in Hn. rewrite nth_error_nil in Hn. discriminate. + - destruct (Classes.eq_dec n m); [subst|]. + * simpl. rewrite nth_error_app_ge unfold_length // Nat.sub_diag /= //. + * simpl. rewrite nth_error_app_lt ?unfold_length //; try lia. + apply IHm; lia. + - simpl in Hn. eapply nth_error_Some_length in Hn. + rewrite app_length /= unfold_length in Hn. lia. +Qed. + +Lemma nth_error_unfold_inv {A} (f : nat -> A) m n t : nth_error (unfold m f) n = Some t -> t = (f n). +Proof. + induction m in n |- *; intros Hn; try lia. + - simpl in Hn. rewrite nth_error_nil in Hn. discriminate. + - simpl in Hn. + pose proof (nth_error_Some_length Hn). + rewrite app_length /= unfold_length in H. + destruct (Classes.eq_dec n m); [subst|]. + * simpl. revert Hn. rewrite nth_error_app_ge unfold_length // Nat.sub_diag /= //; congruence. + * simpl. revert Hn. rewrite nth_error_app_lt ?unfold_length //; try lia. auto. +Qed. + +Lemma In_unfold_inj {A} (f : nat -> A) n i : + (forall i j, f i = f j -> i = j) -> + In (f i) (unfold n f) -> i < n. +Proof. + intros hf. + induction n; simpl => //. + intros H; apply in_app_or in H. + destruct H. + - specialize (IHn H). lia. + - simpl in H. destruct H; [apply hf in H|]. + * subst. auto. + * destruct H. +Qed. + +Lemma forallb_unfold {A} (f : A -> bool) (g : nat -> A) n : + (forall x, x < n -> f (g x)) -> + forallb f (unfold n g). +Proof. + intros fg. + induction n; simpl; auto. + rewrite forallb_app IHn //. + intros x lt; rewrite fg //. lia. + rewrite /= fg //. +Qed. + +Lemma forallb_mapi {A B} (p : B -> bool) (f : nat -> B) l : + (forall i, i < #|l| -> p (f i)) -> + forallb p (mapi (fun i (x : A) => f i) l). +Proof. + intros Hp. rewrite (mapi_unfold f). + induction #|l| in *; simpl; auto. + rewrite forallb_app. simpl. now rewrite Hp // !andb_true_r. +Qed. + Lemma fold_left_andb_forallb {A} P l x : fold_left (fun b x => P x && b) l (P x) = @forallb A P (x :: l). Proof. diff --git a/template-coq/theories/utils/MCPrelude.v b/template-coq/theories/utils/MCPrelude.v index 4760a85b8..5b8768dea 100644 --- a/template-coq/theories/utils/MCPrelude.v +++ b/template-coq/theories/utils/MCPrelude.v @@ -1,7 +1,14 @@ -Require Import Lia. +Require Import String ZArith Lia. +From Equations Require Import Equations. +Set Equations Transparent. + +Derive NoConfusion EqDec for Strings.Ascii.ascii string positive Z. Declare Scope metacoq_scope. +(** We cannot use ssrbool currently as it breaks extraction. *) +Coercion is_true : bool >-> Sortclass. + Notation "'eta_compose'" := (fun g f x => g (f x)). (* \circ *) @@ -26,6 +33,7 @@ Notation "x .π2" := (@projT2 _ _ x) (at level 3, format "x '.π2'"). Create HintDb terms. Ltac arith_congr := repeat (try lia; progress f_equal). +Ltac lia_f_equal := repeat (lia || f_equal). Ltac easy0 := let rec use_hyp H := @@ -54,4 +62,6 @@ Hint Extern 10 (_ < _)%nat => lia : terms. Hint Extern 10 (_ <= _)%nat => lia : terms. Hint Extern 10 (@eq nat _ _) => lia : terms. -Ltac easy ::= easy0 || solve [intuition eauto 3 with core terms]. \ No newline at end of file +Ltac easy ::= easy0 || solve [intuition eauto 3 with core terms]. + +Ltac inv H := inversion_clear H. diff --git a/template-coq/theories/utils/MCProd.v b/template-coq/theories/utils/MCProd.v index 2e32a2f2c..5631a042c 100644 --- a/template-coq/theories/utils/MCProd.v +++ b/template-coq/theories/utils/MCProd.v @@ -1,4 +1,7 @@ -From Coq Require Import Bool. +From Coq Require Import Bool RelationClasses. +Require Import ssreflect. + +Local Coercion is_true : bool >-> Sortclass. Declare Scope pair_scope. @@ -46,11 +49,12 @@ Qed. Definition on_pi2 {A B C} (f : B -> B) (p : A * B * C) : A * B * C := (fst (fst p), f (snd (fst p)), snd p). - -Lemma andb_and b b' : is_true (b && b') <-> is_true b /\ is_true b'. +(** It would be tempting to import ssrbool here, however + https://github.com/coq/coq/issues/13486 prevents this. *) +Lemma andb_and b b' : b && b' <-> b /\ b'. Proof. apply andb_true_iff. Qed. -Lemma andP {b b'} : is_true (b && b') -> is_true b /\ is_true b'. +Lemma andb_andI {b b'} : b && b' -> b /\ b'. Proof. apply andb_and. Qed. Definition fst_eq {A B} {x x' : A} {y y' : B} diff --git a/template-coq/theories/utils/MCRelations.v b/template-coq/theories/utils/MCRelations.v index 00f5877b4..f24839fc0 100644 --- a/template-coq/theories/utils/MCRelations.v +++ b/template-coq/theories/utils/MCRelations.v @@ -1,3 +1,6 @@ +(* Distributed under the terms of the MIT license. *) +Require Import ssreflect. +Require Import Equations.Type.Relation Equations.Type.Relation_Properties. Require Import CRelationClasses. Infix "<~>" := iffT (at level 90). @@ -16,6 +19,163 @@ Notation Trel_conj R S := Notation on_Trel_eq R f g := (fun x y => (R (f x) (f y) * (g x = g y)))%type. +Section Flip. + Local Set Universe Polymorphism. + Context {A : Type} (R : crelation A). + + Lemma flip_Reflexive : Reflexive R -> Reflexive (flip R). + Proof. + intros HR x. unfold flip. apply reflexivity. + Qed. + + Lemma flip_Symmetric : Symmetric R -> Symmetric (flip R). + Proof. + intros HR x y. unfold flip. apply symmetry. + Qed. + + Lemma flip_Transitive : Transitive R -> Transitive (flip R). + Proof. + intros HR x y z xy yz. + unfold flip in *. eapply HR; eassumption. + Qed. + +End Flip. + +Definition commutes {A} (R S : relation A) := + forall x y z, R x y -> S x z -> { w & S y w * R z w}%type. + +Lemma clos_t_rt {A} {R : A -> A -> Type} x y : trans_clos R x y -> clos_refl_trans R x y. +Proof. + induction 1; try solve [econstructor; eauto]. +Qed. + + +Arguments rt_step {A} {R} {x y}. +Polymorphic Hint Resolve rt_refl rt_step : core. + + +Definition clos_rt_monotone {A} (R S : relation A) : + inclusion R S -> inclusion (clos_refl_trans R) (clos_refl_trans S). +Proof. + move => incls x y. + induction 1; solve [econstructor; eauto]. +Qed. + +Lemma relation_equivalence_inclusion {A} (R S : relation A) : + inclusion R S -> inclusion S R -> relation_equivalence R S. +Proof. firstorder. Qed. + +Lemma clos_rt_disjunction_left {A} (R S : relation A) : + inclusion (clos_refl_trans R) + (clos_refl_trans (relation_disjunction R S)). +Proof. + apply clos_rt_monotone. + intros x y H; left; exact H. +Qed. + +Lemma clos_rt_disjunction_right {A} (R S : relation A) : + inclusion (clos_refl_trans S) + (clos_refl_trans (relation_disjunction R S)). +Proof. + apply clos_rt_monotone. + intros x y H; right; exact H. +Qed. + +Global Instance clos_rt_trans A R : Transitive (@clos_refl_trans A R). +Proof. + intros x y z H H'. econstructor 3; eauto. +Qed. + +Global Instance clos_rt_refl A R : Reflexive (@clos_refl_trans A R). +Proof. intros x. constructor 2. Qed. + +Lemma clos_refl_trans_prod_l {A B} (R : relation A) (S : relation (A * B)) : + (forall x y b, R x y -> S (x, b) (y, b)) -> + forall (x y : A) b, + clos_refl_trans R x y -> + clos_refl_trans S (x, b) (y, b). +Proof. + intros. induction X0; try solve [econstructor; eauto]. +Qed. + +Lemma clos_refl_trans_prod_r {A B} (R : relation B) (S : relation (A * B)) a : + (forall x y, R x y -> S (a, x) (a, y)) -> + forall (x y : B), + clos_refl_trans R x y -> + clos_refl_trans S (a, x) (a, y). +Proof. + intros. induction X0; try solve [econstructor; eauto]. +Qed. + +Lemma clos_rt_t_incl {A} {R : relation A} `{Reflexive A R} : + inclusion (clos_refl_trans R) (trans_clos R). +Proof. + intros x y. induction 1; try solve [econstructor; eauto]. +Qed. + +Lemma clos_t_rt_incl {A} {R : relation A} `{Reflexive A R} : + inclusion (trans_clos R) (clos_refl_trans R). +Proof. + intros x y. induction 1; try solve [econstructor; eauto]. +Qed. + +Lemma clos_t_rt_equiv {A} {R} `{Reflexive A R} : + relation_equivalence (trans_clos R) (clos_refl_trans R). +Proof. + apply relation_equivalence_inclusion. + apply clos_t_rt_incl. + apply clos_rt_t_incl. +Qed. + +Global Instance relation_disjunction_refl_l {A} {R S : relation A} : + Reflexive R -> Reflexive (relation_disjunction R S). +Proof. + intros HR x. left; auto. +Qed. + +Global Instance relation_disjunction_refl_r {A} {R S : relation A} : + Reflexive S -> Reflexive (relation_disjunction R S). +Proof. + intros HR x. right; auto. +Qed. + +Global Instance clos_rt_trans_Symmetric A R : + Symmetric R -> Symmetric (@clos_refl_trans A R). +Proof. + intros X x y H. induction H; eauto. + eapply clos_rt_trans; eassumption. +Qed. + +Definition clos_sym {A} (R : relation A) := relation_disjunction R (flip R). + +Global Instance clos_sym_Symmetric A R : + Symmetric (@clos_sym A R). +Proof. + intros x y []; [right|left]; assumption. +Qed. + +Global Instance clos_refl_sym_trans_Symmetric A R : + Symmetric (@clos_refl_sym_trans A R) + := rst_sym R. + +Global Instance clos_refl_sym_trans_Reflexive A R : + Reflexive (@clos_refl_sym_trans A R) + := rst_refl R. + +Global Instance clos_refl_sym_trans_Transitive A R : + Transitive (@clos_refl_sym_trans A R) + := rst_trans R. + +Global Instance relation_disjunction_Symmetric {A} {R S : relation A} : + Symmetric R -> Symmetric S -> Symmetric (relation_disjunction R S). +Proof. + intros HR HS x y [X|X]; [left|right]; eauto. +Qed. + +Ltac rst_induction h := + induction h; [constructor|reflexivity|etransitivity]. + + (* Definition MR {T M} (R : M -> M -> Prop) (m : T -> M) (x y: T): Prop := R (m x) (m y). *) (* From measure_wf of Program.Wf *) diff --git a/template-coq/theories/utils/MCSquash.v b/template-coq/theories/utils/MCSquash.v index 3169de128..20749333e 100644 --- a/template-coq/theories/utils/MCSquash.v +++ b/template-coq/theories/utils/MCSquash.v @@ -4,6 +4,11 @@ Record squash (A : Type) : Prop := sq { _ : A }. Notation "∥ T ∥" := (squash T) (at level 10). Arguments sq {_} _. +Lemma map_squash {A B} (f : A -> B) : ∥ A ∥ -> ∥ B ∥. +Proof. + intros []; constructor; auto. +Qed. + Ltac sq := repeat match goal with | H : ∥ _ ∥ |- _ => case H; clear H; intro H diff --git a/template-coq/theories/utils/MCString.v b/template-coq/theories/utils/MCString.v index 2784e436f..c395a6941 100644 --- a/template-coq/theories/utils/MCString.v +++ b/template-coq/theories/utils/MCString.v @@ -1,4 +1,4 @@ -From Coq Require Import String. +From Coq Require Import String Decimal DecimalString ZArith. From MetaCoq.Template Require Import MCCompare. Local Open Scope string_scope. @@ -26,62 +26,19 @@ Definition parens (top : bool) (s : string) := if top then s else "(" ++ s ++ ")". Definition string_of_nat n : string := - match n with - | 0 => "0" - | 1 => "1" - | 2 => "2" - | 3 => "3" - | 4 => "4" - | 5 => "5" - | 6 => "6" - | 7 => "7" - | 8 => "8" - | 9 => "9" - | 10 => "10" - | 11 => "11" - | 12 => "12" - | 13 => "13" - | 14 => "14" - | 15 => "15" - | 16 => "16" - | 17 => "17" - | 18 => "18" - | 19 => "19" - | 20 => "20" - | 21 => "21" - | 22 => "22" - | 23 => "23" - | 24 => "24" - | 25 => "25" - | 26 => "26" - | 27 => "27" - | 28 => "28" - | 29 => "29" - | 30 => "30" - | 31 => "31" - | 32 => "32" - | 33 => "33" - | 34 => "34" - | 35 => "35" - | 36 => "36" - | 37 => "37" - | 38 => "38" - | 39 => "39" - | 40 => "40" - | 41 => "41" - | 42 => "42" - | 43 => "43" - | 44 => "44" - | 45 => "45" - | 46 => "46" - | 47 => "47" - | 48 => "48" - | 49 => "49" - | _ => "todo string_of_nat" - end. + DecimalString.NilEmpty.string_of_uint (Nat.to_uint n). Hint Resolve String.string_dec : eq_dec. +Definition string_of_positive p := + string_of_nat (Pos.to_nat p). + +Definition string_of_Z (z : Z) : string := + match z with + | Z0 => "0" + | Zpos p => string_of_positive p + | Zneg p => "-" ++ string_of_positive p + end. Definition eq_string s s' := match string_compare s s' with diff --git a/template-coq/theories/utils/MCUtils.v b/template-coq/theories/utils/MCUtils.v index be3d02c58..e6ebf182f 100644 --- a/template-coq/theories/utils/MCUtils.v +++ b/template-coq/theories/utils/MCUtils.v @@ -51,9 +51,6 @@ Ltac rdest := | |- sigT _ => eexists end. -(** We cannot use ssrbool as it breaks extraction. *) -Coercion is_true : bool >-> Sortclass. - Tactic Notation "toProp" ident(H) := match type of H with | is_true (_ apply PeanoNat.Nat.ltb_lt in H diff --git a/template-coq/theories/utils/MC_ExtrOCamlInt63.v b/template-coq/theories/utils/MC_ExtrOCamlInt63.v new file mode 100644 index 000000000..a82fb2502 --- /dev/null +++ b/template-coq/theories/utils/MC_ExtrOCamlInt63.v @@ -0,0 +1,56 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) +(* bool [ true false ]. +Extract Inductive prod => "( * )" [ "" ]. +(* Extract Inductive comparison => int [ "0" "(-1)" "1" ]. *) +Extract Inductive DoubleType.carry => "Uint63.carry" [ "Uint63.C0" "Uint63.C1" ]. + +(** Primitive types and operators. *) +Extract Constant Int63.int => "Uint63.t". +Extraction Inline Int63.int. +(* Otherwise, the name conflicts with the primitive OCaml type [int] *) + +Extract Constant Int63.lsl => "Uint63.l_sl". +Extract Constant Int63.lsr => "Uint63.l_sr". +Extract Constant Int63.land => "Uint63.l_and". +Extract Constant Int63.lor => "Uint63.l_or". +Extract Constant Int63.lxor => "Uint63.l_xor". + +Extract Constant Int63.add => "Uint63.add". +Extract Constant Int63.sub => "Uint63.sub". +Extract Constant Int63.mul => "Uint63.mul". +Extract Constant Int63.mulc => "Uint63.mulc". +Extract Constant Int63.div => "Uint63.div". +Extract Constant Int63.mod => "Uint63.rem". + +Extract Constant Int63.eqb => "Uint63.equal". +Extract Constant Int63.ltb => "Uint63.lt". +Extract Constant Int63.leb => "Uint63.le". + +Extract Constant Int63.addc => "Uint63.addc". +Extract Constant Int63.addcarryc => "Uint63.addcarryc". +Extract Constant Int63.subc => "Uint63.subc". +Extract Constant Int63.subcarryc => "Uint63.subcarryc". + +Extract Constant Int63.diveucl => "Uint63.diveucl". +Extract Constant Int63.diveucl_21 => "Uint63.div21". +Extract Constant Int63.addmuldiv => "Uint63.addmuldiv". + +Extract Constant Int63.compare => "fun _x _y -> failwith ""not yet implemented""". + +Extract Constant Int63.head0 => "Uint63.head0". +Extract Constant Int63.tail0 => "Uint63.tail0". diff --git a/template-coq/theories/utils/wGraph.v b/template-coq/theories/utils/wGraph.v index 0bd96a0dc..4a4220d42 100644 --- a/template-coq/theories/utils/wGraph.v +++ b/template-coq/theories/utils/wGraph.v @@ -491,12 +491,12 @@ Module WeightedGraph (V : UsualOrderedTypeWithLeibniz) (VSet : MSetList.SWithLei Next Obligation. split. - eapply VSetProp.Add_add. - - apply andP in Hp0 as [h1 h2]. + - apply andb_andI in Hp0 as [h1 h2]. apply negb_true_iff in h1. apply VSetFact.not_mem_iff in h1. assumption. Defined. Next Obligation. - apply andP in Hp0 as [? ?]. auto. + apply andb_andI in Hp0 as [? ?]. auto. Defined. @@ -507,23 +507,6 @@ Module WeightedGraph (V : UsualOrderedTypeWithLeibniz) (VSet : MSetList.SWithLei reflexivity. specialize (IHp q); intuition. Qed. - - (* Lemma DisjointAdd_add {s s' x y} (H : DisjointAdd x s s') (H' : x <> y) *) - (* : DisjointAdd x (VSet.add y s) (VSet.add y s'). *) - (* Proof. *) - (* repeat split. 2: intros [H0|H0]. *) - (* - intro H0. apply VSet.add_spec in H0. *) - (* destruct H0 as [H0|H0]. *) - (* right; subst; apply VSet.add_spec; left; reflexivity. *) - (* apply H in H0. destruct H0 as [H0|H0]; [left; assumption |right]. *) - (* apply VSet.add_spec; right; assumption. *) - (* - subst. apply VSet.add_spec; right. apply H; left; reflexivity. *) - (* - apply VSet.add_spec in H0; apply VSet.add_spec; destruct H0 as [H0|H0]. *) - (* left; assumption. right. apply H. right; assumption. *) - (* - intro H0. apply VSet.add_spec in H0; destruct H0 as [H0|H0]. *) - (* contradiction. now apply H. *) - (* Qed. *) - Lemma DisjointAdd_add1 {s0 s1 s2 x y} (H1 : DisjointAdd x s0 s1) (H2 : DisjointAdd y s1 s2) : DisjointAdd x (VSet.add y s0) s2. diff --git a/template-coq/update_plugin.sh b/template-coq/update_plugin.sh index 1ea7d22ec..90c6e355e 100755 --- a/template-coq/update_plugin.sh +++ b/template-coq/update_plugin.sh @@ -4,7 +4,7 @@ TOCOPY="ast_denoter.ml ast_quoter.ml denoter.ml plugin_core.ml plugin_core.mli r # Test is gen-src is older than src if [[ "gen-src" -ot "src" || ! -f "gen-src/denoter.ml" || ! -f "gen-src/metacoq_template_plugin.cmxs" || - "gen-src/extractable.ml" -nt "gen-src/metacoq_template_plugin.cmxs" ]] + "gen-src/extractable.ml" -nt "gen-src/metacoq_template_plugin.cmxs" || "$1" = "force" ]] then echo "Updating gen-src from src" mkdir -p build @@ -12,10 +12,11 @@ then for x in ${TOCOPY}; do rm -f gen-src/$x; cp src/$x gen-src/$x; done echo "Renaming files to camelCase" (cd gen-src; ./to-lower.sh) - rm -f gen-src/*.d + rm -f gen-src/*.d gen-src/*.cm* # Fix an extraction bug: wrong type annotation on eq_equivalence ${SED} -i.bak 's/\r//g' gen-src/cRelationClasses.mli patch -N -p0 < extraction.patch + patch -N -p0 < specFloat.patch exit 0 else echo "Extracted code is up-to-date" diff --git a/test-suite/TypingTests.v b/test-suite/TypingTests.v deleted file mode 100644 index c32b957f9..000000000 --- a/test-suite/TypingTests.v +++ /dev/null @@ -1,164 +0,0 @@ -Require Import Recdef Morphisms. -From MetaCoq.Template Require Import utils All. -Require Import MetaCoq.Checker.All. - -Unset Template Cast Propositions. - -MetaCoq Quote Recursively Definition idq := @Coq.Classes.Morphisms.Proper. - -Existing Instance config.default_checker_flags. -Existing Instance Checker.default_fuel. - -Definition test0 := Eval vm_compute in typecheck_program idq. -Definition nAnon := {| binder_name := nAnon; binder_relevance := Relevant |}. - -Definition timpl x y := tProd nAnon x (LiftSubst.lift0 1 y). - -MetaCoq Quote Recursively Definition four := (2 + 2). -Unset Printing Matching. - -Ltac typecheck := try red; cbn; intros; - match goal with - |- ?Σ ;;; ?Γ |- ?t : ?T => - refine (infer_correct Σ _ _ Γ t T _ _); [reflexivity|constructor|vm_compute; reflexivity] - end. -Ltac infer := try red; - match goal with - |- ?Σ ;;; ?Γ |- ?t : ?T => - refine (infer_correct Σ _ _ Γ t T _ _); [reflexivity|constructor| - let t' := eval vm_compute in (infer' Σ Γ t) in - change (t' = Checked T); reflexivity] - end. - -MetaCoq Quote Definition natr := nat. - -Definition type_program (p : program) (ty : term) := - let Σ := empty_ext (fst p) in - Σ ;;; [] |- snd p : ty. - -Example typecheck_four : type_program four natr:= ltac:(typecheck). - -Goal { ty & type_program four ty }. -Proof. - eexists. infer. -Qed. - -(* Uase template-coq to make a [program] from function defined above *) - -(* Eval native_compute in typecheck_program p_Plus1. *) - -Definition test_reduction (p : program) := - let Σ := empty_ext (fst p) in - reduce (fst Σ) [] (snd p). - -Definition string_of_env_error e := - match e with - | IllFormedDecl s e => ("IllFormedDecl " ^ s ^ "\nType error: " ^ string_of_type_error e)%string - | AlreadyDeclared s => ("Alreadydeclared " ^ s)%string - end. - -Definition out_typing c := - match c with - | Checked t => t - | TypeError e => tVar ("Typing error")%string - end. - -Definition out_check c := - match c with - | CorrectDecl t => t - | EnvError e => tVar ("Check error: " ++ string_of_env_error e)%string - end. - -Ltac interp_red c := - let t:= eval vm_compute in (out_typing (test_reduction c)) in exact t. - -Ltac interp_infer c := - let t:= eval vm_compute in (out_check (typecheck_program c)) in exact t. - -Ltac term_type c := - let X := type of c in exact X. - -Ltac quote_type c := - let X := type of c in - quote_term X ltac:(fun Xast => exact Xast). - -Notation convertible x y := (@eq_refl _ x : x = y). - -Module Test1. - Definition term := (Nat.mul 2 62). - Load "test_term.v". -End Test1. - -Module Test2. - Definition term := (fun (f : nat -> nat) (x : nat) => f (f x)). - Load "test_term.v". -End Test2. - -Module Test3. - (* Definition term := (id 0). *) - (* Load "test_term.v". *) -End Test3. - -Module Test4. - Definition term := @id. - Set Printing Universes. - Load "test_term.v". -End Test4. - -Module Test5. - - (** A function defined using measure or well-founded relation **) - (* Function Plus1 (n: nat) {measure id n} : nat := - match n with - | 0 => 1 - | S p => S (Plus1 p) - end. - - intros. unfold id. abstract omega. - Defined. *) - - (* Time Template Check Plus1. *) - (* Time Template Check Coq.ZArith.BinInt.Z.succ_pred. (* -> 16 s *) *) - (* MetaCoq Quote Recursively Definition plop := Coq.ZArith.BinInt.Z.succ_pred. *) - (* Eval native_compute in (typecheck_program plop). (* -> 31 min!! *) *) - - - (* (* Too long with universes on *) *) - (* MetaCoq Quote Recursively Definition p_Plus1 := Plus1. *) - - (* Definition term := Plus1. *) - (* Definition ast := p_Plus1. *) - (* Set Printing Universes. *) - (* (** Check typing *) *) - - (* (* Yay! Typechecking an actually non-trivial term. (173s) *) *) - - (* MetaCoq Unquote Definition inferred_type := ltac:(interp_infer ast). *) - (* Definition inferred_type' := Eval cbv delta in inferred_type. *) - (* Print inferred_type'. *) - (* Check convertible ltac:(term_type term) inferred_type. *) -End Test5. - -Universe i j. - -Definition f1 := (forall (A:Type@{i}) (B: Prop), A -> B -> A). -(* : Type@{Set+1, i+1} *) - -Definition f2 := (forall (A:Type@{i}) (B: Prop), A -> B -> B). -(* : Prop *) - -MetaCoq Quote Definition f1' := (forall (A:Type@{i}) (B: Prop), A -> B -> A). - -Definition test1 := Eval lazy in infer' (empty_ext []) nil f1'. - -MetaCoq Quote Definition f2' := (forall (A:Type@{i}) (B: Prop), A -> B -> B). - -Definition test2 := Eval lazy in infer' (empty_ext []) nil f2'. - -Definition f := (forall (A:Type@{i}) (B: Type@{j}), A -> B -> A). -(* : Type@{i+1, j+1} *) - -MetaCoq Quote Definition f' := (forall (A:Type@{i}) (B:Type@{j}), A -> B -> A). - -MetaCoq Quote Definition f'' := (forall (B: Type@{j}), B -> B). - -Definition test3 := Eval lazy in infer' (empty_ext []) nil f'. diff --git a/test-suite/_CoqProject b/test-suite/_CoqProject index 86f9eaa69..8f0e6f89a 100644 --- a/test-suite/_CoqProject +++ b/test-suite/_CoqProject @@ -1,9 +1,7 @@ -I ../template-coq/build --I ../checker/src -I ../safechecker/src -I ../erasure/src -Q ../template-coq/theories MetaCoq.Template --Q ../checker/theories MetaCoq.Checker -Q ../pcuic/theories MetaCoq.PCUIC -Q ../safechecker/theories MetaCoq.SafeChecker -Q ../erasure/theories MetaCoq.Erasure @@ -22,6 +20,7 @@ case.v castprop.v # CheckerTest.v disabled cofix.v +vs.v erasure_live_test.v erasure_test.v evars.v @@ -36,13 +35,11 @@ opaque.v proj.v run_in_tactic.v safechecker_test.v -# test_term.v -> loaded in TypingTests.v tmExistingInstance.v tmFreshName.v tmInferInstance.v -TypingTests.v unfold.v univ.v -vs.v tmVariable.v order_rec.v +primitive.v \ No newline at end of file diff --git a/test-suite/erasure_live_test.v b/test-suite/erasure_live_test.v index 4144a6acf..3a9d162bf 100644 --- a/test-suite/erasure_live_test.v +++ b/test-suite/erasure_live_test.v @@ -1,5 +1,7 @@ From Coq Require Import Recdef. -From MetaCoq.Template Require Import Loader. +From MetaCoq.Template Require Import TemplateMonad Loader. +From MetaCoq.SafeChecker Require Import SafeTemplateChecker. +From MetaCoq.PCUIC Require Import PCUICAstUtils. From MetaCoq.Erasure Require Import Erasure. From Coq Require Import String. Local Open Scope string_scope. @@ -10,6 +12,7 @@ Definition test (p : Ast.program) : string := erase_and_print_template_program p. MetaCoq Quote Recursively Definition zero := 0. + Definition zerocst := Eval lazy in test zero. MetaCoq Quote Recursively Definition exproof := I. @@ -18,21 +21,26 @@ Definition exprooftest := Eval lazy in test exproof. MetaCoq Quote Recursively Definition exintro := (@exist _ _ 0 (@eq_refl _ 0) : {x : nat | x = 0}). Definition exintrotest := Eval lazy in test exintro. -MetaCoq Quote Recursively Definition idnat := ((fun (X : Set) (x : X) => x) nat). -Definition test_idnat := Eval lazy in test idnat. +Definition idnat := ((fun (X : Set) (x : X) => x) nat). + +MetaCoq Quote Recursively Definition idnatc := idnat. +Definition test_idnat := Eval lazy in test idnatc. (** Check that optimization of singleton pattern-matchings work *) -MetaCoq Quote Recursively Definition singlelim := ((fun (X : Set) (x : X) (e : x = x) => +Definition singlelim := ((fun (X : Set) (x : X) (e : x = x) => match e in eq _ x' return bool with | eq_refl => true end)). -Time Definition singelim_test := Eval lazy in test singlelim. - -MetaCoq Quote Recursively Definition plusr := (plus 0 1). - -Time Definition plusrtest := Eval lazy in test plusr. +Definition erase {A} (a : A) : TemplateMonad unit := + aq <- tmQuoteRec a ;; + s <- tmEval lazy (erase_and_print_template_program aq) ;; + tmMsg s. +MetaCoq Run (erase 0). +MetaCoq Run (tmEval hnf idnat >>= erase). +MetaCoq Run (tmEval hnf singlelim >>= erase). +MetaCoq Run (erase (plus 0 1)). (** vector addition **) Require Coq.Vectors.Vector. @@ -44,21 +52,7 @@ Definition v01 : Vector.t nat 2 := Definition v23 : Vector.t nat 2 := (Vector.cons nat 2 1 (Vector.cons nat 3 0 (Vector.nil nat))). Definition vplus0123 := (vplus v01 v23). -MetaCoq Quote Recursively Definition cbv_vplus0123 := (* [program] of Coq's answer *) - ltac:(let t:=(eval cbv in (vplus0123)) in exact t). - -(* [Term] of Coq's answer *) -Definition ans_vplus0123 := Eval lazy in test cbv_vplus0123. - -(* [program] of the program *) -MetaCoq Quote Recursively Definition p_vplus0123 := vplus0123. -Time Definition test_p_vplus0123 := Eval lazy in test p_vplus0123. (* 5s *) -(* - Time Eval vm_compute in test p_vplus0123. (* 3.54s *) -Time Eval native_compute in test p_vplus0123. (* 23.54s on first run *) - -Time Eval native_compute in test p_vplus0123. (* 2.8s on second run *) -*) +MetaCoq Run (tmEval hnf vplus0123 >>= erase). (** Ackermann **) Fixpoint ack (n m:nat) {struct n} : nat := diff --git a/test-suite/erasure_test.v b/test-suite/erasure_test.v index 6af5454da..4c74003d4 100644 --- a/test-suite/erasure_test.v +++ b/test-suite/erasure_test.v @@ -41,7 +41,8 @@ Require Import List. Import ListNotations. MetaCoq Erase (map negb [true; false]). -Definition bignat := 10000. +Set Warnings "-abstract-large-number". +Definition bignat := Eval compute in 10000. MetaCoq Erase bignat. Require Import vs. diff --git a/test-suite/evars.v b/test-suite/evars.v index 3e76b17d4..7ac22f79f 100644 --- a/test-suite/evars.v +++ b/test-suite/evars.v @@ -1,4 +1,8 @@ -Require Import MetaCoq.Template.Loader. +From Coq Require Import String List. +From MetaCoq.Template Require Import Ast Loader. +Import ListNotations. +Definition bAnon := {| binder_name := nAnon; binder_relevance := Relevant |}. +Definition bNamed s := {| binder_name := nNamed s; binder_relevance := Relevant |}. Goal True. evar (n : nat). @@ -10,4 +14,31 @@ Goal True. end. exact I. Unshelve. exact 0. -Qed. \ No newline at end of file +Qed. + +(** Evars *) +MetaCoq Quote Definition lnil := @nil. +MetaCoq Quote Definition listnat := (list nat). +MetaCoq Unquote Definition foo := (tCast (tApp lnil [hole]) Cast listnat). + +Local Open Scope string_scope. + +Goal list nat. + (* Back and forth *) + let x := open_constr:(nil) in quote_term x (fun qt => + ltac:(denote_term qt (fun unqt => set (e := eq_refl : unqt = x :> list bool)))). + (* Creation of evars by denotation of 'hole' *) + let x := eval cbv in (tApp lnil [hole]) in + denote_term x (fun k => exact k). +Defined. + +Goal True. + let x := eval cbv in (tLambda bAnon listnat hole) in + (* We check that created evars are in the right context *) + denote_term x (fun k => set (ev := k); revert ev). + instantiate (1 := list nat). + instantiate (1 := l). + exact I. +Qed. + + diff --git a/test-suite/int.v b/test-suite/int.v new file mode 100644 index 000000000..9288926c1 --- /dev/null +++ b/test-suite/int.v @@ -0,0 +1,14 @@ +Require Import MetaCoq.Template.Loader. +Require Import Int63. + +Definition n : Int63.int := 42. +Import List.ListNotations. +Definition ns : list Int63.int := [n]%list. + + +MetaCoq Quote Recursively Definition q_n := n. +MetaCoq Unquote Definition n' := (snd q_n). + +MetaCoq Quote Recursively Definition q_ns := ns. +MetaCoq Unquote Definition ns' := (snd q_ns). + diff --git a/test-suite/primitive.v b/test-suite/primitive.v new file mode 100644 index 000000000..6a085b9e6 --- /dev/null +++ b/test-suite/primitive.v @@ -0,0 +1,37 @@ +From Coq Require Import String. +From MetaCoq.Template Require Import monad_utils All. +From Coq.Numbers.Cyclic Require Import Int63. + +Local Open Scope string_scope. +Import MonadNotation. + +Definition bigint : Int63.int := 542985047%int63. + +Notation eval_hnf := (tmEval hnf). +Notation eval := (tmEval all). + +MetaCoq Run (eval_hnf bigint >>= + (fun x => tmQuote (x + 1)%int63) >>= + tmMkDefinition "foo"). + +Print foo. + +MetaCoq Run (eval_hnf bigint >>= + (fun x => tmQuote (x + 1)%int63 >>= fun q => + tmUnquoteTyped int q >>= fun unq => + tmPrint unq >>= fun _ => + tmLemma "foo'" (bigint + 1 = unq)%int63 >>= + fun x => tmPrint x)). + +From Coq Require Import Int63 PrimFloat. + +Definition f := (of_int63 bigint / 3)%float. +Eval lazy in f. +MetaCoq Run (tmEval lazy f >>= + (fun x => tmQuote (x + 1)%float) >>= + tmMkDefinition "fplus1"). + +MetaCoq Run (tmUnquoteTyped float (tFloat f) >>= + (fun x : float => tmPrint x >>= + fun _ => tmQuote x >>= tmMkDefinition "somefloat")). + Print somefloat. \ No newline at end of file diff --git a/test-suite/test_term.v b/test-suite/test_term.v deleted file mode 100644 index 88f561c2b..000000000 --- a/test-suite/test_term.v +++ /dev/null @@ -1,13 +0,0 @@ -(** Check reduction *) -MetaCoq Quote Recursively Definition ast := term. -MetaCoq Unquote Definition normal_form := ltac:(interp_red ast). - -Definition normal_form' := Eval vm_compute in normal_form. -(* Print normal_form'. *) -Check convertible term normal_form. - -(** Check typing *) -MetaCoq Unquote Definition inferred_type := ltac:(interp_infer ast). -Definition inferred_type' := Eval cbv delta in inferred_type. -(* Print inferred_type'. *) -Check convertible ltac:(term_type term) inferred_type. diff --git a/translations/param_binary.v b/translations/param_binary.v index 54228deec..05e5ae322 100644 --- a/translations/param_binary.v +++ b/translations/param_binary.v @@ -141,6 +141,7 @@ Fixpoint tsl_rec1_app (app : list term) (E : tsl_table) (t : term) : term := | tFix _ _ | tCoFix _ _ => todo "tsl" | tVar _ | tEvar _ _ => todo "tsl" | tLambda _ _ _ => tVar "impossible" + | tInt _ | tFloat _ => todo "impossible" end in apply app t1 end. diff --git a/translations/param_cheap_packed.v b/translations/param_cheap_packed.v index 57055a733..e7e746b07 100644 --- a/translations/param_cheap_packed.v +++ b/translations/param_cheap_packed.v @@ -1,6 +1,5 @@ (* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import utils All. -From MetaCoq.Checker Require Import All. +From MetaCoq.Template Require Import utils All Checker. From MetaCoq.Translations Require Import translation_utils sigma. Local Existing Instance config.default_checker_flags. diff --git a/translations/param_generous_packed.v b/translations/param_generous_packed.v index f4e44a5f6..f1052ded5 100644 --- a/translations/param_generous_packed.v +++ b/translations/param_generous_packed.v @@ -1,6 +1,5 @@ (* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import utils All. -From MetaCoq.Checker Require Import All. +From MetaCoq.Template Require Import utils Checker All. From MetaCoq.Translations Require Import translation_utils MiniHoTT_paths. diff --git a/translations/param_original.v b/translations/param_original.v index ead473d68..b4d6f7b61 100644 --- a/translations/param_original.v +++ b/translations/param_original.v @@ -103,6 +103,7 @@ Fixpoint tsl_rec1_app (app : option term) (E : tsl_table) (t : term) : term := | tFix _ _ | tCoFix _ _ => todo "tsl" | tVar _ | tEvar _ _ => todo "tsl" | tLambda _ _ _ => tVar "impossible" + | tInt _ | tFloat _ => todo "impossible" end in match app with Some t' => mkApp t1 (t' {3 := tRel 1} {2 := tRel 0}) | None => t1 end @@ -403,6 +404,9 @@ Module Axioms. intros H A B []; exact (H A A 1). Defined. + (* The import of CRelationClasses breaks rewrite *) + Set Typeclasses Depth 4. + Theorem Univalence'_provably_parametric : forall h : Univalence', Univalence'ᵗ h. Proof. unfold Univalence', Univalence'ᵗ. @@ -417,7 +421,7 @@ Module Axioms. intros x xᵗ; cbn. eapply pathsᵗ_ok2. set (coh x). set (q1 x) in *. set (q2 x) in *. clearbody p p1 p0; clear; cbn in *. - set (g x) in *. clearbody a. + set (g x) in *. clearbody a. simpl. rewrite transport_pp. destruct p1. cbn in *. match goal with diff --git a/translations/times_bool_fun.v b/translations/times_bool_fun.v index 7c828f648..92ad91548 100644 --- a/translations/times_bool_fun.v +++ b/translations/times_bool_fun.v @@ -1,6 +1,5 @@ (* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import utils All. -From MetaCoq Require Import Checker.All. +From MetaCoq.Template Require Import utils All Checker. From MetaCoq.Translations Require Import translation_utils MiniHoTT. diff --git a/translations/translation_utils.v b/translations/translation_utils.v index 61865e073..b9acffcf9 100644 --- a/translations/translation_utils.v +++ b/translations/translation_utils.v @@ -1,6 +1,5 @@ (* Distributed under the terms of the MIT license. *) -From MetaCoq.Template Require Import utils All. -Require Import MetaCoq.Checker.All. +From MetaCoq.Template Require Import utils Checker All. (* Should be in AstUtils probably *) Fixpoint subst_app (t : term) (us : list term) : term :=