From 46454e2d998e420a11b37941ecbe42f09f24160c Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Sun, 23 Apr 2023 03:33:58 -0700 Subject: [PATCH 01/61] Fix monad_map_branches_k name (#953) --- pcuic/theories/PCUICMonadAst.v | 4 ++-- template-coq/theories/MonadAst.v | 4 ++-- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/pcuic/theories/PCUICMonadAst.v b/pcuic/theories/PCUICMonadAst.v index 649235aed..00ec117ee 100644 --- a/pcuic/theories/PCUICMonadAst.v +++ b/pcuic/theories/PCUICMonadAst.v @@ -59,6 +59,6 @@ Section with_monad. Definition monad_map_branch_k {term term'} (f : nat -> term -> T term') g k b := @monad_map_branch term term' (f (#|bcontext b| + k)) g b. - Notation map_branches_k f h k brs := - (monad_map (monad_map_branch_k f h k) brs). + Definition monad_map_branches_k {term term'} f h k brs := + (monad_map (@monad_map_branch_k term term' f h k) brs). End with_monad. diff --git a/template-coq/theories/MonadAst.v b/template-coq/theories/MonadAst.v index a7277c039..e5d8a3049 100644 --- a/template-coq/theories/MonadAst.v +++ b/template-coq/theories/MonadAst.v @@ -62,6 +62,6 @@ Section with_monad. Definition monad_map_branches {term B} (f : term -> T B) l := monad_map (monad_map_branch f) l. - Notation monad_map_branches_k f k brs := - (monad_map (fun b => monad_map_branch (f (#|b.(bcontext)| + k)) b) brs). + Definition monad_map_branches_k {term term'} f k brs := + (monad_map (fun b => @monad_map_branch term term' (f (#|b.(bcontext)| + k)) b) brs). End with_monad. From d8f14791ac74e7d1257b4b9cfe1ec7c78586d73e Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Sun, 23 Apr 2023 03:38:12 -0700 Subject: [PATCH 02/61] Add boolean versions of the varieties of `extends` (#954) --- common/_CoqProject.in | 1 + common/theories/Environment.v | 46 ++++++- common/theories/EnvironmentReflect.v | 119 ++++++++++++++++++ common/theories/Universes.v | 59 +++++++-- pcuic/theories/Syntax/PCUICReflect.v | 5 + quotation/theories/ToPCUIC/Common/Universes.v | 4 +- .../theories/ToTemplate/Common/Universes.v | 4 +- template-coq/theories/ReflectAst.v | 6 +- utils/theories/MCOption.v | 27 +++- 9 files changed, 251 insertions(+), 20 deletions(-) create mode 100644 common/theories/EnvironmentReflect.v diff --git a/common/_CoqProject.in b/common/_CoqProject.in index 7e3410b00..6530d5162 100644 --- a/common/_CoqProject.in +++ b/common/_CoqProject.in @@ -11,5 +11,6 @@ theories/MonadBasicAst.v theories/Environment.v theories/Reflect.v theories/EnvironmentTyping.v +theories/EnvironmentReflect.v theories/EnvMap.v theories/Transform.v diff --git a/common/theories/Environment.v b/common/theories/Environment.v index 2f7da713b..fbfc1cf67 100644 --- a/common/theories/Environment.v +++ b/common/theories/Environment.v @@ -2,7 +2,7 @@ From Coq Require Import ssreflect ssrbool ssrfun Morphisms Setoid. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import BasicAst Primitive Universes. -From Equations.Prop Require Import Classes. +From Equations.Prop Require Import Classes EqDecInstances. Module Type Term. @@ -28,8 +28,13 @@ End Term. Module Type TermDecide (Import T : Term). #[export] Declare Instance term_eq_dec : EqDec term. + #[export] Hint Extern 0 (ReflectEq term) => exact (@EqDec_ReflectEq term term_eq_dec) : typeclass_instances. End TermDecide. +Module TermDecideReflectInstances (Import T : Term) (Import TDec : TermDecide T). + #[export] Hint Extern 0 (ReflectEq term) => exact (@EqDec_ReflectEq term term_eq_dec) : typeclass_instances. +End TermDecideReflectInstances. + Module Retroknowledge. Record t := mk_retroknowledge { @@ -44,6 +49,20 @@ Module Retroknowledge. option_extends x.(retro_float64) y.(retro_float64). Existing Class extends. + Definition extendsb (x y : t) := + option_extendsb x.(retro_int63) y.(retro_int63) && + option_extendsb x.(retro_float64) y.(retro_float64). + + Lemma extendsT x y : reflect (extends x y) (extendsb x y). + Proof. + rewrite /extends/extendsb; do 2 case: option_extendsT; cbn; constructor; intuition. + Qed. + + Lemma extends_spec x y : extendsb x y <-> extends x y. + Proof. + rewrite /extends/extendsb -!option_extends_spec /is_true !Bool.andb_true_iff //=. + Qed. + #[global] Instance extends_refl x : extends x x. Proof. split; apply option_extends_refl. @@ -54,7 +73,15 @@ Module Retroknowledge. intros x y z [] []; split; cbn; now etransitivity; tea. Qed. + #[export,program] Instance reflect_t : ReflectEq t := { + eqb x y := (x.(retro_int63) == y.(retro_int63)) && + (x.(retro_float64) == y.(retro_float64)) + }. + Next Obligation. + do 2 case: eqb_spec; destruct x, y; cbn; intros; subst; constructor; congruence. + Qed. End Retroknowledge. +Export (hints) Retroknowledge. Module Environment (T : Term). @@ -983,8 +1010,25 @@ Module Type EnvironmentDecide (T : Term) (Import E : EnvironmentSig T). #[export] Declare Instance mutual_inductive_body_eq_dec : EqDec mutual_inductive_body. #[export] Declare Instance constant_body_eq_dec : EqDec constant_body. #[export] Declare Instance global_decl_eq_dec : EqDec global_decl. + #[export] Hint Extern 0 (ReflectEq context) => exact (@EqDec_ReflectEq context context_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq constructor_body) => exact (@EqDec_ReflectEq constructor_body constructor_body_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq projection_body) => exact (@EqDec_ReflectEq projection_body projection_body_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq one_inductive_body) => exact (@EqDec_ReflectEq one_inductive_body one_inductive_body_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq mutual_inductive_body) => exact (@EqDec_ReflectEq mutual_inductive_body mutual_inductive_body_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq constant_body) => exact (@EqDec_ReflectEq constant_body constant_body_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq global_decl) => exact (@EqDec_ReflectEq global_decl global_decl_eq_dec) : typeclass_instances. End EnvironmentDecide. +Module EnvironmentDecideReflectInstances (T : Term) (Import E : EnvironmentSig T) (Import EDec : EnvironmentDecide T E). + #[export] Hint Extern 0 (ReflectEq context) => exact (@EqDec_ReflectEq context context_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq constructor_body) => exact (@EqDec_ReflectEq constructor_body constructor_body_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq projection_body) => exact (@EqDec_ReflectEq projection_body projection_body_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq one_inductive_body) => exact (@EqDec_ReflectEq one_inductive_body one_inductive_body_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq mutual_inductive_body) => exact (@EqDec_ReflectEq mutual_inductive_body mutual_inductive_body_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq constant_body) => exact (@EqDec_ReflectEq constant_body constant_body_eq_dec) : typeclass_instances. + #[export] Hint Extern 0 (ReflectEq global_decl) => exact (@EqDec_ReflectEq global_decl global_decl_eq_dec) : typeclass_instances. +End EnvironmentDecideReflectInstances. + Module Type TermUtils (T: Term) (E: EnvironmentSig T). Import T E. diff --git a/common/theories/EnvironmentReflect.v b/common/theories/EnvironmentReflect.v new file mode 100644 index 000000000..63c5f9571 --- /dev/null +++ b/common/theories/EnvironmentReflect.v @@ -0,0 +1,119 @@ +(* Distributed under the terms of the MIT license. *) +From Coq Require Import ssreflect ssrbool ssrfun Morphisms Setoid. +From MetaCoq.Utils Require Import utils. +From MetaCoq.Common Require Import BasicAst Primitive Universes Environment Reflect. +From Equations.Prop Require Import Classes EqDecInstances. + +Module EnvironmentReflect (T : Term) (Import E : EnvironmentSig T) (Import TDec : TermDecide T) (Import EDec : EnvironmentDecide T E). + + Local Notation extendsb_decls_part Σ Σ' + := (forallb (fun d => let c := d.1 in skipn (#|lookup_envs Σ' c| - #|lookup_envs Σ c|) (lookup_envs Σ' c) == lookup_envs Σ c) (declarations Σ)) (only parsing). + Local Notation strictly_extendsb_decls_part Σ Σ' + := (skipn (#|Σ'.(declarations)| - #|Σ.(declarations)|) Σ'.(declarations) == Σ.(declarations)) (only parsing). + + Lemma extends_decls_partT (Σ Σ' : global_env) + : reflectT (forall c, ∑ decls, lookup_envs Σ' c = decls ++ lookup_envs Σ c) (extendsb_decls_part Σ Σ'). + Proof. + case: (@forallbP _ _ _ _ (fun _ => eqb_specT _ _)) => H; constructor. + all: setoid_rewrite Forall_forall in H. + { move => c. + specialize (fun d => H (c, d)). + cbn [fst] in *. + specialize (fun pf : { d : _ | In (c, d) _ } => H (proj1_sig pf) (proj2_sig pf)). + destruct (lookup_envs Σ c) eqn:H'. + { eexists; rewrite app_nil_r; reflexivity. } + rewrite <- H; clear H. + { eexists; symmetry; apply firstn_skipn. } + unfold lookup_envs in *. + move: H'; elim: (declarations Σ); cbn [lookup_globals] => //= ??. + case: eqb_specT => ?; subst. + all: move => H H'; try destruct (H H'); rdest; eauto. } + { intro H'; apply H; clear H; intros [c ?]; specialize (H' c). + destruct H' as [decls H']. + cbn [fst]. + rewrite H' app_length Nat.add_sub skipn_all_app //. } + Qed. + + Lemma strictly_extends_decls_partT (Σ Σ' : global_env) + : reflectT (∑ Σ'', declarations Σ' = Σ'' ++ declarations Σ) (strictly_extendsb_decls_part Σ Σ'). + Proof. + case: eqb_specT => H; constructor. + { rewrite -H. + eexists; symmetry; apply firstn_skipn. } + { move => [Σ'' H']. + move: H. rewrite H' app_length Nat.add_sub skipn_all_app //. } + Qed. + + Definition extendsb (Σ Σ' : global_env) : bool + := Σ.(universes) ⊂?_cs Σ'.(universes) + && extendsb_decls_part Σ Σ' + && Retroknowledge.extendsb Σ.(retroknowledge) Σ'.(retroknowledge). + + Definition extends_declsb (Σ Σ' : global_env) : bool + := (Σ.(universes) == Σ'.(universes)) + && extendsb_decls_part Σ Σ' + && (Σ.(retroknowledge) == Σ'.(retroknowledge)). + + Definition extends_strictly_on_declsb (Σ Σ' : global_env) : bool + := (Σ.(universes) ⊂?_cs Σ'.(universes)) + && strictly_extendsb_decls_part Σ Σ' + && Retroknowledge.extendsb Σ.(retroknowledge) Σ'.(retroknowledge). + + Definition strictly_extends_declsb (Σ Σ' : global_env) : bool + := (Σ.(universes) == Σ'.(universes)) + && strictly_extendsb_decls_part Σ Σ' + && (Σ.(retroknowledge) == Σ'.(retroknowledge)). + + Lemma extendsT Σ Σ' : reflectT (extends Σ Σ') (extendsb Σ Σ'). + Proof. + rewrite /extends/extendsb. + case: extends_decls_partT; case: Retroknowledge.extendsT; case: ContextSet.subsetP; cbn; + constructor; intuition. + Qed. + + Lemma extends_declsT Σ Σ' : reflectT (extends_decls Σ Σ') (extends_declsb Σ Σ'). + Proof. + rewrite /extends_decls/extends_declsb. + case: extends_decls_partT; case: eqb_specT; case: eqb_specT; cbn; + constructor; intuition. + Qed. + + Lemma extends_strictly_on_declsT Σ Σ' : reflectT (extends_strictly_on_decls Σ Σ') (extends_strictly_on_declsb Σ Σ'). + Proof. + rewrite /extends_strictly_on_decls/extends_strictly_on_declsb. + case: strictly_extends_decls_partT; case: Retroknowledge.extendsT; case: ContextSet.subsetP; cbn; + constructor; intuition. + Qed. + + Lemma strictly_extends_declsT Σ Σ' : reflectT (strictly_extends_decls Σ Σ') (strictly_extends_declsb Σ Σ'). + Proof. + rewrite /strictly_extends_decls/strictly_extends_declsb. + case: strictly_extends_decls_partT; case: eqb_specT; case: eqb_specT; cbn; + constructor; intuition. + Qed. + + Lemma extends_spec Σ Σ' : extendsb Σ Σ' <~> extends Σ Σ'. + Proof. + case: extendsT; split; intuition congruence. + Qed. + + Lemma extends_decls_spec Σ Σ' : extends_declsb Σ Σ' <~> extends_decls Σ Σ'. + Proof. + case: extends_declsT; split; intuition congruence. + Qed. + + Lemma extends_strictly_on_decls_spec Σ Σ' : extends_strictly_on_declsb Σ Σ' <~> extends_strictly_on_decls Σ Σ'. + Proof. + case: extends_strictly_on_declsT; split; intuition congruence. + Qed. + + Lemma strictly_extends_decls_spec Σ Σ' : strictly_extends_declsb Σ Σ' <~> strictly_extends_decls Σ Σ'. + Proof. + case: strictly_extends_declsT; split; intuition congruence. + Qed. + +End EnvironmentReflect. + +Module Type EnvironmentReflectSig (T : Term) (E : EnvironmentSig T) (TDec : TermDecide T) (EDec : EnvironmentDecide T E). + Include EnvironmentReflect T E TDec EDec. +End EnvironmentReflectSig. diff --git a/common/theories/Universes.v b/common/theories/Universes.v index 6c31dd7ce..71e71ac40 100644 --- a/common/theories/Universes.v +++ b/common/theories/Universes.v @@ -172,6 +172,8 @@ Module LS := LevelSet. Ltac lsets := LevelSetDecide.fsetdec. Notation "(=_lset)" := LevelSet.Equal (at level 0). Infix "=_lset" := LevelSet.Equal (at level 30). +Notation "(==_lset)" := LevelSet.equal (at level 0). +Infix "==_lset" := LevelSet.equal (at level 30). Section LevelSetMoreFacts. @@ -1496,6 +1498,8 @@ Ltac csets := ConstraintSetDecide.fsetdec. Notation "(=_cset)" := ConstraintSet.Equal (at level 0). Infix "=_cset" := ConstraintSet.Equal (at level 30). +Notation "(==_cset)" := ConstraintSet.equal (at level 0). +Infix "==_cset" := ConstraintSet.equal (at level 30). Definition declared_cstr_levels levels (cstr : UnivConstraint.t) := let '(l1,_,l2) := cstr in @@ -1600,21 +1604,28 @@ Module ContextSet. Definition is_empty (uctx : t) := LevelSet.is_empty (fst uctx) && ConstraintSet.is_empty (snd uctx). - Definition equal (x y : t) : Prop := + Definition Equal (x y : t) : Prop := x.1 =_lset y.1 /\ x.2 =_cset y.2. - Definition subset (x y : t) : Prop := + Definition equal (x y : t) : bool := + x.1 ==_lset y.1 && x.2 ==_cset y.2. + + Definition Subset (x y : t) : Prop := LevelSet.Subset (levels x) (levels y) /\ ConstraintSet.Subset (constraints x) (constraints y). + Definition subset (x y : t) : bool := + LevelSet.subset (levels x) (levels y) && + ConstraintSet.subset (constraints x) (constraints y). + Definition inter (x y : t) : t := (LevelSet.inter (levels x) (levels y), ConstraintSet.inter (constraints x) (constraints y)). Definition inter_spec (x y : t) : - subset (inter x y) x /\ - subset (inter x y) y /\ - forall z, subset z x -> subset z y -> subset z (inter x y). + Subset (inter x y) x /\ + Subset (inter x y) y /\ + forall z, Subset z x -> Subset z y -> Subset z (inter x y). Proof. split; last split. 1,2: split=> ?; [move=> /LevelSet.inter_spec [//]|move=> /ConstraintSet.inter_spec [//]]. @@ -1626,21 +1637,45 @@ Module ContextSet. (LevelSet.union (levels x) (levels y), ConstraintSet.union (constraints x) (constraints y)). Definition union_spec (x y : t) : - subset x (union x y) /\ - subset y (union x y) /\ - forall z, subset x z -> subset y z -> subset (union x y) z. + Subset x (union x y) /\ + Subset y (union x y) /\ + forall z, Subset x z -> Subset y z -> Subset (union x y) z. Proof. split; last split. 1,2: split=> ??; [apply/LevelSet.union_spec|apply/ConstraintSet.union_spec ]; by constructor. move=> ? [??] [??]; split=> ?; [move=>/LevelSet.union_spec|move=>/ConstraintSet.union_spec]=>-[]; auto. Qed. + + Lemma equal_spec s s' : equal s s' <-> Equal s s'. + Proof. + rewrite /equal/Equal/is_true Bool.andb_true_iff LevelSet.equal_spec ConstraintSet.equal_spec. + reflexivity. + Qed. + + Lemma subset_spec s s' : subset s s' <-> Subset s s'. + Proof. + rewrite /subset/Subset/is_true Bool.andb_true_iff LevelSet.subset_spec ConstraintSet.subset_spec. + reflexivity. + Qed. + + Lemma subsetP s s' : reflect (Subset s s') (subset s s'). + Proof. + generalize (subset_spec s s'). + destruct subset; case; constructor; intuition. + Qed. End ContextSet. -Notation "(=_cs)" := ContextSet.equal (at level 0). -Notation "(⊂_cs)" := ContextSet.subset (at level 0). -Infix "=_cs" := ContextSet.equal (at level 30). -Infix "⊂_cs" := ContextSet.subset (at level 30). +Export (hints) ContextSet. + +Notation "(=_cs)" := ContextSet.Equal (at level 0). +Notation "(⊂_cs)" := ContextSet.Subset (at level 0). +Infix "=_cs" := ContextSet.Equal (at level 30). +Infix "⊂_cs" := ContextSet.Subset (at level 30). +Notation "(==_cs)" := ContextSet.equal (at level 0). +Notation "(⊂?_cs)" := ContextSet.subset (at level 0). +Infix "==_cs" := ContextSet.equal (at level 30). +Infix "⊂?_cs" := ContextSet.subset (at level 30). Lemma incl_cs_refl cs : cs ⊂_cs cs. Proof. diff --git a/pcuic/theories/Syntax/PCUICReflect.v b/pcuic/theories/Syntax/PCUICReflect.v index 51100825c..a28cedb55 100644 --- a/pcuic/theories/Syntax/PCUICReflect.v +++ b/pcuic/theories/Syntax/PCUICReflect.v @@ -4,6 +4,7 @@ From Equations Require Import Equations. From MetaCoq.PCUIC Require Import PCUICAst PCUICInduction. From MetaCoq.Utils Require Import utils. +From MetaCoq.Common Require Import EnvironmentReflect. From MetaCoq.Common Require Export Reflect. Open Scope pcuic. @@ -375,6 +376,7 @@ Qed. Module PCUICTermDecide <: TermDecide PCUICTerm. #[export] Instance term_eq_dec : EqDec term := _. + Include TermDecideReflectInstances PCUICTerm. End PCUICTermDecide. Module PCUICEnvironmentDecide <: EnvironmentDecide PCUICTerm PCUICEnvironment. @@ -385,4 +387,7 @@ Module PCUICEnvironmentDecide <: EnvironmentDecide PCUICTerm PCUICEnvironment. #[export] Instance mutual_inductive_body_eq_dec : EqDec mutual_inductive_body := _. #[export] Instance constant_body_eq_dec : EqDec constant_body := _. #[export] Instance global_decl_eq_dec : EqDec global_decl := _. + Include EnvironmentDecideReflectInstances PCUICTerm PCUICEnvironment. End PCUICEnvironmentDecide. + +Module PCUICEnvironmentReflect := EnvironmentReflect PCUICTerm PCUICEnvironment PCUICTermDecide PCUICEnvironmentDecide. diff --git a/quotation/theories/ToPCUIC/Common/Universes.v b/quotation/theories/ToPCUIC/Common/Universes.v index de10eb793..b5d367728 100644 --- a/quotation/theories/ToPCUIC/Common/Universes.v +++ b/quotation/theories/ToPCUIC/Common/Universes.v @@ -63,8 +63,8 @@ Export (hints) QuoteUniverses1. UContext.t AUContext.t ContextSet.t - ContextSet.equal - ContextSet.subset + ContextSet.Equal + ContextSet.Subset : quotation. #[export] Typeclasses Transparent diff --git a/quotation/theories/ToTemplate/Common/Universes.v b/quotation/theories/ToTemplate/Common/Universes.v index 29c7034df..f89016683 100644 --- a/quotation/theories/ToTemplate/Common/Universes.v +++ b/quotation/theories/ToTemplate/Common/Universes.v @@ -63,8 +63,8 @@ Export (hints) QuoteUniverses1. UContext.t AUContext.t ContextSet.t - ContextSet.equal - ContextSet.subset + ContextSet.Equal + ContextSet.Subset : quotation. #[export] Typeclasses Transparent diff --git a/template-coq/theories/ReflectAst.v b/template-coq/theories/ReflectAst.v index 30029511e..3123f91e7 100644 --- a/template-coq/theories/ReflectAst.v +++ b/template-coq/theories/ReflectAst.v @@ -2,7 +2,7 @@ (* For primitive integers and floats *) From Coq Require Numbers.Cyclic.Int63.Uint63 Floats.PrimFloat Floats.FloatAxioms. From MetaCoq.Utils Require Import utils. -From MetaCoq.Common Require Import BasicAst Reflect Environment. +From MetaCoq.Common Require Import BasicAst Reflect Environment EnvironmentReflect. From MetaCoq.Template Require Import AstUtils Ast Induction. Require Import ssreflect. From Equations Require Import Equations. @@ -248,6 +248,7 @@ Defined. Module TemplateTermDecide <: TermDecide TemplateTerm. #[export] Instance term_eq_dec : EqDec term := _. + Include TermDecideReflectInstances TemplateTerm. End TemplateTermDecide. Module EnvDecide <: EnvironmentDecide TemplateTerm Env. @@ -258,4 +259,7 @@ Module EnvDecide <: EnvironmentDecide TemplateTerm Env. #[export] Instance mutual_inductive_body_eq_dec : EqDec mutual_inductive_body := _. #[export] Instance constant_body_eq_dec : EqDec constant_body := _. #[export] Instance global_decl_eq_dec : EqDec global_decl := _. + Include EnvironmentDecideReflectInstances TemplateTerm Env. End EnvDecide. + +Module EnvReflect := EnvironmentReflect TemplateTerm Env TemplateTermDecide EnvDecide. diff --git a/utils/theories/MCOption.v b/utils/theories/MCOption.v index f1d0dfd19..05bf8efd0 100644 --- a/utils/theories/MCOption.v +++ b/utils/theories/MCOption.v @@ -1,6 +1,6 @@ -From Coq Require Import List ssreflect Arith Morphisms Relation_Definitions. +From Coq Require Import List ssreflect ssrbool Arith Morphisms Relation_Definitions. -From MetaCoq.Utils Require Import MCPrelude MCList MCProd MCReflect. +From MetaCoq.Utils Require Import MCPrelude MCList MCProd MCReflect ReflectEq. Definition option_get {A} (default : A) (x : option A) : A := match x with @@ -202,6 +202,29 @@ Inductive option_extends {A} : relation (option A) := | option_ext_non : option_extends None None. Derive Signature for option_extends. +Definition option_extendsb {A} {Aeq:ReflectEq A} (x y : option A) : bool := + match x, y with + | None, Some _ => true + | Some t1, Some t2 => t1 == t2 + | None, None => true + | _, _ => false + end. + +Lemma option_extendsT {A} {Aeq : ReflectEq A} x y : reflect (@option_extends A x y) (option_extendsb x y). +Proof. + destruct x as [x|], y as [y|]; cbn. + 1: generalize (eqb_spec x y); case: eqb => H. + all: constructor. + 1: replace y with x by now inversion H. + all: first [ constructor | inversion 1 ]. + { inversion H; congruence. } +Qed. + +Lemma option_extends_spec {A} {Aeq : ReflectEq A} x y : option_extendsb x y <-> @option_extends A x y. +Proof. + case: option_extendsT; intuition congruence. +Qed. + #[export] Instance option_extends_refl {A} : RelationClasses.Reflexive (@option_extends A). Proof. intros []; constructor. Qed. From dd8c80d4bfc86efd74a2271702e1081cc1c2adee Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Thu, 20 Apr 2023 20:56:20 -0700 Subject: [PATCH 03/61] Add a merge operation for the global env This is useful for managing global envs around Gallina quotation. --- common/theories/Environment.v | 117 ++++++++++++++++++++++++++++++++++ 1 file changed, 117 insertions(+) diff --git a/common/theories/Environment.v b/common/theories/Environment.v index fbfc1cf67..725265ea5 100644 --- a/common/theories/Environment.v +++ b/common/theories/Environment.v @@ -80,6 +80,39 @@ Module Retroknowledge. Next Obligation. do 2 case: eqb_spec; destruct x, y; cbn; intros; subst; constructor; congruence. Qed. + + (** This operation is asymmetric; it perfers the first argument when the arguments are incompatible, but otherwise takes the join *) + Definition merge (r1 r2 : t) : t + := {| retro_int63 := match r1.(retro_int63) with Some v => Some v | None => r2.(retro_int63) end + ; retro_float64 := match r1.(retro_float64) with Some v => Some v | None => r2.(retro_float64) end |}. + + Lemma extends_l_merge r1 r2 + : extends r1 (merge r1 r2). + Proof. + rewrite /extends/merge; destruct r1, r2; cbn; repeat destruct ?; subst; + repeat constructor; clear; destruct_head' option; constructor. + Qed. + + Lemma extends_merge_idempotent r1 r2 + : extends r1 r2 -> merge r1 r2 = r2. + Proof. + rewrite /extends/merge; destruct r1, r2; cbn. + intro; rdest; destruct_head' (@option_extends); reflexivity. + Qed. + + Definition compatible (x y : t) : bool + := match x.(retro_int63), y.(retro_int63) with Some x, Some y => x == y | _, _ => true end + && match x.(retro_float64), y.(retro_float64) with Some x, Some y => x == y | _, _ => true end. + + Lemma extends_r_merge r1 r2 + : compatible r1 r2 -> extends r2 (merge r1 r2). + Proof. + rewrite /extends/merge/compatible; destruct r1, r2; cbn; repeat destruct ?; subst. + all: repeat case: eqb_spec => //=. + all: intros; subst. + all: repeat constructor; clear; destruct_head' option; constructor. + Qed. + End Retroknowledge. Export (hints) Retroknowledge. @@ -616,6 +649,90 @@ Module Environment (T : Term). #[global] Instance extends_trans : CRelationClasses.Transitive extends. Proof. extends_trans_t. Qed. + (** Merge two lists of global_declarations, assuming that any globals sharing a name are identical *) + Definition merge_globals (Σ Σ' : global_declarations) : global_declarations + := let known_kns := List.fold_right KernameSet.add KernameSet.empty (List.map fst Σ) in + List.filter (fun '(kn, _) => negb (KernameSet.mem kn known_kns)) Σ' ++ Σ. + + Definition merge_global_envs (Σ Σ' : global_env) : global_env + := {| universes := ContextSet.union Σ.(universes) Σ'.(universes) + ; declarations := merge_globals Σ.(declarations) Σ'.(declarations) + ; retroknowledge := Retroknowledge.merge Σ.(retroknowledge) Σ'.(retroknowledge) |}. + + Definition compatible_globals (Σ Σ' : global_declarations) : Prop + := forall c, lookup_globals Σ c <> [] -> lookup_globals Σ' c <> [] -> lookup_globals Σ c = lookup_globals Σ' c. + + Lemma extends_strictly_on_decls_l_merge Σ Σ' + : extends_strictly_on_decls Σ (merge_global_envs Σ Σ'). + Proof. + rewrite /extends_strictly_on_decls/merge_global_envs/merge_globals; cbn. + split; + try first [ apply ContextSet.union_spec + | apply Retroknowledge.extends_l_merge + | now eexists ]. + Qed. + + Definition compatible (Σ Σ' : global_env) + := Retroknowledge.compatible Σ.(retroknowledge) Σ'.(retroknowledge) + /\ compatible_globals Σ.(declarations) Σ'.(declarations). + + Lemma lookup_globals_filter p Σ c + : lookup_globals (filter (fun '(kn, _) => p kn) Σ) c = if p c then lookup_globals Σ c else []. + Proof. + induction Σ as [|?? IH]; cbn; rdest; cbn; try now repeat destruct ?. + case: eqb_spec => ?; repeat destruct ?; subst => //=. + all: rewrite ?eqb_refl. + all: try case: eqb_spec => ?; subst. + all: rewrite IH //=. + all: try congruence. + Qed. + + Lemma extends_r_merge_globals Σ Σ' + : compatible_globals Σ Σ' -> + forall c, + ∑ decls, lookup_globals (merge_globals Σ Σ') c = decls ++ lookup_globals Σ' c. + Proof. + rewrite /merge_globals. + intro H2; cbn. + cbv [compatible_globals] in *. + intro c. + specialize (H2 c). + rewrite lookup_globals_app lookup_globals_filter. + set (s := fold_right _ _ _). + assert (Hs : forall kn, KernameSet.mem kn s <-> List.In kn (map fst Σ)). + { setoid_rewrite <- KernameSetFact.mem_iff. + clear. + generalize dependent Σ; clear. + elim => //=; try setoid_rewrite KernameSetFact.empty_iff => //=. + move => [? ?] ? IH kn //=. + rewrite KernameSet.add_spec. + intuition. } + clearbody s. + specialize (Hs c). + destruct (in_dec eq_dec c (map fst Σ')) as [H'|H']; + [ exists nil | exists (lookup_globals Σ c) ]. + 2: apply lookup_globals_nil in H'; rewrite H'; clear H'. + 2: now destruct ?; cbn; rewrite app_nil_r. + pose proof (lookup_globals_nil Σ c) as Hc. + rewrite <- !lookup_globals_nil, <- Hs in H2. + rewrite <- Hs in Hc. + destruct KernameSet.mem; cbn in *. + { intuition. } + { destruct Hc as [Hc _]. + rewrite Hc ?app_nil_r //=. } + Qed. + + Lemma extends_r_merge Σ Σ' + : compatible Σ Σ' -> extends Σ' (merge_global_envs Σ Σ'). + Proof. + rewrite /extends/compatible/merge_global_envs/lookup_envs. + intros [H1 H2]. + split; + try first [ apply ContextSet.union_spec + | now apply Retroknowledge.extends_r_merge + | now apply extends_r_merge_globals ]. + Qed. + Definition primitive_constant (Σ : global_env) (p : prim_tag) : option kername := match p with | primInt => Σ.(retroknowledge).(Retroknowledge.retro_int63) From 7ea63fabe6c7c302d5dee4bb3c318a6f9b379b0a Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Fri, 21 Apr 2023 12:55:33 -0700 Subject: [PATCH 04/61] Add computable version of Env compatibility; refactor extension a bit --- common/theories/Environment.v | 152 +++++++++++++++++++-------- common/theories/EnvironmentReflect.v | 46 +++++++- 2 files changed, 154 insertions(+), 44 deletions(-) diff --git a/common/theories/Environment.v b/common/theories/Environment.v index 725265ea5..700d6af68 100644 --- a/common/theories/Environment.v +++ b/common/theories/Environment.v @@ -467,24 +467,36 @@ Module Environment (T : Term). └-----------------------------┴------------------------┴---------------------------┘ >>> *) + Notation extends_decls_part_globals Σ Σ' + := (forall c, ∑ decls, lookup_globals Σ' c = decls ++ lookup_globals Σ c) + (only parsing). + Notation strictly_extends_decls_part_globals Σ Σ' + := (∑ Σ'', Σ' = Σ'' ++ Σ) + (only parsing). + Notation extends_decls_part Σ Σ' + := (forall c, ∑ decls, lookup_envs Σ' c = decls ++ lookup_envs Σ c) + (only parsing). + Notation strictly_extends_decls_part Σ Σ' + := (strictly_extends_decls_part_globals Σ.(declarations) Σ'.(declarations)) + (only parsing). Definition extends (Σ Σ' : global_env) := [× Σ.(universes) ⊂_cs Σ'.(universes), - forall c, ∑ decls, lookup_envs Σ' c = decls ++ lookup_envs Σ c & + extends_decls_part Σ Σ' & Retroknowledge.extends Σ.(retroknowledge) Σ'.(retroknowledge)]. Definition extends_decls (Σ Σ' : global_env) := [× Σ.(universes) = Σ'.(universes), - forall c, ∑ decls, lookup_envs Σ' c = decls ++ lookup_envs Σ c & + extends_decls_part Σ Σ' & Σ.(retroknowledge) = Σ'.(retroknowledge)]. Definition extends_strictly_on_decls (Σ Σ' : global_env) := [× Σ.(universes) ⊂_cs Σ'.(universes), - ∑ Σ'', Σ'.(declarations) = Σ'' ++ Σ.(declarations) & + strictly_extends_decls_part Σ Σ' & Retroknowledge.extends Σ.(retroknowledge) Σ'.(retroknowledge)]. Definition strictly_extends_decls (Σ Σ' : global_env) := [× Σ.(universes) = Σ'.(universes), - ∑ Σ'', Σ'.(declarations) = Σ'' ++ Σ.(declarations) & + strictly_extends_decls_part Σ Σ' & Σ.(retroknowledge) = Σ'.(retroknowledge)]. Existing Class extends. @@ -551,17 +563,16 @@ Module Environment (T : Term). all: intuition congruence. Qed. - Lemma lookup_env_extends_NoDup Σ Σ' k d : - NoDup (List.map fst Σ'.(declarations)) -> - lookup_env Σ k = Some d -> - extends Σ Σ' -> lookup_env Σ' k = Some d. + Lemma lookup_global_extends_NoDup Σ Σ' k d : + NoDup (List.map fst Σ') -> + lookup_global Σ k = Some d -> + extends_decls_part_globals Σ Σ' -> lookup_global Σ' k = Some d. Proof. - rewrite /lookup_env /= -!hd_error_lookup_globals => Hnd. + rewrite /= -!hd_error_lookup_globals => Hnd. move: (@NoDup_length_lookup_globals _ Hnd k); clear Hnd. rewrite -hd_error_lookup_globals. - move=> H Hd [? eq]. + move=> H Hd eq. move: (eq k); clear eq. - rewrite /lookup_envs. case => ls eq. move: eq Hd H => ->. case: ls => //= ?. @@ -569,6 +580,15 @@ Module Environment (T : Term). case: lookup_globals => //=. Qed. + Lemma lookup_env_extends_NoDup Σ Σ' k d : + NoDup (List.map fst Σ'.(declarations)) -> + lookup_env Σ k = Some d -> + extends Σ Σ' -> lookup_env Σ' k = Some d. + Proof. + move => Hnd Hd; case => *. + eapply lookup_global_extends_NoDup; tea. + Qed. + Lemma lookup_globals_app Σ Σ' kn : lookup_globals (Σ ++ Σ') kn = lookup_globals Σ kn ++ lookup_globals Σ' kn. Proof. @@ -577,14 +597,25 @@ Module Environment (T : Term). case: eqb_spec => //= -> -> //=. Qed. - #[global] Instance strictly_extends_decls_extends_decls Σ Σ' : strictly_extends_decls Σ Σ' -> extends_decls Σ Σ'. + Lemma strictly_extends_decls_extends_part_globals Σ Σ' + : strictly_extends_decls_part_globals Σ Σ' -> extends_decls_part_globals Σ Σ'. Proof. - destruct Σ, Σ'; case => //= -> [Σ'' ->] ->. - rewrite /extends_decls/lookup_envs //=; split => //= ?. + case => //= ? -> c. rewrite lookup_globals_app. eexists; reflexivity. Qed. + Lemma strictly_extends_decls_extends_part Σ Σ' + : strictly_extends_decls_part Σ Σ' -> extends_decls_part Σ Σ'. + Proof. apply strictly_extends_decls_extends_part_globals. Qed. + + #[global] Instance strictly_extends_decls_extends_decls Σ Σ' : strictly_extends_decls Σ Σ' -> extends_decls Σ Σ'. + Proof. + destruct Σ, Σ'; case => //= -> ? ->. + rewrite /extends_decls; split; try reflexivity. + now apply strictly_extends_decls_extends_part. + Qed. + #[global] Instance strictly_extends_decls_extends_strictly_on_decls Σ Σ' : strictly_extends_decls Σ Σ' -> extends_strictly_on_decls Σ Σ'. Proof. destruct Σ, Σ'; intros []. cbn in *; subst. split => //=. @@ -599,10 +630,9 @@ Module Environment (T : Term). #[global] Instance extends_strictly_on_decls_extends Σ Σ' : extends_strictly_on_decls Σ Σ' -> extends Σ Σ'. Proof. - destruct Σ, Σ'; case => //= ? [Σ'' ->] ?. - rewrite /extends/lookup_envs //=; split => //= ?. - rewrite lookup_globals_app. - eexists; reflexivity. + destruct Σ, Σ'; case => //= ? ? ?. + rewrite /extends; split => //=. + now apply strictly_extends_decls_extends_part. Qed. #[global] Instance strictly_extends_decls_extends_decls_subrel : CRelationClasses.subrelation strictly_extends_decls extends_decls := strictly_extends_decls_extends_decls. @@ -628,18 +658,48 @@ Module Environment (T : Term). Proof. apply extends_refl. Qed. *) + Lemma extends_decls_part_globals_refl Σ : extends_decls_part_globals Σ Σ. + Proof. now exists [] => //. Qed. + + Lemma extends_decls_part_refl Σ : extends_decls_part Σ Σ. + Proof. apply extends_decls_part_globals_refl. Qed. + + Lemma strictly_extends_decls_part_globals_refl (Σ : global_declarations) + : strictly_extends_decls_part_globals Σ Σ. + Proof. now exists [] => //. Qed. + + Lemma strictly_extends_decls_part_refl Σ : strictly_extends_decls_part Σ Σ. + Proof. apply strictly_extends_decls_part_globals_refl. Qed. + + Lemma extends_decls_part_globals_trans Σ Σ' Σ'' + : extends_decls_part_globals Σ Σ' -> extends_decls_part_globals Σ' Σ'' -> extends_decls_part_globals Σ Σ''. + Proof. + move => H1 H2 c; move: (H1 c) (H2 c) => [? ->] [? ->]. + now eexists; rewrite app_assoc. + Qed. + + Lemma extends_decls_part_trans Σ Σ' Σ'' + : extends_decls_part Σ Σ' -> extends_decls_part Σ' Σ'' -> extends_decls_part Σ Σ''. + Proof. apply extends_decls_part_globals_trans. Qed. + + Lemma strictly_extends_decls_part_globals_trans (Σ Σ' Σ'' : global_declarations) + : strictly_extends_decls_part_globals Σ Σ' -> strictly_extends_decls_part_globals Σ' Σ'' -> strictly_extends_decls_part_globals Σ Σ''. + Proof. + move => [? ->] [? ->]. + now eexists; rewrite app_assoc. + Qed. + + Lemma strictly_extends_decls_part_trans Σ Σ' Σ'' + : strictly_extends_decls_part Σ Σ' -> strictly_extends_decls_part Σ' Σ'' -> strictly_extends_decls_part Σ Σ''. + Proof. apply strictly_extends_decls_part_globals_trans. Qed. + Local Ltac extends_trans_t := intros [?] [?] [?] [?] [?]; red; cbn in *; split; - try solve [ etransitivity; eassumption | eapply incl_cs_trans; eassumption ]; - repeat first [ progress subst - | match goal with - | [ H : ∑ x : _, _ |- _ ] => destruct H - | [ H : forall c : kername, _, kn : kername |- _ ] => specialize (H kn) - | [ H : ?x = _ |- context[?x] ] => rewrite H - end - | split - | intro - | now eexists; rewrite app_assoc ]. + try solve [ etransitivity; eassumption + | eapply incl_cs_trans; eassumption + | eapply strictly_extends_decls_part_globals_trans; eassumption + | eapply extends_decls_part_globals_trans; eassumption ]. + #[global] Instance strictly_extends_decls_trans : CRelationClasses.Transitive strictly_extends_decls. Proof. extends_trans_t. Qed. #[global] Instance extends_decls_trans : CRelationClasses.Transitive extends_decls. @@ -650,8 +710,11 @@ Module Environment (T : Term). Proof. extends_trans_t. Qed. (** Merge two lists of global_declarations, assuming that any globals sharing a name are identical *) + Definition declared_kername_set (Σ : global_declarations) : KernameSet.t + := List.fold_right KernameSet.add KernameSet.empty (List.map fst Σ). + Definition merge_globals (Σ Σ' : global_declarations) : global_declarations - := let known_kns := List.fold_right KernameSet.add KernameSet.empty (List.map fst Σ) in + := let known_kns := declared_kername_set Σ in List.filter (fun '(kn, _) => negb (KernameSet.mem kn known_kns)) Σ' ++ Σ. Definition merge_global_envs (Σ Σ' : global_env) : global_env @@ -687,6 +750,22 @@ Module Environment (T : Term). all: try congruence. Qed. + Lemma declared_kername_set_spec + : forall Σ c, KernameSet.In c (declared_kername_set Σ) <-> List.In c (map fst Σ). + Proof. + elim => //=; try setoid_rewrite KernameSetFact.empty_iff => //=. + move => [? ?] ? IH c //=. + rewrite KernameSet.add_spec. + intuition. + Qed. + + Lemma declared_kername_set_mem_iff Σ c + : KernameSet.mem c (declared_kername_set Σ) <-> List.In c (map fst Σ). + Proof. + setoid_rewrite <- KernameSetFact.mem_iff. + apply declared_kername_set_spec. + Qed. + Lemma extends_r_merge_globals Σ Σ' : compatible_globals Σ Σ' -> forall c, @@ -698,24 +777,13 @@ Module Environment (T : Term). intro c. specialize (H2 c). rewrite lookup_globals_app lookup_globals_filter. - set (s := fold_right _ _ _). - assert (Hs : forall kn, KernameSet.mem kn s <-> List.In kn (map fst Σ)). - { setoid_rewrite <- KernameSetFact.mem_iff. - clear. - generalize dependent Σ; clear. - elim => //=; try setoid_rewrite KernameSetFact.empty_iff => //=. - move => [? ?] ? IH kn //=. - rewrite KernameSet.add_spec. - intuition. } - clearbody s. - specialize (Hs c). destruct (in_dec eq_dec c (map fst Σ')) as [H'|H']; [ exists nil | exists (lookup_globals Σ c) ]. 2: apply lookup_globals_nil in H'; rewrite H'; clear H'. 2: now destruct ?; cbn; rewrite app_nil_r. pose proof (lookup_globals_nil Σ c) as Hc. - rewrite <- !lookup_globals_nil, <- Hs in H2. - rewrite <- Hs in Hc. + rewrite <- !lookup_globals_nil in H2. + rewrite <- (declared_kername_set_mem_iff Σ) in *. destruct KernameSet.mem; cbn in *. { intuition. } { destruct Hc as [Hc _]. diff --git a/common/theories/EnvironmentReflect.v b/common/theories/EnvironmentReflect.v index 63c5f9571..10ea249c3 100644 --- a/common/theories/EnvironmentReflect.v +++ b/common/theories/EnvironmentReflect.v @@ -12,7 +12,7 @@ Module EnvironmentReflect (T : Term) (Import E : EnvironmentSig T) (Import TDec := (skipn (#|Σ'.(declarations)| - #|Σ.(declarations)|) Σ'.(declarations) == Σ.(declarations)) (only parsing). Lemma extends_decls_partT (Σ Σ' : global_env) - : reflectT (forall c, ∑ decls, lookup_envs Σ' c = decls ++ lookup_envs Σ c) (extendsb_decls_part Σ Σ'). + : reflectT (extends_decls_part Σ Σ') (extendsb_decls_part Σ Σ'). Proof. case: (@forallbP _ _ _ _ (fun _ => eqb_specT _ _)) => H; constructor. all: setoid_rewrite Forall_forall in H. @@ -35,7 +35,7 @@ Module EnvironmentReflect (T : Term) (Import E : EnvironmentSig T) (Import TDec Qed. Lemma strictly_extends_decls_partT (Σ Σ' : global_env) - : reflectT (∑ Σ'', declarations Σ' = Σ'' ++ declarations Σ) (strictly_extendsb_decls_part Σ Σ'). + : reflectT (strictly_extends_decls_part Σ Σ') (strictly_extendsb_decls_part Σ Σ'). Proof. case: eqb_specT => H; constructor. { rewrite -H. @@ -112,6 +112,48 @@ Module EnvironmentReflect (T : Term) (Import E : EnvironmentSig T) (Import TDec case: strictly_extends_declsT; split; intuition congruence. Qed. + Definition compatible_globalsb (Σ Σ' : global_declarations) : bool + := KernameSet.for_all + (fun c => lookup_globals Σ c == lookup_globals Σ' c) + (KernameSet.inter (declared_kername_set Σ) (declared_kername_set Σ')). + + Definition compatibleb (Σ Σ' : global_env) : bool + := Retroknowledge.compatible Σ.(retroknowledge) Σ'.(retroknowledge) + && compatible_globalsb Σ.(declarations) Σ'.(declarations). + + Lemma compatible_globalsP Σ Σ' : reflect (compatible_globals Σ Σ') (compatible_globalsb Σ Σ'). + Proof. + rewrite /compatible_globals/compatible_globalsb. + lazymatch goal with + | [ |- context[KernameSet.for_all ?f ?s] ] + => generalize (@KernameSet.for_all_spec s f _) + end. + rewrite /KernameSet.For_all. + case: KernameSet.for_all; move => [H1 H2]; constructor. + all: first [ specialize (H1 eq_refl); rename H1 into H | clear H1 ]. + all: try pose proof (fun x => diff_false_true (H2 x)) as H; clear H2. + all: move: H. + all: setoid_rewrite KernameSet.inter_spec. + all: setoid_rewrite declared_kername_set_spec. + all: setoid_rewrite <- lookup_globals_nil. + 1: move => H c; move: (H c); clear H. + 2: move => H nH; apply H => c; move: (nH c); clear H nH. + all: case: eqb_spec. + all: intuition. + Qed. + + Lemma compatible_globals_spec Σ Σ' : compatible_globalsb Σ Σ' <-> compatible_globalsb Σ Σ'. + Proof. case: compatible_globalsP; intuition. Qed. + + Lemma compatibleP Σ Σ' : reflect (compatible Σ Σ') (compatibleb Σ Σ'). + Proof. + rewrite /compatible/compatibleb. + case: compatible_globalsP; destruct Retroknowledge.compatible; cbn; constructor; intuition. + Qed. + + Lemma compatible_spec Σ Σ' : compatibleb Σ Σ' <-> compatibleb Σ Σ'. + Proof. case: compatibleP; intuition. Qed. + End EnvironmentReflect. Module Type EnvironmentReflectSig (T : Term) (E : EnvironmentSig T) (TDec : TermDecide T) (EDec : EnvironmentDecide T E). From 6ddfd1392c8df6960b9fdf6517288dca4d7b6585 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Sun, 23 Apr 2023 22:12:36 -0700 Subject: [PATCH 05/61] More extension lemmas --- common/theories/Environment.v | 54 +++++++++++++++++++++++++++-------- 1 file changed, 42 insertions(+), 12 deletions(-) diff --git a/common/theories/Environment.v b/common/theories/Environment.v index 700d6af68..15d94d79b 100644 --- a/common/theories/Environment.v +++ b/common/theories/Environment.v @@ -537,6 +537,20 @@ Module Environment (T : Term). rewrite (proj1 (@lookup_global_None _ _)) => //= -> //=. Qed. + Lemma NoDup_lookup_globals_eq Σ + : NoDup (List.map fst Σ) + -> forall kn, lookup_globals Σ kn = match lookup_global Σ kn with + | Some v => [v] + | None => [] + end. + Proof. + move => H kn. + move: (NoDup_length_lookup_globals Σ H kn) (hd_error_lookup_globals Σ kn). + repeat destruct ?; subst. + all: case: lookup_globals; cbn; try congruence. + move => ? [|]; cbn; congruence. + Qed. + Lemma lookup_globals_In Σ kn decl : In (kn, decl) Σ <-> In decl (lookup_globals Σ kn). Proof. @@ -725,16 +739,6 @@ Module Environment (T : Term). Definition compatible_globals (Σ Σ' : global_declarations) : Prop := forall c, lookup_globals Σ c <> [] -> lookup_globals Σ' c <> [] -> lookup_globals Σ c = lookup_globals Σ' c. - Lemma extends_strictly_on_decls_l_merge Σ Σ' - : extends_strictly_on_decls Σ (merge_global_envs Σ Σ'). - Proof. - rewrite /extends_strictly_on_decls/merge_global_envs/merge_globals; cbn. - split; - try first [ apply ContextSet.union_spec - | apply Retroknowledge.extends_l_merge - | now eexists ]. - Qed. - Definition compatible (Σ Σ' : global_env) := Retroknowledge.compatible Σ.(retroknowledge) Σ'.(retroknowledge) /\ compatible_globals Σ.(declarations) Σ'.(declarations). @@ -750,6 +754,33 @@ Module Environment (T : Term). all: try congruence. Qed. + Lemma strictly_extends_decls_l_merge_globals Σ Σ' + : strictly_extends_decls_part_globals Σ (merge_globals Σ Σ'). + Proof. now eexists. Qed. + + Lemma extends_l_merge_globals Σ Σ' + : extends_decls_part_globals Σ (merge_globals Σ Σ'). + Proof. + rewrite /merge_globals. + intro c. + rewrite lookup_globals_app lookup_globals_filter. + eexists; reflexivity. + Qed. + + #[export] Instance extends_strictly_on_decls_l_merge Σ Σ' + : extends_strictly_on_decls Σ (merge_global_envs Σ Σ'). + Proof. + rewrite /extends_strictly_on_decls/merge_global_envs/merge_globals; cbn. + split; + try first [ apply ContextSet.union_spec + | apply Retroknowledge.extends_l_merge + | apply strictly_extends_decls_l_merge_globals ]. + Qed. + + Lemma extends_l_merge Σ Σ' + : extends Σ (merge_global_envs Σ Σ'). + Proof. exact _. Qed. + Lemma declared_kername_set_spec : forall Σ c, KernameSet.In c (declared_kername_set Σ) <-> List.In c (map fst Σ). Proof. @@ -768,8 +799,7 @@ Module Environment (T : Term). Lemma extends_r_merge_globals Σ Σ' : compatible_globals Σ Σ' -> - forall c, - ∑ decls, lookup_globals (merge_globals Σ Σ') c = decls ++ lookup_globals Σ' c. + extends_decls_part_globals Σ' (merge_globals Σ Σ'). Proof. rewrite /merge_globals. intro H2; cbn. From 63074670193f3bf655f6d82c176c68a77518ef08 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Sun, 23 Apr 2023 23:33:39 -0700 Subject: [PATCH 06/61] Fix issue with hint breaking a proof --- common/theories/Environment.v | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/common/theories/Environment.v b/common/theories/Environment.v index 15d94d79b..aaa015c1b 100644 --- a/common/theories/Environment.v +++ b/common/theories/Environment.v @@ -767,7 +767,7 @@ Module Environment (T : Term). eexists; reflexivity. Qed. - #[export] Instance extends_strictly_on_decls_l_merge Σ Σ' + Lemma extends_strictly_on_decls_l_merge Σ Σ' : extends_strictly_on_decls Σ (merge_global_envs Σ Σ'). Proof. rewrite /extends_strictly_on_decls/merge_global_envs/merge_globals; cbn. @@ -776,6 +776,7 @@ Module Environment (T : Term). | apply Retroknowledge.extends_l_merge | apply strictly_extends_decls_l_merge_globals ]. Qed. + #[export] Hint Extern 0 (extends_strictly_on_decls _ (merge_global_envs ?Σ ?Σ')) => simple apply (@extends_strictly_on_decls_l_merge Σ Σ') : typeclass_instances. Lemma extends_l_merge Σ Σ' : extends Σ (merge_global_envs Σ Σ'). From 6c63eacb5ccd00defbdb24b82382325e50730f4b Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Wed, 3 May 2023 06:06:38 -0700 Subject: [PATCH 07/61] Add union and inter checker flags (#957) This will allow programatically combining checker flags from multiple constants in Gallina quotation. --- common/theories/config.v | 72 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 72 insertions(+) diff --git a/common/theories/config.v b/common/theories/config.v index edfc01c93..eb89636e3 100644 --- a/common/theories/config.v +++ b/common/theories/config.v @@ -71,3 +71,75 @@ Qed. impl_refl impl_crefl impl_trans impl_ctrans | 100. + +Definition laxest_checker_flags : checker_flags := {| + check_univs := false ; + prop_sub_type := true; + indices_matter := false; + lets_in_constructor_types := true +|}. + +Definition strictest_checker_flags : checker_flags := {| + check_univs := true ; + prop_sub_type := false; + indices_matter := true; + lets_in_constructor_types := false +|}. + +Lemma laxest_checker_flags_laxest : forall cf, impl cf laxest_checker_flags. +Proof. + intro cf; cbv; destruct cf. + repeat match goal with + | [ |- context[match ?x with _ => _ end] ] => case_eq x + end; try reflexivity; congruence. +Qed. + +Lemma strictest_checker_flags_strictest : forall cf, impl strictest_checker_flags cf. +Proof. + intro cf; cbv; destruct cf. + repeat match goal with + | [ |- context[match ?x with _ => _ end] ] => case_eq x + end; try reflexivity; congruence. +Qed. + +(** can check everything that either one can check *) +Definition union_checker_flags (cf1 cf2 : checker_flags) : checker_flags + := {| check_univs := andb cf2.(@check_univs) cf1.(@check_univs) + ; prop_sub_type := orb cf1.(@prop_sub_type) cf2.(@prop_sub_type) + ; indices_matter := andb cf2.(@indices_matter) cf1.(@indices_matter) + ; lets_in_constructor_types := orb cf1.(@lets_in_constructor_types) cf2.(@lets_in_constructor_types) |}. + +(** can check everything that both can check *) +Definition inter_checker_flags (cf1 cf2 : checker_flags) : checker_flags + := {| check_univs := orb cf2.(@check_univs) cf1.(@check_univs) + ; prop_sub_type := andb cf1.(@prop_sub_type) cf2.(@prop_sub_type) + ; indices_matter := orb cf2.(@indices_matter) cf1.(@indices_matter) + ; lets_in_constructor_types := andb cf1.(@lets_in_constructor_types) cf2.(@lets_in_constructor_types) |}. + +Lemma union_checker_flags_spec cf1 cf2 (cf' := union_checker_flags cf1 cf2) + : impl cf1 cf' /\ impl cf2 cf' /\ (forall cf'', impl cf1 cf'' -> impl cf2 cf'' -> impl cf' cf''). +Proof. + destruct cf1, cf2; subst cf'. + cbv. + repeat split. + 3: intro cf''; destruct cf''. + all: repeat first [ match goal with + | [ |- context[match ?x with _ => _ end] ] => is_var x; destruct x + end + | reflexivity + | congruence ]. +Qed. + +Lemma inter_checker_flags_spec cf1 cf2 (cf' := inter_checker_flags cf1 cf2) + : impl cf' cf1 /\ impl cf' cf2 /\ (forall cf'', impl cf'' cf1 -> impl cf'' cf2 -> impl cf'' cf'). +Proof. + destruct cf1, cf2; subst cf'. + cbv. + repeat split. + 3: intro cf''; destruct cf''. + all: repeat first [ match goal with + | [ |- context[match ?x with _ => _ end] ] => is_var x; destruct x + end + | reflexivity + | congruence ]. +Qed. From 198e2023bec459b1ebae5e3c6d38f4fdc8cd819d Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 5 May 2023 08:26:45 +0200 Subject: [PATCH 08/61] Fix sprop_level/prop_level requring lazy evaluation --- template-coq/src/constr_quoter.ml | 32 +++++++++++++++---------------- 1 file changed, 16 insertions(+), 16 deletions(-) diff --git a/template-coq/src/constr_quoter.ml b/template-coq/src/constr_quoter.ml index f796d4df4..8d2240daa 100644 --- a/template-coq/src/constr_quoter.ml +++ b/template-coq/src/constr_quoter.ml @@ -134,7 +134,7 @@ struct match Constr.kind ty with | Constr.Ind (ind, u) -> ind | _ -> failwith "byte_ind : tByte is not bound to an inductive type") - + let quote_char i = Constr.mkConstruct (Lazy.force byte_ind, (i+1)) @@ -168,7 +168,7 @@ struct match r with | Sorts.Relevant -> Lazy.force tRelevant | Sorts.Irrelevant -> Lazy.force tIrrelevant - + let quote_name n = match n with Names.Name id -> constr_mkApp (nNamed, [| quote_ident id |]) @@ -287,7 +287,7 @@ struct | Some var -> let var' = quote_cuminfo_variance var in constr_mkApp (cSome, [| listvar; var' |]) *) - + let quote_abstract_univ_context uctx = let arr = (AbstractContext.names uctx) in let idents = to_coq_listl tname (CArray.map_to_list quote_name arr) in @@ -299,10 +299,10 @@ struct let mkPolymorphic_ctx t = constr_mkApp (cPolymorphic_ctx, [|t|]) - let mkMonomorphic_entry ctx = + let mkMonomorphic_entry ctx = constr_mkApp (cMonomorphic_entry, [| ctx |]) - - let mkPolymorphic_entry ctx = + + let mkPolymorphic_entry ctx = constr_mkApp (cPolymorphic_entry, [| ctx |]) let quote_inductive_universes uctx = @@ -323,14 +323,14 @@ struct let sprop_level = - constr_mkApp (cInl, [|Lazy.force tproplevel;Lazy.force tlevel;Lazy.force tlevelSProp |]) + lazy (constr_mkApp (cInl, [|Lazy.force tproplevel;Lazy.force tlevel;Lazy.force tlevelSProp |])) let prop_level = - constr_mkApp (cInl, [|Lazy.force tproplevel;Lazy.force tlevel;Lazy.force tlevelProp |]) + lazy (constr_mkApp (cInl, [|Lazy.force tproplevel;Lazy.force tlevel;Lazy.force tlevelProp |])) let quote_sort s = match s with | Sorts.Set -> quote_universe Universe.type0 - | Sorts.Prop -> constr_mkApp (tof_levels, [| prop_level |]) - | Sorts.SProp -> constr_mkApp (tof_levels, [| sprop_level |]) + | Sorts.Prop -> constr_mkApp (tof_levels, [| Lazy.force prop_level |]) + | Sorts.SProp -> constr_mkApp (tof_levels, [| Lazy.force sprop_level |]) | Sorts.Type u -> quote_universe u let quote_sort_family = function @@ -344,14 +344,14 @@ struct let quote_context ctx = to_coq_list (constr_mkAppl (tcontext_decl, [| tTerm |])) ctx - + let mk_ctor_list ls = - let ctors = List.map (fun (a,b,c,d,e) -> + let ctors = List.map (fun (a,b,c,d,e) -> constr_mkApp (tBuild_constructor_body, [| a ; b ; to_coq_listl tTerm c ; d ; e |])) ls in to_coq_listl tconstructor_body ctors let mk_proj_list ps = - let projs = List.map (fun (a,b,c) -> + let projs = List.map (fun (a,b,c) -> constr_mkApp (tBuild_projection_body, [| a ; b ; c |])) ps in to_coq_listl tprojection_body projs @@ -408,11 +408,11 @@ struct let pair = pairl tkername tglobal_decl kn d in constr_mkApp (c_cons, [| global_pairty (); pair; l|]) - type pre_quoted_retroknowledge = + type pre_quoted_retroknowledge = { retro_int63 : quoted_kernel_name option; retro_float64 : quoted_kernel_name option } - let quote_retroknowledge r = + let quote_retroknowledge r = let rint63 = to_coq_option (Lazy.force tkername) (fun x -> x) r.retro_int63 in let rfloat64 = to_coq_option (Lazy.force tkername) (fun x -> x) r.retro_float64 in constr_mkApp (tmk_retroknowledge, [| rint63; rfloat64 |]) @@ -432,7 +432,7 @@ struct let consnames = to_coq_listl tident consnames in let constypes = to_coq_listl tTerm constypes in constr_mkApp (tBuild_one_inductive_entry, [| iname; arity; consnames; constypes |]) - + let quote_mutual_inductive_entry (mf, mp, is, mpol, var) = let is = to_coq_listl tOne_inductive_entry (List.map make_one_inductive_entry is) in let mpr = constr_mkAppl (cNone, [|bool_type|]) in From c7b8d8fd90a8a612493f45cea7cc15c46d08e993 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 5 May 2023 08:29:08 +0200 Subject: [PATCH 09/61] Fix use of lazy for (s)prop_level --- template-coq/src/constr_quoter.ml | 32 +++++++++++++++---------------- 1 file changed, 16 insertions(+), 16 deletions(-) diff --git a/template-coq/src/constr_quoter.ml b/template-coq/src/constr_quoter.ml index f796d4df4..8d2240daa 100644 --- a/template-coq/src/constr_quoter.ml +++ b/template-coq/src/constr_quoter.ml @@ -134,7 +134,7 @@ struct match Constr.kind ty with | Constr.Ind (ind, u) -> ind | _ -> failwith "byte_ind : tByte is not bound to an inductive type") - + let quote_char i = Constr.mkConstruct (Lazy.force byte_ind, (i+1)) @@ -168,7 +168,7 @@ struct match r with | Sorts.Relevant -> Lazy.force tRelevant | Sorts.Irrelevant -> Lazy.force tIrrelevant - + let quote_name n = match n with Names.Name id -> constr_mkApp (nNamed, [| quote_ident id |]) @@ -287,7 +287,7 @@ struct | Some var -> let var' = quote_cuminfo_variance var in constr_mkApp (cSome, [| listvar; var' |]) *) - + let quote_abstract_univ_context uctx = let arr = (AbstractContext.names uctx) in let idents = to_coq_listl tname (CArray.map_to_list quote_name arr) in @@ -299,10 +299,10 @@ struct let mkPolymorphic_ctx t = constr_mkApp (cPolymorphic_ctx, [|t|]) - let mkMonomorphic_entry ctx = + let mkMonomorphic_entry ctx = constr_mkApp (cMonomorphic_entry, [| ctx |]) - - let mkPolymorphic_entry ctx = + + let mkPolymorphic_entry ctx = constr_mkApp (cPolymorphic_entry, [| ctx |]) let quote_inductive_universes uctx = @@ -323,14 +323,14 @@ struct let sprop_level = - constr_mkApp (cInl, [|Lazy.force tproplevel;Lazy.force tlevel;Lazy.force tlevelSProp |]) + lazy (constr_mkApp (cInl, [|Lazy.force tproplevel;Lazy.force tlevel;Lazy.force tlevelSProp |])) let prop_level = - constr_mkApp (cInl, [|Lazy.force tproplevel;Lazy.force tlevel;Lazy.force tlevelProp |]) + lazy (constr_mkApp (cInl, [|Lazy.force tproplevel;Lazy.force tlevel;Lazy.force tlevelProp |])) let quote_sort s = match s with | Sorts.Set -> quote_universe Universe.type0 - | Sorts.Prop -> constr_mkApp (tof_levels, [| prop_level |]) - | Sorts.SProp -> constr_mkApp (tof_levels, [| sprop_level |]) + | Sorts.Prop -> constr_mkApp (tof_levels, [| Lazy.force prop_level |]) + | Sorts.SProp -> constr_mkApp (tof_levels, [| Lazy.force sprop_level |]) | Sorts.Type u -> quote_universe u let quote_sort_family = function @@ -344,14 +344,14 @@ struct let quote_context ctx = to_coq_list (constr_mkAppl (tcontext_decl, [| tTerm |])) ctx - + let mk_ctor_list ls = - let ctors = List.map (fun (a,b,c,d,e) -> + let ctors = List.map (fun (a,b,c,d,e) -> constr_mkApp (tBuild_constructor_body, [| a ; b ; to_coq_listl tTerm c ; d ; e |])) ls in to_coq_listl tconstructor_body ctors let mk_proj_list ps = - let projs = List.map (fun (a,b,c) -> + let projs = List.map (fun (a,b,c) -> constr_mkApp (tBuild_projection_body, [| a ; b ; c |])) ps in to_coq_listl tprojection_body projs @@ -408,11 +408,11 @@ struct let pair = pairl tkername tglobal_decl kn d in constr_mkApp (c_cons, [| global_pairty (); pair; l|]) - type pre_quoted_retroknowledge = + type pre_quoted_retroknowledge = { retro_int63 : quoted_kernel_name option; retro_float64 : quoted_kernel_name option } - let quote_retroknowledge r = + let quote_retroknowledge r = let rint63 = to_coq_option (Lazy.force tkername) (fun x -> x) r.retro_int63 in let rfloat64 = to_coq_option (Lazy.force tkername) (fun x -> x) r.retro_float64 in constr_mkApp (tmk_retroknowledge, [| rint63; rfloat64 |]) @@ -432,7 +432,7 @@ struct let consnames = to_coq_listl tident consnames in let constypes = to_coq_listl tTerm constypes in constr_mkApp (tBuild_one_inductive_entry, [| iname; arity; consnames; constypes |]) - + let quote_mutual_inductive_entry (mf, mp, is, mpol, var) = let is = to_coq_listl tOne_inductive_entry (List.map make_one_inductive_entry is) in let mpr = constr_mkAppl (cNone, [|bool_type|]) in From f91a1e7798cd6d2b62a22d6416ae82d8ca8ab2fd Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Fri, 5 May 2023 12:36:39 +0200 Subject: [PATCH 10/61] close some computational obligations with defined in erase_global_decls --- erasure/theories/ErasureFunction.v | 20 ++++++++++---------- 1 file changed, 10 insertions(+), 10 deletions(-) diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index fba458ddf..421898b47 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -924,7 +924,7 @@ Qed. Next Obligation. cbv [id]. unshelve eapply (normalization_in 0); cbn; try reflexivity; try lia. -Qed. +Defined. Next Obligation. pose proof (abstract_env_ext_wf _ H) as [wf]. pose proof (abstract_env_exists (abstract_pop_decls X)) as [[? HX']]. @@ -941,7 +941,7 @@ Next Obligation. Qed. Next Obligation. unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Qed. +Defined. Next Obligation. pose proof (abstract_env_exists X) as [[? HX]]. assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). @@ -950,7 +950,7 @@ Next Obligation. Qed. Next Obligation. unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Qed. +Defined. Next Obligation. pose proof (abstract_env_exists X) as [[? HX]]. assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). @@ -959,7 +959,7 @@ now pose proof (abstract_pop_decls_correct X decls prop' _ _ HX H). Qed. Next Obligation. unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Qed. +Defined. Next Obligation. pose proof (abstract_env_exists X) as [[? HX]]. assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). @@ -968,7 +968,7 @@ Next Obligation. Qed. Next Obligation. unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Qed. +Defined. Next Obligation. pose proof (abstract_env_exists X) as [[? HX]]. assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). @@ -2719,7 +2719,7 @@ Qed. Next Obligation. cbv [id]. unshelve eapply (normalization_in 0); cbn; try reflexivity; try lia. -Qed. +Defined. Next Obligation. pose proof (abstract_env_ext_wf _ H) as [?]. pose proof (abstract_env_exists X) as [[? wf]]. @@ -2732,28 +2732,28 @@ Next Obligation. Qed. Next Obligation. unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Qed. +Defined. Next Obligation. specialize_Σ H. sq; red. destruct prop as [Σ' ->]. eexists (Σ' ++ [(kn, ConstantDecl cb)]); rewrite -app_assoc //. Qed. Next Obligation. unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Qed. +Defined. Next Obligation. specialize_Σ H. sq; red. destruct prop as [Σ' ->]. eexists (Σ' ++ [(kn, ConstantDecl cb)]); rewrite -app_assoc //. Qed. Next Obligation. unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Qed. +Defined. Next Obligation. specialize_Σ H. sq; red. destruct prop as [Σ' ->]. eexists (Σ' ++ [(kn, _)]); rewrite -app_assoc //. Qed. Next Obligation. unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Qed. +Defined. Next Obligation. specialize_Σ H. sq; red. destruct prop as [Σ' ->]. eexists (Σ' ++ [(kn, _)]); rewrite -app_assoc //. From e04bc87cf93427acd8d76c937cc0efb0e2c23a3b Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Fri, 5 May 2023 12:42:39 -0700 Subject: [PATCH 11/61] Add MCListable class for enumerating finite types This will be useful for proving that quotation of byte is well-typed with the safechecker. --- utils/_CoqProject | 1 + utils/theories/MCListable.v | 169 ++++++++++++++++++++++++++++++++++++ 2 files changed, 170 insertions(+) create mode 100644 utils/theories/MCListable.v diff --git a/utils/_CoqProject b/utils/_CoqProject index 6ab957809..9f0b08c4e 100644 --- a/utils/_CoqProject +++ b/utils/_CoqProject @@ -17,6 +17,7 @@ theories/MCEquality.v theories/MCFSets.v theories/LibHypsNaming.v theories/MCList.v +theories/MCListable.v theories/MCMSets.v theories/MCOption.v theories/MCProd.v diff --git a/utils/theories/MCListable.v b/utils/theories/MCListable.v new file mode 100644 index 000000000..dee720445 --- /dev/null +++ b/utils/theories/MCListable.v @@ -0,0 +1,169 @@ +From Coq Require Import Lia Arith Bool List Program. +From MetaCoq.Utils Require Import MCPrelude ReflectEq. +From Coq.ssr Require Import ssreflect. +From Equations Require Import Equations. +Import ListNotations. + +Set Universe Polymorphism. +Set Polymorphic Inductive Cumulativity. + +Class Listable T := { list_all : list T ; find_index : T -> nat ; cardinality := length list_all }. + +Arguments find_index {T} {_} _. +Arguments list_all T {_}. +Arguments cardinality T {_}. + +Class FinitelyListable T {l : Listable T} + := { find_index_ok : forall t, nth_error (list_all T) (find_index t) = Some t + ; enumerate_smallb : forallb (fun '(n, t) => n == find_index t) + (combine (seq 0 (cardinality T)) (list_all T)) }. + +Class FinitelyListablePackage T + := { FinitelyListablePackage_Listable : Listable T + ; FinitelyListablePackage_FinitelyListable : FinitelyListable T }. +#[export] Existing Instances FinitelyListablePackage_Listable FinitelyListablePackage_FinitelyListable. +#[global] Coercion FinitelyListablePackage_Listable : FinitelyListablePackage >-> Listable. +#[global] Coercion FinitelyListablePackage_FinitelyListable : FinitelyListablePackage >-> FinitelyListable. + +Ltac prove_FinitelyListable := + split; + [ let t := fresh in + intro t; destruct t; + vm_compute; try reflexivity + | vm_compute; reflexivity ]. + +Ltac push_S_refine_num ev := + lazymatch ev with + | S ?ev + => refine (S _); push_S_refine_num ev + | ?ev => let __ := open_constr:(eq_refl : ev = S _) in + refine O + end. + +Ltac force_num ev := + lazymatch ev with + | S ?ev => force_num ev + | ?ev => unify ev O + end. + +Ltac partially_prove_Listable := + let fin := open_constr:(_) in + unshelve esplit; + [ + | repeat (let t := fresh in intro t; case t; clear t); + push_S_refine_num fin ]; + force_num fin. + +Ltac finish_prove_FinitelyListable := + split; + [ cbv [list_all find_index]; + repeat (let t := fresh in intro t; case t; clear t); + unshelve + (repeat match goal with + | [ |- nth_error (?x :: _) O = ?rhs ] => change (Some x = rhs) + | [ |- nth_error (_ :: ?xs) (S ?n) = ?rhs ] => change (nth_error xs n = rhs) + | [ |- nth_error ?ev _ = ?rhs ] + => is_evar ev; + let __ := open_constr:(eq_refl : ev = cons _ _) in + idtac + | [ |- Some _ = Some _ ] => reflexivity + end); + try exact nil + | try instantiate (1:=nil) (* for empty types *); + vm_compute; reflexivity ]. + +Ltac prove_ListableDerive := + lazymatch goal with + | [ |- @FinitelyListable ?T ?H ] + => instantiate (1:=ltac:(partially_prove_Listable)) in (value of H); + cbv [H]; + finish_prove_FinitelyListable + end. + +Ltac prove_ListablePackage T := + refine (_ : { l : Listable T | FinitelyListable T }); + unshelve econstructor; + [ partially_prove_Listable + | finish_prove_FinitelyListable ]. + +Ltac prove_FinitelyListablePackage_goal _ := + unshelve econstructor; + [ partially_prove_Listable + | finish_prove_FinitelyListable ]. +Ltac prove_FinitelyListablePackage T := + refine (_ : FinitelyListablePackage T); + prove_FinitelyListablePackage_goal (). + +Ltac get_ListablePackage T := + constr:(ltac:(prove_ListablePackage T)). + +Section with_listable. + Context {T} + {Listable_T : Listable T} + {FinitelyListable_T : FinitelyListable T}. + + Lemma find_index_iff x n : find_index x = n <-> List.nth_error (list_all T) n = Some x. + Proof using FinitelyListable_T. + split; [ intro; subst; apply find_index_ok | ]. + generalize enumerate_smallb. + cbv [cardinality]. + set (offset := O). + change n with (offset + n)%nat at 2. + clearbody offset. + rename x into v. + intros Hfold Hnth; revert offset n Hfold Hnth. + induction (list_all T) as [|x xs IHxs]; intros. + { destruct n; cbn in *; congruence. } + { cbn in *. + move: Hfold; case: eqb_specT => //= Hfold1 Hfold2. + specialize (IHxs (S offset) (pred n) Hfold2). + subst offset; destruct n as [|n]; cbn in *. + { inversion Hnth; subst; lia. } + { specialize (IHxs Hnth); lia. } } + Qed. + + Lemma enumerate_unique n m x y + (Hn : List.nth_error (list_all T) n = Some x) + (Hm : List.nth_error (list_all T) m = Some y) + (Hxy : x = y) + : n = m. + Proof using FinitelyListable_T. rewrite <- !find_index_iff in *; subst; reflexivity. Qed. + + Lemma find_index_inj x y : find_index x = find_index y -> x = y. + Proof using FinitelyListable_T. rewrite find_index_iff find_index_ok; congruence. Qed. + + Definition eqb_of_listable : T -> T -> bool := fun x y => Nat.eqb (find_index x) (find_index y). + Lemma eqb_of_listable_refl x : eqb_of_listable x x = true. + Proof using Type. cbv [eqb_of_listable]. apply Nat.eqb_refl. Qed. + Lemma eqb_of_listable_lb x y : x = y -> eqb_of_listable x y = true. + Proof using Type. intros; subst; apply eqb_of_listable_refl. Qed. + Lemma eqb_of_listable_bl x y : eqb_of_listable x y = true -> x = y. + Proof using FinitelyListable_T. cbv [eqb_of_listable]; rewrite Nat.eqb_eq. apply find_index_inj. Qed. + Lemma eqb_of_listable_eq x y : eqb_of_listable x y = true <-> x = y. + Proof using FinitelyListable_T. split; auto using eqb_of_listable_lb, eqb_of_listable_bl. Qed. + Lemma eqb_of_listable_spec x y : reflectProp (x = y) (eqb_of_listable x y). + Proof using FinitelyListable_T. generalize (eqb_of_listable_eq x y); case: eqb_of_listable; case; constructor; auto. Qed. + + #[local] Instance ReflectEq_of_listable : ReflectEq T + := { eqb := eqb_of_listable + ; eqb_spec := eqb_of_listable_spec }. +End with_listable. + +#[export] Existing Instance ReflectEq_of_listable. + +Local Obligation Tactic := try prove_FinitelyListablePackage_goal (). +#[export,program] Instance unit_FinitelyListablePackage : FinitelyListablePackage unit := _. +#[export,program] Instance Empty_set_FinitelyListablePackage : FinitelyListablePackage Empty_set := _. +#[export,program] Instance True_FinitelyListablePackage : FinitelyListablePackage True := _. +#[export,program] Instance False_FinitelyListablePackage : FinitelyListablePackage False := _. +#[export,program] Instance bool_FinitelyListablePackage : FinitelyListablePackage bool := _. +#[export] Instance byte_FinitelyListablePackage : FinitelyListablePackage Byte.byte. +Proof. + unshelve econstructor. + { unshelve esplit; [ | exact Byte.to_nat ]. + { let ls := (eval cbv in (List.flat_map (fun v => match Byte.of_nat v with Some v => [v] | None => [] end) (seq 0 256))) in + exact ls. } } + { split. + { intro x; destruct x; vm_compute; reflexivity. } + { vm_compute; reflexivity. } } +Defined. From e578a027c9cc9da2b9e84bd331dea5609e7b1be0 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 9 May 2023 09:13:47 +0200 Subject: [PATCH 12/61] Move nix actions away --- ...ix-action-coq-8.16-macos.yml => nix-action-coq-8.17-macos.yml} | 0 ...-action-coq-8.16-ubuntu.yml => nix-action-coq-8.17-ubuntu.yml} | 0 2 files changed, 0 insertions(+), 0 deletions(-) rename .github/{workflows/nix-action-coq-8.16-macos.yml => nix-action-coq-8.17-macos.yml} (100%) rename .github/{workflows/nix-action-coq-8.16-ubuntu.yml => nix-action-coq-8.17-ubuntu.yml} (100%) diff --git a/.github/workflows/nix-action-coq-8.16-macos.yml b/.github/nix-action-coq-8.17-macos.yml similarity index 100% rename from .github/workflows/nix-action-coq-8.16-macos.yml rename to .github/nix-action-coq-8.17-macos.yml diff --git a/.github/workflows/nix-action-coq-8.16-ubuntu.yml b/.github/nix-action-coq-8.17-ubuntu.yml similarity index 100% rename from .github/workflows/nix-action-coq-8.16-ubuntu.yml rename to .github/nix-action-coq-8.17-ubuntu.yml From 14ad024c41da54c73dc95588f99e0a84a66b71f0 Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Fri, 2 Jun 2023 14:09:49 +0200 Subject: [PATCH 13/61] Invariants in named recursion rule (#967) --- erasure/theories/EWcbvEvalNamed.v | 238 +++++++++++++++--------------- 1 file changed, 119 insertions(+), 119 deletions(-) diff --git a/erasure/theories/EWcbvEvalNamed.v b/erasure/theories/EWcbvEvalNamed.v index 6bdbeae78..08e77b7d4 100644 --- a/erasure/theories/EWcbvEvalNamed.v +++ b/erasure/theories/EWcbvEvalNamed.v @@ -1,7 +1,6 @@ (* Distributed under the terms of the MIT license. *) From Coq Require Import Utf8 Program. -From MetaCoq.Utils Require Import utils. -From MetaCoq.Common Require Import config BasicAst. +From MetaCoq Require Import config utils BasicAst. From MetaCoq.PCUIC Require PCUICWcbvEval. From MetaCoq.Erasure Require Import EAst EAstUtils ELiftSubst ECSubst EReflect EGlobalEnv EWellformed EWcbvEval. @@ -37,7 +36,7 @@ Inductive value : Set := Definition environment := list (ident * value). Definition add : ident -> value -> environment -> environment := fun (x : ident) (v : value) env => (x, v) :: env. -Fixpoint add_multiple (xs : list ident) (vs : list value) (env : environment) : environment := +Fixpoint add_multiple (xs : list ident) (vs : list value) (env : environment) {struct vs} : environment := match xs, vs with | x :: xs, v :: vs => add x v (add_multiple xs vs env) | _, _ => env @@ -114,10 +113,12 @@ Section Wcbv. (** Fix unfolding, without guard *) | eval_fix_unfold f mfix idx a na na' av fn res Γ' : + forallb (λ x, isLambda (snd x)) mfix -> + ~ In na (map fst mfix) -> eval Γ f (vRecClos mfix idx Γ') -> NoDup (map fst mfix) -> nth_error mfix idx = Some (na', tLambda (nNamed na) fn) -> - eval (add na av (add_multiple (map fst mfix) (fix_env mfix Γ') Γ')) fn res -> + eval (add na av (add_multiple (List.rev (map fst mfix)) (fix_env mfix Γ') (Γ'))) fn res -> eval Γ a av -> eval Γ (tApp f a) res @@ -137,7 +138,7 @@ Section Wcbv. (** Constructor congruence: we do not allow over-applications *) | eval_construct_block ind c mdecl idecl cdecl args args' : lookup_constructor Σ ind c = Some (mdecl, idecl, cdecl) -> - #|args| <= cstr_arity mdecl cdecl -> + #|args| <= cstr_nargs cdecl -> All2_Set (eval Γ) args args' -> eval Γ (tConstruct ind c args) (vConstruct ind c args') @@ -186,6 +187,8 @@ Section Wcbv. (res : value) (Γ' : list (ident × value)) + (Hbodies : forallb (λ x, isLambda (snd x)) mfix) + (Hfresh : ~ In na (map fst mfix)) (e : eval Γ f5 (vRecClos mfix idx Γ')), @@ -199,16 +202,14 @@ Section Wcbv. (na', tLambda (nNamed na) fn)) (e1 : eval - (add na av - (add_multiple - (map fst mfix) - (fix_env mfix Γ') Γ')) fn + (add na av (add_multiple + (List.rev (map fst mfix)) + (fix_env mfix Γ') (Γ'))) fn res), P - (add na av - (add_multiple - (map fst mfix) - (fix_env mfix Γ') Γ')) fn + (add na av (add_multiple + (List.rev (map fst mfix)) + (fix_env mfix Γ') (Γ'))) fn res e1 → forall e2 : eval Γ a av, @@ -216,7 +217,7 @@ Section Wcbv. P Γ (tApp f5 a) res (eval_fix_unfold Γ f5 mfix - idx a na na' av fn res Γ' e + idx a na na' av fn res Γ' Hbodies Hfresh e n e0 e1 e2)) (f6 : ∀ (Γ : environment) (mfix : list (def term)) (idx : nat) (nms : list ident) (Hlen : (idx < #|mfix|)) (Hbodies : List.forallb (isLambda ∘ dbody) mfix) (n : NoDup nms) (f6 : Forall2 (λ (d : def term) (n0 : ident), nNamed n0 = dname d) mfix nms), @@ -240,7 +241,7 @@ Section Wcbv. (eval_delta Γ c decl body isdecl res e e0)) (f8 : ∀ (Γ : environment) (ind : inductive) (c : nat) (mdecl : mutual_inductive_body) (idecl : one_inductive_body) (cdecl : constructor_body) - (args : list term) (args' : list value) (e : lookup_constructor Σ ind c = Some (mdecl, idecl, cdecl)) (l : #|args| ≤ cstr_arity mdecl cdecl) + (args : list term) (args' : list value) (e : lookup_constructor Σ ind c = Some (mdecl, idecl, cdecl)) (l : #|args| ≤ cstr_nargs cdecl) (a : All2_Set (eval Γ) args args') (IHa : All2_over a (P Γ)), P Γ (tConstruct ind c args) (vConstruct ind c args') (eval_construct_block Γ ind c mdecl idecl cdecl args args' e l a)) (f9 : ∀ (Γ : environment) (ind : inductive) (c : nat) (mdecl : mutual_inductive_body) (idecl : one_inductive_body) (cdecl : constructor_body) (e : lookup_constructor Σ ind c = Some (mdecl, idecl, cdecl)), @@ -254,8 +255,8 @@ Section Wcbv. | @eval_zeta _ na b0 b0' b1 res e0 e1 => f3 Γ na b0 b0' b1 res e0 (F Γ b0 b0' e0) e1 (F (add na b0' Γ) b1 res e1) | @eval_iota_block _ ind cdecl discr c args brs br res nms e0 e1 e2 e3 e4 f10 e5 n => f4 Γ ind cdecl discr c args brs br res nms e0 (F Γ discr (vConstruct ind c args) e0) e1 e2 e3 e4 f10 e5 (F (add_multiple (List.rev nms) args Γ) br.2 res e5) n - | @eval_fix_unfold _ f10 mfix idx a na na' av fn res Γ' e0 n e1 e2 e3 => - f5 Γ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ + | @eval_fix_unfold _ f10 mfix idx a na na' av fn res Γ' Hbodies Hfresh e0 n e1 e2 e3 => + f5 Γ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ | @eval_fix _ mfix idx nms Hlen Hbodies n f10 => f6 Γ mfix idx nms Hlen Hbodies n f10 | @eval_delta _ c decl body isdecl res e0 e1 => f7 Γ c decl body isdecl res e0 e1 (F [] body res e1) | @eval_construct_block _ ind c mdecl idecl cdecl args args' e0 l a => f8 Γ ind c mdecl idecl cdecl args args' e0 l a _ @@ -290,14 +291,14 @@ Inductive represents : list ident -> environment -> term -> term -> Set := List.forallb (isLambda ∘ dbody) mfix -> NoDup nms -> All2_Set (fun d n => d.(dname) = nNamed n) mfix nms -> - All2_Set (fun d d' => (nms ++ Γ) ;;; E ⊩ d.(dbody) ~ d'.(dbody)) mfix mfix' -> + All2_Set (fun d d' => (List.rev nms ++ Γ) ;;; E ⊩ d.(dbody) ~ d'.(dbody)) mfix mfix' -> Γ ;;; E ⊩ tFix mfix idx ~ tFix mfix' idx with represents_value : value -> term -> Set := | represents_value_tBox : represents_value vBox tBox | represents_value_tClos na E s t na' : [na] ;;; E ⊩ s ~ t -> represents_value (vClos na s E) (tLambda na' t) | represents_value_tConstruct vs ts ind c : All2_Set represents_value vs ts -> represents_value (vConstruct ind c vs) (tConstruct ind c ts) | represents_value_tFix vfix i E mfix : - All2_Set (fun v d => isLambda (snd v) × (map fst vfix ;;; E ⊩ snd v ~ d.(dbody))) vfix mfix -> represents_value (vRecClos vfix i E) (tFix mfix i) + All2_Set (fun v d => isLambda (snd v) × (List.rev (map fst vfix) ;;; E ⊩ snd v ~ d.(dbody))) vfix mfix -> represents_value (vRecClos vfix i E) (tFix mfix i) where "Γ ;;; E ⊩ s ~ t" := (represents Γ E s t). Program Definition represents_ind := @@ -362,9 +363,9 @@ Program Definition represents_ind := (λ (d : def term) (n : ident), dname d = nNamed n) mfix nms) (a0 : All2_Set - (λ d d' : def term, (nms ++ Γ);;; E ⊩ dbody d ~ dbody d') + (λ d d' : def term, (List.rev nms ++ Γ);;; E ⊩ dbody d ~ dbody d') mfix mfix') - (IH : All2_over a0 (fun t t' : def term => P (nms ++ Γ) E (dbody t) (dbody t'))), + (IH : All2_over a0 (fun t t' : def term => P (List.rev nms ++ Γ) E (dbody t) (dbody t'))), P Γ E (tFix mfix idx) (tFix mfix' idx) (represents_tFix Γ E mfix mfix' idx nms Hbodies Hnodup a a0)) (f9 : P0 vBox tBox represents_value_tBox) (f10 : @@ -386,8 +387,8 @@ Program Definition represents_ind := (represents_value_tConstruct vs ts ind c a)) (f12 : ∀ (vfix : list (ident × term)) (i : nat) (E : list (ident × value)) (mfix : list (def term)) - (a : All2_Set (λ (v : ident × term) (d : def term), isLambda v.2 × map fst vfix ;;; E ⊩ v.2 ~ dbody d) vfix mfix) - (IH : All2_over a (fun v d H => P (map fst vfix) E v.2 (dbody d) (snd H) ) ), + (a : All2_Set (λ (v : ident × term) (d : def term), isLambda v.2 × List.rev (map fst vfix) ;;; E ⊩ v.2 ~ dbody d) vfix mfix) + (IH : All2_over a (fun v d H => P (List.rev (map fst vfix)) E v.2 (dbody d) (snd H) ) ), P0 (vRecClos vfix i E) (tFix mfix i) (represents_value_tFix vfix i E mfix a)), fix F @@ -506,9 +507,9 @@ Program Definition represents_value_ind := (λ (d : def term) (n : ident), dname d = nNamed n) mfix nms) (a0 : All2_Set - (λ d d' : def term, (nms ++ Γ);;; E ⊩ dbody d ~ dbody d') + (λ d d' : def term, (List.rev nms ++ Γ);;; E ⊩ dbody d ~ dbody d') mfix mfix') - (IH : All2_over a0 (fun t t' : def term => P (nms ++ Γ) E (dbody t) (dbody t'))), + (IH : All2_over a0 (fun t t' : def term => P (List.rev nms ++ Γ) E (dbody t) (dbody t'))), P Γ E (tFix mfix idx) (tFix mfix' idx) (represents_tFix Γ E mfix mfix' idx nms Hbodies Hnodup a a0)) (f9 : P0 vBox tBox represents_value_tBox) (f10 : @@ -532,8 +533,8 @@ Program Definition represents_value_ind := (E : list (ident × value)) (mfix : list (def term)) (a : All2_Set (λ (v : ident × term) (d : def term), isLambda v.2 × - (map fst vfix) ;;; E ⊩ v.2 ~ d.(dbody)) vfix mfix) - (IH : All2_over a (fun v d H => P (map fst vfix) E v.2 (dbody d) (snd H) ) ), + (List.rev (map fst vfix)) ;;; E ⊩ v.2 ~ d.(dbody)) vfix mfix) + (IH : All2_over a (fun v d H => P (List.rev (map fst vfix)) E v.2 (dbody d) (snd H) ) ), P0 (vRecClos vfix i E) (tFix mfix i) (represents_value_tFix vfix i E mfix a)), fix F @@ -652,9 +653,9 @@ Definition rep_ind := (λ (d : def term) (n : ident), dname d = nNamed n) mfix nms) (a0 : All2_Set - (λ d d' : def term, (nms ++ Γ);;; E ⊩ dbody d ~ dbody d') + (λ d d' : def term, (List.rev nms ++ Γ);;; E ⊩ dbody d ~ dbody d') mfix mfix') - (IH : All2_over a0 (fun t t' : def term => P (nms ++ Γ) E (dbody t) (dbody t'))), + (IH : All2_over a0 (fun t t' : def term => P (List.rev nms ++ Γ) E (dbody t) (dbody t'))), P Γ E (tFix mfix idx) (tFix mfix' idx) (represents_tFix Γ E mfix mfix' idx nms Hbodies Hnodup a a0)) (f9 : P0 vBox tBox represents_value_tBox) (f10 : @@ -678,10 +679,10 @@ Definition rep_ind := (E : list (ident × value)) (mfix : list (def term)) (a : All2_Set (λ (v : ident × term) (d : def term), isLambda v.2 × - (map fst vfix) ;;; + (List.rev (map fst vfix)) ;;; E ⊩ v.2 ~ dbody d) vfix mfix) - (IH : All2_over a (fun v d H => P (map fst vfix) E v.2 (dbody d) (snd H) ) ), + (IH : All2_over a (fun v d H => P (List.rev (map fst vfix)) E v.2 (dbody d) (snd H) ) ), P0 (vRecClos vfix i E) (tFix mfix i) (represents_value_tFix vfix i E mfix a)), (represents_ind P P0 f f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12, @@ -691,7 +692,7 @@ Local Notation "'⊩' v ~ s" := (represents_value v s) (at level 50). Local Hint Constructors represents : core. Local Hint Constructors represents_value : core. -From MetaCoq.Utils Require Import bytestring MCString. +From MetaCoq Require Import bytestring MCString. Require Import BinaryString. Import String. @@ -875,7 +876,7 @@ Fixpoint annotate (s : list ident) (u : term) {struct u} : term := | tProj p c => tProj p (annotate s c) | tFix mfix idx => let nms := gen_many_fresh s (map dname mfix) in - let mfix' := map2 (fun d na => map_def_name _ (fun _ => nNamed na) (annotate (gen_many_fresh s (map dname mfix) ++ s)) d) mfix nms in + let mfix' := map2 (fun d na => map_def_name _ (fun _ => nNamed na) (annotate (List.rev (gen_many_fresh s ((map dname mfix))) ++ s)) d) mfix nms in tFix mfix' idx | _ => u end. @@ -951,7 +952,7 @@ Proof. * eapply p. rewrite app_length gen_many_fresh_length. eapply p. * eapply NoDup_gen_many_fresh. - eapply represents_tFix with (nms := gen_many_fresh Γ (map dname m)). - 1:{ solve_all. generalize (gen_many_fresh Γ (map dname m) ++ Γ). clear - H. + 1:{ solve_all. generalize (List.rev (gen_many_fresh Γ (map dname m)) ++ Γ). clear - H. induction H in Γ; cbn. econstructor. intros. destruct x, dname; cbn. all: econstructor. - cbn in *. destruct p, dbody; cbn in *; try congruence. - eapply IHAll. @@ -959,21 +960,21 @@ Proof. - eapply IHAll. } 1:{ eapply NoDup_gen_many_fresh. } - { clear. generalize (((gen_many_fresh Γ (map dname m) ++ Γ))). - generalize Γ. induction m; cbn. + { clear. generalize ((List.rev (gen_many_fresh Γ ( (map dname m))) ++ Γ)). + induction m in Γ |- *; cbn. - econstructor. - - intros. destruct a; cbn. destruct dname; cbn; econstructor; eauto. + - intros. destruct a; cbn. destruct dname; cbn; try econstructor; eauto. } { solve_all. unfold wf_fix in *. rtoProp. solve_all. clear H0. unfold test_def in *. cbn in *. eapply All_impl in H1. 2:{ intros ? [[] ]. - specialize (r (gen_many_fresh Γ (map dname m) ++ Γ)). - revert r. rewrite app_length gen_many_fresh_length map_length. intros r. eapply r in i0. exact i0. + specialize (r (List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)). + revert r. rewrite ?List.rev_length app_length ?List.rev_length gen_many_fresh_length ?List.rev_length map_length. intros r. eapply r in i0. exact i0. } revert H1. - generalize (((gen_many_fresh Γ (map dname m) ++ Γ))). + generalize ((List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)). intros. induction H1 in Γ |- *. - econstructor. - - cbn. destruct x; cbn. destruct dname; cbn; econstructor; eauto. + - cbn. destruct x; cbn. destruct dname; cbn; econstructor; eauto. } Qed. @@ -995,11 +996,11 @@ Proof. 2:{ eapply All2_length in a. clear IH. eapply All2_Set_All2, All2_length in a0. len. } clear Hbodies a. induction a0; cbn in *; econstructor. - rewrite app_length in IH. eapply IH. + rewrite app_length List.rev_length in IH. eapply IH. eapply IHa0. eapply IH. - solve_all. induction a; cbn in *; rtoProp; eauto. econstructor. cbn in *. eapply IH. eapply IHa. eapply IH. - - rewrite map_length in IH. + - rewrite List.rev_length map_length in IH. assert (Hlen : #|vfix| = #|mfix|). { clear IH. eapply All2_Set_All2, All2_length in a. lia. } rewrite Hlen in IH. revert IH. generalize (#|mfix|). induction a; intros n H; cbn in *; rtoProp; split. @@ -1085,9 +1086,9 @@ Proof. + inversion a0; subst. cbn. eauto. + cbn. depelim a0. cbn. invs IH. econstructor. * cbn. - specialize (H (nms ++ Γ1) Γ2). + specialize (H (List.rev nms ++ Γ1) Γ2). rewrite app_length in H. - rewrite <- !app_assoc in *. eapply H; eauto. + rewrite <- !app_assoc in *. rewrite List.rev_length in H. eapply H; eauto. * eapply IHmfix; eauto. Qed. @@ -1316,6 +1317,40 @@ Proof. - eapply IHn; eauto. Qed. +Lemma fresh_subset Γ1 Γ2 i : + fresh i Γ1 -> incl Γ2 Γ1 -> fresh i Γ2. +Proof. + unfold fresh. repeat destruct in_dec; eauto; cbn. + intros. firstorder. +Qed. + +Local Hint Resolve fresh_subset : core. + +Lemma sunny_subset Γ1 Γ2 t : + sunny Γ1 t -> incl Γ2 Γ1 -> sunny Γ2 t. +Proof. + intros H Hincl. + induction t using EInduction.term_forall_list_ind in Γ1, Γ2, H, Hincl |- * ; cbn in *; eauto. + - cbn in H. solve_all. + - destruct n; eauto. cbn in *. rtoProp. split. 1: eapply fresh_subset; eauto. + eapply IHt; eauto. intros ? []; cbn; eauto. + - destruct n; eauto. cbn in *. rtoProp. repeat split; eauto. + eapply IHt2; eauto. eapply incl_cons; cbn; eauto. + eapply incl_tl; eauto. + - rtoProp; split; eauto. + - solve_all. + - rtoProp; split; solve_all; rtoProp; eauto; repeat split; solve_all. + + destruct x0; eauto. + + eapply a0; eauto. + intros ? ?. rewrite -> in_app_iff in *. + destruct H3; eauto. + - rtoProp; repeat split; solve_all. destruct x; cbn in *. + { destruct dname; cbn in *; eauto. } + eapply a; eauto. + intros ? ?. rewrite -> in_app_iff in *. + destruct H; eauto. +Qed. + Lemma eval_wf Σ E s v : Forall (fun d => match d.2 with ConstantDecl (Build_constant_body (Some d)) => sunny [] d | _ => true end) Σ -> All (fun x => wf (snd x)) E -> @@ -1357,12 +1392,13 @@ Proof. eapply X0. + econstructor; cbn; eauto. eapply wf_add_multiple; eauto. - now rewrite map_length fix_env_length. + now rewrite List.rev_length map_length fix_env_length. eapply wf_fix_env; eauto. + let X2 := multimatch goal with H : All _ _ |- _ => H end in eapply All_nth_error in X2; eauto; cbn in X2; rtoProp; - rewrite map_fst_add_multiple; first [ now rewrite map_length fix_env_length - | eauto ]. + rewrite map_fst_add_multiple; first [ now rewrite List.rev_length map_length fix_env_length | eauto ]. + eapply sunny_subset. eauto. + intros ?. cbn. rewrite !in_app_iff. now rewrite <- in_rev. - assert (map fst (MCList.map2 (λ (n : ident) (d0 : def term), (n, EAst.dbody d0)) nms mfix) = nms) as EE. { clear - f6. induction f6; cbn; f_equal; eauto. } econstructor. @@ -1425,7 +1461,7 @@ Qed. Lemma eval_construct_All2 Σ E ind c args vs mdecl idecl cdecl : lookup_constructor Σ ind c = Some (mdecl, idecl, cdecl) -> - #|args| <= cstr_arity mdecl cdecl -> + #|args| <= cstr_nargs cdecl -> All2 (eval Σ E) args vs -> eval Σ E (tConstruct ind c args) (vConstruct ind c vs). Proof. intros Hind. @@ -1438,7 +1474,7 @@ Proof. + eapply All2_All2_Set, All2_app. eapply H1; eauto. econstructor; eauto. Qed. -From MetaCoq.Erasure Require Import EWcbvEvalCstrsAsBlocksFixLambdaInd. +From MetaCoq Require Import EWcbvEvalCstrsAsBlocksFixLambdaInd. Lemma lookup_in_env Σ Σ' ind i : All2 (fun d d' => d.1 = d'.1 × match d.2 with ConstantDecl (Build_constant_body (Some body)) => @@ -1478,40 +1514,6 @@ Proof. Qed. *) -Lemma fresh_subset Γ1 Γ2 i : - fresh i Γ1 -> incl Γ2 Γ1 -> fresh i Γ2. -Proof. - unfold fresh. repeat destruct in_dec; eauto; cbn. - intros. firstorder. -Qed. - -Local Hint Resolve fresh_subset : core. - -Lemma sunny_subset Γ1 Γ2 t : - sunny Γ1 t -> incl Γ2 Γ1 -> sunny Γ2 t. -Proof. - intros H Hincl. - induction t using EInduction.term_forall_list_ind in Γ1, Γ2, H, Hincl |- * ; cbn in *; eauto. - - cbn in H. solve_all. - - destruct n; eauto. cbn in *. rtoProp. split. 1: eapply fresh_subset; eauto. - eapply IHt; eauto. intros ? []; cbn; eauto. - - destruct n; eauto. cbn in *. rtoProp. repeat split; eauto. - eapply IHt2; eauto. eapply incl_cons; cbn; eauto. - eapply incl_tl; eauto. - - rtoProp; split; eauto. - - solve_all. - - rtoProp; split; solve_all; rtoProp; eauto; repeat split; solve_all. - + destruct x0; eauto. - + eapply a0; eauto. - intros ? ?. rewrite -> in_app_iff in *. - destruct H3; eauto. - - rtoProp; repeat split; solve_all. destruct x; cbn in *. - { destruct dname; cbn in *; eauto. } - eapply a; eauto. - intros ? ?. rewrite -> in_app_iff in *. - destruct H; eauto. -Qed. - Lemma represents_add_fresh nms Γ E vs s t : Γ ;;; E ⊩ s ~ t -> sunny nms s -> @@ -1590,8 +1592,7 @@ Proof. edestruct s1 as (? & ? & ?); eauto. invs Hv1. eexists; split; eauto. econstructor; eauto. - invs Hrep. - + let H3 := match goal with H : ⊩ _ ~ tApp _ _ |- _ => H end in - invs H3. + + invs H3. + cbn in Hsunny. rtoProp. edestruct s0 as (v1 & Hv1 & Hv2). 3: eauto. eauto. eauto. invs Hv1. @@ -1615,7 +1616,6 @@ Proof. - assert (pars = 0) as ->. { unfold constructor_isprop_pars_decl in *. destruct lookup_constructor as [[[[] [? ?]] ?] | ] eqn:EE; cbn in *; try congruence. - let H0 := match goal with H : Some _ = Some _ |- _ => H end in invs H0. destruct lookup_env eqn:EEE; try congruence. eapply lookup_env_wellformed in EEE; eauto. @@ -1634,8 +1634,7 @@ Proof. eapply All2_Set_All2 in H14. eapply All2_nth_error_Some_right in H14 as He2. 2: eauto. destruct He2 as (br' & Hnth & nms & Hbrs & Hbr & Hnodup). invs Hv1. edestruct s1 as (v2 & Hv1_ & Hv2_). - 1: { let H6 := match goal with H : is_true (forallb _ _) |- _ => H end in - eapply forallb_nth_error in H6; setoid_rewrite Hnth in H6; cbn in H6. rtoProp. + 1: { eapply forallb_nth_error in H6. setoid_rewrite Hnth in H6. cbn in H6. rtoProp. assert (nms = flat_map (λ x : name, match x with | nAnon => [] | nNamed na => [na] @@ -1644,9 +1643,6 @@ Proof. { rewrite Heqnms flat_map_concat_map. intros ? (? & ([] & <- & ?) % in_map_iff & Hd) % in_concat; cbn in *; eauto. destruct Hd; subst; eauto. - rename H6 into H6_old; - let H6' := match goal with H : is_true (forallb _ _) |- _ => H end in - rename H6' into H6. eapply forallb_forall in H6; eauto. cbn in H6. unfold fresh in H6. destruct in_dec; cbn in *; congruence. } { subst. unfold dupfree in H9. destruct dupfree_dec_ident; cbn in *; congruence. } @@ -1664,10 +1660,7 @@ Proof. eapply All2_Set_All2, All2_length in H10; eauto. eapply All2_Set_All2, All2_length in Hbrs; eauto. rewrite -> !skipn_length in *. lia. - -- rename H6 into H6_old; - let H6' := match goal with H : is_true (forallb _ _) |- _ => H end in - rename H6' into H6. - eapply forallb_nth_error in H6. setoid_rewrite Hnth in H6. cbn in H6. rtoProp. + -- eapply forallb_nth_error in H6. setoid_rewrite Hnth in H6. cbn in H6. rtoProp. enough (nms = flat_map (λ x : name, match x with | nAnon => [] | nNamed na => [na] @@ -1689,7 +1682,7 @@ Proof. rewrite -> !skipn_length in *. lia. * solve_all. * now rewrite rev_involutive in Hv2_. - - eauto. (* artifact of the induction being weird and having a trivial assumption to not mess up proof script. FIXME! *) + - eapply X; eauto. (* artifact of the induction being weird and having a trivial assumption to not mess up proof script. FIXME! *) - invs Hrep. + let H5 := match goal with H : ⊩ _ ~ tApp _ _ |- _ => H end in invs H5. @@ -1716,35 +1709,35 @@ Proof. edestruct s2 as (v'' & IH1'' & IH2''). - eapply represents_substl_rev with (vs := _ :: fix_env vfix E0) (nms := na0 :: map fst vfix); eauto. (* having represents_substl_rev with nms ++ Gamma created an order problem here. changing it there fixes the problem here. but is this correct? *) + eapply represents_substl_rev with (vs := _ :: fix_env vfix E0) (nms := na0 :: List.rev (map fst vfix)); eauto. (* having represents_substl_rev with nms ++ Gamma created an order problem here. changing it there fixes the problem here. but is this correct? *) 8:{ eexists. split. eauto. eapply eval_fix_unfold; eauto. - (* eapply eval_wf in IH2; eauto. + solve_all. + eapply eval_wf in IH2; eauto. invs IH2. - Lemma NoDup_app_inv {A} (l1 l2 : list A) : - NoDup (l1 ++ l2) -> NoDup l1 /\ NoDup l2. - Proof. - induction l1; cbn; eauto using NoDup. - inversion 1; subst; firstorder. - rewrite in_app_iff in H2. econstructor; eauto. - Qed. eapply NoDup_app_inv in H9 as []; eauto. *) + eapply All_nth_error in Hb1. 2: exact X. cbn in *. + rtoProp. + intros Hna. + unfold fresh in H2. + destruct in_dec; cbn in *; try congruence. + eapply n. rewrite in_app_iff. eauto. } { intros ? [-> | ?]. - eapply All_nth_error in sunny_in_vfix; eauto. cbn in sunny_in_vfix. rtoProp. unfold fresh in *. destruct in_dec; cbn in *; eauto. rewrite in_app_iff in n. eauto. - - eapply distinct_vfix_E0; eauto. + - eapply distinct_vfix_E0; eauto. now eapply in_rev. } { econstructor. - eapply All_nth_error in sunny_in_vfix; eauto. cbn in *. rtoProp. unfold fresh in *. destruct in_dec; cbn in *; eauto. - rewrite in_app_iff in n. eauto. - - eauto. + rewrite in_app_iff in n. rewrite <- in_rev. eauto. + - eapply NoDup_rev. eauto. } - { cbn. now rewrite map_length fix_env_length. } + { cbn. now rewrite List.rev_length map_length fix_env_length. } econstructor; eauto. { clear - H13. unfold fix_env, fix_subst. eapply All2_length in H13 as Hlen. rewrite Hlen. clear Hlen. @@ -1753,15 +1746,17 @@ Proof. } now rewrite app_nil_r. { cbn. rewrite map_fst_add_multiple. - now rewrite map_length fix_env_length. + now rewrite List.rev_length map_length fix_env_length. eapply All_nth_error in sunny_in_vfix; eauto. cbn in sunny_in_vfix. - now rtoProp. + rtoProp. + eapply sunny_subset; eauto. + intros ?; cbn in *. rewrite !in_app_iff. rewrite <- !in_rev. eauto. } { econstructor. - cbn. eapply eval_wf. 4: eauto. all:eauto. - eapply wf_add_multiple. + eauto. - + now rewrite map_length fix_env_length. + + now rewrite List.rev_length map_length fix_env_length. + eapply eval_wf in IH2; eauto. eapply wf_fix_env; eauto. } @@ -1801,8 +1796,7 @@ Proof. edestruct s0 as (v & Hv1 & Hv2). 1: eauto. eauto. econstructor. eexists. split. eauto. econstructor; eauto. - invs Hrep. - + let H2 := multimatch goal with H : _ |- _ => H end in - now invs H2. + + invs H2. + cbn in Hsunny. rtoProp. eapply eval_to_value in ev as Hval. invs Hval. * destruct f'; cbn in *; try congruence. @@ -1850,8 +1844,14 @@ Proof. eexists. split. econstructor. { instantiate (1 := vs). clear - Hvs; induction Hvs; econstructor; eauto. eapply r. } econstructor. erewrite <- lookup_in_env. 2: solve_all. - eapply H. eapply All2_length in H5. lia. - clear - Hvs; induction Hvs; econstructor; eauto. eapply r. + eapply H. eapply All2_length in H5. + destruct lookup_env as [ [] | ] eqn:Elo; try congruence. + epose proof (@lookup_env_wellformed _ Σ (inductive_mind ind) _ HwfΣ Elo). + cbn in H1. unfold wf_minductive in H1. rtoProp. cbn in *. eapply Nat.eqb_eq in H1. unfold cstr_arity in *. + destruct nth_error eqn:E1; cbn in *; try congruence. + destruct (nth_error (ind_ctors o) i) eqn:E2; cbn in *; try congruence. + unfold cstr_arity in *. invs H. lia. + clear - Hvs; induction Hvs; econstructor; eauto. eapply r. - invs Hrep; cbn in *; try congruence; rtoProp. + econstructor. split; eauto. econstructor. + destruct args'; congruence. @@ -1866,7 +1866,7 @@ Proof. symmetry. eapply All2_length; eauto. intros. cbn in *. destruct H1 as (([? []] & ?) & ?). - rewrite app_nil_r in r. eauto. + rewrite app_nil_r in r. all: eauto. - Unshelve. all: repeat econstructor. + Unshelve. all: repeat econstructor. Qed. From 713dbde5c9a3702acd64e72fa1742db3a3e4d161 Mon Sep 17 00:00:00 2001 From: "dependabot[bot]" <49699333+dependabot[bot]@users.noreply.github.com> Date: Mon, 19 Jun 2023 07:06:06 -0700 Subject: [PATCH 14/61] Bump cachix/install-nix-action from 20 to 21 (#966) Bumps [cachix/install-nix-action](https://github.com/cachix/install-nix-action) from 20 to 21. - [Release notes](https://github.com/cachix/install-nix-action/releases) - [Commits](https://github.com/cachix/install-nix-action/compare/v20...v21) --- updated-dependencies: - dependency-name: cachix/install-nix-action dependency-type: direct:production update-type: version-update:semver-major ... Signed-off-by: dependabot[bot] Co-authored-by: dependabot[bot] <49699333+dependabot[bot]@users.noreply.github.com> --- .github/workflows/nix-action-coq-8.16-macos.yml | 6 +++--- .github/workflows/nix-action-coq-8.16-ubuntu.yml | 6 +++--- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/.github/workflows/nix-action-coq-8.16-macos.yml b/.github/workflows/nix-action-coq-8.16-macos.yml index 1e89ebab4..73b8a4969 100644 --- a/.github/workflows/nix-action-coq-8.16-macos.yml +++ b/.github/workflows/nix-action-coq-8.16-macos.yml @@ -19,7 +19,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v20 + uses: cachix/install-nix-action@v21 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -59,7 +59,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v20 + uses: cachix/install-nix-action@v21 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -104,7 +104,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v20 + uses: cachix/install-nix-action@v21 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq diff --git a/.github/workflows/nix-action-coq-8.16-ubuntu.yml b/.github/workflows/nix-action-coq-8.16-ubuntu.yml index c9705f2ba..62ef6ba7c 100644 --- a/.github/workflows/nix-action-coq-8.16-ubuntu.yml +++ b/.github/workflows/nix-action-coq-8.16-ubuntu.yml @@ -19,7 +19,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v20 + uses: cachix/install-nix-action@v21 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -59,7 +59,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v20 + uses: cachix/install-nix-action@v21 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -104,7 +104,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v20 + uses: cachix/install-nix-action@v21 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq From 6899a1a4b4db72ff5f28e5114f745e35c848d11e Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 28 Jun 2023 14:06:25 +0200 Subject: [PATCH 15/61] Fix a sneaky Prop <= Type use messing up certified erasure --- Makefile | 2 ++ erasure/theories/ErasureFunction.v | 4 ++-- template-coq/update_plugin.sh | 4 ++-- 3 files changed, 6 insertions(+), 4 deletions(-) diff --git a/Makefile b/Makefile index 1d5d99253..a99854e73 100644 --- a/Makefile +++ b/Makefile @@ -36,6 +36,7 @@ install: all translations $(MAKE) -C quotation install $(MAKE) -C safechecker-plugin install $(MAKE) -C erasure install + $(MAKE) -C erasure-plugin install $(MAKE) -C translations install uninstall: @@ -48,6 +49,7 @@ uninstall: $(MAKE) -C quotation uninstall $(MAKE) -C safechecker-plugin uninstall $(MAKE) -C erasure uninstall + $(MAKE) -C erasure-plugin uninstall $(MAKE) -C translations uninstall html: all diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index fba458ddf..8fc9429cf 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -895,7 +895,7 @@ Program Fixpoint erase_global_decls let X' := abstract_pop_decls X in if KernameSet.mem kn deps then let Xext := abstract_make_wf_env_ext X' (cst_universes cb) _ in - let normalization_in' : id _ := _ in (* hack to avoid program erroring out on unresolved tc *) + let normalization_in' : @id Prop _ := _ in (* hack to avoid program erroring out on unresolved tc *) let cb' := @erase_constant_body X_type Xext normalization_in' cb _ in let deps := KernameSet.union deps (snd cb') in let X'' := erase_global_decls deps X' decls _ in @@ -2694,7 +2694,7 @@ Program Fixpoint erase_global_decls_fast (deps : KernameSet.t) | (kn, ConstantDecl cb) :: decls => if KernameSet.mem kn deps then let Xext' := abstract_make_wf_env_ext X (cst_universes cb) _ in - let normalization_in' : id _ := _ in (* hack to avoid program erroring out on unresolved tc *) + let normalization_in' : @id Prop _ := _ in (* hack to avoid program erroring out on unresolved tc *) let cb' := @erase_constant_body X_type Xext' normalization_in' cb _ in let deps := KernameSet.union deps (snd cb') in let Σ' := erase_global_decls_fast deps X_type X decls _ in diff --git a/template-coq/update_plugin.sh b/template-coq/update_plugin.sh index 097a2e93b..1842cfdeb 100755 --- a/template-coq/update_plugin.sh +++ b/template-coq/update_plugin.sh @@ -19,8 +19,8 @@ then mv -f tmp $f done # Fix an extraction bug: wrong type annotation on eq_equivalence - patch -N -p0 < extraction.patch || exit $? - patch -N -p0 < specFloat.patch || exit $? + patch -N -p0 < extraction.patch + patch -N -p0 < specFloat.patch exit 0 else echo "Extracted code is up-to-date" From c065ed6faedd1e226af967631bc2154db0bc99ae Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 2 Jun 2023 17:06:10 +0200 Subject: [PATCH 16/61] Update Makefile to install erasure plugin --- .gitignore | 2 ++ Makefile | 2 ++ 2 files changed, 4 insertions(+) diff --git a/.gitignore b/.gitignore index 704758b38..0558c7fc7 100644 --- a/.gitignore +++ b/.gitignore @@ -404,3 +404,5 @@ erasure-plugin/Makefile.erasureplugin erasure-plugin/Makefile.erasureplugin.conf erasure-plugin/Makefile.plugin erasure-plugin/Makefile.plugin.conf +erasure-plugin/META +safechecker-plugin/META diff --git a/Makefile b/Makefile index 1d5d99253..a99854e73 100644 --- a/Makefile +++ b/Makefile @@ -36,6 +36,7 @@ install: all translations $(MAKE) -C quotation install $(MAKE) -C safechecker-plugin install $(MAKE) -C erasure install + $(MAKE) -C erasure-plugin install $(MAKE) -C translations install uninstall: @@ -48,6 +49,7 @@ uninstall: $(MAKE) -C quotation uninstall $(MAKE) -C safechecker-plugin uninstall $(MAKE) -C erasure uninstall + $(MAKE) -C erasure-plugin uninstall $(MAKE) -C translations uninstall html: all From e0794e31790fa719d4a9bf34a0f508d17e43da95 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 28 Jun 2023 14:14:44 +0200 Subject: [PATCH 17/61] Fix sneaky, useless Prop <= Type use that was breaking certified erasure --- erasure/theories/ErasureFunction.v | 5 +++-- template-coq/Makefile.plugin.local | 2 +- template-coq/update_plugin.sh | 4 ++-- 3 files changed, 6 insertions(+), 5 deletions(-) diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 421898b47..b1d033023 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -895,7 +895,7 @@ Program Fixpoint erase_global_decls let X' := abstract_pop_decls X in if KernameSet.mem kn deps then let Xext := abstract_make_wf_env_ext X' (cst_universes cb) _ in - let normalization_in' : id _ := _ in (* hack to avoid program erroring out on unresolved tc *) + let normalization_in' : @id Prop _ := _ in (* hack to avoid program erroring out on unresolved tc *) let cb' := @erase_constant_body X_type Xext normalization_in' cb _ in let deps := KernameSet.union deps (snd cb') in let X'' := erase_global_decls deps X' decls _ in @@ -2637,6 +2637,7 @@ Notation NormalizationIn_erase_global_decls_fast X decls end) (only parsing). + Section EraseGlobalFast. Import PCUICEnvironment. @@ -2694,7 +2695,7 @@ Program Fixpoint erase_global_decls_fast (deps : KernameSet.t) | (kn, ConstantDecl cb) :: decls => if KernameSet.mem kn deps then let Xext' := abstract_make_wf_env_ext X (cst_universes cb) _ in - let normalization_in' : id _ := _ in (* hack to avoid program erroring out on unresolved tc *) + let normalization_in' : @id Prop _ := _ in (* hack to avoid program erroring out on unresolved tc *) let cb' := @erase_constant_body X_type Xext' normalization_in' cb _ in let deps := KernameSet.union deps (snd cb') in let Σ' := erase_global_decls_fast deps X_type X decls _ in diff --git a/template-coq/Makefile.plugin.local b/template-coq/Makefile.plugin.local index 9700c197b..3d1b9c40a 100644 --- a/template-coq/Makefile.plugin.local +++ b/template-coq/Makefile.plugin.local @@ -5,7 +5,7 @@ CAMLFLAGS+=-w -39 # Unused rec flag CAMLFLAGS+=-w -26 # Unused variable CAMLFLAGS+=-w -34 # Unused type CAMLFLAGS+=-w -60 # Unused module -CAMLFLAGS+=-w -8 # Non-exhaustive pattern-matchings (BEWARE, just for extracted code) +# CAMLFLAGS+=-w -8 # Non-exhaustive pattern-matchings (BEWARE, just for extracted code) CAMLFLAGS+=-bin-annot # For merlin CAMLPKGS+=-package stdlib-shims INSTALLDEFAULTROOT=MetaCoq/Template diff --git a/template-coq/update_plugin.sh b/template-coq/update_plugin.sh index 097a2e93b..1842cfdeb 100755 --- a/template-coq/update_plugin.sh +++ b/template-coq/update_plugin.sh @@ -19,8 +19,8 @@ then mv -f tmp $f done # Fix an extraction bug: wrong type annotation on eq_equivalence - patch -N -p0 < extraction.patch || exit $? - patch -N -p0 < specFloat.patch || exit $? + patch -N -p0 < extraction.patch + patch -N -p0 < specFloat.patch exit 0 else echo "Extracted code is up-to-date" From f5967d8f2e3043779b6d0ee67edf720b88279850 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Wed, 5 Jul 2023 06:22:39 -0700 Subject: [PATCH 18/61] Update dependabot.yml to also update other branches --- .github/dependabot.yml | 24 ++++++++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/.github/dependabot.yml b/.github/dependabot.yml index dcd7c9ef7..b378be58f 100644 --- a/.github/dependabot.yml +++ b/.github/dependabot.yml @@ -9,3 +9,27 @@ updates: interval: "daily" labels: - "dependencies" + - package-ecosystem: "github-actions" + directory: "/" + target-branch: "master" + schedule: + # Check for updates to GitHub Actions every weekday + interval: "daily" + labels: + - "dependencies" + - package-ecosystem: "github-actions" + directory: "/" + target-branch: "coq-8.17" + schedule: + # Check for updates to GitHub Actions every weekday + interval: "daily" + labels: + - "dependencies" + - package-ecosystem: "github-actions" + directory: "/" + target-branch: "coq-8.16" + schedule: + # Check for updates to GitHub Actions every weekday + interval: "daily" + labels: + - "dependencies" From a738c553d8a7953d5bf0d81beafadbcb34da48f5 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Wed, 5 Jul 2023 06:24:14 -0700 Subject: [PATCH 19/61] Update dependabot.yml to check main, not master --- .github/dependabot.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/dependabot.yml b/.github/dependabot.yml index b378be58f..ea49cfd2a 100644 --- a/.github/dependabot.yml +++ b/.github/dependabot.yml @@ -11,7 +11,7 @@ updates: - "dependencies" - package-ecosystem: "github-actions" directory: "/" - target-branch: "master" + target-branch: "main" schedule: # Check for updates to GitHub Actions every weekday interval: "daily" From e8d0b36017e20a8fb9b7fb32bea74a6dc8d96814 Mon Sep 17 00:00:00 2001 From: "dependabot[bot]" <49699333+dependabot[bot]@users.noreply.github.com> Date: Wed, 5 Jul 2023 06:25:05 -0700 Subject: [PATCH 20/61] Bump cachix/install-nix-action from 21 to 22 (#970) Bumps [cachix/install-nix-action](https://github.com/cachix/install-nix-action) from 21 to 22. - [Release notes](https://github.com/cachix/install-nix-action/releases) - [Commits](https://github.com/cachix/install-nix-action/compare/v21...v22) --- updated-dependencies: - dependency-name: cachix/install-nix-action dependency-type: direct:production update-type: version-update:semver-major ... Signed-off-by: dependabot[bot] Co-authored-by: dependabot[bot] <49699333+dependabot[bot]@users.noreply.github.com> --- .github/workflows/nix-action-coq-8.16-macos.yml | 6 +++--- .github/workflows/nix-action-coq-8.16-ubuntu.yml | 6 +++--- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/.github/workflows/nix-action-coq-8.16-macos.yml b/.github/workflows/nix-action-coq-8.16-macos.yml index 73b8a4969..d4910f21d 100644 --- a/.github/workflows/nix-action-coq-8.16-macos.yml +++ b/.github/workflows/nix-action-coq-8.16-macos.yml @@ -19,7 +19,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v21 + uses: cachix/install-nix-action@v22 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -59,7 +59,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v21 + uses: cachix/install-nix-action@v22 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -104,7 +104,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v21 + uses: cachix/install-nix-action@v22 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq diff --git a/.github/workflows/nix-action-coq-8.16-ubuntu.yml b/.github/workflows/nix-action-coq-8.16-ubuntu.yml index 62ef6ba7c..81a42bc26 100644 --- a/.github/workflows/nix-action-coq-8.16-ubuntu.yml +++ b/.github/workflows/nix-action-coq-8.16-ubuntu.yml @@ -19,7 +19,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v21 + uses: cachix/install-nix-action@v22 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -59,7 +59,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v21 + uses: cachix/install-nix-action@v22 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -104,7 +104,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v21 + uses: cachix/install-nix-action@v22 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq From 943cc72bd12ab41d1f67af96f5744c32e40338bf Mon Sep 17 00:00:00 2001 From: "dependabot[bot]" <49699333+dependabot[bot]@users.noreply.github.com> Date: Wed, 5 Jul 2023 06:25:41 -0700 Subject: [PATCH 21/61] Bump cachix/install-nix-action from 21 to 22 (#973) Bumps [cachix/install-nix-action](https://github.com/cachix/install-nix-action) from 21 to 22. - [Release notes](https://github.com/cachix/install-nix-action/releases) - [Commits](https://github.com/cachix/install-nix-action/compare/v21...v22) --- updated-dependencies: - dependency-name: cachix/install-nix-action dependency-type: direct:production update-type: version-update:semver-major ... Signed-off-by: dependabot[bot] Co-authored-by: dependabot[bot] <49699333+dependabot[bot]@users.noreply.github.com> From 3e0f857109da9530c572edb293a737bc5fe25943 Mon Sep 17 00:00:00 2001 From: "dependabot[bot]" <49699333+dependabot[bot]@users.noreply.github.com> Date: Mon, 4 Sep 2023 18:20:32 +0000 Subject: [PATCH 22/61] Bump actions/checkout from 3 to 4 (#978) --- .github/workflows/build.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 3b85ce6de..0c54ab5d6 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -11,7 +11,7 @@ jobs: runs-on: ubuntu-latest steps: - name: Checkout code - uses: actions/checkout@v3 + uses: actions/checkout@v4 with: fetch-depth: 1 @@ -37,7 +37,7 @@ jobs: steps: - name: Checkout code - uses: actions/checkout@v3 + uses: actions/checkout@v4 with: fetch-depth: 1 From 555b2394100846c758c17fc8b7450a4d2b728baf Mon Sep 17 00:00:00 2001 From: "dependabot[bot]" <49699333+dependabot[bot]@users.noreply.github.com> Date: Mon, 4 Sep 2023 18:21:32 +0000 Subject: [PATCH 23/61] Bump actions/checkout from 3 to 4 (#977) --- .github/workflows/build.yml | 4 ++-- .github/workflows/nix-action-coq-8.16-macos.yml | 6 +++--- .github/workflows/nix-action-coq-8.16-ubuntu.yml | 6 +++--- 3 files changed, 8 insertions(+), 8 deletions(-) diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index cb1ea1d6f..e789f6ed3 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -11,7 +11,7 @@ jobs: runs-on: ubuntu-latest steps: - name: Checkout code - uses: actions/checkout@v3 + uses: actions/checkout@v4 with: fetch-depth: 1 @@ -37,7 +37,7 @@ jobs: steps: - name: Checkout code - uses: actions/checkout@v3 + uses: actions/checkout@v4 with: fetch-depth: 1 diff --git a/.github/workflows/nix-action-coq-8.16-macos.yml b/.github/workflows/nix-action-coq-8.16-macos.yml index d4910f21d..b1fbcb83e 100644 --- a/.github/workflows/nix-action-coq-8.16-macos.yml +++ b/.github/workflows/nix-action-coq-8.16-macos.yml @@ -14,7 +14,7 @@ jobs: \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - name: Git checkout - uses: actions/checkout@v3 + uses: actions/checkout@v4 with: fetch-depth: 0 ref: ${{ env.tested_commit }} @@ -54,7 +54,7 @@ jobs: \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - name: Git checkout - uses: actions/checkout@v3 + uses: actions/checkout@v4 with: fetch-depth: 0 ref: ${{ env.tested_commit }} @@ -99,7 +99,7 @@ jobs: \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - name: Git checkout - uses: actions/checkout@v3 + uses: actions/checkout@v4 with: fetch-depth: 0 ref: ${{ env.tested_commit }} diff --git a/.github/workflows/nix-action-coq-8.16-ubuntu.yml b/.github/workflows/nix-action-coq-8.16-ubuntu.yml index 81a42bc26..c2652234d 100644 --- a/.github/workflows/nix-action-coq-8.16-ubuntu.yml +++ b/.github/workflows/nix-action-coq-8.16-ubuntu.yml @@ -14,7 +14,7 @@ jobs: \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - name: Git checkout - uses: actions/checkout@v3 + uses: actions/checkout@v4 with: fetch-depth: 0 ref: ${{ env.tested_commit }} @@ -54,7 +54,7 @@ jobs: \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - name: Git checkout - uses: actions/checkout@v3 + uses: actions/checkout@v4 with: fetch-depth: 0 ref: ${{ env.tested_commit }} @@ -99,7 +99,7 @@ jobs: \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - name: Git checkout - uses: actions/checkout@v3 + uses: actions/checkout@v4 with: fetch-depth: 0 ref: ${{ env.tested_commit }} From 77234210ad9593d7304f9f0d453f70424b2c8f90 Mon Sep 17 00:00:00 2001 From: "dependabot[bot]" <49699333+dependabot[bot]@users.noreply.github.com> Date: Mon, 4 Sep 2023 19:54:27 +0000 Subject: [PATCH 24/61] Bump cachix/install-nix-action from 22 to 23 (#979) --- .github/workflows/nix-action-coq-8.16-macos.yml | 6 +++--- .github/workflows/nix-action-coq-8.16-ubuntu.yml | 6 +++--- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/.github/workflows/nix-action-coq-8.16-macos.yml b/.github/workflows/nix-action-coq-8.16-macos.yml index b1fbcb83e..ae2df21fe 100644 --- a/.github/workflows/nix-action-coq-8.16-macos.yml +++ b/.github/workflows/nix-action-coq-8.16-macos.yml @@ -19,7 +19,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v22 + uses: cachix/install-nix-action@v23 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -59,7 +59,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v22 + uses: cachix/install-nix-action@v23 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -104,7 +104,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v22 + uses: cachix/install-nix-action@v23 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq diff --git a/.github/workflows/nix-action-coq-8.16-ubuntu.yml b/.github/workflows/nix-action-coq-8.16-ubuntu.yml index c2652234d..3da6051c9 100644 --- a/.github/workflows/nix-action-coq-8.16-ubuntu.yml +++ b/.github/workflows/nix-action-coq-8.16-ubuntu.yml @@ -19,7 +19,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v22 + uses: cachix/install-nix-action@v23 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -59,7 +59,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v22 + uses: cachix/install-nix-action@v23 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq @@ -104,7 +104,7 @@ jobs: fetch-depth: 0 ref: ${{ env.tested_commit }} - name: Cachix install - uses: cachix/install-nix-action@v22 + uses: cachix/install-nix-action@v23 with: nix_path: nixpkgs=channel:nixpkgs-unstable - name: Cachix setup metacoq From adba7945a55e4590802318f27b0c56d6ddf48f7f Mon Sep 17 00:00:00 2001 From: nicolas tabareau Date: Tue, 26 Sep 2023 14:49:58 +0200 Subject: [PATCH 25/61] improve stengthening to get cumul info on type --- .../theories/Bidirectional/BDStrengthening.v | 73 ++++++++++++++----- pcuic/theories/PCUICFirstorder.v | 8 -- utils/theories/All_Forall.v | 5 +- 3 files changed, 56 insertions(+), 30 deletions(-) diff --git a/pcuic/theories/Bidirectional/BDStrengthening.v b/pcuic/theories/Bidirectional/BDStrengthening.v index 2a2f754c5..827f91eaf 100644 --- a/pcuic/theories/Bidirectional/BDStrengthening.v +++ b/pcuic/theories/Bidirectional/BDStrengthening.v @@ -3,7 +3,7 @@ From Coq Require String. From MetaCoq.Utils Require Import utils monad_utils. From MetaCoq.Common Require Import config. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICGlobalEnv - PCUICTactics + PCUICTactics PCUICCumulativity PCUICInduction PCUICLiftSubst PCUICUnivSubst PCUICEquality PCUICUtils PCUICPosition PCUICTyping PCUICSigmaCalculus PCUICOnFreeVars PCUICClosed PCUICConfluence PCUICSpine PCUICInductiveInversion PCUICParallelReductionConfluence PCUICWellScopedCumulativity PCUICClosed PCUICRenameDef PCUICInstConv PCUICClosedTyp PCUICWeakeningEnvTyp PCUICRenameTyp PCUICRenameConv PCUICGuardCondition PCUICWeakeningConv. @@ -896,21 +896,6 @@ Qed. End BDRenaming. -Theorem typing_renaming_cond_P `{checker_flags} {P f Σ Γ Δ t T} {wfΣ : wf Σ.1} : - renaming P Σ Γ Δ f -> - on_ctx_free_vars P Γ -> - on_free_vars P t -> - Σ ;;; Γ |- t : T -> - ∑ T', Σ ;;; Δ |- rename f t : T'. -Proof. - move => [ur wfΔ] fvΓ fvt tyt. - move: (tyt) => /typing_wf_local wfΓ. - move: (tyt) => /typing_infering [T' [inf cum]]. - exists (rename f T'). - apply infering_typing => //. - eapply bidirectional_renaming ; eassumption. -Qed. - Lemma closedn_ctx_lift_inv n k k' Γ : k <= k' -> closedn_ctx (k' + n) (lift_context n k Γ) -> closedn_ctx k' Γ. @@ -996,9 +981,29 @@ Proof. all: lia. Qed. + +Theorem typing_renaming_cond_P `{checker_flags} {P f Σ Γ Δ t T} {wfΣ : wf Σ.1} : + renaming P Σ Γ Δ f -> + on_ctx_free_vars P Γ -> + on_free_vars P t -> + on_free_vars P T -> + Σ ;;; Γ |- t : T -> + ∑ T', (Σ ;;; Δ |- rename f t : T') * (Σ ;;; Δ |- T' <= rename f T). +Proof. + move => [ur wfΔ] fvΓ fvt fvT tyt. + move: (tyt) => /typing_wf_local wfΓ. + move: (tyt) => /typing_infering [T' [inf cum]]. + exists (rename f T'). split; eauto. sq. + apply infering_typing => //. + eapply bidirectional_renaming ; eassumption. + eapply cumul_renameP with (Γ := Γ); eauto. + 2: { eapply ws_cumul_pb_forget_cumul; eauto. } + eapply infering_on_free_vars; try exact inf; eauto. +Qed. + Lemma strengthening `{cf: checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ Γ' Γ'' t T : - Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : T -> - ∑ T', Σ ;;; Γ ,,, Γ'' |- t : T'. + Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : lift #|Γ'| #|Γ''| T -> + ∑ T', (Σ ;;; Γ ,,, Γ'' |- t : T') * (Σ ;;; Γ ,,, Γ'' |- T' <= T). Proof. intros Hty. assert (wf_local Σ Γ) by @@ -1029,9 +1034,9 @@ Proof. eassumption. } - erewrite <- (lift_unlift_term t). + erewrite <- (lift_unlift_term t). erewrite <- (lift_unlift_term T). eapply typing_renaming_cond_P. - 4: eassumption. + 5: eassumption. - split => //. apply urenaming_strengthen. - move: wfΓ'' => /wf_local_closed_context. @@ -1046,4 +1051,32 @@ Proof. 2: lia. eapply closedn_on_free_vars. eassumption. + - rewrite on_free_vars_lift -shiftnP_xpredT. + move: Hty => /type_closed. + len. + rewrite -[X in _ + X]Nat.add_comm Nat.add_assoc => Ht. + eapply closedn_lift_inv in Ht. + 2: lia. + eapply closedn_on_free_vars. + eassumption. Qed. + + +Lemma strengthening_type `{cf: checker_flags} {Σ : global_env_ext} {wfΣ : PCUICTyping.wf Σ} Γ Γ' Γ'' t s : + Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : tSort s -> + ∑ s', (Σ ;;; Γ ,,, Γ'' |- t : tSort s') * (compare_universe Cumul Σ s' s). +Proof. + assert (liftSort : lift #|Γ'| #|Γ''| (tSort s) = tSort s) by reflexivity. rewrite <- liftSort; clear liftSort. + intros H; eapply strengthening in H. destruct H as [T' [HT' Hcumul]]. + pose proof (Hcumul' := Hcumul). + pose proof (type_closed HT'). + eapply closedn_on_free_vars in H. + pose proof (wf_local_closed_context (typing_wf_local HT')). + eapply into_ws_cumul_pb in Hcumul; eauto. + eapply PCUICConversion.ws_cumul_pb_Sort_r_inv in Hcumul as [s' [Hred Hcumul]]. + eapply closed_red_red in Hred. pose proof (HT'':=HT'). eapply PCUICValidity.validity in HT' as [? ?]. + exists s'; split; eauto. + eapply type_Cumul; eauto. eapply PCUICSR.subject_reduction; eauto. + eapply PCUICConversion.cumulAlgo_cumulSpec. eapply PCUICConversion.ws_cumul_pb_red_r_inv; eauto. + eapply ws_cumul_pb_refl; eauto. +Qed. \ No newline at end of file diff --git a/pcuic/theories/PCUICFirstorder.v b/pcuic/theories/PCUICFirstorder.v index 47f5e6198..cf5bc91af 100644 --- a/pcuic/theories/PCUICFirstorder.v +++ b/pcuic/theories/PCUICFirstorder.v @@ -48,14 +48,6 @@ Section firstorder. | tRel i => (k <=? i) && (i false end. - (* - Definition firstorder_type (t : term) := - match (PCUICAstUtils.decompose_app t).1 with - | tInd (mkInd nm i) _ => match (plookup_env Σb nm) with - | Some l => nth i l false | None => false - end - | _ => false - end. *) Definition firstorder_con mind (c : constructor_body) := let inds := #|mind.(ind_bodies)| in diff --git a/utils/theories/All_Forall.v b/utils/theories/All_Forall.v index 848c19f3d..1e408d0aa 100644 --- a/utils/theories/All_Forall.v +++ b/utils/theories/All_Forall.v @@ -1,6 +1,7 @@ From Coq Require Import List Bool Arith ssreflect Morphisms Lia Utf8. From MetaCoq.Utils Require Import MCPrelude MCReflect MCList MCRelations MCProd MCOption. From Equations Require Import Equations. + Import ListNotations. Derive Signature for Forall Forall2. @@ -3016,11 +3017,11 @@ Proof. Qed. Lemma Forall2_sym : - forall A (P : A -> A -> Prop) l l', + forall A B (P : A -> B -> Prop) l l', Forall2 P l l' -> Forall2 (fun x y => P y x) l' l. Proof. - intros A P l l' h. + intros A B P l l' h. induction h. - constructor. - constructor. all: auto. From 9fd6f611eeca836f34807eb706e314e56ee7a01a Mon Sep 17 00:00:00 2001 From: nicolas tabareau Date: Tue, 26 Sep 2023 17:58:59 +0200 Subject: [PATCH 26/61] more general strengthening --- .../theories/Bidirectional/BDStrengthening.v | 62 ++++++++++--------- 1 file changed, 33 insertions(+), 29 deletions(-) diff --git a/pcuic/theories/Bidirectional/BDStrengthening.v b/pcuic/theories/Bidirectional/BDStrengthening.v index 827f91eaf..3882a6364 100644 --- a/pcuic/theories/Bidirectional/BDStrengthening.v +++ b/pcuic/theories/Bidirectional/BDStrengthening.v @@ -986,24 +986,26 @@ Theorem typing_renaming_cond_P `{checker_flags} {P f Σ Γ Δ t T} {wfΣ : wf Σ renaming P Σ Γ Δ f -> on_ctx_free_vars P Γ -> on_free_vars P t -> - on_free_vars P T -> Σ ;;; Γ |- t : T -> - ∑ T', (Σ ;;; Δ |- rename f t : T') * (Σ ;;; Δ |- T' <= rename f T). + ∑ T', [× (on_free_vars P T'), (Σ ;;; Γ |- t : T'), (Σ ;;; Γ |- T' <= T) & + (Σ ;;; Δ |- rename f t : rename f T')]. Proof. - move => [ur wfΔ] fvΓ fvt fvT tyt. + move => [ur wfΔ] fvΓ fvt tyt. move: (tyt) => /typing_wf_local wfΓ. move: (tyt) => /typing_infering [T' [inf cum]]. - exists (rename f T'). split; eauto. sq. - apply infering_typing => //. - eapply bidirectional_renaming ; eassumption. - eapply cumul_renameP with (Γ := Γ); eauto. - 2: { eapply ws_cumul_pb_forget_cumul; eauto. } - eapply infering_on_free_vars; try exact inf; eauto. + unshelve epose proof (infering_on_free_vars _ _ _ _ _ _ _ _ inf); eauto. + pose proof (inf' := inf). apply infering_typing in inf; eauto. + exists T'. split; eauto. + eapply ws_cumul_pb_forget_cumul; eauto. + eapply bidirectional_renaming in inf'; eauto. eapply infering_typing; eauto. Qed. Lemma strengthening `{cf: checker_flags} {Σ : global_env_ext} {wfΣ : wf Σ} Γ Γ' Γ'' t T : - Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : lift #|Γ'| #|Γ''| T -> - ∑ T', (Σ ;;; Γ ,,, Γ'' |- t : T') * (Σ ;;; Γ ,,, Γ'' |- T' <= T). + Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : T -> + ∑ T', [× on_free_vars (strengthenP #|Γ''| #|Γ'| xpredT) T', + (Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : T') , + (Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- T' <= T) & + (Σ ;;; Γ ,,, Γ'' |- t : unlift #|Γ'| #|Γ''| T')]. Proof. intros Hty. assert (wf_local Σ Γ) by @@ -1034,11 +1036,10 @@ Proof. eassumption. } - erewrite <- (lift_unlift_term t). erewrite <- (lift_unlift_term T). - eapply typing_renaming_cond_P. - 5: eassumption. - - split => //. - apply urenaming_strengthen. + eapply typing_renaming_cond_P in Hty as [T' [? ?]]; eauto. + exists T'. split; eauto. + - erewrite <- (lift_unlift_term t); eauto. + - split; eauto. eapply urenaming_strengthen. - move: wfΓ'' => /wf_local_closed_context. rewrite on_free_vars_ctx_app => /andP [? ?]. apply on_ctx_free_vars_strengthenP. @@ -1051,14 +1052,6 @@ Proof. 2: lia. eapply closedn_on_free_vars. eassumption. - - rewrite on_free_vars_lift -shiftnP_xpredT. - move: Hty => /type_closed. - len. - rewrite -[X in _ + X]Nat.add_comm Nat.add_assoc => Ht. - eapply closedn_lift_inv in Ht. - 2: lia. - eapply closedn_on_free_vars. - eassumption. Qed. @@ -1066,17 +1059,28 @@ Lemma strengthening_type `{cf: checker_flags} {Σ : global_env_ext} {wfΣ : PCUI Σ ;;; Γ ,,, Γ' ,,, lift_context #|Γ'| 0 Γ'' |- lift #|Γ'| #|Γ''| t : tSort s -> ∑ s', (Σ ;;; Γ ,,, Γ'' |- t : tSort s') * (compare_universe Cumul Σ s' s). Proof. - assert (liftSort : lift #|Γ'| #|Γ''| (tSort s) = tSort s) by reflexivity. rewrite <- liftSort; clear liftSort. - intros H; eapply strengthening in H. destruct H as [T' [HT' Hcumul]]. + intros H; pose proof (H' := H); eapply strengthening in H. destruct H as [T' [? HT Hcumul HT']]. pose proof (Hcumul' := Hcumul). - pose proof (type_closed HT'). + pose proof (type_closed HT). eapply closedn_on_free_vars in H. - pose proof (wf_local_closed_context (typing_wf_local HT')). + pose proof (type_closed HT'). + eapply closedn_on_free_vars in H0. + pose proof (wf_local_closed_context (typing_wf_local H')). eapply into_ws_cumul_pb in Hcumul; eauto. eapply PCUICConversion.ws_cumul_pb_Sort_r_inv in Hcumul as [s' [Hred Hcumul]]. eapply closed_red_red in Hred. pose proof (HT'':=HT'). eapply PCUICValidity.validity in HT' as [? ?]. exists s'; split; eauto. + assert (PCUICReduction.red Σ (Γ,,, Γ'') (unlift #|Γ'| #|Γ''| T') (tSort s')). + { assert (tSort s' = unlift #|Γ'| #|Γ''| (tSort s')) by reflexivity. rewrite H2; clear H2. + eapply red_rename. eauto. eapply (urenaming_strengthen _ _ _ _); eauto. + 2-3: eauto. eapply typing_wf_local in HT''. + eapply wf_local_closed_context in HT''. revert HT''. + rewrite on_free_vars_ctx_app => /andP [? ?]. + apply on_ctx_free_vars_strengthenP. + all: eapply on_free_vars_ctx_on_ctx_free_vars_xpredT; eauto. + } eapply type_Cumul; eauto. eapply PCUICSR.subject_reduction; eauto. + assert (liftSort : unlift #|Γ'| #|Γ''| (tSort s') = tSort s') by reflexivity. rewrite <- liftSort; clear liftSort. eapply PCUICConversion.cumulAlgo_cumulSpec. eapply PCUICConversion.ws_cumul_pb_red_r_inv; eauto. - eapply ws_cumul_pb_refl; eauto. + eapply ws_cumul_pb_refl; eauto. eapply PCUICInversion.typing_closed_ctx; eauto. Qed. \ No newline at end of file From be4447721e859796b52b4d9a4819e7a0bf2d64bc Mon Sep 17 00:00:00 2001 From: nicolas tabareau Date: Wed, 27 Sep 2023 16:47:01 +0200 Subject: [PATCH 27/61] remove parameters in firstorder inductive types --- pcuic/theories/PCUICFirstorder.v | 37 +++++++++++++++----------------- 1 file changed, 17 insertions(+), 20 deletions(-) diff --git a/pcuic/theories/PCUICFirstorder.v b/pcuic/theories/PCUICFirstorder.v index cf5bc91af..0942d8141 100644 --- a/pcuic/theories/PCUICFirstorder.v +++ b/pcuic/theories/PCUICFirstorder.v @@ -41,12 +41,12 @@ Section firstorder. end. *) Definition firstorder_type (n k : nat) (t : term) := - match (PCUICAstUtils.decompose_app t).1 with - | tInd (mkInd nm i) u => match (plookup_env Σb nm) with + match (PCUICAstUtils.decompose_app t).1, (PCUICAstUtils.decompose_app t).2 with + | tInd (mkInd nm i) u , [] => match (plookup_env Σb nm) with | Some b => b | None => false end - | tRel i => (k <=? i) && (i false + | tRel i , [] => (k <=? i) && (i false end. Definition firstorder_con mind (c : constructor_body) := @@ -432,6 +432,7 @@ Proof using Type. destruct plookup_env eqn:hl => //. destruct b => //. eapply (plookup_env_extends (Σ:=Σ)) in hl. 2:split; eauto. rewrite eq in hl. rewrite hl //. apply wf. + all: destruct l; eauto. Qed. Lemma firstorder_args {Σ : global_env_ext} {wfΣ : wf Σ} { mind cbody i n ui args u pandi oind} : @@ -480,7 +481,7 @@ Proof using Type. { intros k t. rewrite /firstorder_type. rewrite -PCUICUnivSubstitutionConv.subst_instance_decompose_app /=. - destruct (decompose_app) => //=. destruct t0 => //. } + destruct (decompose_app) => //=. destruct t0 => //; destruct l; eauto. } replace (List.rev c)@[ui] with (List.rev c@[ui]). 2:{ rewrite /subst_instance /subst_instance_context /map_context map_rev //. } revert args. @@ -511,17 +512,15 @@ Proof using Type. unfold firstorder_type; cbn. destruct (decompose_app decl_type) eqn:da. rewrite (decompose_app_inv da) subst_mkApps /=. - destruct t0 => //=. + destruct t0 => //=; destruct l; eauto. { move/andP => [/Nat.leb_le hn /Nat.ltb_lt hn']. destruct (Nat.leb_spec n n0). destruct (n0 - n) eqn:E. lia. - cbn. rewrite nth_error_nil /=. - rewrite decompose_app_mkApps //=. + cbn. rewrite nth_error_nil /=. apply/andP. split. apply Nat.leb_le. lia. apply Nat.ltb_lt. lia. - cbn. rewrite decompose_app_mkApps //=. apply/andP. split. apply Nat.leb_le. lia. apply Nat.ltb_lt. lia. } - { destruct ind => //. rewrite decompose_app_mkApps //. } + { destruct ind => //. } + move=> /andP[] fot foΓ. rewrite subst_context_app /=. rewrite it_mkProd_or_LetIn_app /= /mkProd_or_LetIn /=. @@ -529,9 +528,10 @@ Proof using Type. destruct ((PCUICAstUtils.decompose_app t)) eqn:E. cbn in fot. destruct t0; try solve [inv fot]. * rewrite (decompose_app_inv E) /= subst_mkApps. - rewrite Nat.add_0_r in fot. eapply Nat.ltb_lt in fot. + rewrite Nat.add_0_r in fot. destruct l; try solve [inversion fot]. + eapply Nat.ltb_lt in fot. cbn. rewrite nth_error_inds. lia. cbn. - econstructor. + eapply instantiated_tProd with (args := []). unshelve epose proof (d1_ := declared_minductive_to_gen d1); eauto. { rewrite /firstorder_ind d1_ /= /firstorder_mutind indf H0 //. } intros x. @@ -552,20 +552,19 @@ Proof using Type. unfold firstorder_type; cbn. destruct (decompose_app decl_type) eqn:da. rewrite (decompose_app_inv da) subst_mkApps /=. - destruct t0 => //=. + destruct t0 => //=; destruct l; eauto. { move/andP => [/Nat.leb_le hn /Nat.ltb_lt hn']. destruct (Nat.leb_spec n n0). destruct (n0 - n) eqn:E. lia. cbn. rewrite nth_error_nil /=. - rewrite decompose_app_mkApps //=. apply/andP. split. apply Nat.leb_le. lia. apply Nat.ltb_lt. lia. - cbn. rewrite decompose_app_mkApps //=. apply/andP. split. apply Nat.leb_le. lia. apply Nat.ltb_lt. lia. } - { destruct ind => //. rewrite decompose_app_mkApps //. } + { destruct ind => //. } * rewrite (decompose_app_inv E) subst_mkApps //=. constructor. { unfold firstorder_ind. destruct ind. cbn in *. + destruct l; [|inversion fot]. destruct plookup_env eqn:hp => //. eapply plookup_env_lookup_env in hp as [Σ' [decl [eq [ext he]]]]. rewrite eq. destruct decl; subst b => //. @@ -587,17 +586,15 @@ Proof using Type. unfold firstorder_type; cbn. destruct (decompose_app decl_type) eqn:da. rewrite (decompose_app_inv da) subst_mkApps /=. - destruct t0 => //=. + destruct t0 => //=; destruct l; eauto. { move/andP => [/Nat.leb_le hn /Nat.ltb_lt hn']. destruct (Nat.leb_spec n n0). destruct (n0 - n) eqn:E. lia. cbn. rewrite nth_error_nil /=. - rewrite decompose_app_mkApps //=. apply/andP. split. apply Nat.leb_le. lia. apply Nat.ltb_lt. lia. - cbn. rewrite decompose_app_mkApps //=. apply/andP. split. apply Nat.leb_le. lia. apply Nat.ltb_lt. lia. } - { destruct ind => //. rewrite decompose_app_mkApps //. } + { destruct ind => //. } } cbn in Hi |- *. revert Hi Hspine. cbn. From 00a42993444ad6f77c8f20921e4c3e2cdaf56b66 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Fri, 29 Sep 2023 23:17:25 -0700 Subject: [PATCH 28/61] Drastically speed up ByteCompareSpec Old: ``` theories/ByteCompareSpec.vo (real: 113.67, user: 111.75, sys: 1.57, mem: 1287536 ko) ``` New: ``` COQC theories/ByteCompareSpec.v theories/ByteCompareSpec.vo (real: 9.46, user: 8.89, sys: 0.26, mem: 442460 ko) ``` The new file is also compatible with COQNATIVE, taking only seconds rather than dozens of minutes or more. --- utils/theories/ByteCompareSpec.v | 26 +++++++++++++++++++------- 1 file changed, 19 insertions(+), 7 deletions(-) diff --git a/utils/theories/ByteCompareSpec.v b/utils/theories/ByteCompareSpec.v index 1bf10f910..207588924 100644 --- a/utils/theories/ByteCompareSpec.v +++ b/utils/theories/ByteCompareSpec.v @@ -1,4 +1,4 @@ -From Coq Require Import Strings.Byte NArith. +From Coq Require Import Strings.Byte NArith Eqdep_dec. From MetaCoq.Utils Require Import ReflectEq ByteCompare. From Equations Require Import Equations. @@ -6,19 +6,31 @@ Derive NoConfusion for comparison. Section ByteNoConf. Local Ltac solve_noconf ::= idtac. + (** The NoConfusion definition is quite large, so we define a much more compact version of it that is nevertheless convertible *) + Definition eta_byte {A} (f : forall b : byte, A b) (b : byte) : A b + := Eval cbv in + ltac:(pose (f b) as v; destruct b; let val := (eval cbv in v) in exact val). + Definition uneta_byte {A f b} : @eta_byte A f b = f b + := match b with x00 | _ => eq_refl end. + Definition NoConfusion_byte_alt : byte -> byte -> Prop + := eta_byte (fun b1 b2 => eta_byte (fun b2 => if Byte.eqb b1 b2 then True else False) b2). Derive NoConfusion for byte. Next Obligation. - destruct a; abstract (destruct b; try exact (False_ind _ H); exact eq_refl). - Defined. + change (NoConfusion_byte a b) with (NoConfusion_byte_alt a b) in *. + cbv [NoConfusion_byte_alt] in *; apply byte_dec_bl. + rewrite !uneta_byte in *. + destruct Byte.eqb; [ reflexivity | exfalso; assumption ]. + Qed. Next Obligation. destruct b; cbn; exact I. Defined. Next Obligation. - destruct a; abstract (destruct b; try (exact (False_ind _ e)); destruct e; exact eq_refl). - Defined. + lazymatch goal with |- ?f _ _ ?g = ?e => generalize g; revert e end; intros e g; subst; revert e. + destruct b; vm_compute; intros []; reflexivity. + Qed. Next Obligation. - destruct b; cbn; exact eq_refl. - Defined. + apply UIP_dec, byte_eq_dec. + Qed. End ByteNoConf. Lemma reflect_equiv P Q b : P <-> Q -> Bool.reflect P b -> Bool.reflect Q b. From dce0b9e51ce6335bed80cceca15b11f3087777ef Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Thu, 6 Jul 2023 08:06:27 +0200 Subject: [PATCH 29/61] split out verified erasure pipeline --- Makefile | 3 +- erasure-plugin/theories/ETransform.v | 12 +- erasure-plugin/theories/Erasure.v | 58 +++- erasure/theories/EOptimizePropDiscr.v | 334 ++++++++++----------- erasure/theories/EWcbvEvalNamed.v | 160 +++++++++- erasure/theories/Typed/Extraction.v | 2 +- erasure/theories/Typed/OptimizePropDiscr.v | 34 +-- erasure/theories/Typed/TypeAnnotations.v | 30 +- 8 files changed, 403 insertions(+), 230 deletions(-) diff --git a/Makefile b/Makefile index a99854e73..e16daadc8 100644 --- a/Makefile +++ b/Makefile @@ -1,5 +1,5 @@ -all: printconf template-coq pcuic safechecker erasure examples test-suite translations quotation +all: printconf template-coq pcuic safechecker erasure -include Makefile.conf @@ -37,7 +37,6 @@ install: all translations $(MAKE) -C safechecker-plugin install $(MAKE) -C erasure install $(MAKE) -C erasure-plugin install - $(MAKE) -C translations install uninstall: $(MAKE) -C utils uninstall diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index 822730bdd..521481293 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -358,23 +358,23 @@ Qed. Import EOptimizePropDiscr EWcbvEval. -Program Definition optimize_prop_discr_optimization {fl : WcbvFlags} {wcon : with_constructor_as_block = false} {efl : EEnvFlags} {hastrel : has_tRel} {hastbox : has_tBox} : +Program Definition remove_match_on_box_trans {fl : WcbvFlags} {wcon : with_constructor_as_block = false} {efl : EEnvFlags} {hastrel : has_tRel} {hastbox : has_tBox} : Transform.t eprogram_env eprogram EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram (disable_prop_cases fl)) := {| name := "optimize_prop_discr"; - transform p _ := optimize_program p ; + transform p _ := remove_match_on_box_program p ; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; post p := wf_eprogram efl p /\ EEtaExpanded.expanded_eprogram_cstrs p; - obseq g g' v v' := v' = EOptimizePropDiscr.optimize g.1 v |}. + obseq g g' v v' := v' = EOptimizePropDiscr.remove_match_on_box g.1 v |}. Next Obligation. move=> fl wcon efl hastrel hastbox [Σ t] [wfp etap]. cbn in *. split. - - now eapply optimize_program_wf. - - now eapply optimize_program_expanded. + - now eapply remove_match_on_box_program_wf. + - now eapply remove_match_on_box_program_expanded. Qed. Next Obligation. red. move=> fl wcon efl hastrel hastbox [Σ t] /= v [wfe wft] [ev]. - eapply EOptimizePropDiscr.optimize_correct in ev; eauto. + eapply EOptimizePropDiscr.remove_match_on_box_correct in ev; eauto. eexists; split => //. red. sq; auto. cbn. apply wfe. eapply wellformed_closed_env, wfe. eapply wellformed_closed, wfe. diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 05789e1c6..567ca5f4f 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -57,10 +57,10 @@ Next Obligation. apply assume_preservation_template_program_env_expansion in ev as [ev']; eauto. Qed. -Program Definition erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - Transform.t TemplateProgram.template_program EProgram.eprogram - Ast.term EAst.term - TemplateProgram.eval_template_program +Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : + Transform.t pcuic_program EProgram.eprogram + PCUICAst.term EAst.term + PCUICTransform.eval_pcuic_program (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := (* a bunch of nonsense for normalization preconditions *) let K ty (T : ty -> _) p @@ -68,15 +68,6 @@ Program Definition erasure_pipeline {guard : abstract_guard_impl} (efl := EWellf (PCUICTyping.wf_ext p -> PCUICSN.NormalizationIn p) /\ (PCUICTyping.wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p) in let T1 (p:global_env_ext_map) := p in - let T2 (p:global_env_ext_map) := T1 (build_global_env_map (PCUICExpandLets.trans_global_env p.1), p.2) in - let T3 (p:global_env) := T2 (TemplateToPCUIC.trans_global_env p, Monomorphic_ctx) in - let T4 (p:GlobalEnvMap.t) := T3 (eta_expand_global_env p) in - (* Build an efficient lookup map for the following eta-expansion phase *) - build_template_program_env (K _ T4) ▷ - (* Eta-expand constructors and fixpoint *) - eta_expand (K _ T3) ▷ - (* Casts are removed, application is binary, case annotations are inferred from the global environment *) - template_to_pcuic_transform (K _ T2) ▷ (* Branches of cases are expanded to bind only variables, constructor types are expanded accordingly *) pcuic_expand_lets_transform (K _ T1) ▷ (* Erasure of proofs terms in Prop and types *) @@ -90,7 +81,7 @@ Program Definition erasure_pipeline {guard : abstract_guard_impl} (efl := EWellf (* Rebuild the efficient lookup table *) rebuild_wf_env_transform (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) true ▷ (* Remove all cases / projections on propositional content *) - optimize_prop_discr_optimization (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ + remove_match_on_box_trans (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ (* Rebuild the efficient lookup table *) rebuild_wf_env_transform (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) true ▷ (* Inline projections to cases *) @@ -114,6 +105,43 @@ Next Obligation. now eapply ETransform.expanded_eprogram_env_expanded_eprogram_cstrs. Qed. +Program Definition pre_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : + Transform.t TemplateProgram.template_program pcuic_program + Ast.term PCUICAst.term + TemplateProgram.eval_template_program + PCUICTransform.eval_pcuic_program := + (* a bunch of nonsense for normalization preconditions *) + let K ty (T : ty -> _) p + := let p := T p in + (PCUICTyping.wf_ext p -> PCUICSN.NormalizationIn p) /\ + (PCUICTyping.wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p) in + let T1 (p:global_env_ext_map) := p in + let T2 (p:global_env_ext_map) := T1 (build_global_env_map (PCUICExpandLets.trans_global_env p.1), p.2) in + let T3 (p:global_env) := T2 (TemplateToPCUIC.trans_global_env p, Monomorphic_ctx) in + let T4 (p:GlobalEnvMap.t) := T3 (eta_expand_global_env p) in + (* Build an efficient lookup map for the following eta-expansion phase *) + build_template_program_env (K _ T4) ▷ + (* Eta-expand constructors and fixpoint *) + eta_expand (K _ T3) ▷ + (* Casts are removed, application is binary, case annotations are inferred from the global environment *) + template_to_pcuic_transform (K _ T2). + +Program Definition erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : + Transform.t TemplateProgram.template_program EProgram.eprogram + Ast.term EAst.term + TemplateProgram.eval_template_program + (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := + pre_erasure_pipeline ▷ + verified_erasure_pipeline. + +(* At the end of erasure we get a well-formed program (well-scoped globally and localy), without + parameters in inductive declarations. The constructor applications are also transformed to a first-order + "block" application, of the right length, and the evaluation relation does not need to consider guarded + fixpoints or case analyses on propositional content. All fixpoint bodies start with a lambda as well. + Finally, projections are inlined to cases, so no `tProj` remains. *) + +Import EGlobalEnv EWellformed. + Definition run_erase_program {guard : abstract_guard_impl} := run erasure_pipeline. Program Definition erasure_pipeline_fast {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) := @@ -134,7 +162,7 @@ Program Definition erasure_pipeline_fast {guard : abstract_guard_impl} (efl := E guarded_to_unguarded_fix (wcon := eq_refl) eq_refl ▷ remove_params_fast_optimization (wcon := eq_refl) _ ▷ rebuild_wf_env_transform (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) true ▷ - optimize_prop_discr_optimization (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ + remove_match_on_box_trans (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ rebuild_wf_env_transform (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) true ▷ inline_projections_optimization (fl := EWcbvEval.target_wcbv_flags) (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ let efl := EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags) in diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index bbf48f9af..70ad8701c 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -28,7 +28,7 @@ Ltac introdep := let H := fresh in intros H; depelim H. #[global] Hint Constructors eval : core. -Section optimize. +Section remove_match_on_box. Context (Σ : GlobalContextMap.t). Definition isprop_ind Σ (ind:inductive × nat) @@ -37,41 +37,41 @@ Section optimize. | _ => false end. - Fixpoint optimize (t : term) : term := + Fixpoint remove_match_on_box (t : term) : term := match t with | tRel i => tRel i - | tEvar ev args => tEvar ev (List.map optimize args) - | tLambda na M => tLambda na (optimize M) - | tApp u v => tApp (optimize u) (optimize v) - | tLetIn na b b' => tLetIn na (optimize b) (optimize b') + | tEvar ev args => tEvar ev (List.map remove_match_on_box args) + | tLambda na M => tLambda na (remove_match_on_box M) + | tApp u v => tApp (remove_match_on_box u) (remove_match_on_box v) + | tLetIn na b b' => tLetIn na (remove_match_on_box b) (remove_match_on_box b') | tCase ind c brs => - let brs' := List.map (on_snd optimize) brs in + let brs' := List.map (on_snd remove_match_on_box) brs in if isprop_ind Σ ind then match brs' with | [(a, b)] => ECSubst.substl (repeat tBox #|a|) b - | _ => tCase ind (optimize c) brs' + | _ => tCase ind (remove_match_on_box c) brs' end - else tCase ind (optimize c) brs' + else tCase ind (remove_match_on_box c) brs' | tProj p c => match GlobalContextMap.inductive_isprop_and_pars Σ p.(proj_ind) with | Some (true, _) => tBox - | _ => tProj p (optimize c) + | _ => tProj p (remove_match_on_box c) end | tFix mfix idx => - let mfix' := List.map (map_def optimize) mfix in + let mfix' := List.map (map_def remove_match_on_box) mfix in tFix mfix' idx | tCoFix mfix idx => - let mfix' := List.map (map_def optimize) mfix in + let mfix' := List.map (map_def remove_match_on_box) mfix in tCoFix mfix' idx | tBox => t | tVar _ => t | tConst _ => t - | tConstruct ind i args => tConstruct ind i (map optimize args) + | tConstruct ind i args => tConstruct ind i (map remove_match_on_box args) | tPrim _ => t end. - Lemma optimize_mkApps f l : optimize (mkApps f l) = mkApps (optimize f) (map optimize l). + Lemma remove_match_on_box_mkApps f l : remove_match_on_box (mkApps f l) = mkApps (remove_match_on_box f) (map remove_match_on_box l). Proof using Type. induction l using rev_ind; simpl; auto. now rewrite mkApps_app /= IHl map_app /= mkApps_app /=. @@ -82,7 +82,7 @@ Section optimize. now induction n; simpl; auto; rewrite IHn. Qed. - Lemma map_optimize_repeat_box n : map optimize (repeat tBox n) = repeat tBox n. + Lemma map_remove_match_on_box_repeat_box n : map remove_match_on_box (repeat tBox n) = repeat tBox n. Proof using Type. by rewrite map_repeat. Qed. Import ECSubst. @@ -106,7 +106,7 @@ Section optimize. destruct x0; cbn in *. f_equal; auto. Qed. - Lemma closed_optimize t k : closedn k t -> closedn k (optimize t). + Lemma closed_remove_match_on_box t k : closedn k t -> closedn k (remove_match_on_box t). Proof using Type. induction t in k |- * using EInduction.term_forall_list_ind; simpl; auto; intros; try easy; @@ -166,9 +166,9 @@ Section optimize. apply subst_csubst_comm => //. Qed. - Lemma optimize_csubst a k b : + Lemma remove_match_on_box_csubst a k b : closed a -> - optimize (ECSubst.csubst a k b) = ECSubst.csubst (optimize a) k (optimize b). + remove_match_on_box (ECSubst.csubst a k b) = ECSubst.csubst (remove_match_on_box a) k (remove_match_on_box b). Proof using Type. induction b in k |- * using EInduction.term_forall_list_ind; simpl; auto; intros cl; try easy; @@ -186,7 +186,7 @@ Section optimize. rewrite {2}H. rewrite substl_csubst_comm //. solve_all. eapply All_repeat => //. - now eapply closed_optimize. + now eapply closed_remove_match_on_box. * depelim X. depelim X. f_equal; eauto. unfold on_snd; cbn. f_equal; eauto. @@ -199,28 +199,28 @@ Section optimize. now rewrite IHb. Qed. - Lemma optimize_substl s t : + Lemma remove_match_on_box_substl s t : forallb (closedn 0) s -> - optimize (substl s t) = substl (map optimize s) (optimize t). + remove_match_on_box (substl s t) = substl (map remove_match_on_box s) (remove_match_on_box t). Proof using Type. induction s in t |- *; simpl; auto. move/andP => [] cla cls. rewrite IHs //. f_equal. - now rewrite optimize_csubst. + now rewrite remove_match_on_box_csubst. Qed. - Lemma optimize_iota_red pars args br : + Lemma remove_match_on_box_iota_red pars args br : forallb (closedn 0) args -> - optimize (EGlobalEnv.iota_red pars args br) = EGlobalEnv.iota_red pars (map optimize args) (on_snd optimize br). + remove_match_on_box (EGlobalEnv.iota_red pars args br) = EGlobalEnv.iota_red pars (map remove_match_on_box args) (on_snd remove_match_on_box br). Proof using Type. intros cl. unfold EGlobalEnv.iota_red. - rewrite optimize_substl //. + rewrite remove_match_on_box_substl //. rewrite forallb_rev forallb_skipn //. now rewrite map_rev map_skipn. Qed. - Lemma optimize_fix_subst mfix : EGlobalEnv.fix_subst (map (map_def optimize) mfix) = map optimize (EGlobalEnv.fix_subst mfix). + Lemma remove_match_on_box_fix_subst mfix : EGlobalEnv.fix_subst (map (map_def remove_match_on_box) mfix) = map remove_match_on_box (EGlobalEnv.fix_subst mfix). Proof using Type. unfold EGlobalEnv.fix_subst. rewrite map_length. @@ -229,7 +229,7 @@ Section optimize. f_equal; auto. Qed. - Lemma optimize_cofix_subst mfix : EGlobalEnv.cofix_subst (map (map_def optimize) mfix) = map optimize (EGlobalEnv.cofix_subst mfix). + Lemma remove_match_on_box_cofix_subst mfix : EGlobalEnv.cofix_subst (map (map_def remove_match_on_box) mfix) = map remove_match_on_box (EGlobalEnv.cofix_subst mfix). Proof using Type. unfold EGlobalEnv.cofix_subst. rewrite map_length. @@ -238,41 +238,41 @@ Section optimize. f_equal; auto. Qed. - Lemma optimize_cunfold_fix mfix idx n f : + Lemma remove_match_on_box_cunfold_fix mfix idx n f : forallb (closedn 0) (EGlobalEnv.fix_subst mfix) -> cunfold_fix mfix idx = Some (n, f) -> - cunfold_fix (map (map_def optimize) mfix) idx = Some (n, optimize f). + cunfold_fix (map (map_def remove_match_on_box) mfix) idx = Some (n, remove_match_on_box f). Proof using Type. intros hfix. unfold cunfold_fix. rewrite nth_error_map. destruct nth_error. intros [= <- <-] => /=. f_equal. - now rewrite optimize_substl // optimize_fix_subst. + now rewrite remove_match_on_box_substl // remove_match_on_box_fix_subst. discriminate. Qed. - Lemma optimize_cunfold_cofix mfix idx n f : + Lemma remove_match_on_box_cunfold_cofix mfix idx n f : forallb (closedn 0) (EGlobalEnv.cofix_subst mfix) -> cunfold_cofix mfix idx = Some (n, f) -> - cunfold_cofix (map (map_def optimize) mfix) idx = Some (n, optimize f). + cunfold_cofix (map (map_def remove_match_on_box) mfix) idx = Some (n, remove_match_on_box f). Proof using Type. intros hcofix. unfold cunfold_cofix. rewrite nth_error_map. destruct nth_error. intros [= <- <-] => /=. f_equal. - now rewrite optimize_substl // optimize_cofix_subst. + now rewrite remove_match_on_box_substl // remove_match_on_box_cofix_subst. discriminate. Qed. - Lemma optimize_nth {n l d} : - optimize (nth n l d) = nth n (map optimize l) (optimize d). + Lemma remove_match_on_box_nth {n l d} : + remove_match_on_box (nth n l d) = nth n (map remove_match_on_box l) (remove_match_on_box d). Proof using Type. induction l in n |- *; destruct n; simpl; auto. Qed. -End optimize. +End remove_match_on_box. Lemma is_box_inv b : is_box b -> ∑ args, b = mkApps tBox args. Proof. @@ -313,26 +313,26 @@ Proof. eexists; econstructor; eauto. Qed. -Definition optimize_constant_decl Σ cb := - {| cst_body := option_map (optimize Σ) cb.(cst_body) |}. +Definition remove_match_on_box_constant_decl Σ cb := + {| cst_body := option_map (remove_match_on_box Σ) cb.(cst_body) |}. -Definition optimize_decl Σ d := +Definition remove_match_on_box_decl Σ d := match d with - | ConstantDecl cb => ConstantDecl (optimize_constant_decl Σ cb) + | ConstantDecl cb => ConstantDecl (remove_match_on_box_constant_decl Σ cb) | InductiveDecl idecl => d end. -Definition optimize_env Σ := - map (on_snd (optimize_decl Σ)) Σ.(GlobalContextMap.global_decls). +Definition remove_match_on_box_env Σ := + map (on_snd (remove_match_on_box_decl Σ)) Σ.(GlobalContextMap.global_decls). Import EnvMap. -Program Fixpoint optimize_env' Σ : EnvMap.fresh_globals Σ -> global_context := +Program Fixpoint remove_match_on_box_env' Σ : EnvMap.fresh_globals Σ -> global_context := match Σ with | [] => fun _ => [] | hd :: tl => fun HΣ => let Σg := GlobalContextMap.make tl (fresh_globals_cons_inv HΣ) in - on_snd (optimize_decl Σg) hd :: optimize_env' tl (fresh_globals_cons_inv HΣ) + on_snd (remove_match_on_box_decl Σg) hd :: remove_match_on_box_env' tl (fresh_globals_cons_inv HΣ) end. Import EGlobalEnv EExtends. @@ -363,10 +363,10 @@ Proof. destruct nth_error => //. Qed. -Lemma wellformed_optimize_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : +Lemma wellformed_remove_match_on_box_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : forall n, EWellformed.wellformed Σ n t -> forall {Σ' : GlobalContextMap.t}, extends Σ Σ' -> wf_glob Σ' -> - optimize Σ t = optimize Σ' t. + remove_match_on_box Σ t = remove_match_on_box Σ' t. Proof. induction t using EInduction.term_forall_list_ind; cbn -[lookup_constant lookup_inductive lookup_projection @@ -375,7 +375,7 @@ Proof. all:try now f_equal; eauto; solve_all. - destruct cstr_as_blocks; rtoProp; eauto. f_equal. solve_all. destruct args; inv H2. reflexivity. - rewrite !GlobalContextMap.inductive_isprop_and_pars_spec. - assert (map (on_snd (optimize Σ)) l = map (on_snd (optimize Σ')) l) as -> by solve_all. + assert (map (on_snd (remove_match_on_box Σ)) l = map (on_snd (remove_match_on_box Σ')) l) as -> by solve_all. rewrite (extends_inductive_isprop_and_pars H0 H1 H2). destruct inductive_isprop_and_pars as [[[]]|]. destruct map => //. f_equal; eauto. @@ -390,23 +390,23 @@ Proof. all:f_equal; eauto. Qed. -Lemma wellformed_optimize_decl_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : +Lemma wellformed_remove_match_on_box_decl_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : wf_global_decl Σ t -> forall {Σ' : GlobalContextMap.t}, extends Σ Σ' -> wf_glob Σ' -> - optimize_decl Σ t = optimize_decl Σ' t. + remove_match_on_box_decl Σ t = remove_match_on_box_decl Σ' t. Proof. destruct t => /= //. - intros wf Σ' ext wf'. f_equal. unfold optimize_constant_decl. f_equal. + intros wf Σ' ext wf'. f_equal. unfold remove_match_on_box_constant_decl. f_equal. destruct (cst_body c) => /= //. f_equal. - now eapply wellformed_optimize_extends. + now eapply wellformed_remove_match_on_box_extends. Qed. -Lemma lookup_env_optimize_env_Some {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn d : +Lemma lookup_env_remove_match_on_box_env_Some {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn d : wf_glob Σ -> GlobalContextMap.lookup_env Σ kn = Some d -> ∑ Σ' : GlobalContextMap.t, [× extends Σ' Σ, wf_global_decl Σ' d & - lookup_env (optimize_env Σ) kn = Some (optimize_decl Σ' d)]. + lookup_env (remove_match_on_box_env Σ) kn = Some (remove_match_on_box_decl Σ' d)]. Proof. rewrite GlobalContextMap.lookup_env_spec. destruct Σ as [Σ map repr wf]. @@ -417,7 +417,7 @@ Proof. exists (GlobalContextMap.make Σ (fresh_globals_cons_inv wf)). split. now eexists [_]. cbn. now depelim wfg. - f_equal. symmetry. eapply wellformed_optimize_decl_extends. cbn. now depelim wfg. + f_equal. symmetry. eapply wellformed_remove_match_on_box_decl_extends. cbn. now depelim wfg. cbn. now exists [a]. now cbn. - intros _. set (Σ' := GlobalContextMap.make Σ (fresh_globals_cons_inv wf)). @@ -430,7 +430,7 @@ Proof. * rewrite -hl'. f_equal. clear -wfg. eapply map_ext_in => kn hin. unfold on_snd. f_equal. - symmetry. eapply wellformed_optimize_decl_extends => //. cbn. + symmetry. eapply wellformed_remove_match_on_box_decl_extends => //. cbn. eapply lookup_env_In in hin. 2:now depelim wfg. depelim wfg. eapply lookup_env_wellformed; tea. cbn. now exists [a]. @@ -442,48 +442,48 @@ Proof. case: eqb_spec => //. Qed. -Lemma lookup_env_optimize_env_None {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn : +Lemma lookup_env_remove_match_on_box_env_None {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn : GlobalContextMap.lookup_env Σ kn = None -> - lookup_env (optimize_env Σ) kn = None. + lookup_env (remove_match_on_box_env Σ) kn = None. Proof. rewrite GlobalContextMap.lookup_env_spec. destruct Σ as [Σ map repr wf]. cbn. intros hl. rewrite lookup_env_map_snd hl //. Qed. -Lemma lookup_env_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn : +Lemma lookup_env_remove_match_on_box {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn : wf_glob Σ -> - lookup_env (optimize_env Σ) kn = option_map (optimize_decl Σ) (lookup_env Σ kn). + lookup_env (remove_match_on_box_env Σ) kn = option_map (remove_match_on_box_decl Σ) (lookup_env Σ kn). Proof. intros wf. rewrite -GlobalContextMap.lookup_env_spec. destruct (GlobalContextMap.lookup_env Σ kn) eqn:hl. - - eapply lookup_env_optimize_env_Some in hl as [Σ' [ext wf' hl']] => /=. + - eapply lookup_env_remove_match_on_box_env_Some in hl as [Σ' [ext wf' hl']] => /=. rewrite hl'. f_equal. - eapply wellformed_optimize_decl_extends; eauto. auto. + eapply wellformed_remove_match_on_box_decl_extends; eauto. auto. - - cbn. now eapply lookup_env_optimize_env_None in hl. + - cbn. now eapply lookup_env_remove_match_on_box_env_None in hl. Qed. -Lemma is_propositional_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : +Lemma is_propositional_remove_match_on_box {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : wf_glob Σ -> - inductive_isprop_and_pars Σ ind = inductive_isprop_and_pars (optimize_env Σ) ind. + inductive_isprop_and_pars Σ ind = inductive_isprop_and_pars (remove_match_on_box_env Σ) ind. Proof. rewrite /inductive_isprop_and_pars => wf. rewrite /lookup_inductive /lookup_minductive. - rewrite (lookup_env_optimize (inductive_mind ind) wf). + rewrite (lookup_env_remove_match_on_box (inductive_mind ind) wf). rewrite /GlobalContextMap.inductive_isprop_and_pars /GlobalContextMap.lookup_inductive /GlobalContextMap.lookup_minductive. destruct lookup_env as [[decl|]|] => //. Qed. -Lemma is_propositional_cstr_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind c : +Lemma is_propositional_cstr_remove_match_on_box {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind c : wf_glob Σ -> - constructor_isprop_pars_decl Σ ind c = constructor_isprop_pars_decl (optimize_env Σ) ind c. + constructor_isprop_pars_decl Σ ind c = constructor_isprop_pars_decl (remove_match_on_box_env Σ) ind c. Proof. rewrite /constructor_isprop_pars_decl => wf. rewrite /lookup_constructor /lookup_inductive /lookup_minductive. - rewrite (lookup_env_optimize (inductive_mind ind) wf). + rewrite (lookup_env_remove_match_on_box (inductive_mind ind) wf). rewrite /GlobalContextMap.inductive_isprop_and_pars /GlobalContextMap.lookup_inductive /GlobalContextMap.lookup_minductive. destruct lookup_env as [[decl|]|] => //. @@ -511,12 +511,12 @@ Proof. - rewrite mkApps_app /=. now destruct l => /= //; rewrite andb_false_r. Qed. -Lemma lookup_constructor_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} {ind c} : +Lemma lookup_constructor_remove_match_on_box {efl : EEnvFlags} {Σ : GlobalContextMap.t} {ind c} : wf_glob Σ -> - lookup_constructor Σ ind c = lookup_constructor (optimize_env Σ) ind c. + lookup_constructor Σ ind c = lookup_constructor (remove_match_on_box_env Σ) ind c. Proof. intros wfΣ. rewrite /lookup_constructor /lookup_inductive /lookup_minductive. - rewrite lookup_env_optimize // /=. destruct lookup_env => // /=. + rewrite lookup_env_remove_match_on_box // /=. destruct lookup_env => // /=. destruct g => //. Qed. @@ -529,12 +529,12 @@ Proof. destruct nth_error => //. congruence. Qed. -Lemma optimize_correct {efl : EEnvFlags} {fl}{wcon : with_constructor_as_block = false} {Σ : GlobalContextMap.t} t v : +Lemma remove_match_on_box_correct {efl : EEnvFlags} {fl}{wcon : with_constructor_as_block = false} {Σ : GlobalContextMap.t} t v : wf_glob Σ -> closed_env Σ -> @Ee.eval fl Σ t v -> closed t -> - @Ee.eval (disable_prop_cases fl) (optimize_env Σ) (optimize Σ t) (optimize Σ v). + @Ee.eval (disable_prop_cases fl) (remove_match_on_box_env Σ) (remove_match_on_box Σ t) (remove_match_on_box Σ v). Proof. intros wfΣ clΣ ev. induction ev; simpl in *. @@ -544,23 +544,23 @@ Proof. eapply eval_closed in ev2; tea. eapply eval_closed in ev1; tea. econstructor; eauto. - rewrite optimize_csubst // in IHev3. + rewrite remove_match_on_box_csubst // in IHev3. apply IHev3. eapply closed_csubst => //. - - move/andP => [] clb0 clb1. rewrite optimize_csubst in IHev2. + - move/andP => [] clb0 clb1. rewrite remove_match_on_box_csubst in IHev2. now eapply eval_closed in ev1. econstructor; eauto. eapply IHev2, closed_csubst => //. now eapply eval_closed in ev1. - - move/andP => [] cld clbrs. rewrite optimize_mkApps in IHev1. + - move/andP => [] cld clbrs. rewrite remove_match_on_box_mkApps in IHev1. have := (eval_closed _ clΣ _ _ cld ev1); rewrite closedn_mkApps => /andP[] _ clargs. - rewrite optimize_iota_red in IHev2. + rewrite remove_match_on_box_iota_red in IHev2. eapply eval_closed in ev1 => //. unfold isprop_ind. rewrite GlobalContextMap.inductive_isprop_and_pars_spec. rewrite (constructor_isprop_pars_decl_inductive e1). eapply eval_iota; eauto. - now rewrite -is_propositional_cstr_optimize. + now rewrite -is_propositional_cstr_remove_match_on_box. rewrite nth_error_map e2 //. now len. cbn. rewrite -e4. rewrite !skipn_length map_length //. eapply IHev2. @@ -574,24 +574,24 @@ Proof. rewrite GlobalContextMap.inductive_isprop_and_pars_spec. rewrite e0 e1 /=. subst brs. cbn in clbrs. rewrite Nat.add_0_r andb_true_r in clbrs. - rewrite optimize_substl in IHev2. + rewrite remove_match_on_box_substl in IHev2. eapply All_forallb, All_repeat => //. - rewrite map_optimize_repeat_box in IHev2. + rewrite map_remove_match_on_box_repeat_box in IHev2. apply IHev2. eapply closed_substl. eapply All_forallb, All_repeat => //. now rewrite repeat_length Nat.add_0_r. - - move/andP => [] clf cla. rewrite optimize_mkApps in IHev1. + - move/andP => [] clf cla. rewrite remove_match_on_box_mkApps in IHev1. simpl in *. eapply eval_closed in ev1 => //. rewrite closedn_mkApps in ev1. move: ev1 => /andP [] clfix clargs. eapply Ee.eval_fix; eauto. rewrite map_length. - eapply optimize_cunfold_fix; tea. + eapply remove_match_on_box_cunfold_fix; tea. eapply closed_fix_subst. tea. - rewrite optimize_mkApps in IHev3. apply IHev3. + rewrite remove_match_on_box_mkApps in IHev3. apply IHev3. rewrite closedn_mkApps clargs. eapply eval_closed in ev2; tas. rewrite ev2 /= !andb_true_r. eapply closed_cunfold_fix; tea. @@ -601,9 +601,9 @@ Proof. rewrite closedn_mkApps in ev1. move: ev1 => /andP [] clfix clargs. eapply eval_closed in ev2; tas. - rewrite optimize_mkApps in IHev1 |- *. + rewrite remove_match_on_box_mkApps in IHev1 |- *. simpl in *. eapply Ee.eval_fix_value. auto. auto. auto. - eapply optimize_cunfold_fix; eauto. + eapply remove_match_on_box_cunfold_fix; eauto. eapply closed_fix_subst => //. now rewrite map_length. @@ -611,7 +611,7 @@ Proof. eapply eval_closed in ev1 => //. eapply eval_closed in ev2; tas. simpl in *. eapply Ee.eval_fix'. auto. auto. - eapply optimize_cunfold_fix; eauto. + eapply remove_match_on_box_cunfold_fix; eauto. eapply closed_fix_subst => //. eapply IHev2; tea. eapply IHev3. apply/andP; split => //. @@ -625,21 +625,21 @@ Proof. forward IHev2. { rewrite clargs clbrs !andb_true_r. eapply closed_cunfold_cofix; tea. } - rewrite -> optimize_mkApps in IHev1, IHev2. simpl. unfold isprop_ind in *. + rewrite -> remove_match_on_box_mkApps in IHev1, IHev2. simpl. unfold isprop_ind in *. rewrite GlobalContextMap.inductive_isprop_and_pars_spec in IHev2 |- *. destruct EGlobalEnv.inductive_isprop_and_pars as [[[] pars]|] eqn:isp => //. destruct brs as [|[a b] []]; simpl in *; auto. simpl in IHev1. eapply Ee.eval_cofix_case. tea. - apply optimize_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + apply remove_match_on_box_cunfold_cofix; tea. eapply closed_cofix_subst; tea. apply IHev2. eapply Ee.eval_cofix_case; tea. - apply optimize_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + apply remove_match_on_box_cunfold_cofix; tea. eapply closed_cofix_subst; tea. simpl in *. eapply Ee.eval_cofix_case; tea. - apply optimize_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + apply remove_match_on_box_cunfold_cofix; tea. eapply closed_cofix_subst; tea. eapply Ee.eval_cofix_case; tea. - apply optimize_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + apply remove_match_on_box_cunfold_cofix; tea. eapply closed_cofix_subst; tea. - intros cd. specialize (IHev1 cd). move: (eval_closed _ clΣ _ _ cd ev1). @@ -647,15 +647,15 @@ Proof. { rewrite closedn_mkApps clargs andb_true_r. eapply closed_cunfold_cofix; tea. } rewrite GlobalContextMap.inductive_isprop_and_pars_spec in IHev2 |- *. destruct EGlobalEnv.inductive_isprop_and_pars as [[[] pars]|] eqn:isp; auto. - rewrite -> optimize_mkApps in IHev1, IHev2. simpl in *. + rewrite -> remove_match_on_box_mkApps in IHev1, IHev2. simpl in *. econstructor; eauto. - apply optimize_cunfold_cofix; tea. eapply closed_cofix_subst; tea. - rewrite -> optimize_mkApps in IHev1, IHev2. simpl in *. + apply remove_match_on_box_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + rewrite -> remove_match_on_box_mkApps in IHev1, IHev2. simpl in *. econstructor; eauto. - apply optimize_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + apply remove_match_on_box_cunfold_cofix; tea. eapply closed_cofix_subst; tea. - rewrite /declared_constant in isdecl. - move: (lookup_env_optimize c wfΣ). + move: (lookup_env_remove_match_on_box c wfΣ). rewrite isdecl /= //. intros hl. econstructor; tea. cbn. rewrite e //. @@ -668,10 +668,10 @@ Proof. move: ev1; rewrite closedn_mkApps /= => clargs. rewrite GlobalContextMap.inductive_isprop_and_pars_spec. rewrite (constructor_isprop_pars_decl_inductive e1). - rewrite optimize_mkApps in IHev1. + rewrite remove_match_on_box_mkApps in IHev1. specialize (IHev1 cld). eapply Ee.eval_proj; tea. - now rewrite -is_propositional_cstr_optimize. + now rewrite -is_propositional_cstr_remove_match_on_box. now len. rewrite nth_error_map e3 //. eapply IHev2. eapply nth_error_forallb in e3; tea. @@ -682,10 +682,10 @@ Proof. now rewrite e0. - move/andP=> [] clf cla. - rewrite optimize_mkApps. + rewrite remove_match_on_box_mkApps. eapply eval_construct; tea. - rewrite -lookup_constructor_optimize //. exact e0. - rewrite optimize_mkApps in IHev1. now eapply IHev1. + rewrite -lookup_constructor_remove_match_on_box //. exact e0. + rewrite remove_match_on_box_mkApps in IHev1. now eapply IHev1. now len. now eapply IHev2. @@ -696,11 +696,11 @@ Proof. eapply Ee.eval_app_cong; eauto. eapply Ee.eval_to_value in ev1. destruct ev1; simpl in *; eauto. - * destruct t => //; rewrite optimize_mkApps /=. + * destruct t => //; rewrite remove_match_on_box_mkApps /=. * destruct with_guarded_fix. + move: i. rewrite !negb_or. - rewrite optimize_mkApps !isFixApp_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + rewrite remove_match_on_box_mkApps !isFixApp_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. rewrite !andb_true_r. rtoProp; intuition auto. @@ -709,26 +709,26 @@ Proof. destruct v => /= //. + move: i. rewrite !negb_or. - rewrite optimize_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + rewrite remove_match_on_box_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. destruct v => /= //. - destruct t => //. - all:constructor; eauto. cbn [atom optimize] in i |- *. - rewrite -lookup_constructor_optimize //. destruct l => //. + all:constructor; eauto. cbn [atom remove_match_on_box] in i |- *. + rewrite -lookup_constructor_remove_match_on_box //. destruct l => //. Qed. From MetaCoq.Erasure Require Import EEtaExpanded. -Lemma isLambda_optimize Σ t : isLambda t -> isLambda (optimize Σ t). +Lemma isLambda_remove_match_on_box Σ t : isLambda t -> isLambda (remove_match_on_box Σ t). Proof. destruct t => //. Qed. -Lemma isBox_optimize Σ t : isBox t -> isBox (optimize Σ t). +Lemma isBox_remove_match_on_box Σ t : isBox t -> isBox (remove_match_on_box Σ t). Proof. destruct t => //. Qed. -Lemma optimize_expanded {Σ : GlobalContextMap.t} t : expanded Σ t -> expanded Σ (optimize Σ t). +Lemma remove_match_on_box_expanded {Σ : GlobalContextMap.t} t : expanded Σ t -> expanded Σ (remove_match_on_box Σ t). Proof. induction 1 using expanded_ind. all:try solve[constructor; eauto; solve_all]. - all:rewrite ?optimize_mkApps. + all:rewrite ?remove_match_on_box_mkApps. - eapply expanded_mkApps_expanded => //. solve_all. - cbn -[GlobalContextMap.inductive_isprop_and_pars]. unfold isprop_ind. rewrite GlobalContextMap.inductive_isprop_and_pars_spec. @@ -750,40 +750,40 @@ Proof. destruct inductive_isprop_and_pars as [[[|] _]|] => /= //. constructor. all:constructor; auto. - cbn. eapply expanded_tFix. solve_all. - rewrite isLambda_optimize //. + rewrite isLambda_remove_match_on_box //. - eapply expanded_tConstruct_app; tea. now len. solve_all. Qed. -Lemma optimize_expanded_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded Σ t -> expanded (optimize_env Σ) t. +Lemma remove_match_on_box_expanded_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded Σ t -> expanded (remove_match_on_box_env Σ) t. Proof. intros wf; induction 1 using expanded_ind. all:try solve[constructor; eauto; solve_all]. eapply expanded_tConstruct_app. destruct H as [[H ?] ?]. split => //. split => //. red. - red in H. rewrite lookup_env_optimize // /= H //. 1-2:eauto. auto. solve_all. + red in H. rewrite lookup_env_remove_match_on_box // /= H //. 1-2:eauto. auto. solve_all. Qed. -Lemma optimize_expanded_decl {Σ : GlobalContextMap.t} t : expanded_decl Σ t -> expanded_decl Σ (optimize_decl Σ t). +Lemma remove_match_on_box_expanded_decl {Σ : GlobalContextMap.t} t : expanded_decl Σ t -> expanded_decl Σ (remove_match_on_box_decl Σ t). Proof. destruct t as [[[]]|] => /= //. unfold expanded_constant_decl => /=. - apply optimize_expanded. + apply remove_match_on_box_expanded. Qed. -Lemma optimize_expanded_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded_decl Σ t -> expanded_decl (optimize_env Σ) t. +Lemma remove_match_on_box_expanded_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded_decl Σ t -> expanded_decl (remove_match_on_box_env Σ) t. Proof. destruct t as [[[]]|] => /= //. unfold expanded_constant_decl => /=. - apply optimize_expanded_irrel. + apply remove_match_on_box_expanded_irrel. Qed. -Lemma optimize_env_extends' {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : +Lemma remove_match_on_box_env_extends' {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : extends Σ Σ' -> wf_glob Σ' -> - List.map (on_snd (optimize_decl Σ)) Σ.(GlobalContextMap.global_decls) = - List.map (on_snd (optimize_decl Σ')) Σ.(GlobalContextMap.global_decls). + List.map (on_snd (remove_match_on_box_decl Σ)) Σ.(GlobalContextMap.global_decls) = + List.map (on_snd (remove_match_on_box_decl Σ')) Σ.(GlobalContextMap.global_decls). Proof. intros ext. destruct Σ as [Σ map repr wf]; cbn in *. @@ -798,48 +798,48 @@ Proof. f_equal. 2:{ eapply IHΣ'' => //. destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. } unfold on_snd. cbn. f_equal. - eapply wellformed_optimize_decl_extends => //. cbn. + eapply wellformed_remove_match_on_box_decl_extends => //. cbn. eapply extends_wf_global_decl. 3:tea. eapply extends_wf_glob; tea. destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. Qed. -Lemma optimize_env_eq {efl : EEnvFlags} (Σ : GlobalContextMap.t) : wf_glob Σ -> optimize_env Σ = optimize_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). +Lemma remove_match_on_box_env_eq {efl : EEnvFlags} (Σ : GlobalContextMap.t) : wf_glob Σ -> remove_match_on_box_env Σ = remove_match_on_box_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). Proof. intros wf. - unfold optimize_env. + unfold remove_match_on_box_env. destruct Σ; cbn. cbn in wf. induction global_decls in map, repr, wf0, wf |- * => //. cbn. f_equal. destruct a as [kn d]; unfold on_snd; cbn. f_equal. symmetry. - eapply wellformed_optimize_decl_extends => //. cbn. now depelim wf. cbn. now exists [(kn, d)]. cbn. + eapply wellformed_remove_match_on_box_decl_extends => //. cbn. now depelim wf. cbn. now exists [(kn, d)]. cbn. set (Σg' := GlobalContextMap.make global_decls (fresh_globals_cons_inv wf0)). erewrite <- (IHglobal_decls (GlobalContextMap.map Σg') (GlobalContextMap.repr Σg')). 2:now depelim wf. set (Σg := {| GlobalContextMap.global_decls := _ :: _ |}). - symmetry. eapply (optimize_env_extends' (Σ := Σg') (Σ' := Σg)) => //. + symmetry. eapply (remove_match_on_box_env_extends' (Σ := Σg') (Σ' := Σg)) => //. cbn. now exists [a]. Qed. -Lemma optimize_env_expanded {efl : EEnvFlags} {Σ : GlobalContextMap.t} : - wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (optimize_env Σ). +Lemma remove_match_on_box_env_expanded {efl : EEnvFlags} {Σ : GlobalContextMap.t} : + wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (remove_match_on_box_env Σ). Proof. unfold expanded_global_env; move=> wfg. - rewrite optimize_env_eq //. + rewrite remove_match_on_box_env_eq //. destruct Σ as [Σ map repr wf]. cbn in *. clear map repr. induction 1; cbn; constructor; auto. cbn in IHexpanded_global_declarations. unshelve eapply IHexpanded_global_declarations. now depelim wfg. cbn. set (Σ' := GlobalContextMap.make _ _). - rewrite -(optimize_env_eq Σ'). cbn. now depelim wfg. - eapply (optimize_expanded_decl_irrel (Σ := Σ')). now depelim wfg. - now unshelve eapply (optimize_expanded_decl (Σ:=Σ')). + rewrite -(remove_match_on_box_env_eq Σ'). cbn. now depelim wfg. + eapply (remove_match_on_box_expanded_decl_irrel (Σ := Σ')). now depelim wfg. + now unshelve eapply (remove_match_on_box_expanded_decl (Σ:=Σ')). Qed. -Lemma optimize_wellformed {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : +Lemma remove_match_on_box_wellformed {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : has_tBox -> has_tRel -> - wf_glob Σ -> EWellformed.wellformed Σ n t -> EWellformed.wellformed Σ n (optimize Σ t). + wf_glob Σ -> EWellformed.wellformed Σ n t -> EWellformed.wellformed Σ n (remove_match_on_box Σ t). Proof. intros wfΣ hbox hrel. induction t in n |- * using EInduction.term_forall_list_ind => //. @@ -865,7 +865,7 @@ Proof. - cbn -[GlobalContextMap.inductive_isprop_and_pars lookup_inductive]. move/andP => [] /andP[]hasc hs ht. destruct GlobalContextMap.inductive_isprop_and_pars as [[[|] _]|] => /= //. all:rewrite hasc hs /=; eauto. - - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now eapply isLambda_optimize. now len. + - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now eapply isLambda_remove_match_on_box. now len. unfold test_def in *. len. eauto. - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now len. unfold test_def in *. len. eauto. @@ -873,26 +873,26 @@ Qed. Import EWellformed. -Lemma optimize_wellformed_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : +Lemma remove_match_on_box_wellformed_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> - forall n, wellformed Σ n t -> wellformed (optimize_env Σ) n t. + forall n, wellformed Σ n t -> wellformed (remove_match_on_box_env Σ) n t. Proof. intros wfΣ. induction t using EInduction.term_forall_list_ind; cbn => //. all:try solve [intros; unfold wf_fix_gen in *; rtoProp; intuition eauto; solve_all]. - - rewrite lookup_env_optimize //. + - rewrite lookup_env_remove_match_on_box //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct (cst_body c) => //. - - rewrite lookup_env_optimize //. + - rewrite lookup_env_remove_match_on_box //. destruct lookup_env eqn:hl => // /=; intros; rtoProp; eauto. destruct g eqn:hg => /= //; intros; rtoProp; eauto. repeat split; eauto. destruct cstr_as_blocks; rtoProp; repeat split; len; eauto. 1: solve_all. - - rewrite lookup_env_optimize //. + - rewrite lookup_env_remove_match_on_box //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => /= //. intros; rtoProp; intuition auto; solve_all. - - rewrite lookup_env_optimize //. + - rewrite lookup_env_remove_match_on_box //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. rewrite andb_false_r => //. @@ -900,73 +900,73 @@ Proof. all:intros; rtoProp; intuition auto; solve_all. Qed. -Lemma optimize_wellformed_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} d : +Lemma remove_match_on_box_wellformed_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} d : wf_glob Σ -> - wf_global_decl Σ d -> wf_global_decl (optimize_env Σ) d. + wf_global_decl Σ d -> wf_global_decl (remove_match_on_box_env Σ) d. Proof. intros wf; destruct d => /= //. destruct (cst_body c) => /= //. - now eapply optimize_wellformed_irrel. + now eapply remove_match_on_box_wellformed_irrel. Qed. -Lemma optimize_decl_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : +Lemma remove_match_on_box_decl_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : has_tBox -> has_tRel -> wf_glob Σ -> - forall d, wf_global_decl Σ d -> wf_global_decl (optimize_env Σ) (optimize_decl Σ d). + forall d, wf_global_decl Σ d -> wf_global_decl (remove_match_on_box_env Σ) (remove_match_on_box_decl Σ d). Proof. intros hasb hasr wf d. intros hd. - eapply optimize_wellformed_decl_irrel; tea. + eapply remove_match_on_box_wellformed_decl_irrel; tea. move: hd. destruct d => /= //. destruct (cst_body c) => /= //. - now eapply optimize_wellformed => //. + now eapply remove_match_on_box_wellformed => //. Qed. -Lemma fresh_global_optimize_env {Σ : GlobalContextMap.t} kn : +Lemma fresh_global_remove_match_on_box_env {Σ : GlobalContextMap.t} kn : fresh_global kn Σ -> - fresh_global kn (optimize_env Σ). + fresh_global kn (remove_match_on_box_env Σ). Proof. destruct Σ as [Σ map repr wf]; cbn in *. induction 1; cbn; constructor; auto. now eapply Forall_map; cbn. Qed. -Lemma optimize_env_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : +Lemma remove_match_on_box_env_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : has_tBox -> has_tRel -> - wf_glob Σ -> wf_glob (optimize_env Σ). + wf_glob Σ -> wf_glob (remove_match_on_box_env Σ). Proof. intros hasb hasrel. - intros wfg. rewrite optimize_env_eq //. + intros wfg. rewrite remove_match_on_box_env_eq //. destruct Σ as [Σ map repr wf]; cbn in *. clear map repr. induction wfg; cbn; constructor; auto. - - rewrite /= -(optimize_env_eq (GlobalContextMap.make Σ (fresh_globals_cons_inv wf))) //. - eapply optimize_decl_wf => //. - - rewrite /= -(optimize_env_eq (GlobalContextMap.make Σ (fresh_globals_cons_inv wf))) //. - now eapply fresh_global_optimize_env. + - rewrite /= -(remove_match_on_box_env_eq (GlobalContextMap.make Σ (fresh_globals_cons_inv wf))) //. + eapply remove_match_on_box_decl_wf => //. + - rewrite /= -(remove_match_on_box_env_eq (GlobalContextMap.make Σ (fresh_globals_cons_inv wf))) //. + now eapply fresh_global_remove_match_on_box_env. Qed. -Definition optimize_program (p : eprogram_env) := - (EOptimizePropDiscr.optimize_env p.1, EOptimizePropDiscr.optimize p.1 p.2). +Definition remove_match_on_box_program (p : eprogram_env) := +(remove_match_on_box_env p.1, remove_match_on_box p.1 p.2). -Definition optimize_program_wf {efl} (p : eprogram_env) {hastbox : has_tBox} {hastrel : has_tRel} : - wf_eprogram_env efl p -> wf_eprogram efl (optimize_program p). +Definition remove_match_on_box_program_wf {efl} (p : eprogram_env) {hastbox : has_tBox} {hastrel : has_tRel} : + wf_eprogram_env efl p -> wf_eprogram efl (remove_match_on_box_program p). Proof. intros []; split. - now eapply optimize_env_wf. - cbn. eapply optimize_wellformed_irrel => //. now eapply optimize_wellformed. + now eapply remove_match_on_box_env_wf. + cbn. eapply remove_match_on_box_wellformed_irrel => //. now eapply remove_match_on_box_wellformed. Qed. -Definition optimize_program_expanded {efl} (p : eprogram_env) : +Definition remove_match_on_box_program_expanded {efl} (p : eprogram_env) : wf_eprogram_env efl p -> - expanded_eprogram_env_cstrs p -> expanded_eprogram_cstrs (optimize_program p). + expanded_eprogram_env_cstrs p -> expanded_eprogram_cstrs (remove_match_on_box_program p). Proof. unfold expanded_eprogram_env_cstrs. move=> [wfe wft] /andP[] etae etat. apply/andP; split. - cbn. eapply expanded_global_env_isEtaExp_env, optimize_env_expanded => //. + cbn. eapply expanded_global_env_isEtaExp_env, remove_match_on_box_env_expanded => //. now eapply isEtaExp_env_expanded_global_env. eapply expanded_isEtaExp. - eapply optimize_expanded_irrel => //. - now apply optimize_expanded, isEtaExp_expanded. + eapply remove_match_on_box_expanded_irrel => //. + now apply remove_match_on_box_expanded, isEtaExp_expanded. Qed. diff --git a/erasure/theories/EWcbvEvalNamed.v b/erasure/theories/EWcbvEvalNamed.v index 08e77b7d4..c3578d341 100644 --- a/erasure/theories/EWcbvEvalNamed.v +++ b/erasure/theories/EWcbvEvalNamed.v @@ -78,6 +78,8 @@ Section Wcbv. lookup Γ na = Some v -> eval Γ (tVar na) v + | eval_vbox : eval Γ tBox vBox + | eval_box a t t' : eval Γ a vBox -> eval Γ t t' -> @@ -148,7 +150,10 @@ Section Wcbv. Program Definition eval_ind := λ (P : ∀ (Γ : environment) (t : term) (v : value), eval Γ t v → Type) (f : ∀ (Γ : environment) (na : string) (v : value) (e : lookup Γ na = Some v), - P Γ (tVar na) v (eval_var Γ na v e)) (f0 : ∀ (Γ : environment) (a t : term) + P Γ (tVar na) v (eval_var Γ na v e)) + + (f_vbox : ∀ (Γ : environment) (e : eval Γ tBox vBox), P Γ tBox vBox e) + (f0 : ∀ (Γ : environment) (a t : term) (t' : value) (e : eval Γ a vBox), P Γ a vBox e → ∀ e0 : eval Γ t t', @@ -249,6 +254,7 @@ Section Wcbv. fix F (Γ : environment) (t : term) (v : value) (e : eval Γ t v) {struct e} : P Γ t v e := match e as e0 in (eval _ t0 v0) return (P Γ t0 v0 e0) with | @eval_var _ na v0 e0 => f Γ na v0 e0 + | @eval_vbox _ => f_vbox Γ _ | @eval_box _ a t0 t' e0 e1 => f0 Γ a t0 t' e0 (F Γ a vBox e0) e1 (F Γ t0 t' e1) | @eval_beta _ f10 na b a a' res Γ' e0 e1 e2 => f1 Γ f10 na b a a' res Γ' e0 (F Γ f10 (vClos na b Γ') e0) e1 (F Γ a a' e1) e2 (F (add na a' Γ') b res e2) | @eval_lambda _ na b => f2 Γ na b @@ -963,7 +969,7 @@ Proof. { clear. generalize ((List.rev (gen_many_fresh Γ ( (map dname m))) ++ Γ)). induction m in Γ |- *; cbn. - econstructor. - - intros. destruct a; cbn. destruct dname; cbn; try econstructor; eauto. + - intros. destruct a; cbn. destruct dname; cbn; try econstructor; eauto. } { solve_all. unfold wf_fix in *. rtoProp. solve_all. clear H0. unfold test_def in *. cbn in *. eapply All_impl in H1. 2:{ intros ? [[] ]. @@ -974,7 +980,7 @@ Proof. generalize ((List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)). intros. induction H1 in Γ |- *. - econstructor. - - cbn. destruct x; cbn. destruct dname; cbn; econstructor; eauto. + - cbn. destruct x; cbn. destruct dname; cbn; econstructor; eauto. } Qed. @@ -1226,7 +1232,9 @@ Definition dupfree (Γ : list ident) := Fixpoint sunny Γ (t : term) : bool := match t with - | tRel i => true + | tBox => true + | tRel _ => true + | tVar na => true | tEvar ev args => List.forallb (sunny Γ) args | tLambda (nNamed na) M => fresh na Γ && sunny (na :: Γ) M | tApp u v => sunny Γ u && sunny Γ v @@ -1244,8 +1252,7 @@ Fixpoint sunny Γ (t : term) : bool := dupfree names && forallb (test_def (sunny (names ++ Γ))) mfix | tConstruct _ _ args => forallb (sunny Γ) args - | tConst _ => true - | _ => false + | _ => true end. Inductive wf : value -> Type := @@ -1363,6 +1370,13 @@ Proof. revert HE Hsunny. pattern E, s, v, Heval. revert E s v Heval. eapply eval_ind; intros; eauto; cbn in *; rtoProp; eauto using Forall. + - induction HE. + + cbn in *. congruence. + + unfold lookup in e. cbn in e. destruct x. + destruct eqb eqn:E. + * invs e. cbn in *; eauto. + * eapply IHHE. exact e. + - econstructor. - let X := match reverse goal with H : All _ _ -> _ -> _ |- _ => H end in do 2 forward X; eauto; inv X. let X1 := multimatch goal with H : _ |- _ => H end in @@ -1514,6 +1528,8 @@ Proof. Qed. *) +(*) + Lemma represents_add_fresh nms Γ E vs s t : Γ ;;; E ⊩ s ~ t -> sunny nms s -> @@ -1523,6 +1539,7 @@ Proof. revert Γ E s t. refine (@represents_ind _ (fun _ _ _ => True) _ _ _ _ _ _ _ _ _ _ _ _ _ _). all: intros; cbn in *; rtoProp; eauto. + - econstructor; eauto. - econstructor; eauto. eapply H; eauto. eapply sunny_subset; eauto. firstorder. - econstructor; eauto. eapply H0; eauto. @@ -1549,6 +1566,8 @@ Proof. eapply incl_appr, incl_refl. Qed. +*) + Lemma NoDup_In_1 {A} (l1 l2 : list A) a : NoDup (l1 ++ l2) -> In a l1 -> ~ In a l2. Proof. @@ -1853,6 +1872,8 @@ Proof. unfold cstr_arity in *. invs H. lia. clear - Hvs; induction Hvs; econstructor; eauto. eapply r. - invs Hrep; cbn in *; try congruence; rtoProp. + + econstructor. split; eauto. eapply eval_vbox. + + econstructor. split; eauto. econstructor. eauto. + econstructor. split; eauto. econstructor. + destruct args'; congruence. + solve_all. eexists. split. 2: econstructor; eauto. 4: solve_all. @@ -1868,5 +1889,130 @@ Proof. cbn in *. destruct H1 as (([? []] & ?) & ?). rewrite app_nil_r in r. all: eauto. - Unshelve. all: repeat econstructor. + Unshelve. all: repeat econstructor. Qed. + +Lemma concat_sing {X} l : + List.concat (map (fun x :X => [x]) l) = l. +Proof. + induction l; cbn in *; try congruence. +Qed. + +Lemma sunny_annotate Γ s : + sunny Γ (annotate Γ s). +Proof. + induction s in Γ |- * using EInduction.term_forall_list_ind; cbn; eauto; rtoProp. + - destruct nth_error; cbn; eauto. + - solve_all. + - split; eauto. destruct n; eauto. + unfold fresh. destruct in_dec; cbn; eauto. + eapply gen_fresh_fresh in i; eauto. + destruct in_dec; cbn; eauto. + unfold fresh. destruct in_dec; cbn; eauto. + eapply gen_fresh_fresh in i1; eauto. + unfold fresh. destruct in_dec; cbn; eauto. exfalso; eauto. + - split; eauto. destruct n; eauto. + unfold fresh. destruct in_dec; cbn; eauto. + eapply gen_fresh_fresh in i; eauto. + destruct in_dec; cbn; eauto. + unfold fresh. destruct in_dec; cbn; eauto. + eapply gen_fresh_fresh in i1; eauto. + unfold fresh. destruct in_dec; cbn; eauto. exfalso; eauto. + - split; eauto. + - solve_all. + - split; eauto. solve_all. rtoProp. repeat split; eauto. + + solve_all. eapply In_All. intros. + unfold fresh. destruct in_dec; cbn; eauto. + exfalso. eapply NoDup_gen_many_fresh; eauto. + + unfold dupfree. destruct dupfree_dec_ident; eauto. + exfalso. eapply n. + rewrite flat_map_concat_map. rewrite map_map. + rewrite concat_sing. + eapply NoDup_gen_many_fresh; eauto. + + rewrite flat_map_concat_map map_map concat_sing. + eapply sunny_subset. eapply H. + intros ? ? %in_app_iff. eapply in_app_iff. rewrite <- in_rev in H0. + eauto. + - repeat split; eauto. + 1:{ unfold map_def_name. generalize (annotate (List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)). + cbn. intros. solve_all. + clear. induction m in Γ |- *. + - cbn. econstructor. + - cbn. destruct a; cbn. destruct dname; cbn. + + econstructor; cbn. + unfold fresh. destruct in_dec; eauto; cbn. exfalso. + eapply gen_fresh_fresh; eauto. + specialize (IHm (gen_fresh "wildcard" Γ :: Γ)). + solve_all. destruct x, dname; cbn in *; try congruence. + unfold fresh in *. repeat (destruct in_dec; cbn in *; try congruence). + exfalso. eapply n. eauto. + + econstructor; cbn. + { destruct in_dec; cbn. + + unfold fresh. destruct in_dec; eauto; cbn. exfalso. + eapply gen_fresh_fresh; eauto. + + unfold fresh. destruct in_dec; eauto; cbn. exfalso. eauto. } + specialize (IHm ((if is_left (in_dec (ReflectEq_EqDec IdentOT.reflect_eq_string) i Γ) then gen_fresh i Γ else i) :: Γ)). + solve_all. destruct x, dname; cbn in *; try congruence. + destruct in_dec; cbn in *. + * unfold fresh in *. + repeat (destruct in_dec; cbn in *; try congruence). + exfalso. eapply n. eauto. + * unfold fresh in *. + repeat (destruct in_dec; cbn in *; try congruence). + exfalso. eapply n0. eauto. + } + { unfold map_def_name. generalize (annotate (List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)). + intros t0. + assert ((flat_map (λ d : def term, match dname d with + | nAnon => [] + | nNamed i => [i] + end) +(map2 (λ (d : def term) (na : ident), {| dname := nNamed na; dbody := t0 (dbody d); rarg := rarg d |}) m (gen_many_fresh Γ (map dname m)))) = (gen_many_fresh Γ (map dname m))) as ->. + { pose proof (gen_many_fresh_length Γ (map dname m)). rewrite map_length in H. revert H. + generalize (gen_many_fresh Γ (map dname m)). clear. induction m; destruct l; cbn; try congruence. + intros. f_equal. eapply IHm. lia. } + + unfold dupfree. intros. destruct dupfree_dec_ident; cbn; eauto. + exfalso. apply n0. + eapply NoDup_gen_many_fresh. + } + { + unfold map_def_name. + assert ((flat_map (λ d : def term, match dname d with + | nAnon => [] + | nNamed i => [i] + end) +(map2 (λ (d : def term) (na : ident), {| dname := nNamed na; dbody := annotate (List.rev (gen_many_fresh Γ (map dname m)) ++ Γ) (dbody d); rarg := rarg d |}) m (gen_many_fresh Γ (map dname m)))) = (gen_many_fresh Γ (map dname m))) as ->. + { generalize (annotate (List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)). + pose proof (gen_many_fresh_length Γ (map dname m)). rewrite map_length in H. revert H. + generalize (gen_many_fresh Γ (map dname m)). clear. induction m; destruct l; cbn; try congruence. + intros. f_equal. eapply IHm. lia. } + + solve_all. + eapply (All_impl(P := (λ x : def term, test_def (sunny (List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)) x))). + 2:{ intros x H. unfold test_def. eapply sunny_subset. eassumption. + intros ? ? % in_app_iff. eapply in_app_iff. rewrite <- in_rev. eauto. } + + generalize (List.rev (gen_many_fresh Γ (map dname m))). + pose proof (gen_many_fresh_length Γ (map dname m)). rewrite map_length in H. revert H. + generalize (gen_many_fresh Γ (map dname m)). clear - X. induction X. + + intros l. destruct l; cbn in *; try congruence. econstructor. + + intros []; cbn in *; try congruence. intros. econstructor; eauto. + } +Qed. + +Lemma eval_to_eval_named_full (Σ Σ' : global_context) s t : + wf_glob Σ -> + Forall (fun d => match d.2 with ConstantDecl (Build_constant_body (Some d)) => sunny [] d | _ => true end) Σ' -> + All2 (fun d d' => d.1 = d'.1 × match d.2 with ConstantDecl (Build_constant_body (Some body)) => + ∑ body', d'.2 = ConstantDecl (Build_constant_body (Some body')) × [] ;;; [] ⊩ body' ~ body + | decl => d'.2 = decl + end) Σ Σ' -> + EWcbvEval.eval Σ s t -> + wellformed Σ 0 s -> + ∑ v, ⊩ v ~ t × eval Σ' [] (annotate [] s) v. +Proof. + intros. eapply eval_to_eval_named; eauto. + eapply sunny_annotate. + eapply nclosed_represents. eauto. +Qed. \ No newline at end of file diff --git a/erasure/theories/Typed/Extraction.v b/erasure/theories/Typed/Extraction.v index 07d2524ff..2e7c75a9d 100644 --- a/erasure/theories/Typed/Extraction.v +++ b/erasure/theories/Typed/Extraction.v @@ -89,7 +89,7 @@ Program Definition extract_pcuic_env (ignore : kername -> bool) : result ExAst.global_env _ := let Σ := timed "Erasure" (fun _ => erase_global_decls_deps_recursive (declarations Σ) (universes Σ) (retroknowledge Σ) wfΣ seeds ignore) in if optimize_prop_discr params then - let Σ := timed "Removal of prop discrimination" (fun _ => OptimizePropDiscr.optimize_env Σ _) in + let Σ := timed "Removal of prop discrimination" (fun _ => OptimizePropDiscr.remove_match_on_box_env Σ _) in compose_transforms (extract_transforms params) Σ else Ok Σ. diff --git a/erasure/theories/Typed/OptimizePropDiscr.v b/erasure/theories/Typed/OptimizePropDiscr.v index a2fd1965c..b5036ae65 100644 --- a/erasure/theories/Typed/OptimizePropDiscr.v +++ b/erasure/theories/Typed/OptimizePropDiscr.v @@ -3,13 +3,13 @@ From MetaCoq.Erasure.Typed Require Import ExAst. From MetaCoq.Erasure Require Import EOptimizePropDiscr. -Definition optimize_constant_body Σ cst := +Definition remove_match_on_box_constant_body Σ cst := {| cst_type := cst_type cst; - cst_body := option_map (optimize Σ) (cst_body cst) |}. + cst_body := option_map (remove_match_on_box Σ) (cst_body cst) |}. -Definition optimize_decl Σ d := +Definition remove_match_on_box_decl Σ d := match d with - | ConstantDecl cst => ConstantDecl (optimize_constant_body Σ cst) + | ConstantDecl cst => ConstantDecl (remove_match_on_box_constant_body Σ cst) | _ => d end. @@ -36,21 +36,21 @@ Proof. now apply trans_env_fresh_global. Qed. -Program Definition optimize_env (Σ : global_env) (fgΣ : fresh_globals Σ) : global_env := - List.map (MCProd.on_snd (optimize_decl (EEnvMap.GlobalContextMap.make (trans_env Σ) _))) Σ. +Program Definition remove_match_on_box_env (Σ : global_env) (fgΣ : fresh_globals Σ) : global_env := + List.map (MCProd.on_snd (remove_match_on_box_decl (EEnvMap.GlobalContextMap.make (trans_env Σ) _))) Σ. Next Obligation. now apply trans_env_fresh_globals. Qed. Module Ee := EWcbvEval. -Lemma trans_env_optimize_env Σ fgΣ : - trans_env (optimize_env Σ fgΣ) = - EOptimizePropDiscr.optimize_env (EEnvMap.GlobalContextMap.make (trans_env Σ) (trans_env_fresh_globals _ fgΣ)). +Lemma trans_env_remove_match_on_box_env Σ fgΣ : + trans_env (remove_match_on_box_env Σ fgΣ) = + EOptimizePropDiscr.remove_match_on_box_env (EEnvMap.GlobalContextMap.make (trans_env Σ) (trans_env_fresh_globals _ fgΣ)). Proof. unfold trans_env. - unfold EOptimizePropDiscr.optimize_env. - unfold optimize_env. + unfold EOptimizePropDiscr.remove_match_on_box_env. + unfold remove_match_on_box_env. unfold MCProd.on_snd. cbn. rewrite !List.map_map. apply List.map_ext. @@ -62,7 +62,7 @@ Proof. now destruct o. Qed. -Lemma optimize_correct `{EWellformed.EEnvFlags} `{Ee.WcbvFlags} Σ fgΣ t v : +Lemma remove_match_on_box_correct `{EWellformed.EEnvFlags} `{Ee.WcbvFlags} Σ fgΣ t v : EWcbvEval.with_constructor_as_block = false -> ELiftSubst.closed t = true -> EGlobalEnv.closed_env (trans_env Σ) = true -> @@ -70,12 +70,12 @@ Lemma optimize_correct `{EWellformed.EEnvFlags} `{Ee.WcbvFlags} Σ fgΣ t v : @Prelim.Ee.eval _ (trans_env Σ) t v -> @Prelim.Ee.eval (EWcbvEval.disable_prop_cases _) - (trans_env (optimize_env Σ fgΣ)) - (optimize (EEnvMap.GlobalContextMap.make (trans_env Σ) (trans_env_fresh_globals _ fgΣ)) t) - (optimize (EEnvMap.GlobalContextMap.make (trans_env Σ) (trans_env_fresh_globals _ fgΣ)) v). + (trans_env (remove_match_on_box_env Σ fgΣ)) + (remove_match_on_box (EEnvMap.GlobalContextMap.make (trans_env Σ) (trans_env_fresh_globals _ fgΣ)) t) + (remove_match_on_box (EEnvMap.GlobalContextMap.make (trans_env Σ) (trans_env_fresh_globals _ fgΣ)) v). Proof. intros ? cl_t cl_env wfg ev. - rewrite trans_env_optimize_env. + rewrite trans_env_remove_match_on_box_env. remember (EEnvMap.GlobalContextMap.make _ _) as Σ0. - unshelve eapply EOptimizePropDiscr.optimize_correct;subst;cbn;eauto. + unshelve eapply EOptimizePropDiscr.remove_match_on_box_correct;subst;cbn;eauto. Qed. diff --git a/erasure/theories/Typed/TypeAnnotations.v b/erasure/theories/Typed/TypeAnnotations.v index 3815a4a8e..e6b195cd1 100644 --- a/erasure/theories/Typed/TypeAnnotations.v +++ b/erasure/theories/Typed/TypeAnnotations.v @@ -699,7 +699,7 @@ Module AnnotOptimizePropDiscr. intros ? a'; exact (f _ _ a'). Defined. - Definition annot_optimize Σ {t} (a : annots box_type t) : annots box_type (optimize Σ t). + Definition annot_remove_match_on_box Σ {t} (a : annots box_type t) : annots box_type (remove_match_on_box Σ t). Proof. revert t a. fix f 1. @@ -708,7 +708,7 @@ Module AnnotOptimizePropDiscr. - exact (a.1, f _ a.2). - exact (a.1, (f _ a.2.1, f _ a.2.2)). - exact (a.1, (f _ a.2.1, f _ a.2.2)). - - assert (br_annots : bigprod (fun br => annots box_type br.2) (map (on_snd (optimize Σ)) l)). + - assert (br_annots : bigprod (fun br => annots box_type br.2) (map (on_snd (remove_match_on_box Σ)) l)). { refine (bigprod_map _ a.2.2). intros ? a'; apply (f _ a'). } unfold isprop_ind. @@ -743,32 +743,32 @@ Module AnnotOptimizePropDiscr. intros ? a'; exact (f _ a'). Defined. - Definition annot_optimize_constant_body Σ {cst} (a : constant_body_annots box_type cst) : - constant_body_annots box_type (optimize_constant_body Σ cst). + Definition annot_remove_match_on_box_constant_body Σ {cst} (a : constant_body_annots box_type cst) : + constant_body_annots box_type (remove_match_on_box_constant_body Σ cst). Proof. - unfold constant_body_annots, optimize_constant_body in *. + unfold constant_body_annots, remove_match_on_box_constant_body in *. cbn. destruct ExAst.cst_body; cbn; [|exact tt]. - apply annot_optimize. + apply annot_remove_match_on_box. exact a. Defined. - Definition annot_optimize_decl Σ {decl} (a : global_decl_annots box_type decl) : - global_decl_annots box_type (optimize_decl Σ decl). + Definition annot_remove_match_on_box_decl Σ {decl} (a : global_decl_annots box_type decl) : + global_decl_annots box_type (remove_match_on_box_decl Σ decl). Proof. - unfold global_decl_annots, optimize_decl in *. + unfold global_decl_annots, remove_match_on_box_decl in *. destruct decl; [|exact tt|exact tt]. - apply annot_optimize_constant_body. + apply annot_remove_match_on_box_constant_body. exact a. Defined. - Definition annot_optimize_env Σ fgΣ (a : env_annots box_type Σ) : - env_annots box_type (optimize_env Σ fgΣ). + Definition annot_remove_match_on_box_env Σ fgΣ (a : env_annots box_type Σ) : + env_annots box_type (remove_match_on_box_env Σ fgΣ). Proof. - unfold env_annots, optimize_env. + unfold env_annots, remove_match_on_box_env. apply bigprod_map. - intros. - apply annot_optimize_decl. + apply annot_remove_match_on_box_decl. exact X. - exact a. Defined. @@ -803,7 +803,7 @@ Proof. unfold extract_pcuic_env. destruct optimize_prop_discr. + apply annot_compose_transforms; [|exact all]. - apply AnnotOptimizePropDiscr.annot_optimize_env. + apply AnnotOptimizePropDiscr.annot_remove_match_on_box_env. apply annotate_types_erase_global_decls_deps_recursive. + apply annotate_types_erase_global_decls_deps_recursive. Defined. From a813bffeff2c7167a096eb78e38a6eb5fb43150b Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Thu, 6 Jul 2023 15:08:46 +0200 Subject: [PATCH 30/61] Add pass implementing boxes as fixpoints --- erasure-plugin/theories/ETransform.v | 23 ++ erasure/_CoqProject.in | 1 + erasure/theories/EImplementBox.v | 584 +++++++++++++++++++++++++++ 3 files changed, 608 insertions(+) create mode 100644 erasure/theories/EImplementBox.v diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index 521481293..7ca17ed50 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -431,3 +431,26 @@ Next Obligation. cbn -[transform_blocks]. eapply transform_blocks_eval; cbn; eauto. Qed. + +From MetaCoq.Erasure Require Import EImplementBox. + +Program Definition implement_box_transformation (efl : EEnvFlags) + {has_app : has_tApp} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = true} : + Transform.t eprogram eprogram EAst.term EAst.term (eval_eprogram block_wcbv_flags) (eval_eprogram block_wcbv_flags) := + {| name := "transforming to constuctors as blocks"; + transform p _ := EImplementBox.implement_box_program p ; + pre p := wf_eprogram efl p ; + post p := wf_eprogram (switch_off_box efl) p ; + obseq g g' v v' := v' = implement_box v |}. + +Next Obligation. + intros. cbn in *. destruct p. split. + - eapply transform_wf_global; eauto. + - eapply transform_wellformed; eauto. +Qed. +Next Obligation. + red. intros. destruct pr. destruct H. + eexists. split; [ | eauto]. + econstructor. + eapply implement_box_eval; cbn; eauto. +Qed. \ No newline at end of file diff --git a/erasure/_CoqProject.in b/erasure/_CoqProject.in index ed4879dcb..ed2993839 100644 --- a/erasure/_CoqProject.in +++ b/erasure/_CoqProject.in @@ -35,6 +35,7 @@ theories/EInlineProjections.v theories/EConstructorsAsBlocks.v theories/EWcbvEvalCstrsAsBlocksInd.v theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v +theories/EImplementBox.v theories/Typed/Annotations.v theories/Typed/ClosedAux.v diff --git a/erasure/theories/EImplementBox.v b/erasure/theories/EImplementBox.v new file mode 100644 index 000000000..3d854bbb3 --- /dev/null +++ b/erasure/theories/EImplementBox.v @@ -0,0 +1,584 @@ +(* Distributed under the terms of the MIT license. *) +From Coq Require Import Utf8 Program. +From MetaCoq.Utils Require Import utils. +From MetaCoq.Common Require Import config Kernames BasicAst EnvMap. +From MetaCoq.Erasure Require Import EAst EAstUtils EInduction EArities + ELiftSubst ESpineView EGlobalEnv EWellformed EEnvMap + EWcbvEval EEtaExpanded ECSubst EWcbvEvalEtaInd EProgram. + +Local Open Scope string_scope. +Set Asymmetric Patterns. +Import MCMonadNotation. + +From Equations Require Import Equations. +Set Equations Transparent. +Local Set Keyed Unification. +From Coq Require Import ssreflect ssrbool. + +(** We assume [Prop size x < size y) := + | tRel i => EAst.tRel i + | tEvar ev args => EAst.tEvar ev (map_InP args (fun x H => implement_box x)) + | tLambda na M => EAst.tLambda na (implement_box M) + | tApp u v := tApp (implement_box u) (implement_box v) + | tLetIn na b b' => EAst.tLetIn na (implement_box b) (implement_box b') + | tCase ind c brs => + let brs' := map_InP brs (fun x H => (x.1, implement_box x.2)) in + EAst.tCase (ind.1, 0) (implement_box c) brs' + | tProj p c => EAst.tProj {| proj_ind := p.(proj_ind); proj_npars := 0; proj_arg := p.(proj_arg) |} (implement_box c) + | tFix mfix idx => + let mfix' := map_InP mfix (fun d H => {| dname := dname d; dbody := implement_box d.(dbody); rarg := d.(rarg) |}) in + EAst.tFix mfix' idx + | tCoFix mfix idx => + let mfix' := map_InP mfix (fun d H => {| dname := dname d; dbody := implement_box d.(dbody); rarg := d.(rarg) |}) in + EAst.tCoFix mfix' idx + | tBox => iBox + | tVar n => EAst.tVar n + | tConst n => EAst.tConst n + | tConstruct ind i block_args => EAst.tConstruct ind i (map_InP block_args (fun d H => implement_box d)) + | tPrim p => EAst.tPrim p. + Proof. + all:try lia. + all:try apply (In_size); tea. + all:try lia. + - now apply (In_size id size). + - now eapply (In_size id size). + - eapply (In_size snd size) in H. cbn in *. lia. + Qed. + + End Def. + + Hint Rewrite @map_InP_spec : implement_box. + + Arguments eqb : simpl never. + + Opaque implement_box. + Opaque isEtaExp. + Opaque isEtaExp_unfold_clause_1. + + Lemma chop_firstn_skipn {A} n (l : list A) : chop n l = (firstn n l, skipn n l). + Proof using Type. + induction n in l |- *; destruct l; simpl; auto. + now rewrite IHn skipn_S. + Qed. + + Lemma chop_eq {A} n (l : list A) l1 l2 : chop n l = (l1, l2) -> l = l1 ++ l2. + Proof using Type. + rewrite chop_firstn_skipn. intros [= <- <-]. + now rewrite firstn_skipn. + Qed. + + Lemma closed_implement_box t k : closedn k t -> closedn k (implement_box t). + Proof using Type. + funelim (implement_box t); simp implement_box; rewrite <-?implement_box_equation_1; toAll; simpl; + intros; try easy; + rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + unfold test_def in *; + simpl closed in *; + try solve [simpl; subst; simpl closed; f_equal; auto; rtoProp; solve_all; solve_all]; try easy. + Qed. + + Hint Rewrite @forallb_InP_spec : isEtaExp. + Transparent isEtaExp_unfold_clause_1. + + Lemma implement_box_csubst a k b : + implement_box (ECSubst.csubst a k b) = ECSubst.csubst (implement_box a) k (implement_box b). + Proof using Type. + revert a k. + funelim (implement_box b); intros a k; cbn; simp implement_box; try easy. + - destruct Nat.compare => //. + - f_equal. rewrite !map_map_compose. solve_all. + eapply In_All. eauto. + - cbn. f_equal. rewrite !map_map. solve_all. + eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. rewrite H; eauto. + - cbn. rewrite H0. f_equal. rewrite !map_map. solve_all. + eapply In_All. cbn. intros ? ?. f_equal. rewrite H; eauto. + - cbn. f_equal. rewrite !map_map. solve_all. + eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. rewrite H; eauto. + f_equal. now rewrite map_length. + - cbn. f_equal. rewrite !map_map. solve_all. + eapply In_All. intros ? ?. unfold map_def. cbn. f_equal. rewrite H; eauto. + f_equal. now rewrite map_length. + Qed. + + Lemma implement_box_substl s t : + implement_box (substl s t) = substl (map implement_box s) (implement_box t). + Proof using Type. + induction s in t |- *; simpl; auto. + rewrite IHs //. f_equal. + now rewrite implement_box_csubst. + Qed. + + Lemma implement_box_iota_red pars args br : + implement_box (EGlobalEnv.iota_red pars args br) = EGlobalEnv.iota_red pars (map implement_box args) (on_snd implement_box br). + Proof using Type. + unfold EGlobalEnv.iota_red. + rewrite implement_box_substl //. + now rewrite map_rev map_skipn. + Qed. + + Lemma implement_box_fix_subst mfix : EGlobalEnv.fix_subst (map (map_def implement_box) mfix) = map implement_box (EGlobalEnv.fix_subst mfix). + Proof using Type. + unfold EGlobalEnv.fix_subst. + rewrite map_length. + generalize #|mfix|. + induction n; simpl; auto. + f_equal; auto. now simp implement_box. + Qed. + + Lemma implement_box_cofix_subst mfix : EGlobalEnv.cofix_subst (map (map_def implement_box) mfix) = map implement_box (EGlobalEnv.cofix_subst mfix). + Proof using Type. + unfold EGlobalEnv.cofix_subst. + rewrite map_length. + generalize #|mfix|. + induction n; simpl; auto. + f_equal; auto. now simp implement_box. + Qed. + + Lemma implement_box_cunfold_fix mfix idx n f : + cunfold_fix mfix idx = Some (n, f) -> + cunfold_fix (map (map_def implement_box) mfix) idx = Some (n, implement_box f). + Proof using Type. + unfold cunfold_fix. + rewrite nth_error_map. + destruct nth_error eqn:heq. + intros [= <- <-] => /=. f_equal. f_equal. + rewrite implement_box_substl //. 2:congruence. + f_equal. f_equal. apply implement_box_fix_subst. + Qed. + + Lemma implement_box_cunfold_cofix mfix idx n f : + cunfold_cofix mfix idx = Some (n, f) -> + cunfold_cofix (map (map_def implement_box) mfix) idx = Some (n, implement_box f). + Proof using Type. + unfold cunfold_cofix. + rewrite nth_error_map. + destruct nth_error eqn:heq. + intros [= <- <-] => /=. f_equal. + rewrite implement_box_substl //. 2:congruence. + f_equal. f_equal. apply implement_box_cofix_subst. + Qed. + + Lemma implement_box_nth {n l d} : + implement_box (nth n l d) = nth n (map implement_box l) (implement_box d). + Proof using Type. + induction l in n |- *; destruct n; simpl; auto. + Qed. + +End implement_box. + +Definition implement_box_constant_decl cb := + {| cst_body := option_map implement_box cb.(cst_body) |}. + +Definition implement_box_decl d := + match d with + | ConstantDecl cb => ConstantDecl (implement_box_constant_decl cb) + | InductiveDecl idecl => d + end. + +Definition implement_box_env (Σ : global_declarations) := + map (on_snd (implement_box_decl)) Σ. + +Definition implement_box_program (p : eprogram) := + (implement_box_env p.1, implement_box p.2). + +Definition block_wcbv_flags := + {| with_prop_case := false ; with_guarded_fix := false ; with_constructor_as_block := true |}. + +Local Hint Resolve wellformed_closed : core. + +Lemma wellformed_lookup_inductive_pars {efl : EEnvFlags} Σ kn mdecl : + has_cstr_params = false -> + wf_glob Σ -> + lookup_minductive Σ kn = Some mdecl -> mdecl.(ind_npars) = 0. +Proof. + intros hasp. + induction 1; cbn => //. + case: eqb_spec => [|]. + - intros ->. destruct d => //. intros [= <-]. + cbn in H0. unfold wf_minductive in H0. + rtoProp. cbn in H0. rewrite hasp in H0; now eapply eqb_eq in H0. + - intros _. eapply IHwf_glob. +Qed. + +Lemma wellformed_lookup_constructor_pars {efl : EEnvFlags} {Σ kn c mdecl idecl cdecl} : + has_cstr_params = false -> + wf_glob Σ -> + lookup_constructor Σ kn c = Some (mdecl, idecl, cdecl) -> mdecl.(ind_npars) = 0. +Proof. + intros hasp wf. cbn -[lookup_minductive]. + destruct lookup_minductive eqn:hl => //. + do 2 destruct nth_error => //. + eapply wellformed_lookup_inductive_pars in hl => //. congruence. +Qed. + +Lemma lookup_constructor_pars_args_spec {efl : EEnvFlags} {Σ ind n mdecl idecl cdecl} : + wf_glob Σ -> + lookup_constructor Σ ind n = Some (mdecl, idecl, cdecl) -> + lookup_constructor_pars_args Σ ind n = Some (mdecl.(ind_npars), cdecl.(cstr_nargs)). +Proof. + cbn -[lookup_constructor] => wfΣ. + destruct lookup_constructor as [[[mdecl' idecl'] [pars args]]|] eqn:hl => //. + intros [= -> -> <-]. cbn. f_equal. +Qed. + +Lemma wellformed_lookup_constructor_pars_args {efl : EEnvFlags} {Σ ind k n block_args} : + wf_glob Σ -> + has_cstr_params = false -> + wellformed Σ k (EAst.tConstruct ind n block_args) -> + ∑ args, lookup_constructor_pars_args Σ ind n = Some (0, args). +Proof. + intros wfΣ hasp wf. cbn -[lookup_constructor] in wf. + destruct lookup_constructor as [[[mdecl idecl] cdecl]|] eqn:hl => //. + exists cdecl.(cstr_nargs). + pose proof (wellformed_lookup_constructor_pars hasp wfΣ hl). + eapply lookup_constructor_pars_args_spec in hl => //. congruence. + destruct has_tConstruct => //. +Qed. + +Lemma constructor_isprop_pars_decl_params {efl : EEnvFlags} {Σ ind c b pars cdecl} : + has_cstr_params = false -> wf_glob Σ -> + constructor_isprop_pars_decl Σ ind c = Some (b, pars, cdecl) -> pars = 0. +Proof. + intros hasp hwf. + rewrite /constructor_isprop_pars_decl /lookup_constructor /lookup_inductive. + destruct lookup_minductive as [mdecl|] eqn:hl => /= //. + do 2 destruct nth_error => //. + eapply wellformed_lookup_inductive_pars in hl => //. congruence. +Qed. + +Lemma skipn_ge m {A} (l : list A) : + m >= length l -> skipn m l = []. +Proof. + induction l in m |- *. + - destruct m; reflexivity. + - cbn. destruct m; try lia. intros H. + eapply IHl. lia. +Qed. + +Lemma chop_all {A} (l : list A) m : + m >= length l -> chop m l = (l, nil). +Proof. + intros Hl. rewrite chop_firstn_skipn. + rewrite firstn_ge; try lia. rewrite skipn_ge; try lia. + eauto. +Qed. + +Lemma implement_box_isConstructApp {efl : EEnvFlags} {Σ : global_declarations} t : + isConstructApp (implement_box t) = isConstructApp t. +Proof. + induction t; try now cbn; eauto. + simp implement_box. + rewrite (isConstructApp_mkApps _ [t2]). + rewrite (isConstructApp_mkApps _ [_]). eauto. +Qed. + +Lemma implement_box_isPrimApp {efl : EEnvFlags} {Σ : global_declarations} t : + isPrimApp (implement_box t) = isPrimApp t. +Proof. + induction t; try now cbn; eauto. + simp implement_box. + rewrite (isPrimApp_mkApps _ [t2]). + rewrite (isPrimApp_mkApps _ [_]). eauto. +Qed. + +Lemma lookup_env_implement_box {Σ : global_declarations} kn : + lookup_env (implement_box_env Σ) kn = + option_map (implement_box_decl) (lookup_env Σ kn). +Proof. + unfold implement_box_env. + induction Σ at 1 2; simpl; auto. + case: eqb_spec => //. +Qed. + +Lemma implement_box_declared_constant {Σ : global_declarations} c decl : + declared_constant Σ c decl -> + declared_constant (implement_box_env Σ) c (implement_box_constant_decl decl). +Proof. + intros H. red in H; red. + rewrite lookup_env_implement_box H //. +Qed. + +Lemma lookup_constructor_implement_box Σ ind c : + lookup_constructor (implement_box_env Σ) ind c = + lookup_constructor Σ ind c. +Proof. + unfold lookup_constructor, lookup_inductive, lookup_minductive in *. + rewrite lookup_env_implement_box. + destruct lookup_env as [ [] | ]; cbn; congruence. +Qed. + +Lemma isLambda_implement_box c : isLambda c -> isLambda (implement_box c). +Proof. destruct c => //. Qed. + +Definition disable_box_term_flags (et : ETermFlags) := + {| has_tBox := false + ; has_tRel := true + ; has_tVar := has_tVar + ; has_tEvar := has_tEvar + ; has_tLambda := true + ; has_tLetIn := has_tLetIn + ; has_tApp := has_tApp + ; has_tConst := has_tConst + ; has_tConstruct := has_tConstruct + ; has_tCase := has_tCase + ; has_tProj := has_tProj + ; has_tFix := true + ; has_tCoFix := has_tCoFix + ; has_tPrim := has_tPrim + |}. + +Definition switch_off_box (efl : EEnvFlags) := + {| has_axioms := efl.(has_axioms) ; has_cstr_params := efl.(has_cstr_params) ; term_switches := disable_box_term_flags efl.(term_switches) ; cstr_as_blocks := efl.(cstr_as_blocks) |}. + +Lemma transform_wellformed' {efl : EEnvFlags} {Σ : list (kername × global_decl)} n t : + has_tApp -> + wf_glob Σ -> + @wellformed efl Σ n t -> + @wellformed (switch_off_box efl) Σ n (implement_box t). +Proof. + intros hasa. + revert n. funelim (implement_box t); simp_eta; cbn -[implement_box + lookup_inductive lookup_constructor lookup_constructor_pars_args + GlobalContextMap.lookup_constructor_pars_args isEtaExp]; intros m Hwf Hw; rtoProp; try split; eauto. + all: rewrite ?map_InP_spec; toAll; eauto; try now solve_all. + - unfold lookup_constructor_pars_args in *. + destruct (lookup_constructor Σ) as [[[]] | ]; try congruence; cbn - [implement_box]. + all: destruct cstr_as_blocks; cbn in *; rtoProp; try split; eauto. + + now rewrite map_length. + + solve_all. + + destruct block_args; cbn in *; eauto. + - repeat split; eauto. solve_all. + - unfold wf_fix in *. rtoProp. solve_all. solve_all. now eapply isLambda_implement_box. + - unfold wf_fix in *. rtoProp. solve_all. + len. solve_all. len. destruct x. + cbn -[implement_box isEtaExp] in *. rtoProp. eauto. + - unfold wf_fix in *. len. solve_all. rtoProp; intuition auto. + solve_all. +Qed. + +Lemma transform_wellformed_decl' {efl : EEnvFlags} {Σ : global_declarations} d : + has_tApp -> + wf_glob Σ -> + @wf_global_decl efl Σ d -> + @wf_global_decl (switch_off_box efl) Σ (implement_box_decl d). +Proof. + intros. + destruct d => //=. cbn. + destruct c as [[]] => //=. + eapply transform_wellformed'; tea. +Qed. + +From MetaCoq.Erasure Require Import EGenericMapEnv. + +Lemma transform_wellformed {efl : EEnvFlags} {Σ : global_declarations} n t : + has_tApp -> + wf_glob Σ -> + @wellformed efl Σ n t -> + @wellformed (switch_off_box efl) (implement_box_env Σ) n (implement_box t). +Proof. + intros. eapply gen_transform_wellformed_irrel; eauto. + eapply transform_wellformed'; eauto. +Qed. + +Lemma optimize_decl_wf {efl : EEnvFlags} {Σ : global_declarations} : + has_tApp -> + wf_glob (efl := efl) Σ -> + forall d, + wf_global_decl (efl := efl) Σ d -> + wf_global_decl (efl := switch_off_box efl) (implement_box_env Σ) (implement_box_decl d). +Proof. + intros. + destruct d => /= //. + rewrite /isEtaExp_constant_decl. cbn in H1. + destruct (cst_body c) => /= //. + eapply transform_wellformed => //. +Qed. + +Lemma fresh_global_optimize_env {Σ : global_declarations} kn : + fresh_global kn Σ -> + fresh_global kn (implement_box_env Σ). +Proof. + induction 1; cbn; constructor; auto. +Qed. + +Lemma fresh_global_map_on_snd Σ f kn : + fresh_global kn Σ -> + fresh_global kn (map (on_snd f) Σ). +Proof. + induction 1; cbn; constructor; auto. +Qed. + +Lemma transform_wf_global {efl : EEnvFlags} {Σ : global_declarations} : + has_tApp -> + wf_glob (efl := efl) Σ -> wf_glob (efl := switch_off_box efl) (implement_box_env Σ). +Proof. + intros hasapp wfg. clear - hasapp wfg. + assert (extends Σ Σ). now exists []. + revert H wfg. generalize Σ at 2 3. + induction Σ; cbn; constructor; auto. + - eapply IHΣ; rtoProp; intuition eauto. + destruct H. subst Σ0. exists (x ++ [a]). now rewrite -app_assoc. + - epose proof (EExtends.extends_wf_glob _ H wfg); tea. + depelim H0. cbn. + eapply gen_transform_wellformed_decl_irrel; trea. + destruct d; cbn in *; eauto. + destruct (cst_body c); eauto. + cbn -[implement_box]. cbn in H1. + eapply transform_wellformed'; eauto. + - eapply fresh_global_map_on_snd. + eapply EExtends.extends_wf_glob in wfg; tea. now depelim wfg. +Qed. + +Transparent implement_box. + +Lemma fst_decompose_app_rec t l : fst (decompose_app_rec t l) = fst (decompose_app t). +Proof. + induction t in l |- *; simpl; auto. rewrite IHt1. + unfold decompose_app. simpl. now rewrite (IHt1 [t2]). +Qed. + +Lemma head_tapp t1 t2 : head (tApp t1 t2) = head t1. +Proof. rewrite /head /decompose_app /= fst_decompose_app_rec //. Qed. +Lemma tApp_mkApps f a : tApp f a = mkApps f [a]. +Proof. reflexivity. Qed. + +Require Import EWcbvEvalCstrsAsBlocksInd. + +Lemma implement_box_mkApps f args : + implement_box (mkApps f args) = mkApps (implement_box f) (map implement_box args). +Proof. + induction args in f |- *; simp implement_box. + - reflexivity. + - rewrite IHargs. now simp implement_box. +Qed. + +Lemma implement_box_eval {efl : EEnvFlags} (fl := block_wcbv_flags) : + with_constructor_as_block = true -> + has_cstr_params = false -> + has_tApp -> + forall (Σ : global_declarations), @wf_glob efl Σ -> + forall t t', + @wellformed efl Σ 0 t -> + EWcbvEval.eval Σ t t' -> + @EWcbvEval.eval block_wcbv_flags (implement_box_env Σ) (implement_box t) (implement_box t'). +Proof. + intros cstrbl prms hasapp Σ wfΣ. + eapply + (EWcbvEvalCstrsAsBlocksInd.eval_preserve_mkApps_ind fl cstrbl (efl := efl) Σ _ + (wellformed Σ) (Qpres := Qpreserves_wellformed efl _ wfΣ)) => //. + { intros. eapply EWcbvEval.eval_wellformed => //; tea. } + all:intros *. + - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. + simp implement_box in *. + eapply eval_fix' => //. + + eauto. + + cbn. reflexivity. + + eauto. + + econstructor. econstructor => //. + eapply value_final. eapply eval_to_value; eauto. + unfold iBox. cbn -[implement_box]. unfold map_def. cbn -[implement_box]. + econstructor. cbn. eauto. + - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. + simp implement_box in *. + econstructor; eauto. + now rewrite -implement_box_csubst. + - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. + simp implement_box in *. + econstructor; eauto. + now rewrite -implement_box_csubst. + - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. + simp implement_box in *. + assert (pars = 0) as -> by now (eapply constructor_isprop_pars_decl_params; eauto). + pose proof (Hcon := H1). + rewrite /constructor_isprop_pars_decl in H1. + destruct lookup_constructor as [[[]] | ] eqn:eqc; cbn in H1; invs H1. + eapply eval_iota_block => //. + + cbn [fst]. eapply e0. + + unfold constructor_isprop_pars_decl. + rewrite lookup_constructor_implement_box. cbn [fst]. + rewrite eqc //= H7 //. + + now rewrite map_InP_spec nth_error_map H2; eauto. + + rewrite map_InP_spec. len. + + rewrite H8. rewrite map_InP_spec. len. + + rewrite map_InP_spec. rewrite -implement_box_iota_red. now rewrite H8. + - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. + simp implement_box in *. + eapply eval_fix' => //; eauto. rewrite map_InP_spec. + now eapply implement_box_cunfold_fix. + - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. + simp implement_box in *. + eapply eval_cofix_case. + 2: now eapply implement_box_cunfold_cofix. + { rewrite implement_box_mkApps in e0. + simp implement_box in e0. rewrite map_InP_spec in e0. + cbn -[implement_box] in e0. + exact e0. + } + rewrite implement_box_mkApps in e. + simp implement_box in e. + - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. + simp implement_box in *. + eapply eval_cofix_proj. + 2: now eapply implement_box_cunfold_cofix. { rewrite implement_box_mkApps in e0. + simp implement_box in e0. rewrite map_InP_spec in e0. + cbn -[implement_box] in e0. + exact e0. + } + rewrite implement_box_mkApps in e. + simp implement_box in e. + - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. + simp implement_box in *. + econstructor. + eapply implement_box_declared_constant; eauto. + destruct decl. cbn in *. now rewrite H0. + eauto. + - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. + simp implement_box in *. rewrite map_InP_spec in e0. + unfold constructor_isprop_pars_decl in *. + destruct lookup_constructor as [[[mdecl idecl] cdecl']|] eqn:hc => //. + noconf H2. + pose proof (lookup_constructor_pars_args_cstr_arity _ _ _ _ _ _ hc). + assert (ind_npars mdecl = 0). + { eapply wellformed_lookup_constructor_pars; tea. } + eapply eval_proj_block => //; tea. cbn. + + unfold constructor_isprop_pars_decl. + rewrite lookup_constructor_implement_box. cbn [fst]. + rewrite hc //= H1 H7. reflexivity. + + len. + + rewrite nth_error_map /=. rewrite H7 in H2; rewrite -H2 in H4; rewrite H4; eauto. + - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. + simp implement_box in *. + eapply eval_app_cong; eauto. + revert H1. + destruct f'; try now cbn; tauto. + intros H. cbn in H. + rewrite implement_box_isConstructApp; eauto. + rewrite implement_box_isPrimApp; eauto. + - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. + simp implement_box in *. rewrite !map_InP_spec. + eapply eval_construct_block; tea. eauto. + 2: len; eassumption. + rewrite lookup_constructor_implement_box; eauto. + eapply All2_All2_Set. + solve_all. now destruct b. + - intros. destruct t; try solve [constructor; cbn in H, H0 |- *; try congruence]. + cbn -[lookup_constructor] in H |- *. destruct l => //. +Qed. From 5bc51f5cb9f708ea3a6c8b6fcc1abe8120beed9a Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Thu, 6 Jul 2023 15:14:02 +0200 Subject: [PATCH 31/61] adapt name pass to not consider box --- erasure/theories/EWcbvEvalNamed.v | 56 ++++++------------------------- 1 file changed, 11 insertions(+), 45 deletions(-) diff --git a/erasure/theories/EWcbvEvalNamed.v b/erasure/theories/EWcbvEvalNamed.v index c3578d341..bab199936 100644 --- a/erasure/theories/EWcbvEvalNamed.v +++ b/erasure/theories/EWcbvEvalNamed.v @@ -28,7 +28,6 @@ Local Ltac inv H := inversion H; subst. (** ** Big step *named* version of weak cbv beta-zeta-iota-fix-delta reduction. *) Inductive value : Set := -| vBox | vClos (na : ident) (b : term) (env : list (ident * value)) | vConstruct (ind : inductive) (c : nat) (args : list (value)) | vRecClos (b : list (ident * term)) (idx : nat) (env : list (ident * value)). @@ -78,13 +77,6 @@ Section Wcbv. lookup Γ na = Some v -> eval Γ (tVar na) v - | eval_vbox : eval Γ tBox vBox - - | eval_box a t t' : - eval Γ a vBox -> - eval Γ t t' -> - eval Γ (tApp a t) vBox - (** Beta *) | eval_beta f na b a a' res Γ' : eval Γ f (vClos na b Γ') -> @@ -152,14 +144,6 @@ Section Wcbv. λ (P : ∀ (Γ : environment) (t : term) (v : value), eval Γ t v → Type) (f : ∀ (Γ : environment) (na : string) (v : value) (e : lookup Γ na = Some v), P Γ (tVar na) v (eval_var Γ na v e)) - (f_vbox : ∀ (Γ : environment) (e : eval Γ tBox vBox), P Γ tBox vBox e) - (f0 : ∀ (Γ : environment) (a t : term) - (t' : value) (e : eval Γ a vBox), - P Γ a vBox e - → ∀ e0 : eval Γ t t', - P Γ t t' e0 - → P Γ (tApp a t) vBox - (eval_box Γ a t t' e e0)) (f1 : ∀ (Γ : environment) (f1 : term) (na : ident) (b a : term) (a' res : value) (Γ' : list (ident × value)) (e : eval Γ f1 (vClos na b Γ')), P Γ f1 (vClos na b Γ') e → ∀ e0 : eval Γ a a', P Γ a a' e0 → ∀ e1 : eval (add na a' Γ') b res, P (add na a' Γ') b res e1 → P Γ (tApp f1 a) res (eval_beta Γ f1 na b a a' res Γ' e e0 e1)) @@ -254,8 +238,6 @@ Section Wcbv. fix F (Γ : environment) (t : term) (v : value) (e : eval Γ t v) {struct e} : P Γ t v e := match e as e0 in (eval _ t0 v0) return (P Γ t0 v0 e0) with | @eval_var _ na v0 e0 => f Γ na v0 e0 - | @eval_vbox _ => f_vbox Γ _ - | @eval_box _ a t0 t' e0 e1 => f0 Γ a t0 t' e0 (F Γ a vBox e0) e1 (F Γ t0 t' e1) | @eval_beta _ f10 na b a a' res Γ' e0 e1 e2 => f1 Γ f10 na b a a' res Γ' e0 (F Γ f10 (vClos na b Γ') e0) e1 (F Γ a a' e1) e2 (F (add na a' Γ') b res e2) | @eval_lambda _ na b => f2 Γ na b | @eval_zeta _ na b0 b0' b1 res e0 e1 => f3 Γ na b0 b0' b1 res e0 (F Γ b0 b0' e0) e1 (F (add na b0' Γ) b1 res e1) @@ -280,7 +262,6 @@ Definition ident_of_name (na : name) : ident := Reserved Notation "Γ ;;; E ⊩ s ~ t" (at level 50, E, s, t at next level). Inductive represents : list ident -> environment -> term -> term -> Set := -| represents_tBox Γ E : Γ ;;; E ⊩ tBox ~ tBox | represents_bound_tRel Γ n na E : nth_error Γ n = Some na -> Γ ;;; E ⊩ tVar na ~ tRel n | represents_unbound_tRel E na v Γ s : lookup E na = Some v -> represents_value v s -> Γ ;;; E ⊩ tVar na ~ s | represents_tLambda Γ E na na' b b' : (na :: Γ) ;;; E ⊩ b ~ b' -> Γ ;;; E ⊩ tLambda (nNamed na) b ~ tLambda na' b' @@ -300,7 +281,6 @@ Inductive represents : list ident -> environment -> term -> term -> Set := All2_Set (fun d d' => (List.rev nms ++ Γ) ;;; E ⊩ d.(dbody) ~ d'.(dbody)) mfix mfix' -> Γ ;;; E ⊩ tFix mfix idx ~ tFix mfix' idx with represents_value : value -> term -> Set := -| represents_value_tBox : represents_value vBox tBox | represents_value_tClos na E s t na' : [na] ;;; E ⊩ s ~ t -> represents_value (vClos na s E) (tLambda na' t) | represents_value_tConstruct vs ts ind c : All2_Set represents_value vs ts -> represents_value (vConstruct ind c vs) (tConstruct ind c ts) | represents_value_tFix vfix i E mfix : @@ -311,8 +291,7 @@ Program Definition represents_ind := (λ (P : ∀ (l : list ident) (e : environment) (t t0 : term), l;;; e ⊩ t ~ t0 → Type) (P0 : ∀ (v : value) (t : term), represents_value v t → Type) - (f : ∀ (Γ : list ident) (E : environment), - P Γ E tBox tBox (represents_tBox Γ E)) (f0 : + (f0 : ∀ (Γ : list ident) (n : nat) (na : ident) @@ -374,7 +353,7 @@ Program Definition represents_ind := (IH : All2_over a0 (fun t t' : def term => P (List.rev nms ++ Γ) E (dbody t) (dbody t'))), P Γ E (tFix mfix idx) (tFix mfix' idx) (represents_tFix Γ E mfix mfix' idx nms Hbodies Hnodup a a0)) - (f9 : P0 vBox tBox represents_value_tBox) (f10 : + (f10 : ∀ (na : ident) (E : environment) (s t : term) @@ -401,7 +380,6 @@ Program Definition represents_ind := (l : list ident) (e : environment) (t t0 : term) (r : l;;; e ⊩ t ~ t0) {struct r} : P l e t t0 r := match r as r0 in (l0;;; e0 ⊩ t1 ~ t2) return (P l0 e0 t1 t2 r0) with - | represents_tBox Γ E => f Γ E | represents_bound_tRel Γ n na E e0 => f0 Γ n na E e0 | represents_unbound_tRel E na v Γ s e0 r0 => f1 E na v Γ s e0 r0 (F0 v s r0) @@ -423,7 +401,6 @@ Program Definition represents_ind := with F0 (v : value) (t : term) (r : represents_value v t) {struct r} : P0 v t r := match r as r0 in (represents_value v0 t0) return (P0 v0 t0 r0) with - | represents_value_tBox => f9 | represents_value_tClos na E s t0 na' r0 => f10 na E s t0 na' r0 (F [na] E s t0 r0) | represents_value_tConstruct vs ts ind c a => f11 vs ts ind c a _ @@ -451,8 +428,7 @@ Program Definition represents_value_ind := (λ (P : ∀ (l : list ident) (e : environment) (t t0 : term), l;;; e ⊩ t ~ t0 → Type) (P0 : ∀ (v : value) (t : term), represents_value v t → Type) - (f : ∀ (Γ : list ident) (E : environment), - P Γ E tBox tBox (represents_tBox Γ E)) (f0 : + (f0 : ∀ (Γ : list ident) (n : nat) (na : ident) @@ -518,7 +494,7 @@ Program Definition represents_value_ind := (IH : All2_over a0 (fun t t' : def term => P (List.rev nms ++ Γ) E (dbody t) (dbody t'))), P Γ E (tFix mfix idx) (tFix mfix' idx) (represents_tFix Γ E mfix mfix' idx nms Hbodies Hnodup a a0)) - (f9 : P0 vBox tBox represents_value_tBox) (f10 : + (f10 : ∀ (na : ident) (E : environment) (s t : term) @@ -547,7 +523,6 @@ Program Definition represents_value_ind := (l : list ident) (e : environment) (t t0 : term) (r : l;;; e ⊩ t ~ t0) {struct r} : P l e t t0 r := match r as r0 in (l0;;; e0 ⊩ t1 ~ t2) return (P l0 e0 t1 t2 r0) with - | represents_tBox Γ E => f Γ E | represents_bound_tRel Γ n na E e0 => f0 Γ n na E e0 | represents_unbound_tRel E na v Γ s e0 r0 => f1 E na v Γ s e0 r0 (F0 v s r0) @@ -569,7 +544,6 @@ Program Definition represents_value_ind := with F0 (v : value) (t : term) (r : represents_value v t) {struct r} : P0 v t r := match r as r0 in (represents_value v0 t0) return (P0 v0 t0 r0) with - | represents_value_tBox => f9 | represents_value_tClos na E s t0 na' r0 => f10 na E s t0 na' r0 (F [na] E s t0 r0) | represents_value_tConstruct vs ts ind c a => f11 vs ts ind c a _ @@ -597,8 +571,7 @@ Definition rep_ind := (λ (P : ∀ (l : list ident) (e : environment) (t t0 : term), l;;; e ⊩ t ~ t0 → Type) (P0 : ∀ (v : value) (t : term), represents_value v t → Type) - (f : ∀ (Γ : list ident) (E : environment), - P Γ E tBox tBox (represents_tBox Γ E)) (f0 : + (f0 : ∀ (Γ : list ident) (n : nat) (na : ident) @@ -664,7 +637,7 @@ Definition rep_ind := (IH : All2_over a0 (fun t t' : def term => P (List.rev nms ++ Γ) E (dbody t) (dbody t'))), P Γ E (tFix mfix idx) (tFix mfix' idx) (represents_tFix Γ E mfix mfix' idx nms Hbodies Hnodup a a0)) - (f9 : P0 vBox tBox represents_value_tBox) (f10 : + (f10 : ∀ (na : ident) (E : environment) (s t : term) @@ -691,8 +664,8 @@ Definition rep_ind := (IH : All2_over a (fun v d H => P (List.rev (map fst vfix)) E v.2 (dbody d) (snd H) ) ), P0 (vRecClos vfix i E) (tFix mfix i) (represents_value_tFix vfix i E mfix a)), - (represents_ind P P0 f f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12, - represents_value_ind P P0 f f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12)). + (represents_ind P P0 f0 f1 f2 f3 f4 f5 f6 f7 f8 f10 f11 f12, + represents_value_ind P P0 f0 f1 f2 f3 f4 f5 f6 f7 f8 f10 f11 f12)). Local Notation "'⊩' v ~ s" := (represents_value v s) (at level 50). Local Hint Constructors represents : core. @@ -895,7 +868,7 @@ Fixpoint annotate_env Γ (Σ : global_declarations) := end. Definition extraction_term_flags := - {| has_tBox := true + {| has_tBox := false ; has_tRel := true ; has_tVar := false ; has_tEvar := false @@ -988,7 +961,7 @@ Lemma unfolds_bound : (forall Γ E s t, Γ ;;; E ⊩ s ~ t -> closedn #|Γ| t) × (forall v s, ⊩ v ~ s -> closedn 0 s). Proof. - refine (rep_ind _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); intros; cbn; rtoProp; eauto. + refine (rep_ind _ _ _ _ _ _ _ _ _ _ _ _ _ _); intros; cbn; rtoProp; eauto. - eapply Nat.ltb_lt, nth_error_Some. congruence. - eapply closed_upwards; eauto. lia. - solve_all. induction a; cbn in *; econstructor; firstorder. @@ -1256,7 +1229,6 @@ Fixpoint sunny Γ (t : term) : bool := end. Inductive wf : value -> Type := -| wf_vBox : wf vBox | wf_vClos na b E : ~ In na (map fst E) -> sunny (na :: map fst E) b -> All (fun v => wf (snd v)) E -> wf (vClos na b E) | wf_vConstruct ind c args : All wf args -> wf (vConstruct ind c args) | wf_vRecClos vfix idx E : @@ -1376,7 +1348,6 @@ Proof. destruct eqb eqn:E. * invs e. cbn in *; eauto. * eapply IHHE. exact e. - - econstructor. - let X := match reverse goal with H : All _ _ -> _ -> _ |- _ => H end in do 2 forward X; eauto; inv X. let X1 := multimatch goal with H : _ |- _ => H end in @@ -1605,11 +1576,7 @@ Proof. all: try congruence. all: intros; rtoProp. all: repeat match reverse goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. - - invs Hrep. - + invs H2. - + cbn in Hsunny. rtoProp. edestruct s0 as (v & Hv1 & Hv2). 3: eauto. eauto. eauto. - edestruct s1 as (? & ? & ?); eauto. - invs Hv1. eexists; split; eauto. econstructor; eauto. + - cbn in i0. congruence. - invs Hrep. + invs H3. + cbn in Hsunny. rtoProp. @@ -1872,7 +1839,6 @@ Proof. unfold cstr_arity in *. invs H. lia. clear - Hvs; induction Hvs; econstructor; eauto. eapply r. - invs Hrep; cbn in *; try congruence; rtoProp. - + econstructor. split; eauto. eapply eval_vbox. + econstructor. split; eauto. econstructor. eauto. + econstructor. split; eauto. econstructor. + destruct args'; congruence. From 7565dd43f65a0bd0268c2404b6dc929ed8bc28ab Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Thu, 6 Jul 2023 15:38:40 +0200 Subject: [PATCH 32/61] remove assumption --- pcuic/theories/PCUICFirstorder.v | 18 +++++++++--------- 1 file changed, 9 insertions(+), 9 deletions(-) diff --git a/pcuic/theories/PCUICFirstorder.v b/pcuic/theories/PCUICFirstorder.v index 0942d8141..758a4a2cc 100644 --- a/pcuic/theories/PCUICFirstorder.v +++ b/pcuic/theories/PCUICFirstorder.v @@ -622,19 +622,18 @@ Proof using Type. * intros hl. now eapply invert_cumul_prod_ind in hl. Qed. -Lemma firstorder_value_spec (Σ:global_env_ext) t i u args mind : +Lemma firstorder_value_spec (Σ:global_env_ext) t i u args : wf Σ -> wf_local Σ [] -> Σ ;;; [] |- t : mkApps (tInd i u) args -> PCUICWcbvEval.value Σ t -> - lookup_env Σ (i.(inductive_mind)) = Some (InductiveDecl mind) -> @firstorder_ind Σ (firstorder_env Σ) i -> firstorder_value Σ [] t. Proof using Type. intros Hwf Hwfl Hty Hvalue. - revert mind i u args Hty. + revert i u args Hty. induction Hvalue as [ t Hvalue | t args' Hhead Hargs IH ] using PCUICWcbvEval.value_values_ind; - intros mind i u args Hty Hlookup Hfo. + intros i u args Hty Hfo. - destruct t; inversion_clear Hvalue. + exfalso. eapply inversion_Sort in Hty as (? & ? & Hcumul); eauto. now eapply invert_cumul_sort_ind in Hcumul. @@ -667,9 +666,10 @@ Proof using Type. destruct unfold_fix as [ [] | ]; auto. eapply nth_error_nil. + exfalso. eapply (typing_cofix_coind (args := [])) in Hty. red in Hty. red in Hfo. unfold firstorder_ind in Hfo. - rewrite Hlookup in Hfo. + destruct lookup_env eqn:E; try congruence. + destruct g; try congruence. eapply andb_true_iff in Hfo as [Hfo _]. - rewrite /check_recursivity_kind Hlookup in Hty. + rewrite /check_recursivity_kind E in Hty. now eapply (negb_False _ Hfo). + eapply inversion_Prim in Hty as [prim_ty [cdecl [wf hp hdecl [s []] cum]]]; eauto. now eapply invert_cumul_axiom_ind in cum; tea. @@ -697,7 +697,6 @@ Proof using Type. econstructor. inv X. eapply H0. tea. unshelve eapply declared_inductive_to_gen in d0; eauto. - exact i3. inv X. eapply IHspine; eauto. + exfalso. destruct PCUICWcbvEval.cunfold_fix as [[] | ] eqn:E; inversion H. @@ -708,9 +707,10 @@ Proof using Type. eapply nth_error_None. lia. + exfalso. eapply (typing_cofix_coind (args := args')) in Hty. red in Hfo. unfold firstorder_ind in Hfo. - rewrite Hlookup in Hfo. + destruct lookup_env eqn:E; try congruence. + destruct g; try congruence. eapply andb_true_iff in Hfo as [Hfo _]. - rewrite /check_recursivity_kind Hlookup in Hty. + rewrite /check_recursivity_kind E in Hty. now eapply (negb_False _ Hfo). Qed. From a35996168cf4dd16b13da4e1dcf9f4803cfb9f87 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 7 Jul 2023 18:28:57 +0200 Subject: [PATCH 33/61] Fix definition of Transform.t's observational equivalence preservation --- common/theories/Transform.v | 13 ++++++++----- erasure-plugin/theories/ETransform.v | 18 +++++++++--------- erasure-plugin/theories/Erasure.v | 2 +- template-coq/theories/TemplateProgram.v | 2 +- template-pcuic/theories/PCUICTransform.v | 10 ++-------- 5 files changed, 21 insertions(+), 24 deletions(-) diff --git a/common/theories/Transform.v b/common/theories/Transform.v index cf56d02d6..5a4aeace8 100644 --- a/common/theories/Transform.v +++ b/common/theories/Transform.v @@ -25,11 +25,12 @@ Module Transform. Context {eval : program -> value -> Prop}. Context {eval' : program' -> value' -> Prop}. - Definition preserves_eval pre (transform : forall p : program, pre p -> program') obseq := + Definition preserves_eval pre (transform : forall p : program, pre p -> program') + (obseq : forall p : program, pre p -> program' -> value -> value' -> Prop) := forall p v (pr : pre p), eval p v -> let p' := transform p pr in - exists v', eval' p' v' /\ obseq p p' v v'. + exists v', eval' p' v' /\ obseq p pr p' v v'. Record t := { name : string; @@ -37,7 +38,7 @@ Module Transform. transform : forall p : program, pre p -> program'; post : program' -> Prop; correctness : forall input (p : pre input), post (transform input p); - obseq : program -> program' -> value -> value' -> Prop; + obseq : forall p : program, pre p -> program' -> value -> value' -> Prop; preservation : preserves_eval pre transform obseq; }. Definition run (x : t) (p : program) (pr : pre x p) : program' := @@ -63,7 +64,9 @@ Module Transform. transform p hp := run o' (run o p hp) (hpp _ (o.(correctness) _ hp)); pre := o.(pre); post := o'.(post); - obseq g g' v v' := exists g'' v'', o.(obseq) g g'' v v'' × o'.(obseq) g'' g' v'' v' + obseq g preg g' v v' := + exists v'', o.(obseq) g preg (run o g preg) v v'' × + o'.(obseq) (run o g preg) (hpp _ (o.(correctness) _ preg)) g' v'' v' |}. Next Obligation. intros o o' hpp inp pre. @@ -78,7 +81,7 @@ Module Transform. specialize (H0 (hpp _ (o.(correctness) _ pr)) ev). destruct H0 as [v'' [ev' obs']]. exists v''. constructor => //. - exists (transform o p pr), v'. now split. + exists v'. now split. Qed. End Comp. diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index 7ca17ed50..352a50101 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -204,7 +204,7 @@ Program Definition erase_transform {guard : abstract_guard_impl} : Transform.t p transform p hp := let nhs := proj2 (proj2 hp) in @erase_program guard p (proj1 nhs) (proj2 nhs) (proj1 hp) ; post p := [/\ wf_eprogram_env all_env_flags p & EEtaExpandedFix.expanded_eprogram_env p]; - obseq g g' v v' := let Σ := g.1 in Σ ;;; [] |- v ⇝ℇ v' |}. + obseq p hp p' v v' := let Σ := p.1 in Σ ;;; [] |- v ⇝ℇ v' |}. Next Obligation. cbn -[erase_program]. @@ -271,7 +271,7 @@ Program Definition guarded_to_unguarded_fix {fl : EWcbvEval.WcbvFlags} {wcon : E transform p pre := p; pre p := wf_eprogram_env efl p /\ EEtaExpandedFix.expanded_eprogram_env p; post p := wf_eprogram_env efl p /\ EEtaExpandedFix.expanded_eprogram_env p; - obseq g g' v v' := v' = v |}. + obseq p hp p' v v' := v' = v |}. Next Obligation. cbn. eauto. Qed. Next Obligation. cbn. @@ -292,7 +292,7 @@ Program Definition rebuild_wf_env_transform {fl : EWcbvEval.WcbvFlags} {efl} (wi pre p := wf_eprogram efl p /\ (with_exp ==> EEtaExpanded.expanded_eprogram_cstrs p); transform p pre := rebuild_wf_env p (proj1 pre); post p := wf_eprogram_env efl p /\ (with_exp ==> EEtaExpanded.expanded_eprogram_env_cstrs p); - obseq g g' v v' := v = v' |}. + obseq p hp p' v v' := v = v' |}. Next Obligation. cbn. intros fl efl [] input [wf exp]; cbn; split => //. Qed. @@ -307,7 +307,7 @@ Program Definition remove_params_optimization {fl : EWcbvEval.WcbvFlags} {wcon : transform p pre := ERemoveParams.strip_program p; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; post p := wf_eprogram (switch_no_params efl) p /\ EEtaExpanded.expanded_eprogram_cstrs p; - obseq g g' v v' := v' = (ERemoveParams.strip g.1 v) |}. + obseq p hp p' v v' := v' = (ERemoveParams.strip p.1 v) |}. Next Obligation. move=> fl wcon efl [Σ t] [wfp etap]. simpl. @@ -333,7 +333,7 @@ Program Definition remove_params_fast_optimization (fl : EWcbvEval.WcbvFlags) {w transform p _ := (ERemoveParams.Fast.strip_env p.1, ERemoveParams.Fast.strip p.1 [] p.2); pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; post p := wf_eprogram (switch_no_params efl) p /\ EEtaExpanded.expanded_eprogram_cstrs p; - obseq g g' v v' := v' = (ERemoveParams.strip g.1 v) |}. + obseq p hp p' v v' := v' = (ERemoveParams.strip p.1 v) |}. Next Obligation. move=> fl wcon efl [Σ t] [wfp etap]. simpl. @@ -364,7 +364,7 @@ Program Definition remove_match_on_box_trans {fl : WcbvFlags} {wcon : with_const transform p _ := remove_match_on_box_program p ; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; post p := wf_eprogram efl p /\ EEtaExpanded.expanded_eprogram_cstrs p; - obseq g g' v v' := v' = EOptimizePropDiscr.remove_match_on_box g.1 v |}. + obseq p hp p' v v' := v' = EOptimizePropDiscr.remove_match_on_box p.1 v |}. Next Obligation. move=> fl wcon efl hastrel hastbox [Σ t] [wfp etap]. @@ -390,7 +390,7 @@ Program Definition inline_projections_optimization {fl : WcbvFlags} {wcon : EWcb transform p _ := EInlineProjections.optimize_program p ; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; post p := wf_eprogram (disable_projections_env_flag efl) p /\ EEtaExpanded.expanded_eprogram_cstrs p; - obseq g g' v v' := v' = EInlineProjections.optimize g.1 v |}. + obseq p hp p' v v' := v' = EInlineProjections.optimize p.1 v |}. Next Obligation. move=> fl wcon efl hastrel hastbox [Σ t] [wfp etap]. @@ -414,7 +414,7 @@ Program Definition constructors_as_blocks_transformation (efl : EEnvFlags) transform p _ := EConstructorsAsBlocks.transform_blocks_program p ; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; post p := wf_eprogram (switch_cstr_as_blocks efl) p ; - obseq g g' v v' := v' = EConstructorsAsBlocks.transform_blocks g.1 v |}. + obseq p hp p' v v' := v' = EConstructorsAsBlocks.transform_blocks p.1 v |}. Next Obligation. move=> efl hasapp haspars hascstrs [Σ t] [] [wftp wft] /andP [etap etat]. @@ -441,7 +441,7 @@ Program Definition implement_box_transformation (efl : EEnvFlags) transform p _ := EImplementBox.implement_box_program p ; pre p := wf_eprogram efl p ; post p := wf_eprogram (switch_off_box efl) p ; - obseq g g' v v' := v' = implement_box v |}. + obseq p hp p' v v' := v' = implement_box v |}. Next Obligation. intros. cbn in *. destruct p. split. diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 567ca5f4f..71d754645 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -47,7 +47,7 @@ Program Definition eta_expand K : Transform.t template_program_env template_prog pre := fun p => ∥ wt_template_program_env p ∥ /\ K (eta_expand_global_env p.1) ; transform p _ := EtaExpand.eta_expand_program p ; post := fun p => ∥ wt_template_program p ∥ /\ EtaExpand.expanded_program p /\ K p.1; - obseq p p' v v' := v' = EtaExpand.eta_expand p.1 [] v |}. + obseq p hp p' v v' := v' = EtaExpand.eta_expand p.1 [] v |}. Next Obligation. let p := match goal with H : template_program_env |- _ => H end in destruct p. split; [|split]; auto; now apply assume_welltyped_template_program_expansion. diff --git a/template-coq/theories/TemplateProgram.v b/template-coq/theories/TemplateProgram.v index ed305267e..e873a3d6e 100644 --- a/template-coq/theories/TemplateProgram.v +++ b/template-coq/theories/TemplateProgram.v @@ -52,7 +52,7 @@ Program Definition build_template_program_env {cf : checker_flags} K : pre p := ∥ wt_template_program p ∥ /\ forall pf, K (GlobalEnvMap.make p.1 (wt_template_program_fresh p pf)) : Prop ; transform p pre := make_template_program_env p (proj1 pre); post p := ∥ wt_template_program_env p ∥ /\ K p.1; - obseq g g' v v' := v = v' |}. + obseq g preg g' v v' := v = v' |}. Next Obligation. cbn. exists v. cbn; split; auto. Qed. diff --git a/template-pcuic/theories/PCUICTransform.v b/template-pcuic/theories/PCUICTransform.v index 64ee2007f..359f8998b 100644 --- a/template-pcuic/theories/PCUICTransform.v +++ b/template-pcuic/theories/PCUICTransform.v @@ -18,9 +18,6 @@ Import Transform TemplateToPCUIC. Definition eval_pcuic_program (p : pcuic_program) (v : term) := ∥ PCUICWcbvEval.eval p.1.(trans_env_env) p.2 v ∥. -Definition template_to_pcuic_obseq (p : template_program) (p' : pcuic_program) (v : Ast.term) (v' : term) := - let Σ := Ast.Env.empty_ext p.1 in v' = trans (trans_global Σ) v. - Lemma trans_template_program_wt {cf : checker_flags} p (wtp : wt_template_program p) : wt_pcuic_program (trans_template_program p). Proof. move: p wtp. @@ -43,7 +40,7 @@ Program Definition template_to_pcuic_transform {cf : checker_flags} K : pre p := ∥ wt_template_program p ∥ /\ EtaExpand.expanded_program p /\ K (trans_global (Ast.Env.empty_ext p.1)) ; transform p _ := trans_template_program p ; post p := ∥ wt_pcuic_program p ∥ /\ PCUICEtaExpand.expanded_pcuic_program p /\ K p.1; - obseq := template_to_pcuic_obseq |}. + obseq p _ p' v v':= v' = trans p'.1 v |}. Next Obligation. cbn. intros cf K p [[wtp] [etap ?]]. split; split. @@ -68,16 +65,13 @@ From MetaCoq.PCUIC Require Import PCUICExpandLets PCUICExpandLetsCorrectness. expansion of the old one, which is the identiy on normal forms. *) -Definition let_expansion_obseq (p : pcuic_program) (p' : pcuic_program) (v : term) (v' : term) := - v' = PCUICExpandLets.trans v. - Program Definition pcuic_expand_lets_transform {cf : checker_flags} K : self_transform pcuic_program term eval_pcuic_program eval_pcuic_program := {| name := "let expansion in branches/constructors"; pre p := ∥ wt_pcuic_program p ∥ /\ PCUICEtaExpand.expanded_pcuic_program p /\ K (build_global_env_map (trans_global_env p.1.1), p.1.2) ; transform p hp := expand_lets_program p; post p := ∥ wt_pcuic_program (cf:=PCUICExpandLetsCorrectness.cf' cf) p ∥ /\ PCUICEtaExpand.expanded_pcuic_program p /\ K p.1; - obseq := let_expansion_obseq |}. + obseq p _ p' v v' := v' = PCUICExpandLets.trans v |}. Next Obligation. intros cf K [Σ t] [[[wfext ht]] etap]. cbn. split. sq. unfold build_global_env_map. unfold global_env_ext_map_global_env_ext. simpl. From 06ca492ccf37fad84a62778f1517c0683515c5ee Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Tue, 11 Jul 2023 16:45:41 +0200 Subject: [PATCH 34/61] prove wellformedness of annotate --- erasure/theories/EWcbvEvalNamed.v | 106 +++++++++++++++++++++++++----- utils/theories/MCList.v | 16 +++-- 2 files changed, 99 insertions(+), 23 deletions(-) diff --git a/erasure/theories/EWcbvEvalNamed.v b/erasure/theories/EWcbvEvalNamed.v index bab199936..8496db6e2 100644 --- a/erasure/theories/EWcbvEvalNamed.v +++ b/erasure/theories/EWcbvEvalNamed.v @@ -289,7 +289,7 @@ where "Γ ;;; E ⊩ s ~ t" := (represents Γ E s t). Program Definition represents_ind := (λ (P : ∀ (l : list ident) (e : environment) (t t0 : term), - l;;; e ⊩ t ~ t0 → Type) (P0 : ∀ (v : value) (t : term), + l;;; e ⊩ t ~ t0 → Type) (P0 : ∀ (v : value) (t : term), represents_value v t → Type) (f0 : ∀ (Γ : list ident) @@ -426,7 +426,7 @@ Defined. Program Definition represents_value_ind := (λ (P : ∀ (l : list ident) (e : environment) (t t0 : term), - l;;; e ⊩ t ~ t0 → Type) (P0 : ∀ (v : value) (t : term), + l;;; e ⊩ t ~ t0 → Type) (P0 : ∀ (v : value) (t : term), represents_value v t → Type) (f0 : ∀ (Γ : list ident) @@ -569,7 +569,7 @@ Defined. Definition rep_ind := (λ (P : ∀ (l : list ident) (e : environment) (t t0 : term), - l;;; e ⊩ t ~ t0 → Type) (P0 : ∀ (v : value) (t : term), + l;;; e ⊩ t ~ t0 → Type) (P0 : ∀ (v : value) (t : term), represents_value v t → Type) (f0 : ∀ (Γ : list ident) @@ -809,21 +809,21 @@ Definition map_def_name := {| dname := g (dname d); dbody := f (dbody d); rarg := rarg d |}. -Section Map2. +(* Section Map2. *) - Context {A B C : Type}. - Variable (f : A → B → C). - Fixpoint map2 (l : list A) (l' : list B) {struct l} : - list C := - match l with - | [] => [] - | hd :: tl => - match l' with - | [] => [] - | hd' :: tl' => f hd hd' :: map2 tl tl' - end - end. -End Map2. +(* Context {A B C : Type}. *) +(* Variable (f : A → B → C). *) +(* Fixpoint map2 (l : list A) (l' : list B) {struct l} : *) +(* list C := *) +(* match l with *) +(* | [] => [] *) +(* | hd :: tl => *) +(* match l' with *) +(* | [] => [] *) +(* | hd' :: tl' => f hd hd' :: map2 tl tl' *) +(* end *) +(* end. *) +(* End Map2. *) Fixpoint annotate (s : list ident) (u : term) {struct u} : term := match u with @@ -912,6 +912,76 @@ Proof. * intros x [<- | ?]. 1: eauto. eapply IHl in H. firstorder. Qed. +Definition named_extraction_term_flags := + {| has_tBox := false + ; has_tRel := false + ; has_tVar := true + ; has_tEvar := false + ; has_tLambda := true + ; has_tLetIn := true + ; has_tApp := true + ; has_tConst := true + ; has_tConstruct := true + ; has_tCase := true + ; has_tProj := false + ; has_tFix := true + ; has_tCoFix := false + ; has_tPrim := false + |}. + +Definition named_extraction_env_flags := + {| has_axioms := false; + term_switches := named_extraction_term_flags; + has_cstr_params := false ; + cstr_as_blocks := true |}. + +Lemma wellformed_annotate Σ Γ s : + wellformed (efl := extraction_env_flags) Σ #|Γ| s -> wellformed (efl := named_extraction_env_flags) Σ #|Γ| (annotate Γ s). +Proof. + intros Hwf. + induction s in Γ, Hwf |- * using EInduction.term_forall_list_ind; cbn in *; rtoProp; unshelve eauto. + - eapply Nat.ltb_lt in Hwf. destruct nth_error eqn:Eq; eauto. + eapply nth_error_None in Eq. lia. + - destruct n; cbn. 2: destruct in_dec. + all: eapply (IHs (_ :: Γ)); cbn; eauto. + - split; eauto. destruct n; cbn. 2: destruct in_dec; cbn. + all: eapply (IHs2 (_ :: Γ)); cbn; eauto. + - destruct lookup_env as [ [] | ]; cbn in *; eauto. + destruct nth_error as [ [] | ]; cbn in *; eauto. + destruct nth_error as [ [] | ]; cbn in *; eauto. + repeat split. len. solve_all. + - destruct lookup_env as [ [] | ]; cbn in *; eauto. + destruct nth_error as [ [] | ]; cbn in *; eauto. + repeat split. eauto. + solve_all. rewrite map_length. rewrite <- app_length. + eapply a0. len. rewrite gen_many_fresh_length. eauto. + - split. + { clear - H. generalize ((List.rev (gen_many_fresh Γ ( (map dname m))) ++ Γ)). + induction m in Γ, H |- *; cbn. + - econstructor. + - intros. destruct a; cbn in *. destruct dname; cbn; rtoProp; repeat split; eauto. + all: destruct dbody; cbn in *; eauto. + + } + { solve_all. unfold wf_fix in *. rtoProp. split. + rewrite map2_length gen_many_fresh_length map_length. + { eapply Nat.ltb_lt in H0. eapply Nat.ltb_lt. lia. } + solve_all. clear H0. unfold test_def in *. cbn in *. + eapply All_impl in H1. 2:{ intros ? [[] ]. + specialize (i (List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)). + revert i. rewrite ?List.rev_length app_length ?List.rev_length gen_many_fresh_length ?List.rev_length map_length. intros r. eapply r in i1. exact i1. + } + revert H1. + generalize ((List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)). + intros. rewrite map2_length gen_many_fresh_length map_length Nat.min_id. + revert H1. generalize (#|m| + #|Γ|). + intros. + induction m in Γ, n, n0, l, H1 |- *. + - econstructor. + - invs H1. cbn. destruct a; cbn. destruct dname; cbn; econstructor; eauto. + } +Qed. + Lemma nclosed_represents Σ Γ E s : wellformed Σ #|Γ| s -> Γ ;;; E ⊩ annotate Γ s ~ s. Proof. @@ -1981,4 +2051,4 @@ Proof. intros. eapply eval_to_eval_named; eauto. eapply sunny_annotate. eapply nclosed_represents. eauto. -Qed. \ No newline at end of file +Qed. diff --git a/utils/theories/MCList.v b/utils/theories/MCList.v index 4bf7b7194..b3900805b 100644 --- a/utils/theories/MCList.v +++ b/utils/theories/MCList.v @@ -57,11 +57,17 @@ Next Obligation. simpl in H. auto with arith. Defined. -Fixpoint map2 {A B C} (f : A -> B -> C) (l : list A) (l' : list B) : list C := - match l, l' with - | hd :: tl, hd' :: tl' => f hd hd' :: map2 f tl tl' - | _, _ => [] - end. +Section map2. + + Context {A B C} (f : A -> B -> C). + + Fixpoint map2 (l : list A) (l' : list B) : list C := + match l, l' with + | hd :: tl, hd' :: tl' => f hd hd' :: map2 tl tl' + | _, _ => [] + end. + +End map2. Lemma map2_ext {A B C} (f g : A -> B -> C) (l : list A) (l' : list B) : (forall x y, f x y = g x y) -> From f1a254857c202fba06cd74d12b365e591321c2c2 Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Wed, 12 Jul 2023 16:08:14 +0200 Subject: [PATCH 35/61] wellformed_annotate_env --- erasure/theories/EWcbvEvalNamed.v | 50 ++++++++++++++++++++++++------- 1 file changed, 40 insertions(+), 10 deletions(-) diff --git a/erasure/theories/EWcbvEvalNamed.v b/erasure/theories/EWcbvEvalNamed.v index 8496db6e2..a6ca1034e 100644 --- a/erasure/theories/EWcbvEvalNamed.v +++ b/erasure/theories/EWcbvEvalNamed.v @@ -862,7 +862,7 @@ Fixpoint annotate (s : list ident) (u : term) {struct u} : term := Fixpoint annotate_env Γ (Σ : global_declarations) := match Σ with - | (na, ConstantDecl (Build_constant_body (Some b))) :: Σ => (na, ConstantDecl (Build_constant_body (Some (annotate Γ b)))) :: annotate_env (string_of_kername na :: Γ) Σ + | (na, ConstantDecl (Build_constant_body (Some b))) :: Σ => (na, ConstantDecl (Build_constant_body (Some (annotate Γ b)))) :: annotate_env (Γ) Σ | d :: Σ => d :: annotate_env Γ Σ | nil => nil end. @@ -935,33 +935,57 @@ Definition named_extraction_env_flags := has_cstr_params := false ; cstr_as_blocks := true |}. -Lemma wellformed_annotate Σ Γ s : - wellformed (efl := extraction_env_flags) Σ #|Γ| s -> wellformed (efl := named_extraction_env_flags) Σ #|Γ| (annotate Γ s). +Lemma lookup_annotate_env Γ Σ i d : + lookup_env Σ i = Some d -> + lookup_env (annotate_env Γ Σ) i = Some + match d with + | ConstantDecl {| cst_body := b |} => ConstantDecl {| cst_body := option_map (annotate Γ) b |} + | InductiveDecl m => InductiveDecl m + end. Proof. - intros Hwf. - induction s in Γ, Hwf |- * using EInduction.term_forall_list_ind; cbn in *; rtoProp; unshelve eauto. + induction Σ in i |- *; cbn; try congruence. + destruct a. cbn. + destruct (eqb_spec i k). + + subst. intros [= ->]. + destruct d as [ [ []] | ]; cbn; now rewrite eqb_refl. + + intros Hi % IHΣ. destruct d as [ [] | ]; cbn in *; + destruct g as [ [[]]| ]; cbn; destruct (eqb_spec i k); try congruence. +Qed. + +Local Hint Resolve incl_tl incl_appr incl_appl : core. + +Lemma wellformed_annotate' Σ Γ Γ' s : + incl Γ' Γ -> + wellformed (efl := extraction_env_flags) Σ #|Γ| s -> wellformed (efl := named_extraction_env_flags) (annotate_env Γ' Σ) #|Γ| (annotate Γ s). +Proof. + intros Hincl Hwf. + induction s in Γ, Hwf, Γ', Hincl |- * using EInduction.term_forall_list_ind; cbn in *; rtoProp; unshelve eauto. - eapply Nat.ltb_lt in Hwf. destruct nth_error eqn:Eq; eauto. eapply nth_error_None in Eq. lia. - destruct n; cbn. 2: destruct in_dec. all: eapply (IHs (_ :: Γ)); cbn; eauto. - split; eauto. destruct n; cbn. 2: destruct in_dec; cbn. all: eapply (IHs2 (_ :: Γ)); cbn; eauto. - - destruct lookup_env as [ [] | ]; cbn in *; eauto. + - destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. + destruct c, cst_body; cbn in *; try congruence. + erewrite lookup_annotate_env; eauto. cbn; eauto. + - destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. + erewrite lookup_annotate_env; eauto. cbn. destruct nth_error as [ [] | ]; cbn in *; eauto. destruct nth_error as [ [] | ]; cbn in *; eauto. repeat split. len. solve_all. - - destruct lookup_env as [ [] | ]; cbn in *; eauto. + - destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. + erewrite lookup_annotate_env; eauto. cbn. destruct nth_error as [ [] | ]; cbn in *; eauto. repeat split. eauto. solve_all. rewrite map_length. rewrite <- app_length. - eapply a0. len. rewrite gen_many_fresh_length. eauto. + eapply a; eauto. len. rewrite gen_many_fresh_length. eauto. - split. { clear - H. generalize ((List.rev (gen_many_fresh Γ ( (map dname m))) ++ Γ)). induction m in Γ, H |- *; cbn. - econstructor. - intros. destruct a; cbn in *. destruct dname; cbn; rtoProp; repeat split; eauto. all: destruct dbody; cbn in *; eauto. - } { solve_all. unfold wf_fix in *. rtoProp. split. rewrite map2_length gen_many_fresh_length map_length. @@ -970,6 +994,7 @@ Proof. eapply All_impl in H1. 2:{ intros ? [[] ]. specialize (i (List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)). revert i. rewrite ?List.rev_length app_length ?List.rev_length gen_many_fresh_length ?List.rev_length map_length. intros r. eapply r in i1. exact i1. + eapply incl_appr. eauto. } revert H1. generalize ((List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)). @@ -982,6 +1007,12 @@ Proof. } Qed. +Lemma wellformed_annotate Σ Γ s : + wellformed (efl := extraction_env_flags) Σ #|Γ| s -> wellformed (efl := named_extraction_env_flags) (annotate_env Γ Σ) #|Γ| (annotate Γ s). +Proof. + eapply wellformed_annotate'. firstorder. +Qed. + Lemma nclosed_represents Σ Γ E s : wellformed Σ #|Γ| s -> Γ ;;; E ⊩ annotate Γ s ~ s. Proof. @@ -1385,7 +1416,6 @@ Proof. eapply IHt; eauto. intros ? []; cbn; eauto. - destruct n; eauto. cbn in *. rtoProp. repeat split; eauto. eapply IHt2; eauto. eapply incl_cons; cbn; eauto. - eapply incl_tl; eauto. - rtoProp; split; eauto. - solve_all. - rtoProp; split; solve_all; rtoProp; eauto; repeat split; solve_all. From 6a1428acb3f8073dc9f84a3b3f42e6fe8d6f8d1e Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 17 Jul 2023 09:58:23 +0200 Subject: [PATCH 36/61] WIP stronger global weakening theorems for optimizations --- erasure/_CoqProject.in | 1 + erasure/theories/EEnvMap.v | 48 +++- erasure/theories/EEtaExpanded.v | 7 +- erasure/theories/EEtaExpandedFix.v | 5 +- erasure/theories/EExtends.v | 18 +- erasure/theories/EGenericGlobalMap.v | 390 ++++++++++++++++++++++++++ erasure/theories/EGenericMapEnv.v | 59 ++-- erasure/theories/EGlobalEnv.v | 5 +- erasure/theories/EGlobalFragment.v | 143 ++++++++++ erasure/theories/EInlineProjections.v | 211 ++++++-------- erasure/theories/EOptimizePropDiscr.v | 17 +- erasure/theories/EWellformed.v | 34 ++- 12 files changed, 776 insertions(+), 162 deletions(-) create mode 100644 erasure/theories/EGenericGlobalMap.v create mode 100644 erasure/theories/EGlobalFragment.v diff --git a/erasure/_CoqProject.in b/erasure/_CoqProject.in index ed2993839..3a16a4112 100644 --- a/erasure/_CoqProject.in +++ b/erasure/_CoqProject.in @@ -13,6 +13,7 @@ theories/EWcbvEvalNamed.v # theories/EWtAst.v theories/EGlobalEnv.v theories/EGenericMapEnv.v +theories/EGenericGlobalMap.v theories/EWellformed.v theories/EEnvMap.v theories/EWcbvEvalInd.v diff --git a/erasure/theories/EEnvMap.v b/erasure/theories/EEnvMap.v index c2c3875ad..23fa1759f 100644 --- a/erasure/theories/EEnvMap.v +++ b/erasure/theories/EEnvMap.v @@ -1,4 +1,4 @@ -From Coq Require Import ssreflect ssrbool. +From Coq Require Import ssreflect ssrbool Morphisms Setoid. From Equations Require Import Equations. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import Kernames EnvMap BasicAst. @@ -15,6 +15,52 @@ Module GlobalContextMap. repr : EnvMap.repr global_decls map; wf : EnvMap.fresh_globals global_decls }. + Program Definition empty : t := + {| global_decls := []; + map := EnvMap.empty |}. + Next Obligation. + constructor. + Qed. + + Program Definition add (Σ : GlobalContextMap.t) kn d (hpre : fresh_global kn Σ) : GlobalContextMap.t := + {| global_decls := (kn,d) :: Σ.(global_decls); + map := EnvMap.add kn d Σ.(map) |}. + Next Obligation. + simpl. unfold KernameMapFact.uncurry. + pose proof (Σ.(repr)). unfold EnvMap.repr in H. + now rewrite H. + Qed. + Next Obligation. + constructor => //. apply Σ.(wf). + Qed. + + Definition global_context_map_equal Σ Σ' := + Σ.(global_decls) = Σ'.(global_decls). + + #[global] Instance gceq : DefaultRelation global_context_map_equal := {}. + + Lemma global_context_map_ind (P : GlobalContextMap.t -> Prop) + (proper : Morphisms.Proper (global_context_map_equal ==> iff) P) + (Pnil : P empty) + (Pcons : + forall Σ kn d fr, + P Σ -> + P (add Σ kn d fr)) + Σ : P Σ. + Proof. + destruct Σ as [Σ ? ? ?]. + induction Σ in map0, repr0, wf0 |- *; set (Σ':= {| global_decls := _ |}). + - setoid_replace Σ' with empty => //. + - set (Σg := {| global_decls := Σ; map := EnvMap.of_global_env Σ; repr := EnvMap.repr_global_env _; + wf := fresh_globals_cons_inv wf0 |}). + destruct a as [kn d]. + assert (fresh_global kn Σ). now depelim wf0. + setoid_replace Σ' with (add Σg kn d H). + apply Pcons. apply IHΣ. + red. unfold Σ', add. + cbn. reflexivity. + Qed. + Definition lookup_env Σ kn := EnvMap.lookup kn Σ.(map). Lemma lookup_env_spec (Σ : t) kn : lookup_env Σ kn = EGlobalEnv.lookup_env Σ kn. diff --git a/erasure/theories/EEtaExpanded.v b/erasure/theories/EEtaExpanded.v index 90415ebc3..654e4e733 100644 --- a/erasure/theories/EEtaExpanded.v +++ b/erasure/theories/EEtaExpanded.v @@ -5,7 +5,7 @@ From Coq Require Import Utf8 Program. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config Kernames EnvMap BasicAst. -From MetaCoq.Erasure Require Import EAst EAstUtils EInduction EGlobalEnv EWellformed ELiftSubst ESpineView ECSubst EWcbvEval EWcbvEvalInd EProgram. +From MetaCoq.Erasure Require Import EAst EAstUtils EInduction EGlobalEnv EExtends EWellformed ELiftSubst ESpineView ECSubst EWcbvEval EWcbvEvalInd EProgram. Local Open Scope string_scope. Set Asymmetric Patterns. @@ -372,7 +372,7 @@ Fixpoint isEtaExp_env (Σ : global_declarations) := Lemma isEtaExp_lookup_ext {Σ} {kn d}: isEtaExp_env Σ -> lookup_env Σ kn = Some d -> - ∑ Σ', extends Σ' Σ × isEtaExp_decl Σ' d. + ∑ Σ', extends_prefix Σ' Σ × isEtaExp_decl Σ' d. Proof. induction Σ; cbn. - move=> _; rewrite /declared_constant /lookup_env /= //. @@ -442,7 +442,8 @@ Lemma isEtaExp_lookup {efl : EEnvFlags} {Σ kn d}: Proof. move=> etaΣ wfΣ. move/(isEtaExp_lookup_ext etaΣ) => [Σ' []] ext hd. - now eapply isEtaExp_extends_decl. + eapply isEtaExp_extends_decl; eauto. + now eapply extends_prefix_extends. Qed. Section expanded. diff --git a/erasure/theories/EEtaExpandedFix.v b/erasure/theories/EEtaExpandedFix.v index 57656b29b..8c14b356c 100644 --- a/erasure/theories/EEtaExpandedFix.v +++ b/erasure/theories/EEtaExpandedFix.v @@ -946,7 +946,7 @@ Qed. Lemma isEtaExp_lookup_ext {Σ} {kn d}: isEtaExp_env Σ -> lookup_env Σ kn = Some d -> - ∑ Σ', extends Σ' Σ × isEtaExp_decl Σ' d. + ∑ Σ', extends_prefix Σ' Σ × isEtaExp_decl Σ' d. Proof. induction Σ; cbn. - move=> _; rewrite /declared_constant /lookup_env /= //. @@ -1017,7 +1017,8 @@ Lemma isEtaExp_lookup {efl : EEnvFlags} {Σ kn d}: Proof. move=> etaΣ wfΣ. move/(isEtaExp_lookup_ext etaΣ) => [Σ' []] ext hd. - now eapply isEtaExp_extends_decl. + eapply isEtaExp_extends_decl; eauto. + now eapply extends_prefix_extends. Qed. Arguments lookup_inductive_pars_constructor_pars_args {Σ ind n pars args}. diff --git a/erasure/theories/EExtends.v b/erasure/theories/EExtends.v index 869687c06..a431160c6 100644 --- a/erasure/theories/EExtends.v +++ b/erasure/theories/EExtends.v @@ -45,13 +45,29 @@ Section EEnvFlags. split; eauto. eapply weakening_env_declared_inductive; eauto. Qed. - Lemma extends_wf_glob {Σ Σ'} : extends Σ Σ' -> wf_glob Σ' -> wf_glob Σ. + Lemma extends_wf_glob {Σ Σ'} : extends_prefix Σ Σ' -> wf_glob Σ' -> wf_glob Σ. Proof using Type. intros [? ->]. induction x; cbn; auto. intros wf; depelim wf. eauto. Qed. + Lemma extends_prefix_extends Σ Σ' : + extends_prefix Σ Σ' -> wf_glob Σ' -> + extends Σ Σ'. + Proof. + intros [Σ'' ->]. + induction Σ''; cbn; auto. + - unfold extends; tauto. + - intros wf; depelim wf. + specialize (IHΣ'' wf). + red in IHΣ''. intros ? ? H'; specialize (IHΣ'' _ _ H'). + cbn. + case: (eqb_spec kn0 kn). + + intros <-. now eapply lookup_env_Some_fresh in IHΣ''. + + now intros _. + Qed. + Definition global_subset (Σ Σ' : global_declarations) := forall kn d, lookup_env Σ kn = Some d -> lookup_env Σ' kn = Some d. diff --git a/erasure/theories/EGenericGlobalMap.v b/erasure/theories/EGenericGlobalMap.v new file mode 100644 index 000000000..4c80c5a2a --- /dev/null +++ b/erasure/theories/EGenericGlobalMap.v @@ -0,0 +1,390 @@ +(* Distributed under the terms of the MIT license. *) +From Coq Require Import Utf8 Program. +From MetaCoq.Utils Require Import utils. +From MetaCoq.Common Require Import config Kernames BasicAst EnvMap. +From MetaCoq.Erasure Require Import EAst EAstUtils EInduction EArities + ELiftSubst ESpineView EGlobalEnv EWellformed EEnvMap + EWcbvEval EEtaExpanded ECSubst EWcbvEvalEtaInd EProgram. + +Local Open Scope string_scope. +Set Asymmetric Patterns. +Import MCMonadNotation. + +From Equations Require Import Equations. +Set Equations Transparent. +Local Set Keyed Unification. +Require Import ssreflect ssrbool. + +Class GenTransform := + gen_transform : GlobalContextMap.t -> term -> term. + +Section GenTransformEnv. + Context {GT : GenTransform}. + + Definition gen_transform_constant_decl Σ cb := + {| cst_body := option_map (gen_transform Σ) cb.(cst_body) |}. + + Definition gen_transform_decl Σ d := + match d with + | ConstantDecl cb => ConstantDecl (gen_transform_constant_decl Σ cb) + | InductiveDecl idecl => d + end. + + Definition gen_transform_env Σ := + map (on_snd (gen_transform_decl Σ)) Σ.(GlobalContextMap.global_decls). + + Program Fixpoint gen_transform_env' Σ : EnvMap.fresh_globals Σ -> global_context := + match Σ with + | [] => fun _ => [] + | hd :: tl => fun HΣ => + let Σg := GlobalContextMap.make tl (fresh_globals_cons_inv HΣ) in + on_snd (gen_transform_decl Σg) hd :: gen_transform_env' tl (fresh_globals_cons_inv HΣ) + end. +End GenTransformEnv. + +Class GenTransformExtends (efl efl' : EEnvFlags) (GT : GenTransform) := + wellformed_gen_transform_extends : forall {Σ : GlobalContextMap.t} t, + forall n, @EWellformed.wellformed efl Σ n t -> + forall {Σ' : GlobalContextMap.t}, extends Σ Σ' -> + @wf_glob efl Σ -> @wf_glob efl Σ' -> + gen_transform Σ t = gen_transform Σ' t. + +Class GenTransformWf (efl efl' : EEnvFlags) (GT : GenTransform) := +{ gen_transform_pre : GlobalContextMap.t -> term -> Prop; + gtwf_axioms_efl : forall _ : is_true (@has_axioms efl), is_true (@has_axioms efl'); + gtwf_cstrs_efl : forall _ : is_true (@has_cstr_params efl), is_true (@has_cstr_params efl'); + + gen_transform_wellformed : forall {Σ : GlobalContextMap.t} n t, + has_tBox -> has_tRel -> gen_transform_pre Σ t -> + @wf_glob efl Σ -> @EWellformed.wellformed efl Σ n t -> + EWellformed.wellformed (efl := efl') (gen_transform_env Σ) n (gen_transform Σ t) }. + +Section sec. + Context {efl efl' gen_transform} {gt : GenTransformExtends efl efl' gen_transform}. + +Import EGlobalEnv EExtends. + + +Lemma wellformed_gen_transform_decl_extends {Σ : GlobalContextMap.t} t : + @wf_global_decl efl Σ t -> + forall {Σ' : GlobalContextMap.t}, extends Σ Σ' -> @wf_glob efl Σ -> @wf_glob efl Σ' -> + gen_transform_decl Σ t = gen_transform_decl Σ' t. +Proof. +destruct t => /= //. +intros wf Σ' ext wfΣ wf'. f_equal. unfold gen_transform_constant_decl. f_equal. +destruct (cst_body c) => /= //. f_equal. +now eapply wellformed_gen_transform_extends. +Qed. + +Lemma lookup_env_gen_transform_env_Some {Σ : GlobalContextMap.t} kn d : + @wf_glob efl Σ -> + lookup_env Σ kn = Some d -> + ∑ Σ' : GlobalContextMap.t, + [× extends_prefix Σ' Σ, @wf_global_decl efl Σ' d & + lookup_env (gen_transform_env Σ) kn = Some (gen_transform_decl Σ' d)]. +Proof. + destruct Σ as [Σ ? ? ?]. cbn. + revert map repr wf. + induction Σ in |- *; simpl; auto => //. + intros map repr fr wfg. + case: eqb_specT => //. + - intros ->. cbn. intros [= <-]. + exists (GlobalContextMap.make Σ (fresh_globals_cons_inv fr)). split. + now eexists [_]. + cbn. now depelim wfg. + f_equal. symmetry. eapply wellformed_gen_transform_decl_extends. cbn. now depelim wfg. + eapply (@extends_prefix_extends efl) => //. + cbn. now exists [a]. now depelim wfg. auto. + - intros _. + cbn in IHΣ. + specialize (IHΣ (EnvMap.of_global_env Σ) (EnvMap.repr_global_env _) (fresh_globals_cons_inv fr)). + forward IHΣ. now depelim wfg. + intros hl. specialize (IHΣ hl) as [Σ'' [ext wfgd hl']]. + exists Σ''. split => //. + * destruct ext as [? ->]. + now exists (a :: x). + * rewrite -hl'. f_equal. + clear -wfg gt. + eapply map_ext_in => kn hin. unfold on_snd. f_equal. + symmetry. eapply wellformed_gen_transform_decl_extends => //. cbn. + eapply lookup_env_In in hin. 2:now depelim wfg. + depelim wfg. eapply lookup_env_wellformed; tea. + eapply (@extends_prefix_extends efl) => //. + cbn. now exists [a]. now depelim wfg. +Qed. + +Lemma lookup_env_map_snd Σ f kn : lookup_env (List.map (on_snd f) Σ) kn = option_map f (lookup_env Σ kn). +Proof. +induction Σ; cbn; auto. +case: eqb_spec => //. +Qed. + +Lemma lookup_env_gen_transform_env_None {Σ : GlobalContextMap.t} kn : +lookup_env Σ kn = None -> +lookup_env (gen_transform_env Σ) kn = None. +Proof. +cbn. intros hl. rewrite lookup_env_map_snd hl //. +Qed. + +Lemma lookup_env_gen_transform {Σ : GlobalContextMap.t} kn : +@wf_glob efl Σ -> +lookup_env (gen_transform_env Σ) kn = option_map (gen_transform_decl Σ) (lookup_env Σ kn). +Proof. +intros wf. +destruct (lookup_env Σ kn) eqn:hl. +- eapply lookup_env_gen_transform_env_Some in hl as [Σ' [ext wf' hl']] => /=. + rewrite hl'. f_equal. + eapply wellformed_gen_transform_decl_extends; eauto. + eapply (@extends_prefix_extends efl); auto. + eapply extends_wf_glob; eauto. + auto. + +- cbn. now eapply lookup_env_gen_transform_env_None in hl. +Qed. + + +Lemma is_propositional_gen_transform {Σ : GlobalContextMap.t} ind : + @wf_glob efl Σ -> + inductive_isprop_and_pars Σ ind = inductive_isprop_and_pars (gen_transform_env Σ) ind. +Proof. + rewrite /inductive_isprop_and_pars => wf. + rewrite /lookup_inductive /lookup_minductive. + rewrite (lookup_env_gen_transform (inductive_mind ind) wf). + rewrite /GlobalContextMap.inductive_isprop_and_pars /GlobalContextMap.lookup_inductive + /GlobalContextMap.lookup_minductive. + destruct lookup_env as [[decl|]|] => //. +Qed. + +Lemma is_propositional_cstr_gen_transform {Σ : GlobalContextMap.t} ind c : + @wf_glob efl Σ -> + constructor_isprop_pars_decl Σ ind c = constructor_isprop_pars_decl (gen_transform_env Σ) ind c. +Proof. + rewrite /constructor_isprop_pars_decl => wf. + rewrite /lookup_constructor /lookup_inductive /lookup_minductive. + rewrite (lookup_env_gen_transform (inductive_mind ind) wf). + rewrite /GlobalContextMap.inductive_isprop_and_pars /GlobalContextMap.lookup_inductive + /GlobalContextMap.lookup_minductive. + destruct lookup_env as [[decl|]|] => //. +Qed. + +Lemma isFix_mkApps t l : isFix (mkApps t l) = isFix t && match l with [] => true | _ => false end. +Proof. + induction l using rev_ind; cbn. + - now rewrite andb_true_r. + - rewrite mkApps_app /=. now destruct l => /= //; rewrite andb_false_r. +Qed. + +Lemma lookup_constructor_gen_transform {Σ : GlobalContextMap.t} {ind c} : + @wf_glob efl Σ -> + lookup_constructor Σ ind c = lookup_constructor (gen_transform_env Σ) ind c. +Proof. + intros wfΣ. rewrite /lookup_constructor /lookup_inductive /lookup_minductive. + rewrite lookup_env_gen_transform // /=. destruct lookup_env => // /=. + destruct g => //. +Qed. + +Lemma lookup_projection_gen_transform {Σ : GlobalContextMap.t} {p} : + @wf_glob efl Σ -> + lookup_projection Σ p = lookup_projection (gen_transform_env Σ) p. +Proof. + intros wfΣ. rewrite /lookup_projection. + rewrite -lookup_constructor_gen_transform //. +Qed. + +Lemma constructor_isprop_pars_decl_inductive {Σ ind c} {prop pars cdecl} : + constructor_isprop_pars_decl Σ ind c = Some (prop, pars, cdecl) -> + inductive_isprop_and_pars Σ ind = Some (prop, pars). +Proof. + rewrite /constructor_isprop_pars_decl /inductive_isprop_and_pars /lookup_constructor. + destruct lookup_inductive as [[mdecl idecl]|]=> /= //. + destruct nth_error => //. congruence. +Qed. + +Lemma constructor_isprop_pars_decl_constructor {Σ ind c} {mdecl idecl cdecl} : + lookup_constructor Σ ind c = Some (mdecl, idecl, cdecl) -> + constructor_isprop_pars_decl Σ ind c = Some (ind_propositional idecl, ind_npars mdecl, cdecl). +Proof. + rewrite /constructor_isprop_pars_decl. intros -> => /= //. +Qed. + +Lemma wf_mkApps (ha : has_tApp) Σ k f args : reflect (wellformed Σ k f /\ forallb (wellformed Σ k) args) (wellformed Σ k (mkApps f args)). +Proof. + rewrite wellformed_mkApps //. eapply andP. +Qed. + +Lemma extends_cons_inv Σ kn d Σ' : + extends ((kn, d) :: Σ) Σ' -> + fresh_global kn Σ -> + extends Σ Σ'. +Proof. + intros ext fr kn' d' hl. + apply ext. rewrite elookup_env_cons_fresh => //. + intros <-. now apply lookup_env_Some_fresh in hl. +Qed. + +Lemma extends_cons_wf Σ a : + @wf_glob efl (a :: Σ) -> + extends Σ (a :: Σ). +Proof. + intros hwf; depelim hwf. + now apply extends_fresh. +Qed. + +Lemma gen_transform_env_extends' {Σ Σ' : GlobalContextMap.t} : + extends Σ Σ' -> + @wf_glob efl Σ -> + @wf_glob efl Σ' -> + List.map (on_snd (gen_transform_decl Σ)) Σ.(GlobalContextMap.global_decls) = + List.map (on_snd (gen_transform_decl Σ')) Σ.(GlobalContextMap.global_decls). +Proof. + intros ext. + move=> wfΣ wfΣ'. + assert (Hext : extends Σ Σ); auto. red; tauto. + generalize wfΣ. + revert Hext. + generalize Σ at 1 3 5 6. + intros [Σ'' ? ? ?]. revert map repr wf. + induction Σ'' => //; intros map repr wf. cbn. + intros hin wfg. depelim wfg. + f_equal. + 2:{ eapply (IHΣ'' (EnvMap.of_global_env Σ'') (EnvMap.repr_global_env _) (fresh_globals_cons_inv wf)). + now apply extends_cons_inv in hin. now cbn. } + unfold on_snd. cbn. f_equal. + eapply wellformed_gen_transform_decl_extends => //. cbn. + eapply extends_wf_global_decl => //. 2:tea. + now apply extends_cons_inv in hin. +Qed. + +Lemma gen_transform_env_eq (Σ : GlobalContextMap.t) : @wf_glob efl Σ -> + gen_transform_env Σ = gen_transform_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). +Proof. + intros wf. + unfold gen_transform_env. + destruct Σ as [Σ ? ? ?]; cbn in *. + induction Σ in map, repr, wf0, wf |- * => //. + cbn. f_equal. + destruct a as [kn d]; unfold on_snd; cbn. f_equal. symmetry. + eapply wellformed_gen_transform_decl_extends => //. cbn. now depelim wf. + now apply extends_cons_wf. now depelim wf. + erewrite <- (IHΣ (EnvMap.of_global_env _) (EnvMap.repr_global_env _) (fresh_globals_cons_inv wf0)). + set (nΣ := {| GlobalContextMap.global_decls := Σ|}). + 2:now depelim wf. + symmetry. cbn. + change Σ with nΣ.(GlobalContextMap.global_decls). + eapply gen_transform_env_extends'; eauto. + now apply extends_cons_wf. now depelim wf. +Qed. + +Import EWellformed. + +Lemma gen_transform_wellformed_irrel {Σ : GlobalContextMap.t} t : + @wf_glob efl Σ -> + forall n, wellformed (efl := efl) Σ n t -> + wellformed (efl := efl) (gen_transform_env Σ) n t. +Proof. + intros wfΣ. induction t using EInduction.term_forall_list_ind; cbn => //. + all:try solve [intros; unfold wf_fix_gen in *; rtoProp; intuition eauto; solve_all]. + - rewrite lookup_env_gen_transform //. + destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. destruct (cst_body c); cbn; eauto. + - rewrite lookup_env_gen_transform //. + destruct lookup_env eqn:hl => // /=; intros; rtoProp; eauto. + destruct g eqn:hg => /= //; intros; rtoProp; eauto. + repeat split; eauto. destruct cstr_as_blocks; rtoProp; repeat split; eauto. solve_all. + - rewrite lookup_env_gen_transform //. + destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. subst g. + destruct nth_error => /= //. + intros; rtoProp; intuition auto; solve_all. + - rewrite lookup_env_gen_transform //. + destruct lookup_env eqn:hl => // /=; intros; rtoProp; repeat split; eauto. + destruct g eqn:hg => /= //. +Qed. + +Lemma gen_transform_wellformed_decl_irrel {Σ : GlobalContextMap.t} d : + @wf_glob efl Σ -> + wf_global_decl (efl:= efl) Σ d -> + wf_global_decl (efl := efl) (gen_transform_env Σ) d. +Proof. + intros wf; destruct d => /= //. + destruct (cst_body c) => /= //. + now eapply gen_transform_wellformed_irrel. +Qed. + +Context {GTWF : GenTransformWf efl efl' gen_transform}. + + +Definition Pre_decl Σ d := match d with ConstantDecl cb => match cb.(cst_body) with + | Some b => gen_transform_pre Σ b | _ => True end | _ => True end. + +Lemma gen_transform_decl_wf {Σ : GlobalContextMap.t} : + has_tBox -> has_tRel -> @wf_glob efl Σ -> + forall d, @wf_global_decl efl Σ d -> Pre_decl Σ d -> + wf_global_decl (efl := efl') (gen_transform_env Σ) (gen_transform_decl Σ d). +Proof. + intros hasb hasr wf d. + intros hd. intros pre. + move: hd. + destruct d => /= //. cbn in pre. + destruct (cst_body c) => /= //. + intros hwf. eapply gen_transform_wellformed => //. apply gtwf_axioms_efl. auto. + destruct m => //. cbn. unfold wf_minductive. + cbn. move/andP => [] hp //. rtoProp. solve_all. + eapply orb_true_iff. eapply orb_true_iff in hp as []; eauto. + left. eapply gtwf_cstrs_efl. now rewrite H. +Qed. + +Lemma fresh_global_gen_transform_env {Σ : GlobalContextMap.t} kn : + fresh_global kn Σ -> + fresh_global kn (gen_transform_env Σ). +Proof. + destruct Σ as [Σ ? ? ?]. + induction Σ in map, repr, wf |- *; auto. + intros fr; depelim fr. red. + now eapply Forall_map; cbn. +Qed. + +Fixpoint Pre_glob Σ : EnvMap.fresh_globals Σ -> Prop := + match Σ with + | nil => fun _ => True + | (kn, d) :: Σ => fun HΣ => + let Σg := GlobalContextMap.make Σ (fresh_globals_cons_inv HΣ) in + Pre_decl Σg d /\ Pre_glob Σ (fresh_globals_cons_inv HΣ) + end. + +Import GlobalContextMap (repr, map, global_decls, wf). +Lemma gen_transform_env_wf {Σ : GlobalContextMap.t} : + has_tBox -> has_tRel -> + @wf_glob efl Σ -> Pre_glob Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf) -> wf_glob (efl := efl') (gen_transform_env Σ). +Proof. + intros hasb hasrel wfg pre. + rewrite gen_transform_env_eq //. + destruct Σ as [Σ ? ? ?]. cbn in *. revert pre. + induction Σ in map0, repr0, wf0, wfg |- *; auto; cbn; constructor; auto. + - eapply IHΣ. eapply EnvMap.repr_global_env. now depelim wfg. + destruct a. apply pre. + - cbn. + set (Σp := GlobalContextMap.make Σ _). + specialize (IHΣ (GlobalContextMap.map Σp) (GlobalContextMap.repr Σp) (fresh_globals_cons_inv wf0)). + rewrite /= -(gen_transform_env_eq Σp) => //. now depelim wfg. + eapply gen_transform_decl_wf => //. now depelim wfg. cbn. now depelim wfg. + destruct a. apply pre. + - cbn. + set (Σp := {| global_decls := Σ; map := EnvMap.of_global_env Σ; repr := EnvMap.repr_global_env _; + wf := (fresh_globals_cons_inv wf0) |}). + rewrite /= -(gen_transform_env_eq Σp) => //. now depelim wfg. + eapply fresh_global_gen_transform_env. cbn. now depelim wf0. +Qed. + +(* Definition gen_transform_program (p : eprogram_env) := + (gen_transform_env p.1, gen_transform p.1 p.2). + +Definition gen_transform_program_wf (p : eprogram_env) {hastbox : has_tBox} {hastrel : has_tRel} : + wf_eprogram_env efl p -> wf_eprogram (efl') (gen_transform_program p). +Proof. + intros []; split. + now eapply gen_transform_env_wf. + cbn. eapply gen_transform_wellformed_irrel => //. now eapply gen_transform_wellformed. +Qed. *) + +End sec. \ No newline at end of file diff --git a/erasure/theories/EGenericMapEnv.v b/erasure/theories/EGenericMapEnv.v index 85a0f7446..c5a0b4853 100644 --- a/erasure/theories/EGenericMapEnv.v +++ b/erasure/theories/EGenericMapEnv.v @@ -55,16 +55,16 @@ Variable efl : EEnvFlags. Hypothesis wellformed_gen_transform_extends : forall {Σ : global_context} t, forall n, EWellformed.wellformed Σ n t -> -forall {Σ' : global_context}, extends Σ Σ' -> wf_glob Σ' -> +forall {Σ' : global_context}, extends Σ Σ' -> wf_glob Σ -> wf_glob Σ' -> gen_transform Σ t = gen_transform Σ' t. Lemma wellformed_gen_transform_decl_extends {Σ : global_context} t : wf_global_decl Σ t -> -forall {Σ' : global_context}, extends Σ Σ' -> wf_glob Σ' -> +forall {Σ' : global_context}, extends Σ Σ' -> wf_glob Σ -> wf_glob Σ' -> gen_transform_decl Σ t = gen_transform_decl Σ' t. Proof. destruct t => /= //. -intros wf Σ' ext wf'. f_equal. unfold gen_transform_constant_decl. f_equal. +intros wf Σ' ext wfΣ wf'. f_equal. unfold gen_transform_constant_decl. f_equal. destruct (cst_body c) => /= //. f_equal. now eapply wellformed_gen_transform_extends. Qed. @@ -73,7 +73,7 @@ Lemma lookup_env_gen_transform_env_Some {Σ : global_context} kn d : wf_glob Σ -> lookup_env Σ kn = Some d -> ∑ Σ' : global_context, - [× extends Σ' Σ, wf_global_decl Σ' d & + [× extends_prefix Σ' Σ, wf_global_decl Σ' d & lookup_env (gen_transform_env Σ) kn = Some (gen_transform_decl Σ' d)]. Proof. induction Σ in |- *; simpl; auto => //. @@ -83,7 +83,8 @@ case: eqb_specT => //. now eexists [_]. cbn. now depelim wfg. f_equal. symmetry. eapply wellformed_gen_transform_decl_extends. cbn. now depelim wfg. - cbn. now exists [a]. now cbn. + eapply extends_prefix_extends => //. + cbn. now exists [a]. now depelim wfg. auto. - intros _. cbn in IHΣ. forward IHΣ. now depelim wfg. intros hl. specialize (IHΣ hl) as [Σ'' [ext wfgd hl']]. @@ -96,7 +97,8 @@ case: eqb_specT => //. symmetry. eapply wellformed_gen_transform_decl_extends => //. cbn. eapply lookup_env_In in hin. 2:now depelim wfg. depelim wfg. eapply lookup_env_wellformed; tea. - cbn. now exists [a]. + eapply extends_prefix_extends => //. + cbn. now exists [a]. now depelim wfg. Qed. Lemma lookup_env_map_snd Σ f kn : lookup_env (List.map (on_snd f) Σ) kn = option_map f (lookup_env Σ kn). @@ -120,7 +122,10 @@ intros wf. destruct (lookup_env Σ kn) eqn:hl. - eapply lookup_env_gen_transform_env_Some in hl as [Σ' [ext wf' hl']] => /=. rewrite hl'. f_equal. - eapply wellformed_gen_transform_decl_extends; eauto. auto. + eapply wellformed_gen_transform_decl_extends; eauto. + eapply extends_prefix_extends; auto. + eapply extends_wf_glob; eauto. + auto. - cbn. now eapply lookup_env_gen_transform_env_None in hl. Qed. @@ -195,28 +200,45 @@ Proof. rewrite wellformed_mkApps //. eapply andP. Qed. +Lemma extends_cons_inv Σ kn d Σ' : + extends ((kn, d) :: Σ) Σ' -> + fresh_global kn Σ -> + extends Σ Σ'. +Proof. + intros ext fr kn' d' hl. + apply ext. rewrite elookup_env_cons_fresh => //. + intros <-. now apply lookup_env_Some_fresh in hl. +Qed. + +Lemma extends_cons_wf Σ a : + wf_glob (a :: Σ) -> + extends Σ (a :: Σ). +Proof. + intros hwf; depelim hwf. + now apply extends_fresh. +Qed. + Lemma gen_transform_env_extends' {Σ Σ' : global_context} : extends Σ Σ' -> + wf_glob Σ -> wf_glob Σ' -> List.map (on_snd (gen_transform_decl Σ)) Σ = List.map (on_snd (gen_transform_decl Σ')) Σ. Proof. intros ext. - move=> wfΣ. - assert (Hext : extends Σ Σ); auto. now exists []. - assert (Hwfg : wf_glob Σ). - { eapply extends_wf_glob. exact ext. tea. } - revert Hext Hwfg. + move=> wfΣ wfΣ'. + assert (Hext : extends Σ Σ); auto. red; tauto. + generalize wfΣ. + revert Hext. generalize Σ at 1 3 5 6. intros Σ''. induction Σ'' => //. cbn. intros hin wfg. depelim wfg. f_equal. - 2:{ eapply IHΣ'' => //. destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. } + 2:{ eapply IHΣ'' => //. now apply extends_cons_inv in hin. } unfold on_snd. cbn. f_equal. eapply wellformed_gen_transform_decl_extends => //. cbn. - eapply extends_wf_global_decl. 3:tea. - eapply extends_wf_glob; tea. - destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. + eapply extends_wf_global_decl => //. 2:tea. + now apply extends_cons_inv in hin. Qed. Lemma gen_transform_env_eq (Σ : global_context) : wf_glob Σ -> gen_transform_env Σ = gen_transform_env' Σ. @@ -226,11 +248,12 @@ Proof. induction Σ => //. cbn. f_equal. destruct a as [kn d]; unfold on_snd; cbn. f_equal. symmetry. - eapply wellformed_gen_transform_decl_extends => //. cbn. now depelim wf. cbn. now exists [(kn, d)]. cbn. + eapply wellformed_gen_transform_decl_extends => //. cbn. now depelim wf. + now apply extends_cons_wf. now depelim wf. erewrite <- IHΣ. 2:now depelim wf. symmetry. eapply gen_transform_env_extends'; eauto. - cbn. now exists [a]. + now apply extends_cons_wf. now depelim wf. Qed. Variable Pre : global_context -> term -> Prop. diff --git a/erasure/theories/EGlobalEnv.v b/erasure/theories/EGlobalEnv.v index 97b6cf818..18887b1f8 100644 --- a/erasure/theories/EGlobalEnv.v +++ b/erasure/theories/EGlobalEnv.v @@ -183,7 +183,10 @@ Definition closed_env (Σ : EAst.global_declarations) := (** Environment extension and uniqueness of declarations in well-formed global environments *) -Definition extends (Σ Σ' : global_declarations) := ∑ Σ'', Σ' = (Σ'' ++ Σ)%list. +Definition extends (Σ Σ' : global_declarations) := + (forall kn decl, lookup_env Σ kn = Some decl -> lookup_env Σ' kn = Some decl). + +Definition extends_prefix (Σ Σ' : global_declarations) := ∑ Σ'', Σ' = Σ'' ++ Σ. Definition fresh_global kn (Σ : global_declarations) := Forall (fun x => x.1 <> kn) Σ. diff --git a/erasure/theories/EGlobalFragment.v b/erasure/theories/EGlobalFragment.v new file mode 100644 index 000000000..4a287eba6 --- /dev/null +++ b/erasure/theories/EGlobalFragment.v @@ -0,0 +1,143 @@ +From MetaCoq.Utils Require Import utils. +From MetaCoq.Common Require Import config Kernames uGraph. +From MetaCoq.Erasure Require Import EGlobalEnv ErasureFunction. +From MetaCoq.Erasure Require Import EAstUtils EArities Extract Prelim EDeps ErasureProperties ErasureCorrectness. +From MetaCoq.SafeChecker Require Import PCUICErrors PCUICWfEnv PCUICSafeReduce PCUICSafeRetyping PCUICRetypingEnvIrrelevance. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICPrimitive + PCUICReduction PCUICReflect PCUICWeakeningEnv PCUICWeakeningEnvTyp PCUICCasesContexts + PCUICWeakeningConv PCUICWeakeningTyp PCUICContextConversionTyp PCUICTyping PCUICGlobalEnv PCUICInversion PCUICGeneration + PCUICConfluence PCUICConversion PCUICUnivSubstitutionTyp PCUICCumulativity PCUICSR PCUICSafeLemmata PCUICNormalization + PCUICValidity PCUICPrincipality PCUICElimination PCUICOnFreeVars PCUICWellScopedCumulativity PCUICSN PCUICEtaExpand. + +(* Σ is a fragment of Σ' if every lookup in Σ gives the same result as Σ' *) +Definition eglobal_env_fragment Σ Σ' : + (forall kn decl, lookup_env Σ kn = Some decl -> lookup_env Σ' kn = Some decl). + +(* we use the [match] trick to get typeclass resolution to pick up the right instances without leaving any evidence in the resulting term, and without having to pass them manually everywhere *) +Notation NormalizationIn_erase_global_decls X decls + := (match extraction_checker_flags, extraction_normalizing return _ with + | cf, no + => forall n, + n < List.length decls -> + let X' := iter abstract_pop_decls (S n) X in + forall kn cb pf, + List.hd_error (List.skipn n decls) = Some (kn, ConstantDecl cb) -> + let Xext := abstract_make_wf_env_ext X' (cst_universes cb) pf in + forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ + end) + (only parsing). + +Program Fixpoint erase_global + {X_type : abstract_env_impl} (X : X_type.π1) (decls : global_declarations) + {normalization_in : NormalizationIn_erase_global_decls X decls} + (prop : forall Σ : global_env, abstract_env_rel X Σ -> Σ.(declarations) = decls) + : E.global_declarations := + match decls with + | [] => [] + | (kn, ConstantDecl cb) :: decls => + let X' := abstract_pop_decls X in + let Xext := abstract_make_wf_env_ext X' (cst_universes cb) _ in + let normalization_in' : @id Prop _ := _ in (* hack to avoid program erroring out on unresolved tc *) + let cb' := @erase_constant_body X_type Xext normalization_in' cb _ in + let X'' := erase_global X' decls _ in + ((kn, E.ConstantDecl (fst cb')) :: X'') + | (kn, InductiveDecl mib) :: decls => + let X' := abstract_pop_decls X in + let mib' := erase_mutual_inductive_body mib in + let X'' := erase_global X' decls _ in + ((kn, E.InductiveDecl mib') :: X'') + + end. +Next Obligation. + pose proof (abstract_env_wf _ H) as [wf]. + pose proof (abstract_env_exists X) as [[? HX]]. + pose proof (abstract_env_wf _ HX) as [wfX]. + assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). + { now eexists. } + destruct (abstract_pop_decls_correct X decls prop' _ _ HX H) as [? []]. + clear H. specialize (prop _ HX). destruct x, Σ, H0; cbn in *. + subst. sq. destruct wfX. depelim o0. destruct o1. split => //. +Qed. +Next Obligation. + cbv [id]. + unshelve eapply (normalization_in 0); cbn; try reflexivity; try lia. +Defined. +Next Obligation. + pose proof (abstract_env_ext_wf _ H) as [wf]. + pose proof (abstract_env_exists (abstract_pop_decls X)) as [[? HX']]. + pose proof (abstract_env_wf _ HX') as [wfX']. + pose proof (abstract_env_exists X) as [[? HX]]. + pose proof (abstract_env_wf _ HX) as [wfX]. + assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). + { now eexists. } + pose proof (abstract_pop_decls_correct X decls prop' _ _ HX HX') as [? []]. + pose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ HX' H). + clear H HX'. specialize (prop _ HX). destruct x, Σ as [[] u], H0; cbn in *. + subst. sq. inversion H3. subst. clear H3. destruct wfX. cbn in *. + rewrite prop in o0. depelim o0. destruct o1. exact on_global_decl_d. +Qed. +Next Obligation. + unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. +Defined. +Next Obligation. + pose proof (abstract_env_exists X) as [[? HX]]. + assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). + { now eexists. } + now pose proof (abstract_pop_decls_correct X decls prop' _ _ HX H). +Qed. +Next Obligation. + unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. +Defined. +Next Obligation. +pose proof (abstract_env_exists X) as [[? HX]]. +assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). +{ now eexists. } +now pose proof (abstract_pop_decls_correct X decls prop' _ _ HX H). +Qed. +Next Obligation. + unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. +Defined. +Next Obligation. + pose proof (abstract_env_exists X) as [[? HX]]. + assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). + { now eexists. } + now pose proof (abstract_pop_decls_correct X decls prop' _ _ HX H). +Qed. +Next Obligation. + unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. +Defined. +Next Obligation. +pose proof (abstract_env_exists X) as [[? HX]]. +assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). +{ now eexists. } +now pose proof (abstract_pop_decls_correct X decls prop' _ _ HX H). +Qed. + + + +Lemma erase_global_fast_fragment + X_type deps X decls {normalization_in} (prf:forall Σ : global_env, abstract_env_rel X Σ -> declarations Σ = decls) : + eglobal_env_fragment + (@ErasureFunction.erase_global_fast X_type deps X decls normalization_in prf) + (@E) + + + +Lemma erase_program_fst {guard : abstract_guard_impl} (p p' : pcuic_program) + {normalization_in : PCUICTyping.wf_ext p.1 -> PCUICSN.NormalizationIn p.1} + {normalization_in' : PCUICTyping.wf_ext p'.1 -> PCUICSN.NormalizationIn p'.1} +{normalization_in_adjust_universes : PCUICTyping.wf_ext p.1 -> + PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p.1} +{normalization_in_adjust_universes' : PCUICTyping.wf_ext p'.1 -> +PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn + p'.1} wt wt' : +p.1 = p'.1 -> +(erase_program p wt).1 = (erase_program p' wt').1. +Proof. +destruct p, p'; intros; subst. cbn in H. subst g0. +unfold erase_program. +assert (map_squash fst wt = map_squash fst wt') by apply proof_irrelevance. +rewrite -H. cbn. +epose proof ErasureFunction.erase_irrel_global_env. +Abort.Lemma erase_global_fragment Σ + diff --git a/erasure/theories/EInlineProjections.v b/erasure/theories/EInlineProjections.v index bcb098d42..178225ad4 100644 --- a/erasure/theories/EInlineProjections.v +++ b/erasure/theories/EInlineProjections.v @@ -33,6 +33,30 @@ Definition switch_no_params (efl : EEnvFlags) := |}. Definition flags_after_projs := (switch_no_params all_env_flags). Local Existing Instance flags_after_projs. + +Definition disable_projections_term_flags (et : ETermFlags) := + {| has_tBox := has_tBox + ; has_tRel := has_tRel + ; has_tVar := has_tVar + ; has_tEvar := has_tEvar + ; has_tLambda := has_tLambda + ; has_tLetIn := has_tLetIn + ; has_tApp := has_tApp + ; has_tConst := has_tConst + ; has_tConstruct := has_tConstruct + ; has_tCase := true + ; has_tProj := false + ; has_tFix := has_tFix + ; has_tCoFix := has_tCoFix + ; has_tPrim := has_tPrim + |}. + +Definition disable_projections_env_flag (efl : EEnvFlags) := + {| has_axioms := efl.(@has_axioms); + term_switches := disable_projections_term_flags term_switches; + has_cstr_params := efl.(@has_cstr_params) ; + cstr_as_blocks := efl.(@cstr_as_blocks) |}. + Arguments lookup_projection : simpl never. Arguments GlobalContextMap.lookup_projection : simpl never. @@ -323,7 +347,7 @@ Lemma lookup_env_optimize_env_Some {efl : EEnvFlags} {Σ : GlobalContextMap.t} k wf_glob Σ -> GlobalContextMap.lookup_env Σ kn = Some d -> ∑ Σ' : GlobalContextMap.t, - [× extends Σ' Σ, wf_global_decl Σ' d & + [× extends_prefix Σ' Σ, wf_global_decl Σ' d & lookup_env (optimize_env Σ) kn = Some (optimize_decl Σ' d)]. Proof. rewrite GlobalContextMap.lookup_env_spec. @@ -336,7 +360,8 @@ Proof. now eexists [_]. cbn. now depelim wfg. f_equal. symmetry. eapply wellformed_optimize_decl_extends. cbn. now depelim wfg. - cbn. now exists [a]. now cbn. + eapply extends_prefix_extends. + cbn. now exists [a]. now cbn. now cbn. - intros _. set (Σ' := GlobalContextMap.make Σ (fresh_globals_cons_inv wf)). specialize (IHΣ (GlobalContextMap.map Σ') (GlobalContextMap.repr Σ') (GlobalContextMap.wf Σ')). @@ -351,7 +376,8 @@ Proof. symmetry. eapply wellformed_optimize_decl_extends => //. cbn. eapply lookup_env_In in hin. 2:now depelim wfg. depelim wfg. eapply lookup_env_wellformed; tea. - cbn. now exists [a]. + eapply extends_prefix_extends. + cbn. now exists [a]. now cbn. Qed. Lemma lookup_env_map_snd Σ f kn : lookup_env (List.map (on_snd f) Σ) kn = option_map f (lookup_env Σ kn). @@ -378,7 +404,8 @@ Proof. destruct (GlobalContextMap.lookup_env Σ kn) eqn:hl. - eapply lookup_env_optimize_env_Some in hl as [Σ' [ext wf' hl']] => /=. rewrite hl'. f_equal. - eapply wellformed_optimize_decl_extends; eauto. auto. + eapply wellformed_optimize_decl_extends; eauto. + now eapply extends_prefix_extends. auto. - cbn. now eapply lookup_env_optimize_env_None in hl. Qed. @@ -729,88 +756,7 @@ Proof. apply optimize_expanded_irrel. Qed. -Lemma optimize_env_extends' {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : - extends Σ Σ' -> - wf_glob Σ' -> - List.map (on_snd (optimize_decl Σ)) Σ.(GlobalContextMap.global_decls) = - List.map (on_snd (optimize_decl Σ')) Σ.(GlobalContextMap.global_decls). -Proof. - intros ext. - destruct Σ as [Σ map repr wf]; cbn in *. - move=> wfΣ. - assert (extends Σ Σ); auto. now exists []. - assert (wf_glob Σ). - { eapply extends_wf_glob. exact ext. tea. } - revert H H0. - generalize Σ at 1 3 5 6. intros Σ''. - induction Σ'' => //. cbn. - intros hin wfg. depelim wfg. - f_equal. - 2:{ eapply IHΣ'' => //. destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. } - unfold on_snd. cbn. f_equal. - eapply wellformed_optimize_decl_extends => //. cbn. - eapply extends_wf_global_decl. 3:tea. - eapply extends_wf_glob; tea. - destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. -Qed. - -Lemma optimize_env_eq {efl : EEnvFlags} (Σ : GlobalContextMap.t) : wf_glob Σ -> optimize_env Σ = optimize_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). -Proof. - intros wf. - unfold optimize_env. - destruct Σ; cbn. cbn in wf. - induction global_decls in map, repr, wf0, wf |- * => //. - cbn. f_equal. - destruct a as [kn d]; unfold on_snd; cbn. f_equal. symmetry. - eapply wellformed_optimize_decl_extends => //. cbn. now depelim wf. cbn. now exists [(kn, d)]. cbn. - set (Σg' := GlobalContextMap.make global_decls (fresh_globals_cons_inv wf0)). - erewrite <- (IHglobal_decls (GlobalContextMap.map Σg') (GlobalContextMap.repr Σg')). - 2:now depelim wf. - set (Σg := {| GlobalContextMap.global_decls := _ :: _ |}). - symmetry. eapply (optimize_env_extends' (Σ := Σg') (Σ' := Σg)) => //. - cbn. now exists [a]. -Qed. - -Lemma optimize_env_expanded {efl : EEnvFlags} {Σ : GlobalContextMap.t} : - wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (optimize_env Σ). -Proof. - unfold expanded_global_env; move=> wfg. - rewrite optimize_env_eq //. - destruct Σ as [Σ map repr wf]. cbn in *. - clear map repr. - induction 1; cbn; constructor; auto. - cbn in IHexpanded_global_declarations. - unshelve eapply IHexpanded_global_declarations. now depelim wfg. cbn. - set (Σ' := GlobalContextMap.make _ _). - rewrite -(optimize_env_eq Σ'). cbn. now depelim wfg. - eapply (optimize_expanded_decl_irrel (Σ := Σ')). now depelim wfg. - now unshelve eapply (optimize_expanded_decl (Σ:=Σ')). -Qed. - -Definition disable_projections_term_flags (et : ETermFlags) := - {| has_tBox := has_tBox - ; has_tRel := has_tRel - ; has_tVar := has_tVar - ; has_tEvar := has_tEvar - ; has_tLambda := has_tLambda - ; has_tLetIn := has_tLetIn - ; has_tApp := has_tApp - ; has_tConst := has_tConst - ; has_tConstruct := has_tConstruct - ; has_tCase := true - ; has_tProj := false - ; has_tFix := has_tFix - ; has_tCoFix := has_tCoFix - ; has_tPrim := has_tPrim - |}. - -Definition disable_projections_env_flag (efl : EEnvFlags) := - {| has_axioms := efl.(@has_axioms); - term_switches := disable_projections_term_flags term_switches; - has_cstr_params := efl.(@has_cstr_params) ; - cstr_as_blocks := efl.(@cstr_as_blocks) |}. - -Lemma optimize_wellformed {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : +Lemma optimize_wellformed_term {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : has_tBox -> has_tRel -> wf_glob Σ -> EWellformed.wellformed Σ n t -> EWellformed.wellformed (efl := disable_projections_env_flag efl) Σ n (optimize Σ t). @@ -862,52 +808,81 @@ Proof. intros; rtoProp; intuition auto; solve_all. Qed. -Lemma optimize_wellformed_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} d : +Lemma optimize_wellformed {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : + has_tBox -> has_tRel -> + wf_glob Σ -> EWellformed.wellformed Σ n t -> + EWellformed.wellformed (efl := disable_projections_env_flag efl) (optimize_env Σ) n (optimize Σ t). +Proof. + intros. apply optimize_wellformed_irrel => //. + now apply optimize_wellformed_term. +Qed. + +From MetaCoq.Erasure Require Import EGenericGlobalMap. + +#[local] +Instance GT : GenTransform := { gen_transform := optimize }. +#[local] +Instance GTExt efl : GenTransformExtends efl (disable_projections_env_flag efl) GT. +Proof. + intros Σ t n wfΣ Σ' ext wf wf'. + unfold gen_transform, GT. + eapply wellformed_optimize_extends; tea. +Qed. +#[local] +Instance GTWf efl : GenTransformWf efl (disable_projections_env_flag efl) GT. +Proof. + refine {| gen_transform_pre := fun _ _ => True |}; auto. + intros Σ n t hasb hasr _ wfΣ wft. + now apply optimize_wellformed. +Defined. + +Lemma optimize_env_extends' {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : + extends Σ Σ' -> wf_glob Σ -> - wf_global_decl (efl:= disable_projections_env_flag efl) Σ d -> - wf_global_decl (efl := disable_projections_env_flag efl) (optimize_env Σ) d. + wf_glob Σ' -> + List.map (on_snd (optimize_decl Σ)) Σ.(GlobalContextMap.global_decls) = + List.map (on_snd (optimize_decl Σ')) Σ.(GlobalContextMap.global_decls). Proof. - intros wf; destruct d => /= //. - destruct (cst_body c) => /= //. - now eapply optimize_wellformed_irrel. + intros. + epose proof (gen_transform_env_extends' (gt := GTExt efl)). + now eapply H2. Qed. -Lemma optimize_decl_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : - has_tBox -> has_tRel -> wf_glob Σ -> - forall d, wf_global_decl Σ d -> - wf_global_decl (efl := disable_projections_env_flag efl) (optimize_env Σ) (optimize_decl Σ d). +Lemma optimize_env_eq {efl : EEnvFlags} (Σ : GlobalContextMap.t) : wf_glob Σ -> optimize_env Σ = optimize_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). Proof. - intros hasb hasr wf d. - intros hd. - eapply optimize_wellformed_decl_irrel; tea; eauto. - move: hd. - destruct d => /= //. - destruct (cst_body c) => /= //. - intros hwf. eapply optimize_wellformed => //. + intros wf. + now apply (gen_transform_env_eq (gt := GTExt efl)). Qed. -Lemma fresh_global_optimize_env {Σ : GlobalContextMap.t} kn : - fresh_global kn Σ -> - fresh_global kn (optimize_env Σ). +Lemma optimize_env_expanded {efl : EEnvFlags} {Σ : GlobalContextMap.t} : + wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (optimize_env Σ). Proof. - destruct Σ as [Σ map repr wf]; cbn in *. + unfold expanded_global_env; move=> wfg. + rewrite optimize_env_eq //. + destruct Σ as [Σ map repr wf]. cbn in *. + clear map repr. induction 1; cbn; constructor; auto. - now eapply Forall_map; cbn. + cbn in IHexpanded_global_declarations. + unshelve eapply IHexpanded_global_declarations. now depelim wfg. cbn. + set (Σ' := GlobalContextMap.make _ _). + rewrite -(optimize_env_eq Σ'). cbn. now depelim wfg. + eapply (optimize_expanded_decl_irrel (Σ := Σ')). now depelim wfg. + now unshelve eapply (optimize_expanded_decl (Σ:=Σ')). +Qed. + +Lemma Pre_glob efl Σ wf : Pre_glob (GTWF:=GTWf efl) Σ wf. +Proof. + induction Σ => //. destruct a as [kn d]; cbn. + split => //. destruct d as [[[|]]|] => //=. Qed. Lemma optimize_env_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : has_tBox -> has_tRel -> wf_glob Σ -> wf_glob (efl := disable_projections_env_flag efl) (optimize_env Σ). Proof. - intros hasb hasrel. - intros wfg. rewrite optimize_env_eq //. - destruct Σ as [Σ map repr wf]; cbn in *. - clear map repr. - induction wfg; cbn; constructor; auto. - - rewrite /= -(optimize_env_eq (GlobalContextMap.make Σ (fresh_globals_cons_inv wf))) //. - eapply optimize_decl_wf => //. - - rewrite /= -(optimize_env_eq (GlobalContextMap.make Σ (fresh_globals_cons_inv wf))) //. - now eapply fresh_global_optimize_env. + intros hasb hasre wfg. + eapply (gen_transform_env_wf (gt := GTExt efl)) => //. + apply Pre_glob. Qed. Definition optimize_program (p : eprogram_env) := @@ -918,7 +893,7 @@ Definition optimize_program_wf {efl : EEnvFlags} (p : eprogram_env) {hastbox : h Proof. intros []; split. now eapply optimize_env_wf. - cbn. eapply optimize_wellformed_irrel => //. now eapply optimize_wellformed. + cbn. now eapply optimize_wellformed. Qed. Definition optimize_program_expanded {efl} (p : eprogram_env) : diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 70ad8701c..0cd56240e 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -405,7 +405,7 @@ Lemma lookup_env_remove_match_on_box_env_Some {efl : EEnvFlags} {Σ : GlobalCont wf_glob Σ -> GlobalContextMap.lookup_env Σ kn = Some d -> ∑ Σ' : GlobalContextMap.t, - [× extends Σ' Σ, wf_global_decl Σ' d & + [× extends_prefix Σ' Σ, wf_global_decl Σ' d & lookup_env (remove_match_on_box_env Σ) kn = Some (remove_match_on_box_decl Σ' d)]. Proof. rewrite GlobalContextMap.lookup_env_spec. @@ -418,7 +418,8 @@ Proof. now eexists [_]. cbn. now depelim wfg. f_equal. symmetry. eapply wellformed_remove_match_on_box_decl_extends. cbn. now depelim wfg. - cbn. now exists [a]. now cbn. + eapply extends_prefix_extends. + cbn. now exists [a]. now cbn. now cbn. - intros _. set (Σ' := GlobalContextMap.make Σ (fresh_globals_cons_inv wf)). specialize (IHΣ (GlobalContextMap.map Σ') (GlobalContextMap.repr Σ') (GlobalContextMap.wf Σ')). @@ -433,7 +434,7 @@ Proof. symmetry. eapply wellformed_remove_match_on_box_decl_extends => //. cbn. eapply lookup_env_In in hin. 2:now depelim wfg. depelim wfg. eapply lookup_env_wellformed; tea. - cbn. now exists [a]. + cbn. eapply extends_prefix_extends => //. now exists [a]. Qed. Lemma lookup_env_map_snd Σ f kn : lookup_env (List.map (on_snd f) Σ) kn = option_map f (lookup_env Σ kn). @@ -460,7 +461,8 @@ Proof. destruct (GlobalContextMap.lookup_env Σ kn) eqn:hl. - eapply lookup_env_remove_match_on_box_env_Some in hl as [Σ' [ext wf' hl']] => /=. rewrite hl'. f_equal. - eapply wellformed_remove_match_on_box_decl_extends; eauto. auto. + eapply wellformed_remove_match_on_box_decl_extends; eauto. + now eapply extends_prefix_extends. auto. - cbn. now eapply lookup_env_remove_match_on_box_env_None in hl. Qed. @@ -780,7 +782,7 @@ Proof. Qed. Lemma remove_match_on_box_env_extends' {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : - extends Σ Σ' -> + extends_prefix Σ Σ' -> wf_glob Σ' -> List.map (on_snd (remove_match_on_box_decl Σ)) Σ.(GlobalContextMap.global_decls) = List.map (on_snd (remove_match_on_box_decl Σ')) Σ.(GlobalContextMap.global_decls). @@ -788,7 +790,7 @@ Proof. intros ext. destruct Σ as [Σ map repr wf]; cbn in *. move=> wfΣ. - assert (extends Σ Σ); auto. now exists []. + assert (extends_prefix Σ Σ); auto. now exists []. assert (wf_glob Σ). { eapply extends_wf_glob. exact ext. tea. } revert H H0. @@ -801,7 +803,10 @@ Proof. eapply wellformed_remove_match_on_box_decl_extends => //. cbn. eapply extends_wf_global_decl. 3:tea. eapply extends_wf_glob; tea. + eapply extends_prefix_extends. destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. + eapply extends_wf_glob; tea. cbn. + eapply extends_prefix_extends; tea. Qed. Lemma remove_match_on_box_env_eq {efl : EEnvFlags} (Σ : GlobalContextMap.t) : wf_glob Σ -> remove_match_on_box_env Σ = remove_match_on_box_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). diff --git a/erasure/theories/EWellformed.v b/erasure/theories/EWellformed.v index 61a253be2..693297e44 100644 --- a/erasure/theories/EWellformed.v +++ b/erasure/theories/EWellformed.v @@ -415,18 +415,9 @@ Lemma extends_lookup {efl} {Σ Σ' c decl} : 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Σ''. - + now inv wfΣ'. - + intros HΣ. specialize (IHΣ'' HΣ). - inv wfΣ'. simpl in *. - case: eqb_spec; intros e; subst; auto. - apply lookup_env_Some_fresh in IHΣ''; contradiction. + intros wfΣ' Hl; apply Hl. Qed. - Lemma extends_lookup_constructor {efl} {Σ Σ'} : wf_glob Σ' -> extends Σ Σ' -> forall ind c b, lookup_constructor Σ ind c = Some b -> @@ -482,14 +473,33 @@ Proof. now eapply extends_wellformed. Qed. +Lemma extends_fresh Σ kn d : + fresh_global kn Σ -> + extends Σ ((kn, d) :: Σ). +Proof. + induction 1 in d |- *. red. cbn => //. + intros kn' decl. + cbn. + destruct (eqb_spec x.1 kn) => //. + destruct (eqb_spec kn' x.1) => //. + - intros [= <-]. subst kn'. + destruct (eqb_spec x.1 kn) => //. + - intros hl. + destruct (eqb_spec kn' kn) => //. + subst. eapply (IHForall d) in hl. + now cbn in hl; move: hl; rewrite eqb_refl. +Qed. + Lemma lookup_env_wellformed {efl} {Σ kn decl} : wf_glob Σ -> EGlobalEnv.lookup_env Σ kn = Some decl -> wf_global_decl Σ decl. Proof. induction Σ; cbn => //. intros wf. depelim wf => /=. destruct (eqb_spec kn kn0). - move=> [= <-]. eapply extends_wf_global_decl; tea. constructor; auto. now eexists [_]. + move=> [= <-]. eapply extends_wf_global_decl; tea. constructor; auto. + now apply extends_fresh. intros hn. eapply extends_wf_global_decl; tea. 3:eapply IHΣ => //. - now constructor. now eexists [_]. + now constructor. + now apply extends_fresh. Qed. From e09e64d8000785ece4336394a66bb4c501e1cbbb Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 21 Aug 2023 14:44:09 +0200 Subject: [PATCH 37/61] Allow modifying inductive decls in generic weakenability translation proofs. Now parameter stripping is a weakenable translation. --- erasure/theories/EConstructorsAsBlocks.v | 95 ++++++--------- erasure/theories/EGenericGlobalMap.v | 61 ++++++---- erasure/theories/EInlineProjections.v | 19 ++- erasure/theories/EOptimizePropDiscr.v | 146 ++++++++++++----------- erasure/theories/ERemoveParams.v | 84 +++++++------ erasure/theories/EWcbvEvalNamed.v | 2 +- 6 files changed, 208 insertions(+), 199 deletions(-) diff --git a/erasure/theories/EConstructorsAsBlocks.v b/erasure/theories/EConstructorsAsBlocks.v index 98368c0d1..829d0d089 100644 --- a/erasure/theories/EConstructorsAsBlocks.v +++ b/erasure/theories/EConstructorsAsBlocks.v @@ -817,78 +817,55 @@ Proof. eapply transform_wellformed'; eauto. Qed. -(*Lemma optimize_wellformed_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} d : - wf_glob Σ -> - wf_global_decl (efl := env_flags) Σ d -> - wf_global_decl (efl := env_flags_blocks) (transform_blocks_env Σ) d. -Proof. - intros wf; destruct d => /= //. - destruct (cst_body c) => /= //. - intros wf'. eapply transform_wellformed. -Qed. *) +From MetaCoq.Erasure Require Import EGenericGlobalMap. -Lemma optimize_decl_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : - has_cstr_params = false -> - cstr_as_blocks = false -> - has_tApp -> - wf_glob (efl := efl) Σ -> - forall d, - wf_global_decl (efl := efl) Σ d -> - isEtaExp_decl Σ d -> - wf_global_decl (efl := switch_cstr_as_blocks efl) (transform_blocks_env Σ) (transform_blocks_decl Σ d). +#[local] +Instance GT : GenTransform := { gen_transform := transform_blocks }. +#[local] +Instance GTExt efl : has_tApp -> GenTransformExtends efl (switch_cstr_as_blocks efl) GT. Proof. - intros hasp cstrbl hasapp wf d. - destruct d => /= //. - rewrite /isEtaExp_constant_decl. - destruct (cst_body c) => /= //. - intros hwf etat. eapply transform_wellformed => //. + intros Σ t n wfΣ Σ' ext wf wf'. + unfold gen_transform, GT. + eapply transform_blocks_extends; tea. Qed. - -Lemma fresh_global_optimize_env {Σ : GlobalContextMap.t} kn : - fresh_global kn Σ -> - fresh_global kn (transform_blocks_env Σ). +#[local] +Instance GTWf efl : + GenTransformWf efl (switch_cstr_as_blocks efl) GT. Proof. - destruct Σ as [Σ map repr wf]; cbn in *. - induction 1; cbn; constructor; auto. - now eapply Forall_map; cbn. -Qed. - -Lemma fresh_global_map_on_snd Σ f kn : - fresh_global kn Σ -> - fresh_global kn (map (on_snd f) Σ). + refine {| gen_transform_pre := fun Σ t => + has_tApp /\ + has_cstr_params = false /\ + cstr_as_blocks = false /\ + isEtaExp Σ t |}; auto. + intros Σ n t [hasapp [cstrp [cstrb pre]]] wfΣ wft. + apply transform_wellformed => //. +Defined. + +Lemma Pre_glob efl Σ wf : + has_cstr_params = false -> + cstr_as_blocks = false -> + has_tApp -> + EEtaExpanded.isEtaExp_env Σ -> + Pre_glob (GTWF:=GTWf efl) Σ wf. Proof. - induction 1; cbn; constructor; auto. + intros cstrp cstrb happ. + induction Σ => //. destruct a as [kn d]; cbn. + split => //. destruct d as [[[|]]|] => //=. + repeat split => //. move/andP: H => /= [] /=. + now rewrite /isEtaExp_constant_decl /=. + eapply IHΣ. now move/andP: H. Qed. Lemma transform_wf_global {efl : EEnvFlags} {Σ : GlobalContextMap.t} : has_cstr_params = false -> cstr_as_blocks = false -> - has_tApp -> + has_tApp -> has_tBox -> has_tRel -> EEtaExpanded.isEtaExp_env Σ -> wf_glob (efl := efl) Σ -> wf_glob (efl := switch_cstr_as_blocks efl) (transform_blocks_env Σ). Proof. - intros hasp cstrbl hasapp etag wfg. - destruct Σ as [Σ map repr wf]; cbn in *. - revert etag wfg. - assert (extends Σ Σ). now exists []. - revert H. - revert repr wf. generalize Σ at 1 2 4 6 7. - induction Σ; cbn; constructor; auto. - - eapply IHΣ; rtoProp; intuition auto. destruct H. subst Σ0. exists (x ++ [a]). now rewrite -app_assoc. - - set (Σm := {| GlobalContextMap.global_decls := _ |}). - clear IHΣ. - epose proof (EExtends.extends_wf_glob _ H wfg); tea. - depelim H0. - set (Σm' := GlobalContextMap.make Σ (wf_glob_fresh _ H0)). - pose proof (transform_blocks_decl_extends hasapp Σm' _ H1 Σm). - forward H3. cbn. destruct H. subst Σ0. exists (x ++ [(kn, d)]). now rewrite -app_assoc. - specialize (H3 wfg). rewrite -H3. - move/andP: etag => [etad etag]. - unshelve epose proof (@transform_wellformed_decl' _ Σm' d _ _ _ H0 H1 etad); tea. - cbn in H4. unfold Σm. - eapply gen_transform_wellformed_decl_irrel; trea. - - eapply fresh_global_map_on_snd. - eapply EExtends.extends_wf_glob in wfg; tea. now depelim wfg. + intros hasp cstrbl hasapp hasb hasr etag wfg. + unshelve eapply (gen_transform_env_wf (gt := GTExt efl _)) => //. exact hasapp. + eapply Pre_glob => //. Qed. Transparent transform_blocks. diff --git a/erasure/theories/EGenericGlobalMap.v b/erasure/theories/EGenericGlobalMap.v index 4c80c5a2a..4fde99a95 100644 --- a/erasure/theories/EGenericGlobalMap.v +++ b/erasure/theories/EGenericGlobalMap.v @@ -16,7 +16,11 @@ Local Set Keyed Unification. Require Import ssreflect ssrbool. Class GenTransform := - gen_transform : GlobalContextMap.t -> term -> term. + { gen_transform : GlobalContextMap.t -> term -> term; + gen_transform_inductive_decl : mutual_inductive_body -> mutual_inductive_body }. + +Class GenTransformId (G : GenTransform) := + gen_transform_inductive_decl_id : forall idecl, gen_transform_inductive_decl idecl = idecl. Section GenTransformEnv. Context {GT : GenTransform}. @@ -27,7 +31,7 @@ Section GenTransformEnv. Definition gen_transform_decl Σ d := match d with | ConstantDecl cb => ConstantDecl (gen_transform_constant_decl Σ cb) - | InductiveDecl idecl => d + | InductiveDecl idecl => InductiveDecl (gen_transform_inductive_decl idecl) end. Definition gen_transform_env Σ := @@ -52,10 +56,11 @@ Class GenTransformExtends (efl efl' : EEnvFlags) (GT : GenTransform) := Class GenTransformWf (efl efl' : EEnvFlags) (GT : GenTransform) := { gen_transform_pre : GlobalContextMap.t -> term -> Prop; gtwf_axioms_efl : forall _ : is_true (@has_axioms efl), is_true (@has_axioms efl'); - gtwf_cstrs_efl : forall _ : is_true (@has_cstr_params efl), is_true (@has_cstr_params efl'); + gen_transform_inductive_decl_wf : + forall idecl, @wf_minductive efl idecl -> @wf_minductive efl' (gen_transform_inductive_decl idecl); gen_transform_wellformed : forall {Σ : GlobalContextMap.t} n t, - has_tBox -> has_tRel -> gen_transform_pre Σ t -> + gen_transform_pre Σ t -> @wf_glob efl Σ -> @EWellformed.wellformed efl Σ n t -> EWellformed.wellformed (efl := efl') (gen_transform_env Σ) n (gen_transform Σ t) }. @@ -143,6 +148,17 @@ destruct (lookup_env Σ kn) eqn:hl. Qed. +Lemma isFix_mkApps t l : isFix (mkApps t l) = isFix t && match l with [] => true | _ => false end. +Proof. + induction l using rev_ind; cbn. + - now rewrite andb_true_r. + - rewrite mkApps_app /=. now destruct l => /= //; rewrite andb_false_r. +Qed. + +Section GenTransformId. + + Context {gid : GenTransformId gen_transform}. + Lemma is_propositional_gen_transform {Σ : GlobalContextMap.t} ind : @wf_glob efl Σ -> inductive_isprop_and_pars Σ ind = inductive_isprop_and_pars (gen_transform_env Σ) ind. @@ -152,7 +168,8 @@ Proof. rewrite (lookup_env_gen_transform (inductive_mind ind) wf). rewrite /GlobalContextMap.inductive_isprop_and_pars /GlobalContextMap.lookup_inductive /GlobalContextMap.lookup_minductive. - destruct lookup_env as [[decl|]|] => //. + destruct lookup_env as [[decl|]|] => //. cbn. + now rewrite gen_transform_inductive_decl_id. Qed. Lemma is_propositional_cstr_gen_transform {Σ : GlobalContextMap.t} ind c : @@ -164,15 +181,10 @@ Proof. rewrite (lookup_env_gen_transform (inductive_mind ind) wf). rewrite /GlobalContextMap.inductive_isprop_and_pars /GlobalContextMap.lookup_inductive /GlobalContextMap.lookup_minductive. - destruct lookup_env as [[decl|]|] => //. + destruct lookup_env as [[decl|]|] => //=. + now rewrite gen_transform_inductive_decl_id. Qed. -Lemma isFix_mkApps t l : isFix (mkApps t l) = isFix t && match l with [] => true | _ => false end. -Proof. - induction l using rev_ind; cbn. - - now rewrite andb_true_r. - - rewrite mkApps_app /=. now destruct l => /= //; rewrite andb_false_r. -Qed. Lemma lookup_constructor_gen_transform {Σ : GlobalContextMap.t} {ind c} : @wf_glob efl Σ -> @@ -180,7 +192,7 @@ Lemma lookup_constructor_gen_transform {Σ : GlobalContextMap.t} {ind c} : Proof. intros wfΣ. rewrite /lookup_constructor /lookup_inductive /lookup_minductive. rewrite lookup_env_gen_transform // /=. destruct lookup_env => // /=. - destruct g => //. + destruct g => //=. now rewrite gen_transform_inductive_decl_id. Qed. Lemma lookup_projection_gen_transform {Σ : GlobalContextMap.t} {p} : @@ -207,6 +219,8 @@ Proof. rewrite /constructor_isprop_pars_decl. intros -> => /= //. Qed. +End GenTransformId. + Lemma wf_mkApps (ha : has_tApp) Σ k f args : reflect (wellformed Σ k f /\ forallb (wellformed Σ k) args) (wellformed Σ k (mkApps f args)). Proof. rewrite wellformed_mkApps //. eapply andP. @@ -277,7 +291,7 @@ Qed. Import EWellformed. -Lemma gen_transform_wellformed_irrel {Σ : GlobalContextMap.t} t : +Lemma gen_transform_wellformed_irrel {genid : GenTransformId gen_transform} {Σ : GlobalContextMap.t} t : @wf_glob efl Σ -> forall n, wellformed (efl := efl) Σ n t -> wellformed (efl := efl) (gen_transform_env Σ) n t. @@ -290,18 +304,21 @@ Proof. - rewrite lookup_env_gen_transform //. destruct lookup_env eqn:hl => // /=; intros; rtoProp; eauto. destruct g eqn:hg => /= //; intros; rtoProp; eauto. + rewrite gen_transform_inductive_decl_id. repeat split; eauto. destruct cstr_as_blocks; rtoProp; repeat split; eauto. solve_all. - rewrite lookup_env_gen_transform //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. + rewrite gen_transform_inductive_decl_id. destruct nth_error => /= //. intros; rtoProp; intuition auto; solve_all. - rewrite lookup_env_gen_transform //. destruct lookup_env eqn:hl => // /=; intros; rtoProp; repeat split; eauto. destruct g eqn:hg => /= //. + now rewrite gen_transform_inductive_decl_id. Qed. -Lemma gen_transform_wellformed_decl_irrel {Σ : GlobalContextMap.t} d : +Lemma gen_transform_wellformed_decl_irrel {genid : GenTransformId gen_transform} {Σ : GlobalContextMap.t} d : @wf_glob efl Σ -> wf_global_decl (efl:= efl) Σ d -> wf_global_decl (efl := efl) (gen_transform_env Σ) d. @@ -313,25 +330,22 @@ Qed. Context {GTWF : GenTransformWf efl efl' gen_transform}. - Definition Pre_decl Σ d := match d with ConstantDecl cb => match cb.(cst_body) with | Some b => gen_transform_pre Σ b | _ => True end | _ => True end. Lemma gen_transform_decl_wf {Σ : GlobalContextMap.t} : - has_tBox -> has_tRel -> @wf_glob efl Σ -> + @wf_glob efl Σ -> forall d, @wf_global_decl efl Σ d -> Pre_decl Σ d -> wf_global_decl (efl := efl') (gen_transform_env Σ) (gen_transform_decl Σ d). Proof. - intros hasb hasr wf d. + intros wf d. intros hd. intros pre. move: hd. destruct d => /= //. cbn in pre. destruct (cst_body c) => /= //. intros hwf. eapply gen_transform_wellformed => //. apply gtwf_axioms_efl. auto. - destruct m => //. cbn. unfold wf_minductive. - cbn. move/andP => [] hp //. rtoProp. solve_all. - eapply orb_true_iff. eapply orb_true_iff in hp as []; eauto. - left. eapply gtwf_cstrs_efl. now rewrite H. + destruct m => //. cbn. + eapply gen_transform_inductive_decl_wf. Qed. Lemma fresh_global_gen_transform_env {Σ : GlobalContextMap.t} kn : @@ -354,10 +368,9 @@ Fixpoint Pre_glob Σ : EnvMap.fresh_globals Σ -> Prop := Import GlobalContextMap (repr, map, global_decls, wf). Lemma gen_transform_env_wf {Σ : GlobalContextMap.t} : - has_tBox -> has_tRel -> @wf_glob efl Σ -> Pre_glob Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf) -> wf_glob (efl := efl') (gen_transform_env Σ). Proof. - intros hasb hasrel wfg pre. + intros wfg pre. rewrite gen_transform_env_eq //. destruct Σ as [Σ ? ? ?]. cbn in *. revert pre. induction Σ in map0, repr0, wf0, wfg |- *; auto; cbn; constructor; auto. diff --git a/erasure/theories/EInlineProjections.v b/erasure/theories/EInlineProjections.v index 178225ad4..8f65984c3 100644 --- a/erasure/theories/EInlineProjections.v +++ b/erasure/theories/EInlineProjections.v @@ -287,7 +287,7 @@ Definition optimize_constant_decl Σ cb := Definition optimize_decl Σ d := match d with | ConstantDecl cb => ConstantDecl (optimize_constant_decl Σ cb) - | InductiveDecl idecl => d + | InductiveDecl idecl => InductiveDecl idecl end. Definition optimize_env Σ := @@ -820,7 +820,12 @@ Qed. From MetaCoq.Erasure Require Import EGenericGlobalMap. #[local] -Instance GT : GenTransform := { gen_transform := optimize }. +Instance GT : GenTransform := { gen_transform := optimize; gen_transform_inductive_decl := id }. +#[local] +Instance GTID : GenTransformId GT. +Proof. + red. reflexivity. +Qed. #[local] Instance GTExt efl : GenTransformExtends efl (disable_projections_env_flag efl) GT. Proof. @@ -831,8 +836,8 @@ Qed. #[local] Instance GTWf efl : GenTransformWf efl (disable_projections_env_flag efl) GT. Proof. - refine {| gen_transform_pre := fun _ _ => True |}; auto. - intros Σ n t hasb hasr _ wfΣ wft. + refine {| gen_transform_pre := fun _ _ => has_tBox /\ has_tRel |}; auto. + intros Σ n t [] wfΣ wft. now apply optimize_wellformed. Defined. @@ -870,8 +875,10 @@ Proof. now unshelve eapply (optimize_expanded_decl (Σ:=Σ')). Qed. -Lemma Pre_glob efl Σ wf : Pre_glob (GTWF:=GTWf efl) Σ wf. +Lemma Pre_glob efl Σ wf : + has_tBox -> has_tRel -> Pre_glob (GTWF:=GTWf efl) Σ wf. Proof. + intros hasb hasr. induction Σ => //. destruct a as [kn d]; cbn. split => //. destruct d as [[[|]]|] => //=. Qed. @@ -882,7 +889,7 @@ Lemma optimize_env_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : Proof. intros hasb hasre wfg. eapply (gen_transform_env_wf (gt := GTExt efl)) => //. - apply Pre_glob. + now apply Pre_glob. Qed. Definition optimize_program (p : eprogram_env) := diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 0cd56240e..565cdcfbd 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -780,68 +780,6 @@ Proof. unfold expanded_constant_decl => /=. apply remove_match_on_box_expanded_irrel. Qed. - -Lemma remove_match_on_box_env_extends' {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : - extends_prefix Σ Σ' -> - wf_glob Σ' -> - List.map (on_snd (remove_match_on_box_decl Σ)) Σ.(GlobalContextMap.global_decls) = - List.map (on_snd (remove_match_on_box_decl Σ')) Σ.(GlobalContextMap.global_decls). -Proof. - intros ext. - destruct Σ as [Σ map repr wf]; cbn in *. - move=> wfΣ. - assert (extends_prefix Σ Σ); auto. now exists []. - assert (wf_glob Σ). - { eapply extends_wf_glob. exact ext. tea. } - revert H H0. - generalize Σ at 1 3 5 6. intros Σ''. - induction Σ'' => //. cbn. - intros hin wfg. depelim wfg. - f_equal. - 2:{ eapply IHΣ'' => //. destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. } - unfold on_snd. cbn. f_equal. - eapply wellformed_remove_match_on_box_decl_extends => //. cbn. - eapply extends_wf_global_decl. 3:tea. - eapply extends_wf_glob; tea. - eapply extends_prefix_extends. - destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. - eapply extends_wf_glob; tea. cbn. - eapply extends_prefix_extends; tea. -Qed. - -Lemma remove_match_on_box_env_eq {efl : EEnvFlags} (Σ : GlobalContextMap.t) : wf_glob Σ -> remove_match_on_box_env Σ = remove_match_on_box_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). -Proof. - intros wf. - unfold remove_match_on_box_env. - destruct Σ; cbn. cbn in wf. - induction global_decls in map, repr, wf0, wf |- * => //. - cbn. f_equal. - destruct a as [kn d]; unfold on_snd; cbn. f_equal. symmetry. - eapply wellformed_remove_match_on_box_decl_extends => //. cbn. now depelim wf. cbn. now exists [(kn, d)]. cbn. - set (Σg' := GlobalContextMap.make global_decls (fresh_globals_cons_inv wf0)). - erewrite <- (IHglobal_decls (GlobalContextMap.map Σg') (GlobalContextMap.repr Σg')). - 2:now depelim wf. - set (Σg := {| GlobalContextMap.global_decls := _ :: _ |}). - symmetry. eapply (remove_match_on_box_env_extends' (Σ := Σg') (Σ' := Σg)) => //. - cbn. now exists [a]. -Qed. - -Lemma remove_match_on_box_env_expanded {efl : EEnvFlags} {Σ : GlobalContextMap.t} : - wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (remove_match_on_box_env Σ). -Proof. - unfold expanded_global_env; move=> wfg. - rewrite remove_match_on_box_env_eq //. - destruct Σ as [Σ map repr wf]. cbn in *. - clear map repr. - induction 1; cbn; constructor; auto. - cbn in IHexpanded_global_declarations. - unshelve eapply IHexpanded_global_declarations. now depelim wfg. cbn. - set (Σ' := GlobalContextMap.make _ _). - rewrite -(remove_match_on_box_env_eq Σ'). cbn. now depelim wfg. - eapply (remove_match_on_box_expanded_decl_irrel (Σ := Σ')). now depelim wfg. - now unshelve eapply (remove_match_on_box_expanded_decl (Σ:=Σ')). -Qed. - Lemma remove_match_on_box_wellformed {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : has_tBox -> has_tRel -> wf_glob Σ -> EWellformed.wellformed Σ n t -> EWellformed.wellformed Σ n (remove_match_on_box Σ t). @@ -905,6 +843,67 @@ Proof. all:intros; rtoProp; intuition auto; solve_all. Qed. +From MetaCoq.Erasure Require Import EGenericGlobalMap. + +#[local] +Instance GT : GenTransform := { gen_transform := remove_match_on_box; gen_transform_inductive_decl := id }. +#[local] +Instance GTID : GenTransformId GT. +Proof. + red. reflexivity. +Qed. +#[local] +Instance GTExt efl : GenTransformExtends efl efl GT. +Proof. + intros Σ t n wfΣ Σ' ext wf wf'. + unfold gen_transform, GT. + eapply wellformed_remove_match_on_box_extends; tea. +Qed. +#[local] +Instance GTWf efl : GenTransformWf efl efl GT. +Proof. + refine {| gen_transform_pre := fun _ _ => has_tBox /\ has_tRel |}; auto. + intros Σ n t [] wfΣ wft. + eapply remove_match_on_box_wellformed_irrel => //. + now eapply remove_match_on_box_wellformed. +Defined. + +Lemma remove_match_on_box_env_extends' {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : + extends Σ Σ' -> + wf_glob Σ -> + wf_glob Σ' -> + List.map (on_snd (remove_match_on_box_decl Σ)) Σ.(GlobalContextMap.global_decls) = + List.map (on_snd (remove_match_on_box_decl Σ')) Σ.(GlobalContextMap.global_decls). +Proof. + intros ext wf. + eapply (gen_transform_env_extends' (gt := GTExt efl)) => //. +Qed. + +Lemma remove_match_on_box_env_eq {efl : EEnvFlags} (Σ : GlobalContextMap.t) : + wf_glob Σ -> + remove_match_on_box_env Σ = remove_match_on_box_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). +Proof. + intros wf. + eapply (gen_transform_env_eq (gt := GTExt efl)) => //. +Qed. + +Lemma remove_match_on_box_env_expanded {efl : EEnvFlags} {Σ : GlobalContextMap.t} : + wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (remove_match_on_box_env Σ). +Proof. + unfold expanded_global_env; move=> wfg. + rewrite remove_match_on_box_env_eq //. + destruct Σ as [Σ map repr wf]. cbn in *. + clear map repr. + induction 1; cbn; constructor; auto. + cbn in IHexpanded_global_declarations. + unshelve eapply IHexpanded_global_declarations. now depelim wfg. cbn. + set (Σ' := GlobalContextMap.make _ _). + rewrite -(remove_match_on_box_env_eq Σ'). cbn. now depelim wfg. + eapply (remove_match_on_box_expanded_decl_irrel (Σ := Σ')). now depelim wfg. + now unshelve eapply (remove_match_on_box_expanded_decl (Σ:=Σ')). +Qed. + + Lemma remove_match_on_box_wellformed_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} d : wf_glob Σ -> wf_global_decl Σ d -> wf_global_decl (remove_match_on_box_env Σ) d. @@ -936,23 +935,26 @@ Proof. now eapply Forall_map; cbn. Qed. + +Lemma Pre_glob efl Σ wf : + has_tBox -> has_tRel -> + Pre_glob (GTWF:=GTWf efl) Σ wf. +Proof. + intros hasb hasr. induction Σ => //. destruct a as [kn d]; cbn. + split => //. destruct d as [[[|]]|] => //=. +Qed. + Lemma remove_match_on_box_env_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : has_tBox -> has_tRel -> wf_glob Σ -> wf_glob (remove_match_on_box_env Σ). Proof. - intros hasb hasrel. - intros wfg. rewrite remove_match_on_box_env_eq //. - destruct Σ as [Σ map repr wf]; cbn in *. - clear map repr. - induction wfg; cbn; constructor; auto. - - rewrite /= -(remove_match_on_box_env_eq (GlobalContextMap.make Σ (fresh_globals_cons_inv wf))) //. - eapply remove_match_on_box_decl_wf => //. - - rewrite /= -(remove_match_on_box_env_eq (GlobalContextMap.make Σ (fresh_globals_cons_inv wf))) //. - now eapply fresh_global_remove_match_on_box_env. + intros hasb hasrel wfg. + eapply (gen_transform_env_wf (gt := GTExt efl)) => //. + now apply Pre_glob. Qed. Definition remove_match_on_box_program (p : eprogram_env) := -(remove_match_on_box_env p.1, remove_match_on_box p.1 p.2). + (remove_match_on_box_env p.1, remove_match_on_box p.1 p.2). Definition remove_match_on_box_program_wf {efl} (p : eprogram_env) {hastbox : has_tBox} {hastrel : has_tRel} : wf_eprogram_env efl p -> wf_eprogram efl (remove_match_on_box_program p). diff --git a/erasure/theories/ERemoveParams.v b/erasure/theories/ERemoveParams.v index c9f891c5e..bc5a6b529 100644 --- a/erasure/theories/ERemoveParams.v +++ b/erasure/theories/ERemoveParams.v @@ -1011,22 +1011,30 @@ Proof. now intros -> => /=. Qed. +Definition switch_no_params (efl : EEnvFlags) := + {| has_axioms := has_axioms; + has_cstr_params := false; + term_switches := term_switches ; + cstr_as_blocks := cstr_as_blocks + |}. (* Stripping preserves well-formedness directly, not caring about eta-expansion *) -Lemma strip_wellformed {Σ : GlobalContextMap.t} n t : - @wf_glob all_env_flags Σ -> - @wellformed all_env_flags Σ n t -> - @wellformed all_env_flags (strip_env Σ) n (strip Σ t). +Lemma strip_wellformed {efl} {Σ : GlobalContextMap.t} n t : + cstr_as_blocks = false -> + has_tApp -> + @wf_glob efl Σ -> + @wellformed efl Σ n t -> + @wellformed (switch_no_params efl) (strip_env Σ) n (strip Σ t). Proof. - intros wfΣ. + intros cab hasapp wfΣ. revert n. funelim (strip Σ t); try intros n. all:cbn -[strip lookup_constructor lookup_inductive]; simp_strip; intros. all:try solve[unfold wf_fix_gen in *; rtoProp; intuition auto; toAll; solve_all]. - cbn -[strip]; simp_strip. intros; rtoProp; intuition auto. rewrite lookup_env_strip. destruct lookup_env eqn:hl => // /=. - destruct g => /= //. - - cbn -[strip] in *. + destruct g => /= //. destruct (cst_body c) => //. + - rewrite cab in H |- *. cbn -[strip] in *. rewrite lookup_env_strip. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => //. destruct nth_error => //. @@ -1035,25 +1043,27 @@ Proof. destruct g eqn:hg => /= //. subst g. destruct nth_error => //. rtoProp; intuition auto. simp_strip. toAll; solve_all. + toAll. solve_all. - cbn -[strip] in H0 |- *. rewrite lookup_env_strip. destruct lookup_env eqn:hl => // /=. - destruct g eqn:hg => /= //. subst g. + destruct g eqn:hg => /= //. subst g. cbn in H0. now rtoProp. destruct nth_error => //. all:rtoProp; intuition auto. destruct EAst.ind_ctors => //. destruct nth_error => //. + all: eapply H; auto. - unfold wf_fix_gen in *. rewrite map_length. rtoProp; intuition auto. toAll; solve_all. now rewrite -strip_isLambda. toAll; solve_all. - unfold wf_fix in *. rewrite map_length; rtoProp; intuition auto. toAll; solve_all. - move:H1; rewrite !wellformed_mkApps //. rtoProp; intuition auto. - toAll; solve_all. + toAll; solve_all. toAll; solve_all. - move:H0; rewrite !wellformed_mkApps //. rtoProp; intuition auto. - move: H1. cbn. + move: H1. cbn. rewrite cab. rewrite lookup_env_strip. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => //. destruct nth_error => //. toAll; solve_all. eapply All_skipn. solve_all. - move:H0; rewrite !wellformed_mkApps //. rtoProp; intuition auto. - move: H1. cbn. + move: H1. cbn. rewrite cab. rewrite lookup_env_strip. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => //. destruct nth_error => //. @@ -1102,13 +1112,6 @@ Proof. now eapply strip_wellformed_irrel. Qed. -Definition switch_no_params (efl : EEnvFlags) := - {| has_axioms := has_axioms; - has_cstr_params := false; - term_switches := term_switches ; - cstr_as_blocks := false - |}. - Lemma strip_decl_wf (efl := all_env_flags) {Σ : GlobalContextMap.t} : wf_glob Σ -> forall d, wf_global_decl Σ d -> @@ -1212,30 +1215,36 @@ Proof. intros _ => //. Qed. +From MetaCoq.Erasure Require Import EGenericGlobalMap. + +#[local] +Instance GT : GenTransform := { gen_transform := strip; gen_transform_inductive_decl := strip_inductive_decl }. +#[local] +Instance GTExt efl : has_tApp -> GenTransformExtends efl efl GT. +Proof. + intros hasapp Σ t n wfΣ Σ' ext wf wf'. + unfold gen_transform, GT. + eapply strip_extends; tea. +Qed. +#[local] +Instance GTWf efl : GenTransformWf efl (switch_no_params efl) GT. +Proof. + refine {| gen_transform_pre := fun _ _ => has_tApp /\ cstr_as_blocks = false |}; auto. + - unfold wf_minductive; intros []. cbn. repeat rtoProp. intuition auto. + - intros Σ n t [? ?] wfΣ wft. unfold gen_transform_env, gen_transform. simpl. + eapply strip_wellformed => //. +Defined. + Lemma strip_extends' {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : has_tApp -> extends Σ Σ' -> + wf_glob Σ -> wf_glob Σ' -> List.map (on_snd (strip_decl Σ)) Σ.(GlobalContextMap.global_decls) = List.map (on_snd (strip_decl Σ')) Σ.(GlobalContextMap.global_decls). Proof. - intros hast ext. - destruct Σ as [Σ map repr wf]; cbn in *. - move=> wfΣ. - assert (extends Σ Σ); auto. now exists []. - assert (wf_glob Σ). - { eapply EExtends.extends_wf_glob. exact ext. tea. } - revert H H0. - generalize Σ at 1 3 5 6. intros Σ''. - induction Σ'' => //. cbn. - intros hin wfg. depelim wfg. - f_equal. - 2:{ eapply IHΣ'' => //. destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. } - unfold on_snd. cbn. f_equal. - eapply strip_decl_extends => //. - eapply extends_wf_global_decl. 3:tea. auto. cbn. - eapply EExtends.extends_wf_glob; tea. - destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. + intros hast ext wf. + now apply (gen_transform_env_extends' (gt := GTExt efl hast) ext). Qed. Lemma strip_env_eq (efl := all_env_flags) (Σ : GlobalContextMap.t) : wf_glob Σ -> strip_env Σ = strip_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). @@ -1246,13 +1255,14 @@ Proof. induction global_decls in map, repr, wf0, wf |- * => //. cbn. f_equal. destruct a as [kn d]; unfold on_snd; cbn. f_equal. symmetry. - eapply strip_decl_extends => //. cbn. cbn. now exists [(kn, d)]. cbn. now depelim wf. + eapply strip_decl_extends => //. cbn. cbn. eapply EExtends.extends_prefix_extends. now exists [(kn, d)]. auto. cbn. now depelim wf. set (Σg' := GlobalContextMap.make global_decls (fresh_globals_cons_inv wf0)). erewrite <- (IHglobal_decls (GlobalContextMap.map Σg') (GlobalContextMap.repr Σg')). 2:now depelim wf. set (Σg := {| GlobalContextMap.global_decls := _ :: _ |}). symmetry. eapply (strip_extends' (Σ := Σg') (Σ' := Σg)) => //. - cbn. now exists [a]. + cbn. eapply EExtends.extends_prefix_extends => //. now exists [a]. + cbn. now depelim wf. Qed. Lemma strip_env_wf (efl := all_env_flags) {Σ : GlobalContextMap.t} : diff --git a/erasure/theories/EWcbvEvalNamed.v b/erasure/theories/EWcbvEvalNamed.v index a6ca1034e..044e57def 100644 --- a/erasure/theories/EWcbvEvalNamed.v +++ b/erasure/theories/EWcbvEvalNamed.v @@ -955,7 +955,7 @@ Qed. Local Hint Resolve incl_tl incl_appr incl_appl : core. Lemma wellformed_annotate' Σ Γ Γ' s : - incl Γ' Γ -> + incl Γ' Γ -> wellformed (efl := extraction_env_flags) Σ #|Γ| s -> wellformed (efl := named_extraction_env_flags) (annotate_env Γ' Σ) #|Γ| (annotate Γ s). Proof. intros Hincl Hwf. From 805bbdad44325d51864e4df611efdcebce859616 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 24 Aug 2023 01:42:55 +0200 Subject: [PATCH 38/61] Prove a more precise result about erase_global_decls with dependencies --- .vscode/metacoq.code-workspace | 2 + Makefile | 2 +- common/theories/Transform.v | 8 +- erasure-plugin/_CoqProject.in | 1 - erasure-plugin/theories/ETransform.v | 231 ++++++- erasure-plugin/theories/Erasure.v | 55 +- erasure/theories/EConstructorsAsBlocks.v | 91 +-- erasure/theories/EImplementBox.v | 77 +-- erasure/theories/EOptimizePropDiscr.v | 2 +- erasure/theories/EProgram.v | 42 ++ erasure/theories/ErasureFunction.v | 653 ++++++++++++++++---- erasure/theories/Typed/ErasureCorrectness.v | 5 +- template-coq/theories/TemplateProgram.v | 2 +- 13 files changed, 962 insertions(+), 209 deletions(-) diff --git a/.vscode/metacoq.code-workspace b/.vscode/metacoq.code-workspace index 2ce1e12f6..6a0a8cce3 100644 --- a/.vscode/metacoq.code-workspace +++ b/.vscode/metacoq.code-workspace @@ -10,6 +10,7 @@ "coqtop.args": [ + "-R", "safechecker-plugin/theories", "MetaCoq.SafeCheckerPlugin", "-R", "utils/theories", "MetaCoq.Utils", @@ -27,6 +28,7 @@ "-R", "safechecker/theories", "MetaCoq.SafeChecker", "-I", "erasure/src", "-R", "erasure/theories", "MetaCoq.Erasure", + "-R", "erasure-plugin/theories", "MetaCoq.ErasurePlugin", "-R", "translations", "MetaCoq.Translations", "-R", "test-suite/plugin-demo/theories", "MetaCoq.ExtractedPluginDemo", "-R", "test-suite", "MetaCoq.TestSuite" diff --git a/Makefile b/Makefile index e16daadc8..6da8df03b 100644 --- a/Makefile +++ b/Makefile @@ -1,5 +1,5 @@ -all: printconf template-coq pcuic safechecker erasure +all: printconf template-coq pcuic safechecker erasure erasure-plugin -include Makefile.conf diff --git a/common/theories/Transform.v b/common/theories/Transform.v index 5a4aeace8..bbeb25d22 100644 --- a/common/theories/Transform.v +++ b/common/theories/Transform.v @@ -39,7 +39,7 @@ Module Transform. post : program' -> Prop; correctness : forall input (p : pre input), post (transform input p); obseq : forall p : program, pre p -> program' -> value -> value' -> Prop; - preservation : preserves_eval pre transform obseq; }. + preservation : preserves_eval pre transform obseq }. Definition run (x : t) (p : program) (pr : pre x p) : program' := time x.(name) (fun _ => x.(transform) p pr) tt. @@ -58,7 +58,8 @@ Module Transform. Local Obligation Tactic := idtac. Program Definition compose (o : t program program' value value' eval eval') (o' : t program' program'' value' value'' eval' eval'') - (hpp : (forall p, o.(post) p -> o'.(pre) p)) : t program program'' value value'' eval eval'' := + (hpp : (forall p, o.(post) p -> o'.(pre) p)) + : t program program'' value value'' eval eval'' := {| name := (o.(name) ^ " -> " ^ o'.(name))%bs; transform p hp := run o' (run o p hp) (hpp _ (o.(correctness) _ hp)); @@ -79,10 +80,11 @@ Module Transform. cbn in ev. destruct ev as [v' [ev]]. epose proof (o'.(preservation) (o.(transform) p pr) v'). specialize (H0 (hpp _ (o.(correctness) _ pr)) ev). - destruct H0 as [v'' [ev' obs']]. + destruct H0 as [v'' [ev' obs'']]. exists v''. constructor => //. exists v'. now split. Qed. + End Comp. Declare Scope transform_scope. diff --git a/erasure-plugin/_CoqProject.in b/erasure-plugin/_CoqProject.in index 24d1735f9..49d459d43 100644 --- a/erasure-plugin/_CoqProject.in +++ b/erasure-plugin/_CoqProject.in @@ -1,6 +1,5 @@ -R theories MetaCoq.ErasurePlugin - theories/ETransform.v theories/Erasure.v theories/Extraction.v diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index 352a50101..ba2e8c17f 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -256,6 +256,155 @@ Next Obligation. cbn; intros. sq. now subst. Qed. +Definition extends_eprogram (p p' : eprogram) := + extends p.1 p'.1 /\ p.2 = p'.2. + +Definition extends_eprogram_env (p p' : eprogram_env) := + extends p.1 p'.1 /\ p.2 = p'.2. + +Section PCUICEnv. (* Locally reuse the short names for PCUIC environment handling *) +Import PCUICAst.PCUICEnvironment. + +Lemma build_wf_env_from_env_eq {cf : checker_flags} (guard : abstract_guard_impl) (Σ : global_env_ext_map) (wfΣ : ∥ PCUICTyping.wf_ext Σ ∥) : + let wfe := build_wf_env_from_env Σ (map_squash (PCUICTyping.wf_ext_wf Σ) wfΣ) in + forall Σ' : global_env, Σ' ∼ wfe -> declarations Σ' = declarations Σ. +Proof. + cbn; intros. rewrite H. reflexivity. +Qed. + +Definition extends_global_env (Σ Σ' : global_env_ext_map) := + [/\ (forall kn decl, lookup_env Σ kn = Some decl ->lookup_env Σ' kn = Some decl), + Σ.(universes) = Σ'.(universes), Σ.2 = Σ'.2 & Σ.(retroknowledge) = Σ'.(retroknowledge)]. + +Definition extends_pcuic_program (p p' : pcuic_program) := + extends_global_env p.1 p'.1 /\ p.2 = p'.2. +Import ErasureFunction. +Import PCUICAst. +Lemma fresh_global_app kn Σ Σ' : fresh_global kn (Σ ++ Σ') -> + fresh_global kn Σ /\ fresh_global kn Σ'. +Proof. + induction Σ; cbn; intuition auto. + - constructor. + - depelim H. constructor => //. + now eapply Forall_app in H0. + - depelim H. + now eapply Forall_app in H0. +Qed. + +Lemma lookup_global_app_wf Σ Σ' kn : EnvMap.EnvMap.fresh_globals (Σ ++ Σ') -> + forall decl, lookup_global Σ' kn = Some decl -> lookup_global (Σ ++ Σ') kn = Some decl. +Proof. + induction Σ in kn |- *; cbn; auto. + intros fr; depelim fr. + intros decl hd; cbn. + destruct (eqb_spec kn kn0). + eapply PCUICWeakeningEnv.lookup_global_Some_fresh in hd. + subst. now eapply fresh_global_app in H as [H H']. + now eapply IHΣ. +Qed. + +Lemma strictly_extends_lookups {cf:checker_flags} (X X' : wf_env) (Σ Σ' : global_env) : + (forall (kn : kername) (decl decl' : global_decl), lookup_env Σ' kn = Some decl -> lookup_env Σ kn = Some decl' -> decl = decl') -> + retroknowledge Σ = retroknowledge Σ' -> + strictly_extends_decls X Σ -> strictly_extends_decls X' Σ' -> + PCUICTyping.wf Σ -> PCUICTyping.wf Σ' -> + equiv_env_inter X' X. +Proof. + intros hl hr [] [] wf wf'. + split. + - intros kn d. + unfold lookup_env in *. + destruct s as [? eq], s0 as [? eq']. + rewrite eq' eq in hl. + intros decl' hl' hl''. + specialize (hl kn d decl'). + eapply wf_fresh_globals in wf. + eapply wf_fresh_globals in wf'. + rewrite eq in wf; rewrite eq' in wf'. + eapply lookup_global_app_wf in hl'; tea. + eapply lookup_global_app_wf in hl''; tea. + eauto. + - rewrite /primitive_constant. now rewrite e0 e2 hr. +Qed. + +#[global] +Instance erase_transform_extends {guard : abstract_guard_impl} : + TransformExt.t (erase_transform (guard := guard)) extends_pcuic_program extends_eprogram_env. +Proof. + red. intros p p' pr pr' [ext eq]. rewrite /transform /=. + destruct pr. + epose proof (erase_pcuic_program_normalization_helper p (map_squash fst s)). + destruct p as [Σ t]. destruct p' as [Σ' t']. + simpl in *. subst t'. red. split. + - destruct ext. red. + cbn. rewrite /ErasureFunction.erase_global_fast. + unshelve erewrite ErasureFunction.erase_global_deps_fast_spec. + { intros. intuition auto. eapply H11; tea. admit. admit. } + { intros. red in H4. cbn in H4. subst. reflexivity. } + destruct Σ'. cbn in *. subst u. + set (fst := ErasureFunction.erase _ _ _ _ _). + set (snd:= ErasureFunction.erase _ _ _ _ _). + assert (fst = snd). + { subst fst snd. symmetry. + eapply ErasureFunction.erase_irrel_global_env. + eapply ErasureFunction.equiv_env_inter_hlookup. + intros ? ? -> ->. cbn. + eapply ErasureFunction.equiv_env_inter_sym. + eapply ErasureFunction.extends_global_env_equiv_env. + cbn. split => //. + sq. unfold primitive_constant. now rewrite H3. } + clearbody fst snd. subst snd. + intros. + set (X := build_wf_env_from_env g _) in *. + destruct (lookup_env Σ kn) eqn:E. + unshelve epose proof (abstract_env_exists X). 3:tc. tc. + destruct H4 as [[Σ' ΣX]]. + unshelve epose proof (build_wf_env_from_env_eq _ (g, Σ.2) _ _ ΣX). cbn in H4. + epose proof (hl := ErasureFunction.lookup_env_erase_decl (optimized_abstract_env_impl) _ (EAstUtils.term_global_deps fst) _ _ _ kn g0 _ ΣX). + unfold lookup_env in E, hl. rewrite H4 in hl. specialize (hl (H0 _ _ E)). + forward hl. admit. + destruct g0. + + destruct hl as [X' [nin [pf [eq ext]]]]. + unshelve erewrite ErasureFunction.erase_global_deps_fast_spec. shelve. + { intros. red in H5. cbn in H5. subst. reflexivity. } + erewrite eq. + set (Xs := build_wf_env_from_env Σ.1 _) in *. + unshelve epose proof (abstract_env_exists Xs). 3:tc. tc. + destruct H5 as [[Σs Hs]]. + epose proof (hl' := ErasureFunction.lookup_env_erase_decl (optimized_abstract_env_impl) _ (EAstUtils.term_global_deps fst) _ _ _ kn _ _ Hs). + unshelve epose proof (build_wf_env_from_env_eq _ _ _ _ Hs). rewrite /lookup_env H5 in hl'. specialize (hl' E). + forward hl'. admit. cbn in hl'. + destruct hl' as [X'' [nin' [pf' [eq' ext']]]]. + erewrite eq' in H2. rewrite -H2. f_equal. f_equal. + f_equal. unfold erase_constant_decl. + eapply erase_constant_body_irrel. + eapply ErasureFunction.equiv_env_inter_hlookup. cbn. + intros ? ? H6 H7. + cbn in ext. specialize (ext _ _ eq_refl eq_refl) as [ext]. + specialize (ext' _ _ eq_refl eq_refl) as [ext']. + cbn in H6, H7. subst Σ0 Σ'0. cbn. split => //. + { intros kn' decl'. rewrite /lookup_env. + destruct (map_squash Datatypes.fst s). + destruct pr' as [pr' ?]. + destruct (map_squash Datatypes.fst pr'). + eapply strictly_extends_lookups. 2-3:tea. + intros kn'' ? ? hg hΣ. eapply H0 in hΣ. congruence. auto. + apply X0. apply X1. } + { rewrite /primitive_constant. + destruct ext as [_ _ ->]. + destruct ext' as [_ _ ->]. + now rewrite H3. } + + + - symmetry. + eapply ErasureFunction.erase_irrel_global_env. + intros Σg Σg' eq eq'. + destruct ext as [hl hu hr]. + destruct Σ as [m univs]. destruct Σ' as [m' univs']. simpl in *. + subst. cbn. destruct m, m'. cbn in *. split => //. sq. red. split => //. + unfold primitive_constant. now rewrite H. +Qed. +End PCUICEnv. + Obligation Tactic := idtac. (** This transformation is the identity on terms but changes the evaluation relation to one @@ -283,6 +432,13 @@ Next Obligation. now eapply EEtaExpandedFix.expanded_isEtaExp. Qed. +#[global] +Instance guarded_to_unguarded_fix_extends {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} {efl : EEnvFlags} (wguard : with_guarded_fix) : + TransformExt.t (guarded_to_unguarded_fix (wcon:=wcon) wguard) extends_eprogram_env term snd. +Proof. + red. intros p p' pr pr' [ext eq]. now rewrite /transform /=. +Qed. + Definition rebuild_wf_env {efl} (p : eprogram) (hwf : wf_eprogram efl p): eprogram_env := (GlobalContextMap.make p.1 (wf_glob_fresh p.1 (proj1 hwf)), p.2). @@ -300,6 +456,13 @@ Next Obligation. cbn. intros fl efl [] input v [] ev p'; exists v; split => //. Qed. +#[global] +Instance rebuild_wf_env_extends {fl : EWcbvEval.WcbvFlags} {efl : EEnvFlags} with_exp : + TransformExt.t (rebuild_wf_env_transform with_exp) extends_eprogram term snd. +Proof. + red. intros p p' pr pr' [ext eq]. now rewrite /transform /=. +Qed. + Program Definition remove_params_optimization {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := all_env_flags): Transform.t eprogram_env eprogram EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram fl) := @@ -326,7 +489,16 @@ Next Obligation. now move/andP: etap. Qed. -Program Definition remove_params_fast_optimization (fl : EWcbvEval.WcbvFlags) {wcon : EWcbvEval.with_constructor_as_block = false} +#[global] +Instance remove_params_extends {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} + (efl := all_env_flags): + TransformExt.t (remove_params_optimization (wcon:=wcon)) extends_eprogram_env term snd. +Proof. + red. intros p p' pr pr' [ext eq]. rewrite /transform /=. rewrite eq. + eapply strip_extends => //. apply pr'. rewrite -eq. apply pr. +Qed. + +Program Definition remove_params_fast_optimization {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := all_env_flags) : Transform.t eprogram_env eprogram EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram fl) := {| name := "stripping constructor parameters (faster?)"; @@ -356,6 +528,16 @@ Next Obligation. now move/andP: etap. Qed. +#[global] +Instance remove_params_fast_extends {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} + (efl := all_env_flags): + TransformExt.t (remove_params_fast_optimization (wcon:=wcon)) extends_eprogram_env term snd. +Proof. + red. intros p p' pr pr' [ext eq]. rewrite /transform /=. rewrite eq. + rewrite -!ERemoveParams.Fast.strip_fast. + eapply strip_extends => //. apply pr'. rewrite -eq. apply pr. +Qed. + Import EOptimizePropDiscr EWcbvEval. Program Definition remove_match_on_box_trans {fl : WcbvFlags} {wcon : with_constructor_as_block = false} {efl : EEnvFlags} {hastrel : has_tRel} {hastbox : has_tBox} : @@ -381,6 +563,14 @@ Next Obligation. Unshelve. eauto. Qed. +#[global] +Instance remove_match_on_box_extends {fl : WcbvFlags} {wcon : with_constructor_as_block = false} {efl : EEnvFlags} {hastrel : has_tRel} {hastbox : has_tBox} : + TransformExt.t (remove_match_on_box_trans (wcon:=wcon) (hastrel := hastrel) (hastbox := hastbox)) extends_eprogram_env term snd. +Proof. + red. intros p p' pr pr' [ext eq]. rewrite /transform /=. rewrite -eq. + eapply wellformed_remove_match_on_box_extends; eauto. apply pr. apply pr'. +Qed. + From MetaCoq.Erasure Require Import EInlineProjections. Program Definition inline_projections_optimization {fl : WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := switch_no_params all_env_flags) @@ -405,10 +595,19 @@ Next Obligation. cbn. eapply wfe. Unshelve. auto. Qed. +#[global] +Instance inline_projections_optimization_extends {fl : WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := switch_no_params all_env_flags) + {hastrel : has_tRel} {hastbox : has_tBox} : + TransformExt.t (inline_projections_optimization (wcon:=wcon) (hastrel := hastrel) (hastbox := hastbox)) extends_eprogram_env term snd. +Proof. + red. intros p p' pr pr' [ext eq]. rewrite /transform /=. rewrite -eq. + eapply wellformed_optimize_extends; eauto. apply pr. apply pr'. +Qed. + From MetaCoq.Erasure Require Import EConstructorsAsBlocks. -Program Definition constructors_as_blocks_transformation (efl : EEnvFlags) - {has_app : has_tApp} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = false} : +Program Definition constructors_as_blocks_transformation {efl : EEnvFlags} + {has_app : has_tApp} {has_rel : has_tRel} {hasbox : has_tBox} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = false} : Transform.t eprogram_env eprogram EAst.term EAst.term (eval_eprogram_env target_wcbv_flags) (eval_eprogram block_wcbv_flags) := {| name := "transforming to constuctors as blocks"; transform p _ := EConstructorsAsBlocks.transform_blocks_program p ; @@ -417,13 +616,13 @@ Program Definition constructors_as_blocks_transformation (efl : EEnvFlags) obseq p hp p' v v' := v' = EConstructorsAsBlocks.transform_blocks p.1 v |}. Next Obligation. - move=> efl hasapp haspars hascstrs [Σ t] [] [wftp wft] /andP [etap etat]. + move=> efl hasapp hasrel hasbox haspars hascstrs [Σ t] [] [wftp wft] /andP [etap etat]. cbn in *. split. - eapply transform_wf_global; eauto. - eapply transform_wellformed; eauto. Qed. Next Obligation. - red. move=> efl hasapp haspars hascstrs [Σ t] /= v [[wfe1 wfe2] wft] [ev]. + red. move=> efl hasapp hasrel hasbox haspars hascstrs [Σ t] /= v [[wfe1 wfe2] wft] [ev]. eexists. split; [ | eauto]. unfold EEtaExpanded.expanded_eprogram_env_cstrs in *. revert wft. move => /andP // [e1 e2]. @@ -432,9 +631,18 @@ Next Obligation. eapply transform_blocks_eval; cbn; eauto. Qed. +#[global] +Instance constructors_as_blocks_extends (efl : EEnvFlags) + {has_app : has_tApp} {has_rel : has_tRel} {hasbox : has_tBox} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = false} : + TransformExt.t (constructors_as_blocks_transformation (has_app := has_app) (has_rel := has_rel) (hasbox := hasbox) (has_pars := has_pars) (has_cstrblocks := has_cstrblocks)) extends_eprogram_env term snd. +Proof. + red. intros p p' pr pr' [ext eq]. rewrite /transform /=. rewrite -eq. + eapply transform_blocks_extends; eauto. apply pr. apply pr'. +Qed. + From MetaCoq.Erasure Require Import EImplementBox. -Program Definition implement_box_transformation (efl : EEnvFlags) +Program Definition implement_box_transformation {efl : EEnvFlags} {has_app : has_tApp} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = true} : Transform.t eprogram eprogram EAst.term EAst.term (eval_eprogram block_wcbv_flags) (eval_eprogram block_wcbv_flags) := {| name := "transforming to constuctors as blocks"; @@ -446,11 +654,18 @@ Program Definition implement_box_transformation (efl : EEnvFlags) Next Obligation. intros. cbn in *. destruct p. split. - eapply transform_wf_global; eauto. - - eapply transform_wellformed; eauto. + - now eapply transform_wellformed'. Qed. Next Obligation. red. intros. destruct pr. destruct H. eexists. split; [ | eauto]. econstructor. eapply implement_box_eval; cbn; eauto. -Qed. \ No newline at end of file +Qed. + +#[global] +Instance implement_box_extends (efl : EEnvFlags) {has_app : has_tApp} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = true} : + TransformExt.t (implement_box_transformation (has_app := has_app) (has_pars := has_pars) (has_cstrblocks := has_cstrblocks)) extends_eprogram term snd. +Proof. + red. intros p p' pr pr' [ext eq]. rewrite /transform /=. now rewrite -eq. +Qed. diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 71d754645..ce6e68ef0 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -6,6 +6,7 @@ From MetaCoq.Template Require Import EtaExpand TemplateProgram. From MetaCoq.PCUIC Require PCUICAst PCUICAstUtils PCUICProgram. From MetaCoq.SafeChecker Require Import PCUICErrors PCUICWfEnvImpl. From MetaCoq.Erasure Require EAstUtils ErasureFunction ErasureCorrectness EPretty Extract. +From MetaCoq.Erasure Require Import EProgram. From MetaCoq.ErasurePlugin Require Import ETransform. Import PCUICProgram. @@ -90,7 +91,7 @@ Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl (* Rebuild the efficient lookup table *) rebuild_wf_env_transform (efl := efl) true ▷ (* First-order constructor representation *) - constructors_as_blocks_transformation efl (has_app := eq_refl) (has_pars := eq_refl) (has_cstrblocks := eq_refl). + constructors_as_blocks_transformation (efl := efl) (has_app := eq_refl) (has_pars := eq_refl) (has_rel := eq_refl) (hasbox := eq_refl) (has_cstrblocks := eq_refl). (* At the end of erasure we get a well-formed program (well-scoped globally and localy), without parameters in inductive declarations. The constructor applications are also transformed to a first-order @@ -105,6 +106,54 @@ Next Obligation. now eapply ETransform.expanded_eprogram_env_expanded_eprogram_cstrs. Qed. + +Definition transform_compose + {program program' program'' value value' value'' : Type} + {eval : program -> value -> Prop} {eval' : program' -> value' -> Prop} + {eval'' : program'' -> value'' -> Prop} + (o : t program program' value value' eval eval') + (o' : t program' program'' value' value'' eval' eval'') + (pre : forall p : program', post o p -> pre o' p) : + forall x p1, exists p3, + transform (compose o o' pre) x p1 = transform o' (transform o x p1) p3. +Proof. + cbn. intros. + unfold run, time. + eexists;reflexivity. +Qed. + +#[global] +Instance pcuic_expand_lets_transform_ext {cf : checker_flags} K : + TransformExt.t (pcuic_expand_lets_transform K) extends_pcuic_program PCUICAst.term snd. +Proof. + red. intros p p' pre pre' []. cbn. now rewrite H0. +Qed. + +Lemma verified_erasure_pipeline_extends {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : + TransformExt.t verified_erasure_pipeline extends_pcuic_program EAst.term snd. +Proof. + unfold verified_erasure_pipeline. + eapply TransformExt.compose. + 2:eapply TransformExt.compose. + 2: typeclasses eauto. 2:typeclasses eauto. + 2:{ auto. } + repeat try (eapply TransformExt.compose; [|typeclasses eauto|]). + typeclasses eauto. + all:unfold extends_pcuic_program; cbn. + all:intros; intuition auto; try congruence. + red. cbn. + + + + + + + + + } + red. + + Program Definition pre_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : Transform.t TemplateProgram.template_program pcuic_program Ast.term PCUICAst.term @@ -160,14 +209,14 @@ Program Definition erasure_pipeline_fast {guard : abstract_guard_impl} (efl := E pcuic_expand_lets_transform (K _ T1) ▷ erase_transform ▷ guarded_to_unguarded_fix (wcon := eq_refl) eq_refl ▷ - remove_params_fast_optimization (wcon := eq_refl) _ ▷ + remove_params_fast_optimization (wcon := eq_refl) ▷ rebuild_wf_env_transform (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) true ▷ remove_match_on_box_trans (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ rebuild_wf_env_transform (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) true ▷ inline_projections_optimization (fl := EWcbvEval.target_wcbv_flags) (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ let efl := EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags) in rebuild_wf_env_transform (efl := efl) true ▷ - constructors_as_blocks_transformation efl (has_app := eq_refl) (has_pars := eq_refl) (has_cstrblocks := eq_refl). + constructors_as_blocks_transformation (efl := efl) (has_app := eq_refl) (has_pars := eq_refl) (has_rel := eq_refl) (hasbox := eq_refl) (has_cstrblocks := eq_refl). Next Obligation. destruct H; split => //. now eapply ETransform.expanded_eprogram_env_expanded_eprogram_cstrs. Qed. diff --git a/erasure/theories/EConstructorsAsBlocks.v b/erasure/theories/EConstructorsAsBlocks.v index 829d0d089..a8cd21c0a 100644 --- a/erasure/theories/EConstructorsAsBlocks.v +++ b/erasure/theories/EConstructorsAsBlocks.v @@ -432,7 +432,7 @@ Definition transform_blocks_constant_decl Σ cb := Definition transform_blocks_decl Σ d := match d with | ConstantDecl cb => ConstantDecl (transform_blocks_constant_decl Σ cb) - | InductiveDecl idecl => d + | InductiveDecl idecl => InductiveDecl idecl end. Definition transform_blocks_env Σ := @@ -657,6 +657,15 @@ Proof. rewrite lookup_env_transform_blocks H //. Qed. +Lemma lookup_inductive_transform_blocks Σ ind : + lookup_inductive (transform_blocks_env Σ) ind = + lookup_inductive Σ ind. +Proof. + unfold lookup_inductive, lookup_minductive in *. + rewrite lookup_env_transform_blocks. + destruct lookup_env as [ [] | ]; cbn; congruence. +Qed. + Lemma lookup_constructor_transform_blocks Σ ind c : lookup_constructor (transform_blocks_env Σ) ind c = lookup_constructor Σ ind c. @@ -669,45 +678,57 @@ Qed. Lemma isLambda_transform_blocks Σ c : isLambda c -> isLambda (transform_blocks Σ c). Proof. destruct c => //. Qed. -Lemma transform_wellformed' {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : +Lemma transform_wellformed {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : has_cstr_params = false -> cstr_as_blocks = false -> has_tApp -> wf_glob Σ -> @wellformed efl Σ n t -> isEtaExp Σ t -> - @wellformed (switch_cstr_as_blocks efl) Σ n (transform_blocks Σ t). + @wellformed (switch_cstr_as_blocks efl) (transform_blocks_env Σ) n (transform_blocks Σ t). Proof. intros hasp cstrbl hasa. - revert n. funelim (transform_blocks Σ t); simp_eta; cbn -[transform_blocks - lookup_inductive lookup_constructor lookup_constructor_pars_args + revert n. move t after hasa. move Σ after hasa. + funelim (transform_blocks Σ t); simp_eta; cbn -[transform_blocks transform_blocks_env + lookup_inductive lookup_constructor lookup_constructor_pars_args lookup_constant GlobalContextMap.lookup_constructor_pars_args isEtaExp]; intros m Hwf Hw; rtoProp; try split; eauto. - all: rewrite ?map_InP_spec; toAll; eauto; try now solve_all. - - rewrite cstrbl in H0. destruct H2. unfold isEtaExp_app in H2. unfold lookup_constructor_pars_args in *. + all: rewrite ?map_InP_spec; toAll; intuition eauto; try now solve_all. + - rewrite /lookup_constant lookup_env_transform_blocks in H0 *. + destruct lookup_env => //=; cbn in H0. destruct g => //; rewrite /transform_blocks_decl //=. + destruct (cst_body c) => //. + - rewrite cstrbl in H0. + rewrite lookup_constructor_transform_blocks. + unfold isEtaExp_app in H3. unfold lookup_constructor_pars_args in *. destruct (lookup_constructor Σ) as [[[]] | ]; try congruence; cbn - [transform_blocks]. - 2: eauto. split; auto. cbn in H2. eapply Nat.leb_le in H2. - apply/eqb_spec. lia. - - destruct H5. solve_all. solve_all. + - rewrite cstrbl in H0. + unfold isEtaExp_app in H3. + rewrite /lookup_constructor_pars_args lookup_constructor_transform_blocks //=. + rewrite /lookup_constructor_pars_args in H3. + destruct lookup_constructor as [[[] ?]|]=> //. cbn in H3. + eapply Nat.leb_le in H3. intuition auto. apply/eqb_spec. lia. + - now rewrite lookup_inductive_transform_blocks. + - now rewrite lookup_constructor_transform_blocks. - unfold wf_fix in *. rtoProp. solve_all. solve_all. now eapply isLambda_transform_blocks. - unfold wf_fix in *. rtoProp. solve_all. len. solve_all. len. destruct x. cbn -[transform_blocks isEtaExp] in *. rtoProp. eauto. - unfold wf_fix in *. len. solve_all. rtoProp; intuition auto. solve_all. - - rewrite !wellformed_mkApps in Hw |- * => //. rtoProp. intros. - eapply isEtaExp_mkApps in H3. rewrite decompose_app_mkApps in H3; eauto. + - rewrite !wellformed_mkApps in Hw |- * => //. rtoProp. + eapply isEtaExp_mkApps in H1. rewrite decompose_app_mkApps in H1; eauto. destruct construct_viewc; eauto. cbn in d. eauto. - rtoProp. eauto. repeat solve_all. + split; rtoProp; intuition eauto. solve_all; intuition eauto. - Opaque isEtaExp. destruct chop eqn:Ec. rewrite !wellformed_mkApps in Hw |- * => //. rtoProp. rewrite GlobalContextMap.lookup_constructor_pars_args_spec in Heq. - cbn -[lookup_constructor transform_blocks ] in *. intros. rtoProp. - rewrite isEtaExp_Constructor in H2. - rtoProp. unfold isEtaExp_app in *. unfold lookup_constructor_pars_args in H2. + cbn -[lookup_constructor transform_blocks ] in *. rewrite cstrbl in H1. + rewrite lookup_constructor_transform_blocks. intros. rtoProp. + rewrite isEtaExp_Constructor in H0. + rtoProp. unfold isEtaExp_app in *. unfold lookup_constructor_pars_args in H0. repeat split; eauto; rewrite ?lookup_constructor_transform_blocks; eauto. * destruct lookup_constructor as [ [[]] | ] eqn:E; cbn -[transform_blocks] in *; eauto. invs Heq. rewrite chop_firstn_skipn in Ec. invs Ec. - rewrite firstn_length. len. eapply Nat.leb_le in H2. + rewrite firstn_length. len. eapply Nat.leb_le in H0. apply/eqb_spec. assert (ind_npars m0 = 0). { destruct lookup_env as [ [] | ] eqn:E'; try congruence. @@ -725,25 +746,30 @@ Proof. solve_all. eapply All_skipn. solve_all. - rewrite wellformed_mkApps in Hw; eauto. rtoProp. cbn in *. rtoProp. cbn in *. destruct lookup_env as [[] | ]; cbn in *; eauto; try congruence. + - rewrite lookup_constructor_transform_blocks. rewrite wellformed_mkApps in Hw => //. + move/andP: Hw => [wf _]. + simpl in wf. rtoProp; intuition auto. - rewrite isEtaExp_Constructor in H0. rtoProp. rewrite GlobalContextMap.lookup_constructor_pars_args_spec in Heq; unfold lookup_constructor_pars_args in *. - destruct lookup_constructor as [ [[]] | ]; cbn in Heq; try congruence. - cbn. split; eauto. rewrite wellformed_mkApps in Hw; eauto. rtoProp. solve_all. + rewrite lookup_constructor_transform_blocks. + now destruct lookup_constructor as [ [[]] | ]; cbn in Heq; try congruence. + - rewrite wellformed_mkApps in Hw => //. apply isEtaExp_mkApps in H0. + rewrite decompose_app_mkApps in H0 => //. cbn in H0. rtoProp; intuition auto. solve_all. Qed. -Lemma transform_wellformed_decl' {efl : EEnvFlags} {Σ : GlobalContextMap.t} d : +Lemma transform_wellformed_decl {efl : EEnvFlags} {Σ : GlobalContextMap.t} d : has_cstr_params = false -> cstr_as_blocks = false -> has_tApp -> wf_glob Σ -> @wf_global_decl efl Σ d -> isEtaExp_decl Σ d -> - @wf_global_decl (switch_cstr_as_blocks efl) Σ (transform_blocks_decl Σ d). + @wf_global_decl (switch_cstr_as_blocks efl) (transform_blocks_env Σ) (transform_blocks_decl Σ d). Proof. intros wf wfd etad. destruct d => //=. cbn. destruct c as [[]] => //=. - eapply transform_wellformed'; tea. + eapply transform_wellformed; tea. Qed. From MetaCoq.Erasure Require Import EGenericMapEnv. @@ -804,23 +830,13 @@ Proof. eapply H0. Qed. -Lemma transform_wellformed {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : - has_cstr_params = false -> - cstr_as_blocks = false -> - has_tApp -> - wf_glob Σ -> - @wellformed efl Σ n t -> - isEtaExp Σ t -> - @wellformed (switch_cstr_as_blocks efl) (transform_blocks_env Σ) n (transform_blocks Σ t). -Proof. - intros. eapply gen_transform_wellformed_irrel; eauto. - eapply transform_wellformed'; eauto. -Qed. - From MetaCoq.Erasure Require Import EGenericGlobalMap. #[local] -Instance GT : GenTransform := { gen_transform := transform_blocks }. +Instance GT : GenTransform := { gen_transform := transform_blocks; gen_transform_inductive_decl := id }. +#[local] +Instance GTID : GenTransformId GT. +Proof. red; reflexivity. Qed. #[local] Instance GTExt efl : has_tApp -> GenTransformExtends efl (switch_cstr_as_blocks efl) GT. Proof. @@ -838,9 +854,10 @@ Proof. cstr_as_blocks = false /\ isEtaExp Σ t |}; auto. intros Σ n t [hasapp [cstrp [cstrb pre]]] wfΣ wft. - apply transform_wellformed => //. + eapply transform_wellformed => //. Defined. + Lemma Pre_glob efl Σ wf : has_cstr_params = false -> cstr_as_blocks = false -> diff --git a/erasure/theories/EImplementBox.v b/erasure/theories/EImplementBox.v index 3d854bbb3..e843cd6ba 100644 --- a/erasure/theories/EImplementBox.v +++ b/erasure/theories/EImplementBox.v @@ -190,7 +190,7 @@ Definition implement_box_constant_decl cb := Definition implement_box_decl d := match d with | ConstantDecl cb => ConstantDecl (implement_box_constant_decl cb) - | InductiveDecl idecl => d + | InductiveDecl idecl => InductiveDecl idecl end. Definition implement_box_env (Σ : global_declarations) := @@ -325,6 +325,21 @@ Proof. destruct lookup_env as [ [] | ]; cbn; congruence. Qed. +Lemma lookup_inductive_implement_box Σ ind : + lookup_inductive (implement_box_env Σ) ind = + lookup_inductive Σ ind. +Proof. + unfold lookup_constructor, lookup_inductive, lookup_minductive in *. + rewrite lookup_env_implement_box. + destruct lookup_env as [ [] | ]; cbn; congruence. +Qed. + +Lemma lookup_constructor_pars_args_implement_box {efl : EEnvFlags} {Σ ind n} : + lookup_constructor_pars_args (implement_box_env Σ) ind n = lookup_constructor_pars_args Σ ind n. +Proof. + cbn -[lookup_constructor]. now rewrite lookup_constructor_implement_box. +Qed. + Lemma isLambda_implement_box c : isLambda c -> isLambda (implement_box c). Proof. destruct c => //. Qed. @@ -352,20 +367,23 @@ Lemma transform_wellformed' {efl : EEnvFlags} {Σ : list (kername × global_decl has_tApp -> wf_glob Σ -> @wellformed efl Σ n t -> - @wellformed (switch_off_box efl) Σ n (implement_box t). + @wellformed (switch_off_box efl) (implement_box_env Σ) n (implement_box t). Proof. intros hasa. revert n. funelim (implement_box t); simp_eta; cbn -[implement_box lookup_inductive lookup_constructor lookup_constructor_pars_args GlobalContextMap.lookup_constructor_pars_args isEtaExp]; intros m Hwf Hw; rtoProp; try split; eauto. all: rewrite ?map_InP_spec; toAll; eauto; try now solve_all. + - rewrite lookup_env_implement_box. destruct (lookup_env Σ n) => //. destruct g => //=. + destruct (cst_body c) => //=. - unfold lookup_constructor_pars_args in *. - destruct (lookup_constructor Σ) as [[[]] | ]; try congruence; cbn - [implement_box]. - all: destruct cstr_as_blocks; cbn in *; rtoProp; try split; eauto. - + now rewrite map_length. + rewrite lookup_constructor_implement_box. rewrite H2. intuition auto. + - rewrite lookup_constructor_pars_args_implement_box. len. + all: destruct cstr_as_blocks; rtoProp; try split; eauto. + solve_all. + destruct block_args; cbn in *; eauto. - - repeat split; eauto. solve_all. + - rewrite lookup_inductive_implement_box. intuition auto. solve_all. + - rewrite lookup_constructor_implement_box. intuition auto. - unfold wf_fix in *. rtoProp. solve_all. solve_all. now eapply isLambda_implement_box. - unfold wf_fix in *. rtoProp. solve_all. len. solve_all. len. destruct x. @@ -378,7 +396,7 @@ Lemma transform_wellformed_decl' {efl : EEnvFlags} {Σ : global_declarations} d has_tApp -> wf_glob Σ -> @wf_global_decl efl Σ d -> - @wf_global_decl (switch_off_box efl) Σ (implement_box_decl d). + @wf_global_decl (switch_off_box efl) (implement_box_env Σ) (implement_box_decl d). Proof. intros. destruct d => //=. cbn. @@ -386,40 +404,7 @@ Proof. eapply transform_wellformed'; tea. Qed. -From MetaCoq.Erasure Require Import EGenericMapEnv. - -Lemma transform_wellformed {efl : EEnvFlags} {Σ : global_declarations} n t : - has_tApp -> - wf_glob Σ -> - @wellformed efl Σ n t -> - @wellformed (switch_off_box efl) (implement_box_env Σ) n (implement_box t). -Proof. - intros. eapply gen_transform_wellformed_irrel; eauto. - eapply transform_wellformed'; eauto. -Qed. - -Lemma optimize_decl_wf {efl : EEnvFlags} {Σ : global_declarations} : - has_tApp -> - wf_glob (efl := efl) Σ -> - forall d, - wf_global_decl (efl := efl) Σ d -> - wf_global_decl (efl := switch_off_box efl) (implement_box_env Σ) (implement_box_decl d). -Proof. - intros. - destruct d => /= //. - rewrite /isEtaExp_constant_decl. cbn in H1. - destruct (cst_body c) => /= //. - eapply transform_wellformed => //. -Qed. - -Lemma fresh_global_optimize_env {Σ : global_declarations} kn : - fresh_global kn Σ -> - fresh_global kn (implement_box_env Σ). -Proof. - induction 1; cbn; constructor; auto. -Qed. - -Lemma fresh_global_map_on_snd Σ f kn : +Lemma fresh_global_map_on_snd {Σ : global_context} f kn : fresh_global kn Σ -> fresh_global kn (map (on_snd f) Σ). Proof. @@ -430,19 +415,15 @@ Lemma transform_wf_global {efl : EEnvFlags} {Σ : global_declarations} : has_tApp -> wf_glob (efl := efl) Σ -> wf_glob (efl := switch_off_box efl) (implement_box_env Σ). Proof. - intros hasapp wfg. clear - hasapp wfg. - assert (extends Σ Σ). now exists []. + intros hasapp wfg. + assert (extends_prefix Σ Σ). now exists []. revert H wfg. generalize Σ at 2 3. induction Σ; cbn; constructor; auto. - eapply IHΣ; rtoProp; intuition eauto. destruct H. subst Σ0. exists (x ++ [a]). now rewrite -app_assoc. - epose proof (EExtends.extends_wf_glob _ H wfg); tea. depelim H0. cbn. - eapply gen_transform_wellformed_decl_irrel; trea. - destruct d; cbn in *; eauto. - destruct (cst_body c); eauto. - cbn -[implement_box]. cbn in H1. - eapply transform_wellformed'; eauto. + now eapply transform_wellformed_decl'. - eapply fresh_global_map_on_snd. eapply EExtends.extends_wf_glob in wfg; tea. now depelim wfg. Qed. diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 565cdcfbd..af091df81 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -319,7 +319,7 @@ Definition remove_match_on_box_constant_decl Σ cb := Definition remove_match_on_box_decl Σ d := match d with | ConstantDecl cb => ConstantDecl (remove_match_on_box_constant_decl Σ cb) - | InductiveDecl idecl => d + | InductiveDecl idecl => InductiveDecl idecl end. Definition remove_match_on_box_env Σ := diff --git a/erasure/theories/EProgram.v b/erasure/theories/EProgram.v index 699164a66..cd2d7cc49 100644 --- a/erasure/theories/EProgram.v +++ b/erasure/theories/EProgram.v @@ -51,3 +51,45 @@ Proof. intros wf; depelim wf. constructor; auto. Qed. +Module TransformExt. + Section Opt. + Context {program program' : Type}. + Context {value value' : Type}. + Context {eval : program -> value -> Prop}. + Context {eval' : program' -> value' -> Prop}. + Context (o : Transform.t program program' value value' eval eval'). + Context (extends : program -> program -> Prop). + Context (extends' : program' -> program' -> Prop). + + Class t := preserves_obs : forall p p' (pr : o.(pre) p) (pr' : o.(pre) p'), + extends p p' -> extends' (o.(transform) p pr) (o.(transform) p' pr'). + + End Opt. + + Section Comp. + Context {program program' program'' : Type}. + Context {value value' value'' : Type}. + Context {eval : program -> value -> Prop}. + Context {eval' : program' -> value' -> Prop}. + Context {eval'' : program'' -> value'' -> Prop}. + + Context {extends : program -> program -> Prop}. + Context {extends' : program' -> program' -> Prop}. + Context {extends'' : program'' -> program'' -> Prop}. + + Local Obligation Tactic := idtac. + #[global] + Instance compose (o : Transform.t program program' value value' eval eval') + (o' : Transform.t program' program'' value' value'' eval' eval'') + (oext : t o extends extends') + (o'ext : t o' extends' extends'') + (hpp : (forall p, o.(post) p -> o'.(pre) p)) + : t (Transform.compose o o' hpp) extends extends''. + Proof. red. + intros p p' pr pr' ext. + eapply o'ext. now apply oext. + Qed. + + End Comp. + +End TransformExt. diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index b1d033023..10eceaf4b 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -764,25 +764,16 @@ Proof. destruct x => //. Qed. +Definition extends_global_env (Σ Σ' : PCUICAst.PCUICEnvironment.global_env) := + (forall kn decl, PCUICAst.PCUICEnvironment.lookup_env Σ kn = Some decl -> + PCUICAst.PCUICEnvironment.lookup_env Σ' kn = Some decl) /\ + (forall tag, primitive_constant Σ tag = primitive_constant Σ' tag). + Lemma is_erasableb_irrel_global_env {X_type X} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} {X_type' X'} {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} {Γ t wt wt'} : - (forall Σ Σ' : global_env_ext, abstract_env_ext_rel X Σ -> - abstract_env_ext_rel X' Σ' -> ∥ extends_decls Σ' Σ ∥ /\ (Σ.2 = Σ'.2)) -> + forall (hl : Hlookup X_type X X_type' X'), is_erasableb X_type X Γ t wt = is_erasableb X_type' X' Γ t wt'. Proof. - intro ext. - (* ext eqext. *) - assert (hl : Hlookup X_type X X_type' X'). - { red. intros Σ H Σ' H0. specialize (ext _ _ H H0) as [[X0] H1]. - split. intros kn decl decl' H2 H3. - rewrite -(abstract_env_lookup_correct' _ _ H). - rewrite -(abstract_env_lookup_correct' _ _ H0). - rewrite H2 H3. pose proof (abstract_env_ext_wf _ H) as [X1]. - eapply lookup_env_extends_NoDup in H3; try eapply NoDup_on_global_decls; try eapply X1; eauto; tc. cbv [PCUICEnvironment.fst_ctx] in *. congruence. - destruct X0. intros tag. - rewrite (abstract_primitive_constant_correct _ _ _ H). - rewrite (abstract_primitive_constant_correct _ _ _ H0). - unfold primitive_constant. - rewrite e0. congruence. } + intros hl. simp is_erasableb. set (obl := is_erasableb_obligation_2 _ _ _ _). clearbody obl. set(ty := (type_of_typing X_type' _ _ _ wt')) in *. @@ -809,12 +800,28 @@ Ltac iserasableb_irrel_env := generalize dependent H; rewrite (@is_erasableb_irrel_global_env _ _ _ _ _ _ _ _ wt wt') //; intros; rewrite Heq end. + (* + (forall Σ Σ' : global_env_ext, abstract_env_ext_rel X Σ -> + abstract_env_ext_rel X' Σ' -> ∥ extends_decls Σ' Σ ∥ /\ (Σ.2 = Σ'.2)) -> + + assert (hl : Hlookup X_type X X_type' X'). + { red. intros Σ H Σ' H0. specialize (ext _ _ H H0) as [[X0] H1]. + split. intros kn decl decl' H2 H3. + rewrite -(abstract_env_lookup_correct' _ _ H). + rewrite -(abstract_env_lookup_correct' _ _ H0). + rewrite H2 H3. + red in X0. specialize (proj1 X0 kn decl' H3). congruence. + destruct X0. intros tag. + rewrite (abstract_primitive_constant_correct _ _ _ H). + rewrite (abstract_primitive_constant_correct _ _ _ H0). + now symmetry. } +*) + Lemma erase_irrel_global_env {X_type X} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} {X_type' X'} {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} {Γ t wt wt'} : - (forall Σ Σ' : global_env_ext, abstract_env_ext_rel X Σ -> - abstract_env_ext_rel X' Σ' -> ∥ extends_decls Σ' Σ ∥ /\ (Σ.2 = Σ'.2)) -> + forall (hl : Hlookup X_type X X_type' X'), erase X_type X Γ t wt = erase X_type' X' Γ t wt'. Proof. - intros ext. + intros hl. move: wt'. eapply (erase_elim X_type X (fun Γ t wt e => @@ -840,10 +847,9 @@ Proof. all:bang. Qed. -Lemma erase_constant_body_suffix {X_type X} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} {X_type' X'} {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} {cb ondecl ondecl'} : -(forall Σ Σ' : global_env_ext, abstract_env_ext_rel X Σ -> - abstract_env_ext_rel X' Σ' -> ∥ extends_decls Σ' Σ ∥ /\ (Σ.2 = Σ'.2)) -> -erase_constant_body X_type X cb ondecl = erase_constant_body X_type' X' cb ondecl'. +Lemma erase_constant_body_irrel {X_type X} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} {X_type' X'} {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} {cb ondecl ondecl'} : + forall (hl : Hlookup X_type X X_type' X'), + erase_constant_body X_type X cb ondecl = erase_constant_body X_type' X' cb ondecl'. Proof. intros ext. destruct cb => //=. @@ -884,22 +890,57 @@ Notation NormalizationIn_erase_global_decls X decls end) (only parsing). +Lemma normalization_in_inv {X_type : abstract_env_impl} (X : X_type.π1) + (decls : global_declarations) kn cb + (normalization_in : NormalizationIn_erase_global_decls X ((kn, ConstantDecl cb) :: decls)) : + NormalizationIn_erase_global_decls (abstract_pop_decls X) decls. +Proof. + cbv [id]. intros. + unshelve eapply (normalization_in (S n)); cbn; tea. lia. +Defined. + +Program Definition erase_constant_decl {X_type : abstract_env_impl} (X : X_type.π1) + (cb : constant_body) + (normalization_in : forall pf, + let Xext := abstract_make_wf_env_ext X (cst_universes cb) pf in + forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ) + (oncb : forall Σ : global_env, Σ ∼ X -> ∥ wf_ext (Σ, cst_universes cb) * on_constant_decl (lift_typing typing) (Σ, cst_universes cb) cb ∥) := + let Xext := abstract_make_wf_env_ext X (cst_universes cb) _ in + let normalization_in' : @id Prop _ := _ in (* hack to avoid program erroring out on unresolved tc *) + @erase_constant_body X_type Xext normalization_in' cb _. +Next Obligation. +pose proof (abstract_env_wf _ H) as [wf]. +specialize (oncb _ H). destruct oncb. sq. intuition auto. +Qed. + +Next Obligation. + unfold id. intros. + unshelve eapply normalization_in; eauto. +Defined. + +Next Obligation. + destruct Σ as [Σ ext]. + pose proof (abstract_env_exists X) as [[? HX]]. + pose proof (abstract_env_wf _ HX) as [wfX]. + pose proof (abstract_make_wf_env_ext_correct X (cst_universes cb) _ _ _ HX H). + noconf H0. + clear H; specialize (oncb _ HX). destruct oncb as [[]]; sq; auto. +Qed. + Program Fixpoint erase_global_decls {X_type : abstract_env_impl} (deps : KernameSet.t) (X : X_type.π1) (decls : global_declarations) {normalization_in : NormalizationIn_erase_global_decls X decls} (prop : forall Σ : global_env, abstract_env_rel X Σ -> Σ.(declarations) = decls) - : E.global_declarations := + : E.global_declarations * KernameSet.t := match decls with - | [] => [] + | [] => ([], deps) | (kn, ConstantDecl cb) :: decls => let X' := abstract_pop_decls X in if KernameSet.mem kn deps then - let Xext := abstract_make_wf_env_ext X' (cst_universes cb) _ in - let normalization_in' : @id Prop _ := _ in (* hack to avoid program erroring out on unresolved tc *) - let cb' := @erase_constant_body X_type Xext normalization_in' cb _ in + let cb' := @erase_constant_decl X_type X' cb _ _ in let deps := KernameSet.union deps (snd cb') in let X'' := erase_global_decls deps X' decls _ in - ((kn, E.ConstantDecl (fst cb')) :: X'') + (((kn, E.ConstantDecl (fst cb')) :: fst X''), snd X'') else erase_global_decls deps X' decls _ @@ -907,37 +948,26 @@ Program Fixpoint erase_global_decls let X' := abstract_pop_decls X in if KernameSet.mem kn deps then let mib' := erase_mutual_inductive_body mib in - let X'' := erase_global_decls deps X' decls _ in - ((kn, E.InductiveDecl mib') :: X'') + let X'':= erase_global_decls deps X' decls _ in + (((kn, E.InductiveDecl mib') :: fst X''), snd X'') else erase_global_decls deps X' decls _ end. -Next Obligation. - pose proof (abstract_env_wf _ H) as [wf]. - pose proof (abstract_env_exists X) as [[? HX]]. - pose proof (abstract_env_wf _ HX) as [wfX]. - assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). - { now eexists. } - destruct (abstract_pop_decls_correct X decls prop' _ _ HX H) as [? []]. - clear H. specialize (prop _ HX). destruct x, Σ, H0; cbn in *. - subst. sq. destruct wfX. depelim o0. destruct o1. split => //. -Qed. Next Obligation. cbv [id]. - unshelve eapply (normalization_in 0); cbn; try reflexivity; try lia. + unshelve eapply (normalization_in 0); tea; cbn; try reflexivity; try lia. Defined. Next Obligation. - pose proof (abstract_env_ext_wf _ H) as [wf]. - pose proof (abstract_env_exists (abstract_pop_decls X)) as [[? HX']]. - pose proof (abstract_env_wf _ HX') as [wfX']. pose proof (abstract_env_exists X) as [[? HX]]. pose proof (abstract_env_wf _ HX) as [wfX]. assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). { now eexists. } - pose proof (abstract_pop_decls_correct X decls prop' _ _ HX HX') as [? []]. - pose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ HX' H). - clear H HX'. specialize (prop _ HX). destruct x, Σ as [[] u], H0; cbn in *. - subst. sq. inversion H3. subst. clear H3. destruct wfX. cbn in *. - rewrite prop in o0. depelim o0. destruct o1. exact on_global_decl_d. + pose proof (abstract_pop_decls_correct X decls prop' _ _ HX H). + destruct x, Σ as [[] u], H0; cbn in *. + subst. sq. inversion H1. subst. destruct wfX. cbn in *. + specialize (prop _ HX). + cbn in prop. + rewrite prop in o0. depelim o0. destruct o1. + split. 2:exact on_global_decl_d. split => //. Qed. Next Obligation. unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. @@ -984,16 +1014,30 @@ Lemma erase_global_decls_irr erase_global_decls deps X decls (normalization_in:=normalization_in') prf'. Proof. revert X deps normalization_in normalization_in' prf prf'. - induction decls; eauto. intros. destruct a as [kn []]; cbn. - - destruct KernameSet.mem; cbn. - erewrite erase_constant_body_suffix. f_equal. - eapply IHdecls. intros. - pose proof (abstract_env_exists (abstract_pop_decls X)) as [[env wf]]. - rewrite (abstract_make_wf_env_ext_correct (abstract_pop_decls X) (cst_universes c) _ _ _ wf H). - now rewrite (abstract_make_wf_env_ext_correct (abstract_pop_decls X) (cst_universes c) _ _ _ wf H0). - apply IHdecls. - - destruct KernameSet.mem; cbn. f_equal. eapply IHdecls. - apply IHdecls. + induction decls; eauto. intros. destruct a as [kn []]. + - cbn. + set (ec := erase_constant_decl _ _ _ _). + set (ec' := erase_constant_decl _ _ _ _). + assert (ec = ec') as <-. + { subst ec ec'. + unfold erase_constant_decl. + erewrite erase_constant_body_irrel. reflexivity. + pose proof (abstract_env_exists (abstract_pop_decls X)) as [[env wf]]. + red. intros. + pose proof (abstract_make_wf_env_ext_correct (abstract_pop_decls X) (cst_universes c) _ _ _ wf H). + pose proof (abstract_make_wf_env_ext_correct (abstract_pop_decls X) (cst_universes c) _ _ _ wf H0). + split => //. + { intros. + rewrite -(abstract_env_lookup_correct' _ _ H). + rewrite -(abstract_env_lookup_correct' _ _ H0). now rewrite H1 H2. } + intros. + { rewrite (abstract_primitive_constant_correct _ _ _ H). + rewrite (abstract_primitive_constant_correct _ _ _ H0). + now rewrite H1 H2. } } + destruct KernameSet.mem; cbn. + f_equal; f_equal. f_equal; eapply IHdecls. + apply IHdecls. apply IHdecls. + - cbn. destruct KernameSet.mem; cbn. f_equal. f_equal. 1-2:f_equal. all:eapply IHdecls. Qed. Definition global_erased_with_deps (Σ : global_env) (Σ' : EAst.global_declarations) kn := @@ -1508,7 +1552,7 @@ Lemma erase_global_includes X_type (X:X_type.π1) deps decls {normalization_in} KernameSet.subset deps' deps -> forall Σ : global_env, abstract_env_rel X Σ -> let Σ' := erase_global_decls deps X decls (normalization_in:=normalization_in) prf in - includes_deps Σ Σ' deps'. + includes_deps Σ (fst Σ') deps'. Proof. intros ? sub Σ wfΣ; cbn. induction decls in X, H, sub, prf, deps, deps', Σ , wfΣ, normalization_in |- *. - simpl. intros i hin. elimtype False. @@ -1524,6 +1568,7 @@ Proof. elim: H0. intros -> [= <-]. * destruct d as [|]; [left|right]. { cbn. set (Xpop := abstract_pop_decls X). + unfold erase_constant_decl. set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _). epose proof (abstract_env_exists Xpop) as [[Σpop wfpop]]. epose proof (abstract_env_ext_exists Xmake) as [[Σmake wfmake]]. @@ -1594,6 +1639,7 @@ Proof. * intros ikn Hi. destruct d as [|]. ++ simpl. destruct (KernameSet.mem kn deps) eqn:eqkn. + unfold erase_constant_decl. set (Xpop := abstract_pop_decls X). set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _). epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. @@ -1667,7 +1713,7 @@ Lemma erase_correct (wfl := Ee.default_wcbv_flags) X_type (X : X_type.π1) erase_global_decls deps X decls (normalization_in:=normalization_in) prf = Σ' -> (forall Σ : global_env, abstract_env_rel X Σ -> Σ |-p t ⇓ v) -> forall Σ : global_env_ext, abstract_env_ext_rel Xext Σ -> - exists v', Σ ;;; [] |- v ⇝ℇ v' /\ ∥ Σ' ⊢ t' ⇓ v' ∥. + exists v', Σ ;;; [] |- v ⇝ℇ v' /\ ∥ fst Σ' ⊢ t' ⇓ v' ∥. Proof. intros wt. intros HΣ' Hsub Ht' ev Σext wfΣex. @@ -1676,7 +1722,7 @@ Proof. destruct (wt _ wfΣex) as [T wt']. pose proof (abstract_env_ext_wf _ wfΣex) as [wfΣ]. specialize (H _ wfΣex). - unshelve epose proof (erase_global_erases_deps (Σ' := Σ') wfΣ wt' H _); cycle 2. + unshelve epose proof (erase_global_erases_deps (Σ' := fst Σ') wfΣ wt' H _); cycle 2. rewrite <- Ht'. eapply erase_global_includes; eauto. intros. eapply term_global_deps_spec in H; eauto. @@ -1802,7 +1848,7 @@ Local Hint Constructors expanded : expanded. Local Arguments erase_global_decls _ _ _ _ _ : clear implicits. -Lemma lookup_env_erase X_type X deps decls normalization_in prf kn d : +(*Lemma lookup_env_erase X_type X deps decls normalization_in prf kn d : KernameSet.In kn deps -> forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = Some (InductiveDecl d) -> EGlobalEnv.lookup_env (erase_global_decls X_type deps X decls normalization_in prf) kn = Some (EAst.InductiveDecl (erase_mutual_inductive_body d)). @@ -1834,18 +1880,352 @@ Proof. eapply IHdecls => //; eauto. eapply IHdecls => //; eauto. Qed. +*) + +Lemma lookup_env_erase_decl_None X_type X deps decls normalization_in prf kn : + forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = None -> + EGlobalEnv.lookup_env (erase_global_decls X_type deps X decls normalization_in prf).1 kn = None. +Proof. + intros Σ wfΣ. + pose proof (prf _ wfΣ). + unfold lookup_env. rewrite H. clear H. + rewrite /lookup_env. + induction decls in X, Σ , wfΣ ,prf, deps, normalization_in |- *. + - move=> /= //. + - destruct a as [kn' d']. + cbn -[erase_global_decls]. + case: (eqb_spec kn kn'); intros e'; subst. + intros [= ->]. + cbn. destruct d'. cbn. + case_eq (KernameSet.mem kn' deps); move => hin'; cbn. + + destruct (eqb_spec kn kn') => //. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + now eapply IHdecls. + + intros. set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + now eapply IHdecls. + + intros hin. + case_eq (KernameSet.mem kn' deps); move => hin'; cbn. + destruct (eqb_spec kn kn') => //. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + now eapply IHdecls. + intros. set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + now eapply IHdecls. +Qed. + +Lemma lookup_env_erase_global_diff X_type X deps decls kn kn' d' normalization_in prf : + kn <> kn' -> + exists deps' nin' prf', + KernameSet.Subset deps deps' /\ + EGlobalEnv.lookup_env (erase_global_decls X_type deps X ((kn', d') :: decls) normalization_in prf).1 kn = + EGlobalEnv.lookup_env (erase_global_decls X_type deps' (abstract_pop_decls X) decls nin' prf').1 kn /\ + (KernameSet.In kn (erase_global_decls X_type deps X ((kn', d') :: decls) normalization_in prf).2 -> + KernameSet.In kn (erase_global_decls X_type deps' (abstract_pop_decls X) decls nin' prf').2). +Proof. + intros hd. + simpl. destruct d'. + + destruct KernameSet.mem. cbn. + destruct (eqb_spec kn kn') => //. + do 3 eexists. split; [|split;[reflexivity|]]. KernameSetDecide.fsetdec. + auto. + do 3 eexists. split; [|split;[reflexivity|]]. KernameSetDecide.fsetdec. + auto. + + destruct KernameSet.mem. cbn. + destruct (eqb_spec kn kn') => //. + do 3 eexists. split; [|split;[reflexivity|]]. KernameSetDecide.fsetdec. auto. + do 3 eexists. split; [|split;[reflexivity|]]. KernameSetDecide.fsetdec. auto. +Qed. + +Lemma lookup_env_In_map_fst Σ kn decl : lookup_global Σ kn = Some decl -> In kn (map fst Σ). +Proof. + induction Σ; cbn => //. + case: eqb_spec. + + intros -> [= <-]. now left. + + intros _ hl. eauto. +Qed. + +Lemma erase_constant_decl_deps {X_type : abstract_env_impl} (X : X_type.π1) (cb : constant_body) normalization_in oncb : + forall Σ : global_env, abstract_env_rel X Σ -> + forall kn, KernameSet.In kn (erase_constant_decl X cb normalization_in oncb).2 -> In kn (map fst (declarations Σ)). +Proof. + intros Σ wfΣ kn. + unfold erase_constant_decl. + set (ec := erase_constant_body _ _ _ _). + destruct cb. destruct cst_body0. + 2:{ subst ec; cbn. KernameSetDecide.fsetdec. } + intros hc. + epose proof (erase_constant_body_correct' (@eq_refl _ (EAst.cst_body ec.1))). + subst ec. + set (ec := erase_constant_body _ _ _ _) in *. + set (X' := abstract_make_wf_env_ext _ _ _) in *. + pose proof (abstract_env_ext_exists X') as [[Σ' H']]. + pose proof (abstract_env_ext_wf _ H') as [wfΣ']. + specialize (H _ H') as [[t0 [T []]]]. + rewrite -e in hc. + destruct p as [hty her]. + epose proof (term_global_deps_spec wfΣ' hty her _ hc). + epose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ wfΣ H'). cbn in H0. subst Σ'. + red in H. destruct H as [[decl hl]]. + now eapply lookup_env_In_map_fst in hl. +Qed. + +Lemma in_erase_global_decls_acc X_type X deps decls kn normalization_in prf : + KernameSet.In kn (erase_global_decls X_type deps X decls normalization_in prf).2 -> + KernameSet.In kn deps \/ In kn (map fst decls). +Proof. + induction decls in X, prf, deps, normalization_in |- *. + cbn; auto. destruct a as [kn' []]. + + cbn. set (ec := erase_constant_decl _ _ _ _). + set (eg := erase_global_decls _ _ _ _ _ _). + set (eg' := erase_global_decls _ _ _ _ _ _). + case_eq (KernameSet.mem kn' deps). + * cbn; intros. specialize (IHdecls _ _ _ _ H0). + destruct IHdecls as [Hu|Hm]; auto. + eapply KernameSet.union_spec in Hu. destruct Hu; auto. + unfold ec in H1; cbn in H1. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + eapply erase_constant_decl_deps in H1; tea. + pose proof (abstract_env_exists X) as [[ΣX HX]]. + unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ HX wfpop). + intros. rewrite prf => //. now eexists. + now destruct H2 as [<- _]. + * intros hdeps hin. + eapply IHdecls in hin. intuition auto. + + cbn; set (eg := erase_global_decls _ _ _ _ _ _); + set (eg' := erase_global_decls _ _ _ _ _ _). + case_eq (KernameSet.mem kn' deps). + * cbn; intros. specialize (IHdecls _ _ _ _ H0). + destruct IHdecls as [Hu|Hm]; auto. + * intros hdeps hin. + eapply IHdecls in hin. intuition auto. +Qed. + +Lemma wf_fresh_globals {cf : checker_flags} (Σ : global_env) : wf Σ -> EnvMap.EnvMap.fresh_globals Σ.(declarations). +Proof. + destruct Σ as [univs Σ]; cbn. + move=> [] onu; cbn. induction 1; try destruct o; constructor; auto; constructor; eauto. +Qed. + +Lemma lookup_env_erase_decl X_type X deps decls normalization_in prf kn decl : + forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = Some decl -> + let er := erase_global_decls X_type deps X decls normalization_in prf in + KernameSet.In kn er.2 -> + match decl with + | ConstantDecl cb => + exists (X' : X_type.π1) nin pf, + (EGlobalEnv.lookup_env er.1 kn = + Some (EAst.ConstantDecl (fst (erase_constant_decl X' cb nin pf)))) /\ + (forall Σ Σ', Σ ∼ X -> Σ' ∼ X' -> ∥ strictly_extends_decls Σ' Σ ∥) + | InductiveDecl mib => + EGlobalEnv.lookup_env er.1 kn = + Some (EAst.InductiveDecl (erase_mutual_inductive_body mib)) + end. +Proof. + intros Σ wfΣ. pose proof (prf _ wfΣ). + unfold lookup_env. rewrite H. clear H. + rewrite /lookup_env. + induction decls in X, Σ , wfΣ ,prf, deps, decl, normalization_in |- *. + - move=> /= //. + - destruct a as [kn' d']. + cbn -[erase_global_decls]. + case: (eqb_spec kn kn'); intros e'; subst. + intros [= ->]. + destruct decl. + + cbn; set (ec := erase_constant_decl _ _ _ _); + set (eg := erase_global_decls _ _ _ _ _ _); + set (eg' := erase_global_decls _ _ _ _ _ _). + case_eq (KernameSet.mem kn' deps). + move/KernameSet.mem_spec => hin'; cbn. + ++ intros hin. cbn; rewrite eqb_refl; + do 3 eexists. split; [reflexivity|]. + intros ? ? hx hpop. + unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ hx hpop). intros. rewrite prf => //. now eexists. + destruct H as [? []]. unfold strictly_extends_decls. rewrite H. sq; split => //. clear ec eg eg' hin. specialize (prf _ hx). rewrite prf. + now eexists [_]. + ++ intros kn'deps kn'eg'. + eapply in_erase_global_decls_acc in kn'eg'. + destruct kn'eg'. eapply KernameSet.mem_spec in H. congruence. + elimtype False. clear -prf wfΣ H. + specialize (prf _ wfΣ). + pose proof (abstract_env_wf _ wfΣ) as []. eapply wf_fresh_globals in X0. + rewrite prf in X0. depelim X0. + apply EnvMap.EnvMap.fresh_global_iff_not_In in H0. contradiction. + + cbn; + set (eg := erase_global_decls _ _ _ _ _ _); + set (eg' := erase_global_decls _ _ _ _ _ _). + case_eq (KernameSet.mem kn' deps). + move/KernameSet.mem_spec => hin'; cbn. + ++ intros hin. cbn; rewrite eqb_refl; + do 3 eexists. + ++ intros kn'deps kn'eg'. + eapply in_erase_global_decls_acc in kn'eg'. + destruct kn'eg'. eapply KernameSet.mem_spec in H. congruence. + elimtype False. clear -prf wfΣ H. + specialize (prf _ wfΣ). + pose proof (abstract_env_wf _ wfΣ) as []. eapply wf_fresh_globals in X0. + rewrite prf in X0. depelim X0. + apply EnvMap.EnvMap.fresh_global_iff_not_In in H0. contradiction. + + intros hl hin. + epose proof (lookup_env_erase_global_diff X_type X deps decls kn kn' d' _ _ e'). + destruct H as [deps' [nin' [prf' [sub [hl' hin']]]]]. + eapply hin' in hin. + erewrite hl'. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + specialize (IHdecls Xpop deps'). specialize (IHdecls nin' prf' decl _ wfpop hl hin). + destruct decl; intuition auto. + destruct IHdecls as [X' [nin [pf []]]]. exists X', nin, pf. split => //. + intros. + unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ H1 wfpop). intros. rewrite prf => //. now eexists. + specialize (H0 _ _ wfpop H2). destruct H0; sq. move: X0; rewrite /strictly_extends_decls. destruct H3 as [? []]. + rewrite H4 H0 H3. rewrite -(prf' _ wfpop). pose proof (prf _ H1). rewrite -H0 in H5. rewrite H5. + intros []; split => //. destruct s as [Σ'' eq]. eexists ((kn', d') :: Σ''). now rewrite eq. +Qed. + +(* +Lemma lookup_env_erase_decl X_type X deps decls normalization_in prf kn decl : + forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = Some decl -> + if KernameSet.mem kn deps then + match decl with + | ConstantDecl cb => + exists (X' : X_type.π1) nin pf, + EGlobalEnv.lookup_env (erase_global_decls X_type deps X decls normalization_in prf) kn = + Some (EAst.ConstantDecl (fst (erase_constant_decl X' cb nin pf))) + | InductiveDecl mib => + EGlobalEnv.lookup_env (erase_global_decls X_type deps X decls normalization_in prf) kn = + Some (EAst.InductiveDecl (erase_mutual_inductive_body mib)) + end + else + EGlobalEnv.lookup_env (erase_global_decls X_type deps X decls normalization_in prf) kn = None. +Proof. + intros Σ wfΣ. pose proof (prf _ wfΣ). + unfold lookup_env. rewrite H. clear H. + rewrite /lookup_env. + induction decls in X, Σ , wfΣ ,prf, deps, decl, normalization_in |- *. + - move=> /= //. + - destruct a as [kn' d']. + cbn -[erase_global_decls]. + case: (eqb_spec kn kn'); intros e'; subst. + intros [= ->]. + case_eq (KernameSet.mem kn' deps); move => hin'; cbn; rewrite hin'; cbn. + + destruct decl; cbn; rewrite eqb_refl //. + do 3 eexists. reflexivity. + + specialize (IHdecls (abstract_pop_decls X) deps). + destruct decl. + * match goal with + |- context [ erase_global_decls _ _ _ _ ?nin ?prf ] => specialize (IHdecls nin prf) + end. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + case_eq (lookup_global decls kn'). + intros g hg. + specialize (IHdecls g _ wfpop hg). now rewrite hin' in IHdecls. + epose proof (lookup_env_erase_decl_None X_type (abstract_pop_decls X) deps decls _ _ _ _ wfpop). + intros hl. eapply H. unfold lookup_env. + unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wfΣ wfpop). + intros. pose proof (prf _ H0). now eexists. + destruct H0. now rewrite H0. + * rewrite hin' in IHdecls. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + case_eq (lookup_global decls kn'). + intros g hg. + eapply IHdecls; tea. + epose proof (lookup_env_erase_decl_None X_type (abstract_pop_decls X) deps decls _ _ _ _ wfpop). + intros hl; eapply H. + unfold lookup_env. + unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wfΣ wfpop). + intros. pose proof (prf _ H0). now eexists. + destruct H0. now rewrite H0. + + intros hl. + epose proof (lookup_env_erase_global_diff X_type X deps decls kn kn' d' _ _ e'). + destruct H as [deps' [nin' [prf' [sub H]]]]. + erewrite H. + case_eq (KernameSet.mem kn deps) => //. + * move/KernameSet.mem_spec => hin. + assert (kndeps' : KernameSet.mem kn deps'). + { now apply KernameSet.mem_spec; eapply sub. } + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + specialize (IHdecls Xpop deps'). + epose proof (IHdecls _ _ decl _ wfpop hl) ; tea. rewrite kndeps' in H0. eapply H0. + * intros hin. rewrite hin in IHdecls. + + + ** intros g hl. eapply IHdecls. + destruct KernameSet.mem. cbn. + rewrite (negbTE (proj2 (neqb _ _) e')). + eapply IHdecls => //; eauto. eapply KernameSet.union_spec. left => //. + eapply IHdecls => //; eauto. + simpl. + + destruct KernameSet.mem. cbn. + rewrite (negbTE (proj2 (neqb _ _) e')). + eapply IHdecls => //; eauto. + eapply IHdecls => //; eauto. +Qed.*) + +Lemma in_deps_in_erase_global_decls X_type X deps decls normalization_in prf kn decl : + forall Σ, Σ ∼ X -> + lookup_env Σ kn = Some decl -> + KernameSet.In kn deps -> + KernameSet.In kn (erase_global_decls X_type deps X decls normalization_in prf).2. +Proof. + induction decls in X, deps, normalization_in, prf |- *. + - intros. rewrite /lookup_env prf in H0 => //. + - intros. cbn. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ H wfpop). intros. rewrite prf => //. now eexists. + destruct H2 as [Hd [Hu Hr]]. + destruct a. destruct g. + + set (eg := erase_global_decls _ _ _ _ _ _); + set (eg' := erase_global_decls _ _ _ _ _ _). + case_eq (KernameSet.mem k deps). + * cbn. intros hm. + destruct (eqb_spec k kn). + ++ subst k. admit. + ++ eapply IHdecls. tea. + rewrite /lookup_env Hd. + rewrite /lookup_env (prf _ H) in H0. cbn in H0. + admit. + KernameSetDecide.fsetdec. + * intros hnin. eapply IHdecls; tea. + rewrite /lookup_env Hd. + rewrite /lookup_env (prf _ H) in H0. cbn in H0. admit. + + set (eg := erase_global_decls _ _ _ _ _ _); + set (eg' := erase_global_decls _ _ _ _ _ _). + case_eq (KernameSet.mem k deps). + * cbn. intros hm. + destruct (eqb_spec k kn). + ++ subst k. admit. + ++ eapply IHdecls. tea. + rewrite /lookup_env Hd. + rewrite /lookup_env (prf _ H) in H0. cbn in H0. + admit. + KernameSetDecide.fsetdec. + * intros hnin. eapply IHdecls; tea. + rewrite /lookup_env Hd. + rewrite /lookup_env (prf _ H) in H0. cbn in H0. admit. + +Admitted. Lemma erase_global_declared_constructor X_type X ind c mind idecl cdecl deps decls normalization_in prf: forall Σ : global_env, abstract_env_rel X Σ -> declared_constructor Σ (ind, c) mind idecl cdecl -> KernameSet.In ind.(inductive_mind) deps -> - EGlobalEnv.declared_constructor (erase_global_decls X_type deps X decls normalization_in prf) (ind, c) + EGlobalEnv.declared_constructor (erase_global_decls X_type deps X decls normalization_in prf).1 (ind, c) (erase_mutual_inductive_body mind) (erase_one_inductive_body idecl) (EAst.mkConstructor cdecl.(cstr_name) cdecl.(cstr_arity)). Proof. intros Σ wfΣ [[]] Hin. pose (abstract_env_wf _ wfΣ). sq. cbn in *. split. split. - unshelve eapply declared_minductive_to_gen in H; eauto. - red in H |- *. eapply lookup_env_erase; eauto. + red in H |- *. eapply (lookup_env_erase_decl _ _ _ _ _ _ _ (InductiveDecl mind)); eauto. + eapply in_deps_in_erase_global_decls; tea. - cbn. now eapply map_nth_error. - cbn. erewrite map_nth_error; eauto. @@ -1855,7 +2235,7 @@ Import EGlobalEnv. Lemma erase_global_decls_fresh X_type kn deps X decls normalization_in heq : let Σ' := erase_global_decls X_type deps X decls normalization_in heq in PCUICAst.fresh_global kn decls -> - fresh_global kn Σ'. + fresh_global kn Σ'.1. Proof. cbn. revert deps X normalization_in heq. @@ -1942,7 +2322,7 @@ Lemma erases_deps_erase (cf := config.extraction_checker_flags) let et := erase X_type X' Γ t wt in let deps := EAstUtils.term_global_deps et in forall Σ, (abstract_env_rel X Σ) -> - erases_deps Σ (erase_global_decls X_type deps X decls normalization_in prf) et. + erases_deps Σ (erase_global_decls X_type deps X decls normalization_in prf).1 et. Proof. intros et deps Σ wf. pose proof (abstract_env_ext_exists X') as [[Σ' wfΣ']]. @@ -1972,7 +2352,7 @@ Lemma erases_deps_erase_weaken (cf := config.extraction_checker_flags) let et := erase X_type X' Γ t wt in let tdeps := EAstUtils.term_global_deps et in forall Σ, (abstract_env_rel X Σ) -> - erases_deps Σ (erase_global_decls X_type (KernameSet.union deps tdeps) X decls normalization_in prf) et. + erases_deps Σ (erase_global_decls X_type (KernameSet.union deps tdeps) X decls normalization_in prf).1 et. Proof. intros et tdeps Σ wf. pose proof (abstract_env_ext_exists X') as [[Σ' wfΣ']]. @@ -2010,13 +2390,14 @@ Qed. Lemma erase_global_closed X_type X deps decls {normalization_in} prf : let X' := erase_global_decls X_type deps X decls normalization_in prf in - EGlobalEnv.closed_env X'. + EGlobalEnv.closed_env X'.1. Proof. revert X normalization_in prf deps. induction decls; [cbn; auto|]. intros X normalization_in prf deps. destruct a as [kn d]. destruct d as []; simpl; destruct KernameSet.mem; + unfold erase_constant_decl; set (Xpop := abstract_pop_decls X); try (set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _); assert (normalization_in_Xmake : forall Σ : global_env_ext, wf_ext Σ -> Σ ∼_ext Xmake -> NormalizationIn Σ) @@ -2025,6 +2406,7 @@ Proof. rewrite [forallb _ _](IHdecls) // andb_true_r. rewrite /test_snd /EGlobalEnv.closed_decl /=. destruct c as [ty [] univs]; cbn. + unfold erase_constant_decl. set (obl := erase_constant_body_obligation_1 _ _ _ _ _). unfold erase_constant_body. cbn. clearbody obl. cbn in obl. 2:auto. @@ -2037,7 +2419,6 @@ Proof. now eapply PCUICClosedTyp.subject_closed in X0. + eapply IHdecls => //. + cbn [EGlobalEnv.closed_env forallb]. - rewrite {1}/test_snd {1}/EGlobalEnv.closed_decl /=. eapply IHdecls => //. + eapply IHdecls => //. Unshelve. eauto. @@ -2122,7 +2503,7 @@ Lemma erase_wellformed (efl := all_env_flags) {X_type X} decls {normalization_in (X' := abstract_make_wf_env_ext X univs wfΣ) {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} (t' := erase X_type X' Γ t wt) : -wellformed (erase_global_decls X_type (term_global_deps t') X decls normalization_in prf) #|Γ| t'. +wellformed (erase_global_decls X_type (term_global_deps t') X decls normalization_in prf).1 #|Γ| t'. Proof. cbn. epose proof (@erases_deps_erase X_type X univs wfΣ decls normalization_in prf _ Γ t wt). @@ -2144,7 +2525,7 @@ Lemma erase_wellformed_weaken (efl := all_env_flags) {X_type X} decls {normaliza {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} deps: let t' := erase X_type X' Γ t wt in - wellformed (erase_global_decls X_type (KernameSet.union deps (term_global_deps t')) X decls normalization_in prf) #|Γ| t'. + wellformed (erase_global_decls X_type (KernameSet.union deps (term_global_deps t')) X decls normalization_in prf).1 #|Γ| t'. Proof. set (t' := erase _ _ _ _ _). cbn. epose proof (@erases_deps_erase_weaken _ X univs wfΣ decls _ prf _ Γ t wt deps). @@ -2170,7 +2551,7 @@ Lemma erase_constant_body_correct'' {X_type X} {cb} {decls normalization_in prf} forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> ∥ ∑ t T, (Σ' ;;; [] |- t : T) * (Σ' ;;; [] |- t ⇝ℇ body) * (term_global_deps body = snd (erase_constant_body X_type X' cb onc)) * - wellformed (efl:=all_env_flags) (erase_global_decls X_type (KernameSet.union deps (term_global_deps body)) X decls normalization_in prf) 0 body ∥. + wellformed (efl:=all_env_flags) (erase_global_decls X_type (KernameSet.union deps (term_global_deps body)) X decls normalization_in prf).1 0 body ∥. Proof. intros ? Σ' wfΣ'. pose proof (abstract_env_exists X) as [[Σ wf]]. destruct cb as [name [bod|] udecl]; simpl. @@ -2197,7 +2578,7 @@ Lemma erase_global_cst_decl_wf_glob X_type X deps decls normalization_in heq : forall normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ, let ecb := erase_constant_body X_type X' cb hcb in let Σ' := erase_global_decls X_type (KernameSet.union deps ecb.2) X decls normalization_in heq in - (@wf_global_decl all_env_flags Σ' (EAst.ConstantDecl ecb.1) : Prop). + (@wf_global_decl all_env_flags Σ'.1 (EAst.ConstantDecl ecb.1) : Prop). Proof. intros cb wfΣ hcb X' normalization_in' ecb Σ'. unfold wf_global_decl. cbn. @@ -2213,7 +2594,7 @@ Lemma erase_global_ind_decl_wf_glob {X_type X} {deps decls normalization_in kn m (forall Σ : global_env, abstract_env_rel X Σ -> on_inductive cumulSpec0 (lift_typing typing) (Σ, ind_universes m) kn m) -> let m' := erase_mutual_inductive_body m in let Σ' := erase_global_decls X_type deps X decls normalization_in heq in - @wf_global_decl all_env_flags Σ' (EAst.InductiveDecl m'). + @wf_global_decl all_env_flags Σ'.1 (EAst.InductiveDecl m'). Proof. set (m' := erase_mutual_inductive_body m). set (Σ' := erase_global_decls _ _ _ _ _ _). simpl. @@ -2228,7 +2609,7 @@ Qed. Lemma erase_global_decls_wf_glob X_type X deps decls normalization_in heq : let Σ' := erase_global_decls X_type deps X decls normalization_in heq in - @wf_glob all_env_flags Σ'. + @wf_glob all_env_flags Σ'.1. Proof. cbn. pose proof (abstract_env_exists X) as [[Σ wf]]. @@ -2261,7 +2642,7 @@ Qed. Lemma lookup_erase_global (cf := config.extraction_checker_flags) {X_type X} {deps deps' decls normalization_in prf} : KernameSet.Subset deps deps' -> - EExtends.global_subset (erase_global_decls X_type deps X decls normalization_in prf) (erase_global_decls X_type deps' X decls normalization_in prf). + EExtends.global_subset (erase_global_decls X_type deps X decls normalization_in prf).1 (erase_global_decls X_type deps' X decls normalization_in prf).1. Proof. revert deps deps' X normalization_in prf. induction decls. cbn => //. @@ -2322,8 +2703,8 @@ Qed. Lemma expanded_weakening_global X_type X deps deps' decls normalization_in prf Γ t : KernameSet.Subset deps deps' -> - expanded (erase_global_decls X_type deps X decls normalization_in prf) Γ t -> - expanded (erase_global_decls X_type deps' X decls normalization_in prf) Γ t. + expanded (erase_global_decls X_type deps X decls normalization_in prf).1 Γ t -> + expanded (erase_global_decls X_type deps' X decls normalization_in prf).1 Γ t. Proof. intros hs. eapply expanded_ind; intros; try solve [econstructor; eauto]. @@ -2342,7 +2723,7 @@ Lemma expanded_erase (cf := config.extraction_checker_flags) forall normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ, let et := (erase X_type X' [] t wtp) in let deps := EAstUtils.term_global_deps et in - expanded (erase_global_decls X_type deps X decls normalization_in prf) [] et. + expanded (erase_global_decls X_type deps X decls normalization_in prf).1 [] et. Proof. intros Σ wf hexp X' normalization_in'. pose proof (abstract_env_wf _ wf) as [wf']. @@ -2358,7 +2739,7 @@ Lemma expanded_erase_global (cf := config.extraction_checker_flags) deps {X_type X decls normalization_in prf} : forall Σ : global_env, abstract_env_rel X Σ -> PCUICEtaExpand.expanded_global_env Σ -> - expanded_global_env (erase_global_decls X_type deps X decls normalization_in prf). + expanded_global_env (erase_global_decls X_type deps X decls normalization_in prf).1. Proof. intros Σ wf etaΣ. pose proof (prf _ wf). destruct Σ as [univ decls']. @@ -2373,6 +2754,7 @@ Proof. destruct (KernameSet.mem kn deps) eqn:eqkn; simpl; rewrite eqkn. constructor; [eapply IHetaΣ; auto|]. red. cbn. red. cbn in *. + unfold erase_constant_decl. set (deps' := KernameSet.union _ _). hidebody deps'. set (wfext := abstract_make_wf_env_ext _ _ _). hidebody wfext. destruct H. @@ -2412,11 +2794,11 @@ Lemma erase_eval_to_box (wfl := Ee.default_wcbv_flags) forall Σext : global_env_ext, abstract_env_ext_rel Xext Σext -> (forall Σ : global_env, abstract_env_rel X Σ -> PCUICWcbvEval.eval Σ t v) -> - @Ee.eval Ee.default_wcbv_flags Σ' t' E.tBox -> ∥ isErasable Σext [] t ∥. + @Ee.eval Ee.default_wcbv_flags Σ'.1 t' E.tBox -> ∥ isErasable Σext [] t ∥. Proof. intros wt. intros. - destruct (erase_correct X_type X univs wfext _ _ Σ' _ _ decls prf wt H H0 H1 X0 _ H2) as [ev [eg [eg']]]. + destruct (erase_correct X_type X univs wfext _ _ Σ'.1 _ _ decls prf wt H H0 H1 X0 _ H2) as [ev [eg [eg']]]. pose proof (Ee.eval_deterministic H3 eg'). subst. pose proof (abstract_env_exists X) as [[? wf]]. destruct (wfext _ wf). destruct (wt _ H2) as [T wt']. @@ -2438,7 +2820,7 @@ Lemma erase_eval_to_box_eager (wfl := Ee.default_wcbv_flags) forall Σext : global_env_ext, abstract_env_ext_rel Xext Σext -> (forall Σ : global_env, abstract_env_rel X Σ -> PCUICWcbvEval.eval Σ t v) -> - @Ee.eval Ee.default_wcbv_flags Σ' t' E.tBox -> t' = E.tBox. + @Ee.eval Ee.default_wcbv_flags Σ'.1 t' E.tBox -> t' = E.tBox. Proof. intros wt. intros. @@ -2544,7 +2926,7 @@ forall Σ, abstract_env_ext_rel Xext Σ -> erase_global_decls X_type deps X decls normalization_in prf = Σ' -> red Σ [] t v -> ¬ { t' & Σ ;;; [] |- v ⇝ t'} -> - forall wt', ∥ Σ' ⊢ t' ⇓ erase X_type Xext [] v wt' ∥. + forall wt', ∥ Σ'.1 ⊢ t' ⇓ erase X_type Xext [] v wt' ∥. Proof. intros wt Σ Hrel Hax Hty Hfo <- Hsub <- Hred Hirred wt'. destruct (firstorder_lookup_inv Hfo) as [mind Hdecl]. @@ -2575,7 +2957,7 @@ forall Σ, abstract_env_ext_rel Xext Σ -> KernameSet.subset (term_global_deps t') deps -> red Σ [] t v -> ¬ { v' & Σ ;;; [] |- v ⇝ v'} -> - exists wt', ∥ Σ' ⊢ t' ⇓ erase X_type Xext [] v wt' ∥. + exists wt', ∥ Σ'.1 ⊢ t' ⇓ erase X_type Xext [] v wt' ∥. Proof. intros wt Σ Hrel Hax Hty Hdecl Hfo Hsub Hred Hirred. unshelve eexists. @@ -2617,7 +2999,7 @@ forall Σ (wfΣ : abstract_env_ext_rel Xext Σ) @firstorder_ind Σ (firstorder_env Σ) i -> red Σ [] t v -> ¬ { v' & Σ ;;; [] |- v ⇝ v'} -> - exists wt', ∥ Σ' ⊢ t' ⇓ erase X_type Xext [] v wt' ∥. + exists wt', ∥ Σ'.1 ⊢ t' ⇓ erase X_type Xext [] v wt' ∥. Proof. intros. eapply erase_correct_strong; eauto. @@ -2691,7 +3073,7 @@ Program Fixpoint erase_global_decls_fast (deps : KernameSet.t) {normalization_in : NormalizationIn_erase_global_decls_fast X decls} (prop : forall Σ : global_env, abstract_env_rel X Σ -> ∥ decls_prefix decls Σ ∥) : E.global_declarations := match decls with - | [] => [] + | [] => ([],deps) | (kn, ConstantDecl cb) :: decls => if KernameSet.mem kn deps then let Xext' := abstract_make_wf_env_ext X (cst_universes cb) _ in @@ -2699,7 +3081,7 @@ Program Fixpoint erase_global_decls_fast (deps : KernameSet.t) let cb' := @erase_constant_body X_type Xext' normalization_in' cb _ in let deps := KernameSet.union deps (snd cb') in let Σ' := erase_global_decls_fast deps X_type X decls _ in - ((kn, E.ConstantDecl (fst cb')) :: Σ') + (((kn, E.ConstantDecl (fst cb')) :: Σ'.1), Σ'.2) else erase_global_decls_fast deps X_type X decls _ @@ -2707,7 +3089,7 @@ Program Fixpoint erase_global_decls_fast (deps : KernameSet.t) if KernameSet.mem kn deps then let mib' := erase_mutual_inductive_body mib in let Σ' := erase_global_decls_fast deps X_type X decls _ in - ((kn, E.InductiveDecl mib') :: Σ') + (((kn, E.InductiveDecl mib') :: Σ'.1), Σ'.2) else erase_global_decls_fast deps X_type X decls _ end. Next Obligation. @@ -2832,7 +3214,7 @@ Definition same_principal_type {cf} {X_type' : abstract_env_impl} {X' : X_type'.π2.π1} {Γ : context} {t} (p : PCUICSafeRetyping.principal_type X_type X Γ t) (p' : PCUICSafeRetyping.principal_type X_type' X' Γ t) := p.π1 = p'.π1. - +(* Definition Hlookup {cf} (X_type : abstract_env_impl) (X : X_type.π2.π1) (X_type' : abstract_env_impl) (X' : X_type'.π2.π1) := forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> @@ -2840,7 +3222,38 @@ Definition Hlookup {cf} (X_type : abstract_env_impl) (X : X_type.π2.π1) forall kn decl decl', lookup_env Σ kn = Some decl -> lookup_env Σ' kn = Some decl' -> - abstract_env_lookup X kn = abstract_env_lookup X' kn. + abstract_env_lookup X kn = abstract_env_lookup X' kn. *) + + +(* When two environments agree on the intersection of their declarations *) +Definition equiv_env_inter Σ Σ' := + (forall kn decl decl', lookup_env Σ kn = Some decl -> lookup_env Σ' kn = Some decl' -> decl = decl') /\ + (forall tag : Primitive.prim_tag, primitive_constant Σ tag = primitive_constant Σ' tag). + +Lemma equiv_env_inter_sym Σ Σ' : + equiv_env_inter Σ Σ' <-> equiv_env_inter Σ' Σ. +Proof. + unfold equiv_env_inter. + intuition auto; symmetry; eauto. +Qed. + +(* When two environments agree on the intersection of their declarations *) +Lemma equiv_env_inter_hlookup {cf:checker_flags} (X_type : abstract_env_impl) (X : X_type.π2.π1) + (X_type' : abstract_env_impl) (X' : X_type'.π2.π1) : + (forall Σ Σ', Σ ∼_ext X -> Σ' ∼_ext X' -> equiv_env_inter Σ Σ') -> Hlookup X_type X X_type' X'. +Proof. + intros hl Σ HX Σ' HX'. + specialize (hl _ _ HX HX') as [hl hr]. + split. + - intros kn decl decl'. + specialize (hl kn decl decl'). + rewrite <-(abstract_env_lookup_correct' _ _ HX). + rewrite <-(abstract_env_lookup_correct' _ _ HX'). + intros Hl Hl'. + rewrite Hl Hl'. f_equal. eauto. + - intros tag; rewrite (abstract_primitive_constant_correct _ _ _ HX) (abstract_primitive_constant_correct _ _ _ HX'). + apply hr. +Qed. (*Lemma erase_global_deps_suffix {deps} {Σ Σ' : wf_env} {decls hprefix hprefix'} : wf Σ -> wf Σ' -> @@ -2883,14 +3296,42 @@ Proof using Type. { clear -equ hprefix'. move: hprefix'. intros [? ?]. split => //. eexists (x ++ [(kn, ConstantDecl c)]). cbn. now rewrite -app_assoc. } - rewrite (erase_constant_body_suffix (Σ := wfΣx) (Σ' := wfΣdw) (ondecl' := ond') wfΣ extd') //. - rewrite (erase_constant_body_suffix (Σ := wfΣ'x) (Σ' := wfΣdw) (ondecl' := ond') wfΣ' extd'') //. } + rewrite (erase_constant_suffix (Σ := wfΣx) (Σ' := wfΣdw) (ondecl' := ond') wfΣ extd') //. + rewrite (erase_constant_body_irrel (Σ := wfΣ'x) (Σ' := wfΣdw) (ondecl' := ond') wfΣ' extd'') //. } destruct KernameSet.mem => //. f_equal. eapply IHdecls. destruct KernameSet.mem => //. f_equal. eapply IHdecls. Qed.*) +Lemma strictly_extends_decls_extends_global_env Σ Σ' : + wf Σ' -> + strictly_extends_decls Σ Σ' -> + extends_global_env Σ Σ'. +Proof. + intros wfΣ' [uni [Σ'' de] retr]. red. + cbn. unfold lookup_env. rewrite de /primitive_constant retr. + split => //. intros kn decl. + intros hl. + apply lookup_global_Some_iff_In_NoDup. + apply EnvMap.EnvMap.fresh_globals_iff_NoDup. + rewrite -de. now apply wf_fresh_globals. + apply lookup_global_Some_iff_In_NoDup in hl. + apply in_app_iff. now right. + apply EnvMap.EnvMap.fresh_globals_iff_NoDup. + apply wf_fresh_globals in wfΣ'. rewrite de in wfΣ'. + clear -wfΣ'. induction Σ''; auto. apply IHΣ''. + now depelim wfΣ'. +Qed. + +Lemma extends_global_env_equiv_env Σ Σ' : + extends_global_env Σ Σ' -> equiv_env_inter Σ Σ'. +Proof. + intros ext. split. intros kn decl decl' hl hl'. + destruct ext. apply H in hl. congruence. + now destruct ext. +Qed. + Lemma erase_global_deps_fast_spec_gen {deps} {X_type X X'} {decls normalization_in normalization_in' hprefix hprefix'} : (forall Σ Σ', abstract_env_rel X Σ -> abstract_env_rel X' Σ' -> universes Σ = universes Σ' /\ retroknowledge Σ = retroknowledge Σ') -> @@ -2907,7 +3348,7 @@ Proof using Type. { now eexists. } destruct a as [kn []]. - - cbn. + - cbn. unfold erase_constant_decl. set (obl := (erase_global_decls_fast_obligation_1 X_type X ((kn, ConstantDecl c) :: decls) _ hprefix kn c decls eq_refl)). set (eb := erase_constant_body _ _ _ _). @@ -2915,16 +3356,20 @@ Proof using Type. assert (eb = eb') as <-. { subst eb eb'. destruct (hprefix _ wfΣ) as [[Σ'' eq]]. - eapply erase_constant_body_suffix; cbn => //. - intros. + eapply erase_constant_body_irrel; cbn => //. + eapply equiv_env_inter_hlookup. + intros ? ? H2 H3. epose proof (abstract_make_wf_env_ext_correct X (cst_universes c) _ _ _ wfΣ H2). epose proof (abstract_make_wf_env_ext_correct (abstract_pop_decls X') (cst_universes c) _ _ _ wfpop H3). - subst. split => //. - sq. apply strictly_extends_decls_extends_decls. red. cbn. - rewrite eq. rewrite <- H0, <- H1. split. symmetry. apply equ; eauto. - eexists (Σ'' ++ [(kn, ConstantDecl c)]). subst. now rewrite -app_assoc. subst. - symmetry. now apply equ. - } + subst. eapply equiv_env_inter_sym. + eapply extends_global_env_equiv_env. + pose proof (abstract_env_wf _ wfΣ) as []. + apply strictly_extends_decls_extends_global_env. + apply X0. red. + rewrite eq. rewrite <- H0, <- H1. split. cbn. symmetry; apply equ; eauto. + cbn. + eexists (Σ'' ++ [(kn, ConstantDecl c)]). cbn. subst. now rewrite -app_assoc. subst. + symmetry. now apply equ. } destruct KernameSet.mem => //; f_equal; eapply IHdecls. intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. diff --git a/erasure/theories/Typed/ErasureCorrectness.v b/erasure/theories/Typed/ErasureCorrectness.v index 2d3079eee..70272cbf9 100644 --- a/erasure/theories/Typed/ErasureCorrectness.v +++ b/erasure/theories/Typed/ErasureCorrectness.v @@ -167,11 +167,12 @@ Proof. destruct decl; [left|right]. all: unfold declared_constant, declared_minductive, P.declared_constant, P.declared_minductive; cbn. + unfold Erasure.erase_constant_decl. all: unfold eq_kername in *. all: try rewrite eq_kername_refl. + eexists; split; [left;reflexivity|]. - unfold erase_constant_decl. - destruct flag_of_type; cbn in *. + unfold erase_constant_decl. + destruct flag_of_type; cbn in *. destruct conv_ar; cbn in *. destruct wfΣprev as [wfΣprev]. * destruct c eqn:Hc; cbn in *. diff --git a/template-coq/theories/TemplateProgram.v b/template-coq/theories/TemplateProgram.v index e873a3d6e..2895858f6 100644 --- a/template-coq/theories/TemplateProgram.v +++ b/template-coq/theories/TemplateProgram.v @@ -47,7 +47,7 @@ Definition make_template_program_env {cf : checker_flags} (p : template_program) (** We kludge the normalization assumptions by parameterizing over a continuation of "what will be done to the program later" as well as what properties we'll need of it *) Program Definition build_template_program_env {cf : checker_flags} K : - Transform.t template_program template_program_env Ast.term Ast.term eval_template_program eval_template_program_env := + Transform.t template_program template_program_env Ast.term Ast.term eval_template_program eval_template_program_env := {| name := "rebuilding environment lookup table"; pre p := ∥ wt_template_program p ∥ /\ forall pf, K (GlobalEnvMap.make p.1 (wt_template_program_fresh p pf)) : Prop ; transform p pre := make_template_program_env p (proj1 pre); From 8710ece55c005349ab135dccec7c99ce306cf11f Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 28 Aug 2023 10:02:22 +0200 Subject: [PATCH 39/61] No need for ext proof for erasure --- erasure-plugin/theories/ETransform.v | 73 ++++--- erasure/theories/ErasureFunction.v | 291 +++++++++++++-------------- 2 files changed, 188 insertions(+), 176 deletions(-) diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index ba2e8c17f..3dcc50ac2 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -148,7 +148,7 @@ Program Definition erase_pcuic_program {guard : abstract_guard_impl} (p : pcuic_ (fun Σ wfΣ => let '(sq (T; ty)) := wt in PCUICTyping.iswelltyped ty) in let Σ' := ErasureFunction.erase_global_fast (normalization_in:=_) optimized_abstract_env_impl (EAstUtils.term_global_deps t) wfe (p.1.(PCUICAst.PCUICEnvironment.declarations)) _ in - (EEnvMap.GlobalContextMap.make Σ' _, t). + (EEnvMap.GlobalContextMap.make Σ'.1 _, t). Next Obligation. unshelve edestruct erase_pcuic_program_normalization_helper; cbn in *; eauto. Qed. Next Obligation. @@ -280,28 +280,6 @@ Definition extends_pcuic_program (p p' : pcuic_program) := extends_global_env p.1 p'.1 /\ p.2 = p'.2. Import ErasureFunction. Import PCUICAst. -Lemma fresh_global_app kn Σ Σ' : fresh_global kn (Σ ++ Σ') -> - fresh_global kn Σ /\ fresh_global kn Σ'. -Proof. - induction Σ; cbn; intuition auto. - - constructor. - - depelim H. constructor => //. - now eapply Forall_app in H0. - - depelim H. - now eapply Forall_app in H0. -Qed. - -Lemma lookup_global_app_wf Σ Σ' kn : EnvMap.EnvMap.fresh_globals (Σ ++ Σ') -> - forall decl, lookup_global Σ' kn = Some decl -> lookup_global (Σ ++ Σ') kn = Some decl. -Proof. - induction Σ in kn |- *; cbn; auto. - intros fr; depelim fr. - intros decl hd; cbn. - destruct (eqb_spec kn kn0). - eapply PCUICWeakeningEnv.lookup_global_Some_fresh in hd. - subst. now eapply fresh_global_app in H as [H H']. - now eapply IHΣ. -Qed. Lemma strictly_extends_lookups {cf:checker_flags} (X X' : wf_env) (Σ Σ' : global_env) : (forall (kn : kername) (decl decl' : global_decl), lookup_env Σ' kn = Some decl -> lookup_env Σ kn = Some decl' -> decl = decl') -> @@ -327,6 +305,49 @@ Proof. - rewrite /primitive_constant. now rewrite e0 e2 hr. Qed. + +Lemma lookup_env_In_map_fst Σ kn decl : EGlobalEnv.lookup_env Σ kn = Some decl -> In kn (map fst Σ). +Proof. + induction Σ; cbn => //. + case: eqb_spec. + + intros -> [= <-]. now left. + + intros _ hl. eauto. +Qed. + +Lemma in_erase_global_decls_acc {X_type : abstract_env_impl} (X : X_type.π1) deps decls kn normalization_in prf : + In kn (map fst (erase_global_decls deps X decls (normalization_in:=normalization_in) prf).1) -> + KernameSet.In kn (erase_global_decls deps X decls (normalization_in:=normalization_in) prf).2. +Proof. + induction decls in X, prf, deps, normalization_in |- *. + cbn; auto. destruct a as [kn' []]. + + cbn. set (ec := erase_constant_decl _ _ _ _). + set (eg := erase_global_decls _ _ _ _). + set (eg' := erase_global_decls _ _ _ _). + case_eq (KernameSet.mem kn' deps). + * cbn; intros. destruct H0. subst; auto. + eapply KernameSet.mem_spec in H. + +Admitted. + (*auto. + specialize (IHdecls _ _ _ _ H0). + unfold ec in H0; cbn in H0. + set (Xpop := abstract_pop_decls X) in *. + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + epose proof (abstract_env_exists X) as [[ΣX HX]]. + unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ HX wfpop). + intros. rewrite prf => //. now eexists. + destruct H1 as [<- ?]. auto. + * intros hdeps hin. + eapply IHdecls in hin. intuition auto. + + cbn; set (eg := erase_global_decls _ _ _ _); + set (eg' := erase_global_decls _ _ _ _). + case_eq (KernameSet.mem kn' deps). + * cbn; intros. intuition auto. now specialize (IHdecls _ _ _ _ H1). + * intros hdeps hin. + eapply IHdecls in hin. intuition auto. + Unshelve. all:tc. +Qed. +*) #[global] Instance erase_transform_extends {guard : abstract_guard_impl} : TransformExt.t (erase_transform (guard := guard)) extends_pcuic_program extends_eprogram_env. @@ -360,14 +381,14 @@ Proof. unshelve epose proof (abstract_env_exists X). 3:tc. tc. destruct H4 as [[Σ' ΣX]]. unshelve epose proof (build_wf_env_from_env_eq _ (g, Σ.2) _ _ ΣX). cbn in H4. - epose proof (hl := ErasureFunction.lookup_env_erase_decl (optimized_abstract_env_impl) _ (EAstUtils.term_global_deps fst) _ _ _ kn g0 _ ΣX). + epose proof (hl := ErasureFunction.lookup_env_erase_decl (optimized_abstract_env_impl) _ (EAstUtils.term_global_deps fst) (declarations g) _ _ kn g0 _ ΣX). unfold lookup_env in E, hl. rewrite H4 in hl. specialize (hl (H0 _ _ E)). - forward hl. admit. destruct g0. - + destruct hl as [X' [nin [pf [eq ext]]]]. + + destruct hl as [X' [nin [pf [eq ext]]]]; revgoals. unshelve erewrite ErasureFunction.erase_global_deps_fast_spec. shelve. { intros. red in H5. cbn in H5. subst. reflexivity. } erewrite eq. + 2:{ eapply lookup_env_In_map_fst in H2. eapply in_erase_global_decls_acc in H2. } set (Xs := build_wf_env_from_env Σ.1 _) in *. unshelve epose proof (abstract_env_exists Xs). 3:tc. tc. destruct H5 as [[Σs Hs]]. diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 10eceaf4b..5edab27a3 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -927,6 +927,71 @@ Next Obligation. clear H; specialize (oncb _ HX). destruct oncb as [[]]; sq; auto. Qed. +(* When two environments agree on the intersection of their declarations *) +Definition equiv_decls_inter Σ Σ' := + (forall kn decl decl', EGlobalEnv.lookup_env Σ kn = Some decl -> EGlobalEnv.lookup_env Σ' kn = Some decl' -> decl = decl'). + +Definition equiv_env_inter Σ Σ' := + (forall kn decl decl', lookup_env Σ kn = Some decl -> lookup_env Σ' kn = Some decl' -> decl = decl') /\ + (forall tag : Primitive.prim_tag, primitive_constant Σ tag = primitive_constant Σ' tag). + +Lemma equiv_env_inter_sym Σ Σ' : + equiv_env_inter Σ Σ' <-> equiv_env_inter Σ' Σ. +Proof. + unfold equiv_env_inter. + intuition auto; symmetry; eauto. +Qed. + +(* When two environments agree on the intersection of their declarations *) +Lemma equiv_env_inter_hlookup {cf:checker_flags} (X_type : abstract_env_impl) (X : X_type.π2.π1) + (X_type' : abstract_env_impl) (X' : X_type'.π2.π1) : + (forall Σ Σ', Σ ∼_ext X -> Σ' ∼_ext X' -> equiv_env_inter Σ Σ') -> Hlookup X_type X X_type' X'. +Proof. + intros hl Σ HX Σ' HX'. + specialize (hl _ _ HX HX') as [hl hr]. + split. + - intros kn decl decl'. + specialize (hl kn decl decl'). + rewrite <-(abstract_env_lookup_correct' _ _ HX). + rewrite <-(abstract_env_lookup_correct' _ _ HX'). + intros Hl Hl'. + rewrite Hl Hl'. f_equal. eauto. + - intros tag; rewrite (abstract_primitive_constant_correct _ _ _ HX) (abstract_primitive_constant_correct _ _ _ HX'). + apply hr. +Qed. + + +Definition Hlookup_env {cf:checker_flags} (X_type : abstract_env_impl) (X : X_type.π1) (X_type' : abstract_env_impl) + (X' : X_type'.π1) := + forall Σ : global_env, abstract_env_rel X Σ -> + forall Σ' : global_env, abstract_env_rel X' Σ' -> + equiv_env_inter Σ Σ'. + +Lemma erase_constant_decl_irrel {X_type : abstract_env_impl} (X : X_type.π1) (cb : constant_body) normalization_in oncb + (X' : X_type.π1) normalization_in' oncb' : + Hlookup_env X_type X X_type X' -> + erase_constant_decl X cb normalization_in oncb = erase_constant_decl X' cb normalization_in' oncb'. +Proof. + unfold erase_constant_decl. + intros hl. + apply erase_constant_body_irrel. + pose proof (abstract_env_exists X) as [[env wf]]. + pose proof (abstract_env_exists X') as [[env' wf']]. + red. intros. + specialize (hl _ wf _ wf') as [hl hp]. + pose proof (abstract_make_wf_env_ext_correct X (cst_universes cb) _ _ _ wf H). + pose proof (abstract_make_wf_env_ext_correct X' (cst_universes cb) _ _ _ wf' H0). + split => //. + { intros. + rewrite -(abstract_env_lookup_correct' _ _ H). + rewrite -(abstract_env_lookup_correct' _ _ H0). rewrite H3 H4. subst Σ Σ'. + now specialize (hl _ _ _ H3 H4). } + intros. + { rewrite (abstract_primitive_constant_correct _ _ _ H). + rewrite (abstract_primitive_constant_correct _ _ _ H0). + now rewrite H1 H2. } +Qed. + Program Fixpoint erase_global_decls {X_type : abstract_env_impl} (deps : KernameSet.t) (X : X_type.π1) (decls : global_declarations) {normalization_in : NormalizationIn_erase_global_decls X decls} @@ -1972,7 +2037,7 @@ Proof. now eapply lookup_env_In_map_fst in hl. Qed. -Lemma in_erase_global_decls_acc X_type X deps decls kn normalization_in prf : +Lemma in_erase_global_decls_acc X_type X deps decls kn normalization_in prf : KernameSet.In kn (erase_global_decls X_type deps X decls normalization_in prf).2 -> KernameSet.In kn deps \/ In kn (map fst decls). Proof. @@ -2086,132 +2151,85 @@ Proof. intros []; split => //. destruct s as [Σ'' eq]. eexists ((kn', d') :: Σ''). now rewrite eq. Qed. -(* -Lemma lookup_env_erase_decl X_type X deps decls normalization_in prf kn decl : - forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = Some decl -> - if KernameSet.mem kn deps then - match decl with - | ConstantDecl cb => - exists (X' : X_type.π1) nin pf, - EGlobalEnv.lookup_env (erase_global_decls X_type deps X decls normalization_in prf) kn = - Some (EAst.ConstantDecl (fst (erase_constant_decl X' cb nin pf))) - | InductiveDecl mib => - EGlobalEnv.lookup_env (erase_global_decls X_type deps X decls normalization_in prf) kn = - Some (EAst.InductiveDecl (erase_mutual_inductive_body mib)) - end - else - EGlobalEnv.lookup_env (erase_global_decls X_type deps X decls normalization_in prf) kn = None. +Lemma fresh_global_app kn Σ Σ' : fresh_global kn (Σ ++ Σ') -> + fresh_global kn Σ /\ fresh_global kn Σ'. Proof. - intros Σ wfΣ. pose proof (prf _ wfΣ). - unfold lookup_env. rewrite H. clear H. - rewrite /lookup_env. - induction decls in X, Σ , wfΣ ,prf, deps, decl, normalization_in |- *. - - move=> /= //. - - destruct a as [kn' d']. - cbn -[erase_global_decls]. - case: (eqb_spec kn kn'); intros e'; subst. - intros [= ->]. - case_eq (KernameSet.mem kn' deps); move => hin'; cbn; rewrite hin'; cbn. - + destruct decl; cbn; rewrite eqb_refl //. - do 3 eexists. reflexivity. - + specialize (IHdecls (abstract_pop_decls X) deps). - destruct decl. - * match goal with - |- context [ erase_global_decls _ _ _ _ ?nin ?prf ] => specialize (IHdecls nin prf) - end. - set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - case_eq (lookup_global decls kn'). - intros g hg. - specialize (IHdecls g _ wfpop hg). now rewrite hin' in IHdecls. - epose proof (lookup_env_erase_decl_None X_type (abstract_pop_decls X) deps decls _ _ _ _ wfpop). - intros hl. eapply H. unfold lookup_env. - unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wfΣ wfpop). - intros. pose proof (prf _ H0). now eexists. - destruct H0. now rewrite H0. - * rewrite hin' in IHdecls. - set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - case_eq (lookup_global decls kn'). - intros g hg. - eapply IHdecls; tea. - epose proof (lookup_env_erase_decl_None X_type (abstract_pop_decls X) deps decls _ _ _ _ wfpop). - intros hl; eapply H. - unfold lookup_env. - unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wfΣ wfpop). - intros. pose proof (prf _ H0). now eexists. - destruct H0. now rewrite H0. - + intros hl. - epose proof (lookup_env_erase_global_diff X_type X deps decls kn kn' d' _ _ e'). - destruct H as [deps' [nin' [prf' [sub H]]]]. - erewrite H. - case_eq (KernameSet.mem kn deps) => //. - * move/KernameSet.mem_spec => hin. - assert (kndeps' : KernameSet.mem kn deps'). - { now apply KernameSet.mem_spec; eapply sub. } - set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - specialize (IHdecls Xpop deps'). - epose proof (IHdecls _ _ decl _ wfpop hl) ; tea. rewrite kndeps' in H0. eapply H0. - * intros hin. rewrite hin in IHdecls. - + induction Σ; cbn; intuition auto. + - constructor. + - depelim H. constructor => //. + now eapply Forall_app in H0. + - depelim H. + now eapply Forall_app in H0. +Qed. - ** intros g hl. eapply IHdecls. - destruct KernameSet.mem. cbn. - rewrite (negbTE (proj2 (neqb _ _) e')). - eapply IHdecls => //; eauto. eapply KernameSet.union_spec. left => //. - eapply IHdecls => //; eauto. - simpl. +Lemma lookup_global_app_wf Σ Σ' kn : EnvMap.EnvMap.fresh_globals (Σ ++ Σ') -> + forall decl, lookup_global Σ' kn = Some decl -> lookup_global (Σ ++ Σ') kn = Some decl. +Proof. + induction Σ in kn |- *; cbn; auto. + intros fr; depelim fr. + intros decl hd; cbn. + destruct (eqb_spec kn kn0). + eapply PCUICWeakeningEnv.lookup_global_Some_fresh in hd. + subst. now eapply fresh_global_app in H as [H H']. + now eapply IHΣ. +Qed. - destruct KernameSet.mem. cbn. - rewrite (negbTE (proj2 (neqb _ _) e')). - eapply IHdecls => //; eauto. - eapply IHdecls => //; eauto. -Qed.*) +Lemma strictly_extends_decls_lookup {Σ Σ'} : + wf Σ' -> + strictly_extends_decls Σ Σ' -> + forall kn d, lookup_env Σ kn = Some d -> + lookup_env Σ' kn = Some d. +Proof. + intros. + destruct X0 as [? []]. + rewrite /lookup_env in H *. rewrite e0. + erewrite lookup_global_app_wf. reflexivity. 2:eauto. + eapply wf_fresh_globals in X. now rewrite -e0. +Qed. -Lemma in_deps_in_erase_global_decls X_type X deps decls normalization_in prf kn decl : - forall Σ, Σ ∼ X -> - lookup_env Σ kn = Some decl -> +Lemma in_deps_in_erase_global_decls X_type X deps decls normalization_in prf kn : KernameSet.In kn deps -> KernameSet.In kn (erase_global_decls X_type deps X decls normalization_in prf).2. Proof. induction decls in X, deps, normalization_in, prf |- *. - - intros. rewrite /lookup_env prf in H0 => //. + - intros. now cbn. - intros. cbn. + pose proof (abstract_env_exists X) as [[Σ H']]. set (Xpop := abstract_pop_decls X). epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ H wfpop). intros. rewrite prf => //. now eexists. - destruct H2 as [Hd [Hu Hr]]. + unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ H' wfpop). intros. rewrite prf => //. now eexists. + destruct H0 as [Hd [Hu Hr]]. destruct a. destruct g. + set (eg := erase_global_decls _ _ _ _ _ _); set (eg' := erase_global_decls _ _ _ _ _ _). case_eq (KernameSet.mem k deps). * cbn. intros hm. destruct (eqb_spec k kn). - ++ subst k. admit. - ++ eapply IHdecls. tea. - rewrite /lookup_env Hd. - rewrite /lookup_env (prf _ H) in H0. cbn in H0. - admit. - KernameSetDecide.fsetdec. + ++ subst k. eapply IHdecls. KernameSetDecide.fsetdec. + ++ eapply IHdecls. + KernameSetDecide.fsetdec. * intros hnin. eapply IHdecls; tea. - rewrite /lookup_env Hd. - rewrite /lookup_env (prf _ H) in H0. cbn in H0. admit. + set (eg := erase_global_decls _ _ _ _ _ _); set (eg' := erase_global_decls _ _ _ _ _ _). case_eq (KernameSet.mem k deps). * cbn. intros hm. destruct (eqb_spec k kn). - ++ subst k. admit. + ++ subst k. eapply KernameSet.mem_spec in hm. + now eapply IHdecls. ++ eapply IHdecls. tea. - rewrite /lookup_env Hd. - rewrite /lookup_env (prf _ H) in H0. cbn in H0. - admit. - KernameSetDecide.fsetdec. * intros hnin. eapply IHdecls; tea. - rewrite /lookup_env Hd. - rewrite /lookup_env (prf _ H) in H0. cbn in H0. admit. +Qed. + +Lemma erase_global_decls_eval {efl : EWcbvEval.WcbvFlags} (X_type : abstract_env_impl) (X:X_type.π1) ext wfX {normalization_in} decls nin prf Γ t wt v wv : + let Xext := abstract_make_wf_env_ext X ext wfX in + let t' := erase X_type Xext (normalization_in := normalization_in) Γ t wt in + let v' := erase X_type Xext (normalization_in := normalization_in) Γ v wv in + let eg := erase_global_decls X_type (term_global_deps t') X decls nin prf in + let eg' := erase_global_decls X_type (term_global_deps v') X decls nin prf in + eg.1 ⊢ t' ⇓ v' -> KernameSet.Subset eg'.2 eg.2. +Proof. + intros. clearbody t' v'. Admitted. Lemma erase_global_declared_constructor X_type X ind c mind idecl cdecl deps decls normalization_in prf: @@ -2603,7 +2621,7 @@ Proof. pose proof (abstract_env_wf _ wf) as [wfΣ]. assert (erases_mutual_inductive_body m (erase_mutual_inductive_body m)). { eapply (erases_mutual (mdecl:=kn)); tea. } - eapply (erases_mutual_inductive_body_wf (univs := Σ.(universes)) (retro := Σ.(retroknowledge)) (Σ := decls) (kn := kn) (Σ' := Σ')) in H; tea. + eapply (erases_mutual_inductive_body_wf (univs := Σ.(universes)) (retro := Σ.(retroknowledge)) (Σ := decls) (kn := kn) (Σ' := Σ'.1)) in H; tea. rewrite -(heq _ wf). now destruct Σ. Qed. @@ -2798,7 +2816,7 @@ Lemma erase_eval_to_box (wfl := Ee.default_wcbv_flags) Proof. intros wt. intros. - destruct (erase_correct X_type X univs wfext _ _ Σ'.1 _ _ decls prf wt H H0 H1 X0 _ H2) as [ev [eg [eg']]]. + destruct (erase_correct X_type X univs wfext _ _ Σ' _ _ decls prf wt H H0 H1 X0 _ H2) as [ev [eg [eg']]]. pose proof (Ee.eval_deterministic H3 eg'). subst. pose proof (abstract_env_exists X) as [[? wf]]. destruct (wfext _ wf). destruct (wt _ H2) as [T wt']. @@ -3071,7 +3089,7 @@ Qed. Program Fixpoint erase_global_decls_fast (deps : KernameSet.t) X_type (X:X_type.π1) (decls : global_declarations) {normalization_in : NormalizationIn_erase_global_decls_fast X decls} - (prop : forall Σ : global_env, abstract_env_rel X Σ -> ∥ decls_prefix decls Σ ∥) : E.global_declarations := + (prop : forall Σ : global_env, abstract_env_rel X Σ -> ∥ decls_prefix decls Σ ∥) : E.global_declarations * KernameSet.t := match decls with | [] => ([],deps) | (kn, ConstantDecl cb) :: decls => @@ -3214,46 +3232,6 @@ Definition same_principal_type {cf} {X_type' : abstract_env_impl} {X' : X_type'.π2.π1} {Γ : context} {t} (p : PCUICSafeRetyping.principal_type X_type X Γ t) (p' : PCUICSafeRetyping.principal_type X_type' X' Γ t) := p.π1 = p'.π1. -(* -Definition Hlookup {cf} (X_type : abstract_env_impl) (X : X_type.π2.π1) - (X_type' : abstract_env_impl) (X' : X_type'.π2.π1) := - forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> - forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> - forall kn decl decl', - lookup_env Σ kn = Some decl -> - lookup_env Σ' kn = Some decl' -> - abstract_env_lookup X kn = abstract_env_lookup X' kn. *) - - -(* When two environments agree on the intersection of their declarations *) -Definition equiv_env_inter Σ Σ' := - (forall kn decl decl', lookup_env Σ kn = Some decl -> lookup_env Σ' kn = Some decl' -> decl = decl') /\ - (forall tag : Primitive.prim_tag, primitive_constant Σ tag = primitive_constant Σ' tag). - -Lemma equiv_env_inter_sym Σ Σ' : - equiv_env_inter Σ Σ' <-> equiv_env_inter Σ' Σ. -Proof. - unfold equiv_env_inter. - intuition auto; symmetry; eauto. -Qed. - -(* When two environments agree on the intersection of their declarations *) -Lemma equiv_env_inter_hlookup {cf:checker_flags} (X_type : abstract_env_impl) (X : X_type.π2.π1) - (X_type' : abstract_env_impl) (X' : X_type'.π2.π1) : - (forall Σ Σ', Σ ∼_ext X -> Σ' ∼_ext X' -> equiv_env_inter Σ Σ') -> Hlookup X_type X X_type' X'. -Proof. - intros hl Σ HX Σ' HX'. - specialize (hl _ _ HX HX') as [hl hr]. - split. - - intros kn decl decl'. - specialize (hl kn decl decl'). - rewrite <-(abstract_env_lookup_correct' _ _ HX). - rewrite <-(abstract_env_lookup_correct' _ _ HX'). - intros Hl Hl'. - rewrite Hl Hl'. f_equal. eauto. - - intros tag; rewrite (abstract_primitive_constant_correct _ _ _ HX) (abstract_primitive_constant_correct _ _ _ HX'). - apply hr. -Qed. (*Lemma erase_global_deps_suffix {deps} {Σ Σ' : wf_env} {decls hprefix hprefix'} : wf Σ -> wf Σ' -> @@ -3353,6 +3331,9 @@ Proof using Type. _ hprefix kn c decls eq_refl)). set (eb := erase_constant_body _ _ _ _). set (eb' := erase_constant_body _ _ _ _). + set (eg := erase_global_decls _ _ _ _ _ _). + set (eg' := erase_global_decls _ _ _ _ _ _). + assert (eb = eb') as <-. { subst eb eb'. destruct (hprefix _ wfΣ) as [[Σ'' eq]]. @@ -3370,20 +3351,30 @@ Proof using Type. cbn. eexists (Σ'' ++ [(kn, ConstantDecl c)]). cbn. subst. now rewrite -app_assoc. subst. symmetry. now apply equ. } - destruct KernameSet.mem => //; f_equal; eapply IHdecls. + destruct KernameSet.mem => //; f_equal; try eapply IHdecls. + f_equal. f_equal. eapply IHdecls. + intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. + { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. + f_equal. eapply IHdecls. intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. - cbn. - destruct KernameSet.mem => //; f_equal; eapply IHdecls. + destruct KernameSet.mem => //; f_equal; try eapply IHdecls. + f_equal; f_equal; eapply IHdecls. intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. + intros. + f_equal; eapply IHdecls. intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. + intros; unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. + { now eexists. } + intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. Qed. Lemma erase_global_deps_fast_spec {deps} {X_type X} {decls normalization_in normalization_in' hprefix hprefix'} : @@ -3401,7 +3392,7 @@ Lemma expanded_erase_global_fast (cf := config.extraction_checker_flags) deps {X_type X decls normalization_in} {normalization_in':NormalizationIn_erase_global_decls X decls} {prf} : forall Σ : global_env, abstract_env_rel X Σ -> PCUICEtaExpand.expanded_global_env Σ -> - expanded_global_env (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf). + expanded_global_env (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf).1. Proof using Type. unfold erase_global_fast. erewrite erase_global_deps_fast_spec. @@ -3417,7 +3408,7 @@ Lemma expanded_erase_fast (cf := config.extraction_checker_flags) forall (normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ), let et := (erase X_type X' [] t wtp) in let deps := EAstUtils.term_global_deps et in - expanded (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf) [] et. + expanded (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf).1 [] et. Proof using Type. intros Σ wf hexp X' normalization_in'. pose proof (abstract_env_wf _ wf) as [?]. eapply (expanded_erases (Σ := (Σ, univs))); tea. @@ -3430,7 +3421,7 @@ Proof using Type. Qed. Lemma erase_global_fast_wf_glob X_type deps X decls normalization_in (normalization_in':NormalizationIn_erase_global_decls X decls) prf : - @wf_glob all_env_flags (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf). + @wf_glob all_env_flags (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf).1. Proof using Type. unfold erase_global_fast; erewrite erase_global_deps_fast_spec. eapply erase_global_decls_wf_glob. @@ -3442,7 +3433,7 @@ Lemma erase_wellformed_fast (efl := all_env_flags) (X' := abstract_make_wf_env_ext X univs wfΣ) {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} : let t' := erase X_type X' Γ t wt in - wellformed (erase_global_fast X_type (term_global_deps t') X decls (normalization_in:=normalization_in) prf) #|Γ| t'. + wellformed (erase_global_fast X_type (term_global_deps t') X decls (normalization_in:=normalization_in) prf).1 #|Γ| t'. Proof using Type. intros. cbn. unfold erase_global_fast. erewrite erase_global_deps_fast_spec. From 7158506c5564b3f57b6240c385136a94a3ede002 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 28 Aug 2023 10:47:26 +0200 Subject: [PATCH 40/61] Strenghtened program extension preservation --- erasure-plugin/theories/ETransform.v | 154 +++++------------------ erasure/theories/EConstructorsAsBlocks.v | 10 ++ erasure/theories/EGenericGlobalMap.v | 18 +++ erasure/theories/EImplementBox.v | 19 ++- erasure/theories/EInlineProjections.v | 11 +- erasure/theories/EOptimizePropDiscr.v | 10 ++ erasure/theories/ERemoveParams.v | 11 ++ 7 files changed, 101 insertions(+), 132 deletions(-) diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index 3dcc50ac2..acef38110 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -314,116 +314,6 @@ Proof. + intros _ hl. eauto. Qed. -Lemma in_erase_global_decls_acc {X_type : abstract_env_impl} (X : X_type.π1) deps decls kn normalization_in prf : - In kn (map fst (erase_global_decls deps X decls (normalization_in:=normalization_in) prf).1) -> - KernameSet.In kn (erase_global_decls deps X decls (normalization_in:=normalization_in) prf).2. -Proof. - induction decls in X, prf, deps, normalization_in |- *. - cbn; auto. destruct a as [kn' []]. - + cbn. set (ec := erase_constant_decl _ _ _ _). - set (eg := erase_global_decls _ _ _ _). - set (eg' := erase_global_decls _ _ _ _). - case_eq (KernameSet.mem kn' deps). - * cbn; intros. destruct H0. subst; auto. - eapply KernameSet.mem_spec in H. - -Admitted. - (*auto. - specialize (IHdecls _ _ _ _ H0). - unfold ec in H0; cbn in H0. - set (Xpop := abstract_pop_decls X) in *. - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - epose proof (abstract_env_exists X) as [[ΣX HX]]. - unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ HX wfpop). - intros. rewrite prf => //. now eexists. - destruct H1 as [<- ?]. auto. - * intros hdeps hin. - eapply IHdecls in hin. intuition auto. - + cbn; set (eg := erase_global_decls _ _ _ _); - set (eg' := erase_global_decls _ _ _ _). - case_eq (KernameSet.mem kn' deps). - * cbn; intros. intuition auto. now specialize (IHdecls _ _ _ _ H1). - * intros hdeps hin. - eapply IHdecls in hin. intuition auto. - Unshelve. all:tc. -Qed. -*) -#[global] -Instance erase_transform_extends {guard : abstract_guard_impl} : - TransformExt.t (erase_transform (guard := guard)) extends_pcuic_program extends_eprogram_env. -Proof. - red. intros p p' pr pr' [ext eq]. rewrite /transform /=. - destruct pr. - epose proof (erase_pcuic_program_normalization_helper p (map_squash fst s)). - destruct p as [Σ t]. destruct p' as [Σ' t']. - simpl in *. subst t'. red. split. - - destruct ext. red. - cbn. rewrite /ErasureFunction.erase_global_fast. - unshelve erewrite ErasureFunction.erase_global_deps_fast_spec. - { intros. intuition auto. eapply H11; tea. admit. admit. } - { intros. red in H4. cbn in H4. subst. reflexivity. } - destruct Σ'. cbn in *. subst u. - set (fst := ErasureFunction.erase _ _ _ _ _). - set (snd:= ErasureFunction.erase _ _ _ _ _). - assert (fst = snd). - { subst fst snd. symmetry. - eapply ErasureFunction.erase_irrel_global_env. - eapply ErasureFunction.equiv_env_inter_hlookup. - intros ? ? -> ->. cbn. - eapply ErasureFunction.equiv_env_inter_sym. - eapply ErasureFunction.extends_global_env_equiv_env. - cbn. split => //. - sq. unfold primitive_constant. now rewrite H3. } - clearbody fst snd. subst snd. - intros. - set (X := build_wf_env_from_env g _) in *. - destruct (lookup_env Σ kn) eqn:E. - unshelve epose proof (abstract_env_exists X). 3:tc. tc. - destruct H4 as [[Σ' ΣX]]. - unshelve epose proof (build_wf_env_from_env_eq _ (g, Σ.2) _ _ ΣX). cbn in H4. - epose proof (hl := ErasureFunction.lookup_env_erase_decl (optimized_abstract_env_impl) _ (EAstUtils.term_global_deps fst) (declarations g) _ _ kn g0 _ ΣX). - unfold lookup_env in E, hl. rewrite H4 in hl. specialize (hl (H0 _ _ E)). - destruct g0. - + destruct hl as [X' [nin [pf [eq ext]]]]; revgoals. - unshelve erewrite ErasureFunction.erase_global_deps_fast_spec. shelve. - { intros. red in H5. cbn in H5. subst. reflexivity. } - erewrite eq. - 2:{ eapply lookup_env_In_map_fst in H2. eapply in_erase_global_decls_acc in H2. } - set (Xs := build_wf_env_from_env Σ.1 _) in *. - unshelve epose proof (abstract_env_exists Xs). 3:tc. tc. - destruct H5 as [[Σs Hs]]. - epose proof (hl' := ErasureFunction.lookup_env_erase_decl (optimized_abstract_env_impl) _ (EAstUtils.term_global_deps fst) _ _ _ kn _ _ Hs). - unshelve epose proof (build_wf_env_from_env_eq _ _ _ _ Hs). rewrite /lookup_env H5 in hl'. specialize (hl' E). - forward hl'. admit. cbn in hl'. - destruct hl' as [X'' [nin' [pf' [eq' ext']]]]. - erewrite eq' in H2. rewrite -H2. f_equal. f_equal. - f_equal. unfold erase_constant_decl. - eapply erase_constant_body_irrel. - eapply ErasureFunction.equiv_env_inter_hlookup. cbn. - intros ? ? H6 H7. - cbn in ext. specialize (ext _ _ eq_refl eq_refl) as [ext]. - specialize (ext' _ _ eq_refl eq_refl) as [ext']. - cbn in H6, H7. subst Σ0 Σ'0. cbn. split => //. - { intros kn' decl'. rewrite /lookup_env. - destruct (map_squash Datatypes.fst s). - destruct pr' as [pr' ?]. - destruct (map_squash Datatypes.fst pr'). - eapply strictly_extends_lookups. 2-3:tea. - intros kn'' ? ? hg hΣ. eapply H0 in hΣ. congruence. auto. - apply X0. apply X1. } - { rewrite /primitive_constant. - destruct ext as [_ _ ->]. - destruct ext' as [_ _ ->]. - now rewrite H3. } - + - - symmetry. - eapply ErasureFunction.erase_irrel_global_env. - intros Σg Σg' eq eq'. - destruct ext as [hl hu hr]. - destruct Σ as [m univs]. destruct Σ' as [m' univs']. simpl in *. - subst. cbn. destruct m, m'. cbn in *. split => //. sq. red. split => //. - unfold primitive_constant. now rewrite H. -Qed. End PCUICEnv. Obligation Tactic := idtac. @@ -455,7 +345,7 @@ Qed. #[global] Instance guarded_to_unguarded_fix_extends {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} {efl : EEnvFlags} (wguard : with_guarded_fix) : - TransformExt.t (guarded_to_unguarded_fix (wcon:=wcon) wguard) extends_eprogram_env term snd. + TransformExt.t (guarded_to_unguarded_fix (wcon:=wcon) wguard) extends_eprogram_env extends_eprogram_env. Proof. red. intros p p' pr pr' [ext eq]. now rewrite /transform /=. Qed. @@ -479,7 +369,7 @@ Qed. #[global] Instance rebuild_wf_env_extends {fl : EWcbvEval.WcbvFlags} {efl : EEnvFlags} with_exp : - TransformExt.t (rebuild_wf_env_transform with_exp) extends_eprogram term snd. + TransformExt.t (rebuild_wf_env_transform with_exp) extends_eprogram extends_eprogram_env. Proof. red. intros p p' pr pr' [ext eq]. now rewrite /transform /=. Qed. @@ -513,9 +403,10 @@ Qed. #[global] Instance remove_params_extends {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := all_env_flags): - TransformExt.t (remove_params_optimization (wcon:=wcon)) extends_eprogram_env term snd. + TransformExt.t (remove_params_optimization (wcon:=wcon)) extends_eprogram_env extends_eprogram. Proof. - red. intros p p' pr pr' [ext eq]. rewrite /transform /=. rewrite eq. + red. intros p p' pr pr' [ext eq]. rewrite /transform /= /strip_program. rewrite eq. + red. cbn -[strip_env strip]. split. eapply strip_extends_env => //. apply pr. apply pr'. eapply strip_extends => //. apply pr'. rewrite -eq. apply pr. Qed. @@ -552,10 +443,11 @@ Qed. #[global] Instance remove_params_fast_extends {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := all_env_flags): - TransformExt.t (remove_params_fast_optimization (wcon:=wcon)) extends_eprogram_env term snd. + TransformExt.t (remove_params_fast_optimization (wcon:=wcon)) extends_eprogram_env extends_eprogram. Proof. red. intros p p' pr pr' [ext eq]. rewrite /transform /=. rewrite eq. - rewrite -!ERemoveParams.Fast.strip_fast. + rewrite -!ERemoveParams.Fast.strip_fast -!ERemoveParams.Fast.strip_env_fast. + split => /=. eapply strip_extends_env => //. apply pr. apply pr'. eapply strip_extends => //. apply pr'. rewrite -eq. apply pr. Qed. @@ -586,9 +478,11 @@ Qed. #[global] Instance remove_match_on_box_extends {fl : WcbvFlags} {wcon : with_constructor_as_block = false} {efl : EEnvFlags} {hastrel : has_tRel} {hastbox : has_tBox} : - TransformExt.t (remove_match_on_box_trans (wcon:=wcon) (hastrel := hastrel) (hastbox := hastbox)) extends_eprogram_env term snd. + TransformExt.t (remove_match_on_box_trans (wcon:=wcon) (hastrel := hastrel) (hastbox := hastbox)) extends_eprogram_env extends_eprogram. Proof. - red. intros p p' pr pr' [ext eq]. rewrite /transform /=. rewrite -eq. + red. intros p p' pr pr' [ext eq]. rewrite /transform /= /remove_match_on_box_program; split => /=. + eapply remove_match_on_box_extends_env => //. apply pr. apply pr'. + rewrite -eq. eapply wellformed_remove_match_on_box_extends; eauto. apply pr. apply pr'. Qed. @@ -619,9 +513,12 @@ Qed. #[global] Instance inline_projections_optimization_extends {fl : WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := switch_no_params all_env_flags) {hastrel : has_tRel} {hastbox : has_tBox} : - TransformExt.t (inline_projections_optimization (wcon:=wcon) (hastrel := hastrel) (hastbox := hastbox)) extends_eprogram_env term snd. + TransformExt.t (inline_projections_optimization (wcon:=wcon) (hastrel := hastrel) (hastbox := hastbox)) extends_eprogram_env extends_eprogram. Proof. - red. intros p p' pr pr' [ext eq]. rewrite /transform /=. rewrite -eq. + red. intros p p' pr pr' [ext eq]. rewrite /transform /= /optimize_program /=. + split => /=. + eapply optimize_extends_env => //. apply pr. apply pr'. + rewrite -eq. eapply wellformed_optimize_extends; eauto. apply pr. apply pr'. Qed. @@ -655,9 +552,13 @@ Qed. #[global] Instance constructors_as_blocks_extends (efl : EEnvFlags) {has_app : has_tApp} {has_rel : has_tRel} {hasbox : has_tBox} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = false} : - TransformExt.t (constructors_as_blocks_transformation (has_app := has_app) (has_rel := has_rel) (hasbox := hasbox) (has_pars := has_pars) (has_cstrblocks := has_cstrblocks)) extends_eprogram_env term snd. + TransformExt.t (constructors_as_blocks_transformation (has_app := has_app) (has_rel := has_rel) (hasbox := hasbox) (has_pars := has_pars) (has_cstrblocks := has_cstrblocks)) + extends_eprogram_env extends_eprogram. Proof. - red. intros p p' pr pr' [ext eq]. rewrite /transform /=. rewrite -eq. + red. intros p p' pr pr' [ext eq]. rewrite /transform /=. + split => //. eapply transform_blocks_env_extends => //. apply pr. apply pr'. + unfold transform_blocks_program => /=. + rewrite -eq. eapply transform_blocks_extends; eauto. apply pr. apply pr'. Qed. @@ -674,7 +575,7 @@ Program Definition implement_box_transformation {efl : EEnvFlags} Next Obligation. intros. cbn in *. destruct p. split. - - eapply transform_wf_global; eauto. + - eapply implement_box_env_wf_glob; eauto. - now eapply transform_wellformed'. Qed. Next Obligation. @@ -686,7 +587,10 @@ Qed. #[global] Instance implement_box_extends (efl : EEnvFlags) {has_app : has_tApp} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = true} : - TransformExt.t (implement_box_transformation (has_app := has_app) (has_pars := has_pars) (has_cstrblocks := has_cstrblocks)) extends_eprogram term snd. + TransformExt.t (implement_box_transformation (has_app := has_app) (has_pars := has_pars) (has_cstrblocks := has_cstrblocks)) extends_eprogram extends_eprogram. Proof. - red. intros p p' pr pr' [ext eq]. rewrite /transform /=. now rewrite -eq. + red. intros p p' pr pr' [ext eq]. rewrite /transform /= /implement_box_program /=. + split => /=. + eapply (implement_box_env_extends has_app ext). apply pr. apply pr'. + now rewrite -eq. Qed. diff --git a/erasure/theories/EConstructorsAsBlocks.v b/erasure/theories/EConstructorsAsBlocks.v index a8cd21c0a..1fc9a4ab7 100644 --- a/erasure/theories/EConstructorsAsBlocks.v +++ b/erasure/theories/EConstructorsAsBlocks.v @@ -857,6 +857,16 @@ Proof. eapply transform_wellformed => //. Defined. +Lemma transform_blocks_env_extends {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : + has_tApp -> + extends Σ Σ' -> + wf_glob Σ -> + wf_glob Σ' -> + extends (transform_blocks_env Σ) (transform_blocks_env Σ'). +Proof. + intros hast ext wf. + now apply (gen_transform_extends (gt := GTExt efl hast) ext). +Qed. Lemma Pre_glob efl Σ wf : has_cstr_params = false -> diff --git a/erasure/theories/EGenericGlobalMap.v b/erasure/theories/EGenericGlobalMap.v index 4fde99a95..797eedcf3 100644 --- a/erasure/theories/EGenericGlobalMap.v +++ b/erasure/theories/EGenericGlobalMap.v @@ -289,6 +289,24 @@ Proof. now apply extends_cons_wf. now depelim wf. Qed. +Lemma gen_transform_extends {Σ Σ' : GlobalContextMap.t} : + extends Σ Σ' -> + @wf_glob efl Σ -> + @wf_glob efl Σ' -> + extends (gen_transform_env Σ) (gen_transform_env Σ'). +Proof. + intros ext. + unfold gen_transform_env. + intros wf wf'. + rewrite (gen_transform_env_extends' ext wf wf'). + intros kn d. specialize (ext kn). + rewrite lookup_env_map_snd. + destruct (lookup_env) eqn:E; cbn => //. + intros [= <-]. + specialize (ext _ eq_refl). + now rewrite lookup_env_map_snd ext /=. +Qed. + Import EWellformed. Lemma gen_transform_wellformed_irrel {genid : GenTransformId gen_transform} {Σ : GlobalContextMap.t} t : diff --git a/erasure/theories/EImplementBox.v b/erasure/theories/EImplementBox.v index e843cd6ba..b322d74d8 100644 --- a/erasure/theories/EImplementBox.v +++ b/erasure/theories/EImplementBox.v @@ -411,7 +411,7 @@ Proof. induction 1; cbn; constructor; auto. Qed. -Lemma transform_wf_global {efl : EEnvFlags} {Σ : global_declarations} : +Lemma implement_box_env_wf_glob {efl : EEnvFlags} {Σ : global_declarations} : has_tApp -> wf_glob (efl := efl) Σ -> wf_glob (efl := switch_off_box efl) (implement_box_env Σ). Proof. @@ -428,6 +428,23 @@ Proof. eapply EExtends.extends_wf_glob in wfg; tea. now depelim wfg. Qed. +From MetaCoq.Erasure Require Import EGenericGlobalMap. + +Lemma implement_box_env_extends {efl : EEnvFlags} {Σ Σ' : global_declarations} : + has_tApp -> + extends Σ Σ' -> + wf_glob Σ -> + wf_glob Σ' -> + extends (implement_box_env Σ) (implement_box_env Σ'). +Proof. + intros hast ext wf wf'. + intros kn d. + rewrite !lookup_env_map_snd. + specialize (ext kn). destruct (lookup_env) eqn:E => //=. + intros [= <-]. + rewrite (ext g) => //. +Qed. + Transparent implement_box. Lemma fst_decompose_app_rec t l : fst (decompose_app_rec t l) = fst (decompose_app t). diff --git a/erasure/theories/EInlineProjections.v b/erasure/theories/EInlineProjections.v index 8f65984c3..7d760e3b2 100644 --- a/erasure/theories/EInlineProjections.v +++ b/erasure/theories/EInlineProjections.v @@ -841,16 +841,15 @@ Proof. now apply optimize_wellformed. Defined. -Lemma optimize_env_extends' {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : +Lemma optimize_extends_env {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : + has_tApp -> extends Σ Σ' -> wf_glob Σ -> wf_glob Σ' -> - List.map (on_snd (optimize_decl Σ)) Σ.(GlobalContextMap.global_decls) = - List.map (on_snd (optimize_decl Σ')) Σ.(GlobalContextMap.global_decls). + extends (optimize_env Σ) (optimize_env Σ'). Proof. - intros. - epose proof (gen_transform_env_extends' (gt := GTExt efl)). - now eapply H2. + intros hast ext wf. + now apply (gen_transform_extends (gt := GTExt efl) ext). Qed. Lemma optimize_env_eq {efl : EEnvFlags} (Σ : GlobalContextMap.t) : wf_glob Σ -> optimize_env Σ = optimize_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index af091df81..4e50ee4ca 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -879,6 +879,16 @@ Proof. eapply (gen_transform_env_extends' (gt := GTExt efl)) => //. Qed. +Lemma remove_match_on_box_extends_env {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : + extends Σ Σ' -> + wf_glob Σ -> + wf_glob Σ' -> + extends (remove_match_on_box_env Σ) (remove_match_on_box_env Σ'). +Proof. + intros ext wf. + now apply (gen_transform_extends (gt := GTExt efl) ext). +Qed. + Lemma remove_match_on_box_env_eq {efl : EEnvFlags} (Σ : GlobalContextMap.t) : wf_glob Σ -> remove_match_on_box_env Σ = remove_match_on_box_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). diff --git a/erasure/theories/ERemoveParams.v b/erasure/theories/ERemoveParams.v index bc5a6b529..4d6937bf0 100644 --- a/erasure/theories/ERemoveParams.v +++ b/erasure/theories/ERemoveParams.v @@ -1247,6 +1247,17 @@ Proof. now apply (gen_transform_env_extends' (gt := GTExt efl hast) ext). Qed. +Lemma strip_extends_env {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : + has_tApp -> + extends Σ Σ' -> + wf_glob Σ -> + wf_glob Σ' -> + extends (strip_env Σ) (strip_env Σ'). +Proof. + intros hast ext wf. + now apply (gen_transform_extends (gt := GTExt efl hast) ext). +Qed. + Lemma strip_env_eq (efl := all_env_flags) (Σ : GlobalContextMap.t) : wf_glob Σ -> strip_env Σ = strip_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). Proof. intros wf. From 3856b002c63bf90263b70f3d53d0ffdc470cc281 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 28 Aug 2023 18:02:39 +0200 Subject: [PATCH 41/61] Progress on pipeline theorem proof, main argument remains --- erasure-plugin/theories/Erasure.v | 79 ++-- erasure-plugin/theories/ErasureCorrectness.v | 448 +++++++++++++++++++ 2 files changed, 479 insertions(+), 48 deletions(-) create mode 100644 erasure-plugin/theories/ErasureCorrectness.v diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index ce6e68ef0..b44821921 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -58,25 +58,14 @@ Next Obligation. apply assume_preservation_template_program_env_expansion in ev as [ev']; eauto. Qed. -Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - Transform.t pcuic_program EProgram.eprogram - PCUICAst.term EAst.term - PCUICTransform.eval_pcuic_program - (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := - (* a bunch of nonsense for normalization preconditions *) - let K ty (T : ty -> _) p - := let p := T p in - (PCUICTyping.wf_ext p -> PCUICSN.NormalizationIn p) /\ - (PCUICTyping.wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p) in - let T1 (p:global_env_ext_map) := p in - (* Branches of cases are expanded to bind only variables, constructor types are expanded accordingly *) - pcuic_expand_lets_transform (K _ T1) ▷ - (* Erasure of proofs terms in Prop and types *) - erase_transform ▷ +Program Definition verified_lambdabox_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : + Transform.t EProgram.eprogram_env EProgram.eprogram EAst.term EAst.term + (EProgram.eval_eprogram_env {| with_prop_case := true; with_guarded_fix := true; with_constructor_as_block := false |}) + (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := (* Simulation of the guarded fixpoint rules with a single unguarded one: the only "stuck" fixpoints remaining are unapplied. This translation is a noop on terms and environments. *) - guarded_to_unguarded_fix (wcon := eq_refl) eq_refl ▷ + guarded_to_unguarded_fix (fl := default_wcbv_flags) (wcon := eq_refl) eq_refl ▷ (* Remove all constructor parameters *) remove_params_optimization (wcon := eq_refl) ▷ (* Rebuild the efficient lookup table *) @@ -87,11 +76,12 @@ Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl rebuild_wf_env_transform (efl := ERemoveParams.switch_no_params EWellformed.all_env_flags) true ▷ (* Inline projections to cases *) inline_projections_optimization (fl := EWcbvEval.target_wcbv_flags) (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ - let efl := EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags) in (* Rebuild the efficient lookup table *) - rebuild_wf_env_transform (efl := efl) true ▷ + rebuild_wf_env_transform (efl := EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags)) true ▷ (* First-order constructor representation *) - constructors_as_blocks_transformation (efl := efl) (has_app := eq_refl) (has_pars := eq_refl) (has_rel := eq_refl) (hasbox := eq_refl) (has_cstrblocks := eq_refl). + constructors_as_blocks_transformation + (efl := EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags)) + (has_app := eq_refl) (has_pars := eq_refl) (has_rel := eq_refl) (hasbox := eq_refl) (has_cstrblocks := eq_refl). (* At the end of erasure we get a well-formed program (well-scoped globally and localy), without parameters in inductive declarations. The constructor applications are also transformed to a first-order @@ -107,6 +97,25 @@ Next Obligation. Qed. +Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : + Transform.t pcuic_program EProgram.eprogram + PCUICAst.term EAst.term + PCUICTransform.eval_pcuic_program + (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := + (* a bunch of nonsense for normalization preconditions *) + let K ty (T : ty -> _) p + := let p := T p in + (PCUICTyping.wf_ext p -> PCUICSN.NormalizationIn p) /\ + (PCUICTyping.wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p) in + let T1 (p:global_env_ext_map) := p in + (* Branches of cases are expanded to bind only variables, constructor types are expanded accordingly *) + pcuic_expand_lets_transform (K _ T1) ▷ + (* Erasure of proofs terms in Prop and types *) + erase_transform ▷ + verified_lambdabox_pipeline. + +Import EGlobalEnv EWellformed. + Definition transform_compose {program program' program'' value value' value'' : Type} {eval : program -> value -> Prop} {eval' : program' -> value' -> Prop} @@ -122,38 +131,12 @@ Proof. eexists;reflexivity. Qed. -#[global] -Instance pcuic_expand_lets_transform_ext {cf : checker_flags} K : - TransformExt.t (pcuic_expand_lets_transform K) extends_pcuic_program PCUICAst.term snd. +Lemma verified_lambdabox_pipeline_extends {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : + TransformExt.t verified_lambdabox_pipeline extends_eprogram_env extends_eprogram. Proof. - red. intros p p' pre pre' []. cbn. now rewrite H0. + unfold verified_lambdabox_pipeline. tc. Qed. -Lemma verified_erasure_pipeline_extends {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - TransformExt.t verified_erasure_pipeline extends_pcuic_program EAst.term snd. -Proof. - unfold verified_erasure_pipeline. - eapply TransformExt.compose. - 2:eapply TransformExt.compose. - 2: typeclasses eauto. 2:typeclasses eauto. - 2:{ auto. } - repeat try (eapply TransformExt.compose; [|typeclasses eauto|]). - typeclasses eauto. - all:unfold extends_pcuic_program; cbn. - all:intros; intuition auto; try congruence. - red. cbn. - - - - - - - - - } - red. - - Program Definition pre_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : Transform.t TemplateProgram.template_program pcuic_program Ast.term PCUICAst.term diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v new file mode 100644 index 000000000..01c4ac15e --- /dev/null +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -0,0 +1,448 @@ +(* Distributed under the terms of the MIT license. *) +From Coq Require Import Program ssreflect ssrbool. +From MetaCoq.Common Require Import Transform config. +From MetaCoq.Utils Require Import bytestring utils. +From MetaCoq.PCUIC Require PCUICAst PCUICAstUtils PCUICProgram. +From MetaCoq.SafeChecker Require Import PCUICErrors PCUICWfEnvImpl. +From MetaCoq.Erasure Require EAstUtils ErasureFunction ErasureCorrectness EPretty Extract. +From MetaCoq Require Import ETransform EConstructorsAsBlocks. +From MetaCoq.Erasure Require Import EWcbvEvalNamed. +From MetaCoq.ErasurePlugin Require Import Erasure. +Import PCUICProgram. +(* Import TemplateProgram (template_eta_expand). + *) +Import PCUICTransform (template_to_pcuic_transform, pcuic_expand_lets_transform). + +(* This is the total erasure function + + let-expansion of constructor arguments and case branches + + shrinking of the global environment dependencies + + the optimization that removes all pattern-matches on propositions. *) + +Import Transform. + +Obligation Tactic := program_simpl. + +#[local] Existing Instance extraction_checker_flags. + +Import EWcbvEval. + +Lemma transform_compose_assoc + {program program' program'' program''' value value' value'' value''' : Type} + {eval : program -> value -> Prop} {eval' : program' -> value' -> Prop} + {eval'' : program'' -> value'' -> Prop} + {eval''' : program''' -> value''' -> Prop} + (o : t program program' value value' eval eval') + (o' : t program' program'' value' value'' eval' eval'') + (o'' : t program'' program''' value'' value''' eval'' eval''') + (prec : forall p : program', post o p -> pre o' p) + (prec' : forall p : program'', post o' p -> pre o'' p) : + forall x p1, + transform (compose o (compose o' o'' prec') prec) x p1 = + transform (compose (compose o o' prec) o'' prec') x p1. +Proof. + cbn. intros. + unfold run, time. + f_equal. f_equal. + apply proof_irrelevance. +Qed. + +Lemma obseq_compose_assoc + {program program' program'' program''' value value' value'' value''' : Type} + {eval : program -> value -> Prop} {eval' : program' -> value' -> Prop} + {eval'' : program'' -> value'' -> Prop} + {eval''' : program''' -> value''' -> Prop} + (o : t program program' value value' eval eval') + (o' : t program' program'' value' value'' eval' eval'') + (o'' : t program'' program''' value'' value''' eval'' eval''') + (prec : forall p : program', post o p -> pre o' p) + (prec' : forall p : program'', post o' p -> pre o'' p) : + forall x p1 p2 v1 v2, obseq (compose o (compose o' o'' prec') prec) x p1 p2 v1 v2 <-> + obseq (compose (compose o o' prec) o'' prec') x p1 p2 v1 v2. +Proof. + cbn. intros. + unfold run, time. + intros. firstorder. exists x1. split. + exists x0. split => //. + assert (correctness o' (transform o x p1) + (prec (transform o x p1) (correctness o x p1)) = + (Transform.Transform.compose_obligation_1 o o' prec x p1)). apply proof_irrelevance. + now rewrite -H. + + exists x1. split => //. + exists x0. split => //. + assert (correctness o' (transform o x p1) + (prec (transform o x p1) (correctness o x p1)) = + (Transform.Transform.compose_obligation_1 o o' prec x p1)). apply proof_irrelevance. + now rewrite H. +Qed. + +Import EEnvMap.GlobalContextMap. +Lemma make_irrel Σ fr fr' : EEnvMap.GlobalContextMap.make Σ fr = EEnvMap.GlobalContextMap.make Σ fr'. +Proof. + unfold make. f_equal. + apply proof_irrelevance. +Qed. + +Lemma eval_value {efl : WcbvFlags} Σ v v' : + value Σ v -> eval Σ v v' -> v = v'. +Proof. + intros isv ev. + now pose proof (eval_deterministic ev (value_final _ isv)). +Qed. + +Ltac destruct_compose := + match goal with + |- context [ transform (compose ?x ?y ?pre) ?p ?pre' ] => + let pre'' := fresh in + let H := fresh in + destruct (transform_compose x y pre p pre') as [pre'' H]; + rewrite H; clear H; revert pre'' + (* rewrite H'; clear H'; *) + (* revert pre'' *) + end. + +Ltac destruct_compose_no_clear := + match goal with + |- context [ transform (compose ?x ?y ?pre) ?p ?pre' ] => + let pre'' := fresh in + let H := fresh in + destruct (transform_compose x y pre p pre') as [pre'' H]; + rewrite H; revert pre'' H + end. + +(* +Section TransformValue. + Context {program program' : Type}. + Context {value value' : Type}. + Context {eval : program -> value -> Prop}. + Context {eval' : program' -> value' -> Prop}. + Context (t : Transform.t program program' value value' eval eval'). + + Lemma preserves_value p : value p.1 p.2 (transform t p) + + Definition preserves_eval pre (transform : forall p : program, pre p -> program') + (obseq : forall p : program, pre p -> program' -> value -> value' -> Prop) := + forall p v (pr : pre p), + eval p v -> + let p' := transform p pr in + exists v', eval' p' v' /\ obseq p pr p' v v'. + + Record t := + { name : string; + +Lemma transform_value *) + + +Inductive is_construct_app : EAst.term -> Prop := +| is_cstr_app_cstr kn c args : Forall is_construct_app args -> is_construct_app (EAst.tConstruct kn c args) +| is_cstr_app_app f a : is_construct_app f -> is_construct_app a -> is_construct_app (EAst.tApp f a). + +Section lambdabox_theorem. + + Context (Σ Σ' : EEnvMap.GlobalContextMap.t) (v : EAst.term). + + Context (p : pre verified_lambdabox_pipeline (Σ, v)). + Context (p' : pre verified_lambdabox_pipeline (Σ', v)). + Context (is_value : value (wfl := default_wcbv_flags) Σ v). + + Lemma pres : extends_eprogram_env (Σ, v) (Σ', v) -> + extends_eprogram (transform verified_lambdabox_pipeline (Σ, v) p) + (transform verified_lambdabox_pipeline (Σ', v) p'). + Proof. + epose proof (pres := verified_lambdabox_pipeline_extends). + red in pres. specialize (pres _ _ p p'). auto. + Qed. + + (* Final evaluation flags *) + Definition evflags := {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}. + + Lemma pres_firstorder_value : + is_construct_app v -> + is_construct_app (transform verified_lambdabox_pipeline (Σ, v) p).2. + Proof. + intros isapp. + destruct (preservation verified_lambdabox_pipeline (Σ, v) v p) as [v' [[ev] obs]]. + { red. cbn. sq. eapply value_final, is_value. } + set (transp := transform _ _ p) in *. + assert (value (wfl := evflags) transp.1 transp.2). admit. + eapply eval_value in ev => //. subst v'. + clear -obs isapp. + unfold verified_lambdabox_pipeline in obs. + cbn [obseq compose] in obs. + unfold run, time in obs. + decompose [ex and prod] obs. clear obs. subst. + cbn [obseq compose verified_lambdabox_pipeline] in *. + cbn [obseq compose constructors_as_blocks_transformation] in *. + cbn [obseq run compose rebuild_wf_env_transform] in *. + cbn [obseq compose inline_projections_optimization] in *. + cbn [obseq compose remove_match_on_box_trans] in *. + cbn [obseq compose remove_params_optimization] in *. + cbn [obseq compose guarded_to_unguarded_fix] in *. + cbn [obseq compose verified_lambdabox_pipeline] in *. + subst. + cbn [transform rebuild_wf_env_transform] in *. + cbn [transform constructors_as_blocks_transformation] in *. + cbn [transform inline_projections_optimization] in *. + cbn [transform remove_match_on_box_trans] in *. + cbn [transform remove_params_optimization] in *. + cbn [transform guarded_to_unguarded_fix] in *. + clearbody transp. revert b. intros ->. clear transp. + induction isapp. + cbn in *. constructor. constructor. + Admitted. + + +End lambdabox_theorem. + + +Lemma rebuild_wf_env_irr {efl : EWellformed.EEnvFlags} p wf p' wf' : + p.1 = p'.1 -> + (rebuild_wf_env p wf).1 = (rebuild_wf_env p' wf').1. +Proof. + destruct p as [], p' as []. + cbn. intros <-. + unfold make. f_equal. apply proof_irrelevance. +Qed. + +Lemma obseq_lambdabox (Σt Σ'v : EProgram.eprogram_env) pr pr' p' v' : + EGlobalEnv.extends Σ'v.1 Σt.1 -> + obseq verified_lambdabox_pipeline Σt pr p' Σ'v.2 v' -> + (transform verified_lambdabox_pipeline Σ'v pr').2 = v'. +Proof. + intros ext obseq. + destruct Σt as [Σ t], Σ'v as [Σ' v]. + pose proof verified_lambdabox_pipeline_extends. + red in H. + assert (pr'' : pre verified_lambdabox_pipeline (Σ, v)). + { clear -pr pr' ext. destruct pr as [[] ?], pr' as [[] ?]. + split. red; cbn. split => //. + eapply EWellformed.extends_wellformed; tea. + split. apply H1. cbn. destruct H4; cbn in *. + eapply EEtaExpandedFix.isEtaExp_expanded. + eapply EEtaExpandedFix.isEtaExp_extends; tea. + now eapply EEtaExpandedFix.expanded_isEtaExp. } + destruct (H _ _ pr' pr'') as [ext' ->]. + split => //. + clear H. + move: obseq. + unfold verified_lambdabox_pipeline. + repeat destruct_compose. + cbn [transform rebuild_wf_env_transform] in *. + cbn [transform constructors_as_blocks_transformation] in *. + cbn [transform inline_projections_optimization] in *. + cbn [transform remove_match_on_box_trans] in *. + cbn [transform remove_params_optimization] in *. + cbn [transform guarded_to_unguarded_fix] in *. + intros ? ? ? ? ? ? ?. + unfold run, time. + cbn [obseq compose constructors_as_blocks_transformation] in *. + cbn [obseq run compose rebuild_wf_env_transform] in *. + cbn [obseq compose inline_projections_optimization] in *. + cbn [obseq compose remove_match_on_box_trans] in *. + cbn [obseq compose remove_params_optimization] in *. + cbn [obseq compose guarded_to_unguarded_fix] in *. + intros obs. + decompose [ex and prod] obs. clear obs. subst. + unfold run, time. + unfold transform_blocks_program. cbn [snd]. f_equal. + repeat destruct_compose. + intros. + cbn [transform rebuild_wf_env_transform] in *. + cbn [transform constructors_as_blocks_transformation] in *. + cbn [transform inline_projections_optimization] in *. + cbn [transform remove_match_on_box_trans] in *. + cbn [transform remove_params_optimization] in *. + cbn [transform guarded_to_unguarded_fix] in *. + eapply rebuild_wf_env_irr. + unfold EInlineProjections.optimize_program. cbn [fst snd]. + f_equal. + eapply rebuild_wf_env_irr. + unfold EOptimizePropDiscr.remove_match_on_box_program. cbn [fst snd]. + f_equal. + now eapply rebuild_wf_env_irr. +Qed. + +Lemma extends_erase_pcuic_program {guard : abstract_guard_impl} (Σ : global_env_ext_map) t t' nin nin' nin0 nin0' wf wf' ty ty' : + PCUICWcbvEval.eval Σ t' t -> + EGlobalEnv.extends + (@erase_pcuic_program guard (Σ, t) nin nin' wf ty).1 + (@erase_pcuic_program guard (Σ, t') nin0 nin0' wf' ty').1. +Proof. + intros ev. + cbn. + unfold ErasureFunction.erase_global_fast. + set (eg := ErasureFunction.erase _ _ _ _ _). + set (eg' := ErasureFunction.erase _ _ _ _ _). + erewrite ErasureFunction.erase_global_deps_fast_spec. + erewrite ErasureFunction.erase_global_deps_fast_spec. + epose proof ErasureFunction.erase_global_decls_irr. +Admitted. + +Section pipeline_theorem. + + Fixpoint compile_value_box (t : PCUICAst.term) (acc : list EAst.term) : EAst.term := + match t with + | PCUICAst.tApp f a => compile_value_box f (acc ++ [compile_value_box a []]) + | PCUICAst.tConstruct i n _ => EAst.tConstruct i n acc + | _ => EAst.tVar "error" + end. + + Instance cf : checker_flags := extraction_checker_flags. + Instance nf : PCUICSN.normalizing_flags := PCUICSN.extraction_normalizing. + + Variable Σ : global_env_ext_map. + Variable HΣ : PCUICTyping.wf_ext Σ. + Variable expΣ : PCUICEtaExpand.expanded_global_env Σ.1. + + Variable t : PCUICAst.term. + Variable expt : PCUICEtaExpand.expanded Σ.1 [] t. + + Variable v : PCUICAst.term. + + Variable i : Kernames.inductive. + Variable u : Universes.Instance.t. + Variable args : list PCUICAst.term. + + Variable typing : PCUICTyping.typing Σ [] t (PCUICAst.mkApps (PCUICAst.tInd i u) args). + + Variable fo : @PCUICFirstorder.firstorder_ind Σ (PCUICFirstorder.firstorder_env Σ) i. + + Variable Normalisation : PCUICSN.NormalizationIn Σ. + + Lemma precond : pre verified_erasure_pipeline (Σ, t). + Proof. + hnf. repeat eapply conj; sq; cbn; eauto. + - red. cbn. eauto. + - todo "normalization". + - todo "normalization". + Qed. + + Variable Heval : ∥PCUICWcbvEval.eval Σ t v∥. + + Lemma precond2 : pre verified_erasure_pipeline (Σ, v). + Proof. + cbn. repeat eapply conj; sq; cbn; eauto. + - red. cbn. split; eauto. + eexists. + eapply PCUICClassification.subject_reduction_eval; eauto. + - todo "preservation of eta expandedness". + - todo "normalization". + - todo "normalization". + Qed. + + Let Σ_t := (transform verified_erasure_pipeline (Σ, t) precond).1. + Let t_t := (transform verified_erasure_pipeline (Σ, t) precond).2. + Let v_t := compile_value_box v []. + + Lemma fo_v : PCUICFirstorder.firstorder_value Σ [] v. + Proof. + sq. + eapply PCUICFirstorder.firstorder_value_spec; eauto. + - eapply PCUICClassification.subject_reduction_eval; eauto. + - eapply PCUICWcbvEval.eval_to_value; eauto. + Qed. + + + Lemma v_t_spec : v_t = (transform verified_erasure_pipeline (Σ, v) precond2).2. + Proof. + unfold v_t. generalize fo_v, precond2. clear. + induction 1. + intros. unfold verified_erasure_pipeline. + repeat destruct_compose. intros. + cbn [transform erase_transform]. + cbn [transform pcuic_expand_lets_transform]. + Admitted. + + Import PCUICWfEnv. + + + + Lemma verified_erasure_pipeline_theorem : + ∥ eval (wfl := extraction_wcbv_flags) Σ_t t_t v_t∥. + Proof. + hnf. + pose proof (preservation verified_erasure_pipeline (Σ, t)) as Hcorr. + unshelve eapply Hcorr in Heval as Hev. eapply precond. + destruct Hev as [v' [[H1] H2]]. + move: H2. + + (* repeat match goal with + [ H : obseq _ _ _ _ _ |- _ ] => hnf in H ; decompose [ex and prod] H ; subst + end. *) + rewrite v_t_spec. + subst v_t Σ_t t_t. + revert H1. + unfold verified_erasure_pipeline. + intros. + revert H1 H2. clear Hcorr. + intros ev obs. + cbn [obseq compose] in obs. + unfold run, time in obs. + decompose [ex and prod] obs. clear obs. + subst. + cbn [obseq compose erase_transform] in *. + cbn [obseq compose pcuic_expand_lets_transform] in *. + subst. + move: ev b. + repeat destruct_compose. + intros. + move: b. + cbn [transform rebuild_wf_env_transform] in *. + cbn [transform constructors_as_blocks_transformation] in *. + cbn [transform inline_projections_optimization] in *. + cbn [transform remove_match_on_box_trans] in *. + cbn [transform remove_params_optimization] in *. + cbn [transform guarded_to_unguarded_fix] in *. + cbn [transform erase_transform] in *. + cbn [transform compose pcuic_expand_lets_transform] in *. + unfold run, time. + cbn [obseq compose constructors_as_blocks_transformation] in *. + cbn [obseq run compose rebuild_wf_env_transform] in *. + cbn [obseq compose inline_projections_optimization] in *. + cbn [obseq compose remove_match_on_box_trans] in *. + cbn [obseq compose remove_params_optimization] in *. + cbn [obseq compose guarded_to_unguarded_fix] in *. + cbn [obseq compose erase_transform] in *. + cbn [obseq compose pcuic_expand_lets_transform] in *. + cbn [transform compose pcuic_expand_lets_transform] in *. + cbn [transform erase_transform] in *. + destruct Heval. + pose proof typing as typing'. + eapply PCUICClassification.subject_reduction_eval in typing'; tea. + eapply PCUICExpandLetsCorrectness.pcuic_expand_lets in typing'. + rewrite PCUICExpandLetsCorrectness.trans_mkApps /= in typing'. + destruct H1. + (* pose proof (abstract_make_wf_env_ext) *) + unfold PCUICExpandLets.expand_lets_program. + set (em := build_global_env_map _). + unfold erase_program. + set (f := map_squash _ _). cbn in f. + destruct H. destruct s as [[]]. + set (wfe := build_wf_env_from_env em (map_squash (PCUICTyping.wf_ext_wf (em, Σ.2)) (map_squash fst (conj (sq (w0, s)) a).p1))). + destruct Heval. + eapply (ErasureFunction.firstorder_erases_deterministic optimized_abstract_env_impl wfe Σ.2) in b0. 3:tea. + 2:{ cbn. reflexivity. } + 2:{ eapply PCUICExpandLetsCorrectness.trans_wcbveval. eapply PCUICWcbvEval.eval_closed; tea. apply HΣ. + admit. + eapply PCUICWcbvEval.value_final. now eapply PCUICWcbvEval.eval_to_value in X0. } + 2:{ clear -fo. admit. } + 2:{ apply HΣ. } + 2:{ apply PCUICExpandLetsCorrectness.trans_wf, HΣ. } + rewrite b0. + intros obs. + constructor. + match goal with [ H1 : eval _ _ ?v1 |- eval _ _ ?v2 ] => enough (v2 = v1) as -> by exact ev end. + eapply obseq_lambdabox; revgoals. + unfold erase_pcuic_program. cbn [fst snd]. exact obs. + clear obs b0 ev e w. + eapply extends_erase_pcuic_program. cbn. + eapply (PCUICExpandLetsCorrectness.trans_wcbveval (Σ := (Σ.1, Σ.2))). admit. exact X0. + Admitted. + + Lemma verified_erasure_pipeline_lambda : + PCUICAst.isLambda t -> EAst.isLambda t_t. + Proof. + unfold t_t. clear. + Admitted. + +End pipeline_theorem. From 71819e7bc59d814386abdf74958f56b735da29ce Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 29 Aug 2023 18:17:51 +0200 Subject: [PATCH 42/61] Main argument of the erasure with dependencies proof --- erasure-plugin/theories/ErasureCorrectness.v | 5 +- erasure/theories/ErasureFunction.v | 918 ++++++++++++++++--- 2 files changed, 800 insertions(+), 123 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 01c4ac15e..437a1ab61 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -271,11 +271,12 @@ Proof. intros ev. cbn. unfold ErasureFunction.erase_global_fast. - set (eg := ErasureFunction.erase _ _ _ _ _). - set (eg' := ErasureFunction.erase _ _ _ _ _). + set (et := ErasureFunction.erase _ _ _ _ _). + set (et' := ErasureFunction.erase _ _ _ _ _). erewrite ErasureFunction.erase_global_deps_fast_spec. erewrite ErasureFunction.erase_global_deps_fast_spec. epose proof ErasureFunction.erase_global_decls_irr. + intros kn decl. Admitted. Section pipeline_theorem. diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 5edab27a3..3d4b6710e 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -877,7 +877,7 @@ Definition iter {A} (f : A -> A) : nat -> (A -> A) end. (* we use the [match] trick to get typeclass resolution to pick up the right instances without leaving any evidence in the resulting term, and without having to pass them manually everywhere *) -Notation NormalizationIn_erase_global_decls X decls +Notation NormalizationIn_erase_global_deps X decls := (match extraction_checker_flags, extraction_normalizing return _ with | cf, no => forall n, @@ -892,8 +892,8 @@ Notation NormalizationIn_erase_global_decls X decls Lemma normalization_in_inv {X_type : abstract_env_impl} (X : X_type.π1) (decls : global_declarations) kn cb - (normalization_in : NormalizationIn_erase_global_decls X ((kn, ConstantDecl cb) :: decls)) : - NormalizationIn_erase_global_decls (abstract_pop_decls X) decls. + (normalization_in : NormalizationIn_erase_global_deps X ((kn, ConstantDecl cb) :: decls)) : + NormalizationIn_erase_global_deps (abstract_pop_decls X) decls. Proof. cbv [id]. intros. unshelve eapply (normalization_in (S n)); cbn; tea. lia. @@ -992,9 +992,9 @@ Proof. now rewrite H1 H2. } Qed. -Program Fixpoint erase_global_decls +Program Fixpoint erase_global_deps {X_type : abstract_env_impl} (deps : KernameSet.t) (X : X_type.π1) (decls : global_declarations) - {normalization_in : NormalizationIn_erase_global_decls X decls} + {normalization_in : NormalizationIn_erase_global_deps X decls} (prop : forall Σ : global_env, abstract_env_rel X Σ -> Σ.(declarations) = decls) : E.global_declarations * KernameSet.t := match decls with @@ -1004,18 +1004,18 @@ Program Fixpoint erase_global_decls if KernameSet.mem kn deps then let cb' := @erase_constant_decl X_type X' cb _ _ in let deps := KernameSet.union deps (snd cb') in - let X'' := erase_global_decls deps X' decls _ in + let X'' := erase_global_deps deps X' decls _ in (((kn, E.ConstantDecl (fst cb')) :: fst X''), snd X'') else - erase_global_decls deps X' decls _ + erase_global_deps deps X' decls _ | (kn, InductiveDecl mib) :: decls => let X' := abstract_pop_decls X in if KernameSet.mem kn deps then let mib' := erase_mutual_inductive_body mib in - let X'':= erase_global_decls deps X' decls _ in + let X'':= erase_global_deps deps X' decls _ in (((kn, E.InductiveDecl mib') :: fst X''), snd X'') - else erase_global_decls deps X' decls _ + else erase_global_deps deps X' decls _ end. Next Obligation. cbv [id]. @@ -1071,12 +1071,66 @@ assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(de now pose proof (abstract_pop_decls_correct X decls prop' _ _ HX H). Qed. -Lemma erase_global_decls_irr +Program Fixpoint erase_global {X_type : abstract_env_impl} (X : X_type.π1) (decls : global_declarations) + {normalization_in : NormalizationIn_erase_global_deps X decls} + (prop : forall Σ : global_env, abstract_env_rel X Σ -> Σ.(declarations) = decls) + : E.global_declarations := + match decls with + | [] => [] + | (kn, ConstantDecl cb) :: decls => + let X' := abstract_pop_decls X in + let cb' := @erase_constant_decl X_type X' cb _ _ in + let X'' := erase_global X' decls _ in + ((kn, E.ConstantDecl (fst cb')) :: X'') + + | (kn, InductiveDecl mib) :: decls => + let X' := abstract_pop_decls X in + let mib' := erase_mutual_inductive_body mib in + let X'':= erase_global X' decls _ in + ((kn, E.InductiveDecl mib') :: X'') + end. +Next Obligation. + cbv [id]. + unshelve eapply (normalization_in 0); tea; cbn; try reflexivity; try lia. +Defined. +Next Obligation. + pose proof (abstract_env_exists X) as [[? HX]]. + pose proof (abstract_env_wf _ HX) as [wfX]. + assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). + { now eexists. } + pose proof (abstract_pop_decls_correct X decls prop' _ _ HX H). + destruct x, Σ as [[] u], H0; cbn in *. + subst. sq. inversion H1. subst. destruct wfX. cbn in *. + specialize (prop _ HX). + cbn in prop. + rewrite prop in o0. depelim o0. destruct o1. + split. 2:exact on_global_decl_d. split => //. +Qed. +Next Obligation. + unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. +Defined. +Next Obligation. + pose proof (abstract_env_exists X) as [[? HX]]. + assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). + { now eexists. } + now pose proof (abstract_pop_decls_correct X decls prop' _ _ HX H). +Qed. +Next Obligation. + unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. +Defined. +Next Obligation. +pose proof (abstract_env_exists X) as [[? HX]]. +assert (prop': forall Σ : global_env, abstract_env_rel X Σ -> exists d, Σ.(declarations) = d :: decls). +{ now eexists. } +now pose proof (abstract_pop_decls_correct X decls prop' _ _ HX H). +Qed. + +Lemma erase_global_deps_irr X_type deps (X:X_type.π1) decls - {normalization_in normalization_in' : NormalizationIn_erase_global_decls X decls} + {normalization_in normalization_in' : NormalizationIn_erase_global_deps X decls} prf prf' : - erase_global_decls deps X decls (normalization_in:=normalization_in) prf = - erase_global_decls deps X decls (normalization_in:=normalization_in') prf'. + erase_global_deps deps X decls (normalization_in:=normalization_in) prf = + erase_global_deps deps X decls (normalization_in:=normalization_in') prf'. Proof. revert X deps normalization_in normalization_in' prf prf'. induction decls; eauto. intros. destruct a as [kn []]. @@ -1616,7 +1670,7 @@ Lemma erase_global_includes X_type (X:X_type.π1) deps decls {normalization_in} forall Σ : global_env, abstract_env_rel X Σ -> ∥ ∑ decl, lookup_env Σ d = Some decl ∥) -> KernameSet.subset deps' deps -> forall Σ : global_env, abstract_env_rel X Σ -> - let Σ' := erase_global_decls deps X decls (normalization_in:=normalization_in) prf in + let Σ' := erase_global_deps deps X decls (normalization_in:=normalization_in) prf in includes_deps Σ (fst Σ') deps'. Proof. intros ? sub Σ wfΣ; cbn. induction decls in X, H, sub, prf, deps, deps', Σ , wfΣ, normalization_in |- *. @@ -1629,7 +1683,7 @@ Proof. eapply KernameSet.subset_spec in sub. edestruct (H i hin) as [[decl Hdecl]]; eauto. unfold lookup_env in Hdecl. pose proof (eqb_spec i kn). - rewrite (prf _ wfΣ) in Hdecl. move: Hdecl. cbn -[erase_global_decls]. + rewrite (prf _ wfΣ) in Hdecl. move: Hdecl. cbn -[erase_global_deps]. elim: H0. intros -> [= <-]. * destruct d as [|]; [left|right]. { cbn. set (Xpop := abstract_pop_decls X). @@ -1775,7 +1829,7 @@ Lemma erase_correct (wfl := Ee.default_wcbv_flags) X_type (X : X_type.π1) forall wt : forall Σ, Σ ∼_ext Xext -> welltyped Σ [] t, erase X_type Xext [] t wt = t' -> KernameSet.subset (term_global_deps t') deps -> - erase_global_decls deps X decls (normalization_in:=normalization_in) prf = Σ' -> + erase_global_deps deps X decls (normalization_in:=normalization_in) prf = Σ' -> (forall Σ : global_env, abstract_env_rel X Σ -> Σ |-p t ⇓ v) -> forall Σ : global_env_ext, abstract_env_ext_rel Xext Σ -> exists v', Σ ;;; [] |- v ⇝ℇ v' /\ ∥ fst Σ' ⊢ t' ⇓ v' ∥. @@ -1911,22 +1965,22 @@ Qed. Local Hint Constructors expanded : expanded. -Local Arguments erase_global_decls _ _ _ _ _ : clear implicits. +Local Arguments erase_global_deps _ _ _ _ _ : clear implicits. (*Lemma lookup_env_erase X_type X deps decls normalization_in prf kn d : KernameSet.In kn deps -> forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = Some (InductiveDecl d) -> - EGlobalEnv.lookup_env (erase_global_decls X_type deps X decls normalization_in prf) kn = Some (EAst.InductiveDecl (erase_mutual_inductive_body d)). + EGlobalEnv.lookup_env (erase_global_deps X_type deps X decls normalization_in prf) kn = Some (EAst.InductiveDecl (erase_mutual_inductive_body d)). Proof. intros hin Σ wfΣ. pose proof (prf _ wfΣ). unfold lookup_env. rewrite H. clear H. rewrite /lookup_env. induction decls in X, Σ , wfΣ ,prf, deps, hin, normalization_in |- *. - move=> /= //. - destruct a as [kn' d']. - cbn -[erase_global_decls]. + cbn -[erase_global_deps]. case: (eqb_spec kn kn'); intros e'; subst. intros [= ->]. - unfold erase_global_decls. + unfold erase_global_deps. eapply KernameSet.mem_spec in hin. rewrite hin /=. now rewrite eq_kername_refl. intros hl. destruct d'. simpl. @@ -1949,7 +2003,7 @@ Qed. Lemma lookup_env_erase_decl_None X_type X deps decls normalization_in prf kn : forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = None -> - EGlobalEnv.lookup_env (erase_global_decls X_type deps X decls normalization_in prf).1 kn = None. + EGlobalEnv.lookup_env (erase_global_deps X_type deps X decls normalization_in prf).1 kn = None. Proof. intros Σ wfΣ. pose proof (prf _ wfΣ). @@ -1958,7 +2012,7 @@ Proof. induction decls in X, Σ , wfΣ ,prf, deps, normalization_in |- *. - move=> /= //. - destruct a as [kn' d']. - cbn -[erase_global_decls]. + cbn -[erase_global_deps]. case: (eqb_spec kn kn'); intros e'; subst. intros [= ->]. cbn. destruct d'. cbn. @@ -1985,10 +2039,10 @@ Lemma lookup_env_erase_global_diff X_type X deps decls kn kn' d' normalization_i kn <> kn' -> exists deps' nin' prf', KernameSet.Subset deps deps' /\ - EGlobalEnv.lookup_env (erase_global_decls X_type deps X ((kn', d') :: decls) normalization_in prf).1 kn = - EGlobalEnv.lookup_env (erase_global_decls X_type deps' (abstract_pop_decls X) decls nin' prf').1 kn /\ - (KernameSet.In kn (erase_global_decls X_type deps X ((kn', d') :: decls) normalization_in prf).2 -> - KernameSet.In kn (erase_global_decls X_type deps' (abstract_pop_decls X) decls nin' prf').2). + EGlobalEnv.lookup_env (erase_global_deps X_type deps X ((kn', d') :: decls) normalization_in prf).1 kn = + EGlobalEnv.lookup_env (erase_global_deps X_type deps' (abstract_pop_decls X) decls nin' prf').1 kn /\ + (KernameSet.In kn (erase_global_deps X_type deps X ((kn', d') :: decls) normalization_in prf).2 -> + KernameSet.In kn (erase_global_deps X_type deps' (abstract_pop_decls X) decls nin' prf').2). Proof. intros hd. simpl. destruct d'. @@ -2037,15 +2091,15 @@ Proof. now eapply lookup_env_In_map_fst in hl. Qed. -Lemma in_erase_global_decls_acc X_type X deps decls kn normalization_in prf : - KernameSet.In kn (erase_global_decls X_type deps X decls normalization_in prf).2 -> +Lemma in_erase_global_deps_acc X_type X deps decls kn normalization_in prf : + KernameSet.In kn (erase_global_deps X_type deps X decls normalization_in prf).2 -> KernameSet.In kn deps \/ In kn (map fst decls). Proof. induction decls in X, prf, deps, normalization_in |- *. cbn; auto. destruct a as [kn' []]. + cbn. set (ec := erase_constant_decl _ _ _ _). - set (eg := erase_global_decls _ _ _ _ _ _). - set (eg' := erase_global_decls _ _ _ _ _ _). + set (eg := erase_global_deps _ _ _ _ _ _). + set (eg' := erase_global_deps _ _ _ _ _ _). case_eq (KernameSet.mem kn' deps). * cbn; intros. specialize (IHdecls _ _ _ _ H0). destruct IHdecls as [Hu|Hm]; auto. @@ -2060,8 +2114,8 @@ Proof. now destruct H2 as [<- _]. * intros hdeps hin. eapply IHdecls in hin. intuition auto. - + cbn; set (eg := erase_global_decls _ _ _ _ _ _); - set (eg' := erase_global_decls _ _ _ _ _ _). + + cbn; set (eg := erase_global_deps _ _ _ _ _ _); + set (eg' := erase_global_deps _ _ _ _ _ _). case_eq (KernameSet.mem kn' deps). * cbn; intros. specialize (IHdecls _ _ _ _ H0). destruct IHdecls as [Hu|Hm]; auto. @@ -2077,7 +2131,7 @@ Qed. Lemma lookup_env_erase_decl X_type X deps decls normalization_in prf kn decl : forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = Some decl -> - let er := erase_global_decls X_type deps X decls normalization_in prf in + let er := erase_global_deps X_type deps X decls normalization_in prf in KernameSet.In kn er.2 -> match decl with | ConstantDecl cb => @@ -2096,13 +2150,13 @@ Proof. induction decls in X, Σ , wfΣ ,prf, deps, decl, normalization_in |- *. - move=> /= //. - destruct a as [kn' d']. - cbn -[erase_global_decls]. + cbn -[erase_global_deps]. case: (eqb_spec kn kn'); intros e'; subst. intros [= ->]. destruct decl. + cbn; set (ec := erase_constant_decl _ _ _ _); - set (eg := erase_global_decls _ _ _ _ _ _); - set (eg' := erase_global_decls _ _ _ _ _ _). + set (eg := erase_global_deps _ _ _ _ _ _); + set (eg' := erase_global_deps _ _ _ _ _ _). case_eq (KernameSet.mem kn' deps). move/KernameSet.mem_spec => hin'; cbn. ++ intros hin. cbn; rewrite eqb_refl; @@ -2112,7 +2166,7 @@ Proof. destruct H as [? []]. unfold strictly_extends_decls. rewrite H. sq; split => //. clear ec eg eg' hin. specialize (prf _ hx). rewrite prf. now eexists [_]. ++ intros kn'deps kn'eg'. - eapply in_erase_global_decls_acc in kn'eg'. + eapply in_erase_global_deps_acc in kn'eg'. destruct kn'eg'. eapply KernameSet.mem_spec in H. congruence. elimtype False. clear -prf wfΣ H. specialize (prf _ wfΣ). @@ -2120,14 +2174,14 @@ Proof. rewrite prf in X0. depelim X0. apply EnvMap.EnvMap.fresh_global_iff_not_In in H0. contradiction. + cbn; - set (eg := erase_global_decls _ _ _ _ _ _); - set (eg' := erase_global_decls _ _ _ _ _ _). + set (eg := erase_global_deps _ _ _ _ _ _); + set (eg' := erase_global_deps _ _ _ _ _ _). case_eq (KernameSet.mem kn' deps). move/KernameSet.mem_spec => hin'; cbn. ++ intros hin. cbn; rewrite eqb_refl; do 3 eexists. ++ intros kn'deps kn'eg'. - eapply in_erase_global_decls_acc in kn'eg'. + eapply in_erase_global_deps_acc in kn'eg'. destruct kn'eg'. eapply KernameSet.mem_spec in H. congruence. elimtype False. clear -prf wfΣ H. specialize (prf _ wfΣ). @@ -2187,9 +2241,9 @@ Proof. eapply wf_fresh_globals in X. now rewrite -e0. Qed. -Lemma in_deps_in_erase_global_decls X_type X deps decls normalization_in prf kn : +Lemma in_deps_in_erase_global_deps X_type X deps decls normalization_in prf kn : KernameSet.In kn deps -> - KernameSet.In kn (erase_global_decls X_type deps X decls normalization_in prf).2. + KernameSet.In kn (erase_global_deps X_type deps X decls normalization_in prf).2. Proof. induction decls in X, deps, normalization_in, prf |- *. - intros. now cbn. @@ -2200,8 +2254,8 @@ Proof. unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ H' wfpop). intros. rewrite prf => //. now eexists. destruct H0 as [Hd [Hu Hr]]. destruct a. destruct g. - + set (eg := erase_global_decls _ _ _ _ _ _); - set (eg' := erase_global_decls _ _ _ _ _ _). + + set (eg := erase_global_deps _ _ _ _ _ _); + set (eg' := erase_global_deps _ _ _ _ _ _). case_eq (KernameSet.mem k deps). * cbn. intros hm. destruct (eqb_spec k kn). @@ -2209,8 +2263,8 @@ Proof. ++ eapply IHdecls. KernameSetDecide.fsetdec. * intros hnin. eapply IHdecls; tea. - + set (eg := erase_global_decls _ _ _ _ _ _); - set (eg' := erase_global_decls _ _ _ _ _ _). + + set (eg := erase_global_deps _ _ _ _ _ _); + set (eg' := erase_global_deps _ _ _ _ _ _). case_eq (KernameSet.mem k deps). * cbn. intros hm. destruct (eqb_spec k kn). @@ -2220,22 +2274,357 @@ Proof. * intros hnin. eapply IHdecls; tea. Qed. +Definition constant_decls_deps decl := + match decl with + | {| EAst.cst_body := Some b |} => term_global_deps b + | _ => KernameSet.empty + end. -Lemma erase_global_decls_eval {efl : EWcbvEval.WcbvFlags} (X_type : abstract_env_impl) (X:X_type.π1) ext wfX {normalization_in} decls nin prf Γ t wt v wv : - let Xext := abstract_make_wf_env_ext X ext wfX in - let t' := erase X_type Xext (normalization_in := normalization_in) Γ t wt in - let v' := erase X_type Xext (normalization_in := normalization_in) Γ v wv in - let eg := erase_global_decls X_type (term_global_deps t') X decls nin prf in - let eg' := erase_global_decls X_type (term_global_deps v') X decls nin prf in - eg.1 ⊢ t' ⇓ v' -> KernameSet.Subset eg'.2 eg.2. +Fixpoint filter_deps deps decls := + match decls with + | [] => [] + | (kn, decl) :: decls => + if KernameSet.mem kn deps then + match decl with + | EAst.ConstantDecl cst => + (kn, decl) :: filter_deps (KernameSet.union deps (constant_decls_deps cst)) decls + | _ => (kn, decl) :: filter_deps deps decls + end + else filter_deps deps decls + end. + +Lemma erase_global_deps_erase_global {efl : EWcbvEval.WcbvFlags} (X_type : abstract_env_impl) (X:X_type.π1) deps decls {normalization_in} prf : + (@erase_global_deps X_type deps X decls normalization_in prf).1 = + filter_deps deps (@erase_global X_type X decls normalization_in prf). Proof. - intros. clearbody t' v'. + symmetry. + generalize normalization_in at 1. + generalize prf at 1. + intros prf' normalization_in'. + induction decls in normalization_in, normalization_in', prf, prf', X, deps |- *. + - cbn. auto. + - destruct a as [kn []]. + cbn. + + destruct (KernameSet.mem kn deps). + destruct c as [? [] ? ?]. + set (ec := erase_constant_decl _ _ _ _). + set (ec' := erase_constant_decl _ _ _ _). + assert (ec = ec') as <-. + { unfold ec, ec'. eapply erase_constant_decl_irrel. red. cbn. + intros. pose proof (abstract_env_irr _ H H0). subst. red. split; auto. congruence. } + cbn; f_equal. set (d := KernameSet.union _ _). eapply IHdecls. + cbn. f_equal. eapply IHdecls. + eapply IHdecls. + + cbn; destruct (KernameSet.mem kn deps). + cbn. f_equal. eapply IHdecls. + eapply IHdecls. +Qed. + +Lemma knset_mem_spec kn s : reflect (KernameSet.In kn s) (KernameSet.mem kn s). +Proof. + destruct (KernameSet.mem_spec s kn). + destruct (KernameSet.mem kn s). + - now constructor. + - constructor. intros hin. now apply H0 in hin. +Qed. + +Definition global_decl_deps decl := + match decl with + | EAst.ConstantDecl cst => constant_decls_deps cst + | _ => KernameSet.empty + end. + +Fixpoint global_deps (Σ : EAst.global_declarations) deps := + match Σ with + | [] => deps + | (kn, decl) :: decls => + if KernameSet.mem kn deps then + global_deps decls + (KernameSet.union deps (global_decl_deps decl)) + else global_deps decls deps + end. + +Infix "=kn" := KernameSet.Equal (at level 90). + +Ltac knset := KernameSetDecide.fsetdec. + +Lemma global_deps_union Σ deps deps' : + global_deps Σ (KernameSet.union deps deps') =kn + KernameSet.union (global_deps Σ deps) (global_deps Σ deps'). +Proof. + induction Σ in deps, deps' |- *; auto. + - reflexivity. + - destruct a. cbn -[KernameSet.mem]. + destruct (knset_mem_spec k deps'). + * case: (knset_mem_spec k _). + move/KernameSet.union_spec => hin. + destruct hin. + eapply KernameSet.mem_spec in H. rewrite H. + rewrite !IHΣ. KernameSetDecide.fsetdec. + ** case: (knset_mem_spec k _); intros hin'; + rewrite !IHΣ; KernameSetDecide.fsetdec. + ** intros hin'. + case: (knset_mem_spec k _); intros hin''; + rewrite !IHΣ; KernameSetDecide.fsetdec. + * case: (knset_mem_spec k _); intros hin''. + case: (knset_mem_spec k _); intros hin'''. + rewrite !IHΣ. knset. + rewrite !IHΣ. knset. + case: (knset_mem_spec k _); intros hin'''. + rewrite !IHΣ. knset. + rewrite !IHΣ. knset. +Qed. + +Lemma in_global_deps deps Σ : + KernameSet.Subset deps (global_deps Σ deps). +Proof. + induction Σ => //. + destruct a as [kn [[[]]|]]; cbn; eauto; + case: (knset_mem_spec kn _); intros hin'''; + rewrite ?global_deps_union; + intros h hin; rewrite ?KernameSet.union_spec; try left; knset. +Qed. + +Lemma term_global_deps_csubst a k b : + KernameSet.Subset (term_global_deps (ECSubst.csubst a k b)) + (KernameSet.union (term_global_deps a) (term_global_deps b)). Admitted. +Lemma global_deps_subset Σ deps deps' : + KernameSet.Subset deps deps' -> + KernameSet.Subset (global_deps Σ deps) (global_deps Σ deps'). +Proof. + induction Σ in deps, deps' |- *. + - cbn. auto. + - destruct a; cbn. + intros sub. + case: (knset_mem_spec k deps) => hin. + eapply sub in hin. eapply KernameSet.mem_spec in hin. rewrite hin. + rewrite !global_deps_union. + specialize (IHΣ _ _ sub). knset. + case: (knset_mem_spec k deps') => hin'. + rewrite !global_deps_union. + specialize (IHΣ _ _ sub). knset. + now eapply IHΣ. +Qed. + +Arguments KernameSet.mem : simpl never. + +Lemma lookup_env_In Σ kn d : EGlobalEnv.lookup_env Σ kn = Some d -> In kn (map fst Σ). +Proof. + induction Σ; cbn; auto. + - easy. + - destruct (eqb_spec kn a.1). + intros [= <-]. left; auto. + intros hl; right; auto. +Qed. + + +Lemma term_global_deps_fresh {efl : EWellformed.EEnvFlags} Σ k t : + EWellformed.wellformed Σ k t -> + forall kn, KernameSet.In kn (term_global_deps t) -> In kn (map fst Σ). +Proof. + induction t in k |- * using EInduction.term_forall_list_ind; intros; try cbn in *; try knset; + rtoProp; intuition eauto. + all:try eapply KernameSet.union_spec in H0; intuition eauto. + - apply KernameSet.singleton_spec in H0. subst s. + destruct (EGlobalEnv.lookup_env Σ kn) eqn:E. + destruct g => //. destruct c => //. cbn in H1. + now eapply lookup_env_In in E. easy. + - admit. + - admit. + - admit. + - admit. + - admit. +Admitted. + +Lemma wf_global_decl_deps {efl : EWellformed.EEnvFlags} Σ d : + EWellformed.wf_global_decl Σ d -> + forall kn, KernameSet.In kn (global_decl_deps d) -> In kn (map fst Σ). +Proof. + intros wf kn hin. + destruct d as [[[]]|] => //; cbn in hin. + * eapply term_global_deps_fresh in hin. exact hin. cbn in wf; tea. + * knset. + * knset. +Qed. + +Lemma lookup_global_deps {efl : EWellformed.EEnvFlags} Σ kn decl : + EWellformed.wf_glob Σ -> + EGlobalEnv.lookup_env Σ kn = Some decl -> + forall kn, KernameSet.In kn (global_decl_deps decl) -> In kn (map fst Σ). +Proof. + induction 1 in kn, decl |- *. + - cbn => //. + - cbn. + case: (eqb_spec kn kn0) => heq. + subst kn0; intros [= <-]. + intros kn' hkn'. + * eapply wf_global_decl_deps in H0; tea. now right. + * intros hl kn' kin. + eapply IHwf_glob in hl; tea. now right. +Qed. + +Lemma fresh_global_In kn Σ : EGlobalEnv.fresh_global kn Σ <-> ~ In kn (map fst Σ). +Proof. + split. + - intros fr. + eapply (Forall_map (fun x => x <> kn) fst) in fr. + intros hin. + now eapply PCUICWfUniverses.Forall_In in fr; tea. + - intros nin. + red. eapply In_Forall. intros kn' hin <-. + eapply nin. now eapply (in_map fst). +Qed. + +Lemma global_deps_kn {fl :EWellformed.EEnvFlags} Σ kn decl : + EWellformed.wf_glob Σ -> + EGlobalEnv.declared_constant Σ kn decl -> + KernameSet.Subset (global_deps Σ (constant_decls_deps decl)) (global_deps Σ (KernameSet.singleton kn)). +Proof. + induction 1 in kn, decl |- *. + - cbn. unfold EGlobalEnv.declared_constant. cbn => //. + - cbn. unfold EGlobalEnv.declared_constant. cbn => //. + destruct (eqb_spec kn kn0). subst kn0. cbn. intros [= ->]. + + case: (knset_mem_spec kn _) => hin. + * eapply wf_global_decl_deps in H0; tea. + now eapply fresh_global_In in H1. + * case: (knset_mem_spec kn _) => hin'. + rewrite !global_deps_union. cbn. knset. + knset. + + + intros hl. + case: (knset_mem_spec kn0 _) => hin. + * elimtype False. + eapply lookup_global_deps in hl; tea. + now eapply fresh_global_In in H1. + * case: (knset_mem_spec kn0 _). + ** move/KernameSet.singleton_spec. intros <-. congruence. + ** intros hnin. now eapply IHwf_glob. +Qed. + +Lemma eval_global_deps {fl : EWcbvEval.WcbvFlags} {efl : EWellformed.EEnvFlags} Σ t t' : + EWellformed.wf_glob Σ -> + Σ ⊢ t ⇓ t' -> KernameSet.Subset (global_deps Σ (term_global_deps t')) (global_deps Σ (term_global_deps t)). +Proof. + intros wf. + induction 1 using EWcbvEval.eval_ind; cbn; rewrite ?global_deps_union; try knset. + - cbn in IHeval1. + epose proof (term_global_deps_csubst a' 0 b). + eapply (global_deps_subset Σ) in H2. + rewrite global_deps_union in H2. + knset. + - epose proof (term_global_deps_csubst b0' 0 b1). + eapply (global_deps_subset Σ) in H1. + rewrite global_deps_union in H1. + knset. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - cbn. + epose proof (global_deps_kn Σ c decl wf isdecl). + unfold constant_decls_deps in H0. destruct decl => //. + cbn in e. subst cst_body0. knset. + - admit. +Admitted. + +Lemma in_global_deps_fresh {efl : EWellformed.EEnvFlags} kn Σ deps : + EWellformed.wf_glob Σ -> + KernameSet.In kn (global_deps Σ deps) -> + EGlobalEnv.fresh_global kn Σ -> + KernameSet.In kn deps. +Proof. + induction Σ in deps |- *. + - now cbn. + - intros wf. destruct a as [kn' d]; cbn. depelim wf. + case: (knset_mem_spec kn' deps) => hin. + * rewrite global_deps_union KernameSet.union_spec. + intros [] fr. + ** depelim fr. now eapply IHΣ. + ** depelim fr. elimtype False. eapply IHΣ in H1; eauto. + destruct d as [[[]]|] eqn:eqd; cbn in H. + + cbn in H1. eapply (term_global_deps_fresh Σ) in H1; tea. cbn in H2. + eapply (Forall_map (fun x => x <> kn) fst) in fr. + now eapply PCUICWfUniverses.Forall_In in fr. + + cbn in H1; knset. + + cbn in H1; knset. + * intros hin' fr. depelim fr. now eapply IHΣ. +Qed. + +Lemma global_deps_empty Σ : global_deps Σ KernameSet.empty = KernameSet.empty. +Proof. + induction Σ; cbn; auto. + destruct a as [kn d]. + destruct (knset_mem_spec kn KernameSet.empty). + knset. auto. +Qed. + +Lemma filter_deps_filter {efl : EWellformed.EEnvFlags} deps Σ : + EWellformed.wf_glob Σ -> + (* (forall kn, KernameSet.In kn deps -> In kn (map fst Σ)) -> *) + filter_deps deps Σ = filter (flip KernameSet.mem (global_deps Σ deps) ∘ fst) Σ. +Proof. + induction Σ in deps |- *. + - now cbn. + - destruct a as [kn [[[]]|]]. + + cbn. unfold flip. + destruct (knset_mem_spec kn deps). + * case: (knset_mem_spec kn _). + ** intros hin' hwf. f_equal. depelim hwf. now eapply IHΣ. + ** intros nhin. + elim: nhin. + rewrite global_deps_union. + rewrite KernameSet.union_spec. left. + now eapply in_global_deps. + * intros hwf. + case: (knset_mem_spec kn _); intros hin. + ** elimtype False. + depelim hwf. + eapply in_global_deps_fresh in hin => //. + ** depelim hwf. rewrite -IHΣ => //. + + intros wf; depelim wf. + cbn. unfold flip. + * case: (knset_mem_spec kn _). + ** intros hin'. + case: (knset_mem_spec kn _) => hin''. + ++ f_equal. now eapply IHΣ. + ++ elim: hin''. + rewrite !global_deps_union. + rewrite KernameSet.union_spec. left. + now eapply in_global_deps. + ** intros hwf. + case: (knset_mem_spec kn _); intros hin. + *** elimtype False. + eapply in_global_deps_fresh in hin => //. + *** rewrite -IHΣ => //. + + intros wf; depelim wf. + cbn. unfold flip. + * case: (knset_mem_spec kn _). + ** intros hin'. + case: (knset_mem_spec kn _) => hin''. + ++ f_equal. rewrite IHΣ //. eapply filter_ext. + intros ?. unfold flip. rewrite global_deps_union /=. + rewrite global_deps_empty. + now rewrite KernameSetFact.union_b KernameSetFact.empty_b orb_false_r. + ++ elim: hin''. + rewrite !global_deps_union. + rewrite KernameSet.union_spec. left. + now eapply in_global_deps. + ** intros hwf. + case: (knset_mem_spec kn _); intros hin. + *** elimtype False. + eapply in_global_deps_fresh in hin => //. + *** rewrite -IHΣ => //. +Qed. + Lemma erase_global_declared_constructor X_type X ind c mind idecl cdecl deps decls normalization_in prf: forall Σ : global_env, abstract_env_rel X Σ -> declared_constructor Σ (ind, c) mind idecl cdecl -> KernameSet.In ind.(inductive_mind) deps -> - EGlobalEnv.declared_constructor (erase_global_decls X_type deps X decls normalization_in prf).1 (ind, c) + EGlobalEnv.declared_constructor (erase_global_deps X_type deps X decls normalization_in prf).1 (ind, c) (erase_mutual_inductive_body mind) (erase_one_inductive_body idecl) (EAst.mkConstructor cdecl.(cstr_name) cdecl.(cstr_arity)). Proof. @@ -2243,15 +2632,15 @@ Proof. cbn in *. split. split. - unshelve eapply declared_minductive_to_gen in H; eauto. red in H |- *. eapply (lookup_env_erase_decl _ _ _ _ _ _ _ (InductiveDecl mind)); eauto. - eapply in_deps_in_erase_global_decls; tea. + eapply in_deps_in_erase_global_deps; tea. - cbn. now eapply map_nth_error. - cbn. erewrite map_nth_error; eauto. Qed. Import EGlobalEnv. -Lemma erase_global_decls_fresh X_type kn deps X decls normalization_in heq : - let Σ' := erase_global_decls X_type deps X decls normalization_in heq in +Lemma erase_global_deps_fresh X_type kn deps X decls normalization_in heq : + let Σ' := erase_global_deps X_type deps X decls normalization_in heq in PCUICAst.fresh_global kn decls -> fresh_global kn Σ'.1. Proof. @@ -2340,7 +2729,7 @@ Lemma erases_deps_erase (cf := config.extraction_checker_flags) let et := erase X_type X' Γ t wt in let deps := EAstUtils.term_global_deps et in forall Σ, (abstract_env_rel X Σ) -> - erases_deps Σ (erase_global_decls X_type deps X decls normalization_in prf).1 et. + erases_deps Σ (erase_global_deps X_type deps X decls normalization_in prf).1 et. Proof. intros et deps Σ wf. pose proof (abstract_env_ext_exists X') as [[Σ' wfΣ']]. @@ -2370,7 +2759,7 @@ Lemma erases_deps_erase_weaken (cf := config.extraction_checker_flags) let et := erase X_type X' Γ t wt in let tdeps := EAstUtils.term_global_deps et in forall Σ, (abstract_env_rel X Σ) -> - erases_deps Σ (erase_global_decls X_type (KernameSet.union deps tdeps) X decls normalization_in prf).1 et. + erases_deps Σ (erase_global_deps X_type (KernameSet.union deps tdeps) X decls normalization_in prf).1 et. Proof. intros et tdeps Σ wf. pose proof (abstract_env_ext_exists X') as [[Σ' wfΣ']]. @@ -2407,7 +2796,7 @@ Proof. Qed. Lemma erase_global_closed X_type X deps decls {normalization_in} prf : - let X' := erase_global_decls X_type deps X decls normalization_in prf in + let X' := erase_global_deps X_type deps X decls normalization_in prf in EGlobalEnv.closed_env X'.1. Proof. revert X normalization_in prf deps. @@ -2517,11 +2906,66 @@ Proof. now rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wf'). Qed. -Lemma erase_wellformed (efl := all_env_flags) {X_type X} decls {normalization_in} prf univs wfΣ {Γ t} wt -(X' := abstract_make_wf_env_ext X univs wfΣ) -{normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} -(t' := erase X_type X' Γ t wt) : -wellformed (erase_global_decls X_type (term_global_deps t') X decls normalization_in prf).1 #|Γ| t'. +Lemma erases_global_erase_global (efl := all_env_flags) {X_type X} Σ decls {normalization_in} prf : + Σ ∼ X -> + erases_global Σ (@erase_global X_type X decls normalization_in prf). +Proof. + intros h. pose proof (prf _ h). red. rewrite H. clear H. + induction decls in X, Σ, prf, normalization_in, h |- *. + - cbn. constructor. + - cbn. + set (Xpop := abstract_pop_decls X). + destruct (abstract_env_exists Xpop) as [[Σpop hpop]]. + unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ h hpop). + intros; eexists; now eapply prf. + destruct H as [? []]. + destruct a as [kn []]. + * constructor. cbn. + unfold erase_constant_decl. + set (Σd := {| universes := _ |}). + assert (Σd = Σpop). + { subst Σd. rewrite -H H0 H1. now destruct Σpop. } + subst Σd. + set (ext := abstract_make_wf_env_ext _ _ _). + pose proof (abstract_env_ext_exists ext) as [[Σ' wf']]. + epose proof (abstract_make_wf_env_ext_correct (abstract_pop_decls X) (cst_universes c) _ _ _ hpop wf'). + subst Σ'. + cbn. pose proof (abstract_env_wf _ hpop) as []. + eapply (erase_constant_body_correct _ _ _ _ (Σpop, cst_universes c)) => //. + now rewrite H2. rewrite H2. split => //. now exists []. + rewrite H0 H1. + now eapply IHdecls. + * pose proof (abstract_env_wf _ h) as [wfΣ]. + destruct wfΣ. rewrite (prf _ h) in o0. + depelim o0. depelim o1. + constructor. + eapply erases_mutual. eapply on_global_decl_d. + rewrite H0 H1. + now eapply IHdecls. +Qed. + +Lemma erase_global_wellformed (efl := all_env_flags) {X_type X} decls {normalization_in} prf univs wfΣ {Γ t} wt + (X' := abstract_make_wf_env_ext X univs wfΣ) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} + (t' := erase X_type X' Γ t wt) : + wellformed (@erase_global X_type X decls normalization_in prf) #|Γ| t'. +Proof. + cbn. + epose proof (@erases_erase X_type X' normalization_in' _ _ wt). + pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. + pose proof (abstract_env_ext_exists X') as [[Σ' wf']]. + pose proof (abstract_env_ext_wf _ wf') as [?]. + eapply (erases_wellformed (wt _ wf')) in H; eauto; tea. + eapply erases_global_all_deps. apply X0. + epose proof (abstract_make_wf_env_ext_correct _ _ wfΣ _ _ wf wf'). subst Σ'. + now eapply erases_global_erase_global. +Qed. + +Lemma erase_wellformed_deps (efl := all_env_flags) {X_type X} decls {normalization_in} prf univs wfΣ {Γ t} wt + (X' := abstract_make_wf_env_ext X univs wfΣ) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} + (t' := erase X_type X' Γ t wt) : + wellformed (erase_global_deps X_type (term_global_deps t') X decls normalization_in prf).1 #|Γ| t'. Proof. cbn. epose proof (@erases_deps_erase X_type X univs wfΣ decls normalization_in prf _ Γ t wt). @@ -2543,7 +2987,7 @@ Lemma erase_wellformed_weaken (efl := all_env_flags) {X_type X} decls {normaliza {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} deps: let t' := erase X_type X' Γ t wt in - wellformed (erase_global_decls X_type (KernameSet.union deps (term_global_deps t')) X decls normalization_in prf).1 #|Γ| t'. + wellformed (erase_global_deps X_type (KernameSet.union deps (term_global_deps t')) X decls normalization_in prf).1 #|Γ| t'. Proof. set (t' := erase _ _ _ _ _). cbn. epose proof (@erases_deps_erase_weaken _ X univs wfΣ decls _ prf _ Γ t wt deps). @@ -2569,7 +3013,7 @@ Lemma erase_constant_body_correct'' {X_type X} {cb} {decls normalization_in prf} forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> ∥ ∑ t T, (Σ' ;;; [] |- t : T) * (Σ' ;;; [] |- t ⇝ℇ body) * (term_global_deps body = snd (erase_constant_body X_type X' cb onc)) * - wellformed (efl:=all_env_flags) (erase_global_decls X_type (KernameSet.union deps (term_global_deps body)) X decls normalization_in prf).1 0 body ∥. + wellformed (efl:=all_env_flags) (erase_global_deps X_type (KernameSet.union deps (term_global_deps body)) X decls normalization_in prf).1 0 body ∥. Proof. intros ? Σ' wfΣ'. pose proof (abstract_env_exists X) as [[Σ wf]]. destruct cb as [name [bod|] udecl]; simpl. @@ -2595,7 +3039,7 @@ Lemma erase_global_cst_decl_wf_glob X_type X deps decls normalization_in heq : let X' := abstract_make_wf_env_ext X (cst_universes cb) wfΣ in forall normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ, let ecb := erase_constant_body X_type X' cb hcb in - let Σ' := erase_global_decls X_type (KernameSet.union deps ecb.2) X decls normalization_in heq in + let Σ' := erase_global_deps X_type (KernameSet.union deps ecb.2) X decls normalization_in heq in (@wf_global_decl all_env_flags Σ'.1 (EAst.ConstantDecl ecb.1) : Prop). Proof. intros cb wfΣ hcb X' normalization_in' ecb Σ'. @@ -2611,11 +3055,11 @@ Qed. Lemma erase_global_ind_decl_wf_glob {X_type X} {deps decls normalization_in kn m} heq : (forall Σ : global_env, abstract_env_rel X Σ -> on_inductive cumulSpec0 (lift_typing typing) (Σ, ind_universes m) kn m) -> let m' := erase_mutual_inductive_body m in - let Σ' := erase_global_decls X_type deps X decls normalization_in heq in + let Σ' := erase_global_deps X_type deps X decls normalization_in heq in @wf_global_decl all_env_flags Σ'.1 (EAst.InductiveDecl m'). Proof. set (m' := erase_mutual_inductive_body m). - set (Σ' := erase_global_decls _ _ _ _ _ _). simpl. + set (Σ' := erase_global_deps _ _ _ _ _ _). simpl. intros oni. pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. pose proof (abstract_env_wf _ wf) as [wfΣ]. @@ -2625,8 +3069,8 @@ Proof. rewrite -(heq _ wf). now destruct Σ. Qed. -Lemma erase_global_decls_wf_glob X_type X deps decls normalization_in heq : - let Σ' := erase_global_decls X_type deps X decls normalization_in heq in +Lemma erase_global_deps_wf_glob X_type X deps decls normalization_in heq : + let Σ' := erase_global_deps X_type deps X decls normalization_in heq in @wf_glob all_env_flags Σ'.1. Proof. cbn. @@ -2640,7 +3084,105 @@ Proof. epose proof (abstract_env_exists Xpop) as [[Σpop wfpop]]; pose proof (abstract_env_wf _ wfpop) as [wfΣpop]. + constructor. eapply IHdecls => //; eauto. eapply erase_global_cst_decl_wf_glob; auto. - eapply erase_global_decls_fresh; auto. + eapply erase_global_deps_fresh; auto. + destruct wfΣ. destruct wfΣpop. + rewrite (heq _ wf) in o0. depelim o0. now destruct o3. + + cbn. eapply IHdecls; eauto. + + constructor. eapply IHdecls; eauto. + destruct wfΣ as [[onu ond]]. + rewrite (heq _ wf) in o. depelim o. destruct o0. + eapply (erase_global_ind_decl_wf_glob (kn:=kn)); tea. + intros. rewrite (abstract_env_irr _ H wfpop). + unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wf wfpop) as [? ?]. + {intros; now eexists. } + destruct Σpop, Σ; cbn in *. now subst. + eapply erase_global_deps_fresh. + destruct wfΣ as [[onu ond]]. + rewrite (heq _ wf) in o. depelim o. now destruct o0. + + eapply IHdecls; eauto. +Qed. + +Lemma erase_constant_body_correct_glob {X_type X} {cb} {decls normalization_in prf} {univs wfΣ} + (X' := abstract_make_wf_env_ext X univs wfΣ) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} + {onc : forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> ∥ on_constant_decl (lift_typing typing) Σ' cb ∥} {body} : + EAst.cst_body (fst (erase_constant_body X_type X' cb onc)) = Some body -> + forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> + ∥ ∑ t T, (Σ' ;;; [] |- t : T) * (Σ' ;;; [] |- t ⇝ℇ body) * + (term_global_deps body = snd (erase_constant_body X_type X' cb onc)) * + wellformed (efl:=all_env_flags) (@erase_global X_type X decls normalization_in prf) 0 body ∥. +Proof. + intros ? Σ' wfΣ'. pose proof (abstract_env_exists X) as [[Σ wf]]. + destruct cb as [name [bod|] udecl]; simpl. + simpl in H. noconf H. + set (obl :=(erase_constant_body_obligation_1 X_type X' + {| + cst_type := name; + cst_body := Some bod; + cst_universes := udecl |} onc bod eq_refl)). clearbody obl. + destruct (obl _ wfΣ'). sq. + have er : (Σ, univs);;; [] |- bod ⇝ℇ erase X_type X' [] bod obl. + { eapply (erases_erase (X:=X')). + now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). + } + exists bod, A; intuition auto. + now rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). + eapply (erase_global_wellformed decls _ _ _ (Γ := [])). + now simpl in H. +Qed. + +Lemma erase_global_cst_wf_glob X_type X decls normalization_in heq : + forall cb wfΣ hcb, + let X' := abstract_make_wf_env_ext X (cst_universes cb) wfΣ in + forall normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ, + let ecb := erase_constant_body X_type X' cb hcb in + let Σ' := @erase_global X_type X decls normalization_in heq in + (@wf_global_decl all_env_flags Σ' (EAst.ConstantDecl ecb.1) : Prop). +Proof. + intros cb wfΣ hcb X' normalization_in' ecb Σ'. + unfold wf_global_decl. cbn. + pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. + pose proof (abstract_env_ext_exists X') as [[Σmake wfmake]]. + destruct (wfΣ _ wf), (hcb _ wfmake). red in X1. + destruct EAst.cst_body eqn:hb => /= //. + eapply (erase_constant_body_correct'') in hb; eauto. + destruct hb as [[t0 [T [[] ?]]]]. rewrite e in i. exact i. +Qed. + +Lemma erase_global_ind_decl_wf_glob {X_type X} {deps decls normalization_in kn m} heq : + (forall Σ : global_env, abstract_env_rel X Σ -> on_inductive cumulSpec0 (lift_typing typing) (Σ, ind_universes m) kn m) -> + let m' := erase_mutual_inductive_body m in + let Σ' := erase_global_deps X_type deps X decls normalization_in heq in + @wf_global_decl all_env_flags Σ'.1 (EAst.InductiveDecl m'). +Proof. + set (m' := erase_mutual_inductive_body m). + set (Σ' := erase_global_deps _ _ _ _ _ _). simpl. + intros oni. + pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. + pose proof (abstract_env_wf _ wf) as [wfΣ]. + assert (erases_mutual_inductive_body m (erase_mutual_inductive_body m)). + { eapply (erases_mutual (mdecl:=kn)); tea. } + eapply (erases_mutual_inductive_body_wf (univs := Σ.(universes)) (retro := Σ.(retroknowledge)) (Σ := decls) (kn := kn) (Σ' := Σ'.1)) in H; tea. + rewrite -(heq _ wf). now destruct Σ. +Qed. + +Lemma erase_global_wf_glob X_type X decls normalization_in heq : + let Σ' := @erase_global X_type X decls normalization_in heq in + @wf_glob all_env_flags Σ'. +Proof. + cbn. + pose proof (abstract_env_exists X) as [[Σ wf]]. + pose proof (abstract_env_wf _ wf) as [wfΣ]. + revert X Σ wf wfΣ normalization_in heq. + induction decls; [cbn; auto|]. + { intros. constructor. } + intros. destruct a as [kn []]; simpl; set (Xpop := abstract_pop_decls X); + try set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _); + epose proof (abstract_env_exists Xpop) as [[Σpop wfpop]]; + pose proof (abstract_env_wf _ wfpop) as [wfΣpop]. + + constructor. eapply IHdecls => //; eauto. + eapply erase_global_cst_decl_wf_glob; auto. + eapply erase_global_deps_fresh; auto. destruct wfΣ. destruct wfΣpop. rewrite (heq _ wf) in o0. depelim o0. now destruct o3. + cbn. eapply IHdecls; eauto. @@ -2652,15 +3194,99 @@ Proof. unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wf wfpop) as [? ?]. {intros; now eexists. } destruct Σpop, Σ; cbn in *. now subst. - eapply erase_global_decls_fresh. + eapply erase_global_deps_fresh. destruct wfΣ as [[onu ond]]. rewrite (heq _ wf) in o. depelim o. now destruct o0. + eapply IHdecls; eauto. Qed. + +Lemma extends_cons d Σ Σ' : extends Σ Σ' -> extends (d :: Σ) (d :: Σ'). +Proof. + intros ext kn decl. + cbn. + case: (eqb_spec kn d.1). + - now intros -> [= <-]. + - intros hin. apply ext. +Qed. + +Lemma extends_fresh Σ Σ' kn d : + fresh_global kn Σ' -> + extends Σ Σ' -> + extends Σ ((kn, d) :: Σ'). +Proof. + intros fr ext. + intros kn' decl hl. + cbn. destruct (eqb_spec kn' kn) => //. subst. + eapply ext in hl. now eapply lookup_env_Some_fresh in hl. + now eapply ext. +Qed. + +Lemma Forall_filter {A} {P} f l : @Forall A P l -> Forall P (filter f l). +Proof. + induction 1; try econstructor; auto. + cbn. case E: (f x) => //. now constructor. +Qed. + +Lemma extends_filter_impl {efl : EEnvFlags} (f f' : kername * EAst.global_decl -> bool) Σ : + (forall x, f x -> f' x) -> + wf_glob Σ -> + extends (filter f Σ) (filter f' Σ). +Proof. + intros hf; induction Σ; cbn; auto. + - red; auto. + - intros wf; depelim wf. + case E: (f _); auto. rewrite (hf _ E). + now eapply extends_cons. + case E': (f' _); auto. + eapply extends_fresh. now eapply Forall_filter. + auto. +Qed. + +Lemma extends_filter {efl : EEnvFlags} (f : kername * EAst.global_decl -> bool) Σ : + wf_glob Σ -> + extends (filter f Σ) Σ. +Proof. + intros wf hf; induction Σ; cbn; auto. + intros decl. depelim wf. cbn in *. + case E: (f _); auto; cbn. + case (eqb_spec hf kn) => //. + - intros hd. now eapply IHΣ. + - intros hl; eapply IHΣ in hl => //. rewrite hl. + case (eqb_spec hf kn) => heq. + subst. now eapply lookup_env_Some_fresh in hl. + reflexivity. +Qed. + +Lemma erase_global_deps_eval {efl : EWcbvEval.WcbvFlags} (X_type : abstract_env_impl) (X:X_type.π1) ext wfX {normalization_in} decls nin prf Γ t wt v wv : + let Xext := abstract_make_wf_env_ext X ext wfX in + let t' := erase X_type Xext (normalization_in := normalization_in) Γ t wt in + let v' := erase X_type Xext (normalization_in := normalization_in) Γ v wv in + let eg := erase_global_deps X_type (term_global_deps t') X decls nin prf in + let eg' := erase_global_deps X_type (term_global_deps v') X decls nin prf in + eg.1 ⊢ t' ⇓ v' -> EGlobalEnv.extends eg'.1 eg.1. +Proof. + intros. + move: (erase_global_deps_wf_glob X_type X (term_global_deps t') decls nin prf). + subst eg eg'. + rewrite -!erase_global_deps. + assert (wfer : wf_glob (efl := all_env_flags) (@erase_global X_type X decls nin prf)). + { admit. } + rewrite !(filter_deps_filter (efl := all_env_flags)) //. + cbn. + intros wf. + eapply extends_filter_impl; tea. unfold flip in *. + intros [kn d]; cbn. clear d. + rewrite /is_true !KernameSet.mem_spec. revert kn. + eapply EWcbvEval.weakening_eval_env in H. 2:exact wfer. + now eapply eval_global_deps in H. + rewrite -!erase_global_deps (filter_deps_filter (efl := all_env_flags)) //. + now eapply extends_filter. +Qed. + Lemma lookup_erase_global (cf := config.extraction_checker_flags) {X_type X} {deps deps' decls normalization_in prf} : KernameSet.Subset deps deps' -> - EExtends.global_subset (erase_global_decls X_type deps X decls normalization_in prf).1 (erase_global_decls X_type deps' X decls normalization_in prf).1. + EExtends.global_subset (erase_global_deps X_type deps X decls normalization_in prf).1 (erase_global_deps X_type deps' X decls normalization_in prf).1. Proof. revert deps deps' X normalization_in prf. induction decls. cbn => //. @@ -2669,10 +3295,10 @@ Proof. destruct a as [kn' []]; cbn; ( set (decl := E.ConstantDecl _) || set (decl := E.InductiveDecl _)); hidebody decl; - set (eg := erase_global_decls _ _ _ _ _ _); hidebody eg; - set (eg' := erase_global_decls _ _ _ _ _ _); hidebody eg'; - try (set (eg'' := erase_global_decls _ _ _ _ _ _); hidebody eg''); - try (set (eg''' := erase_global_decls _ _ _ _ _ _); hidebody eg'''). + set (eg := erase_global_deps _ _ _ _ _ _); hidebody eg; + set (eg' := erase_global_deps _ _ _ _ _ _); hidebody eg'; + try (set (eg'' := erase_global_deps _ _ _ _ _ _); hidebody eg''); + try (set (eg''' := erase_global_deps _ _ _ _ _ _); hidebody eg'''). { destruct (KernameSet.mem) eqn:knm => /=. + eapply KernameSet.mem_spec, sub, KernameSet.mem_spec in knm. rewrite knm. apply EExtends.global_subset_cons. eapply IHdecls; eauto. @@ -2681,15 +3307,15 @@ Proof. right => //. + destruct (KernameSet.mem kn' deps') eqn:eq'. eapply EExtends.global_subset_cons_right; eauto. - eapply erase_global_decls_wf_glob. + eapply erase_global_deps_wf_glob. unfold decl. unfold hidebody. - constructor. eapply erase_global_decls_wf_glob. + constructor. eapply erase_global_deps_wf_glob. eapply erase_global_cst_decl_wf_glob. - eapply erase_global_decls_fresh => //. + eapply erase_global_deps_fresh => //. pose proof (abstract_env_wf _ wf) as [wfΣ]. depelim wfΣ. rewrite (prf _ wf) in o0. clear - o0. depelim o0. now destruct o. unfold eg', eg'', hidebody. - erewrite erase_global_decls_irr. + erewrite erase_global_deps_irr. eapply IHdecls. intros x hin. eapply KernameSet.union_spec. left. now eapply sub. @@ -2698,8 +3324,8 @@ Proof. + eapply KernameSet.mem_spec, sub, KernameSet.mem_spec in knm. rewrite knm. apply EExtends.global_subset_cons. eapply IHdecls => //. + destruct (KernameSet.mem kn' deps') eqn:eq'. - eapply EExtends.global_subset_cons_right. eapply erase_global_decls_wf_glob. - constructor. eapply erase_global_decls_wf_glob. + eapply EExtends.global_subset_cons_right. eapply erase_global_deps_wf_glob. + constructor. eapply erase_global_deps_wf_glob. pose proof (abstract_env_wf _ wf) as [wfΣ]. pose proof (prf _ wf) as prf'. eapply (erase_global_ind_decl_wf_glob (kn:=kn')). @@ -2709,20 +3335,20 @@ Proof. destruct Σ, Σ0. cbn in *. rewrite prf' in wfΣ. depelim wfΣ. cbn in *. rewrite <- H1, H0, <- H2. depelim o0. now destruct o1. - eapply erase_global_decls_fresh => //. + eapply erase_global_deps_fresh => //. pose proof (abstract_env_wf _ wf) as [wfΣ]. pose proof (prf _ wf) as prf'. destruct Σ. cbn in *. rewrite prf' in wfΣ. clear -wfΣ. destruct wfΣ. cbn in *. depelim o0. now destruct o1. - unfold eg'', hidebody. erewrite erase_global_decls_irr. + unfold eg'', hidebody. erewrite erase_global_deps_irr. eapply IHdecls. intros x hin. now eapply sub. eapply IHdecls => //. } Qed. Lemma expanded_weakening_global X_type X deps deps' decls normalization_in prf Γ t : KernameSet.Subset deps deps' -> - expanded (erase_global_decls X_type deps X decls normalization_in prf).1 Γ t -> - expanded (erase_global_decls X_type deps' X decls normalization_in prf).1 Γ t. + expanded (erase_global_deps X_type deps X decls normalization_in prf).1 Γ t -> + expanded (erase_global_deps X_type deps' X decls normalization_in prf).1 Γ t. Proof. intros hs. eapply expanded_ind; intros; try solve [econstructor; eauto]. @@ -2741,7 +3367,7 @@ Lemma expanded_erase (cf := config.extraction_checker_flags) forall normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ, let et := (erase X_type X' [] t wtp) in let deps := EAstUtils.term_global_deps et in - expanded (erase_global_decls X_type deps X decls normalization_in prf).1 [] et. + expanded (erase_global_deps X_type deps X decls normalization_in prf).1 [] et. Proof. intros Σ wf hexp X' normalization_in'. pose proof (abstract_env_wf _ wf) as [wf']. @@ -2757,7 +3383,7 @@ Lemma expanded_erase_global (cf := config.extraction_checker_flags) deps {X_type X decls normalization_in prf} : forall Σ : global_env, abstract_env_rel X Σ -> PCUICEtaExpand.expanded_global_env Σ -> - expanded_global_env (erase_global_decls X_type deps X decls normalization_in prf).1. + expanded_global_env (erase_global_deps X_type deps X decls normalization_in prf).1. Proof. intros Σ wf etaΣ. pose proof (prf _ wf). destruct Σ as [univ decls']. @@ -2808,7 +3434,7 @@ Lemma erase_eval_to_box (wfl := Ee.default_wcbv_flags) forall wt : forall Σ : global_env_ext, abstract_env_ext_rel Xext Σ -> welltyped Σ [] t, erase X_type Xext [] t wt = t' -> KernameSet.subset (term_global_deps t') deps -> - erase_global_decls X_type deps X decls normalization_in prf = Σ' -> + erase_global_deps X_type deps X decls normalization_in prf = Σ' -> forall Σext : global_env_ext, abstract_env_ext_rel Xext Σext -> (forall Σ : global_env, abstract_env_rel X Σ -> PCUICWcbvEval.eval Σ t v) -> @@ -2834,7 +3460,7 @@ Lemma erase_eval_to_box_eager (wfl := Ee.default_wcbv_flags) forall wt : forall Σ : global_env_ext, abstract_env_ext_rel Xext Σ -> welltyped Σ [] t, erase X_type Xext [] t wt = t' -> KernameSet.subset (term_global_deps t') deps -> - erase_global_decls X_type deps X decls normalization_in prf = Σ' -> + erase_global_deps X_type deps X decls normalization_in prf = Σ' -> forall Σext : global_env_ext, abstract_env_ext_rel Xext Σext -> (forall Σ : global_env, abstract_env_rel X Σ -> PCUICWcbvEval.eval Σ t v) -> @@ -2941,7 +3567,7 @@ forall Σ, abstract_env_ext_rel Xext Σ -> @firstorder_ind Σ (firstorder_env Σ) i -> erase X_type Xext [] t wt = t' -> KernameSet.subset (term_global_deps t') deps -> - erase_global_decls X_type deps X decls normalization_in prf = Σ' -> + erase_global_deps X_type deps X decls normalization_in prf = Σ' -> red Σ [] t v -> ¬ { t' & Σ ;;; [] |- v ⇝ t'} -> forall wt', ∥ Σ'.1 ⊢ t' ⇓ erase X_type Xext [] v wt' ∥. @@ -2970,7 +3596,7 @@ forall Σ, abstract_env_ext_rel Xext Σ -> axiom_free Σ -> Σ ;;; [] |- t : mkApps (tInd i u) args -> let t' := erase X_type Xext [] t wt in - let Σ' := erase_global_decls X_type deps X decls normalization_in prf in + let Σ' := erase_global_deps X_type deps X decls normalization_in prf in @firstorder_ind Σ (firstorder_env Σ) i -> KernameSet.subset (term_global_deps t') deps -> red Σ [] t v -> @@ -3012,7 +3638,7 @@ forall Σ (wfΣ : abstract_env_ext_rel Xext Σ) axiom_free Σ -> let t' := erase X_type Xext [] t (unique_type wfΣ wt) in let decls := declarations (fst Σ) in - let Σ' := erase_global_decls X_type (term_global_deps t') X decls + let Σ' := erase_global_deps X_type (term_global_deps t') X decls (fun _ _ _ _ _ _ _ => normalization_in _ _) (unique_decls univs wfext wfΣ) in @firstorder_ind Σ (firstorder_env Σ) i -> red Σ [] t v -> @@ -3025,7 +3651,7 @@ Proof. Qed. (* we use the [match] trick to get typeclass resolution to pick up the right instances without leaving any evidence in the resulting term, and without having to pass them manually everywhere *) -Notation NormalizationIn_erase_global_decls_fast X decls +Notation NormalizationIn_erase_global_deps_fast X decls := (match extraction_checker_flags, extraction_normalizing return _ with | cf, no => forall n, @@ -3086,9 +3712,9 @@ Qed. (* This version of global environment erasure keeps the same global environment throughout the whole erasure, while folding over the list of declarations. *) -Program Fixpoint erase_global_decls_fast (deps : KernameSet.t) +Program Fixpoint erase_global_deps_fast (deps : KernameSet.t) X_type (X:X_type.π1) (decls : global_declarations) - {normalization_in : NormalizationIn_erase_global_decls_fast X decls} + {normalization_in : NormalizationIn_erase_global_deps_fast X decls} (prop : forall Σ : global_env, abstract_env_rel X Σ -> ∥ decls_prefix decls Σ ∥) : E.global_declarations * KernameSet.t := match decls with | [] => ([],deps) @@ -3098,17 +3724,17 @@ Program Fixpoint erase_global_decls_fast (deps : KernameSet.t) let normalization_in' : @id Prop _ := _ in (* hack to avoid program erroring out on unresolved tc *) let cb' := @erase_constant_body X_type Xext' normalization_in' cb _ in let deps := KernameSet.union deps (snd cb') in - let Σ' := erase_global_decls_fast deps X_type X decls _ in + let Σ' := erase_global_deps_fast deps X_type X decls _ in (((kn, E.ConstantDecl (fst cb')) :: Σ'.1), Σ'.2) else - erase_global_decls_fast deps X_type X decls _ + erase_global_deps_fast deps X_type X decls _ | (kn, InductiveDecl mib) :: decls => if KernameSet.mem kn deps then let mib' := erase_mutual_inductive_body mib in - let Σ' := erase_global_decls_fast deps X_type X decls _ in + let Σ' := erase_global_deps_fast deps X_type X decls _ in (((kn, E.InductiveDecl mib') :: Σ'.1), Σ'.2) - else erase_global_decls_fast deps X_type X decls _ + else erase_global_deps_fast deps X_type X decls _ end. Next Obligation. pose proof (abstract_env_wf _ H) as [?]. @@ -3236,13 +3862,13 @@ Definition same_principal_type {cf} (*Lemma erase_global_deps_suffix {deps} {Σ Σ' : wf_env} {decls hprefix hprefix'} : wf Σ -> wf Σ' -> universes Σ = universes Σ' -> - erase_global_decls_fast deps Σ decls hprefix = - erase_global_decls_fast deps Σ' decls hprefix'. + erase_global_deps_fast deps Σ decls hprefix = + erase_global_deps_fast deps Σ' decls hprefix'. Proof using Type. intros wfΣ wfΣ' equ. induction decls in deps, hprefix, hprefix' |- *; cbn => //. destruct a as [kn []]. - set (obl := (erase_global_decls_fast_obligation_1 Σ ((kn, ConstantDecl c) :: decls) + set (obl := (erase_global_deps_fast_obligation_1 Σ ((kn, ConstantDecl c) :: decls) hprefix kn c decls eq_refl)). clearbody obl. destruct obl. set (eb := erase_constant_body _ _ _). @@ -3251,7 +3877,7 @@ Proof using Type. { subst eb eb'. set (wfΣx := abstract_make_wf_env_ext _ _ _). set (wfΣ'x := abstract_make_wf_env_ext _ _ _). - set (obl' := erase_global_decls_fast_obligation_2 _ _ _ _ _ _ _). + set (obl' := erase_global_deps_fast_obligation_2 _ _ _ _ _ _ _). clearbody obl'. set (Σd := {| universes := Σ.(universes); declarations := decls |}). assert (wfΣd : wf Σd). @@ -3313,8 +3939,8 @@ Qed. Lemma erase_global_deps_fast_spec_gen {deps} {X_type X X'} {decls normalization_in normalization_in' hprefix hprefix'} : (forall Σ Σ', abstract_env_rel X Σ -> abstract_env_rel X' Σ' -> universes Σ = universes Σ' /\ retroknowledge Σ = retroknowledge Σ') -> - erase_global_decls_fast deps X_type X decls (normalization_in:=normalization_in) hprefix = - erase_global_decls X_type deps X' decls normalization_in' hprefix'. + erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) hprefix = + erase_global_deps X_type deps X' decls normalization_in' hprefix'. Proof using Type. intros equ. induction decls in X, X', equ, deps, normalization_in, normalization_in', hprefix, hprefix' |- * => //. @@ -3327,12 +3953,12 @@ Proof using Type. destruct a as [kn []]. - cbn. unfold erase_constant_decl. - set (obl := (erase_global_decls_fast_obligation_1 X_type X ((kn, ConstantDecl c) :: decls) + set (obl := (erase_global_deps_fast_obligation_1 X_type X ((kn, ConstantDecl c) :: decls) _ hprefix kn c decls eq_refl)). set (eb := erase_constant_body _ _ _ _). set (eb' := erase_constant_body _ _ _ _). - set (eg := erase_global_decls _ _ _ _ _ _). - set (eg' := erase_global_decls _ _ _ _ _ _). + set (eg := erase_global_deps _ _ _ _ _ _). + set (eg' := erase_global_deps _ _ _ _ _ _). assert (eb = eb') as <-. { subst eb eb'. @@ -3378,18 +4004,18 @@ Proof using Type. Qed. Lemma erase_global_deps_fast_spec {deps} {X_type X} {decls normalization_in normalization_in' hprefix hprefix'} : - erase_global_decls_fast deps X_type X decls (normalization_in:=normalization_in) hprefix = - erase_global_decls X_type deps X decls normalization_in' hprefix'. + erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) hprefix = + erase_global_deps X_type deps X decls normalization_in' hprefix'. Proof using Type. eapply erase_global_deps_fast_spec_gen; intros. rewrite (abstract_env_irr _ H H0); eauto. Qed. Definition erase_global_fast X_type deps X decls {normalization_in} (prf:forall Σ : global_env, abstract_env_rel X Σ -> declarations Σ = decls) := - erase_global_decls_fast deps X_type X decls (normalization_in:=normalization_in) (fun _ H => sq ([] ; prf _ H)). + erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) (fun _ H => sq ([] ; prf _ H)). Lemma expanded_erase_global_fast (cf := config.extraction_checker_flags) deps - {X_type X decls normalization_in} {normalization_in':NormalizationIn_erase_global_decls X decls} {prf} : + {X_type X decls normalization_in} {normalization_in':NormalizationIn_erase_global_deps X decls} {prf} : forall Σ : global_env, abstract_env_rel X Σ -> PCUICEtaExpand.expanded_global_env Σ -> expanded_global_env (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf).1. @@ -3401,7 +4027,7 @@ Proof using Type. Qed. Lemma expanded_erase_fast (cf := config.extraction_checker_flags) - {X_type X decls normalization_in prf} {normalization_in'':NormalizationIn_erase_global_decls X decls} univs wfΣ t wtp : + {X_type X decls normalization_in prf} {normalization_in'':NormalizationIn_erase_global_deps X decls} univs wfΣ t wtp : forall Σ : global_env, abstract_env_rel X Σ -> PCUICEtaExpand.expanded Σ [] t -> let X' := abstract_make_wf_env_ext X univs wfΣ in @@ -3420,16 +4046,16 @@ Proof using Type. Unshelve. all: eauto. Qed. -Lemma erase_global_fast_wf_glob X_type deps X decls normalization_in (normalization_in':NormalizationIn_erase_global_decls X decls) prf : +Lemma erase_global_fast_wf_glob X_type deps X decls normalization_in (normalization_in':NormalizationIn_erase_global_deps X decls) prf : @wf_glob all_env_flags (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf).1. Proof using Type. unfold erase_global_fast; erewrite erase_global_deps_fast_spec. - eapply erase_global_decls_wf_glob. + eapply erase_global_deps_wf_glob. Unshelve. all: eauto. Qed. Lemma erase_wellformed_fast (efl := all_env_flags) - {X_type X decls normalization_in prf} {normalization_in'':NormalizationIn_erase_global_decls X decls} univs wfΣ {Γ t} wt + {X_type X decls normalization_in prf} {normalization_in'':NormalizationIn_erase_global_deps X decls} univs wfΣ {Γ t} wt (X' := abstract_make_wf_env_ext X univs wfΣ) {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} : let t' := erase X_type X' Γ t wt in @@ -3442,3 +4068,53 @@ Proof using Type. Qed. End EraseGlobalFast. + +#[local] Instance extends_trans : Transitive extends. +Proof. + Admitted. + +Lemma extends_erase_pcuic_program {wfl : EWcbvEval.WcbvFlags} X_type X decls nin pf : + forall Σ et et', + EWcbvEval.eval Σ et' et -> + Σ = (@erase_global_deps X_type (term_global_deps et') X decls nin pf).1 -> + EGlobalEnv.extends + (@erase_global_deps X_type (term_global_deps et) X decls nin pf).1 + (@erase_global_deps X_type (term_global_deps et') X decls nin pf).1. +Proof. + intros Σ et et'. + induction 1; intros eqΣ. + - cbn. eapply lookup_erase_global. KernameSetDecide.fsetdec. + - cbn in *. etransitivity; tea. + transitivity ((erase_global_deps X_type (term_global_deps (ECSubst.csubst a 0 f0)) X decls nin pf).1). + admit. admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - admit. + - cbn. etransitivity; tea. + intros kn decl'. + intros hl. epose proof (lookup_env_erase_decl X_type X). in hl. + + + +Lemma extends_erase_pcuic_program X_type X t t' Γ : + (forall Σ, Σ ∼ X -> PCUICWcbvEval.eval Σ t' t) -> + forall Xext norm decls nin pf wt wt', + let et := @erase X_type Xext norm Γ t wt in + let et' := @erase X_type Xext norm Γ t' wt' in + EGlobalEnv.extends + (@erase_global_deps X_type (term_global_deps et) X decls nin pf).1 + (@erase_global_deps X_type (term_global_deps et') X decls nin pf).1. +Proof. + intros ev Xext norm decls nin pf wt wt' et et'. + intros kn decl. + pose proof (abstract_env_exists X) as [[Σ H]]. + specialize (ev _ H). + induction ev. + +Admitted. From 93395bfb481db230d41b4aca9239f09e081c7718 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 30 Aug 2023 11:15:21 +0200 Subject: [PATCH 43/61] Finish the proofs in ErasureFunction --- erasure/theories/ErasureFunction.v | 947 +++++++++++++++++++---------- 1 file changed, 622 insertions(+), 325 deletions(-) diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 3d4b6710e..15f45b589 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -68,9 +68,589 @@ Section OnInductives. eapply PCUICInductives.on_minductive_wf_params_indices_inst; tea. Qed. - End OnInductives. +Infix "=kn" := KernameSet.Equal (at level 90). +Ltac knset := KernameSetDecide.fsetdec. + +Lemma term_global_deps_csubst a k b : + KernameSet.Subset (term_global_deps (ECSubst.csubst a k b)) + (KernameSet.union (term_global_deps a) (term_global_deps b)). +Proof. + induction b in k |- * using EInduction.term_forall_list_ind; cbn; try knset; eauto. + - destruct Nat.compare; knset. + - specialize (IHb1 k); specialize (IHb2 (S k)); knset. + - specialize (IHb1 k); specialize (IHb2 k); knset. + - destruct i. intros kn. rewrite knset_in_fold_left. + rewrite KernameSet.union_spec knset_in_fold_left. + intuition auto. destruct H0 as [a' [hin hin']]. + eapply in_map_iff in hin as [cs [<- hin]]. + eapply All_In in X as [X]; tea. specialize (X k). + specialize (X _ hin'). + eapply KernameSet.union_spec in X as []. knset. + right; right. now exists cs. + - destruct p. intros kn. + rewrite !KernameSet.union_spec !knset_in_fold_left. + specialize (IHb k kn). intuition auto. knset. + destruct H as [? []]. + eapply in_map_iff in H as [cs [<- hin]]. + eapply All_In in X as [X]; tea. specialize (X _ _ H0). intuition auto; try knset. + cbn in H0. apply KernameSet.union_spec in X as []; [knset|]. + right; right; right. now exists cs. + - specialize (IHb k). knset. + - intros kn; rewrite knset_in_fold_left !KernameSet.union_spec knset_in_fold_left. + intuition auto. + destruct H0 as [? []]. + eapply in_map_iff in H as [cs [<- hin]]. + eapply All_In in X as [X]; tea. specialize (X _ _ H0). intuition auto; try knset. + cbn in H0. apply KernameSet.union_spec in X as []; [knset|]. + right; right. now exists cs. + - intros kn; rewrite knset_in_fold_left !KernameSet.union_spec knset_in_fold_left. + intuition auto. + destruct H0 as [? []]. + eapply in_map_iff in H as [cs [<- hin]]. + eapply All_In in X as [X]; tea. specialize (X _ _ H0). intuition auto; try knset. + cbn in H0. apply KernameSet.union_spec in X as []; [knset|]. + right; right. now exists cs. +Qed. + + +Arguments KernameSet.mem : simpl never. + +Lemma lookup_env_In Σ kn d : EGlobalEnv.lookup_env Σ kn = Some d -> In kn (map fst Σ). +Proof. + induction Σ; cbn; auto. + - easy. + - destruct (eqb_spec kn a.1). + intros [= <-]. left; auto. + intros hl; right; auto. +Qed. + + +Notation terms_global_deps l := + (fold_left (fun (acc : KernameSet.t) (x : EAst.term) => + KernameSet.union (term_global_deps x) acc) l + KernameSet.empty). + +Lemma term_global_deps_fresh {efl : EWellformed.EEnvFlags} Σ k t : + EWellformed.wellformed Σ k t -> + forall kn, KernameSet.In kn (term_global_deps t) -> In kn (map fst Σ). +Proof. + induction t in k |- * using EInduction.term_forall_list_ind; intros; try (cbn in *; try knset; + rtoProp; intuition eauto). + all:try eapply KernameSet.union_spec in H0; intuition eauto. + - apply KernameSet.singleton_spec in H0. subst s. + destruct (EGlobalEnv.lookup_env Σ kn) eqn:E. + destruct g => //. destruct c => //. cbn in H1. + now eapply lookup_env_In in E. easy. + - destruct i. rewrite knset_in_fold_left in H0. + destruct H0. + apply KernameSet.singleton_spec in H0. subst kn. + destruct (EGlobalEnv.lookup_env Σ _) eqn:E. + destruct g => //. cbn in E. + now eapply lookup_env_In in E. easy. + destruct H0 as [a []]. + eapply All_In in X as [X]; tea. specialize (X k). + destruct EWellformed.cstr_as_blocks. move/andP: H1 => [_ wf]. + solve_all. eapply All_In in wf as [wf]; tea. now specialize (X wf kn). + destruct args => //. + - destruct p. rewrite KernameSet.union_spec knset_in_fold_left in H0. + destruct H0. + apply KernameSet.singleton_spec in H0. subst kn. cbn in *. + destruct (EGlobalEnv.lookup_env Σ _) eqn:E. + destruct g => //. + now eapply lookup_env_In in E. easy. clear H1. + destruct H0. now apply (IHt k H3). + destruct H0 as [a []]. + eapply All_In in X as [X]; tea. + solve_all. eapply All_In in H2 as [wf]; tea. now specialize (X _ wf kn). + - apply KernameSet.singleton_spec in H3. subst kn. cbn in *. + destruct (EGlobalEnv.lookup_env Σ _) eqn:E. + destruct g => //. + now eapply lookup_env_In in E. easy. + - rewrite knset_in_fold_left in H0; destruct H0; [knset|]. + destruct H0 as [a []]. + unfold EWellformed.wf_fix_gen in H1. move/andP: H1 => [_ hm]. solve_all. + eapply All_In in hm as [hm]; tea. intuition auto. + eapply (a0 (#|m| + k)) => //. + - rewrite knset_in_fold_left in H0; destruct H0; [knset|]. + destruct H0 as [a []]. + unfold EWellformed.wf_fix_gen in H1. move/andP: H1 => [_ hm]. solve_all. + eapply All_In in hm as [hm]; tea. intuition auto. + eapply (a0 (#|m| + k)) => //. +Qed. + +Lemma knset_mem_spec kn s : reflect (KernameSet.In kn s) (KernameSet.mem kn s). +Proof. + destruct (KernameSet.mem_spec s kn). + destruct (KernameSet.mem kn s). + - now constructor. + - constructor. intros hin. now apply H0 in hin. +Qed. + +Definition constant_decls_deps decl := + match decl with + | {| EAst.cst_body := Some b |} => term_global_deps b + | _ => KernameSet.empty + end. + +Definition global_decl_deps decl := + match decl with + | EAst.ConstantDecl cst => constant_decls_deps cst + | _ => KernameSet.empty + end. + +Fixpoint global_deps (Σ : EAst.global_declarations) deps := + match Σ with + | [] => deps + | (kn, decl) :: decls => + if KernameSet.mem kn deps then + global_deps decls + (KernameSet.union deps (global_decl_deps decl)) + else global_deps decls deps + end. + +Lemma global_deps_union Σ deps deps' : + global_deps Σ (KernameSet.union deps deps') =kn + KernameSet.union (global_deps Σ deps) (global_deps Σ deps'). +Proof. + induction Σ in deps, deps' |- *; auto. + - reflexivity. + - destruct a. cbn -[KernameSet.mem]. + destruct (knset_mem_spec k deps'). + * case: (knset_mem_spec k _). + move/KernameSet.union_spec => hin. + destruct hin. + eapply KernameSet.mem_spec in H. rewrite H. + rewrite !IHΣ. knset. + ** case: (knset_mem_spec k _); intros hin'; + rewrite !IHΣ; knset. + ** intros hin'. + case: (knset_mem_spec k _); intros hin''; + rewrite !IHΣ; knset. + * case: (knset_mem_spec k _); intros hin''. + case: (knset_mem_spec k _); intros hin'''. + rewrite !IHΣ. knset. + rewrite !IHΣ. knset. + case: (knset_mem_spec k _); intros hin'''. + rewrite !IHΣ. knset. + rewrite !IHΣ. knset. +Qed. + +Lemma in_global_deps deps Σ : + KernameSet.Subset deps (global_deps Σ deps). +Proof. + induction Σ => //. + destruct a as [kn [[[]]|]]; cbn; eauto; + case: (knset_mem_spec kn _); intros hin'''; + rewrite ?global_deps_union; + intros h hin; rewrite ?KernameSet.union_spec; try left; knset. +Qed. + +Lemma global_deps_subset Σ deps deps' : + KernameSet.Subset deps deps' -> + KernameSet.Subset (global_deps Σ deps) (global_deps Σ deps'). +Proof. + induction Σ in deps, deps' |- *. + - cbn. auto. + - destruct a; cbn. + intros sub. + case: (knset_mem_spec k deps) => hin. + eapply sub in hin. eapply KernameSet.mem_spec in hin. rewrite hin. + rewrite !global_deps_union. + specialize (IHΣ _ _ sub). knset. + case: (knset_mem_spec k deps') => hin'. + rewrite !global_deps_union. + specialize (IHΣ _ _ sub). knset. + now eapply IHΣ. +Qed. + +Lemma wf_global_decl_deps {efl : EWellformed.EEnvFlags} Σ d : + EWellformed.wf_global_decl Σ d -> + forall kn, KernameSet.In kn (global_decl_deps d) -> In kn (map fst Σ). +Proof. + intros wf kn hin. + destruct d as [[[]]|] => //; cbn in hin. + * eapply term_global_deps_fresh in hin. exact hin. cbn in wf; tea. + * knset. + * knset. +Qed. + +Lemma lookup_global_deps {efl : EWellformed.EEnvFlags} Σ kn decl : + EWellformed.wf_glob Σ -> + EGlobalEnv.lookup_env Σ kn = Some decl -> + forall kn, KernameSet.In kn (global_decl_deps decl) -> In kn (map fst Σ). +Proof. + induction 1 in kn, decl |- *. + - cbn => //. + - cbn. + case: (eqb_spec kn kn0) => heq. + subst kn0; intros [= <-]. + intros kn' hkn'. + * eapply wf_global_decl_deps in H0; tea. now right. + * intros hl kn' kin. + eapply IHwf_glob in hl; tea. now right. +Qed. + +Lemma fresh_global_In kn Σ : EGlobalEnv.fresh_global kn Σ <-> ~ In kn (map fst Σ). +Proof. + split. + - intros fr. + eapply (Forall_map (fun x => x <> kn) fst) in fr. + intros hin. + now eapply PCUICWfUniverses.Forall_In in fr; tea. + - intros nin. + red. eapply In_Forall. intros kn' hin <-. + eapply nin. now eapply (in_map fst). +Qed. + +Lemma global_deps_kn {fl :EWellformed.EEnvFlags} Σ kn decl : + EWellformed.wf_glob Σ -> + EGlobalEnv.declared_constant Σ kn decl -> + KernameSet.Subset (global_deps Σ (constant_decls_deps decl)) (global_deps Σ (KernameSet.singleton kn)). +Proof. + induction 1 in kn, decl |- *. + - cbn. unfold EGlobalEnv.declared_constant. cbn => //. + - cbn. unfold EGlobalEnv.declared_constant. cbn => //. + destruct (eqb_spec kn kn0). subst kn0. cbn. intros [= ->]. + + case: (knset_mem_spec kn _) => hin. + * eapply wf_global_decl_deps in H0; tea. + now eapply fresh_global_In in H1. + * case: (knset_mem_spec kn _) => hin'. + rewrite !global_deps_union. cbn. knset. + knset. + + + intros hl. + case: (knset_mem_spec kn0 _) => hin. + * elimtype False. + eapply lookup_global_deps in hl; tea. + now eapply fresh_global_In in H1. + * case: (knset_mem_spec kn0 _). + ** move/KernameSet.singleton_spec. intros <-. congruence. + ** intros hnin. now eapply IHwf_glob. +Qed. + +Lemma term_global_deps_mkApps f args : + term_global_deps (EAst.mkApps f args) =kn + KernameSet.union (term_global_deps f) (terms_global_deps args). +Proof. + induction args in f |- *. + - cbn; knset. + - cbn. rewrite IHargs. cbn. + intros kn. + rewrite !KernameSet.union_spec !knset_in_fold_left KernameSet.union_spec. + intuition auto. +Qed. + +From Coq Require Import Morphisms. + +#[export] Instance global_deps_proper : Proper (eq ==> KernameSet.Equal ==> KernameSet.Equal) global_deps. +Proof. + intros ? ? -> kns kns' eq. + induction y in kns, kns', eq |- *; cbn; auto; try knset. + destruct a. setoid_rewrite eq. destruct KernameSet.mem. rewrite IHy; [|reflexivity]. + now rewrite eq. now apply IHy. +Qed. + +#[export] Instance global_deps_proper_subset : Proper (eq ==> KernameSet.Subset ==> KernameSet.Subset) global_deps. +Proof. + intros ? ? -> kns kns' eq. + induction y in kns, kns', eq |- *; cbn; auto; try knset. + destruct a. destruct (knset_mem_spec k kns). eapply eq in i. + eapply KernameSet.mem_spec in i. rewrite i. eapply IHy. knset. + destruct (knset_mem_spec k kns'). specialize (IHy _ _ eq). + rewrite global_deps_union. knset. + eauto. +Qed. + +Lemma term_global_deps_substl defs body : + KernameSet.Subset (term_global_deps (ECSubst.substl defs body)) + (KernameSet.union (terms_global_deps defs) (term_global_deps body)). +Proof. + intros kn. rewrite KernameSet.union_spec. + unfold ECSubst.substl. + induction defs in body |- *; cbn; auto. + - intros hin. eapply IHdefs in hin as []; eauto. left. + rewrite knset_in_fold_left. right. + rewrite knset_in_fold_left in H. + intuition. knset. + rewrite knset_in_fold_left. + eapply term_global_deps_csubst in H. + intuition knset. +Qed. + +Lemma terms_global_deps_rev l : terms_global_deps (List.rev l) =kn terms_global_deps l. +Proof. + intros kn. + rewrite !knset_in_fold_left. intuition auto; try knset. + - destruct H0 as [x []]; right. exists x. now eapply In_rev in H. + - destruct H0 as [x []]; right. exists x. now eapply In_rev in H. +Qed. + +Lemma In_skipn {A} n (l : list A) : forall x, In x (skipn n l) -> In x l. +Proof. + induction l in n |- *; destruct n => //. + rewrite skipn_S. intros x hin; right; eauto. +Qed. + +Lemma terms_global_deps_skipn n l : KernameSet.Subset (terms_global_deps (List.skipn n l)) (terms_global_deps l). +Proof. + intros kn. + rewrite !knset_in_fold_left. intuition auto; try knset. + destruct H0 as [x []]; right. eapply In_skipn in H. now exists x. +Qed. + + +Lemma term_global_deps_iota_red pars args br : + KernameSet.Subset (term_global_deps (EGlobalEnv.iota_red pars args br)) + (KernameSet.union (terms_global_deps args) (term_global_deps br.2)). +Proof. + intros kn hin. + eapply term_global_deps_substl in hin. + rewrite !KernameSet.union_spec in hin *. intuition auto. left. + rewrite terms_global_deps_rev in H. + now apply terms_global_deps_skipn in H. +Qed. + +Lemma fold_left_eq {A} f acc l : + fold_left (fun acc (x : A) => KernameSet.union (f x) acc) l acc =kn + KernameSet.union (fold_left (fun acc x => KernameSet.union (f x) acc) l KernameSet.empty) acc. +Proof. + induction l in acc |- *; cbn. knset. + rewrite IHl. rewrite (IHl (KernameSet.union _ _)). knset. +Qed. + +Lemma term_global_deps_repeat t n : KernameSet.Subset (terms_global_deps (repeat t n)) (term_global_deps t). +Proof. + induction n; cbn; auto; try knset. + intros kn hin. + rewrite fold_left_eq in hin. + rewrite KernameSet.union_spec in hin. intuition auto. + rewrite KernameSet.union_spec in H; intuition auto. + knset. +Qed. + +Notation terms_global_deps_gen f l := + (fold_left + (fun (acc : KernameSet.t) (x : _) => + KernameSet.union (term_global_deps (f x)) acc) l KernameSet.empty). + +Lemma term_global_deps_fix_subst mfix : + KernameSet.Subset (terms_global_deps (EGlobalEnv.fix_subst mfix)) (terms_global_deps_gen EAst.dbody mfix). +Proof. + unfold EGlobalEnv.fix_subst. + induction #|mfix|; cbn. + - knset. + - rewrite fold_left_eq. + intros kn. rewrite KernameSet.union_spec. + intros []. + + now eapply IHn in H. + + knset. +Qed. + +Lemma term_global_deps_cunfold_fix mfix idx n f : + EGlobalEnv.cunfold_fix mfix idx = Some (n, f) -> + KernameSet.Subset (term_global_deps f) (terms_global_deps_gen EAst.dbody mfix). +Proof. + unfold EGlobalEnv.cunfold_fix. + destruct nth_error eqn:E => //. + intros [= <- <-]. + intros kn hin. + eapply term_global_deps_substl in hin. + rewrite KernameSet.union_spec in hin. + destruct hin. + now eapply term_global_deps_fix_subst in H. + eapply nth_error_In in E. + rewrite knset_in_fold_left. right. now exists d. +Qed. + +Lemma term_global_deps_cofix_subst mfix : + KernameSet.Subset (terms_global_deps (EGlobalEnv.cofix_subst mfix)) (terms_global_deps_gen EAst.dbody mfix). +Proof. + unfold EGlobalEnv.cofix_subst. + induction #|mfix|; cbn. + - knset. + - rewrite fold_left_eq. + intros kn. rewrite KernameSet.union_spec. + intros []. + + now eapply IHn in H. + + knset. +Qed. + +Lemma term_global_deps_cunfold_cofix mfix idx n f : + EGlobalEnv.cunfold_cofix mfix idx = Some (n, f) -> + KernameSet.Subset (term_global_deps f) (terms_global_deps_gen EAst.dbody mfix). +Proof. + unfold EGlobalEnv.cunfold_cofix. + destruct nth_error eqn:E => //. + intros [= <- <-]. + intros kn hin. + eapply term_global_deps_substl in hin. + rewrite KernameSet.union_spec in hin. + destruct hin. + now eapply term_global_deps_cofix_subst in H. + eapply nth_error_In in E. + rewrite knset_in_fold_left. right. now exists d. +Qed. + +Lemma global_deps_empty Σ : global_deps Σ KernameSet.empty = KernameSet.empty. +Proof. + induction Σ; cbn; auto. + destruct a as [kn d]. + destruct (knset_mem_spec kn KernameSet.empty). + knset. auto. +Qed. + +Lemma eval_global_deps {fl : EWcbvEval.WcbvFlags} {efl : EWellformed.EEnvFlags} Σ t t' : + EWellformed.wf_glob Σ -> + Σ ⊢ t ⇓ t' -> KernameSet.Subset (global_deps Σ (term_global_deps t')) (global_deps Σ (term_global_deps t)). +Proof. + intros wf. + induction 1 using EWcbvEval.eval_ind; cbn; rewrite ?global_deps_union; try knset. + - cbn in IHeval1. + epose proof (term_global_deps_csubst a' 0 b). + eapply (global_deps_subset Σ) in H2. + rewrite global_deps_union in H2. + knset. + - epose proof (term_global_deps_csubst b0' 0 b1). + eapply (global_deps_subset Σ) in H1. + rewrite global_deps_union in H1. + knset. + - rewrite term_global_deps_mkApps in IHeval1. + rewrite global_deps_union /= in IHeval1. destruct ind; cbn in *. + intros kn hr. + eapply IHeval2 in hr. + eapply global_deps_proper_subset in hr. + 3:eapply (term_global_deps_iota_red pars args br). 2:reflexivity. + rewrite global_deps_union KernameSet.union_spec in hr. + specialize (IHeval1 kn); rewrite !KernameSet.union_spec in IHeval1 *. + right. + destruct hr. + eapply global_deps_proper_subset. reflexivity. + intros kn'. rewrite knset_in_fold_left. intros; left; eauto. + now eapply IHeval1. + rewrite fold_left_eq global_deps_union KernameSet.union_spec. left. + eapply global_deps_proper_subset. reflexivity. + intros kn'. rewrite knset_in_fold_left. intros; right. + eapply nth_error_In in e2. exists br; split. auto. exact H2. + exact H1. + - intros kn hr. + eapply IHeval2 in hr. + eapply global_deps_proper_subset in hr; [|reflexivity|]. + 2:eapply term_global_deps_iota_red. + rewrite global_deps_union KernameSet.union_spec in hr. + rewrite KernameSet.union_spec. + right. + eapply global_deps_proper_subset. reflexivity. + intros kn'. rewrite fold_left_eq. intros h; exact h. + rewrite global_deps_union KernameSet.union_spec. + destruct hr. right. eapply IHeval1. cbn. destruct ind. + now rewrite fold_left_eq global_deps_union KernameSet.union_spec. + left. + eapply global_deps_proper_subset. reflexivity. + intros kn'. rewrite knset_in_fold_left. intros; right. + eapply nth_error_In in e2. exists br; split. auto. exact H2. + exact H1. + - intros kn hr. eapply IHeval2 in hr. + rewrite KernameSet.union_spec. + eapply global_deps_proper_subset in hr; [|reflexivity|]. + 2:eapply term_global_deps_substl. + right. rewrite fold_left_eq global_deps_union KernameSet.union_spec. + subst brs. cbn. left. + rewrite global_deps_union KernameSet.union_spec in hr. destruct hr. + eapply global_deps_proper_subset. reflexivity. 2:exact H1. + setoid_rewrite term_global_deps_repeat. cbn. knset. + now rewrite global_deps_union KernameSet.union_spec. + - intros kn hin. eapply IHeval3 in hin. + cbn in hin. rewrite global_deps_union KernameSet.union_spec in hin. + rewrite term_global_deps_mkApps global_deps_union KernameSet.union_spec in hin. + rewrite KernameSet.union_spec. intuition auto. + + left. eapply IHeval1. + rewrite term_global_deps_mkApps global_deps_union KernameSet.union_spec. left. + cbn. eapply term_global_deps_cunfold_fix in e1. + now setoid_rewrite e1 in H3. + + left. eapply IHeval1. + rewrite term_global_deps_mkApps global_deps_union KernameSet.union_spec. now right. + - intros kn hin. eapply IHeval3 in hin. + cbn in hin. rewrite global_deps_union KernameSet.union_spec in hin. + rewrite KernameSet.union_spec. intuition auto. + left. eapply IHeval1. cbn. + eapply term_global_deps_cunfold_fix in e0. + now setoid_rewrite e0 in H2. + - intros kn hin. eapply IHeval2 in hin. + cbn in hin. destruct ip. + rewrite global_deps_union KernameSet.union_spec in hin. + rewrite global_deps_union KernameSet.union_spec. intuition auto. right. + rewrite fold_left_eq. rewrite fold_left_eq in H1. + rewrite global_deps_union KernameSet.union_spec in H1. + rewrite global_deps_union KernameSet.union_spec. intuition auto. + rewrite term_global_deps_mkApps global_deps_union KernameSet.union_spec in H2. + right. eapply IHeval1. + rewrite term_global_deps_mkApps global_deps_union KernameSet.union_spec. intuition auto. + left. cbn. + eapply term_global_deps_cunfold_cofix in e0. + now setoid_rewrite e0 in H1. + - intros kn hin. eapply IHeval2 in hin. + cbn in hin. rewrite global_deps_union KernameSet.union_spec in hin. + rewrite KernameSet.union_spec. intuition auto. right. + eapply IHeval1. cbn. + rewrite term_global_deps_mkApps global_deps_union KernameSet.union_spec in H1. + rewrite term_global_deps_mkApps global_deps_union KernameSet.union_spec. + intuition auto. left. + eapply term_global_deps_cunfold_cofix in e0. + now setoid_rewrite e0 in H2. + - cbn. + epose proof (global_deps_kn Σ c decl wf isdecl). + unfold constant_decls_deps in H0. destruct decl => //. + cbn in e. subst cst_body0. knset. + - move=> kn /IHeval2 hin. + rewrite KernameSet.union_spec; right. eapply IHeval1. + rewrite term_global_deps_mkApps global_deps_union KernameSet.union_spec. + right. eapply nth_error_In in e3. + eapply global_deps_proper_subset. reflexivity. 2:exact hin. + intros k h. rewrite knset_in_fold_left. right. now exists a. + - eapply nth_error_In in e3. + move=> kn /IHeval2 hin. + rewrite KernameSet.union_spec. right; apply IHeval1; cbn. + destruct proj_ind. rewrite fold_left_eq. + rewrite global_deps_union KernameSet.union_spec. left. + eapply global_deps_proper_subset. reflexivity. + intros kn'. rewrite knset_in_fold_left. intros; right. 2:tea. + now exists a. + - destruct ind. + rewrite fold_left_eq. rewrite (fold_left_eq _ (KernameSet.singleton _)). + intros kn. rewrite !global_deps_union !KernameSet.union_spec. intuition auto. + left. clear -a iha H0. + induction a; auto. + move: H0. cbn. + rewrite fold_left_eq (fold_left_eq _ (KernameSet.union _ _)). + rewrite !global_deps_union !KernameSet.union_spec. + destruct iha as [sub ih]. specialize (IHa ih). intuition eauto. +Qed. + +Lemma in_global_deps_fresh {efl : EWellformed.EEnvFlags} kn Σ deps : + EWellformed.wf_glob Σ -> + KernameSet.In kn (global_deps Σ deps) -> + EGlobalEnv.fresh_global kn Σ -> + KernameSet.In kn deps. +Proof. + induction Σ in deps |- *. + - now cbn. + - intros wf. destruct a as [kn' d]; cbn. depelim wf. + case: (knset_mem_spec kn' deps) => hin. + * rewrite global_deps_union KernameSet.union_spec. + intros [] fr. + ** depelim fr. now eapply IHΣ. + ** depelim fr. elimtype False. eapply IHΣ in H1; eauto. + destruct d as [[[]]|] eqn:eqd; cbn in H. + + cbn in H1. eapply (term_global_deps_fresh Σ) in H1; tea. cbn in H2. + eapply (Forall_map (fun x => x <> kn) fst) in fr. + now eapply PCUICWfUniverses.Forall_In in fr. + + cbn in H1; knset. + + cbn in H1; knset. + * intros hin' fr. depelim fr. now eapply IHΣ. +Qed. + Local Existing Instance extraction_checker_flags. (* Definition wf_ext_wf Σ : wf_ext Σ -> wf Σ := fst. #[global] @@ -2048,14 +2628,14 @@ Proof. simpl. destruct d'. + destruct KernameSet.mem. cbn. destruct (eqb_spec kn kn') => //. - do 3 eexists. split; [|split;[reflexivity|]]. KernameSetDecide.fsetdec. + do 3 eexists. split; [|split;[reflexivity|]]. knset. auto. - do 3 eexists. split; [|split;[reflexivity|]]. KernameSetDecide.fsetdec. + do 3 eexists. split; [|split;[reflexivity|]]. knset. auto. + destruct KernameSet.mem. cbn. destruct (eqb_spec kn kn') => //. - do 3 eexists. split; [|split;[reflexivity|]]. KernameSetDecide.fsetdec. auto. - do 3 eexists. split; [|split;[reflexivity|]]. KernameSetDecide.fsetdec. auto. + do 3 eexists. split; [|split;[reflexivity|]]. knset. auto. + do 3 eexists. split; [|split;[reflexivity|]]. knset. auto. Qed. Lemma lookup_env_In_map_fst Σ kn decl : lookup_global Σ kn = Some decl -> In kn (map fst Σ). @@ -2074,7 +2654,7 @@ Proof. unfold erase_constant_decl. set (ec := erase_constant_body _ _ _ _). destruct cb. destruct cst_body0. - 2:{ subst ec; cbn. KernameSetDecide.fsetdec. } + 2:{ subst ec; cbn. knset. } intros hc. epose proof (erase_constant_body_correct' (@eq_refl _ (EAst.cst_body ec.1))). subst ec. @@ -2259,9 +2839,9 @@ Proof. case_eq (KernameSet.mem k deps). * cbn. intros hm. destruct (eqb_spec k kn). - ++ subst k. eapply IHdecls. KernameSetDecide.fsetdec. + ++ subst k. eapply IHdecls. knset. ++ eapply IHdecls. - KernameSetDecide.fsetdec. + knset. * intros hnin. eapply IHdecls; tea. + set (eg := erase_global_deps _ _ _ _ _ _); set (eg' := erase_global_deps _ _ _ _ _ _). @@ -2274,12 +2854,6 @@ Proof. * intros hnin. eapply IHdecls; tea. Qed. -Definition constant_decls_deps decl := - match decl with - | {| EAst.cst_body := Some b |} => term_global_deps b - | _ => KernameSet.empty - end. - Fixpoint filter_deps deps decls := match decls with | [] => [] @@ -2320,248 +2894,6 @@ Proof. eapply IHdecls. Qed. -Lemma knset_mem_spec kn s : reflect (KernameSet.In kn s) (KernameSet.mem kn s). -Proof. - destruct (KernameSet.mem_spec s kn). - destruct (KernameSet.mem kn s). - - now constructor. - - constructor. intros hin. now apply H0 in hin. -Qed. - -Definition global_decl_deps decl := - match decl with - | EAst.ConstantDecl cst => constant_decls_deps cst - | _ => KernameSet.empty - end. - -Fixpoint global_deps (Σ : EAst.global_declarations) deps := - match Σ with - | [] => deps - | (kn, decl) :: decls => - if KernameSet.mem kn deps then - global_deps decls - (KernameSet.union deps (global_decl_deps decl)) - else global_deps decls deps - end. - -Infix "=kn" := KernameSet.Equal (at level 90). - -Ltac knset := KernameSetDecide.fsetdec. - -Lemma global_deps_union Σ deps deps' : - global_deps Σ (KernameSet.union deps deps') =kn - KernameSet.union (global_deps Σ deps) (global_deps Σ deps'). -Proof. - induction Σ in deps, deps' |- *; auto. - - reflexivity. - - destruct a. cbn -[KernameSet.mem]. - destruct (knset_mem_spec k deps'). - * case: (knset_mem_spec k _). - move/KernameSet.union_spec => hin. - destruct hin. - eapply KernameSet.mem_spec in H. rewrite H. - rewrite !IHΣ. KernameSetDecide.fsetdec. - ** case: (knset_mem_spec k _); intros hin'; - rewrite !IHΣ; KernameSetDecide.fsetdec. - ** intros hin'. - case: (knset_mem_spec k _); intros hin''; - rewrite !IHΣ; KernameSetDecide.fsetdec. - * case: (knset_mem_spec k _); intros hin''. - case: (knset_mem_spec k _); intros hin'''. - rewrite !IHΣ. knset. - rewrite !IHΣ. knset. - case: (knset_mem_spec k _); intros hin'''. - rewrite !IHΣ. knset. - rewrite !IHΣ. knset. -Qed. - -Lemma in_global_deps deps Σ : - KernameSet.Subset deps (global_deps Σ deps). -Proof. - induction Σ => //. - destruct a as [kn [[[]]|]]; cbn; eauto; - case: (knset_mem_spec kn _); intros hin'''; - rewrite ?global_deps_union; - intros h hin; rewrite ?KernameSet.union_spec; try left; knset. -Qed. - -Lemma term_global_deps_csubst a k b : - KernameSet.Subset (term_global_deps (ECSubst.csubst a k b)) - (KernameSet.union (term_global_deps a) (term_global_deps b)). -Admitted. - -Lemma global_deps_subset Σ deps deps' : - KernameSet.Subset deps deps' -> - KernameSet.Subset (global_deps Σ deps) (global_deps Σ deps'). -Proof. - induction Σ in deps, deps' |- *. - - cbn. auto. - - destruct a; cbn. - intros sub. - case: (knset_mem_spec k deps) => hin. - eapply sub in hin. eapply KernameSet.mem_spec in hin. rewrite hin. - rewrite !global_deps_union. - specialize (IHΣ _ _ sub). knset. - case: (knset_mem_spec k deps') => hin'. - rewrite !global_deps_union. - specialize (IHΣ _ _ sub). knset. - now eapply IHΣ. -Qed. - -Arguments KernameSet.mem : simpl never. - -Lemma lookup_env_In Σ kn d : EGlobalEnv.lookup_env Σ kn = Some d -> In kn (map fst Σ). -Proof. - induction Σ; cbn; auto. - - easy. - - destruct (eqb_spec kn a.1). - intros [= <-]. left; auto. - intros hl; right; auto. -Qed. - - -Lemma term_global_deps_fresh {efl : EWellformed.EEnvFlags} Σ k t : - EWellformed.wellformed Σ k t -> - forall kn, KernameSet.In kn (term_global_deps t) -> In kn (map fst Σ). -Proof. - induction t in k |- * using EInduction.term_forall_list_ind; intros; try cbn in *; try knset; - rtoProp; intuition eauto. - all:try eapply KernameSet.union_spec in H0; intuition eauto. - - apply KernameSet.singleton_spec in H0. subst s. - destruct (EGlobalEnv.lookup_env Σ kn) eqn:E. - destruct g => //. destruct c => //. cbn in H1. - now eapply lookup_env_In in E. easy. - - admit. - - admit. - - admit. - - admit. - - admit. -Admitted. - -Lemma wf_global_decl_deps {efl : EWellformed.EEnvFlags} Σ d : - EWellformed.wf_global_decl Σ d -> - forall kn, KernameSet.In kn (global_decl_deps d) -> In kn (map fst Σ). -Proof. - intros wf kn hin. - destruct d as [[[]]|] => //; cbn in hin. - * eapply term_global_deps_fresh in hin. exact hin. cbn in wf; tea. - * knset. - * knset. -Qed. - -Lemma lookup_global_deps {efl : EWellformed.EEnvFlags} Σ kn decl : - EWellformed.wf_glob Σ -> - EGlobalEnv.lookup_env Σ kn = Some decl -> - forall kn, KernameSet.In kn (global_decl_deps decl) -> In kn (map fst Σ). -Proof. - induction 1 in kn, decl |- *. - - cbn => //. - - cbn. - case: (eqb_spec kn kn0) => heq. - subst kn0; intros [= <-]. - intros kn' hkn'. - * eapply wf_global_decl_deps in H0; tea. now right. - * intros hl kn' kin. - eapply IHwf_glob in hl; tea. now right. -Qed. - -Lemma fresh_global_In kn Σ : EGlobalEnv.fresh_global kn Σ <-> ~ In kn (map fst Σ). -Proof. - split. - - intros fr. - eapply (Forall_map (fun x => x <> kn) fst) in fr. - intros hin. - now eapply PCUICWfUniverses.Forall_In in fr; tea. - - intros nin. - red. eapply In_Forall. intros kn' hin <-. - eapply nin. now eapply (in_map fst). -Qed. - -Lemma global_deps_kn {fl :EWellformed.EEnvFlags} Σ kn decl : - EWellformed.wf_glob Σ -> - EGlobalEnv.declared_constant Σ kn decl -> - KernameSet.Subset (global_deps Σ (constant_decls_deps decl)) (global_deps Σ (KernameSet.singleton kn)). -Proof. - induction 1 in kn, decl |- *. - - cbn. unfold EGlobalEnv.declared_constant. cbn => //. - - cbn. unfold EGlobalEnv.declared_constant. cbn => //. - destruct (eqb_spec kn kn0). subst kn0. cbn. intros [= ->]. - + case: (knset_mem_spec kn _) => hin. - * eapply wf_global_decl_deps in H0; tea. - now eapply fresh_global_In in H1. - * case: (knset_mem_spec kn _) => hin'. - rewrite !global_deps_union. cbn. knset. - knset. - - + intros hl. - case: (knset_mem_spec kn0 _) => hin. - * elimtype False. - eapply lookup_global_deps in hl; tea. - now eapply fresh_global_In in H1. - * case: (knset_mem_spec kn0 _). - ** move/KernameSet.singleton_spec. intros <-. congruence. - ** intros hnin. now eapply IHwf_glob. -Qed. - -Lemma eval_global_deps {fl : EWcbvEval.WcbvFlags} {efl : EWellformed.EEnvFlags} Σ t t' : - EWellformed.wf_glob Σ -> - Σ ⊢ t ⇓ t' -> KernameSet.Subset (global_deps Σ (term_global_deps t')) (global_deps Σ (term_global_deps t)). -Proof. - intros wf. - induction 1 using EWcbvEval.eval_ind; cbn; rewrite ?global_deps_union; try knset. - - cbn in IHeval1. - epose proof (term_global_deps_csubst a' 0 b). - eapply (global_deps_subset Σ) in H2. - rewrite global_deps_union in H2. - knset. - - epose proof (term_global_deps_csubst b0' 0 b1). - eapply (global_deps_subset Σ) in H1. - rewrite global_deps_union in H1. - knset. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - cbn. - epose proof (global_deps_kn Σ c decl wf isdecl). - unfold constant_decls_deps in H0. destruct decl => //. - cbn in e. subst cst_body0. knset. - - admit. -Admitted. - -Lemma in_global_deps_fresh {efl : EWellformed.EEnvFlags} kn Σ deps : - EWellformed.wf_glob Σ -> - KernameSet.In kn (global_deps Σ deps) -> - EGlobalEnv.fresh_global kn Σ -> - KernameSet.In kn deps. -Proof. - induction Σ in deps |- *. - - now cbn. - - intros wf. destruct a as [kn' d]; cbn. depelim wf. - case: (knset_mem_spec kn' deps) => hin. - * rewrite global_deps_union KernameSet.union_spec. - intros [] fr. - ** depelim fr. now eapply IHΣ. - ** depelim fr. elimtype False. eapply IHΣ in H1; eauto. - destruct d as [[[]]|] eqn:eqd; cbn in H. - + cbn in H1. eapply (term_global_deps_fresh Σ) in H1; tea. cbn in H2. - eapply (Forall_map (fun x => x <> kn) fst) in fr. - now eapply PCUICWfUniverses.Forall_In in fr. - + cbn in H1; knset. - + cbn in H1; knset. - * intros hin' fr. depelim fr. now eapply IHΣ. -Qed. - -Lemma global_deps_empty Σ : global_deps Σ KernameSet.empty = KernameSet.empty. -Proof. - induction Σ; cbn; auto. - destruct a as [kn d]. - destruct (knset_mem_spec kn KernameSet.empty). - knset. auto. -Qed. Lemma filter_deps_filter {efl : EWellformed.EEnvFlags} deps Σ : EWellformed.wf_glob Σ -> @@ -2659,6 +2991,24 @@ Proof. + eapply IHdecls => //. Qed. +Lemma erase_global_fresh X_type kn X decls normalization_in heq : + let Σ' := @erase_global X_type X decls normalization_in heq in + PCUICAst.fresh_global kn decls -> + fresh_global kn Σ'. +Proof. + cbn. + revert X normalization_in heq. + induction decls; [cbn; auto|]. + - intros. red. constructor. + - destruct a as [kn' d]. intros. depelim H. + cbn in H, H0. cbn. + destruct d as []; simpl. + + cbn [EGlobalEnv.closed_env forallb]. cbn. + constructor => //. eapply IHdecls => //. + + constructor; auto. + eapply IHdecls => //. +Qed. + From MetaCoq.Erasure Require Import EEtaExpandedFix. Lemma erase_brs_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ p ts wt : @@ -3145,24 +3495,24 @@ Proof. pose proof (abstract_env_ext_exists X') as [[Σmake wfmake]]. destruct (wfΣ _ wf), (hcb _ wfmake). red in X1. destruct EAst.cst_body eqn:hb => /= //. - eapply (erase_constant_body_correct'') in hb; eauto. - destruct hb as [[t0 [T [[] ?]]]]. rewrite e in i. exact i. + eapply (erase_constant_body_correct_glob) in hb; eauto. + destruct hb as [[t0 [T [[] ?]]]]. exact i. Qed. -Lemma erase_global_ind_decl_wf_glob {X_type X} {deps decls normalization_in kn m} heq : +Lemma erase_global_ind_decl_wf_glob' {X_type X} {decls normalization_in kn m} heq : (forall Σ : global_env, abstract_env_rel X Σ -> on_inductive cumulSpec0 (lift_typing typing) (Σ, ind_universes m) kn m) -> let m' := erase_mutual_inductive_body m in - let Σ' := erase_global_deps X_type deps X decls normalization_in heq in - @wf_global_decl all_env_flags Σ'.1 (EAst.InductiveDecl m'). + let Σ' := @erase_global X_type X decls normalization_in heq in + @wf_global_decl all_env_flags Σ' (EAst.InductiveDecl m'). Proof. set (m' := erase_mutual_inductive_body m). - set (Σ' := erase_global_deps _ _ _ _ _ _). simpl. + set (Σ' := erase_global _ _ _). simpl. intros oni. pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. pose proof (abstract_env_wf _ wf) as [wfΣ]. assert (erases_mutual_inductive_body m (erase_mutual_inductive_body m)). { eapply (erases_mutual (mdecl:=kn)); tea. } - eapply (erases_mutual_inductive_body_wf (univs := Σ.(universes)) (retro := Σ.(retroknowledge)) (Σ := decls) (kn := kn) (Σ' := Σ'.1)) in H; tea. + eapply (erases_mutual_inductive_body_wf (univs := Σ.(universes)) (retro := Σ.(retroknowledge)) (Σ := decls) (kn := kn) (Σ' := Σ')) in H; tea. rewrite -(heq _ wf). now destruct Σ. Qed. @@ -3181,26 +3531,23 @@ Proof. epose proof (abstract_env_exists Xpop) as [[Σpop wfpop]]; pose proof (abstract_env_wf _ wfpop) as [wfΣpop]. + constructor. eapply IHdecls => //; eauto. - eapply erase_global_cst_decl_wf_glob; auto. - eapply erase_global_deps_fresh; auto. + eapply erase_global_cst_wf_glob; auto. + eapply erase_global_fresh; auto. destruct wfΣ. destruct wfΣpop. rewrite (heq _ wf) in o0. depelim o0. now destruct o3. - + cbn. eapply IHdecls; eauto. + constructor. eapply IHdecls; eauto. destruct wfΣ as [[onu ond]]. rewrite (heq _ wf) in o. depelim o. destruct o0. - eapply (erase_global_ind_decl_wf_glob (kn:=kn)); tea. + eapply (erase_global_ind_decl_wf_glob' (kn:=kn)); tea. intros. rewrite (abstract_env_irr _ H wfpop). unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wf wfpop) as [? ?]. - {intros; now eexists. } + { intros; now eexists. } destruct Σpop, Σ; cbn in *. now subst. - eapply erase_global_deps_fresh. + eapply erase_global_fresh. destruct wfΣ as [[onu ond]]. rewrite (heq _ wf) in o. depelim o. now destruct o0. - + eapply IHdecls; eauto. Qed. - Lemma extends_cons d Σ Σ' : extends Σ Σ' -> extends (d :: Σ) (d :: Σ'). Proof. intros ext kn decl. @@ -3269,9 +3616,9 @@ Proof. intros. move: (erase_global_deps_wf_glob X_type X (term_global_deps t') decls nin prf). subst eg eg'. - rewrite -!erase_global_deps. + rewrite !erase_global_deps_erase_global. assert (wfer : wf_glob (efl := all_env_flags) (@erase_global X_type X decls nin prf)). - { admit. } + { eapply erase_global_wf_glob. } rewrite !(filter_deps_filter (efl := all_env_flags)) //. cbn. intros wf. @@ -3280,7 +3627,7 @@ Proof. rewrite /is_true !KernameSet.mem_spec. revert kn. eapply EWcbvEval.weakening_eval_env in H. 2:exact wfer. now eapply eval_global_deps in H. - rewrite -!erase_global_deps (filter_deps_filter (efl := all_env_flags)) //. + rewrite !erase_global_deps_erase_global (filter_deps_filter (efl := all_env_flags)) //. now eapply extends_filter. Qed. @@ -4063,58 +4410,8 @@ Lemma erase_wellformed_fast (efl := all_env_flags) Proof using Type. intros. cbn. unfold erase_global_fast. erewrite erase_global_deps_fast_spec. - eapply erase_wellformed. + eapply erase_wellformed_deps. Unshelve. all: eauto. Qed. End EraseGlobalFast. - -#[local] Instance extends_trans : Transitive extends. -Proof. - Admitted. - -Lemma extends_erase_pcuic_program {wfl : EWcbvEval.WcbvFlags} X_type X decls nin pf : - forall Σ et et', - EWcbvEval.eval Σ et' et -> - Σ = (@erase_global_deps X_type (term_global_deps et') X decls nin pf).1 -> - EGlobalEnv.extends - (@erase_global_deps X_type (term_global_deps et) X decls nin pf).1 - (@erase_global_deps X_type (term_global_deps et') X decls nin pf).1. -Proof. - intros Σ et et'. - induction 1; intros eqΣ. - - cbn. eapply lookup_erase_global. KernameSetDecide.fsetdec. - - cbn in *. etransitivity; tea. - transitivity ((erase_global_deps X_type (term_global_deps (ECSubst.csubst a 0 f0)) X decls nin pf).1). - admit. admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - admit. - - cbn. etransitivity; tea. - intros kn decl'. - intros hl. epose proof (lookup_env_erase_decl X_type X). in hl. - - - -Lemma extends_erase_pcuic_program X_type X t t' Γ : - (forall Σ, Σ ∼ X -> PCUICWcbvEval.eval Σ t' t) -> - forall Xext norm decls nin pf wt wt', - let et := @erase X_type Xext norm Γ t wt in - let et' := @erase X_type Xext norm Γ t' wt' in - EGlobalEnv.extends - (@erase_global_deps X_type (term_global_deps et) X decls nin pf).1 - (@erase_global_deps X_type (term_global_deps et') X decls nin pf).1. -Proof. - intros ev Xext norm decls nin pf wt wt' et et'. - intros kn decl. - pose proof (abstract_env_exists X) as [[Σ H]]. - specialize (ev _ H). - induction ev. - -Admitted. From 80055db424031c287f6cbf1cbc6efd3dd29e7a0a Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 31 Aug 2023 12:08:47 +0200 Subject: [PATCH 44/61] Finished proof of firstorder correctness for erase_pcuic_program --- erasure-plugin/theories/ETransform.v | 30 +-- erasure-plugin/theories/Erasure.v | 2 +- erasure-plugin/theories/ErasureCorrectness.v | 71 +++++- erasure/theories/ErasureFunction.v | 244 +++++++++++++------ 4 files changed, 247 insertions(+), 100 deletions(-) diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index acef38110..c8c4584c1 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -24,11 +24,11 @@ Definition build_wf_env_from_env {cf : checker_flags} (Σ : global_env_map) (wf Notation NormalizationIn_erase_pcuic_program_1 p - := (@PCUICTyping.wf_ext config.extraction_checker_flags p.1 -> PCUICSN.NormalizationIn (cf:=config.extraction_checker_flags) (no:=PCUICSN.extraction_normalizing) p.1) + := (@PCUICTyping.wf_ext config.extraction_checker_flags p -> PCUICSN.NormalizationIn (cf:=config.extraction_checker_flags) (no:=PCUICSN.extraction_normalizing) p) (only parsing). Notation NormalizationIn_erase_pcuic_program_2 p - := (@PCUICTyping.wf_ext config.extraction_checker_flags p.1 -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn (cf:=config.extraction_checker_flags) (no:=PCUICSN.extraction_normalizing) p.1) + := (@PCUICTyping.wf_ext config.extraction_checker_flags p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn (cf:=config.extraction_checker_flags) (no:=PCUICSN.extraction_normalizing) p) (only parsing). (* TODO: Where should this go? *) @@ -70,11 +70,11 @@ Qed. #[local] Lemma erase_pcuic_program_normalization_helper (cf := config.extraction_checker_flags) (no := PCUICSN.extraction_normalizing) - {guard : abstract_guard_impl} (p : pcuic_program) + {guard : abstract_guard_impl} (p : global_env_ext_map) {normalization_in : NormalizationIn_erase_pcuic_program_1 p} {normalization_in_adjust_universes : NormalizationIn_erase_pcuic_program_2 p} - (wfΣ : ∥ PCUICTyping.wf_ext (H := config.extraction_checker_flags) p.1 ∥) - : (let wfe := build_wf_env_from_env p.1.1 (map_squash (PCUICTyping.wf_ext_wf _) (wfΣ : ∥ PCUICTyping.wf_ext (H := config.extraction_checker_flags) p.1 ∥)) in + (wfΣ : ∥ PCUICTyping.wf_ext (H := config.extraction_checker_flags) p ∥) + : (let wfe := build_wf_env_from_env p.1 (map_squash (PCUICTyping.wf_ext_wf _) (wfΣ : ∥ PCUICTyping.wf_ext (H := config.extraction_checker_flags) p ∥)) in forall n : nat, n < #|PCUICAst.PCUICEnvironment.declarations p.1| -> forall kn cb pf, @@ -87,7 +87,7 @@ Qed. (@optimized_abstract_env_impl extraction_checker_flags _) wfe (PCUICAst.PCUICEnvironment.cst_universes cb) pf -> PCUICSN.NormalizationIn Σ) /\ - (let wfe := build_wf_env_from_env p.1.1 (map_squash (PCUICTyping.wf_ext_wf _) (wfΣ : ∥ PCUICTyping.wf_ext (H := config.extraction_checker_flags) p.1 ∥)) in + (let wfe := build_wf_env_from_env p.1 (map_squash (PCUICTyping.wf_ext_wf _) (wfΣ : ∥ PCUICTyping.wf_ext (H := config.extraction_checker_flags) p ∥)) in forall n : nat, n < #|PCUICAst.PCUICEnvironment.declarations p.1| -> let X' := @@ -138,8 +138,8 @@ Proof. Qed. Program Definition erase_pcuic_program {guard : abstract_guard_impl} (p : pcuic_program) - {normalization_in : NormalizationIn_erase_pcuic_program_1 p} - {normalization_in_adjust_universes : NormalizationIn_erase_pcuic_program_2 p} + {normalization_in : NormalizationIn_erase_pcuic_program_1 p.1} + {normalization_in_adjust_universes : NormalizationIn_erase_pcuic_program_2 p.1} (wfΣ : ∥ PCUICTyping.wf_ext (H := config.extraction_checker_flags) p.1 ∥) (wt : ∥ ∑ T, PCUICTyping.typing (H := config.extraction_checker_flags) p.1 [] p.2 T ∥) : eprogram_env := let wfe := build_wf_env_from_env p.1.1 (map_squash (PCUICTyping.wf_ext_wf _) wfΣ) in @@ -150,14 +150,14 @@ Program Definition erase_pcuic_program {guard : abstract_guard_impl} (p : pcuic_ (EAstUtils.term_global_deps t) wfe (p.1.(PCUICAst.PCUICEnvironment.declarations)) _ in (EEnvMap.GlobalContextMap.make Σ'.1 _, t). -Next Obligation. unshelve edestruct erase_pcuic_program_normalization_helper; cbn in *; eauto. Qed. +Next Obligation. unshelve edestruct erase_pcuic_program_normalization_helper; cbn in *; eauto. Defined. Next Obligation. eapply wf_glob_fresh. eapply ErasureFunction.erase_global_fast_wf_glob. unshelve edestruct erase_pcuic_program_normalization_helper; cbn in *; eauto. -Qed. +Defined. -Obligation Tactic := idtac. +Local Obligation Tactic := idtac. Import Extract. @@ -191,7 +191,7 @@ Proof. - eapply EEtaExpanded.isEtaExpFix_isEtaExp. now eapply EEtaExpandedFix.expanded_isEtaExp. Qed. -Obligation Tactic := try solve [ eauto ]. +Local Obligation Tactic := try solve [ eauto ]. Program Definition erase_transform {guard : abstract_guard_impl} : Transform.t pcuic_program eprogram_env PCUICAst.term EAst.term eval_pcuic_program (eval_eprogram_env EWcbvEval.default_wcbv_flags) := @@ -199,8 +199,8 @@ Program Definition erase_transform {guard : abstract_guard_impl} : Transform.t p pre p := ∥ wt_pcuic_program (cf := config.extraction_checker_flags) p ∥ /\ PCUICEtaExpand.expanded_pcuic_program p - /\ NormalizationIn_erase_pcuic_program_1 p - /\ NormalizationIn_erase_pcuic_program_2 p ; + /\ NormalizationIn_erase_pcuic_program_1 p.1 + /\ NormalizationIn_erase_pcuic_program_2 p.1 ; transform p hp := let nhs := proj2 (proj2 hp) in @erase_program guard p (proj1 nhs) (proj2 nhs) (proj1 hp) ; post p := [/\ wf_eprogram_env all_env_flags p & EEtaExpandedFix.expanded_eprogram_env p]; @@ -316,7 +316,7 @@ Qed. End PCUICEnv. -Obligation Tactic := idtac. +#[local] Obligation Tactic := idtac. (** This transformation is the identity on terms but changes the evaluation relation to one where fixpoints are not guarded. It requires eta-expanded fixpoints and evaluation diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index b44821921..4467c1dae 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -23,7 +23,7 @@ Local Open Scope string_scope2. Import Transform. -Obligation Tactic := program_simpl. +Local Obligation Tactic := program_simpl. #[local] Existing Instance extraction_checker_flags. #[local] Existing Instance PCUICSN.extraction_normalizing. diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 437a1ab61..4558691a0 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -262,22 +262,69 @@ Proof. now eapply rebuild_wf_env_irr. Qed. -Lemma extends_erase_pcuic_program {guard : abstract_guard_impl} (Σ : global_env_ext_map) t t' nin nin' nin0 nin0' wf wf' ty ty' : - PCUICWcbvEval.eval Σ t' t -> - EGlobalEnv.extends - (@erase_pcuic_program guard (Σ, t) nin nin' wf ty).1 - (@erase_pcuic_program guard (Σ, t') nin0 nin0' wf' ty').1. +From MetaCoq.Erasure Require Import Erasure Extract ErasureFunction. +From MetaCoq.PCUIC Require Import PCUICTyping. + +Lemma extends_erase_pcuic_program (efl := EWcbvEval.default_wcbv_flags) {guard : abstract_guard_impl} (Σ : global_env_ext_map) t v nin nin' nin0 nin0' + wf wf' ty ty' i u args : + PCUICWcbvEval.eval Σ t v -> + axiom_free Σ -> + Σ ;;; [] |- t : PCUICAst.mkApps (PCUICAst.tInd i u) args -> + @PCUICFirstorder.firstorder_ind Σ (PCUICFirstorder.firstorder_env Σ) i -> + let pt := @erase_pcuic_program guard (Σ, t) nin0 nin0' wf' ty' in + let pv := @erase_pcuic_program guard (Σ, v) nin nin' wf ty in + EGlobalEnv.extends pv.1 pt.1 /\ ∥ eval pt.1 pt.2 pv.2 ∥. Proof. - intros ev. - cbn. + intros ev axf ht fo. + cbn -[erase_pcuic_program]. + unfold erase_pcuic_program. + set (prf0 := (fun (Σ0 : PCUICAst.PCUICEnvironment.global_env) => _)). + set (prf1 := (fun (Σ0 : PCUICAst.PCUICEnvironment.global_env) => _)). + set (prf2 := (fun (Σ0 : PCUICAst.PCUICEnvironment.global_env) => _)). + set (prf3 := (fun (Σ0 : PCUICAst.PCUICEnvironment.global_env_ext) => _)). + set (prf4 := (fun (Σ0 : PCUICAst.PCUICEnvironment.global_env_ext) => _)). + set (prf5 := (fun (Σ0 : PCUICAst.PCUICEnvironment.global_env_ext) => _)). + set (prf6 := (fun (Σ0 : PCUICAst.PCUICEnvironment.global_env_ext) => _)). + set (env' := build_wf_env_from_env _ _). + set (env := build_wf_env_from_env _ _). + set (X := PCUICWfEnv.abstract_make_wf_env_ext _ _ _). + set (X' := PCUICWfEnv.abstract_make_wf_env_ext _ _ _). unfold ErasureFunction.erase_global_fast. + set (prf7 := (fun (Σ0 : PCUICAst.PCUICEnvironment.global_env) => _)). set (et := ErasureFunction.erase _ _ _ _ _). set (et' := ErasureFunction.erase _ _ _ _ _). - erewrite ErasureFunction.erase_global_deps_fast_spec. - erewrite ErasureFunction.erase_global_deps_fast_spec. - epose proof ErasureFunction.erase_global_decls_irr. - intros kn decl. -Admitted. + destruct Σ as [Σ ext]. + cbn -[et et' PCUICWfEnv.abstract_make_wf_env_ext] in *. + unshelve (epose proof ErasureFunction.erase_global_deps_fast_erase_global_deps as [norm eq]; + erewrite eq). + { cbn. now intros ? ->. } + unshelve (epose proof ErasureFunction.erase_global_deps_fast_erase_global_deps as [norm' eq']; + erewrite eq'). + { cbn. now intros ? ->. } + set (prf := (fun (Σ0 : PCUICAst.PCUICEnvironment.global_env) => _)). cbn in prf. + rewrite (ErasureFunction.erase_global_deps_irr optimized_abstract_env_impl (EAstUtils.term_global_deps et) env' env _ prf prf). + { cbn. now intros ? ? -> ->. } + clearbody prf0 prf1 prf2 prf3 prf4 prf5 prf6 prf7. + epose proof (ErasureFunction.erase_correct_strong optimized_abstract_env_impl (v:=v) env ext prf2 + (PCUICAst.PCUICEnvironment.declarations Σ) norm' prf prf6 X eq_refl axf ht fo). + pose proof wf as []. + forward H by (eapply Kernames.KernameSet.subset_spec; reflexivity). + forward H by unshelve (eapply PCUICClassification.wcbveval_red; tea). + forward H. { + intros [? hr]. + eapply PCUICNormalization.firstorder_value_irred; tea. cbn. + eapply PCUICFirstorder.firstorder_value_spec; tea. apply X0. constructor. + eapply PCUICClassification.subject_reduction_eval; tea. + eapply PCUICWcbvEval.eval_to_value; tea. } + destruct H as [wt' []]. + split => //. + eapply (ErasureFunction.erase_global_deps_eval optimized_abstract_env_impl env env' ext). + unshelve erewrite (ErasureFunction.erase_irrel_global_env (X_type:=optimized_abstract_env_impl) (t:=v)); tea. + red. intros. split; reflexivity. + sq. unfold et', et. + unshelve erewrite (ErasureFunction.erase_irrel_global_env (X_type:=optimized_abstract_env_impl) (t:=v)); tea. + red. intros. split; reflexivity. +Qed. Section pipeline_theorem. diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 15f45b589..8313423d9 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -1706,37 +1706,73 @@ now pose proof (abstract_pop_decls_correct X decls prop' _ _ HX H). Qed. Lemma erase_global_deps_irr - X_type deps (X:X_type.π1) decls - {normalization_in normalization_in' : NormalizationIn_erase_global_deps X decls} + X_type deps (X X':X_type.π1) decls + {normalization_in normalization_in'} prf prf' : + (forall Σ Σ', Σ ∼ X -> Σ' ∼ X' -> forall tag, primitive_constant Σ tag = primitive_constant Σ' tag) -> erase_global_deps deps X decls (normalization_in:=normalization_in) prf = - erase_global_deps deps X decls (normalization_in:=normalization_in') prf'. + erase_global_deps deps X' decls (normalization_in:=normalization_in') prf'. Proof. - revert X deps normalization_in normalization_in' prf prf'. + revert X X' deps normalization_in normalization_in' prf prf'. induction decls; eauto. intros. destruct a as [kn []]. - cbn. set (ec := erase_constant_decl _ _ _ _). set (ec' := erase_constant_decl _ _ _ _). + pose proof (abstract_env_exists X) as [[envX wfX]]. + pose proof (abstract_env_exists X') as [[envX' wfX']]. + pose proof (abstract_env_exists (abstract_pop_decls X)) as [[env wf]]. + pose proof (abstract_env_exists (abstract_pop_decls X')) as [[env' wf']]. + unshelve epose proof (abstract_pop_decls_correct _ _ _ _ _ wfX wf) as [Hd [Hu Hr]]. shelve. + { intros ? h. rewrite (prf _ h). now eexists. } + unshelve epose proof (abstract_pop_decls_correct _ _ _ _ _ wfX' wf') as [Hd' [Hu' Hr']]. shelve. + { intros ? h. rewrite (prf' _ h). now eexists. } assert (ec = ec') as <-. { subst ec ec'. unfold erase_constant_decl. erewrite erase_constant_body_irrel. reflexivity. - pose proof (abstract_env_exists (abstract_pop_decls X)) as [[env wf]]. red. intros. - pose proof (abstract_make_wf_env_ext_correct (abstract_pop_decls X) (cst_universes c) _ _ _ wf H). pose proof (abstract_make_wf_env_ext_correct (abstract_pop_decls X) (cst_universes c) _ _ _ wf H0). + pose proof (abstract_make_wf_env_ext_correct (abstract_pop_decls X') (cst_universes c) _ _ _ wf' H1). split => //. { intros. - rewrite -(abstract_env_lookup_correct' _ _ H). - rewrite -(abstract_env_lookup_correct' _ _ H0). now rewrite H1 H2. } + rewrite -(abstract_env_lookup_correct' _ _ H0). + rewrite -(abstract_env_lookup_correct' _ _ H1). + move: H4 H5. rewrite /lookup_env H2 H3 /= Hd Hd'. congruence. } intros. - { rewrite (abstract_primitive_constant_correct _ _ _ H). - rewrite (abstract_primitive_constant_correct _ _ _ H0). - now rewrite H1 H2. } } + { rewrite (abstract_primitive_constant_correct _ _ _ H0). + rewrite (abstract_primitive_constant_correct _ _ _ H1). + specialize (H _ _ wfX wfX'). + rewrite H2 H3 /primitive_constant /= -Hr -Hr'. apply H. } } + assert ((forall Σ Σ' : global_env, + Σ ∼ abstract_pop_decls X -> Σ' ∼ abstract_pop_decls X' -> + forall tag : Primitive.prim_tag, primitive_constant Σ tag = primitive_constant Σ' tag)) . + { intros Σ Σ' h h' tag. + pose proof (abstract_env_irr _ wf h). subst env. + pose proof (abstract_env_irr _ wf' h'). subst env'. + specialize (H _ _ wfX wfX'). + rewrite /primitive_constant -Hr -Hr'. apply H. } + specialize (IHdecls (abstract_pop_decls X) (abstract_pop_decls X')). destruct KernameSet.mem; cbn. - f_equal; f_equal. f_equal; eapply IHdecls. - apply IHdecls. apply IHdecls. - - cbn. destruct KernameSet.mem; cbn. f_equal. f_equal. 1-2:f_equal. all:eapply IHdecls. + f_equal; f_equal. f_equal; now eapply IHdecls. + now apply IHdecls. now apply IHdecls. + - cbn. + pose proof (abstract_env_exists X) as [[envX wfX]]. + pose proof (abstract_env_exists X') as [[envX' wfX']]. + pose proof (abstract_env_exists (abstract_pop_decls X)) as [[env wf]]. + pose proof (abstract_env_exists (abstract_pop_decls X')) as [[env' wf']]. + unshelve epose proof (abstract_pop_decls_correct _ _ _ _ _ wfX wf) as [Hd [Hu Hr]]. shelve. + { intros ? h. rewrite (prf _ h). now eexists. } + unshelve epose proof (abstract_pop_decls_correct _ _ _ _ _ wfX' wf') as [Hd' [Hu' Hr']]. shelve. + { intros ? h. rewrite (prf' _ h). now eexists. } + assert ((forall Σ Σ' : global_env, + Σ ∼ abstract_pop_decls X -> Σ' ∼ abstract_pop_decls X' -> + forall tag : Primitive.prim_tag, primitive_constant Σ tag = primitive_constant Σ' tag)) . + { intros Σ Σ' h h' tag. + pose proof (abstract_env_irr _ wf h). subst env. + pose proof (abstract_env_irr _ wf' h'). subst env'. + specialize (H _ _ wfX wfX'). + rewrite /primitive_constant -Hr -Hr'. apply H. } + destruct KernameSet.mem; cbn. f_equal. f_equal. 1-2:f_equal. all:now eapply IHdecls. Qed. Definition global_erased_with_deps (Σ : global_env) (Σ' : EAst.global_declarations) kn := @@ -2867,7 +2903,7 @@ Fixpoint filter_deps deps decls := else filter_deps deps decls end. -Lemma erase_global_deps_erase_global {efl : EWcbvEval.WcbvFlags} (X_type : abstract_env_impl) (X:X_type.π1) deps decls {normalization_in} prf : +Lemma erase_global_deps_erase_global (X_type : abstract_env_impl) (X:X_type.π1) deps decls {normalization_in} prf : (@erase_global_deps X_type deps X decls normalization_in prf).1 = filter_deps deps (@erase_global X_type X decls normalization_in prf). Proof. @@ -3605,10 +3641,11 @@ Proof. reflexivity. Qed. -Lemma erase_global_deps_eval {efl : EWcbvEval.WcbvFlags} (X_type : abstract_env_impl) (X:X_type.π1) ext wfX {normalization_in} decls nin prf Γ t wt v wv : +Lemma erase_global_deps_eval {efl : EWcbvEval.WcbvFlags} (X_type : abstract_env_impl) (X X':X_type.π1) ext wfX wfX' {normalization_in} {normalization_in'} decls nin prf Γ t wt v wv : let Xext := abstract_make_wf_env_ext X ext wfX in + let Xext' := abstract_make_wf_env_ext X' ext wfX' in let t' := erase X_type Xext (normalization_in := normalization_in) Γ t wt in - let v' := erase X_type Xext (normalization_in := normalization_in) Γ v wv in + let v' := erase X_type Xext' (normalization_in := normalization_in') Γ v wv in let eg := erase_global_deps X_type (term_global_deps t') X decls nin prf in let eg' := erase_global_deps X_type (term_global_deps v') X decls nin prf in eg.1 ⊢ t' ⇓ v' -> EGlobalEnv.extends eg'.1 eg.1. @@ -3635,61 +3672,16 @@ Lemma lookup_erase_global (cf := config.extraction_checker_flags) {X_type X} {de KernameSet.Subset deps deps' -> EExtends.global_subset (erase_global_deps X_type deps X decls normalization_in prf).1 (erase_global_deps X_type deps' X decls normalization_in prf).1. Proof. - revert deps deps' X normalization_in prf. - induction decls. cbn => //. - intros ? ? ? ? ? sub. - epose proof (abstract_env_exists X) as [[Σ wf]]. - destruct a as [kn' []]; cbn; - ( set (decl := E.ConstantDecl _) || - set (decl := E.InductiveDecl _)); hidebody decl; - set (eg := erase_global_deps _ _ _ _ _ _); hidebody eg; - set (eg' := erase_global_deps _ _ _ _ _ _); hidebody eg'; - try (set (eg'' := erase_global_deps _ _ _ _ _ _); hidebody eg''); - try (set (eg''' := erase_global_deps _ _ _ _ _ _); hidebody eg'''). - { destruct (KernameSet.mem) eqn:knm => /=. - + eapply KernameSet.mem_spec, sub, KernameSet.mem_spec in knm. rewrite knm. - apply EExtends.global_subset_cons. eapply IHdecls; eauto. - intros x hin. eapply KernameSet.union_spec in hin. - eapply KernameSet.union_spec. destruct hin. left. now eapply sub. - right => //. - + destruct (KernameSet.mem kn' deps') eqn:eq'. - eapply EExtends.global_subset_cons_right; eauto. - eapply erase_global_deps_wf_glob. - unfold decl. unfold hidebody. - constructor. eapply erase_global_deps_wf_glob. - eapply erase_global_cst_decl_wf_glob. - eapply erase_global_deps_fresh => //. - pose proof (abstract_env_wf _ wf) as [wfΣ]. - depelim wfΣ. rewrite (prf _ wf) in o0. clear - o0. depelim o0. now destruct o. - unfold eg', eg'', hidebody. - erewrite erase_global_deps_irr. - eapply IHdecls. - intros x hin. - eapply KernameSet.union_spec. left. now eapply sub. - eapply IHdecls => //. } - { destruct (KernameSet.mem) eqn:knm => /=. - + eapply KernameSet.mem_spec, sub, KernameSet.mem_spec in knm. rewrite knm. - apply EExtends.global_subset_cons. eapply IHdecls => //. - + destruct (KernameSet.mem kn' deps') eqn:eq'. - eapply EExtends.global_subset_cons_right. eapply erase_global_deps_wf_glob. - constructor. eapply erase_global_deps_wf_glob. - pose proof (abstract_env_wf _ wf) as [wfΣ]. - pose proof (prf _ wf) as prf'. - eapply (erase_global_ind_decl_wf_glob (kn:=kn')). - intros. - unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wf H) as [? [? ?]]. - { now eexists. } - destruct Σ, Σ0. cbn in *. rewrite prf' in wfΣ. - depelim wfΣ. cbn in *. rewrite <- H1, H0, <- H2. - depelim o0. now destruct o1. - eapply erase_global_deps_fresh => //. - pose proof (abstract_env_wf _ wf) as [wfΣ]. - pose proof (prf _ wf) as prf'. - destruct Σ. cbn in *. rewrite prf' in wfΣ. - clear -wfΣ. destruct wfΣ. cbn in *. depelim o0. now destruct o1. - unfold eg'', hidebody. erewrite erase_global_deps_irr. - eapply IHdecls. intros x hin. now eapply sub. - eapply IHdecls => //. } + intros sub. + rewrite !erase_global_deps_erase_global. + assert (wfer : wf_glob (efl := all_env_flags) (@erase_global X_type X decls normalization_in prf)). + { eapply erase_global_wf_glob. } + rewrite !(filter_deps_filter (efl := all_env_flags)) //. + eapply extends_filter_impl => //. 2:exact wfer. + unfold flip. + move=> x /KernameSet.mem_spec hin. + apply/KernameSet.mem_spec. + eapply global_deps_subset; tea. Qed. Lemma expanded_weakening_global X_type X deps deps' decls normalization_in prf Γ t : @@ -4358,6 +4350,114 @@ Proof using Type. rewrite (abstract_env_irr _ H H0); eauto. Qed. +Lemma cored_extends (Σ Σ' : global_env_ext) Γ x y : + wf Σ' -> + extends Σ Σ' -> + cored Σ Γ x y -> cored Σ' Γ x y. +Proof. + intros wf ext. + induction 1. + - constructor. + eapply PCUICWeakeningEnvConv.weakening_env_red1. apply wf. exact ext. + exact X. + - econstructor 2; tea. + eapply PCUICWeakeningEnvConv.weakening_env_red1. apply wf. + exact ext. + exact X. +Qed. + +Lemma normalization_in_extends (Σ Σ' : global_env) ext : + wf Σ -> wf Σ' -> + extends Σ' Σ -> + NormalizationIn (Σ, ext) -> NormalizationIn (Σ', ext). +Proof. + intros wfΣ wfΣ' extH nin. + red in nin. intros Γ t wt. + assert (welltyped (Σ, ext) Γ t). + { destruct wt. exists A. eapply (env_prop_typing weakening_env) in X; tea; cbn. } + specialize (nin _ _ H). clear H. + induction nin. + constructor. intros y r. + eapply H0. eapply cored_extends in r; tea. + eapply cored_welltyped; tea. +Qed. + +Lemma iter_abstract_pop_decls_correct {cf} {abstract_env_impl abstract_env_ext_impl : Type} + {abstract_env_struct0 : abstract_env_struct abstract_env_impl + abstract_env_ext_impl} : + abstract_env_prop abstract_env_impl abstract_env_ext_impl -> + forall (X : abstract_env_impl) (decls : list (kername × global_decl)) n, + (forall Σ : global_env, + Σ ∼ X -> exists decls', #|decls'| = n /\ declarations Σ = decls' ++ decls) -> + let X' := iter abstract_pop_decls n X in + forall Σ Σ' : global_env, + Σ ∼ X -> + Σ' ∼ X' -> + declarations Σ' = decls /\ universes Σ = universes Σ' /\ retroknowledge Σ = retroknowledge Σ'. +Proof. + intros ap X decls. + induction n in X, decls |- *. cbn. + - intros hdecl Σ Σ' HX HX'. + specialize (hdecl _ HX) as [decls' [hlen hde]]. + destruct decls' => //. cbn in hde. + now rewrite -hde (abstract_env_irr _ HX HX'). + - intros hdecl X' Σ Σ' HX HX'. + specialize (hdecl _ HX) as [decls' [hlen hde]]. + destruct decls' => //; cbn in hde. + specialize (IHn (abstract_pop_decls X) decls). + pose proof (abstract_pop_decls_correct X (decls' ++ decls)). + forward H. + { intros ? h. rewrite -(abstract_env_irr _ HX h). now exists p. } + forward IHn. + { intros ? H0. now specialize (H _ _ HX H0). } + pose proof (abstract_env_exists (abstract_pop_decls X)) as [[Σpop hpop]]. + specialize (IHn _ _ hpop HX'). + destruct IHn. split => //. + pose proof (abstract_pop_decls_correct X (decls' ++ decls)). + forward H2. + { intros ? hx. pose proof (abstract_env_irr _ hx HX). subst Σ0. now exists p. } + specialize (H2 _ _ HX hpop). destruct H2 as [? []]. + destruct H1. + split; congruence. +Qed. + +Lemma skipn_app_prefix {A} n (l l' : list A) : skipn (#|l| + n) (l ++ l') = skipn n l'. +Proof. + induction l; cbn; auto. +Qed. + +Lemma erase_global_deps_fast_erase_global_deps {deps} {X_type X} {decls normalization_in hprefix hprefix'} : + exists normalization_in', + erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) hprefix = + erase_global_deps X_type deps X decls normalization_in' hprefix'. +Proof using Type. + unshelve eexists. + intros. + pose proof (abstract_env_exists X) as [[Σ' H']]. + pose proof (abstract_env_wf _ H') as []. + specialize (hprefix _ H') as [[Σ'' ?]]. + pose proof (abstract_env_exists (iter abstract_pop_decls (S n) X)) as [[? ?]]. + pose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ a H1). noconf H2. + epose proof (iter_abstract_pop_decls_correct _ X (skipn (S n) (Σ'' ++ decls)) (S n)). + forward H2. { intros ? h. rewrite (abstract_env_irr _ h H'). exists (firstn (S n) (Σ'' ++ decls)). + rewrite firstn_skipn. split => //. rewrite firstn_length_le //. rewrite app_length. lia. } + specialize (H2 _ _ H' a) as [? []]. + eapply normalization_in_extends. exact X1. exact X0. + { eapply extends_decls_extends, strictly_extends_decls_extends_decls. + red. rewrite H2 H3 H4. split => //. exists (firstn (S n) (Σ'' ++ decls)). + now rewrite firstn_skipn. } + specialize (normalization_in n H kn cb). + assert (wf_ext (Σ', cst_universes cb)). + { split => //. cbn. rewrite H3. apply X0. } + assert (forall Σ : global_env, Σ ∼ X -> ∥ wf_ext (Σ, cst_universes cb) ∥). + { intros. pose proof (abstract_env_irr _ H' H5). now subst Σ'. } + unshelve eapply normalization_in; tea. + set (foo := abstract_make_wf_env_ext X (cst_universes cb) H5). + pose proof (abstract_env_ext_exists foo) as [[Σext hext]]. + epose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ H' hext). now subst Σext. + erewrite erase_global_deps_fast_spec; tea. reflexivity. +Qed. + Definition erase_global_fast X_type deps X decls {normalization_in} (prf:forall Σ : global_env, abstract_env_rel X Σ -> declarations Σ = decls) := erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) (fun _ H => sq ([] ; prf _ H)). From 0934fc0d92d6e04933fec71f41934ebd92e4ca27 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 31 Aug 2023 18:07:10 +0200 Subject: [PATCH 45/61] Modify Transform class to allow showing preservation of fo values compositionally --- common/theories/Transform.v | 31 +-- erasure-plugin/_CoqProject.in | 1 + erasure-plugin/theories/ETransform.v | 20 +- erasure-plugin/theories/Erasure.v | 27 ++- erasure-plugin/theories/ErasureCorrectness.v | 187 +++++++++++++++---- erasure/theories/EProgram.v | 38 ++-- erasure/theories/ErasureFunction.v | 81 ++++++++ template-coq/theories/TemplateProgram.v | 2 +- template-pcuic/theories/PCUICTransform.v | 4 +- 9 files changed, 299 insertions(+), 92 deletions(-) diff --git a/common/theories/Transform.v b/common/theories/Transform.v index bbeb25d22..d1678fb27 100644 --- a/common/theories/Transform.v +++ b/common/theories/Transform.v @@ -19,14 +19,17 @@ Extract Constant time => "(fun c f x -> let s = Caml_bytestring.caml_string_of_bytestring c in Tm_util.time (Pp.str s) f x)". Module Transform. + Definition program env term := env * term. Section Opt. - Context {program program' : Type}. - Context {value value' : Type}. - Context {eval : program -> value -> Prop}. - Context {eval' : program' -> value' -> Prop}. + Context {env env' : Type}. + Context {term term' : Type}. + Notation program' := (program env' term'). + Notation program := (program env term). + Context {eval : program -> term -> Prop}. + Context {eval' : program' -> term' -> Prop}. Definition preserves_eval pre (transform : forall p : program, pre p -> program') - (obseq : forall p : program, pre p -> program' -> value -> value' -> Prop) := + (obseq : forall p : program, pre p -> program' -> term -> term' -> Prop) := forall p v (pr : pre p), eval p v -> let p' := transform p pr in @@ -38,7 +41,7 @@ Module Transform. transform : forall p : program, pre p -> program'; post : program' -> Prop; correctness : forall input (p : pre input), post (transform input p); - obseq : forall p : program, pre p -> program' -> value -> value' -> Prop; + obseq : forall p : program, pre p -> program' -> term -> term' -> Prop; preservation : preserves_eval pre transform obseq }. Definition run (x : t) (p : program) (pr : pre x p) : program' := @@ -47,19 +50,19 @@ Module Transform. End Opt. Arguments t : clear implicits. - Definition self_transform program value eval eval' := t program program value value eval eval'. + Definition self_transform env term eval eval' := t env env term term eval eval'. Section Comp. - Context {program program' program'' : Type}. - Context {value value' value'' : Type}. - Context {eval : program -> value -> Prop}. - Context {eval' : program' -> value' -> Prop}. - Context {eval'' : program'' -> value'' -> Prop}. + Context {env env' env'' : Type}. + Context {term term' term'' : Type}. + Context {eval : program env term -> term -> Prop}. + Context {eval' : program env' term' -> term' -> Prop}. + Context {eval'' : program env'' term'' -> term'' -> Prop}. Local Obligation Tactic := idtac. - Program Definition compose (o : t program program' value value' eval eval') (o' : t program' program'' value' value'' eval' eval'') + Program Definition compose (o : t env env' term term' eval eval') (o' : t env' env'' term' term'' eval' eval'') (hpp : (forall p, o.(post) p -> o'.(pre) p)) - : t program program'' value value'' eval eval'' := + : t env env'' term term'' eval eval'' := {| name := (o.(name) ^ " -> " ^ o'.(name))%bs; transform p hp := run o' (run o p hp) (hpp _ (o.(correctness) _ hp)); diff --git a/erasure-plugin/_CoqProject.in b/erasure-plugin/_CoqProject.in index 49d459d43..4978832ec 100644 --- a/erasure-plugin/_CoqProject.in +++ b/erasure-plugin/_CoqProject.in @@ -2,4 +2,5 @@ theories/ETransform.v theories/Erasure.v +theories/ErasureCorrectness.v theories/Extraction.v diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index c8c4584c1..bc6694a18 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -137,6 +137,8 @@ Proof. eauto using firstn_skipn. } Qed. +Local Obligation Tactic := program_simpl. + Program Definition erase_pcuic_program {guard : abstract_guard_impl} (p : pcuic_program) {normalization_in : NormalizationIn_erase_pcuic_program_1 p.1} {normalization_in_adjust_universes : NormalizationIn_erase_pcuic_program_2 p.1} @@ -193,7 +195,7 @@ Qed. Local Obligation Tactic := try solve [ eauto ]. -Program Definition erase_transform {guard : abstract_guard_impl} : Transform.t pcuic_program eprogram_env PCUICAst.term EAst.term +Program Definition erase_transform {guard : abstract_guard_impl} : Transform.t _ _ PCUICAst.term EAst.term eval_pcuic_program (eval_eprogram_env EWcbvEval.default_wcbv_flags) := {| name := "erasure"; pre p := @@ -325,7 +327,7 @@ End PCUICEnv. Import EWcbvEval (WcbvFlags, with_prop_case, with_guarded_fix). Program Definition guarded_to_unguarded_fix {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} {efl : EEnvFlags} (wguard : with_guarded_fix) : - Transform.t eprogram_env eprogram_env EAst.term EAst.term + Transform.t _ _ EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram_env (EWcbvEval.switch_unguarded_fix fl)) := {| name := "switching to unguarded fixpoints"; transform p pre := p; @@ -354,7 +356,7 @@ Definition rebuild_wf_env {efl} (p : eprogram) (hwf : wf_eprogram efl p): eprogr (GlobalContextMap.make p.1 (wf_glob_fresh p.1 (proj1 hwf)), p.2). Program Definition rebuild_wf_env_transform {fl : EWcbvEval.WcbvFlags} {efl} (with_exp : bool) : - Transform.t eprogram eprogram_env EAst.term EAst.term (eval_eprogram fl) (eval_eprogram_env fl) := + Transform.t _ _ EAst.term EAst.term (eval_eprogram fl) (eval_eprogram_env fl) := {| name := "rebuilding environment lookup table"; pre p := wf_eprogram efl p /\ (with_exp ==> EEtaExpanded.expanded_eprogram_cstrs p); transform p pre := rebuild_wf_env p (proj1 pre); @@ -376,7 +378,7 @@ Qed. Program Definition remove_params_optimization {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := all_env_flags): - Transform.t eprogram_env eprogram EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram fl) := + Transform.t _ _ EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram fl) := {| name := "stripping constructor parameters"; transform p pre := ERemoveParams.strip_program p; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; @@ -412,7 +414,7 @@ Qed. Program Definition remove_params_fast_optimization {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := all_env_flags) : - Transform.t eprogram_env eprogram EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram fl) := + Transform.t _ _ EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram fl) := {| name := "stripping constructor parameters (faster?)"; transform p _ := (ERemoveParams.Fast.strip_env p.1, ERemoveParams.Fast.strip p.1 [] p.2); pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; @@ -454,7 +456,7 @@ Qed. Import EOptimizePropDiscr EWcbvEval. Program Definition remove_match_on_box_trans {fl : WcbvFlags} {wcon : with_constructor_as_block = false} {efl : EEnvFlags} {hastrel : has_tRel} {hastbox : has_tBox} : - Transform.t eprogram_env eprogram EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram (disable_prop_cases fl)) := + Transform.t _ _ EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram (disable_prop_cases fl)) := {| name := "optimize_prop_discr"; transform p _ := remove_match_on_box_program p ; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; @@ -490,7 +492,7 @@ From MetaCoq.Erasure Require Import EInlineProjections. Program Definition inline_projections_optimization {fl : WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := switch_no_params all_env_flags) {hastrel : has_tRel} {hastbox : has_tBox} : - Transform.t eprogram_env eprogram EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram fl) := + Transform.t _ _ EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram fl) := {| name := "primitive projection inlining"; transform p _ := EInlineProjections.optimize_program p ; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; @@ -526,7 +528,7 @@ From MetaCoq.Erasure Require Import EConstructorsAsBlocks. Program Definition constructors_as_blocks_transformation {efl : EEnvFlags} {has_app : has_tApp} {has_rel : has_tRel} {hasbox : has_tBox} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = false} : - Transform.t eprogram_env eprogram EAst.term EAst.term (eval_eprogram_env target_wcbv_flags) (eval_eprogram block_wcbv_flags) := + Transform.t _ _ EAst.term EAst.term (eval_eprogram_env target_wcbv_flags) (eval_eprogram block_wcbv_flags) := {| name := "transforming to constuctors as blocks"; transform p _ := EConstructorsAsBlocks.transform_blocks_program p ; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; @@ -566,7 +568,7 @@ From MetaCoq.Erasure Require Import EImplementBox. Program Definition implement_box_transformation {efl : EEnvFlags} {has_app : has_tApp} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = true} : - Transform.t eprogram eprogram EAst.term EAst.term (eval_eprogram block_wcbv_flags) (eval_eprogram block_wcbv_flags) := + Transform.t _ _ EAst.term EAst.term (eval_eprogram block_wcbv_flags) (eval_eprogram block_wcbv_flags) := {| name := "transforming to constuctors as blocks"; transform p _ := EImplementBox.implement_box_program p ; pre p := wf_eprogram efl p ; diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 4467c1dae..175f9210f 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -23,13 +23,13 @@ Local Open Scope string_scope2. Import Transform. -Local Obligation Tactic := program_simpl. - #[local] Existing Instance extraction_checker_flags. #[local] Existing Instance PCUICSN.extraction_normalizing. Import EWcbvEval. +Local Obligation Tactic := program_simpl. + Axiom assume_welltyped_template_program_expansion : forall p (wtp : ∥ wt_template_program_env p ∥), let p' := EtaExpand.eta_expand_program p in @@ -42,7 +42,7 @@ Axiom assume_preservation_template_program_env_expansion : (** We kludge the normalization assumptions by parameterizing over a continuation of "what will be done to the program later" as well as what properties we'll need of it *) -Program Definition eta_expand K : Transform.t template_program_env template_program Ast.term Ast.term +Program Definition eta_expand K : Transform.t _ _ Ast.term Ast.term eval_template_program_env eval_template_program := {| name := "eta expand cstrs and fixpoints"; pre := fun p => ∥ wt_template_program_env p ∥ /\ K (eta_expand_global_env p.1) ; @@ -50,7 +50,7 @@ Program Definition eta_expand K : Transform.t template_program_env template_prog post := fun p => ∥ wt_template_program p ∥ /\ EtaExpand.expanded_program p /\ K p.1; obseq p hp p' v v' := v' = EtaExpand.eta_expand p.1 [] v |}. Next Obligation. - let p := match goal with H : template_program_env |- _ => H end in + let p := match goal with H : program _ _ |- _ => H end in destruct p. split; [|split]; auto; now apply assume_welltyped_template_program_expansion. Qed. Next Obligation. @@ -59,7 +59,7 @@ Next Obligation. Qed. Program Definition verified_lambdabox_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - Transform.t EProgram.eprogram_env EProgram.eprogram EAst.term EAst.term + Transform.t _ _ EAst.term EAst.term (EProgram.eval_eprogram_env {| with_prop_case := true; with_guarded_fix := true; with_constructor_as_block := false |}) (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := (* Simulation of the guarded fixpoint rules with a single unguarded one: @@ -98,7 +98,7 @@ Qed. Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - Transform.t pcuic_program EProgram.eprogram + Transform.t _ _ PCUICAst.term EAst.term PCUICTransform.eval_pcuic_program (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := @@ -117,12 +117,11 @@ Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl Import EGlobalEnv EWellformed. Definition transform_compose - {program program' program'' value value' value'' : Type} - {eval : program -> value -> Prop} {eval' : program' -> value' -> Prop} - {eval'' : program'' -> value'' -> Prop} - (o : t program program' value value' eval eval') - (o' : t program' program'' value' value'' eval' eval'') - (pre : forall p : program', post o p -> pre o' p) : + {env env' env'' term term' term'' : Type} + {eval eval' eval''} + (o : t env env' term term' eval eval') + (o' : t env' env'' term' term'' eval' eval'') + (pre : forall p, post o p -> pre o' p) : forall x p1, exists p3, transform (compose o o' pre) x p1 = transform o' (transform o x p1) p3. Proof. @@ -138,7 +137,7 @@ Proof. Qed. Program Definition pre_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - Transform.t TemplateProgram.template_program pcuic_program + Transform.t _ _ Ast.term PCUICAst.term TemplateProgram.eval_template_program PCUICTransform.eval_pcuic_program := @@ -159,7 +158,7 @@ Program Definition pre_erasure_pipeline {guard : abstract_guard_impl} (efl := EW template_to_pcuic_transform (K _ T2). Program Definition erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - Transform.t TemplateProgram.template_program EProgram.eprogram + Transform.t _ _ Ast.term EAst.term TemplateProgram.eval_template_program (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 4558691a0..158bd529c 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -20,22 +20,20 @@ Import PCUICTransform (template_to_pcuic_transform, pcuic_expand_lets_transform) Import Transform. -Obligation Tactic := program_simpl. +#[local] Obligation Tactic := program_simpl. #[local] Existing Instance extraction_checker_flags. Import EWcbvEval. Lemma transform_compose_assoc - {program program' program'' program''' value value' value'' value''' : Type} - {eval : program -> value -> Prop} {eval' : program' -> value' -> Prop} - {eval'' : program'' -> value'' -> Prop} - {eval''' : program''' -> value''' -> Prop} - (o : t program program' value value' eval eval') - (o' : t program' program'' value' value'' eval' eval'') - (o'' : t program'' program''' value'' value''' eval'' eval''') - (prec : forall p : program', post o p -> pre o' p) - (prec' : forall p : program'', post o' p -> pre o'' p) : + {env env' env'' env''' term term' term'' term''' : Type} + {eval eval' eval'' eval'''} + (o : t env env' term term' eval eval') + (o' : t env' env'' term' term'' eval' eval'') + (o'' : t env'' env''' term'' term''' eval'' eval''') + (prec : forall p, post o p -> pre o' p) + (prec' : forall p, post o' p -> pre o'' p) : forall x p1, transform (compose o (compose o' o'' prec') prec) x p1 = transform (compose (compose o o' prec) o'' prec') x p1. @@ -47,15 +45,13 @@ Proof. Qed. Lemma obseq_compose_assoc - {program program' program'' program''' value value' value'' value''' : Type} - {eval : program -> value -> Prop} {eval' : program' -> value' -> Prop} - {eval'' : program'' -> value'' -> Prop} - {eval''' : program''' -> value''' -> Prop} - (o : t program program' value value' eval eval') - (o' : t program' program'' value' value'' eval' eval'') - (o'' : t program'' program''' value'' value''' eval'' eval''') - (prec : forall p : program', post o p -> pre o' p) - (prec' : forall p : program'', post o' p -> pre o'' p) : + {env env' env'' env''' term term' term'' term''' : Type} + {eval eval' eval'' eval'''} + (o : t env env' term term' eval eval') + (o' : t env' env'' term' term'' eval' eval'') + (o'' : t env'' env''' term'' term''' eval'' eval''') + (prec : forall p, post o p -> pre o' p) + (prec' : forall p, post o' p -> pre o'' p) : forall x p1 p2 v1 v2, obseq (compose o (compose o' o'' prec') prec) x p1 p2 v1 v2 <-> obseq (compose (compose o o' prec) o'' prec') x p1 p2 v1 v2. Proof. @@ -326,14 +322,126 @@ Proof. red. intros. split; reflexivity. Qed. -Section pipeline_theorem. +Lemma expand_lets_fo (Σ : global_env_ext_map) t : + PCUICFirstorder.firstorder_value Σ [] t -> + let p := (Σ, t) in + PCUICExpandLets.expand_lets_program p = + (build_global_env_map (PCUICAst.PCUICEnvironment.fst_ctx (PCUICExpandLets.trans_global p.1)), p.1.2, t). +Proof. + intros p. + cbn. unfold PCUICExpandLets.expand_lets_program. f_equal. cbn. + move: p. apply: (PCUICFirstorder.firstorder_value_inds _ _ (fun t => PCUICExpandLets.trans t = t)). + intros i n ui u args pandi ht hf ih isp. + rewrite PCUICExpandLetsCorrectness.trans_mkApps /=. f_equal. + now eapply forall_map_id_spec. +Qed. - Fixpoint compile_value_box (t : PCUICAst.term) (acc : list EAst.term) : EAst.term := - match t with - | PCUICAst.tApp f a => compile_value_box f (acc ++ [compile_value_box a []]) - | PCUICAst.tConstruct i n _ => EAst.tConstruct i n acc - | _ => EAst.tVar "error" - end. +Fixpoint compile_value_box (t : PCUICAst.term) (acc : list EAst.term) : EAst.term := + match t with + | PCUICAst.tApp f a => compile_value_box f (acc ++ [compile_value_box a []]) + | PCUICAst.tConstruct i n _ => EAst.tConstruct i n acc + | _ => EAst.tVar "error" + end. + +From Equations Require Import Equations. + +Lemma erase_transform_fo_gen (p : pcuic_program) pr : + PCUICFirstorder.firstorder_value p.1 [] p.2 -> + forall ep, transform erase_transform p pr = ep -> ep.2 = compile_value_erase p.2 []. +Proof. + destruct p as [Σ t]. cbn. + intros hev ep <-. move: hev pr. + unfold erase_program, erase_pcuic_program; cbn -[erase PCUICWfEnv.abstract_make_wf_env_ext]. + intros fo pr. + set (prf0 := fun (Σ0 : PCUICAst.PCUICEnvironment.global_env_ext) => _). + set (prf1 := fun (Σ0 : PCUICAst.PCUICEnvironment.global_env_ext) => _). + clearbody prf0 prf1. + destruct pr as []. + destruct s as [[]]. + epose proof (erases_erase (X_type := optimized_abstract_env_impl) prf1 _ eq_refl). + eapply erases_firstorder' in H; eauto. +Qed. + +Lemma erase_transform_fo (p : pcuic_program) pr : + PCUICFirstorder.firstorder_value p.1 [] p.2 -> + transform erase_transform p pr = ((transform erase_transform p pr).1, compile_value_erase p.2 []). +Proof. + intros fo. + set (tr := transform _ _ _). + change tr with (tr.1, tr.2). f_equal. + eapply erase_transform_fo_gen; tea. reflexivity. +Qed. + +Inductive firstorder_evalue : EAst.term -> Prop := + | is_fo i n args : Forall firstorder_evalue args -> firstorder_evalue (EAst.tConstruct i n args). + +Lemma compile_fo_value Σ t : + PCUICFirstorder.firstorder_value Σ [] t -> + firstorder_evalue (compile_value_erase t []). +Proof. Admitted. + +Fixpoint compile_evalue_box (t : EAst.term) (acc : list EAst.term) := + match t with + | EAst.tApp f a => compile_evalue_box f (compile_evalue_box a [] :: acc) + | EAst.tConstruct i n _ => EAst.tConstruct i n acc + | _ => EAst.tVar "error" + end. + +Import MetaCoq.Common.Transform. + +Module ETransformPresFO. + Section Opt. + Context {env env' : Type}. + Context {eval : program env EAst.term -> EAst.term -> Prop}. + Context {eval' : program env' EAst.term -> EAst.term -> Prop}. + Context (o : Transform.t _ _ _ _ eval eval'). + Context (compile_fo_value : EAst.term -> EAst.term). + + Class t := + { preserves_fo : forall v, firstorder_evalue v -> firstorder_evalue (compile_fo_value v); + transform_fo : forall v (pr : o.(pre) v), + firstorder_evalue v.2 -> (o.(transform) v pr).2 = compile_fo_value v.2 }. + End Opt. + + Section Comp. + Context {env env' env'' : Type}. + Context {eval : program env EAst.term -> EAst.term -> Prop}. + Context {eval' : program env' EAst.term -> EAst.term -> Prop}. + Context {eval'' : program env'' EAst.term -> EAst.term -> Prop}. + Context (compile_fo_value compile_fo_value' : EAst.term -> EAst.term). + + Local Obligation Tactic := idtac. + #[global] + Instance compose (o : Transform.t _ _ _ _ eval eval') + (o' : Transform.t _ _ _ _ eval' eval'') + (oext : t o compile_fo_value) + (o'ext : t o' compile_fo_value') + (hpp : (forall p, o.(post) p -> o'.(pre) p)) + : t (Transform.compose o o' hpp) (compile_fo_value' ∘ compile_fo_value). + Proof. + split. + - intros. eapply preserves_fo; tea. eapply preserves_fo; tea. + - intros. cbn. unfold run, time. + rewrite o'ext.(transform_fo _ _). + rewrite oext.(transform_fo _ _) //. + eapply preserves_fo; tea. + rewrite oext.(transform_fo _ _) //. + Qed. + End Comp. + +End ETransformPresFO. + +Axiom lambdabox_pres_fo : ETransformPresFO.t verified_lambdabox_pipeline (flip compile_evalue_box []). + +Lemma transform_lambda_box_firstorder (Σer : EEnvMap.GlobalContextMap.t) v pre : + firstorder_evalue v -> + (transform verified_lambdabox_pipeline (Σer, v) pre).2 = (compile_evalue_box v []). +Proof. + intros fo. + now rewrite (ETransformPresFO.transform_fo _ _ (t:=lambdabox_pres_fo)). +Qed. + +Section pipeline_theorem. Instance cf : checker_flags := extraction_checker_flags. Instance nf : PCUICSN.normalizing_flags := PCUICSN.extraction_normalizing. @@ -369,12 +477,12 @@ Section pipeline_theorem. Lemma precond2 : pre verified_erasure_pipeline (Σ, v). Proof. - cbn. repeat eapply conj; sq; cbn; eauto. + cbn. destruct Heval. repeat eapply conj; sq; cbn; eauto. - red. cbn. split; eauto. eexists. eapply PCUICClassification.subject_reduction_eval; eauto. - todo "preservation of eta expandedness". - - todo "normalization". + - cbn. todo "normalization". - todo "normalization". Qed. @@ -384,7 +492,7 @@ Section pipeline_theorem. Lemma fo_v : PCUICFirstorder.firstorder_value Σ [] v. Proof. - sq. + destruct Heval. sq. eapply PCUICFirstorder.firstorder_value_spec; eauto. - eapply PCUICClassification.subject_reduction_eval; eauto. - eapply PCUICWcbvEval.eval_to_value; eauto. @@ -394,11 +502,24 @@ Section pipeline_theorem. Lemma v_t_spec : v_t = (transform verified_erasure_pipeline (Σ, v) precond2).2. Proof. unfold v_t. generalize fo_v, precond2. clear. - induction 1. - intros. unfold verified_erasure_pipeline. - repeat destruct_compose. intros. - cbn [transform erase_transform]. + intros hv pre. + unfold verified_erasure_pipeline. + rewrite -transform_compose_assoc. + destruct_compose. cbn [transform pcuic_expand_lets_transform]. + rewrite (expand_lets_fo _ _ hv). + cbn [fst snd]. + intros h. + destruct_compose. + rewrite erase_transform_fo. cbn. admit. + set (Σer := (transform erase_transform _ _).1). + cbn [fst snd]. intros pre'. + symmetry. + erewrite transform_lambda_box_firstorder; tea. + eapply obseq_lambdabox. intros kn decl H; exact H. + Unshelve. 2:exact pre'. 2:exact (Σer, compile_value_erase v []). + cbn [snd]. + red. cbn. Admitted. Import PCUICWfEnv. diff --git a/erasure/theories/EProgram.v b/erasure/theories/EProgram.v index cd2d7cc49..15cfe1e16 100644 --- a/erasure/theories/EProgram.v +++ b/erasure/theories/EProgram.v @@ -13,7 +13,7 @@ Local Open Scope string_scope2. Import Transform. -Obligation Tactic := program_simpl. +Local Obligation Tactic := program_simpl. Import EGlobalEnv EWellformed. @@ -53,13 +53,14 @@ Qed. Module TransformExt. Section Opt. - Context {program program' : Type}. - Context {value value' : Type}. - Context {eval : program -> value -> Prop}. - Context {eval' : program' -> value' -> Prop}. - Context (o : Transform.t program program' value value' eval eval'). - Context (extends : program -> program -> Prop). - Context (extends' : program' -> program' -> Prop). + Context {env env' env'' : Type}. + Context {term term' term'' : Type}. + Context {eval : program env term -> term -> Prop}. + Context {eval' : program env' term' -> term' -> Prop}. + Context {eval'' : program env'' term'' -> term'' -> Prop}. + Context (o : Transform.t env env' term term' eval eval'). + Context (extends : program env term -> program env term -> Prop). + Context (extends' : program env' term' -> program env' term' -> Prop). Class t := preserves_obs : forall p p' (pr : o.(pre) p) (pr' : o.(pre) p'), extends p p' -> extends' (o.(transform) p pr) (o.(transform) p' pr'). @@ -67,20 +68,19 @@ Module TransformExt. End Opt. Section Comp. - Context {program program' program'' : Type}. - Context {value value' value'' : Type}. - Context {eval : program -> value -> Prop}. - Context {eval' : program' -> value' -> Prop}. - Context {eval'' : program'' -> value'' -> Prop}. - - Context {extends : program -> program -> Prop}. - Context {extends' : program' -> program' -> Prop}. - Context {extends'' : program'' -> program'' -> Prop}. + Context {env env' env'' : Type}. + Context {term term' term'' : Type}. + Context {eval : program env term -> term -> Prop}. + Context {eval' : program env' term' -> term' -> Prop}. + Context {eval'' : program env'' term'' -> term'' -> Prop}. + Context {extends : program env term -> program env term -> Prop}. + Context {extends' : program env' term' -> program env' term' -> Prop}. + Context {extends'' : program env'' term'' -> program env'' term'' -> Prop}. Local Obligation Tactic := idtac. #[global] - Instance compose (o : Transform.t program program' value value' eval eval') - (o' : Transform.t program' program'' value' value'' eval' eval'') + Instance compose (o : Transform.t env env' term term' eval eval') + (o' : Transform.t env' env'' term' term'' eval' eval'') (oext : t o extends extends') (o'ext : t o' extends' extends'') (hpp : (forall p, o.(post) p -> o'.(pre) p)) diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 8313423d9..528be6234 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -4051,6 +4051,8 @@ Qed. (* This version of global environment erasure keeps the same global environment throughout the whole erasure, while folding over the list of declarations. *) +Local Obligation Tactic := program_simpl. + Program Fixpoint erase_global_deps_fast (deps : KernameSet.t) X_type (X:X_type.π1) (decls : global_declarations) {normalization_in : NormalizationIn_erase_global_deps_fast X decls} @@ -4515,3 +4517,82 @@ Proof using Type. Qed. End EraseGlobalFast. + +Fixpoint compile_value_erase (t : PCUICAst.term) (acc : list EAst.term) : EAst.term := + match t with + | PCUICAst.tApp f a => compile_value_erase f (compile_value_erase a [] :: acc) + | PCUICAst.tConstruct i n _ => EAst.mkApps (EAst.tConstruct i n []) acc + | _ => EAst.tVar "error" + end. + +Lemma compile_value_erase_mkApps i n ui args acc : + compile_value_erase (mkApps (tConstruct i n ui) args) acc = + EAst.mkApps (EAst.tConstruct i n []) (map (flip compile_value_erase []) args ++ acc). +Proof. + revert acc; induction args using rev_ind. + - cbn. auto. + - intros acc. rewrite PCUICAstUtils.mkApps_app /=. cbn. + now rewrite IHargs map_app /= -app_assoc /=. +Qed. + +Lemma erases_firstorder Σ Γ t : + PCUICFirstorder.firstorder_value Σ Γ t -> + erases Σ Γ t (compile_value_erase t []). +Proof. + move: t. + apply: (PCUICFirstorder.firstorder_value_inds Σ Γ). + intros i n ui u args pandi hty hfo ih isp. + assert (Forall2 (erases Σ Γ) args (map (flip compile_value_erase []) args)). + { solve_all. eapply All_All2; tea. + cbn. intros x [fo hx]. exact hx. } + unshelve epose proof (erases_mkApps Σ Γ (tConstruct i n ui) (EAst.tConstruct i n []) args _ _ H). + constructor. + { move: isp. rewrite /isPropositional /Extract.isPropositional /PCUICAst.lookup_inductive /= /lookup_inductive_gen /lookup_minductive_gen. + destruct PCUICEnvironment.lookup_env => //. + destruct g => //. destruct nth_error => //. } + now rewrite compile_value_erase_mkApps app_nil_r. +Qed. + +Lemma erases_firstorder' Σ Γ t : + wf_ext Σ -> + PCUICFirstorder.firstorder_value Σ Γ t -> + forall t', erases Σ Γ t t' -> t' = (compile_value_erase t []). +Proof. + intros wf. move: t. + apply: (PCUICFirstorder.firstorder_value_inds Σ Γ). + intros i n ui u args pandi hty hfo ih isp. + assert (Forall2 (erases Σ Γ) args (map (flip compile_value_erase []) args)). + { solve_all. eapply All_All2; tea. + cbn. intros x [fo hx]. now eapply erases_firstorder. } + unshelve epose proof (erases_mkApps Σ Γ (tConstruct i n ui) (EAst.tConstruct i n []) args _ _ H). + constructor. + { move: isp. rewrite /isPropositional /Extract.isPropositional /PCUICAst.lookup_inductive /= /lookup_inductive_gen /lookup_minductive_gen. + destruct PCUICEnvironment.lookup_env => //. + destruct g => //. destruct nth_error => //. } + rewrite compile_value_erase_mkApps app_nil_r. + intros t' ht'. + eapply erases_mkApps_inv in ht'. + destruct ht'. + { destruct H1 as [? [? [? [? [[] ?]]]]]. + eapply isErasable_Propositional in X => //. + clear -isp X. elimtype False. + move: isp X. + rewrite /isPropositional /Extract.isPropositional /PCUICAst.lookup_inductive /= /lookup_inductive_gen /lookup_minductive_gen. + destruct PCUICEnvironment.lookup_env => //. + destruct g => //. destruct nth_error => //. + rewrite /isPropositionalArity. destruct destArity => //. + destruct p. intros ->. auto. } + destruct H1 as [f' [L' [-> [erc erargs]]]]. + depelim erc. f_equal. + solve_all. clear -erargs. + induction erargs => //. rewrite IHerargs. cbn. f_equal. + destruct r as [[] ?]. eapply e in e0. now unfold flip. + eapply (isErasable_Propositional (args := [])) in X => //. + clear -isp X. elimtype False. + move: isp X. + rewrite /isPropositional /Extract.isPropositional /PCUICAst.lookup_inductive /= /lookup_inductive_gen /lookup_minductive_gen. + destruct PCUICEnvironment.lookup_env => //. + destruct g => //. destruct nth_error => //. + rewrite /isPropositionalArity. destruct destArity => //. + destruct p. intros ->. auto. +Qed. \ No newline at end of file diff --git a/template-coq/theories/TemplateProgram.v b/template-coq/theories/TemplateProgram.v index 2895858f6..6db0926ba 100644 --- a/template-coq/theories/TemplateProgram.v +++ b/template-coq/theories/TemplateProgram.v @@ -47,7 +47,7 @@ Definition make_template_program_env {cf : checker_flags} (p : template_program) (** We kludge the normalization assumptions by parameterizing over a continuation of "what will be done to the program later" as well as what properties we'll need of it *) Program Definition build_template_program_env {cf : checker_flags} K : - Transform.t template_program template_program_env Ast.term Ast.term eval_template_program eval_template_program_env := + Transform.t global_env GlobalEnvMap.t Ast.term Ast.term eval_template_program eval_template_program_env := {| name := "rebuilding environment lookup table"; pre p := ∥ wt_template_program p ∥ /\ forall pf, K (GlobalEnvMap.make p.1 (wt_template_program_fresh p pf)) : Prop ; transform p pre := make_template_program_env p (proj1 pre); diff --git a/template-pcuic/theories/PCUICTransform.v b/template-pcuic/theories/PCUICTransform.v index 359f8998b..ef840ddaf 100644 --- a/template-pcuic/theories/PCUICTransform.v +++ b/template-pcuic/theories/PCUICTransform.v @@ -34,7 +34,7 @@ Local Obligation Tactic := idtac. (** We kludge the normalization assumptions by parameterizing over a continuation of "what will be done to the program later" as well as what properties we'll need of it *) Program Definition template_to_pcuic_transform {cf : checker_flags} K : - Transform.t template_program pcuic_program Ast.term term + Transform.t Ast.Env.global_env global_env_ext_map Ast.term term eval_template_program eval_pcuic_program := {| name := "template to pcuic"; pre p := ∥ wt_template_program p ∥ /\ EtaExpand.expanded_program p /\ K (trans_global (Ast.Env.empty_ext p.1)) ; @@ -66,7 +66,7 @@ From MetaCoq.PCUIC Require Import PCUICExpandLets PCUICExpandLetsCorrectness. *) Program Definition pcuic_expand_lets_transform {cf : checker_flags} K : - self_transform pcuic_program term eval_pcuic_program eval_pcuic_program := + self_transform global_env_ext_map term eval_pcuic_program eval_pcuic_program := {| name := "let expansion in branches/constructors"; pre p := ∥ wt_pcuic_program p ∥ /\ PCUICEtaExpand.expanded_pcuic_program p /\ K (build_global_env_map (trans_global_env p.1.1), p.1.2) ; transform p hp := expand_lets_program p; From 1f1248dd3aefd63a926112949d05685b65de4938 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 1 Sep 2023 18:15:51 +0200 Subject: [PATCH 46/61] Proved most of the fo preservation of the lambdabox pipeline --- erasure-plugin/theories/ErasureCorrectness.v | 376 +++++++++++++++++-- erasure/theories/EInlineProjections.v | 11 + erasure/theories/EOptimizePropDiscr.v | 10 + erasure/theories/ERemoveParams.v | 11 + 4 files changed, 371 insertions(+), 37 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 158bd529c..5c0facb2e 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -336,10 +336,15 @@ Proof. now eapply forall_map_id_spec. Qed. -Fixpoint compile_value_box (t : PCUICAst.term) (acc : list EAst.term) : EAst.term := +Fixpoint compile_value_box Σ (t : PCUICAst.term) (acc : list EAst.term) : EAst.term := match t with - | PCUICAst.tApp f a => compile_value_box f (acc ++ [compile_value_box a []]) - | PCUICAst.tConstruct i n _ => EAst.tConstruct i n acc + | PCUICAst.tApp f a => compile_value_box Σ f (compile_value_box Σ a [] :: acc) + | PCUICAst.tConstruct i n _ => + match PCUICAst.PCUICEnvironment.lookup_env Σ (Kernames.inductive_mind i) with + | Some (PCUICAst.PCUICEnvironment.InductiveDecl mdecl) => + EAst.tConstruct i n (skipn mdecl.(PCUICAst.PCUICEnvironment.ind_npars) acc) + | _ => EAst.tVar "error" + end | _ => EAst.tVar "error" end. @@ -372,13 +377,95 @@ Proof. eapply erase_transform_fo_gen; tea. reflexivity. Qed. -Inductive firstorder_evalue : EAst.term -> Prop := - | is_fo i n args : Forall firstorder_evalue args -> firstorder_evalue (EAst.tConstruct i n args). +Inductive firstorder_evalue Σ : EAst.term -> Prop := + | is_fo i n args npars : + EGlobalEnv.lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars -> + Forall (firstorder_evalue Σ) (skipn npars args) -> + firstorder_evalue Σ (EAst.mkApps (EAst.tConstruct i n []) args). -Lemma compile_fo_value Σ t : +Lemma list_size_skipn {A} (l : list A) size n : + list_size size (skipn n l) <= list_size size l. +Proof. + induction n in l |- *. + - rewrite skipn_0 //. + - destruct l. rewrite skipn_nil. now cbn. + rewrite skipn_S. cbn. specialize (IHn l); lia. +Qed. + +Section elim. + Context (Σ : E.global_context). + Context (P : EAst.term -> Prop). + Context (ih : (forall i n args npars, + EGlobalEnv.lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars -> + Forall (firstorder_evalue Σ) (skipn npars args) -> + Forall P (skipn npars args) -> + P (EAst.mkApps (EAst.tConstruct i n []) args))). + + Equations? firstorder_evalue_elim (t : EAst.term) (fo : firstorder_evalue Σ t) : P t by wf t (MR lt EInduction.size) := + { | _, is_fo i n args npars hl hf => _ }. + Proof. + eapply ih; tea. + eapply In_Forall. intros x hin. + eapply PCUICWfUniverses.Forall_In in hf; tea. + apply firstorder_evalue_elim => //. red. + rewrite EInduction.size_mkApps. + eapply (In_size id EInduction.size) in hin. + unfold id in *. pose proof (list_size_skipn args EInduction.size npars). + change (fun x => EInduction.size x) with EInduction.size in hin. clear -hin H. + eapply Nat.lt_le_trans; tea. cbn. lia. + Qed. +End elim. + +(*Lemma firstorder_evalue_elim Σ {P : EAst.term -> Prop} : + (forall i n args npars, + EGlobalEnv.lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars -> + Forall (firstorder_evalue Σ) (skipn npars args) -> + Forall P (skipn npars args) -> + P (EAst.mkApps (EAst.tConstruct i n []) args)) -> + forall t, firstorder_evalue Σ t -> P t. +Proof. + intros Hf. + induction t using term_size + fix aux 2. + intros t fo; destruct fo.x + eapply Hf => //; tea. + clear H. + move: H0. + move: args H0. + fix aux' 2. + intros args []; constructor. + now apply aux. now apply aux'. +Qed.*) + +Inductive firstorder_evalue_block : EAst.term -> Prop := + | is_fo_block i n args : + Forall (firstorder_evalue_block) args -> + firstorder_evalue_block (EAst.tConstruct i n args). + +Lemma firstorder_evalue_block_elim {P : EAst.term -> Prop} : + (forall i n args, + Forall firstorder_evalue_block args -> + Forall P args -> + P (EAst.tConstruct i n args)) -> + forall t, firstorder_evalue_block t -> P t. +Proof. + intros Hf. + fix aux 2. + intros t fo; destruct fo. + eapply Hf => //. + move: args H. + fix aux' 2. + intros args []; constructor. + now apply aux. now apply aux'. +Qed. + +(* Import PCUICAst. + +Lemma compile_fo_value (Σ : global_env_ext) Σ' t : PCUICFirstorder.firstorder_value Σ [] t -> - firstorder_evalue (compile_value_erase t []). -Proof. Admitted. + erases_global + firstorder_evalue Σ (compile_value_erase t []). +Proof. Admitted. *) Fixpoint compile_evalue_box (t : EAst.term) (acc : list EAst.term) := match t with @@ -388,6 +475,7 @@ Fixpoint compile_evalue_box (t : EAst.term) (acc : list EAst.term) := end. Import MetaCoq.Common.Transform. +From Coq Require Import Morphisms. Module ETransformPresFO. Section Opt. @@ -395,50 +483,263 @@ Module ETransformPresFO. Context {eval : program env EAst.term -> EAst.term -> Prop}. Context {eval' : program env' EAst.term -> EAst.term -> Prop}. Context (o : Transform.t _ _ _ _ eval eval'). - Context (compile_fo_value : EAst.term -> EAst.term). + Context (firstorder_value : program env EAst.term -> Prop). + Context (firstorder_value' : program env' EAst.term -> Prop). + Context (compile_fo_value : forall p : program env EAst.term, o.(pre) p -> + firstorder_value p -> program env' EAst.term). Class t := - { preserves_fo : forall v, firstorder_evalue v -> firstorder_evalue (compile_fo_value v); - transform_fo : forall v (pr : o.(pre) v), - firstorder_evalue v.2 -> (o.(transform) v pr).2 = compile_fo_value v.2 }. + { preserves_fo : forall p pr (fo : firstorder_value p), firstorder_value' (compile_fo_value p pr fo); + transform_fo : forall v (pr : o.(pre) v) (fo : firstorder_value v), o.(transform) v pr = compile_fo_value v pr fo }. End Opt. + Section ExtEq. + Context {env env' : Type}. + Context {eval : program env EAst.term -> EAst.term -> Prop}. + Context {eval' : program env' EAst.term -> EAst.term -> Prop}. + Context (o : Transform.t _ _ _ _ eval eval'). + Context (firstorder_value : program env EAst.term -> Prop). + Context (firstorder_value' : program env' EAst.term -> Prop). + + Lemma proper_pres (compile_fo_value compile_fo_value' : forall p : program env EAst.term, o.(pre) p -> firstorder_value p -> program env' EAst.term) : + (forall p pre fo, compile_fo_value p pre fo = compile_fo_value' p pre fo) -> + t o firstorder_value firstorder_value' compile_fo_value <-> + t o firstorder_value firstorder_value' compile_fo_value'. + Proof. + intros Hfg. + split; move=> []; split; eauto. + - now intros ? ? ?; rewrite -Hfg. + - now intros v pr ?; rewrite -Hfg. + - now intros ???; rewrite Hfg. + - now intros ???; rewrite Hfg. + Qed. + End ExtEq. Section Comp. Context {env env' env'' : Type}. Context {eval : program env EAst.term -> EAst.term -> Prop}. Context {eval' : program env' EAst.term -> EAst.term -> Prop}. Context {eval'' : program env'' EAst.term -> EAst.term -> Prop}. - Context (compile_fo_value compile_fo_value' : EAst.term -> EAst.term). + Context (firstorder_value : program env EAst.term -> Prop). + Context (firstorder_value' : program env' EAst.term -> Prop). + Context (firstorder_value'' : program env'' EAst.term -> Prop). + Context (o : Transform.t _ _ _ _ eval eval') (o' : Transform.t _ _ _ _ eval' eval''). + Context compile_fo_value compile_fo_value' + (oext : t o firstorder_value firstorder_value' compile_fo_value) + (o'ext : t o' firstorder_value' firstorder_value'' compile_fo_value') + (hpp : (forall p, o.(post) p -> o'.(pre) p)). Local Obligation Tactic := idtac. + + Definition compose_compile_fo_value (p : program env EAst.term) (pr : o.(pre) p) (fo : firstorder_value p) : program env'' EAst.term := + compile_fo_value' (compile_fo_value p pr fo) (eq_rect_r (o'.(pre)) (hpp _ (correctness o p pr)) (eq_sym (oext.(transform_fo _ _ _ _) _ _ _))) (oext.(preserves_fo _ _ _ _) p pr fo). + #[global] - Instance compose (o : Transform.t _ _ _ _ eval eval') - (o' : Transform.t _ _ _ _ eval' eval'') - (oext : t o compile_fo_value) - (o'ext : t o' compile_fo_value') - (hpp : (forall p, o.(post) p -> o'.(pre) p)) - : t (Transform.compose o o' hpp) (compile_fo_value' ∘ compile_fo_value). + Instance compose + : t (Transform.compose o o' hpp) firstorder_value firstorder_value'' compose_compile_fo_value. Proof. split. - - intros. eapply preserves_fo; tea. eapply preserves_fo; tea. + - intros. eapply o'ext.(preserves_fo _ _ _ _); tea. - intros. cbn. unfold run, time. - rewrite o'ext.(transform_fo _ _). - rewrite oext.(transform_fo _ _) //. - eapply preserves_fo; tea. - rewrite oext.(transform_fo _ _) //. + unfold compose_compile_fo_value. + set (cor := correctness o v pr). clearbody cor. move: cor. + set (foo := eq_sym (transform_fo _ _ _ _ _ _ _)). clearbody foo. + destruct foo. cbn. intros cor. + apply o'ext.(transform_fo _ _ _ _). Qed. End Comp. End ETransformPresFO. -Axiom lambdabox_pres_fo : ETransformPresFO.t verified_lambdabox_pipeline (flip compile_evalue_box []). +Import EWellformed. + +Lemma compile_value_box_mkApps Σ i n ui args acc : + compile_value_box Σ (PCUICAst.mkApps (PCUICAst.tConstruct i n ui) args) acc = + EAst.tConstruct i n (List.map (flip compile_value_box []) args ++ acc). +Proof. + revert acc; induction args using rev_ind. + - cbn. auto. + - intros acc. rewrite PCUICAstUtils.mkApps_app /=. cbn. + now rewrite IHargs map_app /= -app_assoc /=. +Qed. + +Lemma compile_evalue_box_mkApps i n ui args acc : + compile_evalue_box (EAst.mkApps (EAst.tConstruct i n ui) args) acc = + EAst.tConstruct i n (List.map (flip compile_evalue_box []) args ++ acc). +Proof. + revert acc; induction args using rev_ind. + - cbn. auto. + - intros acc. rewrite EAstUtils.mkApps_app /=. cbn. + now rewrite IHargs map_app /= -app_assoc /=. +Qed. + +Lemma compile_evalue_erase Σ v : + PCUICFirstorder.firstorder_value Σ [] v -> + compile_evalue_box (compile_value_erase v []) [] = compile_value_box v []. +Proof. + move=> fo; move: v fo. + apply: PCUICFirstorder.firstorder_value_inds. + intros i n ui u args pandi hty hargs ih isp. + rewrite compile_value_erase_mkApps compile_value_box_mkApps compile_evalue_box_mkApps. + rewrite !app_nil_r. f_equal. + rewrite map_map. + now eapply map_ext_Forall. +Qed. -Lemma transform_lambda_box_firstorder (Σer : EEnvMap.GlobalContextMap.t) v pre : - firstorder_evalue v -> - (transform verified_lambdabox_pipeline (Σer, v) pre).2 = (compile_evalue_box v []). +Lemma compile_evalue_box_firstorder {efl : EEnvFlags} Σ v : + has_cstr_params = false -> + EWellformed.wf_glob Σ -> + firstorder_evalue Σ v -> firstorder_evalue_block (flip compile_evalue_box [] v). +Proof. + intros hpars wf. + move: v; apply: firstorder_evalue_elim. + intros. + rewrite /flip compile_evalue_box_mkApps app_nil_r. + rewrite /EGlobalEnv.lookup_inductive_pars /= in H. + destruct EGlobalEnv.lookup_minductive eqn:e => //. + eapply wellformed_lookup_inductive_pars in hpars; tea => //. + noconf H. rewrite hpars in H1. rewrite skipn_0 in H1. + constructor. ELiftSubst.solve_all. +Qed. + +Definition fo_evalue (p : program E.global_context EAst.term) : Prop := firstorder_evalue p.1 p.2. +Definition fo_evalue_map (p : program EEnvMap.GlobalContextMap.t EAst.term) : Prop := firstorder_evalue p.1 p.2. + +#[global] Instance rebuild_wf_env_transform_pres {fl : WcbvFlags} {efl : EEnvFlags} we : + ETransformPresFO.t + (rebuild_wf_env_transform we) fo_evalue fo_evalue_map (fun p pr fo => rebuild_wf_env p pr.p1). +Proof. split => //. Qed. + +Lemma wf_glob_lookup_inductive_pars {efl : EEnvFlags} (Σ : E.global_context) (kn : Kernames.kername) : + has_cstr_params = false -> + wf_glob Σ -> + forall pars, EGlobalEnv.lookup_inductive_pars Σ kn = Some pars -> pars = 0. +Proof. + intros hasp wf. + rewrite /EGlobalEnv.lookup_inductive_pars. + destruct EGlobalEnv.lookup_minductive eqn:e => //=. + eapply wellformed_lookup_inductive_pars in e => //. congruence. +Qed. + +#[global] Instance inline_projections_optimization_pres {fl : WcbvFlags} + (efl := EInlineProjections.switch_no_params all_env_flags) {wcon : with_constructor_as_block = false} + {has_rel : has_tRel} {has_box : has_tBox} : + ETransformPresFO.t + (inline_projections_optimization (wcon := wcon) (hastrel := has_rel) (hastbox := has_box)) + fo_evalue_map fo_evalue (fun p pr fo => (EInlineProjections.optimize_env p.1, p.2)). +Proof. split => //. + - intros [] pr fo. + cbn in *. + destruct pr as [pr _]. destruct pr as [pr wf]; cbn in *. + clear wf; move: t1 fo. unfold fo_evalue, fo_evalue_map. cbn. + apply: firstorder_evalue_elim; intros. + econstructor. + rewrite EInlineProjections.lookup_inductive_pars_optimize in H => //; tea. auto. + - rewrite /fo_evalue_map. intros [] pr fo. cbn in *. unfold EInlineProjections.optimize_program. cbn. f_equal. + destruct pr as [[pr _] _]. cbn in *. move: t1 fo. + apply: firstorder_evalue_elim; intros. + eapply wf_glob_lookup_inductive_pars in H => //. subst npars; rewrite skipn_0 in H0 H1. + rewrite EInlineProjections.optimize_mkApps /=. f_equal. + ELiftSubst.solve_all. +Qed. + +#[global] Instance remove_match_on_box_pres {fl : WcbvFlags} {efl : EEnvFlags} {wcon : with_constructor_as_block = false} + {has_rel : has_tRel} {has_box : has_tBox} : + has_cstr_params = false -> + ETransformPresFO.t + (remove_match_on_box_trans (wcon := wcon) (hastrel := has_rel) (hastbox := has_box)) + fo_evalue_map fo_evalue (fun p pr fo => (EOptimizePropDiscr.remove_match_on_box_env p.1, p.2)). +Proof. split => //. + - unfold fo_evalue, fo_evalue_map; intros [] pr fo. cbn in *. + destruct pr as [pr _]. destruct pr as [pr wf]; cbn in *. + clear wf; move: t1 fo. + apply: firstorder_evalue_elim; intros. + econstructor; tea. + rewrite EOptimizePropDiscr.lookup_inductive_pars_optimize in H0 => //; tea. + - intros [] pr fo. + cbn in *. + unfold EOptimizePropDiscr.remove_match_on_box_program; cbn. f_equal. + destruct pr as [[pr _] _]; cbn in *; move: t1 fo. + apply: firstorder_evalue_elim; intros. + eapply wf_glob_lookup_inductive_pars in H0 => //. subst npars; rewrite skipn_0 in H2. + rewrite EOptimizePropDiscr.remove_match_on_box_mkApps /=. f_equal. + ELiftSubst.solve_all. +Qed. + +#[global] Instance remove_params_optimization_pres {fl : WcbvFlags} {wcon : with_constructor_as_block = false} : + ETransformPresFO.t + (remove_params_optimization (wcon := wcon)) + fo_evalue_map fo_evalue (fun p pr fo => (ERemoveParams.strip_env p.1, ERemoveParams.strip p.1 p.2)). +Proof. split => //. + intros [] pr fo. + cbn [transform remove_params_optimization] in *. + destruct pr as [[pr _] _]; cbn -[ERemoveParams.strip] in *; move: t1 fo. + apply: firstorder_evalue_elim; intros. + rewrite ERemoveParams.strip_mkApps //. cbn -[EGlobalEnv.lookup_inductive_pars]. rewrite H. + econstructor. cbn -[EGlobalEnv.lookup_inductive_pars]. + now eapply ERemoveParams.lookup_inductive_pars_strip in H; tea. + rewrite skipn_0 /=. + rewrite skipn_map. + ELiftSubst.solve_all. +Qed. + +#[global] Instance constructors_as_blocks_transformation_pres {efl : EWellformed.EEnvFlags} + {has_app : has_tApp} {has_rel : has_tRel} {hasbox : has_tBox} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = false} : + ETransformPresFO.t + (@constructors_as_blocks_transformation efl has_app has_rel hasbox has_pars has_cstrblocks) + fo_evalue_map (fun p => firstorder_evalue_block p.2) + (fun p pr fo => (transform_blocks_env p.1, compile_evalue_box p.2 [])). +Proof. + split. + - intros v pr fo; eapply compile_evalue_box_firstorder; tea. apply pr. + - move=> [Σ v] /= pr fo. rewrite /flip. + clear pr. move: v fo. + apply: firstorder_evalue_elim; intros. + rewrite /transform_blocks_program /=. f_equal. + rewrite EConstructorsAsBlocks.transform_blocks_decompose. + rewrite EAstUtils.decompose_app_mkApps // /=. + rewrite compile_evalue_box_mkApps app_nil_r. + admit. +Admitted. + + +#[global] Instance guarded_to_unguarded_fix_pres {efl : EWellformed.EEnvFlags} + {has_guard : with_guarded_fix} {has_cstrblocks : with_constructor_as_block = false} : + ETransformPresFO.t + (@guarded_to_unguarded_fix default_wcbv_flags has_cstrblocks efl has_guard) + fo_evalue_map fo_evalue_map + (fun p pr fo => p). +Proof. + split => //. +Qed. + +Lemma lambdabox_pres_fo : + exists compile_value, ETransformPresFO.t verified_lambdabox_pipeline fo_evalue_map (fun p => firstorder_evalue_block p.2) compile_value /\ + forall p pr fo, (compile_value p pr fo).2 = compile_evalue_box (ERemoveParams.strip p.1 p.2) []. +Proof. + eexists. + split. + unfold verified_lambdabox_pipeline. + unshelve eapply ETransformPresFO.compose; tc. shelve. + 2:intros p pr fo; unfold ETransformPresFO.compose_compile_fo_value; f_equal. 2:cbn. + unshelve eapply ETransformPresFO.compose; tc. shelve. + 2:unfold ETransformPresFO.compose_compile_fo_value; cbn. + unshelve eapply ETransformPresFO.compose; tc. shelve. + 2:unfold ETransformPresFO.compose_compile_fo_value; cbn. + unshelve eapply ETransformPresFO.compose; tc. shelve. + 2:unfold ETransformPresFO.compose_compile_fo_value; cbn. + unshelve eapply ETransformPresFO.compose. shelve. eapply remove_match_on_box_pres => //. + unfold ETransformPresFO.compose_compile_fo_value; cbn -[ERemoveParams.strip ERemoveParams.strip_env]. + reflexivity. +Qed. + +Lemma transform_lambda_box_firstorder (Σer : EEnvMap.GlobalContextMap.t) p pre : + firstorder_evalue Σer p -> + (transform verified_lambdabox_pipeline (Σer, p) pre).2 = (compile_evalue_box (ERemoveParams.strip Σer p) []). Proof. intros fo. - now rewrite (ETransformPresFO.transform_fo _ _ (t:=lambdabox_pres_fo)). + destruct lambdabox_pres_fo as [fn [tr hfn]]. + rewrite (ETransformPresFO.transform_fo _ _ _ _ (t:=tr)). + nowrewrite hfn. Qed. Section pipeline_theorem. @@ -488,7 +789,7 @@ Section pipeline_theorem. Let Σ_t := (transform verified_erasure_pipeline (Σ, t) precond).1. Let t_t := (transform verified_erasure_pipeline (Σ, t) precond).2. - Let v_t := compile_value_box v []. + Let v_t := compile_value_box Σ v []. Lemma fo_v : PCUICFirstorder.firstorder_value Σ [] v. Proof. @@ -498,7 +799,6 @@ Section pipeline_theorem. - eapply PCUICWcbvEval.eval_to_value; eauto. Qed. - Lemma v_t_spec : v_t = (transform verified_erasure_pipeline (Σ, v) precond2).2. Proof. unfold v_t. generalize fo_v, precond2. clear. @@ -511,16 +811,18 @@ Section pipeline_theorem. cbn [fst snd]. intros h. destruct_compose. - rewrite erase_transform_fo. cbn. admit. + rewrite erase_transform_fo. cbn. unfold global_env_ext_map_global_env_ext; cbn. + todo "expand lets preserves fo values". set (Σer := (transform erase_transform _ _).1). cbn [fst snd]. intros pre'. symmetry. erewrite transform_lambda_box_firstorder; tea. - eapply obseq_lambdabox. intros kn decl H; exact H. - Unshelve. 2:exact pre'. 2:exact (Σer, compile_value_erase v []). - cbn [snd]. - red. cbn. - Admitted. + now eapply compile_evalue_erase. + clear -hv. + move: v hv; eapply PCUICFirstorder.firstorder_value_inds. + intros. rewrite compile_value_erase_mkApps app_nil_r. + econstructor. ELiftSubst.solve_all. + Qed. Import PCUICWfEnv. diff --git a/erasure/theories/EInlineProjections.v b/erasure/theories/EInlineProjections.v index 7d760e3b2..1772da104 100644 --- a/erasure/theories/EInlineProjections.v +++ b/erasure/theories/EInlineProjections.v @@ -422,6 +422,17 @@ Proof. destruct lookup_env as [[decl|]|] => //. Qed. +Lemma lookup_inductive_pars_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : + wf_glob Σ -> + EGlobalEnv.lookup_inductive_pars Σ ind = EGlobalEnv.lookup_inductive_pars (optimize_env Σ) ind. +Proof. + rewrite /lookup_inductive_pars => wf. + rewrite /lookup_inductive /lookup_minductive. + rewrite (lookup_env_optimize ind wf). + rewrite /GlobalContextMap.lookup_inductive /GlobalContextMap.lookup_minductive. + destruct lookup_env as [[decl|]|] => //. +Qed. + Lemma is_propositional_cstr_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind c : wf_glob Σ -> constructor_isprop_pars_decl Σ ind c = constructor_isprop_pars_decl (optimize_env Σ) ind c. diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 4e50ee4ca..0bbc6b016 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -491,6 +491,16 @@ Proof. destruct lookup_env as [[decl|]|] => //. Qed. +Lemma lookup_inductive_pars_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : + wf_glob Σ -> + EGlobalEnv.lookup_inductive_pars Σ ind = EGlobalEnv.lookup_inductive_pars (remove_match_on_box_env Σ) ind. +Proof. + rewrite /lookup_inductive_pars => wf. + rewrite /lookup_inductive /lookup_minductive. + rewrite (lookup_env_remove_match_on_box ind wf). + rewrite /GlobalContextMap.lookup_inductive /GlobalContextMap.lookup_minductive. + destruct lookup_env as [[decl|]|] => //. +Qed. Lemma closed_iota_red pars c args brs br : forallb (closedn 0) args -> diff --git a/erasure/theories/ERemoveParams.v b/erasure/theories/ERemoveParams.v index 4d6937bf0..52a8c3414 100644 --- a/erasure/theories/ERemoveParams.v +++ b/erasure/theories/ERemoveParams.v @@ -430,6 +430,17 @@ Proof. destruct g; simpl; auto. do 2 destruct nth_error => //. Qed. +Lemma lookup_inductive_pars_strip {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : + wf_glob Σ -> + forall pars, lookup_inductive_pars Σ ind = Some pars -> + EGlobalEnv.lookup_inductive_pars (strip_env Σ) ind = Some 0. +Proof. + rewrite /lookup_inductive_pars => wf pars. + rewrite /lookup_inductive /lookup_minductive. + rewrite (lookup_env_strip _ ind). + destruct lookup_env as [[decl|]|] => //=. +Qed. + Arguments eval {wfl}. Arguments isEtaExp : simpl never. From 0e5938a5d1b93ba7b6e1cc3344efd8673dd04c13 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 7 Sep 2023 16:33:11 +0200 Subject: [PATCH 47/61] Finish main erasure correctness proof. Need to move expand_lets reasoning to a first step --- erasure-plugin/theories/ETransform.v | 24 +- erasure-plugin/theories/ErasureCorrectness.v | 412 ++- erasure/_CoqProject.in | 1 + erasure/theories/EArities.v | 21 +- erasure/theories/ESubstitution.v | 8 +- erasure/theories/ErasureCorrectness.v | 26 +- erasure/theories/ErasureFunction.v | 2762 +--------------- erasure/theories/ErasureFunctionProperties.v | 2877 +++++++++++++++++ erasure/theories/Extract.v | 22 +- erasure/theories/Typed/ErasureCorrectness.v | 7 +- .../theories/Typed/ExtractionCorrectness.v | 1 + pcuic/theories/PCUICFirstorder.v | 36 +- 12 files changed, 3225 insertions(+), 2972 deletions(-) create mode 100644 erasure/theories/ErasureFunctionProperties.v diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index bc6694a18..f108396d6 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -8,8 +8,8 @@ Set Warnings "-notation-overridden". From MetaCoq.PCUIC Require PCUICAst PCUICAstUtils PCUICProgram PCUICWeakeningEnvSN. Set Warnings "+notation-overridden". From MetaCoq.SafeChecker Require Import PCUICErrors PCUICWfEnv PCUICWfEnvImpl. -From MetaCoq.Erasure Require EAstUtils ErasureFunction ErasureCorrectness Extract - EOptimizePropDiscr ERemoveParams EProgram. +From MetaCoq.Erasure Require EAstUtils ErasureCorrectness Extract EOptimizePropDiscr ERemoveParams EProgram. +From MetaCoq.Erasure Require Import ErasureFunction ErasureFunctionProperties. From MetaCoq.TemplatePCUIC Require Import PCUICTransform. Import PCUICAst (term) PCUICProgram PCUICTransform (eval_pcuic_program) Extract EProgram @@ -148,14 +148,14 @@ Program Definition erase_pcuic_program {guard : abstract_guard_impl} (p : pcuic_ let wfext := @abstract_make_wf_env_ext _ optimized_abstract_env_impl wfe p.1.2 _ in let t := ErasureFunction.erase (normalization_in:=_) optimized_abstract_env_impl wfext nil p.2 (fun Σ wfΣ => let '(sq (T; ty)) := wt in PCUICTyping.iswelltyped ty) in - let Σ' := ErasureFunction.erase_global_fast (normalization_in:=_) optimized_abstract_env_impl + let Σ' := ErasureFunctionProperties.erase_global_fast (normalization_in:=_) optimized_abstract_env_impl (EAstUtils.term_global_deps t) wfe (p.1.(PCUICAst.PCUICEnvironment.declarations)) _ in (EEnvMap.GlobalContextMap.make Σ'.1 _, t). Next Obligation. unshelve edestruct erase_pcuic_program_normalization_helper; cbn in *; eauto. Defined. Next Obligation. eapply wf_glob_fresh. - eapply ErasureFunction.erase_global_fast_wf_glob. + eapply ErasureFunctionProperties.erase_global_fast_wf_glob. unshelve edestruct erase_pcuic_program_normalization_helper; cbn in *; eauto. Defined. @@ -175,9 +175,9 @@ Lemma expanded_erase_program {guard : abstract_guard_impl} p {normalization_in n Proof. intros [etaenv etat]. split; unfold erase_program, erase_pcuic_program. - eapply ErasureFunction.expanded_erase_global_fast, etaenv; try reflexivity; eauto. + eapply ErasureFunctionProperties.expanded_erase_global_fast, etaenv; try reflexivity; eauto. unshelve edestruct erase_pcuic_program_normalization_helper; cbn in *; eauto. - apply: (ErasureFunction.expanded_erase_fast (X_type:=optimized_abstract_env_impl)). + apply: (ErasureFunctionProperties.expanded_erase_fast (X_type:=optimized_abstract_env_impl)). unshelve edestruct erase_pcuic_program_normalization_helper; cbn in *; eauto. reflexivity. exact etat. Qed. @@ -214,14 +214,14 @@ Next Obligation. destruct erase_program eqn:e. split. - unfold erase_program, erase_pcuic_program in e. - set (egf := ErasureFunction.erase_global_fast _ _ _ _ _) in e. - set (ef := ErasureFunction.erase _ _ _ _ _) in e. + set (egf := erase_global_fast _ _ _ _ _) in e. + set (ef := erase _ _ _ _ _) in e. cbn -[egf ef] in e. injection e. intros <- <-. split. - eapply ErasureFunction.erase_global_fast_wf_glob; eauto; + eapply erase_global_fast_wf_glob; eauto; try match goal with H : _ |- _ => eapply H end. unshelve edestruct erase_pcuic_program_normalization_helper; cbn in *; eauto. - apply: (ErasureFunction.erase_wellformed_fast (X_type:=optimized_abstract_env_impl)); eauto; + apply: (erase_wellformed_fast (X_type:=optimized_abstract_env_impl)); eauto; try match goal with H : _ |- _ => eapply H end. unshelve edestruct erase_pcuic_program_normalization_helper; cbn in *; eauto. - rewrite -e. cbn. @@ -240,8 +240,8 @@ Next Obligation. set (Σ' := build_wf_env_from_env _ _). assert (ev' :forall Σ0 : PCUICAst.PCUICEnvironment.global_env, Σ0 = Σ' -> PCUICWcbvEval.eval Σ0 t v). { intros; now subst. } - eapply (ErasureFunction.erase_correct optimized_abstract_env_impl Σ' Σ.2 _ _ _ _ _ (EAstUtils.term_global_deps _)) in ev'. - 4:{ erewrite <- ErasureFunction.erase_global_deps_fast_spec. reflexivity. } + eapply (erase_correct optimized_abstract_env_impl Σ' Σ.2 _ _ _ _ _ (EAstUtils.term_global_deps _)) in ev'. + 4:{ erewrite <- erase_global_deps_fast_spec. reflexivity. } all:trea. 2:eapply Kernames.KernameSet.subset_spec; reflexivity. destruct ev' as [v' [he [hev]]]. exists v'; split => //. diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 5c0facb2e..5b32b75a6 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -4,9 +4,9 @@ From MetaCoq.Common Require Import Transform config. From MetaCoq.Utils Require Import bytestring utils. From MetaCoq.PCUIC Require PCUICAst PCUICAstUtils PCUICProgram. From MetaCoq.SafeChecker Require Import PCUICErrors PCUICWfEnvImpl. -From MetaCoq.Erasure Require EAstUtils ErasureFunction ErasureCorrectness EPretty Extract. +From MetaCoq.Erasure Require EAstUtils ErasureCorrectness EPretty Extract. From MetaCoq Require Import ETransform EConstructorsAsBlocks. -From MetaCoq.Erasure Require Import EWcbvEvalNamed. +From MetaCoq.Erasure Require Import EWcbvEvalNamed ErasureFunction ErasureFunctionProperties. From MetaCoq.ErasurePlugin Require Import Erasure. Import PCUICProgram. (* Import TemplateProgram (template_eta_expand). @@ -269,7 +269,7 @@ Lemma extends_erase_pcuic_program (efl := EWcbvEval.default_wcbv_flags) {guard : @PCUICFirstorder.firstorder_ind Σ (PCUICFirstorder.firstorder_env Σ) i -> let pt := @erase_pcuic_program guard (Σ, t) nin0 nin0' wf' ty' in let pv := @erase_pcuic_program guard (Σ, v) nin nin' wf ty in - EGlobalEnv.extends pv.1 pt.1 /\ ∥ eval pt.1 pt.2 pv.2 ∥. + EGlobalEnv.extends pv.1 pt.1 /\ ∥ eval pt.1 pt.2 pv.2 ∥ /\ firstorder_evalue pt.1 pv.2. Proof. intros ev axf ht fo. cbn -[erase_pcuic_program]. @@ -285,26 +285,25 @@ Proof. set (env := build_wf_env_from_env _ _). set (X := PCUICWfEnv.abstract_make_wf_env_ext _ _ _). set (X' := PCUICWfEnv.abstract_make_wf_env_ext _ _ _). - unfold ErasureFunction.erase_global_fast. + unfold erase_global_fast. set (prf7 := (fun (Σ0 : PCUICAst.PCUICEnvironment.global_env) => _)). set (et := ErasureFunction.erase _ _ _ _ _). set (et' := ErasureFunction.erase _ _ _ _ _). destruct Σ as [Σ ext]. cbn -[et et' PCUICWfEnv.abstract_make_wf_env_ext] in *. - unshelve (epose proof ErasureFunction.erase_global_deps_fast_erase_global_deps as [norm eq]; + unshelve (epose proof erase_global_deps_fast_erase_global_deps as [norm eq]; erewrite eq). { cbn. now intros ? ->. } - unshelve (epose proof ErasureFunction.erase_global_deps_fast_erase_global_deps as [norm' eq']; + unshelve (epose proof erase_global_deps_fast_erase_global_deps as [norm' eq']; erewrite eq'). { cbn. now intros ? ->. } set (prf := (fun (Σ0 : PCUICAst.PCUICEnvironment.global_env) => _)). cbn in prf. rewrite (ErasureFunction.erase_global_deps_irr optimized_abstract_env_impl (EAstUtils.term_global_deps et) env' env _ prf prf). { cbn. now intros ? ? -> ->. } clearbody prf0 prf1 prf2 prf3 prf4 prf5 prf6 prf7. - epose proof (ErasureFunction.erase_correct_strong optimized_abstract_env_impl (v:=v) env ext prf2 + epose proof (erase_correct_strong optimized_abstract_env_impl (v:=v) env ext prf2 (PCUICAst.PCUICEnvironment.declarations Σ) norm' prf prf6 X eq_refl axf ht fo). pose proof wf as []. - forward H by (eapply Kernames.KernameSet.subset_spec; reflexivity). forward H by unshelve (eapply PCUICClassification.wcbveval_red; tea). forward H. { intros [? hr]. @@ -312,14 +311,18 @@ Proof. eapply PCUICFirstorder.firstorder_value_spec; tea. apply X0. constructor. eapply PCUICClassification.subject_reduction_eval; tea. eapply PCUICWcbvEval.eval_to_value; tea. } - destruct H as [wt' []]. + destruct H as [wt' [[] hfo]]. split => //. - eapply (ErasureFunction.erase_global_deps_eval optimized_abstract_env_impl env env' ext). + eapply (erase_global_deps_eval optimized_abstract_env_impl env env' ext). unshelve erewrite (ErasureFunction.erase_irrel_global_env (X_type:=optimized_abstract_env_impl) (t:=v)); tea. red. intros. split; reflexivity. + split => //. sq. unfold et', et. unshelve erewrite (ErasureFunction.erase_irrel_global_env (X_type:=optimized_abstract_env_impl) (t:=v)); tea. red. intros. split; reflexivity. + subst et et' X X'. + unshelve erewrite (ErasureFunction.erase_irrel_global_env (X_type:=optimized_abstract_env_impl) (t:=v)); tea. + red. intros. split; reflexivity. Qed. Lemma expand_lets_fo (Σ : global_env_ext_map) t : @@ -336,23 +339,170 @@ Proof. now eapply forall_map_id_spec. Qed. +Definition pcuic_lookup_inductive_pars Σ ind := + match PCUICAst.PCUICEnvironment.lookup_env Σ (Kernames.inductive_mind ind) with + | Some (PCUICAst.PCUICEnvironment.InductiveDecl mdecl) => Some mdecl.(PCUICAst.PCUICEnvironment.ind_npars) + | _ => None + end. + Fixpoint compile_value_box Σ (t : PCUICAst.term) (acc : list EAst.term) : EAst.term := match t with | PCUICAst.tApp f a => compile_value_box Σ f (compile_value_box Σ a [] :: acc) | PCUICAst.tConstruct i n _ => - match PCUICAst.PCUICEnvironment.lookup_env Σ (Kernames.inductive_mind i) with - | Some (PCUICAst.PCUICEnvironment.InductiveDecl mdecl) => - EAst.tConstruct i n (skipn mdecl.(PCUICAst.PCUICEnvironment.ind_npars) acc) - | _ => EAst.tVar "error" + match pcuic_lookup_inductive_pars Σ i with + | Some npars => EAst.tConstruct i n (skipn npars acc) + | None => EAst.tVar "error" end | _ => EAst.tVar "error" end. From Equations Require Import Equations. + +Inductive firstorder_evalue_block : EAst.term -> Prop := + | is_fo_block i n args : + Forall (firstorder_evalue_block) args -> + firstorder_evalue_block (EAst.tConstruct i n args). + +Lemma firstorder_evalue_block_elim {P : EAst.term -> Prop} : + (forall i n args, + Forall firstorder_evalue_block args -> + Forall P args -> + P (EAst.tConstruct i n args)) -> + forall t, firstorder_evalue_block t -> P t. +Proof. + intros Hf. + fix aux 2. + intros t fo; destruct fo. + eapply Hf => //. + move: args H. + fix aux' 2. + intros args []; constructor. + now apply aux. now apply aux'. +Qed. + +Import EWcbvEval. +Arguments erase_global_deps _ _ _ _ _ : clear implicits. +Arguments erase_global_deps_fast _ _ _ _ _ _ : clear implicits. + +(*Lemma erase_pcuic_program_spec {guard : abstract_guard_impl} + (p : pcuic_program) + (nin : (wf_ext p.1 -> PCUICSN.NormalizationIn p.1)) + (nin' : (wf_ext p.1 -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p.1)) + (wfext : ∥ wf_ext p.1 ∥) + (wt : ∥ ∑ T : PCUICAst.term, p.1;;; [] |- p.2 : T ∥) : + erase_pcuic_program p nin nin'wfext wt = + let et' := @erase optimized_abstract_env_impl + @erase_global_deps optimized_abstract_env_impl*) + +Section PCUICProof. + Import PCUICAst.PCUICEnvironment. + + Definition erase_preserves_inductives Σ Σ' := + (forall kn decl decl', EGlobalEnv.lookup_env Σ' kn = Some (EAst.InductiveDecl decl) -> + lookup_env Σ kn = Some (PCUICAst.PCUICEnvironment.InductiveDecl decl') -> + decl = erase_mutual_inductive_body decl'). + + Lemma lookup_env_in_erase_global_deps X_type X deps decls kn normalization_in prf decl : + EnvMap.EnvMap.fresh_globals decls -> + EGlobalEnv.lookup_env (erase_global_deps X_type deps X decls normalization_in prf).1 kn = Some (EAst.InductiveDecl decl) -> + exists decl', lookup_global decls kn = Some (InductiveDecl decl') /\ decl = erase_mutual_inductive_body decl'. + Proof. + induction decls in deps, X, normalization_in, prf |- *; cbn [erase_global_deps] => //. + destruct a => //. destruct g => //. + - case: (knset_mem_spec k deps) => // hdeps. + cbn [EGlobalEnv.lookup_env fst lookup_env lookup_global]. + { destruct (eqb_spec kn k) => //. + intros hl. eapply IHdecls. now depelim hl. } + { intros hl. depelim hl. + intros hl'. + eapply IHdecls in hl. destruct hl. + exists x. + cbn. + destruct (eqb_spec kn k) => //. subst k. + destruct H0. + now eapply PCUICWeakeningEnv.lookup_global_Some_fresh in H0. + exact hl'. } + - intros hf; depelim hf. + case: (knset_mem_spec k deps) => // hdeps. + cbn [EGlobalEnv.lookup_env fst lookup_env lookup_global]. + { destruct (eqb_spec kn k) => //. + intros hl. noconf hl. subst k. eexists; split; cbn; eauto. + intros hl'. eapply IHdecls => //. tea. } + { intros hl'. eapply IHdecls in hf; tea. destruct hf. + exists x. + cbn. + destruct (eqb_spec kn k) => //. subst k. + destruct H0. + now eapply PCUICWeakeningEnv.lookup_global_Some_fresh in H0. } + Qed. + + Lemma erase_tranform_firstorder (no := PCUICSN.extraction_normalizing) (wfl := default_wcbv_flags) + {p : Transform.program global_env_ext_map PCUICAst.term} {pr v i u args} + {normalization_in : PCUICSN.NormalizationIn p.1} : + forall (wt : p.1 ;;; [] |- p.2 : PCUICAst.mkApps (PCUICAst.tInd i u) args), + axiom_free p.1 -> + @PCUICFirstorder.firstorder_ind p.1 (PCUICFirstorder.firstorder_env p.1) i -> + PCUICWcbvEval.eval p.1 p.2 v -> + forall ep, transform erase_transform p pr = ep -> + erase_preserves_inductives p.1 ep.1 /\ + ∥ EWcbvEval.eval ep.1 ep.2 (compile_value_erase v []) ∥ /\ + firstorder_evalue ep.1 (compile_value_erase v []). + Proof. + destruct p as [Σ t]; cbn. + intros ht ax fo ev [Σe te]; cbn. + unfold erase_program, erase_pcuic_program. + set (obl := ETransform.erase_pcuic_program_obligation_6 _ _ _ _ _ _). + move: obl. + rewrite /erase_global_fast. + set (prf0 := fun (Σ0 : global_env) => _). + set (prf1 := fun (Σ0 : global_env_ext) => _). + set (prf2 := fun (Σ0 : global_env_ext) => _). + set (prf3 := fun (Σ0 : global_env) => _). + set (prf4 := fun n (H : n < _) => _). + set (gext := PCUICWfEnv.abstract_make_wf_env_ext _ _ _). + set (et := erase _ _ _ _ _). + set (g := build_wf_env_from_env _ _). + assert (hprefix: forall Σ0 : global_env, PCUICWfEnv.abstract_env_rel g Σ0 -> declarations Σ0 = declarations g). + { intros Σ' eq; cbn in eq. rewrite eq; reflexivity. } + destruct (@erase_global_deps_fast_erase_global_deps (EAstUtils.term_global_deps et) optimized_abstract_env_impl g + (declarations Σ) prf4 prf3 hprefix) as [nin' eq]. + cbn [fst snd]. + rewrite eq. + set (eg := erase_global_deps _ _ _ _ _ _). + intros obl. + epose proof (@erase_correct_strong optimized_abstract_env_impl g Σ.2 prf0 t v i u args _ _ hprefix prf1 prf2 Σ eq_refl ax ht fo). + pose proof (proj1 pr) as [[]]. + forward H. eapply PCUICClassification.wcbveval_red; tea. + assert (PCUICFirstorder.firstorder_value Σ [] v). + { eapply PCUICFirstorder.firstorder_value_spec; tea. apply w. constructor. + eapply PCUICClassification.subject_reduction_eval; tea. + eapply PCUICWcbvEval.eval_to_value; tea. } + forward H. + { intros [v' redv]. eapply PCUICNormalization.firstorder_value_irred; tea. } + destruct H as [wt' [ev' fo']]. + assert (erase optimized_abstract_env_impl (PCUICWfEnv.abstract_make_wf_env_ext (X_type:=optimized_abstract_env_impl) g Σ.2 prf0) [] v wt' = + compile_value_erase v []). + { clear -H0. + clearbody prf0 prf1. + destruct pr as []. + destruct s as [[]]. + epose proof (erases_erase (X_type := optimized_abstract_env_impl) wt' _ eq_refl). + eapply erases_firstorder' in H; eauto. } + rewrite H in ev', fo'. + intros [=]; subst te Σe. + split => //. + cbn. subst eg. + intros kn decl decl' hl hl'. + eapply lookup_env_in_erase_global_deps in hl as [decl'' [hl eq']]. + rewrite /lookup_env hl in hl'. now noconf hl'. + eapply wf_fresh_globals, w. + Qed. +End PCUICProof. Lemma erase_transform_fo_gen (p : pcuic_program) pr : PCUICFirstorder.firstorder_value p.1 [] p.2 -> - forall ep, transform erase_transform p pr = ep -> ep.2 = compile_value_erase p.2 []. + forall ep, transform erase_transform p pr = ep -> + ep.2 = compile_value_erase p.2 []. Proof. destruct p as [Σ t]. cbn. intros hev ep <-. move: hev pr. @@ -377,87 +527,6 @@ Proof. eapply erase_transform_fo_gen; tea. reflexivity. Qed. -Inductive firstorder_evalue Σ : EAst.term -> Prop := - | is_fo i n args npars : - EGlobalEnv.lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars -> - Forall (firstorder_evalue Σ) (skipn npars args) -> - firstorder_evalue Σ (EAst.mkApps (EAst.tConstruct i n []) args). - -Lemma list_size_skipn {A} (l : list A) size n : - list_size size (skipn n l) <= list_size size l. -Proof. - induction n in l |- *. - - rewrite skipn_0 //. - - destruct l. rewrite skipn_nil. now cbn. - rewrite skipn_S. cbn. specialize (IHn l); lia. -Qed. - -Section elim. - Context (Σ : E.global_context). - Context (P : EAst.term -> Prop). - Context (ih : (forall i n args npars, - EGlobalEnv.lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars -> - Forall (firstorder_evalue Σ) (skipn npars args) -> - Forall P (skipn npars args) -> - P (EAst.mkApps (EAst.tConstruct i n []) args))). - - Equations? firstorder_evalue_elim (t : EAst.term) (fo : firstorder_evalue Σ t) : P t by wf t (MR lt EInduction.size) := - { | _, is_fo i n args npars hl hf => _ }. - Proof. - eapply ih; tea. - eapply In_Forall. intros x hin. - eapply PCUICWfUniverses.Forall_In in hf; tea. - apply firstorder_evalue_elim => //. red. - rewrite EInduction.size_mkApps. - eapply (In_size id EInduction.size) in hin. - unfold id in *. pose proof (list_size_skipn args EInduction.size npars). - change (fun x => EInduction.size x) with EInduction.size in hin. clear -hin H. - eapply Nat.lt_le_trans; tea. cbn. lia. - Qed. -End elim. - -(*Lemma firstorder_evalue_elim Σ {P : EAst.term -> Prop} : - (forall i n args npars, - EGlobalEnv.lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars -> - Forall (firstorder_evalue Σ) (skipn npars args) -> - Forall P (skipn npars args) -> - P (EAst.mkApps (EAst.tConstruct i n []) args)) -> - forall t, firstorder_evalue Σ t -> P t. -Proof. - intros Hf. - induction t using term_size - fix aux 2. - intros t fo; destruct fo.x - eapply Hf => //; tea. - clear H. - move: H0. - move: args H0. - fix aux' 2. - intros args []; constructor. - now apply aux. now apply aux'. -Qed.*) - -Inductive firstorder_evalue_block : EAst.term -> Prop := - | is_fo_block i n args : - Forall (firstorder_evalue_block) args -> - firstorder_evalue_block (EAst.tConstruct i n args). - -Lemma firstorder_evalue_block_elim {P : EAst.term -> Prop} : - (forall i n args, - Forall firstorder_evalue_block args -> - Forall P args -> - P (EAst.tConstruct i n args)) -> - forall t, firstorder_evalue_block t -> P t. -Proof. - intros Hf. - fix aux 2. - intros t fo; destruct fo. - eapply Hf => //. - move: args H. - fix aux' 2. - intros args []; constructor. - now apply aux. now apply aux'. -Qed. (* Import PCUICAst. @@ -467,13 +536,6 @@ Lemma compile_fo_value (Σ : global_env_ext) Σ' t : firstorder_evalue Σ (compile_value_erase t []). Proof. Admitted. *) -Fixpoint compile_evalue_box (t : EAst.term) (acc : list EAst.term) := - match t with - | EAst.tApp f a => compile_evalue_box f (compile_evalue_box a [] :: acc) - | EAst.tConstruct i n _ => EAst.tConstruct i n acc - | _ => EAst.tVar "error" - end. - Import MetaCoq.Common.Transform. From Coq Require Import Morphisms. @@ -552,40 +614,97 @@ End ETransformPresFO. Import EWellformed. -Lemma compile_value_box_mkApps Σ i n ui args acc : +Fixpoint compile_evalue_box_strip Σ (t : EAst.term) (acc : list EAst.term) := + match t with + | EAst.tApp f a => compile_evalue_box_strip Σ f (compile_evalue_box_strip Σ a [] :: acc) + | EAst.tConstruct i n _ => + match lookup_inductive_pars Σ (Kernames.inductive_mind i) with + | Some npars => EAst.tConstruct i n (skipn npars acc) + | None => EAst.tVar "error" + end + | _ => EAst.tVar "error" + end. + +Fixpoint compile_evalue_box (t : EAst.term) (acc : list EAst.term) := + match t with + | EAst.tApp f a => compile_evalue_box f (compile_evalue_box a [] :: acc) + | EAst.tConstruct i n _ => EAst.tConstruct i n acc + | _ => EAst.tVar "error" + end. + +Lemma compile_value_box_mkApps {Σ i n ui args npars acc} : + pcuic_lookup_inductive_pars Σ i = Some npars -> compile_value_box Σ (PCUICAst.mkApps (PCUICAst.tConstruct i n ui) args) acc = - EAst.tConstruct i n (List.map (flip compile_value_box []) args ++ acc). + EAst.tConstruct i n (skipn npars (List.map (flip (compile_value_box Σ) []) args ++ acc)). Proof. revert acc; induction args using rev_ind. - - cbn. auto. + - intros acc. cbn. intros ->. reflexivity. - intros acc. rewrite PCUICAstUtils.mkApps_app /=. cbn. - now rewrite IHargs map_app /= -app_assoc /=. + intros hl. + now rewrite IHargs // map_app /= -app_assoc /=. Qed. -Lemma compile_evalue_box_mkApps i n ui args acc : +Lemma compile_evalue_box_strip_mkApps {Σ i n ui args acc npars} : + lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars -> + compile_evalue_box_strip Σ (EAst.mkApps (EAst.tConstruct i n ui) args) acc = + EAst.tConstruct i n (skipn npars (List.map (flip (compile_evalue_box_strip Σ) []) args ++ acc)). +Proof. + revert acc; induction args using rev_ind. + - intros acc. cbn. intros ->. auto. + - intros acc hl. rewrite EAstUtils.mkApps_app /=. cbn. + now rewrite IHargs // map_app /= -app_assoc /=. +Qed. + +Lemma compile_evalue_box_mkApps {i n ui args acc} : compile_evalue_box (EAst.mkApps (EAst.tConstruct i n ui) args) acc = EAst.tConstruct i n (List.map (flip compile_evalue_box []) args ++ acc). Proof. revert acc; induction args using rev_ind. - - cbn. auto. + - now intros acc. - intros acc. rewrite EAstUtils.mkApps_app /=. cbn. - now rewrite IHargs map_app /= -app_assoc /=. + now rewrite IHargs // map_app /= -app_assoc /=. Qed. +Derive Signature for firstorder_evalue. -Lemma compile_evalue_erase Σ v : +Lemma compile_evalue_erase (Σ : PCUICAst.PCUICEnvironment.global_env_ext) (Σ' : EEnvMap.GlobalContextMap.t) v : + wf Σ.1 -> PCUICFirstorder.firstorder_value Σ [] v -> - compile_evalue_box (compile_value_erase v []) [] = compile_value_box v []. + firstorder_evalue Σ' (compile_value_erase v []) -> + erase_preserves_inductives (PCUICAst.PCUICEnvironment.fst_ctx Σ) Σ' -> + compile_evalue_box_strip Σ' (compile_value_erase v []) [] = compile_value_box (PCUICAst.PCUICEnvironment.fst_ctx Σ) v []. Proof. - move=> fo; move: v fo. + move=> wf fo fo' hΣ; move: v fo fo'. apply: PCUICFirstorder.firstorder_value_inds. intros i n ui u args pandi hty hargs ih isp. - rewrite compile_value_erase_mkApps compile_value_box_mkApps compile_evalue_box_mkApps. + eapply PCUICInductiveInversion.Construct_Ind_ind_eq' in hty as [mdecl [idecl [cdecl [declc _]]]] => //. + rewrite compile_value_erase_mkApps. + intros fo'. depelim fo'. EAstUtils.solve_discr. noconf H1. + assert (npars = PCUICAst.PCUICEnvironment.ind_npars mdecl). + { destruct declc as [[declm decli] declc]. + unshelve eapply declared_minductive_to_gen in declm. 3:exact wf. red in declm. + rewrite /EGlobalEnv.lookup_inductive_pars /EGlobalEnv.lookup_minductive in H. + destruct (PCUICAst.PCUICEnvironment.lookup_env) eqn:hl => //. + noconf declm. + destruct (EGlobalEnv.lookup_env) eqn:hl' => //. destruct g => //. + red in hΣ. + eapply hΣ in hl'; tea. cbn in H. noconf H. subst m. reflexivity. } + subst npars. + rewrite (compile_value_box_mkApps (npars := PCUICAst.PCUICEnvironment.ind_npars mdecl)). + { destruct declc as [[declm decli] declc]. + unshelve eapply declared_minductive_to_gen in declm. 3:exact wf. + rewrite /PCUICAst.declared_minductive_gen in declm. + rewrite /pcuic_lookup_inductive_pars // declm //. } + rewrite (compile_evalue_box_strip_mkApps (npars := PCUICAst.PCUICEnvironment.ind_npars mdecl)) //. + rewrite lookup_inductive_pars_spec //. rewrite !app_nil_r. f_equal. - rewrite map_map. - now eapply map_ext_Forall. + rewrite app_nil_r skipn_map in H0. + eapply Forall_map_inv in H0. + eapply (Forall_skipn _ (PCUICAst.PCUICEnvironment.ind_npars mdecl)) in ih. + rewrite !skipn_map /flip map_map. + ELiftSubst.solve_all. Qed. -Lemma compile_evalue_box_firstorder {efl : EEnvFlags} Σ v : +Lemma compile_evalue_box_firstorder {efl : EEnvFlags} {Σ : EEnvMap.GlobalContextMap.t} v : has_cstr_params = false -> EWellformed.wf_glob Σ -> firstorder_evalue Σ v -> firstorder_evalue_block (flip compile_evalue_box [] v). @@ -593,7 +712,7 @@ Proof. intros hpars wf. move: v; apply: firstorder_evalue_elim. intros. - rewrite /flip compile_evalue_box_mkApps app_nil_r. + rewrite /flip (compile_evalue_box_mkApps) // ?app_nil_r. rewrite /EGlobalEnv.lookup_inductive_pars /= in H. destruct EGlobalEnv.lookup_minductive eqn:e => //. eapply wellformed_lookup_inductive_pars in hpars; tea => //. @@ -697,7 +816,8 @@ Proof. rewrite /transform_blocks_program /=. f_equal. rewrite EConstructorsAsBlocks.transform_blocks_decompose. rewrite EAstUtils.decompose_app_mkApps // /=. - rewrite compile_evalue_box_mkApps app_nil_r. + rewrite compile_evalue_box_mkApps // ?app_nil_r. + (* rewrite lookup_inductive_pars_spec //. *) admit. Admitted. @@ -739,9 +859,17 @@ Proof. intros fo. destruct lambdabox_pres_fo as [fn [tr hfn]]. rewrite (ETransformPresFO.transform_fo _ _ _ _ (t:=tr)). - nowrewrite hfn. + now rewrite hfn. Qed. +Lemma compile_evalue_strip (Σer : EEnvMap.GlobalContextMap.t) p : + firstorder_evalue Σer p -> + compile_evalue_box (ERemoveParams.strip Σer p) [] = compile_evalue_box_strip Σer p []. +Proof. +Admitted. + +Arguments PCUICFirstorder.firstorder_ind _ _ : clear implicits. + Section pipeline_theorem. Instance cf : checker_flags := extraction_checker_flags. @@ -782,14 +910,14 @@ Section pipeline_theorem. - red. cbn. split; eauto. eexists. eapply PCUICClassification.subject_reduction_eval; eauto. - - todo "preservation of eta expandedness". + - todo "preservation of eta expandedness". - cbn. todo "normalization". - todo "normalization". Qed. Let Σ_t := (transform verified_erasure_pipeline (Σ, t) precond).1. Let t_t := (transform verified_erasure_pipeline (Σ, t) precond).2. - Let v_t := compile_value_box Σ v []. + Let v_t := compile_value_box (PCUICExpandLets.trans_global_env Σ) v []. Lemma fo_v : PCUICFirstorder.firstorder_value Σ [] v. Proof. @@ -801,7 +929,7 @@ Section pipeline_theorem. Lemma v_t_spec : v_t = (transform verified_erasure_pipeline (Σ, v) precond2).2. Proof. - unfold v_t. generalize fo_v, precond2. clear. + unfold v_t. generalize fo_v, precond2. intros hv pre. unfold verified_erasure_pipeline. rewrite -transform_compose_assoc. @@ -811,23 +939,40 @@ Section pipeline_theorem. cbn [fst snd]. intros h. destruct_compose. - rewrite erase_transform_fo. cbn. unfold global_env_ext_map_global_env_ext; cbn. - todo "expand lets preserves fo values". + assert (PCUICFirstorder.firstorder_value (PCUICExpandLets.trans_global_env Σ.1, Σ.2) [] v). + { todo "expand lets preserves fo values". } + assert (Normalisation': PCUICSN.NormalizationIn (PCUICExpandLets.trans_global Σ)). + { destruct h as [[] ?]. apply H0. cbn. apply X. } + set (Σ' := build_global_env_map _). + set (p := transform erase_transform _ _). + pose proof (@erase_tranform_firstorder _ h v i u args Normalisation'). + forward H0. + { todo "preserves typing of fo values". } + forward H0. + { cbn. todo "preserves axiom freeness". } + forward H0. + { cbn. todo "preserves fo ind". } + forward H0. + { cbn. todo "preserves values". } + specialize (H0 _ eq_refl). + rewrite /p. + rewrite erase_transform_fo //. set (Σer := (transform erase_transform _ _).1). cbn [fst snd]. intros pre'. symmetry. + destruct Heval as [Heval']. + assert (firstorder_evalue Σer (compile_value_erase v [])). + { apply H0. } erewrite transform_lambda_box_firstorder; tea. - now eapply compile_evalue_erase. - clear -hv. - move: v hv; eapply PCUICFirstorder.firstorder_value_inds. - intros. rewrite compile_value_erase_mkApps app_nil_r. - econstructor. ELiftSubst.solve_all. + rewrite compile_evalue_strip //. + destruct pre as [[wt] ?]. destruct wt. + apply (compile_evalue_erase (PCUICExpandLets.trans_global Σ) Σer) => //. + { cbn. now eapply (@PCUICExpandLetsCorrectness.trans_wf extraction_checker_flags Σ). } + destruct H0. cbn -[transform erase_transform] in H0. apply H0. Qed. Import PCUICWfEnv. - - Lemma verified_erasure_pipeline_theorem : ∥ eval (wfl := extraction_wcbv_flags) Σ_t t_t v_t∥. Proof. @@ -891,7 +1036,7 @@ Section pipeline_theorem. destruct H. destruct s as [[]]. set (wfe := build_wf_env_from_env em (map_squash (PCUICTyping.wf_ext_wf (em, Σ.2)) (map_squash fst (conj (sq (w0, s)) a).p1))). destruct Heval. - eapply (ErasureFunction.firstorder_erases_deterministic optimized_abstract_env_impl wfe Σ.2) in b0. 3:tea. + eapply (ErasureFunctionProperties.firstorder_erases_deterministic optimized_abstract_env_impl wfe Σ.2) in b0. 3:tea. 2:{ cbn. reflexivity. } 2:{ eapply PCUICExpandLetsCorrectness.trans_wcbveval. eapply PCUICWcbvEval.eval_closed; tea. apply HΣ. admit. @@ -908,6 +1053,7 @@ Section pipeline_theorem. clear obs b0 ev e w. eapply extends_erase_pcuic_program. cbn. eapply (PCUICExpandLetsCorrectness.trans_wcbveval (Σ := (Σ.1, Σ.2))). admit. exact X0. + cbn. 2:cbn. 3:cbn. Admitted. Lemma verified_erasure_pipeline_lambda : diff --git a/erasure/_CoqProject.in b/erasure/_CoqProject.in index 3a16a4112..5397184fa 100644 --- a/erasure/_CoqProject.in +++ b/erasure/_CoqProject.in @@ -26,6 +26,7 @@ theories/EArities.v theories/ErasureProperties.v theories/ErasureCorrectness.v theories/ErasureFunction.v +theories/ErasureFunctionProperties.v theories/EExtends.v theories/EOptimizePropDiscr.v theories/EEtaExpandedFix.v diff --git a/erasure/theories/EArities.v b/erasure/theories/EArities.v index 3b11eeb08..666de6ffb 100644 --- a/erasure/theories/EArities.v +++ b/erasure/theories/EArities.v @@ -10,7 +10,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICWellScopedCumulativity PCUICContextConversion PCUICConversion PCUICClassification PCUICSpine PCUICInductives PCUICInductiveInversion PCUICConfluence - PCUICArities PCUICPrincipality. + PCUICArities PCUICPrincipality PCUICFirstorder. From MetaCoq.Erasure Require Import Extract. @@ -836,7 +836,7 @@ Qed. Lemma isErasable_Propositional {Σ : global_env_ext} {Γ ind n u args} : wf_ext Σ -> - isErasable Σ Γ (mkApps (tConstruct ind n u) args) -> isPropositional Σ ind true. + isErasable Σ Γ (mkApps (tConstruct ind n u) args) -> isPropositional Σ ind. Proof. intros wfΣ ise. eapply tConstruct_no_Type in ise; eauto. @@ -871,7 +871,7 @@ Qed. Lemma nisErasable_Propositional {Σ : global_env_ext} {Γ ind n u} : wf_ext Σ -> welltyped Σ Γ (tConstruct ind n u) -> - (isErasable Σ Γ (tConstruct ind n u) -> False) -> isPropositional Σ ind false. + (isErasable Σ Γ (tConstruct ind n u) -> False) -> ~~ isPropositional Σ ind. Proof. intros wfΣ wt ise. destruct wt as [T HT]. @@ -881,7 +881,7 @@ Proof. pose proof d as [decli ?]. destruct (on_declared_constructor d). destruct p as [onind oib]. - red. unfold lookup_inductive. + unfold isPropositional, lookup_inductive. pose proof (wfΣ' := wfΣ.1). unshelve epose proof (decli_ := declared_inductive_to_gen decli); eauto. rewrite (declared_inductive_lookup_gen decli_). @@ -925,18 +925,18 @@ Lemma isPropositional_propositional Σ {wfΣ: wf Σ} (Σ' : E.global_context) in EGlobalEnv.declared_inductive Σ' ind mdecl' idecl' -> erases_mutual_inductive_body mdecl mdecl' -> erases_one_inductive_body idecl idecl' -> - forall b, isPropositional Σ ind b -> EGlobalEnv.inductive_isprop_and_pars Σ' ind = Some (b, mdecl.(ind_npars)). + EGlobalEnv.inductive_isprop_and_pars Σ' ind = Some (isPropositional Σ ind, mdecl.(ind_npars)). Proof. - intros decli decli' [_ indp] [] b. + intros decli decli' [_ indp] []. unfold isPropositional, EGlobalEnv.inductive_isprop_and_pars. unfold lookup_inductive. unshelve epose proof (decli_ := declared_inductive_to_gen decli); eauto. rewrite (declared_inductive_lookup_gen decli_). rewrite (EGlobalEnv.declared_inductive_lookup decli') /= /isPropositionalArity. - destruct H0 as [_ [_ [_ isP]]]. red in isP. + destruct H0 as [_ [_ [_ isP]]]. unfold isPropositionalArity in isP. destruct destArity as [[ctx s]|] eqn:da => //. - rewrite isP. intros ->. f_equal. f_equal. now rewrite indp. + rewrite isP; congruence. congruence. Qed. Lemma isPropositional_propositional_cstr Σ (Σ' : E.global_context) ind c mdecl idecl cdecl mdecl' idecl' : @@ -945,11 +945,10 @@ Lemma isPropositional_propositional_cstr Σ (Σ' : E.global_context) ind c mdecl EGlobalEnv.declared_inductive Σ' ind mdecl' idecl' -> erases_mutual_inductive_body mdecl mdecl' -> erases_one_inductive_body idecl idecl' -> - forall b, isPropositional Σ ind b -> EGlobalEnv.constructor_isprop_pars_decl Σ' ind c = - Some (b, mdecl.(ind_npars), EAst.mkConstructor cdecl.(cstr_name) (context_assumptions cdecl.(cstr_args))). + Some (isPropositional Σ ind, mdecl.(ind_npars), EAst.mkConstructor cdecl.(cstr_name) (context_assumptions cdecl.(cstr_args))). Proof. - intros wfΣ declc decli' em ei b isp. + intros wfΣ declc decli' em ei. pose proof declc as [decli'' _]. eapply isPropositional_propositional in decli''; tea. move: decli''. diff --git a/erasure/theories/ESubstitution.v b/erasure/theories/ESubstitution.v index e4f7b558a..5ff50816b 100644 --- a/erasure/theories/ESubstitution.v +++ b/erasure/theories/ESubstitution.v @@ -5,7 +5,7 @@ From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config. From MetaCoq.PCUIC Require Import PCUICAst PCUICLiftSubst PCUICTyping PCUICGlobalEnv PCUICWeakeningConv PCUICWeakeningTyp PCUICSubstitution - PCUICWeakeningEnv PCUICWeakeningEnvTyp PCUICOnFreeVars PCUICElimination. + PCUICWeakeningEnv PCUICWeakeningEnvTyp PCUICOnFreeVars PCUICElimination PCUICFirstorder. From MetaCoq.Erasure Require Import EGlobalEnv Extract Prelim. Local Set Keyed Unification. @@ -80,9 +80,11 @@ Proof. all: try now (econstructor; eauto). all: try now (econstructor; eapply Is_type_extends; eauto; tc). - econstructor. - red. red in H4. unfold PCUICAst.lookup_inductive in H4. + move: H4; apply contraNN. + unfold isPropositional. + unfold PCUICAst.lookup_inductive. unshelve eapply declared_constructor_to_gen in isdecl; eauto. - rewrite (PCUICAst.declared_inductive_lookup_gen isdecl.p1) in H4. + rewrite (PCUICAst.declared_inductive_lookup_gen isdecl.p1). destruct isdecl as [decli declc]. eapply declared_inductive_from_gen in decli. eapply PCUICWeakeningEnv.weakening_env_declared_inductive in decli; tea; eauto; tc. diff --git a/erasure/theories/ErasureCorrectness.v b/erasure/theories/ErasureCorrectness.v index ae4ddaec8..8a3c835ba 100644 --- a/erasure/theories/ErasureCorrectness.v +++ b/erasure/theories/ErasureCorrectness.v @@ -269,7 +269,7 @@ Proof. depelim H2. eapply isErasable_Propositional in X0; eauto. rewrite -eq_npars. - eapply isPropositional_propositional; eauto. + erewrite isPropositional_propositional; eauto. now rewrite X0. apply declared_inductive_from_gen; eauto. invs e. cbn in *. rewrite map_length. @@ -355,7 +355,8 @@ Proof. now rewrite closedn_mkApps /=. } solve_all. eapply (erases_closed _ []); tea. } rewrite -eq_npars. - eapply isPropositional_propositional_cstr; eauto. + erewrite isPropositional_propositional_cstr; eauto. + move/negbTE: H13 => ->. reflexivity. eapply declared_constructor_from_gen; eauto. rewrite -(Forall2_length H3) /= e1 //. rewrite skipn_length -(Forall2_length H3) -e6 /= map_length. @@ -369,9 +370,9 @@ Proof. -- eapply Is_type_app in X1 as []; auto. 2:{ eapply subject_reduction_eval. 2:eassumption. eauto. } assert (ispind : inductive_isprop_and_pars Σ' ind = Some (true, ind_npars mdecl)). - { eapply isPropositional_propositional; eauto. - apply declared_inductive_from_gen; eauto. - eapply isErasable_Propositional; eauto. } + { erewrite isPropositional_propositional; eauto. f_equal. f_equal. + eapply isErasable_Propositional; eauto. + apply declared_inductive_from_gen; eauto. } eapply tConstruct_no_Type in X1; auto. eapply H6 in X1 as []; eauto. @@ -492,9 +493,9 @@ Proof. unshelve eapply declared_projection_to_gen in decli; eauto. destruct (declared_inductive_inj decli.p1 d.p1) as [<- <-]. assert (isprop : inductive_isprop_and_pars Σ' p.(proj_ind) = Some (true, ind_npars mdecl)). - { eapply isPropositional_propositional; eauto. - eapply declared_inductive_from_gen; exact d. all:eauto. apply decli'. - eapply isErasable_Propositional; eauto. } + { erewrite isPropositional_propositional; eauto. f_equal. f_equal. + eapply isErasable_Propositional; eauto. + eapply declared_inductive_from_gen; exact d. all:eauto. apply decli'. } split. eapply Is_type_app in X as []; eauto. 2:{ rewrite -mkApps_app. eapply subject_reduction_eval; eauto. } rewrite -mkApps_app in X. @@ -532,15 +533,16 @@ Proof. invs H2. -- exists x9. split; eauto. constructor. eapply Ee.eval_proj; eauto. rewrite -eqpars. - eapply isPropositional_propositional_cstr; eauto. + erewrite isPropositional_propositional_cstr; eauto. + move/negbTE: H9 => ->. reflexivity. eapply declared_constructor_from_gen. apply d0. cbn. eapply decli'. cbn. rewrite -lenx5 //. move: eargs. unfold PCUICWcbvEval.cstr_arity; cbn. now rewrite -eqpars. -- exists EAst.tBox. assert (isprop : inductive_isprop_and_pars Σ' p.(proj_ind) = Some (true, ind_npars mdecl)). - { eapply isPropositional_propositional; eauto. - eapply declared_inductive_from_gen. apply d. apply decli'. - eapply (isErasable_Propositional (args:=[])); eauto. } + { erewrite isPropositional_propositional; eauto. do 2 f_equal. + eapply (isErasable_Propositional (args:=[])); eauto. + eapply declared_inductive_from_gen. apply d. apply decli'. } split. eapply Is_type_app in X as []; eauto. 2:{ eapply subject_reduction_eval; [|eauto]; eauto. } diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 528be6234..75ca06fa8 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -1775,273 +1775,8 @@ Proof. destruct KernameSet.mem; cbn. f_equal. f_equal. 1-2:f_equal. all:now eapply IHdecls. Qed. -Definition global_erased_with_deps (Σ : global_env) (Σ' : EAst.global_declarations) kn := - (exists cst, declared_constant Σ kn cst /\ - exists cst' : EAst.constant_body, - EGlobalEnv.declared_constant Σ' kn cst' /\ - erases_constant_body (Σ, cst_universes cst) cst cst' /\ - (forall body : EAst.term, - EAst.cst_body cst' = Some body -> erases_deps Σ Σ' body)) \/ - (exists mib mib', declared_minductive Σ kn mib /\ - EGlobalEnv.declared_minductive Σ' kn mib' /\ - erases_mutual_inductive_body mib mib'). - -Definition includes_deps (Σ : global_env) (Σ' : EAst.global_declarations) deps := - forall kn, - KernameSet.In kn deps -> - global_erased_with_deps Σ Σ' kn. - -Lemma includes_deps_union (Σ : global_env) (Σ' : EAst.global_declarations) deps deps' : - includes_deps Σ Σ' (KernameSet.union deps deps') -> - includes_deps Σ Σ' deps /\ includes_deps Σ Σ' deps'. -Proof. - intros. - split; intros kn knin; eapply H; rewrite KernameSet.union_spec; eauto. -Qed. - -Lemma includes_deps_fold {A} (Σ : global_env) (Σ' : EAst.global_declarations) brs deps (f : A -> EAst.term) : - includes_deps Σ Σ' - (fold_left - (fun (acc : KernameSet.t) (x : A) => - KernameSet.union (term_global_deps (f x)) acc) brs - deps) -> - includes_deps Σ Σ' deps /\ - (forall x, In x brs -> - includes_deps Σ Σ' (term_global_deps (f x))). -Proof. - intros incl; split. - intros kn knin; apply incl. - rewrite knset_in_fold_left. left; auto. - intros br brin. intros kn knin. apply incl. - rewrite knset_in_fold_left. right. - now exists br. -Qed. - -Definition declared_kn Σ kn := - ∥ ∑ decl, lookup_env Σ kn = Some decl ∥. - -Lemma term_global_deps_spec {cf} {Σ : global_env_ext} {Γ t et T} : - wf Σ.1 -> - Σ ;;; Γ |- t : T -> - Σ;;; Γ |- t ⇝ℇ et -> - forall x, KernameSet.In x (term_global_deps et) -> declared_kn Σ x. -Proof. - intros wf wt er. - induction er in T, wt |- * using erases_forall_list_ind; - cbn in *; try solve [constructor]; intros kn' hin; - repeat match goal with - | [ H : KernameSet.In _ KernameSet.empty |- _ ] => - now apply KernameSet.empty_spec in hin - | [ H : KernameSet.In _ (KernameSet.union _ _) |- _ ] => - apply KernameSet.union_spec in hin as [?|?] - end. - - apply inversion_Lambda in wt as (? & ? & ? & ? & ?); eauto. - - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. - - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. - - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. - - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. - - apply inversion_Const in wt as (? & ? & ? & ? & ?); eauto. - eapply KernameSet.singleton_spec in hin; subst. - unshelve eapply declared_constant_to_gen in d; eauto. - red in d. red. sq. rewrite d. intuition eauto. - - apply inversion_Construct in wt as (? & ? & ? & ? & ? & ? & ?); eauto. - destruct kn. - eapply KernameSet.singleton_spec in hin; subst. - destruct d as [[H' _] _]. - unshelve eapply declared_minductive_to_gen in H'; eauto. - red in H'. simpl in *. - red. sq. rewrite H'. intuition eauto. - - apply inversion_Case in wt as (? & ? & ? & ? & [] & ?); eauto. - destruct ci as [kn i']; simpl in *. - eapply KernameSet.singleton_spec in H1; subst. - destruct x1 as [d _]. - unshelve eapply declared_minductive_to_gen in d; eauto. - red in d. - simpl in *. eexists; intuition eauto. - - apply inversion_Case in wt as (? & ? & ? & ? & [] & ?); eauto. - eapply knset_in_fold_left in H1. - destruct H1. eapply IHer; eauto. - destruct H1 as [br [inbr inkn]]. - eapply Forall2_All2 in H0. - assert (All (fun br => ∑ T, Σ ;;; Γ ,,, inst_case_branch_context p br |- br.(bbody) : T) brs). - eapply All2i_All_right. eapply brs_ty. - simpl. intuition auto. eexists ; eauto. - now rewrite -(PCUICCasesContexts.inst_case_branch_context_eq a). - eapply All2_All_mix_left in H0; eauto. simpl in H0. - eapply All2_In_right in H0; eauto. - destruct H0. - destruct X1 as [br' [[T' HT] ?]]. - eauto. - - - eapply KernameSet.singleton_spec in H0; subst. - apply inversion_Proj in wt as (?&?&?&?&?&?&?&?&?&?); eauto. - unshelve eapply declared_projection_to_gen in d; eauto. - destruct d as [[[d _] _] _]. - red in d. eexists; eauto. - - - apply inversion_Proj in wt as (?&?&?&?&?&?&?&?&?); eauto. - - - apply inversion_Fix in wt as (?&?&?&?&?&?&?); eauto. - eapply knset_in_fold_left in hin as [|]. - now eapply KernameSet.empty_spec in H0. - destruct H0 as [? [ina indeps]]. - 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]]]. - eapply Hdef; eauto. - - - apply inversion_CoFix in wt as (?&?&?&?&?&?&?); eauto. - eapply knset_in_fold_left in hin as [|]. - now eapply KernameSet.empty_spec in H0. - destruct H0 as [? [ina indeps]]. - 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]]]. - eapply Hdef; eauto. -Qed. - -Global Remove Hints erases_deps_eval : core. - -Lemma erase_global_erases_deps {Σ} {Σ' : EAst.global_declarations} {Γ t et T} : - wf_ext Σ -> - Σ;;; Γ |- t : T -> - Σ;;; Γ |- t ⇝ℇ et -> - includes_deps Σ Σ' (term_global_deps et) -> - erases_deps Σ Σ' et. -Proof. - intros wf wt er. - induction er in er, t, T, wf, wt |- * using erases_forall_list_ind; - cbn in *; try solve [constructor]; intros Σer; - repeat match goal with - | [ H : includes_deps _ _ (KernameSet.union _ _ ) |- _ ] => - apply includes_deps_union in H as [? ?] - end. - - now apply inversion_Evar in wt. - - constructor. - now apply inversion_Lambda in wt as (? & ? & ? & ? & ?); eauto. - - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. - constructor; eauto. - - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. - now constructor; eauto. - - apply inversion_Const in wt as (? & ? & ? & ? & ?); eauto. - specialize (Σer kn). - forward Σer. rewrite KernameSet.singleton_spec //. - pose proof (wf' := wf.1). - unshelve eapply declared_constant_to_gen in d; eauto. - destruct Σer as [[c' [declc' (? & ? & ? & ?)]]|]. - unshelve eapply declared_constant_to_gen in declc'; eauto. - pose proof (declared_constant_inj _ _ d declc'). subst x. - econstructor; eauto. eapply declared_constant_from_gen; eauto. - destruct H as [mib [mib' [declm declm']]]. - unshelve eapply declared_minductive_to_gen in declm; eauto. - red in declm, d. unfold declared_minductive_gen in declm. - rewrite d in declm. noconf declm. - - apply inversion_Construct in wt as (? & ? & ? & ? & ? & ?); eauto. - red in Σer. destruct kn. - setoid_rewrite KernameSetFact.singleton_iff in Σer. - pose (wf' := wf.1). - destruct (Σer _ eq_refl) as [ (? & ? & ? & ? & ? & ?) | (? & ? & ? & ? & ?) ]. - + destruct d as [[]]. - unshelve eapply declared_constant_to_gen in H0; eauto. - unshelve eapply declared_minductive_to_gen in H4; eauto. - red in H4, H. cbn in *. unfold PCUICEnvironment.fst_ctx in *. congruence. - + pose proof H2 as H2'. destruct d. cbn in *. - destruct H3. cbn in *. red in H3. - unshelve eapply declared_minductive_to_gen in H0, H3; eauto. - red in H0, H3. rewrite H0 in H3. invs H3. - destruct H2. - red in H1. - eapply Forall2_nth_error_Some_l in H2 as (? & ? & ?); eauto. pose proof H6 as H6'. destruct H6 as (? & ? & ? & ? & ?); eauto. - eapply Forall2_nth_error_Some_l in H6 as ([] & ? & ? & ?); subst; eauto. - econstructor. - eapply declared_constructor_from_gen; eauto. repeat eapply conj; eauto. - repeat eapply conj; cbn; eauto. eauto. eauto. - - apply inversion_Case in wt as (? & ? & ? & ? & [] & ?); eauto. - destruct ci as [[kn i'] ? ?]; simpl in *. - apply includes_deps_fold in H2 as [? ?]. - pose proof (wf' := wf.1). - specialize (H1 kn). forward H1. - now rewrite KernameSet.singleton_spec. red in H1. destruct H1. - elimtype False. destruct H1 as [cst [declc _]]. - { destruct x1 as [d _]. - unshelve eapply declared_minductive_to_gen in d; eauto. - unshelve eapply declared_constant_to_gen in declc; eauto. - red in d, declc. - unfold declared_constant_gen in declc. rewrite d in declc. noconf declc. } - destruct H1 as [mib [mib' [declm [declm' em]]]]. - pose proof em as em'. destruct em'. - destruct x1 as [x1 hnth]. - unshelve eapply declared_minductive_to_gen in x1, declm; eauto. - red in x1, declm. unfold declared_minductive_gen in declm. rewrite x1 in declm. noconf declm. - eapply Forall2_nth_error_left in H1; eauto. destruct H1 as [? [? ?]]. - eapply erases_deps_tCase; eauto. - apply declared_inductive_from_gen; auto. - split; eauto. split; eauto. - destruct H1. - eapply In_Forall in H3. - eapply All_Forall. eapply Forall_All in H3. - eapply Forall2_All2 in H0. - eapply All2_All_mix_right in H0; eauto. - assert (All (fun br => ∑ T, Σ ;;; Γ ,,, inst_case_branch_context p br |- br.(bbody) : T) brs). - eapply All2i_All_right. eapply brs_ty. - simpl. intuition auto. eexists ; eauto. - now rewrite -(PCUICCasesContexts.inst_case_branch_context_eq a). - ELiftSubst.solve_all. destruct a0 as [T' HT]. eauto. - - - apply inversion_Proj in wt as (?&?&?&?&?&?&?&?&?&?); eauto. - destruct (proj1 d). - pose proof (wf' := wf.1). - specialize (H0 (inductive_mind p.(proj_ind))). forward H0. - now rewrite KernameSet.singleton_spec. red in H0. destruct H0. - elimtype False. destruct H0 as [cst [declc _]]. - { - unshelve eapply declared_constant_to_gen in declc; eauto. - red in declc. destruct d as [[[d _] _] _]. - unshelve eapply declared_minductive_to_gen in d; eauto. - red in d. - unfold declared_constant_gen in declc. rewrite d in declc. noconf declc. } - destruct H0 as [mib [mib' [declm [declm' em]]]]. - assert (mib = x0). - { destruct d as [[[]]]. - unshelve eapply declared_minductive_to_gen in declm, H0; eauto. - red in H0, declm. unfold declared_minductive_gen in declm. rewrite H0 in declm. now noconf declm. } - subst x0. - pose proof em as em'. destruct em'. - eapply Forall2_nth_error_left in H0 as (x' & ? & ?); eauto. - 2:{ apply d. } - simpl in *. - destruct (ind_ctors x1) => //. noconf H3. - set (H1' := H5). destruct H1' as []. - set (d' := d). destruct d' as [[[]]]. - eapply Forall2_All2 in H3. eapply All2_nth_error_Some in H3 as [? [? []]]; tea. - destruct H6 as [Hprojs _]. - eapply Forall2_All2 in Hprojs. eapply All2_nth_error_Some in Hprojs as [? []]; tea. - 2:eapply d. - econstructor; tea. all:eauto. - split => //. 2:split; eauto. - split; eauto. split; eauto. - rewrite -H4. symmetry; apply d. - - - constructor. - apply inversion_Fix in wt as (?&?&?&?&?&?&?); eauto. - eapply All_Forall. eapply includes_deps_fold in Σer as [_ Σer]. - eapply In_Forall in Σer. - eapply Forall_All in Σer. - eapply Forall2_All2 in H. - ELiftSubst.solve_all. - - constructor. - apply inversion_CoFix in wt as (?&?&?&?&?&?&?); eauto. - eapply All_Forall. eapply includes_deps_fold in Σer as [_ Σer]. - eapply In_Forall in Σer. - eapply Forall_All in Σer. - eapply Forall2_All2 in H. - ELiftSubst.solve_all. Unshelve. -Qed. - (* TODO: Figure out if this (and [erases]) should use [strictly_extends_decls] or [extends] -Jason Gross *) -Lemma erases_weakeninv_env {Σ Σ' : global_env_ext} {Γ t t' T} : +Lemma erases_weakening_env {Σ Σ' : global_env_ext} {Γ t t' T} : wf Σ -> wf Σ' -> strictly_extends_decls Σ Σ' -> Σ ;;; Γ |- t : T -> erases Σ Γ t t' -> erases (Σ'.1, Σ.2) Γ t t'. @@ -2050,181 +1785,6 @@ Proof. eapply (env_prop_typing ESubstitution.erases_extends); tea. Qed. -Lemma erases_deps_weaken kn d (Σ : global_env) (Σ' : EAst.global_declarations) t : - wf (add_global_decl Σ (kn, d)) -> - erases_deps Σ Σ' t -> - erases_deps (add_global_decl Σ (kn, d)) Σ' t. -Proof. - intros wfΣ er. - assert (wfΣ' : wf Σ) - by now eapply strictly_extends_decls_wf; tea; split => //; eexists [_]. - induction er using erases_deps_forall_ind; try solve [now constructor]. - - inv wfΣ. inv X. - assert (wf Σ) by (inversion H4;econstructor; eauto). - pose proof (H_ := H). unshelve eapply declared_constant_to_gen in H; eauto. - assert (kn <> kn0). - { intros <-. - eapply lookup_env_Some_fresh in H. destruct X1. contradiction. } - eapply erases_deps_tConst with cb cb'; eauto. - eapply declared_constant_from_gen. - red. rewrite /declared_constant_gen /lookup_env lookup_env_cons_fresh //. - red. - red in H1. - destruct (cst_body cb) eqn:cbe; - destruct (E.cst_body cb') eqn:cbe'; auto. - specialize (H3 _ eq_refl). - eapply on_declared_constant in H_; auto. - red in H_. rewrite cbe in H_. simpl in H_. - eapply (erases_weakeninv_env (Σ := (Σ, cst_universes cb)) - (Σ' := (add_global_decl Σ (kn, d), cst_universes cb))); eauto. - simpl. econstructor; eauto. econstructor; eauto. - split => //; eexists [(kn, d)]; intuition eauto. - - econstructor; eauto. eapply weakening_env_declared_constructor; eauto; tc. - eapply extends_decls_extends, strictly_extends_decls_extends_decls. econstructor; try reflexivity. eexists [(_, _)]; reflexivity. - - econstructor; eauto. - eapply declared_inductive_from_gen. - inv wfΣ. inv X. - assert (wf Σ) by (inversion H5;econstructor; eauto). - unshelve eapply declared_inductive_to_gen in H; eauto. - red. destruct H. split; eauto. - red in H. red. - rewrite -H. simpl. unfold lookup_env; simpl. destruct (eqb_spec (inductive_mind p.1) kn); try congruence. - eapply lookup_env_Some_fresh in H. destruct X1. subst kn; contradiction. - - econstructor; eauto. - inv wfΣ. inv X. - assert (wf Σ) by (inversion H3;econstructor; eauto). - eapply declared_projection_from_gen. - unshelve eapply declared_projection_to_gen in H; eauto. - red. destruct H. split; eauto. - destruct H; split; eauto. - destruct H; split; eauto. - red in H |- *. - rewrite -H. simpl. unfold lookup_env; simpl; destruct (eqb_spec (inductive_mind p.(proj_ind)) kn); try congruence. - eapply lookup_env_Some_fresh in H. subst kn. destruct X1. contradiction. -Qed. - -Lemma lookup_env_ext {Σ kn kn' d d'} : - wf (add_global_decl Σ (kn', d')) -> - lookup_env Σ kn = Some d -> - kn <> kn'. -Proof. - intros wf hl. - eapply lookup_env_Some_fresh in hl. - inv wf. inv X. - destruct (eqb_spec kn kn'); subst; destruct X1; congruence. -Qed. - -Lemma lookup_env_cons_disc {Σ kn kn' d} : - kn <> kn' -> - lookup_env (add_global_decl Σ (kn', d)) kn = lookup_env Σ kn. -Proof. - intros Hk. simpl; unfold lookup_env; simpl. - destruct (eqb_spec kn kn'); congruence. -Qed. - -Lemma elookup_env_cons_disc {Σ kn kn' d} : - kn <> kn' -> - EGlobalEnv.lookup_env ((kn', d) :: Σ) kn = EGlobalEnv.lookup_env Σ kn. -Proof. - intros Hk. simpl. - destruct (eqb_spec kn kn'); congruence. -Qed. - -Lemma global_erases_with_deps_cons kn kn' d d' Σ Σ' : - wf (add_global_decl Σ (kn', d)) -> - global_erased_with_deps Σ Σ' kn -> - global_erased_with_deps (add_global_decl Σ (kn', d)) ((kn', d') :: Σ') kn. -Proof. - intros wf [[cst [declc [cst' [declc' [ebody IH]]]]]|]. - red. inv wf. inv X. - assert (wfΣ : PCUICTyping.wf Σ). - { eapply strictly_extends_decls_wf; split; tea ; eauto. econstructor; eauto. - now eexists [_]. - } - left. - exists cst. split. - eapply declared_constant_from_gen. - unshelve eapply declared_constant_to_gen in declc; eauto. - red in declc |- *. unfold declared_constant_gen, lookup_env in *. - rewrite lookup_env_cons_fresh //. - { eapply lookup_env_Some_fresh in declc. destruct X1. - intros <-; contradiction. } - exists cst'. - pose proof (declc_ := declc). - unshelve eapply declared_constant_to_gen in declc; eauto. - unfold EGlobalEnv.declared_constant. rewrite EGlobalEnv.elookup_env_cons_fresh //. - { eapply lookup_env_Some_fresh in declc. destruct X1. - intros <-; contradiction. } - red in ebody. unfold erases_constant_body. - destruct (cst_body cst) eqn:bod; destruct (E.cst_body cst') eqn:bod' => //. - intuition auto. - eapply (erases_weakeninv_env (Σ := (_, cst_universes cst)) (Σ' := (add_global_decl Σ (kn', d), cst_universes cst))); eauto. - constructor; eauto. constructor; eauto. - split => //. exists [(kn', d)]; intuition eauto. - eapply on_declared_constant in declc_; auto. - red in declc_. rewrite bod in declc_. eapply declc_. - noconf H0. - eapply erases_deps_cons; eauto. - constructor; auto. - - right. - pose proof (wf_ := wf). inv wf. inv X. - assert (wf Σ) by (inversion H0;econstructor; eauto). - destruct H as [mib [mib' [? [? ?]]]]. - unshelve eapply declared_minductive_to_gen in H; eauto. - exists mib, mib'. intuition eauto. - * eapply declared_minductive_from_gen. - red. red in H. pose proof (lookup_env_ext wf_ H). unfold declared_minductive_gen. - now rewrite lookup_env_cons_disc. - * red. pose proof (lookup_env_ext wf_ H). - now rewrite elookup_env_cons_disc. -Qed. - -Lemma global_erases_with_deps_weaken kn kn' d Σ Σ' : - wf (add_global_decl Σ (kn', d)) -> - global_erased_with_deps Σ Σ' kn -> - global_erased_with_deps (add_global_decl Σ (kn', d)) Σ' kn. -Proof. - intros wf [[cst [declc [cst' [declc' [ebody IH]]]]]|]. - red. inv wf. inv X. left. - assert (wfΣ : PCUICTyping.wf Σ). - { eapply strictly_extends_decls_wf; split; tea ; eauto. econstructor; eauto. - now eexists [_]. - } - exists cst. split. - eapply declared_constant_from_gen. - unshelve eapply declared_constant_to_gen in declc; eauto. - red in declc |- *. - unfold declared_constant_gen, lookup_env in *. - rewrite lookup_env_cons_fresh //. - { eapply lookup_env_Some_fresh in declc. - intros <-. destruct X1. contradiction. } - exists cst'. - unfold EGlobalEnv.declared_constant. - red in ebody. unfold erases_constant_body. - destruct (cst_body cst) eqn:bod; destruct (E.cst_body cst') eqn:bod' => //. - intuition auto. - eapply (erases_weakeninv_env (Σ := (Σ, cst_universes cst)) (Σ' := (add_global_decl Σ (kn', d), cst_universes cst))). 5:eauto. - split; auto. constructor; eauto. constructor; eauto. - split; auto; exists [(kn', d)]; intuition eauto. - eapply on_declared_constant in declc; auto. - red in declc. rewrite bod in declc. eapply declc. - noconf H0. - apply erases_deps_weaken; auto. - constructor; eauto. constructor; eauto. - - right. destruct H as [mib [mib' [Hm [? ?]]]]. - exists mib, mib'; intuition auto. pose proof (wf_ := wf). - inv wf. inv X. - assert (wf Σ) by (inversion H;econstructor; eauto). - eapply declared_minductive_from_gen. - - unshelve eapply declared_minductive_to_gen in Hm; eauto. - red. unfold declared_minductive_gen, lookup_env in *. - rewrite lookup_env_cons_fresh //. - now epose proof (lookup_env_ext wf_ Hm). -Qed. - (* TODO: Figure out if this (and [erases]) should use [strictly_extends_decls] or [extends] -Jason Gross *) Lemma erase_constant_body_correct X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} cb (onc : forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> ∥ on_constant_decl (lift_typing typing) Σ cb ∥) : @@ -2235,7 +1795,7 @@ Proof. red. destruct cb as [name [bod|] univs]; simpl; eauto. intros. set (ecbo := erase_constant_body_obligation_1 X_type X _ _ _ _). clearbody ecbo. cbn in *. specialize_Σ H. sq. - eapply (erases_weakeninv_env (Σ := Σ) (Σ' := (Σ', univs))); simpl; eauto. + eapply (erases_weakening_env (Σ := Σ) (Σ' := (Σ', univs))); simpl; eauto. now eapply erases_erase. Qed. @@ -2277,2322 +1837,4 @@ Proof. cbn in *. intuition auto. induction ind_projs0; constructor; auto. - unfold isPropositionalArity. - destruct destArity as [[? ?]|] eqn:da; auto. -Qed. - -Lemma erase_global_includes X_type (X:X_type.π1) deps decls {normalization_in} prf deps' : - (forall d, KernameSet.In d deps' -> - forall Σ : global_env, abstract_env_rel X Σ -> ∥ ∑ decl, lookup_env Σ d = Some decl ∥) -> - KernameSet.subset deps' deps -> - forall Σ : global_env, abstract_env_rel X Σ -> - let Σ' := erase_global_deps deps X decls (normalization_in:=normalization_in) prf in - includes_deps Σ (fst Σ') deps'. -Proof. - intros ? sub Σ wfΣ; cbn. induction decls in X, H, sub, prf, deps, deps', Σ , wfΣ, normalization_in |- *. - - simpl. intros i hin. elimtype False. - specialize (H i hin) as [[decl Hdecl]]; eauto. - rewrite /lookup_env (prf _ wfΣ) in Hdecl. noconf Hdecl. - - intros i hin. - destruct (abstract_env_wf _ wfΣ) as [wf]. - destruct a as [kn d]. - eapply KernameSet.subset_spec in sub. - edestruct (H i hin) as [[decl Hdecl]]; eauto. unfold lookup_env in Hdecl. - pose proof (eqb_spec i kn). - rewrite (prf _ wfΣ) in Hdecl. move: Hdecl. cbn -[erase_global_deps]. - elim: H0. intros -> [= <-]. - * destruct d as [|]; [left|right]. - { cbn. set (Xpop := abstract_pop_decls X). - unfold erase_constant_decl. - set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _). - epose proof (abstract_env_exists Xpop) as [[Σpop wfpop]]. - epose proof (abstract_env_ext_exists Xmake) as [[Σmake wfmake]]. - exists c. split; auto. - eapply declared_constant_from_gen. - unfold declared_constant_gen, lookup_env; simpl; rewrite (prf _ wfΣ). cbn. - rewrite eq_kername_refl //. - pose proof (sub _ hin) as indeps. - eapply KernameSet.mem_spec in indeps. - unfold EGlobalEnv.declared_constant. - edestruct (H _ hin) as [[decl hd]]; eauto. - eexists; intuition eauto. - rewrite indeps. cbn. - rewrite eq_kername_refl. reflexivity. - cut (strictly_extends_decls Σpop Σ) => [?|]. - eapply (erase_constant_body_correct _ _ _ _ (Σpop , _)); eauto. - rewrite <- (abstract_make_wf_env_ext_correct Xpop (cst_universes c) _ Σpop Σmake wfpop wfmake); eauto. - { now eapply strictly_extends_decls_wf. } - red. simpl. unshelve epose (abstract_pop_decls_correct X decls _ Σ Σpop wfΣ wfpop). - { intros. now eexists. } - split => //. intuition eauto. - exists [(kn, ConstantDecl c)]; intuition eauto. rewrite H0; eauto. - now destruct a. - rewrite indeps. unshelve epose proof (abstract_pop_decls_correct X decls _ Σ Σpop wfΣ wfpop) as [Hpop [Hpop' Hpop'']]. - { intros. now eexists. } - pose (prf' := prf _ wfΣ). - destruct Σ. cbn in *. rewrite Hpop' Hpop'' prf'. rewrite <- Hpop at 1. - eapply (erases_deps_cons Σpop). - rewrite <- Hpop'. apply wf. - rewrite Hpop. rewrite prf' in wf. destruct wf. now rewrite Hpop'' Hpop' in o0. - - pose proof (erase_constant_body_correct' H0). specialize_Σ wfmake. - sq. destruct H1 as [bod [bodty [[Hbod Hebod] Heqdeps]]]. - rewrite (abstract_make_wf_env_ext_correct Xpop (cst_universes c) _ Σpop Σmake wfpop wfmake) in Hbod, Hebod. - eapply (erase_global_erases_deps (Σ := (Σpop, cst_universes c))); simpl; auto. - { constructor; simpl; auto. depelim wf. rewrite Hpop' Hpop'' in o0. - cbn in o0, o. rewrite prf' in o0. rewrite <- Hpop in o0. rewrite Hpop' in o. clear -o o0. - now depelim o0. - depelim wf. rewrite Hpop' in o0. - cbn in o0, o. rewrite prf' in o0. rewrite <- Hpop in o0. clear -o0. depelim o0. - now destruct o. - } - all: eauto. - apply IHdecls; eauto. - { intros. pose proof (abstract_env_wf _ wfpop) as [wf']. - destruct Σpop. cbn in *. clear prf'. subst. - unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wfΣ H2) as [Hpop Hpop']. - { intros. now eexists. } - destruct Σ. cbn in *. subst. - eapply term_global_deps_spec in Hebod; eauto. } - { eapply KernameSet.subset_spec. - intros x hin'. eapply KernameSet.union_spec. right; auto. - now rewrite -Heqdeps. } } - { eexists m, _; intuition eauto. - eapply declared_minductive_from_gen. - simpl. rewrite /declared_minductive /declared_minductive_gen /lookup_env prf; eauto. - simpl. rewrite eq_kername_refl. reflexivity. - specialize (sub _ hin). - eapply KernameSet.mem_spec in sub. - simpl. rewrite sub. - red. cbn. rewrite eq_kername_refl. - reflexivity. - assert (declared_minductive Σ kn m). - { eapply declared_minductive_from_gen. red. unfold declared_minductive_gen, lookup_env. rewrite prf; eauto. cbn. now rewrite eqb_refl. } - eapply on_declared_minductive in H0; tea. - now eapply erases_mutual. } - - * intros ikn Hi. - destruct d as [|]. - ++ simpl. destruct (KernameSet.mem kn deps) eqn:eqkn. - unfold erase_constant_decl. - set (Xpop := abstract_pop_decls X). - set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - pose proof (abstract_env_wf _ wfpop) as [wfΣp]. - unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wfΣ wfpop) as [Hpop [Hpop' Hpop'']]. - { intros. now eexists. } - pose proof (prf _ wfΣ). destruct Σ. cbn in *. subst. - eapply global_erases_with_deps_cons; eauto. - eapply (IHdecls Xpop _ _ _ (KernameSet.singleton i)); auto. - 3:{ eapply KernameSet.singleton_spec => //. } - intros. - eapply KernameSet.singleton_spec in H0. - pose proof (abstract_env_irr _ H1 wfpop). subst. - sq; exists decl; eauto. - eapply KernameSet.subset_spec. intros ? ?. - eapply KernameSet.union_spec. left. - eapply KernameSet.singleton_spec in H0; subst. - now eapply sub. - - cbn. set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - pose proof (abstract_env_wf _ wfpop) as [wfΣp]. - unshelve epose proof (abstract_pop_decls_correct X decls _ Σ Σp wfΣ wfpop) as [Hpop [Hpop' Hpop'']]. - { intros. now eexists. } - pose proof (prf _ wfΣ). destruct Σ. cbn in *. subst. - eapply global_erases_with_deps_weaken. eauto. - eapply IHdecls => //. - 3:now eapply KernameSet.singleton_spec. - intros d ind%KernameSet.singleton_spec. - intros. pose proof (abstract_env_irr _ H0 wfpop). subst. - sq; eexists; eauto. - eapply KernameSet.subset_spec. - intros ? hin'. eapply sub. eapply KernameSet.singleton_spec in hin'. now subst. - - ++ simpl. set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - pose proof (abstract_env_wf _ wfpop) as [wfΣp]. - unshelve epose proof (abstract_pop_decls_correct X decls _ Σ Σp wfΣ wfpop) as [Hpop [Hpop' Hpop'']]. - { intros. now eexists. } - pose proof (prf _ wfΣ). destruct Σ. cbn in *. subst. - destruct (KernameSet.mem kn deps) eqn:eqkn. - { eapply (global_erases_with_deps_cons _ kn (InductiveDecl _) _ Σp); eauto. - eapply (IHdecls Xpop _ _ _ (KernameSet.singleton i)); auto. - 3:{ eapply KernameSet.singleton_spec => //. } - intros. - eapply KernameSet.singleton_spec in H0. subst. - pose proof (abstract_env_irr _ H1 wfpop). subst. - sq; eexists; eauto. - eapply KernameSet.subset_spec. intros ? ?. - eapply KernameSet.singleton_spec in H0; subst. - now eapply sub. } - - { eapply (global_erases_with_deps_weaken _ kn (InductiveDecl _) Σp). eauto. - eapply (IHdecls Xpop _ _ _ (KernameSet.singleton i)) => //. - 3:now eapply KernameSet.singleton_spec. - intros d ind%KernameSet.singleton_spec. - intros. pose proof (abstract_env_irr _ H0 wfpop). subst. - sq; eexists; eauto. - eapply KernameSet.subset_spec. - intros ? hin'. eapply sub. eapply KernameSet.singleton_spec in hin'. now subst. } -Qed. - -Lemma erase_correct (wfl := Ee.default_wcbv_flags) X_type (X : X_type.π1) - univs wfext t v Σ' t' deps decls {normalization_in} prf - (Xext := abstract_make_wf_env_ext X univs wfext) - {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} - : - forall wt : forall Σ, Σ ∼_ext Xext -> welltyped Σ [] t, - erase X_type Xext [] t wt = t' -> - KernameSet.subset (term_global_deps t') deps -> - erase_global_deps deps X decls (normalization_in:=normalization_in) prf = Σ' -> - (forall Σ : global_env, abstract_env_rel X Σ -> Σ |-p t ⇓ v) -> - forall Σ : global_env_ext, abstract_env_ext_rel Xext Σ -> - exists v', Σ ;;; [] |- v ⇝ℇ v' /\ ∥ fst Σ' ⊢ t' ⇓ v' ∥. -Proof. - intros wt. - intros HΣ' Hsub Ht' ev Σext wfΣex. - pose proof (erases_erase (X := Xext) wt). - rewrite HΣ' in H. - destruct (wt _ wfΣex) as [T wt']. - pose proof (abstract_env_ext_wf _ wfΣex) as [wfΣ]. - specialize (H _ wfΣex). - unshelve epose proof (erase_global_erases_deps (Σ' := fst Σ') wfΣ wt' H _); cycle 2. - rewrite <- Ht'. - eapply erase_global_includes; eauto. - intros. eapply term_global_deps_spec in H; eauto. - now rewrite (abstract_make_wf_env_ext_correct X univs wfext Σ _ H1 wfΣex) in H. - - epose proof (abstract_env_exists X) as [[Σ wfΣX]]. - now rewrite (abstract_make_wf_env_ext_correct X univs wfext Σ _ wfΣX wfΣex). - epose proof (abstract_env_exists X) as [[Σ wfΣX]]. - eapply erases_correct; tea. - - eexists; eauto. - - rewrite (abstract_make_wf_env_ext_correct X univs wfext _ _ wfΣX wfΣex); eauto. -Qed. - -Lemma global_env_ind (P : global_env -> Type) - (Pnil : forall univs retro, P {| universes := univs; declarations := []; retroknowledge := retro |}) - (Pcons : forall (Σ : global_env) d, P Σ -> P (add_global_decl Σ d)) - (Σ : global_env) : P Σ. -Proof. - destruct Σ as [univs Σ]. - induction Σ; cbn. apply Pnil. - now apply (Pcons {| universes := univs; declarations := Σ |} a). -Qed. - -Lemma on_global_env_ind (P : forall Σ : global_env, wf Σ -> Type) - (Pnil : forall univs retro (onu : on_global_univs univs), P {| universes := univs; declarations := []; retroknowledge := retro |} - (onu, globenv_nil _ _ _ _)) - (Pcons : forall (Σ : global_env) kn d (wf : wf Σ) - (Hfresh : fresh_global kn Σ.(declarations)) - (udecl := PCUICLookup.universes_decl_of_decl d) - (onud : on_udecl Σ.(universes) udecl) - (pd : on_global_decl cumulSpec0 (lift_typing typing) - ({| universes := Σ.(universes); declarations := Σ.(declarations); retroknowledge := Σ.(retroknowledge) |}, udecl) kn d), - P Σ wf -> P (add_global_decl Σ (kn, d)) - (fst wf, globenv_decl _ _ Σ.(universes) Σ.(retroknowledge) Σ.(declarations) kn d (snd wf) - {| kn_fresh := Hfresh ; on_udecl_udecl := onud ; on_global_decl_d := pd |})) - (Σ : global_env) (wfΣ : wf Σ) : P Σ wfΣ. -Proof. - destruct Σ as [univs Σ]. destruct wfΣ; cbn in *. - induction o0. apply Pnil. destruct o1. - apply (Pcons {| universes := univs; declarations := Σ |} kn d (o, o0)). - exact IHo0. -Qed. - -Ltac inv_eta := - lazymatch goal with - | [ H : PCUICEtaExpand.expanded _ _ |- _ ] => depelim H - end. - -Lemma leq_term_propositional_sorted_l {Σ Γ v v' u u'} : - wf_ext Σ -> - PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> - Σ;;; Γ |- v : tSort u -> - Σ;;; Γ |- v' : tSort u' -> is_propositional u -> - leq_universe (global_ext_constraints Σ) u' u. -Proof. - intros wfΣ leq hv hv' isp. - eapply orb_true_iff in isp as [isp | isp]. - - eapply leq_term_prop_sorted_l; eauto. - - eapply leq_term_sprop_sorted_l; eauto. -Qed. - -Fixpoint map_erase (X_type:abstract_env_impl) (X : X_type.π2.π1) {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ - (ts : list term) - (H2 : forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> - Forall (welltyped Σ Γ) ts) {struct ts}: list E.term. -destruct ts as [ | t ts]. -- exact []. -- eapply cons. refine (erase X_type X Γ t _). - 2: eapply (map_erase X_type X _ Γ ts). - all: intros; specialize_Σ H; now inversion H2; subst. -Defined. - -Lemma map_erase_irrel X_type X {normalization_in1 normalization_in2 : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ t H1 H2 : - map_erase X_type X Γ t H1 (normalization_in:=normalization_in1) = map_erase X_type X Γ t H2 (normalization_in:=normalization_in2). -Proof. - epose proof (abstract_env_ext_exists X) as [[Σ wfΣX]]. - induction t. - - reflexivity. - - cbn. f_equal. - + eapply erase_irrel. - + eapply IHt. -Qed. - -Arguments map_erase _ _ _ _ _, _ _ {_} _ _ _, _ _ {_} _ _ {_}. - -Lemma erase_mkApps {X_type X} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ t args H2 Ht Hargs : - (forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> wf_local Σ Γ) -> - (forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> ~ ∥ isErasable Σ Γ (mkApps t args) ∥) -> - erase X_type X Γ (mkApps t args) H2 = - E.mkApps (erase X_type X Γ t Ht) (map_erase X_type X Γ args Hargs). -Proof. - epose proof (abstract_env_ext_exists X) as [[Σ wfΣX]]. - pose proof (abstract_env_ext_wf X wfΣX) as [wf]. - intros Hwflocal Herasable. induction args in t, H2, Ht, Hargs, Herasable |- *. - - cbn. eapply erase_irrel. - - cbn [mkApps]. - rewrite IHargs; clear IHargs. - all: intros; specialize_Σ H; try pose proof (abstract_env_ext_wf _ H) as [wfH]. - 1: inversion Hargs; eauto. - 2: eauto. - 2:{ destruct H2. cbn in X0. eapply inversion_mkApps in X0 as (? & ? & ?). - econstructor. eauto. } - etransitivity. simp erase. reflexivity. unfold erase_clause_1. - unfold inspect. unfold erase_clause_1_clause_2. - Unshelve. - elim: is_erasableP. - + intros. exfalso. - eapply Herasable; eauto. specialize_Σ wfΣX. destruct p. - cbn. destruct H2. eapply Is_type_app; eauto. - + cbn [map_erase]. - epose proof (fst (erase_irrel _ _)). cbn. - intros he. f_equal => //. f_equal. - eapply erase_irrel. eapply erase_irrel. - eapply map_erase_irrel. -Qed. - -Lemma map_erase_length X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ t H1 : length (map_erase X_type X Γ t H1) = length t. -Proof. - induction t; cbn; f_equal; eauto. -Qed. - -Local Hint Constructors expanded : expanded. - -Local Arguments erase_global_deps _ _ _ _ _ : clear implicits. - -(*Lemma lookup_env_erase X_type X deps decls normalization_in prf kn d : - KernameSet.In kn deps -> - forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = Some (InductiveDecl d) -> - EGlobalEnv.lookup_env (erase_global_deps X_type deps X decls normalization_in prf) kn = Some (EAst.InductiveDecl (erase_mutual_inductive_body d)). -Proof. - intros hin Σ wfΣ. pose proof (prf _ wfΣ). unfold lookup_env. rewrite H. clear H. - rewrite /lookup_env. - induction decls in X, Σ , wfΣ ,prf, deps, hin, normalization_in |- *. - - move=> /= //. - - destruct a as [kn' d']. - cbn -[erase_global_deps]. - case: (eqb_spec kn kn'); intros e'; subst. - intros [= ->]. - unfold erase_global_deps. - eapply KernameSet.mem_spec in hin. rewrite hin /=. - now rewrite eq_kername_refl. - intros hl. destruct d'. simpl. - set (Xpop := abstract_pop_decls X). - set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - destruct KernameSet.mem. cbn. - rewrite (negbTE (proj2 (neqb _ _) e')). - eapply IHdecls => //; eauto. eapply KernameSet.union_spec. left => //. - eapply IHdecls => //; eauto. - simpl. - set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - destruct KernameSet.mem. cbn. - rewrite (negbTE (proj2 (neqb _ _) e')). - eapply IHdecls => //; eauto. - eapply IHdecls => //; eauto. -Qed. -*) - -Lemma lookup_env_erase_decl_None X_type X deps decls normalization_in prf kn : - forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = None -> - EGlobalEnv.lookup_env (erase_global_deps X_type deps X decls normalization_in prf).1 kn = None. -Proof. - intros Σ wfΣ. - pose proof (prf _ wfΣ). - unfold lookup_env. rewrite H. clear H. - rewrite /lookup_env. - induction decls in X, Σ , wfΣ ,prf, deps, normalization_in |- *. - - move=> /= //. - - destruct a as [kn' d']. - cbn -[erase_global_deps]. - case: (eqb_spec kn kn'); intros e'; subst. - intros [= ->]. - cbn. destruct d'. cbn. - case_eq (KernameSet.mem kn' deps); move => hin'; cbn. - + destruct (eqb_spec kn kn') => //. - set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - now eapply IHdecls. - + intros. set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - now eapply IHdecls. - + intros hin. - case_eq (KernameSet.mem kn' deps); move => hin'; cbn. - destruct (eqb_spec kn kn') => //. - set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - now eapply IHdecls. - intros. set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - now eapply IHdecls. -Qed. - -Lemma lookup_env_erase_global_diff X_type X deps decls kn kn' d' normalization_in prf : - kn <> kn' -> - exists deps' nin' prf', - KernameSet.Subset deps deps' /\ - EGlobalEnv.lookup_env (erase_global_deps X_type deps X ((kn', d') :: decls) normalization_in prf).1 kn = - EGlobalEnv.lookup_env (erase_global_deps X_type deps' (abstract_pop_decls X) decls nin' prf').1 kn /\ - (KernameSet.In kn (erase_global_deps X_type deps X ((kn', d') :: decls) normalization_in prf).2 -> - KernameSet.In kn (erase_global_deps X_type deps' (abstract_pop_decls X) decls nin' prf').2). -Proof. - intros hd. - simpl. destruct d'. - + destruct KernameSet.mem. cbn. - destruct (eqb_spec kn kn') => //. - do 3 eexists. split; [|split;[reflexivity|]]. knset. - auto. - do 3 eexists. split; [|split;[reflexivity|]]. knset. - auto. - + destruct KernameSet.mem. cbn. - destruct (eqb_spec kn kn') => //. - do 3 eexists. split; [|split;[reflexivity|]]. knset. auto. - do 3 eexists. split; [|split;[reflexivity|]]. knset. auto. -Qed. - -Lemma lookup_env_In_map_fst Σ kn decl : lookup_global Σ kn = Some decl -> In kn (map fst Σ). -Proof. - induction Σ; cbn => //. - case: eqb_spec. - + intros -> [= <-]. now left. - + intros _ hl. eauto. -Qed. - -Lemma erase_constant_decl_deps {X_type : abstract_env_impl} (X : X_type.π1) (cb : constant_body) normalization_in oncb : - forall Σ : global_env, abstract_env_rel X Σ -> - forall kn, KernameSet.In kn (erase_constant_decl X cb normalization_in oncb).2 -> In kn (map fst (declarations Σ)). -Proof. - intros Σ wfΣ kn. - unfold erase_constant_decl. - set (ec := erase_constant_body _ _ _ _). - destruct cb. destruct cst_body0. - 2:{ subst ec; cbn. knset. } - intros hc. - epose proof (erase_constant_body_correct' (@eq_refl _ (EAst.cst_body ec.1))). - subst ec. - set (ec := erase_constant_body _ _ _ _) in *. - set (X' := abstract_make_wf_env_ext _ _ _) in *. - pose proof (abstract_env_ext_exists X') as [[Σ' H']]. - pose proof (abstract_env_ext_wf _ H') as [wfΣ']. - specialize (H _ H') as [[t0 [T []]]]. - rewrite -e in hc. - destruct p as [hty her]. - epose proof (term_global_deps_spec wfΣ' hty her _ hc). - epose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ wfΣ H'). cbn in H0. subst Σ'. - red in H. destruct H as [[decl hl]]. - now eapply lookup_env_In_map_fst in hl. -Qed. - -Lemma in_erase_global_deps_acc X_type X deps decls kn normalization_in prf : - KernameSet.In kn (erase_global_deps X_type deps X decls normalization_in prf).2 -> - KernameSet.In kn deps \/ In kn (map fst decls). -Proof. - induction decls in X, prf, deps, normalization_in |- *. - cbn; auto. destruct a as [kn' []]. - + cbn. set (ec := erase_constant_decl _ _ _ _). - set (eg := erase_global_deps _ _ _ _ _ _). - set (eg' := erase_global_deps _ _ _ _ _ _). - case_eq (KernameSet.mem kn' deps). - * cbn; intros. specialize (IHdecls _ _ _ _ H0). - destruct IHdecls as [Hu|Hm]; auto. - eapply KernameSet.union_spec in Hu. destruct Hu; auto. - unfold ec in H1; cbn in H1. - set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - eapply erase_constant_decl_deps in H1; tea. - pose proof (abstract_env_exists X) as [[ΣX HX]]. - unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ HX wfpop). - intros. rewrite prf => //. now eexists. - now destruct H2 as [<- _]. - * intros hdeps hin. - eapply IHdecls in hin. intuition auto. - + cbn; set (eg := erase_global_deps _ _ _ _ _ _); - set (eg' := erase_global_deps _ _ _ _ _ _). - case_eq (KernameSet.mem kn' deps). - * cbn; intros. specialize (IHdecls _ _ _ _ H0). - destruct IHdecls as [Hu|Hm]; auto. - * intros hdeps hin. - eapply IHdecls in hin. intuition auto. -Qed. - -Lemma wf_fresh_globals {cf : checker_flags} (Σ : global_env) : wf Σ -> EnvMap.EnvMap.fresh_globals Σ.(declarations). -Proof. - destruct Σ as [univs Σ]; cbn. - move=> [] onu; cbn. induction 1; try destruct o; constructor; auto; constructor; eauto. -Qed. - -Lemma lookup_env_erase_decl X_type X deps decls normalization_in prf kn decl : - forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = Some decl -> - let er := erase_global_deps X_type deps X decls normalization_in prf in - KernameSet.In kn er.2 -> - match decl with - | ConstantDecl cb => - exists (X' : X_type.π1) nin pf, - (EGlobalEnv.lookup_env er.1 kn = - Some (EAst.ConstantDecl (fst (erase_constant_decl X' cb nin pf)))) /\ - (forall Σ Σ', Σ ∼ X -> Σ' ∼ X' -> ∥ strictly_extends_decls Σ' Σ ∥) - | InductiveDecl mib => - EGlobalEnv.lookup_env er.1 kn = - Some (EAst.InductiveDecl (erase_mutual_inductive_body mib)) - end. -Proof. - intros Σ wfΣ. pose proof (prf _ wfΣ). - unfold lookup_env. rewrite H. clear H. - rewrite /lookup_env. - induction decls in X, Σ , wfΣ ,prf, deps, decl, normalization_in |- *. - - move=> /= //. - - destruct a as [kn' d']. - cbn -[erase_global_deps]. - case: (eqb_spec kn kn'); intros e'; subst. - intros [= ->]. - destruct decl. - + cbn; set (ec := erase_constant_decl _ _ _ _); - set (eg := erase_global_deps _ _ _ _ _ _); - set (eg' := erase_global_deps _ _ _ _ _ _). - case_eq (KernameSet.mem kn' deps). - move/KernameSet.mem_spec => hin'; cbn. - ++ intros hin. cbn; rewrite eqb_refl; - do 3 eexists. split; [reflexivity|]. - intros ? ? hx hpop. - unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ hx hpop). intros. rewrite prf => //. now eexists. - destruct H as [? []]. unfold strictly_extends_decls. rewrite H. sq; split => //. clear ec eg eg' hin. specialize (prf _ hx). rewrite prf. - now eexists [_]. - ++ intros kn'deps kn'eg'. - eapply in_erase_global_deps_acc in kn'eg'. - destruct kn'eg'. eapply KernameSet.mem_spec in H. congruence. - elimtype False. clear -prf wfΣ H. - specialize (prf _ wfΣ). - pose proof (abstract_env_wf _ wfΣ) as []. eapply wf_fresh_globals in X0. - rewrite prf in X0. depelim X0. - apply EnvMap.EnvMap.fresh_global_iff_not_In in H0. contradiction. - + cbn; - set (eg := erase_global_deps _ _ _ _ _ _); - set (eg' := erase_global_deps _ _ _ _ _ _). - case_eq (KernameSet.mem kn' deps). - move/KernameSet.mem_spec => hin'; cbn. - ++ intros hin. cbn; rewrite eqb_refl; - do 3 eexists. - ++ intros kn'deps kn'eg'. - eapply in_erase_global_deps_acc in kn'eg'. - destruct kn'eg'. eapply KernameSet.mem_spec in H. congruence. - elimtype False. clear -prf wfΣ H. - specialize (prf _ wfΣ). - pose proof (abstract_env_wf _ wfΣ) as []. eapply wf_fresh_globals in X0. - rewrite prf in X0. depelim X0. - apply EnvMap.EnvMap.fresh_global_iff_not_In in H0. contradiction. - + intros hl hin. - epose proof (lookup_env_erase_global_diff X_type X deps decls kn kn' d' _ _ e'). - destruct H as [deps' [nin' [prf' [sub [hl' hin']]]]]. - eapply hin' in hin. - erewrite hl'. - set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - specialize (IHdecls Xpop deps'). specialize (IHdecls nin' prf' decl _ wfpop hl hin). - destruct decl; intuition auto. - destruct IHdecls as [X' [nin [pf []]]]. exists X', nin, pf. split => //. - intros. - unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ H1 wfpop). intros. rewrite prf => //. now eexists. - specialize (H0 _ _ wfpop H2). destruct H0; sq. move: X0; rewrite /strictly_extends_decls. destruct H3 as [? []]. - rewrite H4 H0 H3. rewrite -(prf' _ wfpop). pose proof (prf _ H1). rewrite -H0 in H5. rewrite H5. - intros []; split => //. destruct s as [Σ'' eq]. eexists ((kn', d') :: Σ''). now rewrite eq. -Qed. - -Lemma fresh_global_app kn Σ Σ' : fresh_global kn (Σ ++ Σ') -> - fresh_global kn Σ /\ fresh_global kn Σ'. -Proof. - induction Σ; cbn; intuition auto. - - constructor. - - depelim H. constructor => //. - now eapply Forall_app in H0. - - depelim H. - now eapply Forall_app in H0. -Qed. - -Lemma lookup_global_app_wf Σ Σ' kn : EnvMap.EnvMap.fresh_globals (Σ ++ Σ') -> - forall decl, lookup_global Σ' kn = Some decl -> lookup_global (Σ ++ Σ') kn = Some decl. -Proof. - induction Σ in kn |- *; cbn; auto. - intros fr; depelim fr. - intros decl hd; cbn. - destruct (eqb_spec kn kn0). - eapply PCUICWeakeningEnv.lookup_global_Some_fresh in hd. - subst. now eapply fresh_global_app in H as [H H']. - now eapply IHΣ. -Qed. - -Lemma strictly_extends_decls_lookup {Σ Σ'} : - wf Σ' -> - strictly_extends_decls Σ Σ' -> - forall kn d, lookup_env Σ kn = Some d -> - lookup_env Σ' kn = Some d. -Proof. - intros. - destruct X0 as [? []]. - rewrite /lookup_env in H *. rewrite e0. - erewrite lookup_global_app_wf. reflexivity. 2:eauto. - eapply wf_fresh_globals in X. now rewrite -e0. -Qed. - -Lemma in_deps_in_erase_global_deps X_type X deps decls normalization_in prf kn : - KernameSet.In kn deps -> - KernameSet.In kn (erase_global_deps X_type deps X decls normalization_in prf).2. -Proof. - induction decls in X, deps, normalization_in, prf |- *. - - intros. now cbn. - - intros. cbn. - pose proof (abstract_env_exists X) as [[Σ H']]. - set (Xpop := abstract_pop_decls X). - epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. - unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ H' wfpop). intros. rewrite prf => //. now eexists. - destruct H0 as [Hd [Hu Hr]]. - destruct a. destruct g. - + set (eg := erase_global_deps _ _ _ _ _ _); - set (eg' := erase_global_deps _ _ _ _ _ _). - case_eq (KernameSet.mem k deps). - * cbn. intros hm. - destruct (eqb_spec k kn). - ++ subst k. eapply IHdecls. knset. - ++ eapply IHdecls. - knset. - * intros hnin. eapply IHdecls; tea. - + set (eg := erase_global_deps _ _ _ _ _ _); - set (eg' := erase_global_deps _ _ _ _ _ _). - case_eq (KernameSet.mem k deps). - * cbn. intros hm. - destruct (eqb_spec k kn). - ++ subst k. eapply KernameSet.mem_spec in hm. - now eapply IHdecls. - ++ eapply IHdecls. tea. - * intros hnin. eapply IHdecls; tea. -Qed. - -Fixpoint filter_deps deps decls := - match decls with - | [] => [] - | (kn, decl) :: decls => - if KernameSet.mem kn deps then - match decl with - | EAst.ConstantDecl cst => - (kn, decl) :: filter_deps (KernameSet.union deps (constant_decls_deps cst)) decls - | _ => (kn, decl) :: filter_deps deps decls - end - else filter_deps deps decls - end. - -Lemma erase_global_deps_erase_global (X_type : abstract_env_impl) (X:X_type.π1) deps decls {normalization_in} prf : - (@erase_global_deps X_type deps X decls normalization_in prf).1 = - filter_deps deps (@erase_global X_type X decls normalization_in prf). -Proof. - symmetry. - generalize normalization_in at 1. - generalize prf at 1. - intros prf' normalization_in'. - induction decls in normalization_in, normalization_in', prf, prf', X, deps |- *. - - cbn. auto. - - destruct a as [kn []]. - cbn. - + destruct (KernameSet.mem kn deps). - destruct c as [? [] ? ?]. - set (ec := erase_constant_decl _ _ _ _). - set (ec' := erase_constant_decl _ _ _ _). - assert (ec = ec') as <-. - { unfold ec, ec'. eapply erase_constant_decl_irrel. red. cbn. - intros. pose proof (abstract_env_irr _ H H0). subst. red. split; auto. congruence. } - cbn; f_equal. set (d := KernameSet.union _ _). eapply IHdecls. - cbn. f_equal. eapply IHdecls. - eapply IHdecls. - + cbn; destruct (KernameSet.mem kn deps). - cbn. f_equal. eapply IHdecls. - eapply IHdecls. -Qed. - - -Lemma filter_deps_filter {efl : EWellformed.EEnvFlags} deps Σ : - EWellformed.wf_glob Σ -> - (* (forall kn, KernameSet.In kn deps -> In kn (map fst Σ)) -> *) - filter_deps deps Σ = filter (flip KernameSet.mem (global_deps Σ deps) ∘ fst) Σ. -Proof. - induction Σ in deps |- *. - - now cbn. - - destruct a as [kn [[[]]|]]. - + cbn. unfold flip. - destruct (knset_mem_spec kn deps). - * case: (knset_mem_spec kn _). - ** intros hin' hwf. f_equal. depelim hwf. now eapply IHΣ. - ** intros nhin. - elim: nhin. - rewrite global_deps_union. - rewrite KernameSet.union_spec. left. - now eapply in_global_deps. - * intros hwf. - case: (knset_mem_spec kn _); intros hin. - ** elimtype False. - depelim hwf. - eapply in_global_deps_fresh in hin => //. - ** depelim hwf. rewrite -IHΣ => //. - + intros wf; depelim wf. - cbn. unfold flip. - * case: (knset_mem_spec kn _). - ** intros hin'. - case: (knset_mem_spec kn _) => hin''. - ++ f_equal. now eapply IHΣ. - ++ elim: hin''. - rewrite !global_deps_union. - rewrite KernameSet.union_spec. left. - now eapply in_global_deps. - ** intros hwf. - case: (knset_mem_spec kn _); intros hin. - *** elimtype False. - eapply in_global_deps_fresh in hin => //. - *** rewrite -IHΣ => //. - + intros wf; depelim wf. - cbn. unfold flip. - * case: (knset_mem_spec kn _). - ** intros hin'. - case: (knset_mem_spec kn _) => hin''. - ++ f_equal. rewrite IHΣ //. eapply filter_ext. - intros ?. unfold flip. rewrite global_deps_union /=. - rewrite global_deps_empty. - now rewrite KernameSetFact.union_b KernameSetFact.empty_b orb_false_r. - ++ elim: hin''. - rewrite !global_deps_union. - rewrite KernameSet.union_spec. left. - now eapply in_global_deps. - ** intros hwf. - case: (knset_mem_spec kn _); intros hin. - *** elimtype False. - eapply in_global_deps_fresh in hin => //. - *** rewrite -IHΣ => //. -Qed. - -Lemma erase_global_declared_constructor X_type X ind c mind idecl cdecl deps decls normalization_in prf: - forall Σ : global_env, abstract_env_rel X Σ -> declared_constructor Σ (ind, c) mind idecl cdecl -> - KernameSet.In ind.(inductive_mind) deps -> - EGlobalEnv.declared_constructor (erase_global_deps X_type deps X decls normalization_in prf).1 (ind, c) - (erase_mutual_inductive_body mind) (erase_one_inductive_body idecl) - (EAst.mkConstructor cdecl.(cstr_name) cdecl.(cstr_arity)). -Proof. - intros Σ wfΣ [[]] Hin. pose (abstract_env_wf _ wfΣ). sq. - cbn in *. split. split. - - unshelve eapply declared_minductive_to_gen in H; eauto. - red in H |- *. eapply (lookup_env_erase_decl _ _ _ _ _ _ _ (InductiveDecl mind)); eauto. - eapply in_deps_in_erase_global_deps; tea. - - - cbn. now eapply map_nth_error. - - cbn. erewrite map_nth_error; eauto. -Qed. - -Import EGlobalEnv. -Lemma erase_global_deps_fresh X_type kn deps X decls normalization_in heq : - let Σ' := erase_global_deps X_type deps X decls normalization_in heq in - PCUICAst.fresh_global kn decls -> - fresh_global kn Σ'.1. -Proof. - cbn. - revert deps X normalization_in heq. - induction decls; [cbn; auto|]. - - intros. red. constructor. - - destruct a as [kn' d]. intros. depelim H. - cbn in H, H0. - destruct d as []; simpl; destruct KernameSet.mem. - + cbn [EGlobalEnv.closed_env forallb]. cbn. - constructor => //. eapply IHdecls => //. - + eapply IHdecls => //. - + constructor; auto. - eapply IHdecls => //. - + eapply IHdecls => //. -Qed. - -Lemma erase_global_fresh X_type kn X decls normalization_in heq : - let Σ' := @erase_global X_type X decls normalization_in heq in - PCUICAst.fresh_global kn decls -> - fresh_global kn Σ'. -Proof. - cbn. - revert X normalization_in heq. - induction decls; [cbn; auto|]. - - intros. red. constructor. - - destruct a as [kn' d]. intros. depelim H. - cbn in H, H0. cbn. - destruct d as []; simpl. - + cbn [EGlobalEnv.closed_env forallb]. cbn. - constructor => //. eapply IHdecls => //. - + constructor; auto. - eapply IHdecls => //. -Qed. - -From MetaCoq.Erasure Require Import EEtaExpandedFix. - -Lemma erase_brs_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ p ts wt : - erase_brs X_type X Γ p ts wt = - map_All (fun br wt => (erase_context (bcontext br), erase X_type X _ (bbody br) wt)) ts wt. -Proof. - funelim (map_All _ ts wt); cbn; auto. - f_equal => //. f_equal => //. apply erase_irrel. - rewrite -H. eapply erase_irrel. -Qed. - -Lemma erase_fix_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ ts wt : - erase_fix X_type X Γ ts wt = map_All (fun d wt => - let dbody' := erase X_type X _ (dbody d) (fun Σ abs => proj2 (wt Σ abs)) in - let dbody' := if isBox dbody' then - match d.(dbody) with - | tLambda na _ _ => E.tLambda (binder_name na) E.tBox - | _ => dbody' - end else dbody' - in - {| E.dname := d.(dname).(binder_name); E.rarg := d.(rarg); E.dbody := dbody' |}) ts wt. -Proof. - funelim (map_All _ ts wt); cbn; auto. - f_equal => //. f_equal => //. - rewrite (fst (erase_irrel _ _) _ _ _ _ (fun (Σ : global_env_ext) (abs : abstract_env_ext_rel X Σ) => - (map_All_obligation_1 x xs h Σ abs).p2)). - destruct erase => //. - rewrite -H. eapply erase_irrel. -Qed. - -Lemma erase_cofix_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ ts wt : - erase_cofix X_type X Γ ts wt = map_All (fun d wt => - let dbody' := erase X_type X _ (dbody d) wt in - {| E.dname := d.(dname).(binder_name); E.rarg := d.(rarg); E.dbody := dbody' |}) ts wt. -Proof. - funelim (map_All _ ts wt); cbn; auto. - f_equal => //. f_equal => //. - apply erase_irrel. - rewrite -H. eapply erase_irrel. -Qed. - -Lemma isConstruct_erase X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ t wt : - ~ (PCUICEtaExpand.isConstruct t || PCUICEtaExpand.isFix t || PCUICEtaExpand.isRel t) -> - ~ (isConstruct (erase X_type X Γ t wt) || isFix (erase X_type X Γ t wt) || isRel (erase X_type X Γ t wt)). -Proof. - apply (erase_elim X_type X - (fun Γ t wt e => ~ (PCUICEtaExpand.isConstruct t || PCUICEtaExpand.isFix t || PCUICEtaExpand.isRel t) -> ~ (isConstruct e || isFix e || isRel e)) - (fun Γ l awt e => True) - (fun Γ p l awt e => True) - (fun Γ l awt e => True) - (fun Γ l awt e => True)); intros => //. bang. -Qed. - -Lemma apply_expanded Σ Γ Γ' t t' : - expanded Σ Γ t -> Γ = Γ' -> t = t' -> expanded Σ Γ' t'. -Proof. intros; now subst. Qed. - -Lemma isLambda_inv t : isLambda t -> exists na ty bod, t = tLambda na ty bod. -Proof. destruct t => //; eauto. Qed. - -Lemma erases_deps_erase (cf := config.extraction_checker_flags) - {X_type X} univs - (wfΣ : forall Σ, (abstract_env_rel X Σ) -> ∥ wf_ext (Σ, univs) ∥) decls {normalization_in} prf - (X' := abstract_make_wf_env_ext X univs wfΣ) - {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} - Γ t - (wt : forall Σ : global_env_ext, abstract_env_ext_rel X' Σ -> welltyped Σ Γ t) : - let et := erase X_type X' Γ t wt in - let deps := EAstUtils.term_global_deps et in - forall Σ, (abstract_env_rel X Σ) -> - erases_deps Σ (erase_global_deps X_type deps X decls normalization_in prf).1 et. -Proof. - intros et deps Σ wf. - pose proof (abstract_env_ext_exists X') as [[Σ' wfΣ']]. - pose proof (wt _ wfΣ'). destruct H. pose proof (wfΣ _ wf) as [w]. - rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ') in X0. - eapply (erase_global_erases_deps w); tea. - eapply (erases_erase (X := X') (Γ := Γ)). - now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). - eapply erase_global_includes. - intros. rewrite (abstract_env_irr _ H0 wf). - eapply term_global_deps_spec in H; eauto. - assumption. - eapply (erases_erase (X := X') (Γ := Γ)). - now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). - eapply KernameSet.subset_spec. reflexivity. - now cbn. -Qed. - -Lemma erases_deps_erase_weaken (cf := config.extraction_checker_flags) - {X_type X} univs - (wfΣ : forall Σ, (abstract_env_rel X Σ) -> ∥ wf_ext (Σ, univs) ∥) decls {normalization_in} prf - (X' := abstract_make_wf_env_ext X univs wfΣ) - {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} - Γ t - (wt : forall Σ : global_env_ext, abstract_env_ext_rel X' Σ -> welltyped Σ Γ t) - deps : - let et := erase X_type X' Γ t wt in - let tdeps := EAstUtils.term_global_deps et in - forall Σ, (abstract_env_rel X Σ) -> - erases_deps Σ (erase_global_deps X_type (KernameSet.union deps tdeps) X decls normalization_in prf).1 et. -Proof. - intros et tdeps Σ wf. - pose proof (abstract_env_ext_exists X') as [[Σ' wfΣ']]. - pose proof (wt _ wfΣ'). destruct H. pose proof (wfΣ _ wf) as [w]. - rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ') in X0. - eapply (erase_global_erases_deps w); tea. - eapply (erases_erase (X := X') (Γ := Γ)). - now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). - eapply erase_global_includes. - intros. rewrite (abstract_env_irr _ H0 wf). - eapply term_global_deps_spec in H; eauto. - assumption. - eapply (erases_erase (X := X') (Γ := Γ)). - now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). - eapply KernameSet.subset_spec. intros x hin. - eapply KernameSet.union_spec. now right. - now cbn. -Qed. - -Lemma eta_expand_erase {X_type X} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Σ' {Γ t} - (wt : forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> welltyped Σ Γ t) Γ' : - forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> - PCUICEtaExpand.expanded Σ Γ' t -> - erases_global Σ Σ' -> - expanded Σ' Γ' (erase X_type X Γ t wt). -Proof. - intros Σ wfΣ exp deps. - pose proof (abstract_env_ext_wf _ wfΣ) as [wf]. - eapply expanded_erases. apply wf. - eapply erases_erase; eauto. assumption. - pose proof (wt _ wfΣ). destruct H as [T ht]. - eapply erases_global_erases_deps; tea. - eapply erases_erase; eauto. -Qed. - -Lemma erase_global_closed X_type X deps decls {normalization_in} prf : - let X' := erase_global_deps X_type deps X decls normalization_in prf in - EGlobalEnv.closed_env X'.1. -Proof. - revert X normalization_in prf deps. - induction decls; [cbn; auto|]. - intros X normalization_in prf deps. - destruct a as [kn d]. - destruct d as []; simpl; destruct KernameSet.mem; - unfold erase_constant_decl; - set (Xpop := abstract_pop_decls X); - try (set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _); - assert (normalization_in_Xmake : forall Σ : global_env_ext, wf_ext Σ -> Σ ∼_ext Xmake -> NormalizationIn Σ) - by (subst Xmake; unshelve eapply (normalization_in 0); revgoals; try reflexivity; cbn; lia)). - + cbn [EGlobalEnv.closed_env forallb]. cbn. - rewrite [forallb _ _](IHdecls) // andb_true_r. - rewrite /test_snd /EGlobalEnv.closed_decl /=. - destruct c as [ty [] univs]; cbn. - unfold erase_constant_decl. - set (obl := erase_constant_body_obligation_1 _ _ _ _ _). - unfold erase_constant_body. cbn. clearbody obl. cbn in obl. - 2:auto. - pose proof (abstract_env_ext_exists Xmake) as [[Σmake wfmake]]. - pose proof (abstract_env_ext_wf _ wfmake) as [[?]]. - unshelve epose proof (erases_erase (X := Xmake) (obl eq_refl) _ _) as H; eauto. - eapply erases_closed in H. - 1: edestruct erase_irrel as [-> _]; eassumption. - cbn. destruct (obl eq_refl _ wfmake). clear H. - now eapply PCUICClosedTyp.subject_closed in X0. - + eapply IHdecls => //. - + cbn [EGlobalEnv.closed_env forallb]. - eapply IHdecls => //. - + eapply IHdecls => //. - Unshelve. eauto. -Qed. - -Import EWellformed. - -Section wffix. - Import EAst. - Fixpoint wf_fixpoints (t : term) : bool := - match t with - | tRel i => true - | tEvar ev args => List.forallb (wf_fixpoints) args - | tLambda N M => wf_fixpoints M - | tApp u v => wf_fixpoints u && wf_fixpoints v - | tLetIn na b b' => wf_fixpoints b && wf_fixpoints b' - | tCase ind c brs => - let brs' := forallb (wf_fixpoints ∘ snd) brs in - wf_fixpoints c && brs' - | tProj p c => wf_fixpoints c - | tFix mfix idx => - (idx isLambda d.(dbody) && wf_fixpoints d.(dbody)) mfix - | tCoFix mfix idx => - (idx true - | tConstruct ind c _ => true - | tVar _ => true - | tBox => true - | tPrim _ => true - end. - -End wffix. - -Lemma erases_deps_wellformed (cf := config.extraction_checker_flags) (efl := all_env_flags) {Σ} {Σ'} et : - erases_deps Σ Σ' et -> - forall n, ELiftSubst.closedn n et -> - wf_fixpoints et -> - wellformed Σ' n et. -Proof. - intros ed. - induction ed using erases_deps_forall_ind; intros => //; - try solve [cbn in *; unfold wf_fix in *; rtoProp; intuition eauto; solve_all]. - - cbn. red in H0. rewrite H0 //. - - cbn -[lookup_constructor]. - cbn. now destruct H0 as [[-> ->] ->]. - - cbn in *. move/andP: H5 => [] cld clbrs. - cbn. apply/andP; split. apply/andP; split. - * now destruct H0 as [-> ->]. - * now move/andP: H6. - * move/andP: H6; solve_all. - - cbn -[lookup_projection] in *. apply/andP; split; eauto. - now rewrite (declared_projection_lookup H0). -Qed. - -Lemma erases_wf_fixpoints Σ Γ t t' : Σ;;; Γ |- t ⇝ℇ t' -> - ErasureProperties.wellformed Σ t -> wf_fixpoints t'. -Proof. - induction 1 using erases_forall_list_ind; cbn; auto; try solve [rtoProp; repeat solve_all]. - - move/andP => []. rewrite (All2_length X) => -> /=. unfold test_def in *. - eapply Forall2_All2 in H. - eapply All2_All2_mix in X; tea. solve_all. - destruct b0; eapply erases_isLambda in H1; tea. - - move/andP => []. rewrite (All2_length X) => -> /=. - unfold test_def in *. solve_all. -Qed. - -Lemma erase_wf_fixpoints (efl := all_env_flags) {X_type X} univs wfΣ {Γ t} wt - (X' := abstract_make_wf_env_ext X univs wfΣ) {normalization_in} : - let t' := erase X_type X' (normalization_in:=normalization_in) Γ t wt in - wf_fixpoints t'. -Proof. - cbn. pose proof (abstract_env_ext_exists X') as [[Σ' wf']]. - pose proof (abstract_env_exists X) as [[Σ wf]]. - eapply erases_wf_fixpoints. - eapply erases_erase; eauto. - specialize (wt _ wf'). destruct (wfΣ _ wf). - unshelve eapply welltyped_wellformed in wt; eauto. - now rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wf'). -Qed. - -Lemma erases_global_erase_global (efl := all_env_flags) {X_type X} Σ decls {normalization_in} prf : - Σ ∼ X -> - erases_global Σ (@erase_global X_type X decls normalization_in prf). -Proof. - intros h. pose proof (prf _ h). red. rewrite H. clear H. - induction decls in X, Σ, prf, normalization_in, h |- *. - - cbn. constructor. - - cbn. - set (Xpop := abstract_pop_decls X). - destruct (abstract_env_exists Xpop) as [[Σpop hpop]]. - unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ h hpop). - intros; eexists; now eapply prf. - destruct H as [? []]. - destruct a as [kn []]. - * constructor. cbn. - unfold erase_constant_decl. - set (Σd := {| universes := _ |}). - assert (Σd = Σpop). - { subst Σd. rewrite -H H0 H1. now destruct Σpop. } - subst Σd. - set (ext := abstract_make_wf_env_ext _ _ _). - pose proof (abstract_env_ext_exists ext) as [[Σ' wf']]. - epose proof (abstract_make_wf_env_ext_correct (abstract_pop_decls X) (cst_universes c) _ _ _ hpop wf'). - subst Σ'. - cbn. pose proof (abstract_env_wf _ hpop) as []. - eapply (erase_constant_body_correct _ _ _ _ (Σpop, cst_universes c)) => //. - now rewrite H2. rewrite H2. split => //. now exists []. - rewrite H0 H1. - now eapply IHdecls. - * pose proof (abstract_env_wf _ h) as [wfΣ]. - destruct wfΣ. rewrite (prf _ h) in o0. - depelim o0. depelim o1. - constructor. - eapply erases_mutual. eapply on_global_decl_d. - rewrite H0 H1. - now eapply IHdecls. -Qed. - -Lemma erase_global_wellformed (efl := all_env_flags) {X_type X} decls {normalization_in} prf univs wfΣ {Γ t} wt - (X' := abstract_make_wf_env_ext X univs wfΣ) - {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} - (t' := erase X_type X' Γ t wt) : - wellformed (@erase_global X_type X decls normalization_in prf) #|Γ| t'. -Proof. - cbn. - epose proof (@erases_erase X_type X' normalization_in' _ _ wt). - pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. - pose proof (abstract_env_ext_exists X') as [[Σ' wf']]. - pose proof (abstract_env_ext_wf _ wf') as [?]. - eapply (erases_wellformed (wt _ wf')) in H; eauto; tea. - eapply erases_global_all_deps. apply X0. - epose proof (abstract_make_wf_env_ext_correct _ _ wfΣ _ _ wf wf'). subst Σ'. - now eapply erases_global_erase_global. -Qed. - -Lemma erase_wellformed_deps (efl := all_env_flags) {X_type X} decls {normalization_in} prf univs wfΣ {Γ t} wt - (X' := abstract_make_wf_env_ext X univs wfΣ) - {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} - (t' := erase X_type X' Γ t wt) : - wellformed (erase_global_deps X_type (term_global_deps t') X decls normalization_in prf).1 #|Γ| t'. -Proof. - cbn. - epose proof (@erases_deps_erase X_type X univs wfΣ decls normalization_in prf _ Γ t wt). - pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. - pose proof (abstract_env_ext_exists X') as [[Σ' wf']]. - pose proof (abstract_env_ext_wf _ wf') as [[?]]. - epose proof (erases_deps_wellformed _ H #|Γ|). - apply H0. - eapply (erases_closed _ Γ). - destruct (wt _ wf'). - cbn in X. destruct (wfΣ _ wf). - eapply PCUICClosedTyp.subject_closed in X0. eassumption. - eapply erases_erase; eauto. - eapply erase_wf_fixpoints. Unshelve. eauto. -Qed. - -Lemma erase_wellformed_weaken (efl := all_env_flags) {X_type X} decls {normalization_in} prf univs wfΣ {Γ t} wt -(X' := abstract_make_wf_env_ext X univs wfΣ) -{normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} -deps: -let t' := erase X_type X' Γ t wt in - wellformed (erase_global_deps X_type (KernameSet.union deps (term_global_deps t')) X decls normalization_in prf).1 #|Γ| t'. -Proof. - set (t' := erase _ _ _ _ _). cbn. - epose proof (@erases_deps_erase_weaken _ X univs wfΣ decls _ prf _ Γ t wt deps). - pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. - pose proof (abstract_env_ext_exists X') as [[Σ' wf']]. - pose proof (abstract_env_ext_wf _ wf') as [[?]]. - epose proof (erases_deps_wellformed _ H #|Γ|). - apply H0. - eapply (erases_closed _ Γ). - - destruct (wt _ wf'). - destruct (wfΣ _ wf). - eapply PCUICClosedTyp.subject_closed in X0. eassumption. - - eapply erases_erase; eauto. - - eapply erase_wf_fixpoints. - Unshelve. eauto. -Qed. - -Lemma erase_constant_body_correct'' {X_type X} {cb} {decls normalization_in prf} {univs wfΣ} -(X' := abstract_make_wf_env_ext X univs wfΣ) -{normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} -{onc : forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> ∥ on_constant_decl (lift_typing typing) Σ' cb ∥} {body} deps : - EAst.cst_body (fst (erase_constant_body X_type X' cb onc)) = Some body -> - forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> - ∥ ∑ t T, (Σ' ;;; [] |- t : T) * (Σ' ;;; [] |- t ⇝ℇ body) * - (term_global_deps body = snd (erase_constant_body X_type X' cb onc)) * - wellformed (efl:=all_env_flags) (erase_global_deps X_type (KernameSet.union deps (term_global_deps body)) X decls normalization_in prf).1 0 body ∥. -Proof. - intros ? Σ' wfΣ'. pose proof (abstract_env_exists X) as [[Σ wf]]. - destruct cb as [name [bod|] udecl]; simpl. - simpl in H. noconf H. - set (obl :=(erase_constant_body_obligation_1 X_type X' - {| - cst_type := name; - cst_body := Some bod; - cst_universes := udecl |} onc bod eq_refl)). clearbody obl. - destruct (obl _ wfΣ'). sq. - have er : (Σ, univs);;; [] |- bod ⇝ℇ erase X_type X' [] bod obl. - { eapply (erases_erase (X:=X')). - now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). - } - exists bod, A; intuition auto. - now rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). - eapply erase_wellformed_weaken. - now simpl in H. -Qed. - -Lemma erase_global_cst_decl_wf_glob X_type X deps decls normalization_in heq : - forall cb wfΣ hcb, - let X' := abstract_make_wf_env_ext X (cst_universes cb) wfΣ in - forall normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ, - let ecb := erase_constant_body X_type X' cb hcb in - let Σ' := erase_global_deps X_type (KernameSet.union deps ecb.2) X decls normalization_in heq in - (@wf_global_decl all_env_flags Σ'.1 (EAst.ConstantDecl ecb.1) : Prop). -Proof. - intros cb wfΣ hcb X' normalization_in' ecb Σ'. - unfold wf_global_decl. cbn. - pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. - pose proof (abstract_env_ext_exists X') as [[Σmake wfmake]]. - destruct (wfΣ _ wf), (hcb _ wfmake). red in X1. - destruct EAst.cst_body eqn:hb => /= //. - eapply (erase_constant_body_correct'') in hb; eauto. - destruct hb as [[t0 [T [[] ?]]]]. rewrite e in i. exact i. -Qed. - -Lemma erase_global_ind_decl_wf_glob {X_type X} {deps decls normalization_in kn m} heq : - (forall Σ : global_env, abstract_env_rel X Σ -> on_inductive cumulSpec0 (lift_typing typing) (Σ, ind_universes m) kn m) -> - let m' := erase_mutual_inductive_body m in - let Σ' := erase_global_deps X_type deps X decls normalization_in heq in - @wf_global_decl all_env_flags Σ'.1 (EAst.InductiveDecl m'). -Proof. - set (m' := erase_mutual_inductive_body m). - set (Σ' := erase_global_deps _ _ _ _ _ _). simpl. - intros oni. - pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. - pose proof (abstract_env_wf _ wf) as [wfΣ]. - assert (erases_mutual_inductive_body m (erase_mutual_inductive_body m)). - { eapply (erases_mutual (mdecl:=kn)); tea. } - eapply (erases_mutual_inductive_body_wf (univs := Σ.(universes)) (retro := Σ.(retroknowledge)) (Σ := decls) (kn := kn) (Σ' := Σ'.1)) in H; tea. - rewrite -(heq _ wf). now destruct Σ. -Qed. - -Lemma erase_global_deps_wf_glob X_type X deps decls normalization_in heq : - let Σ' := erase_global_deps X_type deps X decls normalization_in heq in - @wf_glob all_env_flags Σ'.1. -Proof. - cbn. - pose proof (abstract_env_exists X) as [[Σ wf]]. - pose proof (abstract_env_wf _ wf) as [wfΣ]. - revert deps X Σ wf wfΣ normalization_in heq. - induction decls; [cbn; auto|]. - { intros. constructor. } - intros. destruct a as [kn []]; simpl; destruct KernameSet.mem; set (Xpop := abstract_pop_decls X); - try set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _); - epose proof (abstract_env_exists Xpop) as [[Σpop wfpop]]; - pose proof (abstract_env_wf _ wfpop) as [wfΣpop]. - + constructor. eapply IHdecls => //; eauto. eapply erase_global_cst_decl_wf_glob; auto. - eapply erase_global_deps_fresh; auto. - destruct wfΣ. destruct wfΣpop. - rewrite (heq _ wf) in o0. depelim o0. now destruct o3. - + cbn. eapply IHdecls; eauto. - + constructor. eapply IHdecls; eauto. - destruct wfΣ as [[onu ond]]. - rewrite (heq _ wf) in o. depelim o. destruct o0. - eapply (erase_global_ind_decl_wf_glob (kn:=kn)); tea. - intros. rewrite (abstract_env_irr _ H wfpop). - unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wf wfpop) as [? ?]. - {intros; now eexists. } - destruct Σpop, Σ; cbn in *. now subst. - eapply erase_global_deps_fresh. - destruct wfΣ as [[onu ond]]. - rewrite (heq _ wf) in o. depelim o. now destruct o0. - + eapply IHdecls; eauto. -Qed. - -Lemma erase_constant_body_correct_glob {X_type X} {cb} {decls normalization_in prf} {univs wfΣ} - (X' := abstract_make_wf_env_ext X univs wfΣ) - {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} - {onc : forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> ∥ on_constant_decl (lift_typing typing) Σ' cb ∥} {body} : - EAst.cst_body (fst (erase_constant_body X_type X' cb onc)) = Some body -> - forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> - ∥ ∑ t T, (Σ' ;;; [] |- t : T) * (Σ' ;;; [] |- t ⇝ℇ body) * - (term_global_deps body = snd (erase_constant_body X_type X' cb onc)) * - wellformed (efl:=all_env_flags) (@erase_global X_type X decls normalization_in prf) 0 body ∥. -Proof. - intros ? Σ' wfΣ'. pose proof (abstract_env_exists X) as [[Σ wf]]. - destruct cb as [name [bod|] udecl]; simpl. - simpl in H. noconf H. - set (obl :=(erase_constant_body_obligation_1 X_type X' - {| - cst_type := name; - cst_body := Some bod; - cst_universes := udecl |} onc bod eq_refl)). clearbody obl. - destruct (obl _ wfΣ'). sq. - have er : (Σ, univs);;; [] |- bod ⇝ℇ erase X_type X' [] bod obl. - { eapply (erases_erase (X:=X')). - now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). - } - exists bod, A; intuition auto. - now rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). - eapply (erase_global_wellformed decls _ _ _ (Γ := [])). - now simpl in H. -Qed. - -Lemma erase_global_cst_wf_glob X_type X decls normalization_in heq : - forall cb wfΣ hcb, - let X' := abstract_make_wf_env_ext X (cst_universes cb) wfΣ in - forall normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ, - let ecb := erase_constant_body X_type X' cb hcb in - let Σ' := @erase_global X_type X decls normalization_in heq in - (@wf_global_decl all_env_flags Σ' (EAst.ConstantDecl ecb.1) : Prop). -Proof. - intros cb wfΣ hcb X' normalization_in' ecb Σ'. - unfold wf_global_decl. cbn. - pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. - pose proof (abstract_env_ext_exists X') as [[Σmake wfmake]]. - destruct (wfΣ _ wf), (hcb _ wfmake). red in X1. - destruct EAst.cst_body eqn:hb => /= //. - eapply (erase_constant_body_correct_glob) in hb; eauto. - destruct hb as [[t0 [T [[] ?]]]]. exact i. -Qed. - -Lemma erase_global_ind_decl_wf_glob' {X_type X} {decls normalization_in kn m} heq : - (forall Σ : global_env, abstract_env_rel X Σ -> on_inductive cumulSpec0 (lift_typing typing) (Σ, ind_universes m) kn m) -> - let m' := erase_mutual_inductive_body m in - let Σ' := @erase_global X_type X decls normalization_in heq in - @wf_global_decl all_env_flags Σ' (EAst.InductiveDecl m'). -Proof. - set (m' := erase_mutual_inductive_body m). - set (Σ' := erase_global _ _ _). simpl. - intros oni. - pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. - pose proof (abstract_env_wf _ wf) as [wfΣ]. - assert (erases_mutual_inductive_body m (erase_mutual_inductive_body m)). - { eapply (erases_mutual (mdecl:=kn)); tea. } - eapply (erases_mutual_inductive_body_wf (univs := Σ.(universes)) (retro := Σ.(retroknowledge)) (Σ := decls) (kn := kn) (Σ' := Σ')) in H; tea. - rewrite -(heq _ wf). now destruct Σ. -Qed. - -Lemma erase_global_wf_glob X_type X decls normalization_in heq : - let Σ' := @erase_global X_type X decls normalization_in heq in - @wf_glob all_env_flags Σ'. -Proof. - cbn. - pose proof (abstract_env_exists X) as [[Σ wf]]. - pose proof (abstract_env_wf _ wf) as [wfΣ]. - revert X Σ wf wfΣ normalization_in heq. - induction decls; [cbn; auto|]. - { intros. constructor. } - intros. destruct a as [kn []]; simpl; set (Xpop := abstract_pop_decls X); - try set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _); - epose proof (abstract_env_exists Xpop) as [[Σpop wfpop]]; - pose proof (abstract_env_wf _ wfpop) as [wfΣpop]. - + constructor. eapply IHdecls => //; eauto. - eapply erase_global_cst_wf_glob; auto. - eapply erase_global_fresh; auto. - destruct wfΣ. destruct wfΣpop. - rewrite (heq _ wf) in o0. depelim o0. now destruct o3. - + constructor. eapply IHdecls; eauto. - destruct wfΣ as [[onu ond]]. - rewrite (heq _ wf) in o. depelim o. destruct o0. - eapply (erase_global_ind_decl_wf_glob' (kn:=kn)); tea. - intros. rewrite (abstract_env_irr _ H wfpop). - unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wf wfpop) as [? ?]. - { intros; now eexists. } - destruct Σpop, Σ; cbn in *. now subst. - eapply erase_global_fresh. - destruct wfΣ as [[onu ond]]. - rewrite (heq _ wf) in o. depelim o. now destruct o0. -Qed. - -Lemma extends_cons d Σ Σ' : extends Σ Σ' -> extends (d :: Σ) (d :: Σ'). -Proof. - intros ext kn decl. - cbn. - case: (eqb_spec kn d.1). - - now intros -> [= <-]. - - intros hin. apply ext. -Qed. - -Lemma extends_fresh Σ Σ' kn d : - fresh_global kn Σ' -> - extends Σ Σ' -> - extends Σ ((kn, d) :: Σ'). -Proof. - intros fr ext. - intros kn' decl hl. - cbn. destruct (eqb_spec kn' kn) => //. subst. - eapply ext in hl. now eapply lookup_env_Some_fresh in hl. - now eapply ext. -Qed. - -Lemma Forall_filter {A} {P} f l : @Forall A P l -> Forall P (filter f l). -Proof. - induction 1; try econstructor; auto. - cbn. case E: (f x) => //. now constructor. -Qed. - -Lemma extends_filter_impl {efl : EEnvFlags} (f f' : kername * EAst.global_decl -> bool) Σ : - (forall x, f x -> f' x) -> - wf_glob Σ -> - extends (filter f Σ) (filter f' Σ). -Proof. - intros hf; induction Σ; cbn; auto. - - red; auto. - - intros wf; depelim wf. - case E: (f _); auto. rewrite (hf _ E). - now eapply extends_cons. - case E': (f' _); auto. - eapply extends_fresh. now eapply Forall_filter. - auto. -Qed. - -Lemma extends_filter {efl : EEnvFlags} (f : kername * EAst.global_decl -> bool) Σ : - wf_glob Σ -> - extends (filter f Σ) Σ. -Proof. - intros wf hf; induction Σ; cbn; auto. - intros decl. depelim wf. cbn in *. - case E: (f _); auto; cbn. - case (eqb_spec hf kn) => //. - - intros hd. now eapply IHΣ. - - intros hl; eapply IHΣ in hl => //. rewrite hl. - case (eqb_spec hf kn) => heq. - subst. now eapply lookup_env_Some_fresh in hl. - reflexivity. -Qed. - -Lemma erase_global_deps_eval {efl : EWcbvEval.WcbvFlags} (X_type : abstract_env_impl) (X X':X_type.π1) ext wfX wfX' {normalization_in} {normalization_in'} decls nin prf Γ t wt v wv : - let Xext := abstract_make_wf_env_ext X ext wfX in - let Xext' := abstract_make_wf_env_ext X' ext wfX' in - let t' := erase X_type Xext (normalization_in := normalization_in) Γ t wt in - let v' := erase X_type Xext' (normalization_in := normalization_in') Γ v wv in - let eg := erase_global_deps X_type (term_global_deps t') X decls nin prf in - let eg' := erase_global_deps X_type (term_global_deps v') X decls nin prf in - eg.1 ⊢ t' ⇓ v' -> EGlobalEnv.extends eg'.1 eg.1. -Proof. - intros. - move: (erase_global_deps_wf_glob X_type X (term_global_deps t') decls nin prf). - subst eg eg'. - rewrite !erase_global_deps_erase_global. - assert (wfer : wf_glob (efl := all_env_flags) (@erase_global X_type X decls nin prf)). - { eapply erase_global_wf_glob. } - rewrite !(filter_deps_filter (efl := all_env_flags)) //. - cbn. - intros wf. - eapply extends_filter_impl; tea. unfold flip in *. - intros [kn d]; cbn. clear d. - rewrite /is_true !KernameSet.mem_spec. revert kn. - eapply EWcbvEval.weakening_eval_env in H. 2:exact wfer. - now eapply eval_global_deps in H. - rewrite !erase_global_deps_erase_global (filter_deps_filter (efl := all_env_flags)) //. - now eapply extends_filter. -Qed. - -Lemma lookup_erase_global (cf := config.extraction_checker_flags) {X_type X} {deps deps' decls normalization_in prf} : - KernameSet.Subset deps deps' -> - EExtends.global_subset (erase_global_deps X_type deps X decls normalization_in prf).1 (erase_global_deps X_type deps' X decls normalization_in prf).1. -Proof. - intros sub. - rewrite !erase_global_deps_erase_global. - assert (wfer : wf_glob (efl := all_env_flags) (@erase_global X_type X decls normalization_in prf)). - { eapply erase_global_wf_glob. } - rewrite !(filter_deps_filter (efl := all_env_flags)) //. - eapply extends_filter_impl => //. 2:exact wfer. - unfold flip. - move=> x /KernameSet.mem_spec hin. - apply/KernameSet.mem_spec. - eapply global_deps_subset; tea. -Qed. - -Lemma expanded_weakening_global X_type X deps deps' decls normalization_in prf Γ t : - KernameSet.Subset deps deps' -> - expanded (erase_global_deps X_type deps X decls normalization_in prf).1 Γ t -> - expanded (erase_global_deps X_type deps' X decls normalization_in prf).1 Γ t. -Proof. - intros hs. - eapply expanded_ind; intros; try solve [econstructor; eauto]. - - eapply expanded_tFix; tea. solve_all. - - eapply expanded_tConstruct_app; tea. - destruct H. split; tea. - destruct d; split => //. - cbn in *. red in H. - eapply lookup_erase_global in H; tea. -Qed. - -Lemma expanded_erase (cf := config.extraction_checker_flags) - {X_type X decls normalization_in prf} univs wfΣ t wtp : - forall Σ : global_env, abstract_env_rel X Σ -> PCUICEtaExpand.expanded Σ [] t -> - let X' := abstract_make_wf_env_ext X univs wfΣ in - forall normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ, - let et := (erase X_type X' [] t wtp) in - let deps := EAstUtils.term_global_deps et in - expanded (erase_global_deps X_type deps X decls normalization_in prf).1 [] et. -Proof. - intros Σ wf hexp X' normalization_in'. - pose proof (abstract_env_wf _ wf) as [wf']. - pose proof (abstract_env_ext_exists X') as [[? wfmake]]. - eapply (expanded_erases (Σ := (Σ, univs))); tea. - eapply (erases_erase (X := X')); eauto. - now erewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ). - cbn. - eapply (erases_deps_erase (X := X) univs wfΣ); eauto. -Qed. - -Lemma expanded_erase_global (cf := config.extraction_checker_flags) - deps {X_type X decls normalization_in prf} : - forall Σ : global_env, abstract_env_rel X Σ -> - PCUICEtaExpand.expanded_global_env Σ -> - expanded_global_env (erase_global_deps X_type deps X decls normalization_in prf).1. -Proof. - intros Σ wf etaΣ. pose proof (prf _ wf). - destruct Σ as [univ decls']. - red. revert wf. red in etaΣ. cbn in *. subst. - revert deps X normalization_in prf. - induction etaΣ; intros deps. intros. constructor. intros. - pose proof (abstract_env_exists (abstract_pop_decls X)) as [[Σpop wfpop]]. - unshelve epose proof (abstract_pop_decls_correct X Σ _ _ _ wf wfpop) as [? [? ?]]. - { now eexists. } - destruct Σpop. cbn in H0, H1, H2. subst. - destruct decl as [kn []]; - destruct (KernameSet.mem kn deps) eqn:eqkn; simpl; rewrite eqkn. - constructor; [eapply IHetaΣ; auto|]. - red. cbn. red. cbn in *. - unfold erase_constant_decl. - set (deps' := KernameSet.union _ _). hidebody deps'. - set (wfext := abstract_make_wf_env_ext _ _ _). hidebody wfext. - destruct H. - destruct c as [cst_na [cst_body|] cst_type cst_rel]. - cbn in *. - eapply expanded_weakening_global. - 2:{ eapply expanded_erase; tea. } - set (et := erase _ _ _ _) in *. - unfold deps'. unfold hidebody. intros x hin. - eapply KernameSet.union_spec. right => //. - now cbn. - eapply IHetaΣ; eauto. - constructor. eapply IHetaΣ; eauto. - red. cbn => //. - eapply IHetaΣ; eauto. -Qed. - -(* Sanity checks: the [erase] function maximally erases terms *) -Lemma erasable_tBox_value (wfl := Ee.default_wcbv_flags) (Σ : global_env_ext) (wfΣ : wf_ext Σ) t T v : - forall wt : Σ ;;; [] |- t : T, - Σ |-p t ⇓ v -> erases Σ [] v E.tBox -> ∥ isErasable Σ [] t ∥. -Proof. - intros. - depind H. - eapply Is_type_eval_inv; eauto. eexists; eauto. -Qed. - -Lemma erase_eval_to_box (wfl := Ee.default_wcbv_flags) - {X_type X} {univs wfext t v Σ' t' deps decls normalization_in prf} - (Xext := abstract_make_wf_env_ext X univs wfext) - {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} - : - forall wt : forall Σ : global_env_ext, abstract_env_ext_rel Xext Σ -> welltyped Σ [] t, - erase X_type Xext [] t wt = t' -> - KernameSet.subset (term_global_deps t') deps -> - erase_global_deps X_type deps X decls normalization_in prf = Σ' -> - forall Σext : global_env_ext, abstract_env_ext_rel Xext Σext -> - (forall Σ : global_env, abstract_env_rel X Σ -> - PCUICWcbvEval.eval Σ t v) -> - @Ee.eval Ee.default_wcbv_flags Σ'.1 t' E.tBox -> ∥ isErasable Σext [] t ∥. -Proof. - intros wt. - intros. - destruct (erase_correct X_type X univs wfext _ _ Σ' _ _ decls prf wt H H0 H1 X0 _ H2) as [ev [eg [eg']]]. - pose proof (Ee.eval_deterministic H3 eg'). subst. - pose proof (abstract_env_exists X) as [[? wf]]. - destruct (wfext _ wf). destruct (wt _ H2) as [T wt']. - pose proof (abstract_env_ext_wf _ H2) as [?]. - eapply erasable_tBox_value; eauto. - pose proof (abstract_make_wf_env_ext_correct X univs wfext _ _ wf H2). subst. - apply X0; eauto. -Qed. - -Lemma erase_eval_to_box_eager (wfl := Ee.default_wcbv_flags) - {X_type X} {univs wfext t v Σ' t' deps decls normalization_in prf} - (Xext := abstract_make_wf_env_ext X univs wfext) - {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} - : - forall wt : forall Σ : global_env_ext, abstract_env_ext_rel Xext Σ -> welltyped Σ [] t, - erase X_type Xext [] t wt = t' -> - KernameSet.subset (term_global_deps t') deps -> - erase_global_deps X_type deps X decls normalization_in prf = Σ' -> - forall Σext : global_env_ext, abstract_env_ext_rel Xext Σext -> - (forall Σ : global_env, abstract_env_rel X Σ -> - PCUICWcbvEval.eval Σ t v) -> - @Ee.eval Ee.default_wcbv_flags Σ'.1 t' E.tBox -> t' = E.tBox. -Proof. - intros wt. - intros. - destruct (erase_eval_to_box wt H H0 H1 _ H2 X0 H3). - subst t'. - destruct (inspect_bool (is_erasableb X_type Xext [] t wt)) eqn:heq. - - simp erase. rewrite heq. - simp erase => //. - - elimtype False. - pose proof (abstract_env_exists X) as [[? wf]]. - destruct (@is_erasableP X_type Xext _ [] t wt) => //. apply n. - intros. sq. now rewrite (abstract_env_ext_irr _ H H2). -Qed. - -From MetaCoq.PCUIC Require Import PCUICFirstorder. - - -Lemma welltyped_mkApps_inv {cf} {Σ : global_env_ext} Γ f args : ∥ wf Σ ∥ -> - welltyped Σ Γ (mkApps f args) -> welltyped Σ Γ f /\ Forall (welltyped Σ Γ) args. -Proof. - intros wf (A & HA). sq. eapply inversion_mkApps in HA as (? & ? & ?). - split. eexists; eauto. - induction t0 in f |- *; econstructor; eauto; econstructor; eauto. -Qed. - -From MetaCoq.PCUIC Require Import PCUICProgress. - -Lemma firstorder_erases_deterministic X_type (X : X_type.π1) - univs wfext {v t' i u args} - (Xext := abstract_make_wf_env_ext X univs wfext) - {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} - : - forall wv : (forall Σ, Σ ∼_ext Xext -> welltyped Σ [] v), - forall Σ, Σ ∼_ext Xext -> - Σ ;;; [] |- v : mkApps (tInd i u) args -> - Σ |-p v ⇓ v -> - @firstorder_ind Σ (firstorder_env Σ) i -> - erases Σ [] v t' -> - t' = erase X_type Xext (normalization_in:=normalization_in) [] v wv. -Proof. - intros wv Σ Hrel Hty Hvalue Hfo Herase. - destruct (firstorder_lookup_inv Hfo) as [mind Hdecl]. - assert (Hext : ∥ wf_ext Σ∥) by now eapply heΣ. - sq. eapply firstorder_value_spec in Hty as Hfov; eauto. - clear - Hrel Hext Hfov Herase. - revert t' wv Herase. - pattern v. - revert v Hfov. - eapply firstorder_value_inds. intros. - rewrite erase_mkApps. - - intros Σ0 HΣ0. pose proof (abstract_env_ext_irr _ HΣ0 Hrel). subst. - eapply PCUICValidity.inversion_mkApps in X0 as (? & XX & Happ). - clear XX. revert Happ. clear. generalize (mkApps (tInd i u) pandi). induction 1. - + econstructor. - + econstructor. econstructor; eauto. eauto. - - intros. eapply erases_mkApps_inv in Herase as [(? & ? & ? & -> & [Herasable] & ? & ? & ->)|(? & ? & -> & ? & ?)]. all:eauto. - + exfalso. eapply isErasable_Propositional in Herasable; eauto. - red in H1, Herasable. unfold PCUICAst.lookup_inductive, lookup_inductive_gen, PCUICAst.lookup_minductive, lookup_minductive_gen, isPropositionalArity in *. - edestruct PCUICEnvironment.lookup_env as [ [] | ], nth_error, destArity as [[] | ]; auto; try congruence. - + inv H2. - * cbn. unfold erase_clause_1. destruct (inspect_bool (is_erasableb X_type Xext [] (tConstruct i n ui) Hyp0)). - -- exfalso. sq. destruct (@is_erasableP _ _ _ [] (tConstruct i n ui) Hyp0) => //. - specialize_Σ Hrel. sq. - eapply (isErasable_Propositional (args := [])) in s; eauto. - red in H1, s. unfold PCUICAst.lookup_inductive, lookup_inductive_gen, PCUICAst.lookup_minductive, lookup_minductive_gen, isPropositionalArity in *. - edestruct PCUICEnvironment.lookup_env as [ [] | ], nth_error, destArity as [[] | ]; auto; congruence. - -- f_equal. eapply Forall2_eq. clear X0 H wv. induction H3. - ++ cbn. econstructor. - ++ cbn. econstructor. - ** inv H0. eapply H5. eauto. - ** inv H0. eapply IHForall2. eauto. - * exfalso. eapply (isErasable_Propositional (args := [])) in X1; eauto. - red in H1, X1. - unfold PCUICAst.lookup_inductive, lookup_inductive_gen, PCUICAst.lookup_minductive, lookup_minductive_gen, isPropositionalArity in *. - edestruct PCUICEnvironment.lookup_env as [ [] | ], nth_error, destArity as [[] | ]; auto; congruence. - - eauto. - - intros ? ? H3. assert (Hext_ : ∥ wf_ext Σ0∥) by now eapply heΣ. - sq. - specialize_Σ H2. - eapply (isErasable_Propositional) in H3; eauto. - pose proof (abstract_env_ext_irr _ H2 Hrel). subst. - red in H1, H3. unfold PCUICAst.lookup_inductive, lookup_inductive_gen, PCUICAst.lookup_minductive, lookup_minductive_gen, isPropositionalArity in *. - edestruct PCUICEnvironment.lookup_env as [ [] | ], nth_error, destArity as [[] | ]; auto; congruence. - - intros. assert (Hext__ : ∥ wf_ext Σ0∥) by now eapply heΣ. - specialize_Σ H2. eapply welltyped_mkApps_inv in wv; eauto. eapply wv. - now sq. - Unshelve. all: try exact False. - - eapply PCUICWcbvEval.eval_to_value; eauto. -Qed. - -Lemma erase_correct_strong' (wfl := Ee.default_wcbv_flags) X_type (X : X_type.π1) -univs wfext {t v Σ' t' deps i u args} decls normalization_in prf -(Xext := abstract_make_wf_env_ext X univs wfext) -{normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} -: -forall wt : (forall Σ, Σ ∼_ext Xext -> welltyped Σ [] t), -forall Σ, abstract_env_ext_rel Xext Σ -> - axiom_free Σ -> - Σ ;;; [] |- t : mkApps (tInd i u) args -> - @firstorder_ind Σ (firstorder_env Σ) i -> - erase X_type Xext [] t wt = t' -> - KernameSet.subset (term_global_deps t') deps -> - erase_global_deps X_type deps X decls normalization_in prf = Σ' -> - red Σ [] t v -> - ¬ { t' & Σ ;;; [] |- v ⇝ t'} -> - forall wt', ∥ Σ'.1 ⊢ t' ⇓ erase X_type Xext [] v wt' ∥. -Proof. - intros wt Σ Hrel Hax Hty Hfo <- Hsub <- Hred Hirred wt'. - destruct (firstorder_lookup_inv Hfo) as [mind Hdecl]. - pose proof (heΣ _ _ _ Hrel) as [Hwf]. eapply wcbv_standardization_fst in Hty as Heval; eauto. - edestruct (erase_correct X_type X univs wfext t v) as [v' [H1 H2]]; eauto. - 1:{ intros ? H_. sq. enough (Σ0 = Σ) as -> by eauto. - pose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ H_ Hrel). now subst. } - eapply firstorder_erases_deterministic in H1; eauto. - { rewrite H1 in H2. eapply H2. } - { eapply subject_reduction; eauto. } - { eapply PCUICWcbvEval.value_final. eapply PCUICWcbvEval.eval_to_value. eauto. } - all: eauto. - Unshelve. eauto. -Qed. - -Lemma erase_correct_strong (wfl := Ee.default_wcbv_flags) X_type (X : X_type.π1) -univs wfext {t v deps i u args} decls normalization_in prf -(Xext := abstract_make_wf_env_ext X univs wfext) -{normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} -: -forall wt : (forall Σ, Σ ∼_ext Xext -> welltyped Σ [] t), -forall Σ, abstract_env_ext_rel Xext Σ -> - axiom_free Σ -> - Σ ;;; [] |- t : mkApps (tInd i u) args -> - let t' := erase X_type Xext [] t wt in - let Σ' := erase_global_deps X_type deps X decls normalization_in prf in - @firstorder_ind Σ (firstorder_env Σ) i -> - KernameSet.subset (term_global_deps t') deps -> - red Σ [] t v -> - ¬ { v' & Σ ;;; [] |- v ⇝ v'} -> - exists wt', ∥ Σ'.1 ⊢ t' ⇓ erase X_type Xext [] v wt' ∥. -Proof. - intros wt Σ Hrel Hax Hty Hdecl Hfo Hsub Hred Hirred. - unshelve eexists. - - abstract (intros Σ_ H_; pose proof (heΣ _ _ _ H_); sq; - pose proof (abstract_env_ext_irr _ H_ Hrel); subst; eapply red_welltyped; eauto; econstructor; eauto). - - eapply erase_correct_strong'; eauto. - all: reflexivity. -Qed. - -Lemma unique_decls {X_type} {X : X_type.π1} univs wfext -(Xext := abstract_make_wf_env_ext X univs wfext) {Σ'} (wfΣ' : Σ' ∼_ext Xext) : - forall Σ, Σ ∼ X -> declarations Σ = declarations Σ'. -Proof. - intros Σ wfΣ. pose (abstract_make_wf_env_ext_correct X univs wfext Σ Σ' wfΣ wfΣ'). - rewrite e. reflexivity. -Qed. - -Lemma unique_type {X_type : abstract_env_impl} {X : X_type.π2.π1} - {Σ:global_env_ext} - (wfΣ : Σ ∼_ext X) {t T}: - Σ ;;; [] |- t : T -> forall Σ, Σ ∼_ext X -> welltyped Σ [] t. -Proof. - intros wt Σ' wfΣ'. erewrite (abstract_env_ext_irr X wfΣ' wfΣ). - now exists T. -Qed. - -Definition erase_correct_firstorder (wfl := Ee.default_wcbv_flags) - X_type (X : X_type.π1) univs wfext {t v i u args} -(Xext := abstract_make_wf_env_ext X univs wfext) -{normalization_in : forall X Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} -: -forall Σ (wfΣ : abstract_env_ext_rel Xext Σ) - (wt : Σ ;;; [] |- t : mkApps (tInd i u) args), - axiom_free Σ -> - let t' := erase X_type Xext [] t (unique_type wfΣ wt) in - let decls := declarations (fst Σ) in - let Σ' := erase_global_deps X_type (term_global_deps t') X decls - (fun _ _ _ _ _ _ _ => normalization_in _ _) (unique_decls univs wfext wfΣ) in - @firstorder_ind Σ (firstorder_env Σ) i -> - red Σ [] t v -> - ¬ { v' & Σ ;;; [] |- v ⇝ v'} -> - exists wt', ∥ Σ'.1 ⊢ t' ⇓ erase X_type Xext [] v wt' ∥. -Proof. - intros. - eapply erase_correct_strong; eauto. - apply KernameSetOrdProp.P.Dec.F.subset_iff. reflexivity. -Qed. - -(* we use the [match] trick to get typeclass resolution to pick up the right instances without leaving any evidence in the resulting term, and without having to pass them manually everywhere *) -Notation NormalizationIn_erase_global_deps_fast X decls - := (match extraction_checker_flags, extraction_normalizing return _ with - | cf, no - => forall n, - n < List.length decls -> - forall kn cb pf, - List.hd_error (List.skipn n decls) = Some (kn, ConstantDecl cb) -> - let Xext := abstract_make_wf_env_ext X (cst_universes cb) pf in - forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ - end) - (only parsing). - - -Section EraseGlobalFast. - - Import PCUICEnvironment. - -Definition decls_prefix decls (Σ' : global_env) := - ∑ Σ'', declarations Σ' = Σ'' ++ decls. - -Lemma on_global_decls_prefix {cf} Pcmp P univs retro decls decls' : - PCUICAst.on_global_decls Pcmp P univs retro (decls ++ decls') -> - PCUICAst.on_global_decls Pcmp P univs retro decls'. -Proof using Type. - induction decls => //. - intros ha; depelim ha. - now apply IHdecls. -Qed. - -Lemma decls_prefix_wf {decls Σ} : - decls_prefix decls Σ -> wf Σ -> wf {| universes := Σ.(universes); declarations := decls; retroknowledge := Σ.(retroknowledge) |}. -Proof using Type. - intros [Σ' hd] wfΣ. - split. apply wfΣ. - cbn. destruct wfΣ as [_ ond]. rewrite hd in ond. - now eapply on_global_decls_prefix in ond. -Qed. - -Lemma incl_cs_refl cs : cs ⊂_cs cs. -Proof using Type. - split; [lsets|csets]. -Qed. - -Lemma weaken_prefix {decls Σ kn decl} : - decls_prefix decls Σ -> - wf Σ -> - lookup_env {| universes := Σ; declarations := decls; retroknowledge := Σ.(retroknowledge) |} kn = Some decl -> - on_global_decl cumulSpec0 (lift_typing typing) (Σ, universes_decl_of_decl decl) kn decl. -Proof using Type. - intros prefix wfΣ. - have wfdecls := decls_prefix_wf prefix wfΣ. - epose proof (weakening_env_lookup_on_global_env (lift_typing typing) _ Σ kn decl - weaken_env_prop_typing wfΣ wfdecls). - forward X. apply extends_strictly_on_decls_extends, strictly_extends_decls_extends_strictly_on_decls. red; split => //. - now apply (X wfdecls). -Qed. - - -(* This version of global environment erasure keeps the same global environment throughout the whole - erasure, while folding over the list of declarations. *) - -Local Obligation Tactic := program_simpl. - -Program Fixpoint erase_global_deps_fast (deps : KernameSet.t) - X_type (X:X_type.π1) (decls : global_declarations) - {normalization_in : NormalizationIn_erase_global_deps_fast X decls} - (prop : forall Σ : global_env, abstract_env_rel X Σ -> ∥ decls_prefix decls Σ ∥) : E.global_declarations * KernameSet.t := - match decls with - | [] => ([],deps) - | (kn, ConstantDecl cb) :: decls => - if KernameSet.mem kn deps then - let Xext' := abstract_make_wf_env_ext X (cst_universes cb) _ in - let normalization_in' : @id Prop _ := _ in (* hack to avoid program erroring out on unresolved tc *) - let cb' := @erase_constant_body X_type Xext' normalization_in' cb _ in - let deps := KernameSet.union deps (snd cb') in - let Σ' := erase_global_deps_fast deps X_type X decls _ in - (((kn, E.ConstantDecl (fst cb')) :: Σ'.1), Σ'.2) - else - erase_global_deps_fast deps X_type X decls _ - - | (kn, InductiveDecl mib) :: decls => - if KernameSet.mem kn deps then - let mib' := erase_mutual_inductive_body mib in - let Σ' := erase_global_deps_fast deps X_type X decls _ in - (((kn, E.InductiveDecl mib') :: Σ'.1), Σ'.2) - else erase_global_deps_fast deps X_type X decls _ - end. -Next Obligation. - pose proof (abstract_env_wf _ H) as [?]. - specialize_Σ H. sq. split. cbn. apply X3. cbn. - eapply decls_prefix_wf in X3; tea. - destruct X3 as [onu ond]. cbn in ond. - depelim ond. now destruct o. -Qed. -Next Obligation. - cbv [id]. - unshelve eapply (normalization_in 0); cbn; try reflexivity; try lia. -Defined. -Next Obligation. - pose proof (abstract_env_ext_wf _ H) as [?]. - pose proof (abstract_env_exists X) as [[? wf]]. - pose proof (abstract_env_wf _ wf) as [?]. - pose proof (prop' := prop _ wf). - sq. eapply (weaken_prefix (kn := kn)) in prop'; tea. - 2:{ cbn. rewrite eqb_refl //. } - epose proof (abstract_make_wf_env_ext_correct X (cst_universes cb) _ _ _ wf H). subst. - apply prop'. -Qed. -Next Obligation. - unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Defined. -Next Obligation. - specialize_Σ H. sq; red. destruct prop as [Σ' ->]. - eexists (Σ' ++ [(kn, ConstantDecl cb)]); rewrite -app_assoc //. -Qed. -Next Obligation. - unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Defined. -Next Obligation. - specialize_Σ H. sq; red. destruct prop as [Σ' ->]. - eexists (Σ' ++ [(kn, ConstantDecl cb)]); rewrite -app_assoc //. -Qed. -Next Obligation. - unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Defined. -Next Obligation. - specialize_Σ H. sq; red. destruct prop as [Σ' ->]. - eexists (Σ' ++ [(kn, _)]); rewrite -app_assoc //. -Qed. -Next Obligation. - unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. -Defined. -Next Obligation. - specialize_Σ H. sq; red. destruct prop as [Σ' ->]. - eexists (Σ' ++ [(kn, _)]); rewrite -app_assoc //. -Qed. - -Import PCUICAst PCUICEnvironment. - -Lemma wf_lookup Σ kn d suffix g : - wf Σ -> - declarations Σ = suffix ++ (kn, d) :: g -> - lookup_env Σ kn = Some d. -Proof using Type. - unfold lookup_env. - intros [_ ond] eq. rewrite eq in ond |- *. - move: ond; clear. - induction suffix. - - cbn. rewrite eqb_refl //. - - cbn. intros ond. - depelim ond. destruct o as [f ? ? ? ]. cbn. red in f. - eapply Forall_app in f as []. depelim H0. cbn in H0. - destruct (eqb_spec kn kn0); [contradiction|]. - now apply IHsuffix. -Qed. - -Definition add_suffix suffix Σ := set_declarations Σ (suffix ++ Σ.(declarations)). - -Lemma add_suffix_cons d suffix Σ : add_suffix (d :: suffix) Σ = add_global_decl (add_suffix suffix Σ) d. -Proof using Type. reflexivity. Qed. - -Lemma global_erased_with_deps_weaken_prefix suffix Σ Σ' kn : - wf (add_suffix suffix Σ) -> - global_erased_with_deps Σ Σ' kn -> - global_erased_with_deps (add_suffix suffix Σ) Σ' kn. -Proof using Type. - induction suffix. - - unfold add_suffix; cbn. intros wf hg. - now rewrite /set_declarations /= -eta_global_env. - - rewrite add_suffix_cons. intros wf H. - destruct a as [kn' d]. eapply global_erases_with_deps_weaken => //. - apply IHsuffix => //. - destruct wf as [onu ond]. now depelim ond. -Qed. - - -(* Using weakening it is trivial to show that a term found to be erasable in Σ - will be found erasable in any well-formed extension. The converse is not so trivial: - some valid types in the extension are not valid in the restricted global context. - So, we will rather show that the erasure function is invariant under extension. *) - -Lemma isErasable_irrel_global_env {Σ Σ' : global_env_ext} {Γ t} : - wf Σ -> - wf Σ' -> - extends Σ' Σ -> - Σ.2 = Σ'.2 -> - isErasable Σ' Γ t -> isErasable Σ Γ t. -Proof using Type. - unfold isErasable. - intros wfΣ wfΣ' ext eqext [T [ht isar]]. - destruct Σ as [Σ decl], Σ' as [Σ' decl']. cbn in eqext; subst decl'. - exists T; split. - eapply (env_prop_typing weakening_env (Σ', decl)) => //=. - destruct isar as [|s]; [left|right] => //. - destruct s as [u [Hu isp]]. - exists u; split => //. - eapply (env_prop_typing weakening_env (Σ', decl)) => //=. Qed. - -Definition reduce_stack_eq {cf} {fl} {no:normalizing_flags} {X_type : abstract_env_impl} {X : X_type.π2.π1} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ t π wi : reduce_stack fl X_type X Γ t π wi = ` (reduce_stack_full fl X_type X Γ t π wi). -Proof using Type. - unfold reduce_stack. destruct reduce_stack_full => //. -Qed. - -Definition same_principal_type {cf} - {X_type : abstract_env_impl} {X : X_type.π2.π1} - {X_type' : abstract_env_impl} {X' : X_type'.π2.π1} - {Γ : context} {t} (p : PCUICSafeRetyping.principal_type X_type X Γ t) (p' : PCUICSafeRetyping.principal_type X_type' X' Γ t) := - p.π1 = p'.π1. - -(*Lemma erase_global_deps_suffix {deps} {Σ Σ' : wf_env} {decls hprefix hprefix'} : - wf Σ -> wf Σ' -> - universes Σ = universes Σ' -> - erase_global_deps_fast deps Σ decls hprefix = - erase_global_deps_fast deps Σ' decls hprefix'. -Proof using Type. - intros wfΣ wfΣ' equ. - induction decls in deps, hprefix, hprefix' |- *; cbn => //. - destruct a as [kn []]. - set (obl := (erase_global_deps_fast_obligation_1 Σ ((kn, ConstantDecl c) :: decls) - hprefix kn c decls eq_refl)). clearbody obl. - destruct obl. - set (eb := erase_constant_body _ _ _). - set (eb' := erase_constant_body _ _ _). - assert (eb = eb') as <-. - { subst eb eb'. - set (wfΣx := abstract_make_wf_env_ext _ _ _). - set (wfΣ'x := abstract_make_wf_env_ext _ _ _). - set (obl' := erase_global_deps_fast_obligation_2 _ _ _ _ _ _ _). - clearbody obl'. - set (Σd := {| universes := Σ.(universes); declarations := decls |}). - assert (wfΣd : wf Σd). - { eapply decls_prefix_wf; tea. - clear -hprefix. move: hprefix. unfold decls_prefix. - intros [Σ'' eq]. exists (Σ'' ++ [(kn, ConstantDecl c)]). - rewrite -app_assoc //. } - set (wfΣdwfe := build_wf_env_from_env _ (sq wfΣd)). - assert (wfextΣd : wf_ext (Σd, cst_universes c)). - { split => //. cbn. apply w. } - set (wfΣdw := abstract_make_wf_env_ext wfΣdwfe (cst_universes c) (sq wfextΣd)). - assert (ond' : ∥ on_constant_decl (lift_typing typing) (Σd, cst_universes c) c ∥ ). - { pose proof (decls_prefix_wf hprefix wfΣ). - destruct X as [onu ond]. depelim ond. sq. apply o0. } - assert (extd' : extends_decls Σd Σ). - { clear -hprefix. move: hprefix. - intros [? ?]. split => //. eexists (x ++ [(kn, ConstantDecl c)]). cbn. - now rewrite -app_assoc. } - assert (extd'' : extends_decls Σd Σ'). - { clear -equ hprefix'. move: hprefix'. - intros [? ?]. split => //. eexists (x ++ [(kn, ConstantDecl c)]). cbn. - now rewrite -app_assoc. } - rewrite (erase_constant_suffix (Σ := wfΣx) (Σ' := wfΣdw) (ondecl' := ond') wfΣ extd') //. - rewrite (erase_constant_body_irrel (Σ := wfΣ'x) (Σ' := wfΣdw) (ondecl' := ond') wfΣ' extd'') //. } - destruct KernameSet.mem => //. f_equal. - eapply IHdecls. - destruct KernameSet.mem => //. f_equal. - eapply IHdecls. -Qed.*) - -Lemma strictly_extends_decls_extends_global_env Σ Σ' : - wf Σ' -> - strictly_extends_decls Σ Σ' -> - extends_global_env Σ Σ'. -Proof. - intros wfΣ' [uni [Σ'' de] retr]. red. - cbn. unfold lookup_env. rewrite de /primitive_constant retr. - split => //. intros kn decl. - intros hl. - apply lookup_global_Some_iff_In_NoDup. - apply EnvMap.EnvMap.fresh_globals_iff_NoDup. - rewrite -de. now apply wf_fresh_globals. - apply lookup_global_Some_iff_In_NoDup in hl. - apply in_app_iff. now right. - apply EnvMap.EnvMap.fresh_globals_iff_NoDup. - apply wf_fresh_globals in wfΣ'. rewrite de in wfΣ'. - clear -wfΣ'. induction Σ''; auto. apply IHΣ''. - now depelim wfΣ'. -Qed. - -Lemma extends_global_env_equiv_env Σ Σ' : - extends_global_env Σ Σ' -> equiv_env_inter Σ Σ'. -Proof. - intros ext. split. intros kn decl decl' hl hl'. - destruct ext. apply H in hl. congruence. - now destruct ext. -Qed. - -Lemma erase_global_deps_fast_spec_gen {deps} - {X_type X X'} {decls normalization_in normalization_in' hprefix hprefix'} : - (forall Σ Σ', abstract_env_rel X Σ -> abstract_env_rel X' Σ' -> universes Σ = universes Σ' /\ retroknowledge Σ = retroknowledge Σ') -> - erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) hprefix = - erase_global_deps X_type deps X' decls normalization_in' hprefix'. -Proof using Type. - intros equ. - induction decls in X, X', equ, deps, normalization_in, normalization_in', hprefix, hprefix' |- * => //. - pose proof (abstract_env_exists X) as [[Σ wfΣ]]. - pose proof (abstract_env_exists X') as [[Σ' wfΣ']]. - pose proof (abstract_env_wf _ wfΣ) as [wf]. - pose proof (abstract_env_exists (abstract_pop_decls X')) as [[? wfpop]]. - unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' wfpop) as [? [? ?]]. - { now eexists. } - - destruct a as [kn []]. - - cbn. unfold erase_constant_decl. - set (obl := (erase_global_deps_fast_obligation_1 X_type X ((kn, ConstantDecl c) :: decls) - _ hprefix kn c decls eq_refl)). - set (eb := erase_constant_body _ _ _ _). - set (eb' := erase_constant_body _ _ _ _). - set (eg := erase_global_deps _ _ _ _ _ _). - set (eg' := erase_global_deps _ _ _ _ _ _). - - assert (eb = eb') as <-. - { subst eb eb'. - destruct (hprefix _ wfΣ) as [[Σ'' eq]]. - eapply erase_constant_body_irrel; cbn => //. - eapply equiv_env_inter_hlookup. - intros ? ? H2 H3. - epose proof (abstract_make_wf_env_ext_correct X (cst_universes c) _ _ _ wfΣ H2). - epose proof (abstract_make_wf_env_ext_correct (abstract_pop_decls X') (cst_universes c) _ _ _ wfpop H3). - subst. eapply equiv_env_inter_sym. - eapply extends_global_env_equiv_env. - pose proof (abstract_env_wf _ wfΣ) as []. - apply strictly_extends_decls_extends_global_env. - apply X0. red. - rewrite eq. rewrite <- H0, <- H1. split. cbn. symmetry; apply equ; eauto. - cbn. - eexists (Σ'' ++ [(kn, ConstantDecl c)]). cbn. subst. now rewrite -app_assoc. subst. - symmetry. now apply equ. } - destruct KernameSet.mem => //; f_equal; try eapply IHdecls. - f_equal. f_equal. eapply IHdecls. - intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. - { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. - f_equal. eapply IHdecls. - intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. - { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. - intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. - { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. - - - cbn. - destruct KernameSet.mem => //; f_equal; try eapply IHdecls. - f_equal; f_equal; eapply IHdecls. - intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. - { now eexists. } - intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. - intros. - f_equal; eapply IHdecls. - intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. - { now eexists. } - intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. - intros; unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. - { now eexists. } - intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. -Qed. - -Lemma erase_global_deps_fast_spec {deps} {X_type X} {decls normalization_in normalization_in' hprefix hprefix'} : - erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) hprefix = - erase_global_deps X_type deps X decls normalization_in' hprefix'. -Proof using Type. - eapply erase_global_deps_fast_spec_gen; intros. - rewrite (abstract_env_irr _ H H0); eauto. -Qed. - -Lemma cored_extends (Σ Σ' : global_env_ext) Γ x y : - wf Σ' -> - extends Σ Σ' -> - cored Σ Γ x y -> cored Σ' Γ x y. -Proof. - intros wf ext. - induction 1. - - constructor. - eapply PCUICWeakeningEnvConv.weakening_env_red1. apply wf. exact ext. - exact X. - - econstructor 2; tea. - eapply PCUICWeakeningEnvConv.weakening_env_red1. apply wf. - exact ext. - exact X. -Qed. - -Lemma normalization_in_extends (Σ Σ' : global_env) ext : - wf Σ -> wf Σ' -> - extends Σ' Σ -> - NormalizationIn (Σ, ext) -> NormalizationIn (Σ', ext). -Proof. - intros wfΣ wfΣ' extH nin. - red in nin. intros Γ t wt. - assert (welltyped (Σ, ext) Γ t). - { destruct wt. exists A. eapply (env_prop_typing weakening_env) in X; tea; cbn. } - specialize (nin _ _ H). clear H. - induction nin. - constructor. intros y r. - eapply H0. eapply cored_extends in r; tea. - eapply cored_welltyped; tea. -Qed. - -Lemma iter_abstract_pop_decls_correct {cf} {abstract_env_impl abstract_env_ext_impl : Type} - {abstract_env_struct0 : abstract_env_struct abstract_env_impl - abstract_env_ext_impl} : - abstract_env_prop abstract_env_impl abstract_env_ext_impl -> - forall (X : abstract_env_impl) (decls : list (kername × global_decl)) n, - (forall Σ : global_env, - Σ ∼ X -> exists decls', #|decls'| = n /\ declarations Σ = decls' ++ decls) -> - let X' := iter abstract_pop_decls n X in - forall Σ Σ' : global_env, - Σ ∼ X -> - Σ' ∼ X' -> - declarations Σ' = decls /\ universes Σ = universes Σ' /\ retroknowledge Σ = retroknowledge Σ'. -Proof. - intros ap X decls. - induction n in X, decls |- *. cbn. - - intros hdecl Σ Σ' HX HX'. - specialize (hdecl _ HX) as [decls' [hlen hde]]. - destruct decls' => //. cbn in hde. - now rewrite -hde (abstract_env_irr _ HX HX'). - - intros hdecl X' Σ Σ' HX HX'. - specialize (hdecl _ HX) as [decls' [hlen hde]]. - destruct decls' => //; cbn in hde. - specialize (IHn (abstract_pop_decls X) decls). - pose proof (abstract_pop_decls_correct X (decls' ++ decls)). - forward H. - { intros ? h. rewrite -(abstract_env_irr _ HX h). now exists p. } - forward IHn. - { intros ? H0. now specialize (H _ _ HX H0). } - pose proof (abstract_env_exists (abstract_pop_decls X)) as [[Σpop hpop]]. - specialize (IHn _ _ hpop HX'). - destruct IHn. split => //. - pose proof (abstract_pop_decls_correct X (decls' ++ decls)). - forward H2. - { intros ? hx. pose proof (abstract_env_irr _ hx HX). subst Σ0. now exists p. } - specialize (H2 _ _ HX hpop). destruct H2 as [? []]. - destruct H1. - split; congruence. -Qed. - -Lemma skipn_app_prefix {A} n (l l' : list A) : skipn (#|l| + n) (l ++ l') = skipn n l'. -Proof. - induction l; cbn; auto. -Qed. - -Lemma erase_global_deps_fast_erase_global_deps {deps} {X_type X} {decls normalization_in hprefix hprefix'} : - exists normalization_in', - erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) hprefix = - erase_global_deps X_type deps X decls normalization_in' hprefix'. -Proof using Type. - unshelve eexists. - intros. - pose proof (abstract_env_exists X) as [[Σ' H']]. - pose proof (abstract_env_wf _ H') as []. - specialize (hprefix _ H') as [[Σ'' ?]]. - pose proof (abstract_env_exists (iter abstract_pop_decls (S n) X)) as [[? ?]]. - pose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ a H1). noconf H2. - epose proof (iter_abstract_pop_decls_correct _ X (skipn (S n) (Σ'' ++ decls)) (S n)). - forward H2. { intros ? h. rewrite (abstract_env_irr _ h H'). exists (firstn (S n) (Σ'' ++ decls)). - rewrite firstn_skipn. split => //. rewrite firstn_length_le //. rewrite app_length. lia. } - specialize (H2 _ _ H' a) as [? []]. - eapply normalization_in_extends. exact X1. exact X0. - { eapply extends_decls_extends, strictly_extends_decls_extends_decls. - red. rewrite H2 H3 H4. split => //. exists (firstn (S n) (Σ'' ++ decls)). - now rewrite firstn_skipn. } - specialize (normalization_in n H kn cb). - assert (wf_ext (Σ', cst_universes cb)). - { split => //. cbn. rewrite H3. apply X0. } - assert (forall Σ : global_env, Σ ∼ X -> ∥ wf_ext (Σ, cst_universes cb) ∥). - { intros. pose proof (abstract_env_irr _ H' H5). now subst Σ'. } - unshelve eapply normalization_in; tea. - set (foo := abstract_make_wf_env_ext X (cst_universes cb) H5). - pose proof (abstract_env_ext_exists foo) as [[Σext hext]]. - epose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ H' hext). now subst Σext. - erewrite erase_global_deps_fast_spec; tea. reflexivity. -Qed. - -Definition erase_global_fast X_type deps X decls {normalization_in} (prf:forall Σ : global_env, abstract_env_rel X Σ -> declarations Σ = decls) := - erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) (fun _ H => sq ([] ; prf _ H)). - -Lemma expanded_erase_global_fast (cf := config.extraction_checker_flags) deps - {X_type X decls normalization_in} {normalization_in':NormalizationIn_erase_global_deps X decls} {prf} : - forall Σ : global_env, abstract_env_rel X Σ -> - PCUICEtaExpand.expanded_global_env Σ -> - expanded_global_env (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf).1. -Proof using Type. - unfold erase_global_fast. - erewrite erase_global_deps_fast_spec. - eapply expanded_erase_global. - Unshelve. all: eauto. -Qed. - -Lemma expanded_erase_fast (cf := config.extraction_checker_flags) - {X_type X decls normalization_in prf} {normalization_in'':NormalizationIn_erase_global_deps X decls} univs wfΣ t wtp : - forall Σ : global_env, abstract_env_rel X Σ -> - PCUICEtaExpand.expanded Σ [] t -> - let X' := abstract_make_wf_env_ext X univs wfΣ in - forall (normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ), - let et := (erase X_type X' [] t wtp) in - let deps := EAstUtils.term_global_deps et in - expanded (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf).1 [] et. -Proof using Type. - intros Σ wf hexp X' normalization_in'. pose proof (abstract_env_wf _ wf) as [?]. - eapply (expanded_erases (Σ := (Σ, univs))); tea. - eapply (erases_erase (X := X')). - pose proof (abstract_env_ext_exists X') as [[? wfmake]]. - now rewrite <- (abstract_make_wf_env_ext_correct X univs _ _ _ wf wfmake). - cbn. unfold erase_global_fast. erewrite erase_global_deps_fast_spec => //. - eapply (erases_deps_erase (X := X) univs wfΣ); eauto. - Unshelve. all: eauto. -Qed. - -Lemma erase_global_fast_wf_glob X_type deps X decls normalization_in (normalization_in':NormalizationIn_erase_global_deps X decls) prf : - @wf_glob all_env_flags (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf).1. -Proof using Type. - unfold erase_global_fast; erewrite erase_global_deps_fast_spec. - eapply erase_global_deps_wf_glob. - Unshelve. all: eauto. -Qed. - -Lemma erase_wellformed_fast (efl := all_env_flags) - {X_type X decls normalization_in prf} {normalization_in'':NormalizationIn_erase_global_deps X decls} univs wfΣ {Γ t} wt - (X' := abstract_make_wf_env_ext X univs wfΣ) - {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} : - let t' := erase X_type X' Γ t wt in - wellformed (erase_global_fast X_type (term_global_deps t') X decls (normalization_in:=normalization_in) prf).1 #|Γ| t'. -Proof using Type. - intros. - cbn. unfold erase_global_fast. erewrite erase_global_deps_fast_spec. - eapply erase_wellformed_deps. - Unshelve. all: eauto. -Qed. - -End EraseGlobalFast. - -Fixpoint compile_value_erase (t : PCUICAst.term) (acc : list EAst.term) : EAst.term := - match t with - | PCUICAst.tApp f a => compile_value_erase f (compile_value_erase a [] :: acc) - | PCUICAst.tConstruct i n _ => EAst.mkApps (EAst.tConstruct i n []) acc - | _ => EAst.tVar "error" - end. - -Lemma compile_value_erase_mkApps i n ui args acc : - compile_value_erase (mkApps (tConstruct i n ui) args) acc = - EAst.mkApps (EAst.tConstruct i n []) (map (flip compile_value_erase []) args ++ acc). -Proof. - revert acc; induction args using rev_ind. - - cbn. auto. - - intros acc. rewrite PCUICAstUtils.mkApps_app /=. cbn. - now rewrite IHargs map_app /= -app_assoc /=. -Qed. - -Lemma erases_firstorder Σ Γ t : - PCUICFirstorder.firstorder_value Σ Γ t -> - erases Σ Γ t (compile_value_erase t []). -Proof. - move: t. - apply: (PCUICFirstorder.firstorder_value_inds Σ Γ). - intros i n ui u args pandi hty hfo ih isp. - assert (Forall2 (erases Σ Γ) args (map (flip compile_value_erase []) args)). - { solve_all. eapply All_All2; tea. - cbn. intros x [fo hx]. exact hx. } - unshelve epose proof (erases_mkApps Σ Γ (tConstruct i n ui) (EAst.tConstruct i n []) args _ _ H). - constructor. - { move: isp. rewrite /isPropositional /Extract.isPropositional /PCUICAst.lookup_inductive /= /lookup_inductive_gen /lookup_minductive_gen. - destruct PCUICEnvironment.lookup_env => //. - destruct g => //. destruct nth_error => //. } - now rewrite compile_value_erase_mkApps app_nil_r. -Qed. - -Lemma erases_firstorder' Σ Γ t : - wf_ext Σ -> - PCUICFirstorder.firstorder_value Σ Γ t -> - forall t', erases Σ Γ t t' -> t' = (compile_value_erase t []). -Proof. - intros wf. move: t. - apply: (PCUICFirstorder.firstorder_value_inds Σ Γ). - intros i n ui u args pandi hty hfo ih isp. - assert (Forall2 (erases Σ Γ) args (map (flip compile_value_erase []) args)). - { solve_all. eapply All_All2; tea. - cbn. intros x [fo hx]. now eapply erases_firstorder. } - unshelve epose proof (erases_mkApps Σ Γ (tConstruct i n ui) (EAst.tConstruct i n []) args _ _ H). - constructor. - { move: isp. rewrite /isPropositional /Extract.isPropositional /PCUICAst.lookup_inductive /= /lookup_inductive_gen /lookup_minductive_gen. - destruct PCUICEnvironment.lookup_env => //. - destruct g => //. destruct nth_error => //. } - rewrite compile_value_erase_mkApps app_nil_r. - intros t' ht'. - eapply erases_mkApps_inv in ht'. - destruct ht'. - { destruct H1 as [? [? [? [? [[] ?]]]]]. - eapply isErasable_Propositional in X => //. - clear -isp X. elimtype False. - move: isp X. - rewrite /isPropositional /Extract.isPropositional /PCUICAst.lookup_inductive /= /lookup_inductive_gen /lookup_minductive_gen. - destruct PCUICEnvironment.lookup_env => //. - destruct g => //. destruct nth_error => //. - rewrite /isPropositionalArity. destruct destArity => //. - destruct p. intros ->. auto. } - destruct H1 as [f' [L' [-> [erc erargs]]]]. - depelim erc. f_equal. - solve_all. clear -erargs. - induction erargs => //. rewrite IHerargs. cbn. f_equal. - destruct r as [[] ?]. eapply e in e0. now unfold flip. - eapply (isErasable_Propositional (args := [])) in X => //. - clear -isp X. elimtype False. - move: isp X. - rewrite /isPropositional /Extract.isPropositional /PCUICAst.lookup_inductive /= /lookup_inductive_gen /lookup_minductive_gen. - destruct PCUICEnvironment.lookup_env => //. - destruct g => //. destruct nth_error => //. - rewrite /isPropositionalArity. destruct destArity => //. - destruct p. intros ->. auto. -Qed. \ No newline at end of file diff --git a/erasure/theories/ErasureFunctionProperties.v b/erasure/theories/ErasureFunctionProperties.v new file mode 100644 index 000000000..b1c0a710e --- /dev/null +++ b/erasure/theories/ErasureFunctionProperties.v @@ -0,0 +1,2877 @@ +(* Distributed under the terms of the MIT license. *) +From Coq Require Import Program ssreflect ssrbool. +From Equations Require Import Equations. +Set Equations Transparent. +From MetaCoq.Utils Require Import utils. +From MetaCoq.Common Require Import config Kernames uGraph. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICPrimitive + PCUICReduction PCUICReflect PCUICWeakeningEnv PCUICWeakeningEnvTyp PCUICCasesContexts + PCUICWeakeningConv PCUICWeakeningTyp PCUICContextConversionTyp PCUICTyping PCUICGlobalEnv PCUICInversion PCUICGeneration + PCUICConfluence PCUICConversion PCUICUnivSubstitutionTyp PCUICCumulativity PCUICSR PCUICSafeLemmata PCUICNormalization + PCUICValidity PCUICPrincipality PCUICElimination PCUICOnFreeVars PCUICWellScopedCumulativity PCUICSN PCUICEtaExpand. + +From MetaCoq.SafeChecker Require Import PCUICErrors PCUICWfEnv PCUICSafeReduce PCUICSafeRetyping PCUICRetypingEnvIrrelevance. +From MetaCoq.Erasure Require Import EAstUtils EArities Extract Prelim EDeps ErasureProperties ErasureCorrectness ErasureFunction. + +Local Open Scope string_scope. +Set Asymmetric Patterns. +Local Set Keyed Unification. +Set Default Proof Using "Type*". +Import MCMonadNotation. + +Implicit Types (cf : checker_flags). + +#[local] Existing Instance extraction_normalizing. +#[local] Existing Instance extraction_checker_flags. + +Notation alpha_eq := (All2 (PCUICEquality.compare_decls eq eq)). + +Definition global_erased_with_deps (Σ : global_env) (Σ' : EAst.global_declarations) kn := + (exists cst, declared_constant Σ kn cst /\ + exists cst' : EAst.constant_body, + EGlobalEnv.declared_constant Σ' kn cst' /\ + erases_constant_body (Σ, cst_universes cst) cst cst' /\ + (forall body : EAst.term, + EAst.cst_body cst' = Some body -> erases_deps Σ Σ' body)) \/ + (exists mib mib', declared_minductive Σ kn mib /\ + EGlobalEnv.declared_minductive Σ' kn mib' /\ + erases_mutual_inductive_body mib mib'). + +Definition includes_deps (Σ : global_env) (Σ' : EAst.global_declarations) deps := + forall kn, + KernameSet.In kn deps -> + global_erased_with_deps Σ Σ' kn. + +Lemma includes_deps_union (Σ : global_env) (Σ' : EAst.global_declarations) deps deps' : + includes_deps Σ Σ' (KernameSet.union deps deps') -> + includes_deps Σ Σ' deps /\ includes_deps Σ Σ' deps'. +Proof. + intros. + split; intros kn knin; eapply H; rewrite KernameSet.union_spec; eauto. +Qed. + +Lemma includes_deps_fold {A} (Σ : global_env) (Σ' : EAst.global_declarations) brs deps (f : A -> EAst.term) : + includes_deps Σ Σ' + (fold_left + (fun (acc : KernameSet.t) (x : A) => + KernameSet.union (term_global_deps (f x)) acc) brs + deps) -> + includes_deps Σ Σ' deps /\ + (forall x, In x brs -> + includes_deps Σ Σ' (term_global_deps (f x))). +Proof. + intros incl; split. + intros kn knin; apply incl. + rewrite knset_in_fold_left. left; auto. + intros br brin. intros kn knin. apply incl. + rewrite knset_in_fold_left. right. + now exists br. +Qed. + +Definition declared_kn Σ kn := + ∥ ∑ decl, lookup_env Σ kn = Some decl ∥. + +Lemma term_global_deps_spec {cf} {Σ : global_env_ext} {Γ t et T} : + wf Σ.1 -> + Σ ;;; Γ |- t : T -> + Σ;;; Γ |- t ⇝ℇ et -> + forall x, KernameSet.In x (term_global_deps et) -> declared_kn Σ x. +Proof. + intros wf wt er. + induction er in T, wt |- * using erases_forall_list_ind; + cbn in *; try solve [constructor]; intros kn' hin; + repeat match goal with + | [ H : KernameSet.In _ KernameSet.empty |- _ ] => + now apply KernameSet.empty_spec in hin + | [ H : KernameSet.In _ (KernameSet.union _ _) |- _ ] => + apply KernameSet.union_spec in hin as [?|?] + end. + - apply inversion_Lambda in wt as (? & ? & ? & ? & ?); eauto. + - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. + - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. + - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. + - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. + - apply inversion_Const in wt as (? & ? & ? & ? & ?); eauto. + eapply KernameSet.singleton_spec in hin; subst. + unshelve eapply declared_constant_to_gen in d; eauto. + red in d. red. sq. rewrite d. intuition eauto. + - apply inversion_Construct in wt as (? & ? & ? & ? & ? & ? & ?); eauto. + destruct kn. + eapply KernameSet.singleton_spec in hin; subst. + destruct d as [[H' _] _]. + unshelve eapply declared_minductive_to_gen in H'; eauto. + red in H'. simpl in *. + red. sq. rewrite H'. intuition eauto. + - apply inversion_Case in wt as (? & ? & ? & ? & [] & ?); eauto. + destruct ci as [kn i']; simpl in *. + eapply KernameSet.singleton_spec in H1; subst. + destruct x1 as [d _]. + unshelve eapply declared_minductive_to_gen in d; eauto. + red in d. + simpl in *. eexists; intuition eauto. + - apply inversion_Case in wt as (? & ? & ? & ? & [] & ?); eauto. + eapply knset_in_fold_left in H1. + destruct H1. eapply IHer; eauto. + destruct H1 as [br [inbr inkn]]. + eapply Forall2_All2 in H0. + assert (All (fun br => ∑ T, Σ ;;; Γ ,,, inst_case_branch_context p br |- br.(bbody) : T) brs). + eapply All2i_All_right. eapply brs_ty. + simpl. intuition auto. eexists ; eauto. + now rewrite -(PCUICCasesContexts.inst_case_branch_context_eq a). + eapply All2_All_mix_left in H0; eauto. simpl in H0. + eapply All2_In_right in H0; eauto. + destruct H0. + destruct X1 as [br' [[T' HT] ?]]. + eauto. + + - eapply KernameSet.singleton_spec in H0; subst. + apply inversion_Proj in wt as (?&?&?&?&?&?&?&?&?&?); eauto. + unshelve eapply declared_projection_to_gen in d; eauto. + destruct d as [[[d _] _] _]. + red in d. eexists; eauto. + + - apply inversion_Proj in wt as (?&?&?&?&?&?&?&?&?); eauto. + + - apply inversion_Fix in wt as (?&?&?&?&?&?&?); eauto. + eapply knset_in_fold_left in hin as [|]. + now eapply KernameSet.empty_spec in H0. + destruct H0 as [? [ina indeps]]. + 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]]]. + eapply Hdef; eauto. + + - apply inversion_CoFix in wt as (?&?&?&?&?&?&?); eauto. + eapply knset_in_fold_left in hin as [|]. + now eapply KernameSet.empty_spec in H0. + destruct H0 as [? [ina indeps]]. + 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]]]. + eapply Hdef; eauto. +Qed. + +Global Remove Hints erases_deps_eval : core. + +Lemma erase_global_erases_deps {Σ} {Σ' : EAst.global_declarations} {Γ t et T} : + wf_ext Σ -> + Σ;;; Γ |- t : T -> + Σ;;; Γ |- t ⇝ℇ et -> + includes_deps Σ Σ' (term_global_deps et) -> + erases_deps Σ Σ' et. +Proof. + intros wf wt er. + induction er in er, t, T, wf, wt |- * using erases_forall_list_ind; + cbn in *; try solve [constructor]; intros Σer; + repeat match goal with + | [ H : includes_deps _ _ (KernameSet.union _ _ ) |- _ ] => + apply includes_deps_union in H as [? ?] + end. + - now apply inversion_Evar in wt. + - constructor. + now apply inversion_Lambda in wt as (? & ? & ? & ? & ?); eauto. + - apply inversion_LetIn in wt as (? & ? & ? & ? & ? & ?); eauto. + constructor; eauto. + - apply inversion_App in wt as (? & ? & ? & ? & ? & ?); eauto. + now constructor; eauto. + - apply inversion_Const in wt as (? & ? & ? & ? & ?); eauto. + specialize (Σer kn). + forward Σer. rewrite KernameSet.singleton_spec //. + pose proof (wf' := wf.1). + unshelve eapply declared_constant_to_gen in d; eauto. + destruct Σer as [[c' [declc' (? & ? & ? & ?)]]|]. + unshelve eapply declared_constant_to_gen in declc'; eauto. + pose proof (declared_constant_inj _ _ d declc'). subst x. + econstructor; eauto. eapply declared_constant_from_gen; eauto. + destruct H as [mib [mib' [declm declm']]]. + unshelve eapply declared_minductive_to_gen in declm; eauto. + red in declm, d. unfold declared_minductive_gen in declm. + rewrite d in declm. noconf declm. + - apply inversion_Construct in wt as (? & ? & ? & ? & ? & ?); eauto. + red in Σer. destruct kn. + setoid_rewrite KernameSetFact.singleton_iff in Σer. + pose (wf' := wf.1). + destruct (Σer _ eq_refl) as [ (? & ? & ? & ? & ? & ?) | (? & ? & ? & ? & ?) ]. + + destruct d as [[]]. + unshelve eapply declared_constant_to_gen in H0; eauto. + unshelve eapply declared_minductive_to_gen in H4; eauto. + red in H4, H. cbn in *. unfold PCUICEnvironment.fst_ctx in *. congruence. + + pose proof H2 as H2'. destruct d. cbn in *. + destruct H3. cbn in *. red in H3. + unshelve eapply declared_minductive_to_gen in H0, H3; eauto. + red in H0, H3. rewrite H0 in H3. invs H3. + destruct H2. + red in H1. + eapply Forall2_nth_error_Some_l in H2 as (? & ? & ?); eauto. pose proof H6 as H6'. destruct H6 as (? & ? & ? & ? & ?); eauto. + eapply Forall2_nth_error_Some_l in H6 as ([] & ? & ? & ?); subst; eauto. + econstructor. + eapply declared_constructor_from_gen; eauto. repeat eapply conj; eauto. + repeat eapply conj; cbn; eauto. eauto. eauto. + - apply inversion_Case in wt as (? & ? & ? & ? & [] & ?); eauto. + destruct ci as [[kn i'] ? ?]; simpl in *. + apply includes_deps_fold in H2 as [? ?]. + pose proof (wf' := wf.1). + specialize (H1 kn). forward H1. + now rewrite KernameSet.singleton_spec. red in H1. destruct H1. + elimtype False. destruct H1 as [cst [declc _]]. + { destruct x1 as [d _]. + unshelve eapply declared_minductive_to_gen in d; eauto. + unshelve eapply declared_constant_to_gen in declc; eauto. + red in d, declc. + unfold declared_constant_gen in declc. rewrite d in declc. noconf declc. } + destruct H1 as [mib [mib' [declm [declm' em]]]]. + pose proof em as em'. destruct em'. + destruct x1 as [x1 hnth]. + unshelve eapply declared_minductive_to_gen in x1, declm; eauto. + red in x1, declm. unfold declared_minductive_gen in declm. rewrite x1 in declm. noconf declm. + eapply Forall2_nth_error_left in H1; eauto. destruct H1 as [? [? ?]]. + eapply erases_deps_tCase; eauto. + apply declared_inductive_from_gen; auto. + split; eauto. split; eauto. + destruct H1. + eapply In_Forall in H3. + eapply All_Forall. eapply Forall_All in H3. + eapply Forall2_All2 in H0. + eapply All2_All_mix_right in H0; eauto. + assert (All (fun br => ∑ T, Σ ;;; Γ ,,, inst_case_branch_context p br |- br.(bbody) : T) brs). + eapply All2i_All_right. eapply brs_ty. + simpl. intuition auto. eexists ; eauto. + now rewrite -(PCUICCasesContexts.inst_case_branch_context_eq a). + ELiftSubst.solve_all. destruct a0 as [T' HT]. eauto. + + - apply inversion_Proj in wt as (?&?&?&?&?&?&?&?&?&?); eauto. + destruct (proj1 d). + pose proof (wf' := wf.1). + specialize (H0 (inductive_mind p.(proj_ind))). forward H0. + now rewrite KernameSet.singleton_spec. red in H0. destruct H0. + elimtype False. destruct H0 as [cst [declc _]]. + { + unshelve eapply declared_constant_to_gen in declc; eauto. + red in declc. destruct d as [[[d _] _] _]. + unshelve eapply declared_minductive_to_gen in d; eauto. + red in d. + unfold declared_constant_gen in declc. rewrite d in declc. noconf declc. } + destruct H0 as [mib [mib' [declm [declm' em]]]]. + assert (mib = x0). + { destruct d as [[[]]]. + unshelve eapply declared_minductive_to_gen in declm, H0; eauto. + red in H0, declm. unfold declared_minductive_gen in declm. rewrite H0 in declm. now noconf declm. } + subst x0. + pose proof em as em'. destruct em'. + eapply Forall2_nth_error_left in H0 as (x' & ? & ?); eauto. + 2:{ apply d. } + simpl in *. + destruct (ind_ctors x1) => //. noconf H3. + set (H1' := H5). destruct H1' as []. + set (d' := d). destruct d' as [[[]]]. + eapply Forall2_All2 in H3. eapply All2_nth_error_Some in H3 as [? [? []]]; tea. + destruct H6 as [Hprojs _]. + eapply Forall2_All2 in Hprojs. eapply All2_nth_error_Some in Hprojs as [? []]; tea. + 2:eapply d. + econstructor; tea. all:eauto. + split => //. 2:split; eauto. + split; eauto. split; eauto. + rewrite -H4. symmetry; apply d. + + - constructor. + apply inversion_Fix in wt as (?&?&?&?&?&?&?); eauto. + eapply All_Forall. eapply includes_deps_fold in Σer as [_ Σer]. + eapply In_Forall in Σer. + eapply Forall_All in Σer. + eapply Forall2_All2 in H. + ELiftSubst.solve_all. + - constructor. + apply inversion_CoFix in wt as (?&?&?&?&?&?&?); eauto. + eapply All_Forall. eapply includes_deps_fold in Σer as [_ Σer]. + eapply In_Forall in Σer. + eapply Forall_All in Σer. + eapply Forall2_All2 in H. + ELiftSubst.solve_all. Unshelve. +Qed. + +Lemma erases_deps_weaken kn d (Σ : global_env) (Σ' : EAst.global_declarations) t : + wf (add_global_decl Σ (kn, d)) -> + erases_deps Σ Σ' t -> + erases_deps (add_global_decl Σ (kn, d)) Σ' t. +Proof. + intros wfΣ er. + assert (wfΣ' : wf Σ) + by now eapply strictly_extends_decls_wf; tea; split => //; eexists [_]. + induction er using erases_deps_forall_ind; try solve [now constructor]. + - inv wfΣ. inv X. + assert (wf Σ) by (inversion H4;econstructor; eauto). + pose proof (H_ := H). unshelve eapply declared_constant_to_gen in H; eauto. + assert (kn <> kn0). + { intros <-. + eapply lookup_env_Some_fresh in H. destruct X1. contradiction. } + eapply erases_deps_tConst with cb cb'; eauto. + eapply declared_constant_from_gen. + red. rewrite /declared_constant_gen /lookup_env lookup_env_cons_fresh //. + red. + red in H1. + destruct (cst_body cb) eqn:cbe; + destruct (E.cst_body cb') eqn:cbe'; auto. + specialize (H3 _ eq_refl). + eapply on_declared_constant in H_; auto. + red in H_. rewrite cbe in H_. simpl in H_. + eapply (erases_weakening_env (Σ := (Σ, cst_universes cb)) + (Σ' := (add_global_decl Σ (kn, d), cst_universes cb))); eauto. + simpl. econstructor; eauto. econstructor; eauto. + split => //; eexists [(kn, d)]; intuition eauto. + - econstructor; eauto. eapply weakening_env_declared_constructor; eauto; tc. + eapply extends_decls_extends, strictly_extends_decls_extends_decls. econstructor; try reflexivity. eexists [(_, _)]; reflexivity. + - econstructor; eauto. + eapply declared_inductive_from_gen. + inv wfΣ. inv X. + assert (wf Σ) by (inversion H5;econstructor; eauto). + unshelve eapply declared_inductive_to_gen in H; eauto. + red. destruct H. split; eauto. + red in H. red. + rewrite -H. simpl. unfold lookup_env; simpl. destruct (eqb_spec (inductive_mind p.1) kn); try congruence. + eapply lookup_env_Some_fresh in H. destruct X1. subst kn; contradiction. + - econstructor; eauto. + inv wfΣ. inv X. + assert (wf Σ) by (inversion H3;econstructor; eauto). + eapply declared_projection_from_gen. + unshelve eapply declared_projection_to_gen in H; eauto. + red. destruct H. split; eauto. + destruct H; split; eauto. + destruct H; split; eauto. + red in H |- *. + rewrite -H. simpl. unfold lookup_env; simpl; destruct (eqb_spec (inductive_mind p.(proj_ind)) kn); try congruence. + eapply lookup_env_Some_fresh in H. subst kn. destruct X1. contradiction. +Qed. + +Lemma lookup_env_ext {Σ kn kn' d d'} : + wf (add_global_decl Σ (kn', d')) -> + lookup_env Σ kn = Some d -> + kn <> kn'. +Proof. + intros wf hl. + eapply lookup_env_Some_fresh in hl. + inv wf. inv X. + destruct (eqb_spec kn kn'); subst; destruct X1; congruence. +Qed. + +Lemma lookup_env_cons_disc {Σ kn kn' d} : + kn <> kn' -> + lookup_env (add_global_decl Σ (kn', d)) kn = lookup_env Σ kn. +Proof. + intros Hk. simpl; unfold lookup_env; simpl. + destruct (eqb_spec kn kn'); congruence. +Qed. + +Lemma elookup_env_cons_disc {Σ kn kn' d} : + kn <> kn' -> + EGlobalEnv.lookup_env ((kn', d) :: Σ) kn = EGlobalEnv.lookup_env Σ kn. +Proof. + intros Hk. simpl. + destruct (eqb_spec kn kn'); congruence. +Qed. + +Lemma global_erases_with_deps_cons kn kn' d d' Σ Σ' : + wf (add_global_decl Σ (kn', d)) -> + global_erased_with_deps Σ Σ' kn -> + global_erased_with_deps (add_global_decl Σ (kn', d)) ((kn', d') :: Σ') kn. +Proof. + intros wf [[cst [declc [cst' [declc' [ebody IH]]]]]|]. + red. inv wf. inv X. + assert (wfΣ : PCUICTyping.wf Σ). + { eapply strictly_extends_decls_wf; split; tea ; eauto. econstructor; eauto. + now eexists [_]. + } + left. + exists cst. split. + eapply declared_constant_from_gen. + unshelve eapply declared_constant_to_gen in declc; eauto. + red in declc |- *. unfold declared_constant_gen, lookup_env in *. + rewrite lookup_env_cons_fresh //. + { eapply lookup_env_Some_fresh in declc. destruct X1. + intros <-; contradiction. } + exists cst'. + pose proof (declc_ := declc). + unshelve eapply declared_constant_to_gen in declc; eauto. + unfold EGlobalEnv.declared_constant. rewrite EGlobalEnv.elookup_env_cons_fresh //. + { eapply lookup_env_Some_fresh in declc. destruct X1. + intros <-; contradiction. } + red in ebody. unfold erases_constant_body. + destruct (cst_body cst) eqn:bod; destruct (E.cst_body cst') eqn:bod' => //. + intuition auto. + eapply (erases_weakening_env (Σ := (_, cst_universes cst)) (Σ' := (add_global_decl Σ (kn', d), cst_universes cst))); eauto. + constructor; eauto. constructor; eauto. + split => //. exists [(kn', d)]; intuition eauto. + eapply on_declared_constant in declc_; auto. + red in declc_. rewrite bod in declc_. eapply declc_. + noconf H0. + eapply erases_deps_cons; eauto. + constructor; auto. + + right. + pose proof (wf_ := wf). inv wf. inv X. + assert (wf Σ) by (inversion H0;econstructor; eauto). + destruct H as [mib [mib' [? [? ?]]]]. + unshelve eapply declared_minductive_to_gen in H; eauto. + exists mib, mib'. intuition eauto. + * eapply declared_minductive_from_gen. + red. red in H. pose proof (lookup_env_ext wf_ H). unfold declared_minductive_gen. + now rewrite lookup_env_cons_disc. + * red. pose proof (lookup_env_ext wf_ H). + now rewrite elookup_env_cons_disc. +Qed. + +Lemma global_erases_with_deps_weaken kn kn' d Σ Σ' : + wf (add_global_decl Σ (kn', d)) -> + global_erased_with_deps Σ Σ' kn -> + global_erased_with_deps (add_global_decl Σ (kn', d)) Σ' kn. +Proof. + intros wf [[cst [declc [cst' [declc' [ebody IH]]]]]|]. + red. inv wf. inv X. left. + assert (wfΣ : PCUICTyping.wf Σ). + { eapply strictly_extends_decls_wf; split; tea ; eauto. econstructor; eauto. + now eexists [_]. + } + exists cst. split. + eapply declared_constant_from_gen. + unshelve eapply declared_constant_to_gen in declc; eauto. + red in declc |- *. + unfold declared_constant_gen, lookup_env in *. + rewrite lookup_env_cons_fresh //. + { eapply lookup_env_Some_fresh in declc. + intros <-. destruct X1. contradiction. } + exists cst'. + unfold EGlobalEnv.declared_constant. + red in ebody. unfold erases_constant_body. + destruct (cst_body cst) eqn:bod; destruct (E.cst_body cst') eqn:bod' => //. + intuition auto. + eapply (erases_weakening_env (Σ := (Σ, cst_universes cst)) (Σ' := (add_global_decl Σ (kn', d), cst_universes cst))). 5:eauto. + split; auto. constructor; eauto. constructor; eauto. + split; auto; exists [(kn', d)]; intuition eauto. + eapply on_declared_constant in declc; auto. + red in declc. rewrite bod in declc. eapply declc. + noconf H0. + apply erases_deps_weaken; auto. + constructor; eauto. constructor; eauto. + + right. destruct H as [mib [mib' [Hm [? ?]]]]. + exists mib, mib'; intuition auto. pose proof (wf_ := wf). + inv wf. inv X. + assert (wf Σ) by (inversion H;econstructor; eauto). + eapply declared_minductive_from_gen. + + unshelve eapply declared_minductive_to_gen in Hm; eauto. + red. unfold declared_minductive_gen, lookup_env in *. + rewrite lookup_env_cons_fresh //. + now epose proof (lookup_env_ext wf_ Hm). +Qed. + +Lemma erase_global_includes X_type (X:X_type.π1) deps decls {normalization_in} prf deps' : + (forall d, KernameSet.In d deps' -> + forall Σ : global_env, abstract_env_rel X Σ -> ∥ ∑ decl, lookup_env Σ d = Some decl ∥) -> + KernameSet.subset deps' deps -> + forall Σ : global_env, abstract_env_rel X Σ -> + let Σ' := erase_global_deps deps X decls (normalization_in:=normalization_in) prf in + includes_deps Σ (fst Σ') deps'. +Proof. + intros ? sub Σ wfΣ; cbn. induction decls in X, H, sub, prf, deps, deps', Σ , wfΣ, normalization_in |- *. + - simpl. intros i hin. elimtype False. + specialize (H i hin) as [[decl Hdecl]]; eauto. + rewrite /lookup_env (prf _ wfΣ) in Hdecl. noconf Hdecl. + - intros i hin. + destruct (abstract_env_wf _ wfΣ) as [wf]. + destruct a as [kn d]. + eapply KernameSet.subset_spec in sub. + edestruct (H i hin) as [[decl Hdecl]]; eauto. unfold lookup_env in Hdecl. + pose proof (eqb_spec i kn). + rewrite (prf _ wfΣ) in Hdecl. move: Hdecl. cbn -[erase_global_deps]. + elim: H0. intros -> [= <-]. + * destruct d as [|]; [left|right]. + { cbn. set (Xpop := abstract_pop_decls X). + unfold erase_constant_decl. + set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _). + epose proof (abstract_env_exists Xpop) as [[Σpop wfpop]]. + epose proof (abstract_env_ext_exists Xmake) as [[Σmake wfmake]]. + exists c. split; auto. + eapply declared_constant_from_gen. + unfold declared_constant_gen, lookup_env; simpl; rewrite (prf _ wfΣ). cbn. + rewrite eq_kername_refl //. + pose proof (sub _ hin) as indeps. + eapply KernameSet.mem_spec in indeps. + unfold EGlobalEnv.declared_constant. + edestruct (H _ hin) as [[decl hd]]; eauto. + eexists; intuition eauto. + rewrite indeps. cbn. + rewrite eq_kername_refl. reflexivity. + cut (strictly_extends_decls Σpop Σ) => [?|]. + eapply (erase_constant_body_correct _ _ _ _ (Σpop , _)); eauto. + rewrite <- (abstract_make_wf_env_ext_correct Xpop (cst_universes c) _ Σpop Σmake wfpop wfmake); eauto. + { now eapply strictly_extends_decls_wf. } + red. simpl. unshelve epose (abstract_pop_decls_correct X decls _ Σ Σpop wfΣ wfpop). + { intros. now eexists. } + split => //. intuition eauto. + exists [(kn, ConstantDecl c)]; intuition eauto. rewrite H0; eauto. + now destruct a. + rewrite indeps. unshelve epose proof (abstract_pop_decls_correct X decls _ Σ Σpop wfΣ wfpop) as [Hpop [Hpop' Hpop'']]. + { intros. now eexists. } + pose (prf' := prf _ wfΣ). + destruct Σ. cbn in *. rewrite Hpop' Hpop'' prf'. rewrite <- Hpop at 1. + eapply (erases_deps_cons Σpop). + rewrite <- Hpop'. apply wf. + rewrite Hpop. rewrite prf' in wf. destruct wf. now rewrite Hpop'' Hpop' in o0. + + pose proof (erase_constant_body_correct' H0). specialize_Σ wfmake. + sq. destruct H1 as [bod [bodty [[Hbod Hebod] Heqdeps]]]. + rewrite (abstract_make_wf_env_ext_correct Xpop (cst_universes c) _ Σpop Σmake wfpop wfmake) in Hbod, Hebod. + eapply (erase_global_erases_deps (Σ := (Σpop, cst_universes c))); simpl; auto. + { constructor; simpl; auto. depelim wf. rewrite Hpop' Hpop'' in o0. + cbn in o0, o. rewrite prf' in o0. rewrite <- Hpop in o0. rewrite Hpop' in o. clear -o o0. + now depelim o0. + depelim wf. rewrite Hpop' in o0. + cbn in o0, o. rewrite prf' in o0. rewrite <- Hpop in o0. clear -o0. depelim o0. + now destruct o. + } + all: eauto. + apply IHdecls; eauto. + { intros. pose proof (abstract_env_wf _ wfpop) as [wf']. + destruct Σpop. cbn in *. clear prf'. subst. + unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wfΣ H2) as [Hpop Hpop']. + { intros. now eexists. } + destruct Σ. cbn in *. subst. + eapply term_global_deps_spec in Hebod; eauto. } + { eapply KernameSet.subset_spec. + intros x hin'. eapply KernameSet.union_spec. right; auto. + now rewrite -Heqdeps. } } + { eexists m, _; intuition eauto. + eapply declared_minductive_from_gen. + simpl. rewrite /declared_minductive /declared_minductive_gen /lookup_env prf; eauto. + simpl. rewrite eq_kername_refl. reflexivity. + specialize (sub _ hin). + eapply KernameSet.mem_spec in sub. + simpl. rewrite sub. + red. cbn. rewrite eq_kername_refl. + reflexivity. + assert (declared_minductive Σ kn m). + { eapply declared_minductive_from_gen. red. unfold declared_minductive_gen, lookup_env. rewrite prf; eauto. cbn. now rewrite eqb_refl. } + eapply on_declared_minductive in H0; tea. + now eapply erases_mutual. } + + * intros ikn Hi. + destruct d as [|]. + ++ simpl. destruct (KernameSet.mem kn deps) eqn:eqkn. + unfold erase_constant_decl. + set (Xpop := abstract_pop_decls X). + set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + pose proof (abstract_env_wf _ wfpop) as [wfΣp]. + unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wfΣ wfpop) as [Hpop [Hpop' Hpop'']]. + { intros. now eexists. } + pose proof (prf _ wfΣ). destruct Σ. cbn in *. subst. + eapply global_erases_with_deps_cons; eauto. + eapply (IHdecls Xpop _ _ _ (KernameSet.singleton i)); auto. + 3:{ eapply KernameSet.singleton_spec => //. } + intros. + eapply KernameSet.singleton_spec in H0. + pose proof (abstract_env_irr _ H1 wfpop). subst. + sq; exists decl; eauto. + eapply KernameSet.subset_spec. intros ? ?. + eapply KernameSet.union_spec. left. + eapply KernameSet.singleton_spec in H0; subst. + now eapply sub. + + cbn. set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + pose proof (abstract_env_wf _ wfpop) as [wfΣp]. + unshelve epose proof (abstract_pop_decls_correct X decls _ Σ Σp wfΣ wfpop) as [Hpop [Hpop' Hpop'']]. + { intros. now eexists. } + pose proof (prf _ wfΣ). destruct Σ. cbn in *. subst. + eapply global_erases_with_deps_weaken. eauto. + eapply IHdecls => //. + 3:now eapply KernameSet.singleton_spec. + intros d ind%KernameSet.singleton_spec. + intros. pose proof (abstract_env_irr _ H0 wfpop). subst. + sq; eexists; eauto. + eapply KernameSet.subset_spec. + intros ? hin'. eapply sub. eapply KernameSet.singleton_spec in hin'. now subst. + + ++ simpl. set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + pose proof (abstract_env_wf _ wfpop) as [wfΣp]. + unshelve epose proof (abstract_pop_decls_correct X decls _ Σ Σp wfΣ wfpop) as [Hpop [Hpop' Hpop'']]. + { intros. now eexists. } + pose proof (prf _ wfΣ). destruct Σ. cbn in *. subst. + destruct (KernameSet.mem kn deps) eqn:eqkn. + { eapply (global_erases_with_deps_cons _ kn (InductiveDecl _) _ Σp); eauto. + eapply (IHdecls Xpop _ _ _ (KernameSet.singleton i)); auto. + 3:{ eapply KernameSet.singleton_spec => //. } + intros. + eapply KernameSet.singleton_spec in H0. subst. + pose proof (abstract_env_irr _ H1 wfpop). subst. + sq; eexists; eauto. + eapply KernameSet.subset_spec. intros ? ?. + eapply KernameSet.singleton_spec in H0; subst. + now eapply sub. } + + { eapply (global_erases_with_deps_weaken _ kn (InductiveDecl _) Σp). eauto. + eapply (IHdecls Xpop _ _ _ (KernameSet.singleton i)) => //. + 3:now eapply KernameSet.singleton_spec. + intros d ind%KernameSet.singleton_spec. + intros. pose proof (abstract_env_irr _ H0 wfpop). subst. + sq; eexists; eauto. + eapply KernameSet.subset_spec. + intros ? hin'. eapply sub. eapply KernameSet.singleton_spec in hin'. now subst. } +Qed. + +Lemma erase_correct (wfl := Ee.default_wcbv_flags) X_type (X : X_type.π1) + univs wfext t v Σ' t' deps decls {normalization_in} prf + (Xext := abstract_make_wf_env_ext X univs wfext) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} + : + forall wt : forall Σ, Σ ∼_ext Xext -> welltyped Σ [] t, + erase X_type Xext [] t wt = t' -> + KernameSet.subset (term_global_deps t') deps -> + erase_global_deps deps X decls (normalization_in:=normalization_in) prf = Σ' -> + (forall Σ : global_env, abstract_env_rel X Σ -> Σ |-p t ⇓ v) -> + forall Σ : global_env_ext, abstract_env_ext_rel Xext Σ -> + exists v', Σ ;;; [] |- v ⇝ℇ v' /\ ∥ fst Σ' ⊢ t' ⇓ v' ∥. +Proof. + intros wt. + intros HΣ' Hsub Ht' ev Σext wfΣex. + pose proof (erases_erase (X := Xext) wt). + rewrite HΣ' in H. + destruct (wt _ wfΣex) as [T wt']. + pose proof (abstract_env_ext_wf _ wfΣex) as [wfΣ]. + specialize (H _ wfΣex). + unshelve epose proof (erase_global_erases_deps (Σ' := fst Σ') wfΣ wt' H _); cycle 2. + rewrite <- Ht'. + eapply erase_global_includes; eauto. + intros. eapply term_global_deps_spec in H; eauto. + now rewrite (abstract_make_wf_env_ext_correct X univs wfext Σ _ H1 wfΣex) in H. + + epose proof (abstract_env_exists X) as [[Σ wfΣX]]. + now rewrite (abstract_make_wf_env_ext_correct X univs wfext Σ _ wfΣX wfΣex). + epose proof (abstract_env_exists X) as [[Σ wfΣX]]. + eapply erases_correct; tea. + - eexists; eauto. + - rewrite (abstract_make_wf_env_ext_correct X univs wfext _ _ wfΣX wfΣex); eauto. +Qed. + +Lemma global_env_ind (P : global_env -> Type) + (Pnil : forall univs retro, P {| universes := univs; declarations := []; retroknowledge := retro |}) + (Pcons : forall (Σ : global_env) d, P Σ -> P (add_global_decl Σ d)) + (Σ : global_env) : P Σ. +Proof. + destruct Σ as [univs Σ]. + induction Σ; cbn. apply Pnil. + now apply (Pcons {| universes := univs; declarations := Σ |} a). +Qed. + +Lemma on_global_env_ind (P : forall Σ : global_env, wf Σ -> Type) + (Pnil : forall univs retro (onu : on_global_univs univs), P {| universes := univs; declarations := []; retroknowledge := retro |} + (onu, globenv_nil _ _ _ _)) + (Pcons : forall (Σ : global_env) kn d (wf : wf Σ) + (Hfresh : fresh_global kn Σ.(declarations)) + (udecl := PCUICLookup.universes_decl_of_decl d) + (onud : on_udecl Σ.(universes) udecl) + (pd : on_global_decl cumulSpec0 (lift_typing typing) + ({| universes := Σ.(universes); declarations := Σ.(declarations); retroknowledge := Σ.(retroknowledge) |}, udecl) kn d), + P Σ wf -> P (add_global_decl Σ (kn, d)) + (fst wf, globenv_decl _ _ Σ.(universes) Σ.(retroknowledge) Σ.(declarations) kn d (snd wf) + {| kn_fresh := Hfresh ; on_udecl_udecl := onud ; on_global_decl_d := pd |})) + (Σ : global_env) (wfΣ : wf Σ) : P Σ wfΣ. +Proof. + destruct Σ as [univs Σ]. destruct wfΣ; cbn in *. + induction o0. apply Pnil. destruct o1. + apply (Pcons {| universes := univs; declarations := Σ |} kn d (o, o0)). + exact IHo0. +Qed. + +Ltac inv_eta := + lazymatch goal with + | [ H : PCUICEtaExpand.expanded _ _ |- _ ] => depelim H + end. + +Lemma leq_term_propositional_sorted_l {Σ Γ v v' u u'} : + wf_ext Σ -> + PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' -> + Σ;;; Γ |- v : tSort u -> + Σ;;; Γ |- v' : tSort u' -> is_propositional u -> + leq_universe (global_ext_constraints Σ) u' u. +Proof. + intros wfΣ leq hv hv' isp. + eapply orb_true_iff in isp as [isp | isp]. + - eapply leq_term_prop_sorted_l; eauto. + - eapply leq_term_sprop_sorted_l; eauto. +Qed. + +Lemma map_erase_irrel X_type X {normalization_in1 normalization_in2 : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ t H1 H2 : + erase_terms X_type X Γ t H1 (normalization_in:=normalization_in1) = erase_terms X_type X Γ t H2 (normalization_in:=normalization_in2). +Proof. + eapply erase_irrel. +Qed. + +Lemma erase_mkApps {X_type X} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ t args H2 Ht Hargs : + (forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> wf_local Σ Γ) -> + (forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> ~ ∥ isErasable Σ Γ (mkApps t args) ∥) -> + erase X_type X Γ (mkApps t args) H2 = + E.mkApps (erase X_type X Γ t Ht) (erase_terms X_type X Γ args Hargs). +Proof. + epose proof (abstract_env_ext_exists X) as [[Σ wfΣX]]. + pose proof (abstract_env_ext_wf X wfΣX) as [wf]. + intros Hwflocal Herasable. induction args in t, H2, Ht, Hargs, Herasable |- *. + - cbn. eapply erase_irrel. + - cbn [mkApps]. + rewrite IHargs; clear IHargs. + all: intros; specialize_Σ H; try pose proof (abstract_env_ext_wf _ H) as [wfH]. + 1: inversion Hargs; eauto. depelim X0. now sq. + 2: eauto. + 2:{ destruct H2. cbn in X0. eapply inversion_mkApps in X0 as (? & ? & ?). + econstructor. eauto. } + etransitivity. simp erase. reflexivity. unfold erase_clause_1. + unfold inspect. unfold erase_clause_1_clause_2. + elim: is_erasableP. + + intros. exfalso. + eapply Herasable; eauto. specialize_Σ wfΣX. destruct p. + cbn. destruct H2. eapply Is_type_app; eauto. + + cbn [erase_terms]. + epose proof (fst (erase_irrel _ _)). cbn. + intros he. f_equal => //. f_equal. + eapply erase_irrel. eapply erase_irrel. + eapply map_erase_irrel. +Qed. + +Lemma map_erase_length X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ t H1 : length (erase_terms X_type X Γ t H1) = length t. +Proof. + induction t; cbn; f_equal; eauto. +Qed. + +Local Hint Constructors expanded : expanded. + +Local Arguments erase_global_deps _ _ _ _ _ : clear implicits. + +(*Lemma lookup_env_erase X_type X deps decls normalization_in prf kn d : + KernameSet.In kn deps -> + forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = Some (InductiveDecl d) -> + EGlobalEnv.lookup_env (erase_global_deps X_type deps X decls normalization_in prf) kn = Some (EAst.InductiveDecl (erase_mutual_inductive_body d)). +Proof. + intros hin Σ wfΣ. pose proof (prf _ wfΣ). unfold lookup_env. rewrite H. clear H. + rewrite /lookup_env. + induction decls in X, Σ , wfΣ ,prf, deps, hin, normalization_in |- *. + - move=> /= //. + - destruct a as [kn' d']. + cbn -[erase_global_deps]. + case: (eqb_spec kn kn'); intros e'; subst. + intros [= ->]. + unfold erase_global_deps. + eapply KernameSet.mem_spec in hin. rewrite hin /=. + now rewrite eq_kername_refl. + intros hl. destruct d'. simpl. + set (Xpop := abstract_pop_decls X). + set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + destruct KernameSet.mem. cbn. + rewrite (negbTE (proj2 (neqb _ _) e')). + eapply IHdecls => //; eauto. eapply KernameSet.union_spec. left => //. + eapply IHdecls => //; eauto. + simpl. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + destruct KernameSet.mem. cbn. + rewrite (negbTE (proj2 (neqb _ _) e')). + eapply IHdecls => //; eauto. + eapply IHdecls => //; eauto. +Qed. +*) + +Lemma lookup_env_erase_decl_None X_type X deps decls normalization_in prf kn : + forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = None -> + EGlobalEnv.lookup_env (erase_global_deps X_type deps X decls normalization_in prf).1 kn = None. +Proof. + intros Σ wfΣ. + pose proof (prf _ wfΣ). + unfold lookup_env. rewrite H. clear H. + rewrite /lookup_env. + induction decls in X, Σ , wfΣ ,prf, deps, normalization_in |- *. + - move=> /= //. + - destruct a as [kn' d']. + cbn -[erase_global_deps]. + case: (eqb_spec kn kn'); intros e'; subst. + intros [= ->]. + cbn. destruct d'. cbn. + case_eq (KernameSet.mem kn' deps); move => hin'; cbn. + + destruct (eqb_spec kn kn') => //. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + now eapply IHdecls. + + intros. set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + now eapply IHdecls. + + intros hin. + case_eq (KernameSet.mem kn' deps); move => hin'; cbn. + destruct (eqb_spec kn kn') => //. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + now eapply IHdecls. + intros. set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + now eapply IHdecls. +Qed. + +Lemma lookup_env_erase_global_diff X_type X deps decls kn kn' d' normalization_in prf : + kn <> kn' -> + exists deps' nin' prf', + KernameSet.Subset deps deps' /\ + EGlobalEnv.lookup_env (erase_global_deps X_type deps X ((kn', d') :: decls) normalization_in prf).1 kn = + EGlobalEnv.lookup_env (erase_global_deps X_type deps' (abstract_pop_decls X) decls nin' prf').1 kn /\ + (KernameSet.In kn (erase_global_deps X_type deps X ((kn', d') :: decls) normalization_in prf).2 -> + KernameSet.In kn (erase_global_deps X_type deps' (abstract_pop_decls X) decls nin' prf').2). +Proof. + intros hd. + simpl. destruct d'. + + destruct KernameSet.mem. cbn. + destruct (eqb_spec kn kn') => //. + do 3 eexists. split; [|split;[reflexivity|]]. knset. + auto. + do 3 eexists. split; [|split;[reflexivity|]]. knset. + auto. + + destruct KernameSet.mem. cbn. + destruct (eqb_spec kn kn') => //. + do 3 eexists. split; [|split;[reflexivity|]]. knset. auto. + do 3 eexists. split; [|split;[reflexivity|]]. knset. auto. +Qed. + +Lemma lookup_env_In_map_fst Σ kn decl : lookup_global Σ kn = Some decl -> In kn (map fst Σ). +Proof. + induction Σ; cbn => //. + case: eqb_spec. + + intros -> [= <-]. now left. + + intros _ hl. eauto. +Qed. + +Lemma erase_constant_decl_deps {X_type : abstract_env_impl} (X : X_type.π1) (cb : constant_body) normalization_in oncb : + forall Σ : global_env, abstract_env_rel X Σ -> + forall kn, KernameSet.In kn (erase_constant_decl X cb normalization_in oncb).2 -> In kn (map fst (declarations Σ)). +Proof. + intros Σ wfΣ kn. + unfold erase_constant_decl. + set (ec := erase_constant_body _ _ _ _). + destruct cb. destruct cst_body0. + 2:{ subst ec; cbn. knset. } + intros hc. + epose proof (erase_constant_body_correct' (@eq_refl _ (EAst.cst_body ec.1))). + subst ec. + set (ec := erase_constant_body _ _ _ _) in *. + set (X' := abstract_make_wf_env_ext _ _ _) in *. + pose proof (abstract_env_ext_exists X') as [[Σ' H']]. + pose proof (abstract_env_ext_wf _ H') as [wfΣ']. + specialize (H _ H') as [[t0 [T []]]]. + rewrite -e in hc. + destruct p as [hty her]. + epose proof (term_global_deps_spec wfΣ' hty her _ hc). + epose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ wfΣ H'). cbn in H0. subst Σ'. + red in H. destruct H as [[decl hl]]. + now eapply lookup_env_In_map_fst in hl. +Qed. + +Lemma in_erase_global_deps_acc X_type X deps decls kn normalization_in prf : + KernameSet.In kn (erase_global_deps X_type deps X decls normalization_in prf).2 -> + KernameSet.In kn deps \/ In kn (map fst decls). +Proof. + induction decls in X, prf, deps, normalization_in |- *. + cbn; auto. destruct a as [kn' []]. + + cbn. set (ec := erase_constant_decl _ _ _ _). + set (eg := erase_global_deps _ _ _ _ _ _). + set (eg' := erase_global_deps _ _ _ _ _ _). + case_eq (KernameSet.mem kn' deps). + * cbn; intros. specialize (IHdecls _ _ _ _ H0). + destruct IHdecls as [Hu|Hm]; auto. + eapply KernameSet.union_spec in Hu. destruct Hu; auto. + unfold ec in H1; cbn in H1. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + eapply erase_constant_decl_deps in H1; tea. + pose proof (abstract_env_exists X) as [[ΣX HX]]. + unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ HX wfpop). + intros. rewrite prf => //. now eexists. + now destruct H2 as [<- _]. + * intros hdeps hin. + eapply IHdecls in hin. intuition auto. + + cbn; set (eg := erase_global_deps _ _ _ _ _ _); + set (eg' := erase_global_deps _ _ _ _ _ _). + case_eq (KernameSet.mem kn' deps). + * cbn; intros. specialize (IHdecls _ _ _ _ H0). + destruct IHdecls as [Hu|Hm]; auto. + * intros hdeps hin. + eapply IHdecls in hin. intuition auto. +Qed. + +Lemma wf_fresh_globals {cf : checker_flags} (Σ : global_env) : wf Σ -> EnvMap.EnvMap.fresh_globals Σ.(declarations). +Proof. + destruct Σ as [univs Σ]; cbn. + move=> [] onu; cbn. induction 1; try destruct o; constructor; auto; constructor; eauto. +Qed. + +Lemma lookup_env_erase_decl X_type X deps decls normalization_in prf kn decl : + forall Σ : global_env, abstract_env_rel X Σ -> lookup_env Σ kn = Some decl -> + let er := erase_global_deps X_type deps X decls normalization_in prf in + KernameSet.In kn er.2 -> + match decl with + | ConstantDecl cb => + exists (X' : X_type.π1) nin pf, + (EGlobalEnv.lookup_env er.1 kn = + Some (EAst.ConstantDecl (fst (erase_constant_decl X' cb nin pf)))) /\ + (forall Σ Σ', Σ ∼ X -> Σ' ∼ X' -> ∥ strictly_extends_decls Σ' Σ ∥) + | InductiveDecl mib => + EGlobalEnv.lookup_env er.1 kn = + Some (EAst.InductiveDecl (erase_mutual_inductive_body mib)) + end. +Proof. + intros Σ wfΣ. pose proof (prf _ wfΣ). + unfold lookup_env. rewrite H. clear H. + rewrite /lookup_env. + induction decls in X, Σ , wfΣ ,prf, deps, decl, normalization_in |- *. + - move=> /= //. + - destruct a as [kn' d']. + cbn -[erase_global_deps]. + case: (eqb_spec kn kn'); intros e'; subst. + intros [= ->]. + destruct decl. + + cbn; set (ec := erase_constant_decl _ _ _ _); + set (eg := erase_global_deps _ _ _ _ _ _); + set (eg' := erase_global_deps _ _ _ _ _ _). + case_eq (KernameSet.mem kn' deps). + move/KernameSet.mem_spec => hin'; cbn. + ++ intros hin. cbn; rewrite eqb_refl; + do 3 eexists. split; [reflexivity|]. + intros ? ? hx hpop. + unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ hx hpop). intros. rewrite prf => //. now eexists. + destruct H as [? []]. unfold strictly_extends_decls. rewrite H. sq; split => //. clear ec eg eg' hin. specialize (prf _ hx). rewrite prf. + now eexists [_]. + ++ intros kn'deps kn'eg'. + eapply in_erase_global_deps_acc in kn'eg'. + destruct kn'eg'. eapply KernameSet.mem_spec in H. congruence. + elimtype False. clear -prf wfΣ H. + specialize (prf _ wfΣ). + pose proof (abstract_env_wf _ wfΣ) as []. eapply wf_fresh_globals in X0. + rewrite prf in X0. depelim X0. + apply EnvMap.EnvMap.fresh_global_iff_not_In in H0. contradiction. + + cbn; + set (eg := erase_global_deps _ _ _ _ _ _); + set (eg' := erase_global_deps _ _ _ _ _ _). + case_eq (KernameSet.mem kn' deps). + move/KernameSet.mem_spec => hin'; cbn. + ++ intros hin. cbn; rewrite eqb_refl; + do 3 eexists. + ++ intros kn'deps kn'eg'. + eapply in_erase_global_deps_acc in kn'eg'. + destruct kn'eg'. eapply KernameSet.mem_spec in H. congruence. + elimtype False. clear -prf wfΣ H. + specialize (prf _ wfΣ). + pose proof (abstract_env_wf _ wfΣ) as []. eapply wf_fresh_globals in X0. + rewrite prf in X0. depelim X0. + apply EnvMap.EnvMap.fresh_global_iff_not_In in H0. contradiction. + + intros hl hin. + epose proof (lookup_env_erase_global_diff X_type X deps decls kn kn' d' _ _ e'). + destruct H as [deps' [nin' [prf' [sub [hl' hin']]]]]. + eapply hin' in hin. + erewrite hl'. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + specialize (IHdecls Xpop deps'). specialize (IHdecls nin' prf' decl _ wfpop hl hin). + destruct decl; intuition auto. + destruct IHdecls as [X' [nin [pf []]]]. exists X', nin, pf. split => //. + intros. + unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ H1 wfpop). intros. rewrite prf => //. now eexists. + specialize (H0 _ _ wfpop H2). destruct H0; sq. move: X0; rewrite /strictly_extends_decls. destruct H3 as [? []]. + rewrite H4 H0 H3. rewrite -(prf' _ wfpop). pose proof (prf _ H1). rewrite -H0 in H5. rewrite H5. + intros []; split => //. destruct s as [Σ'' eq]. eexists ((kn', d') :: Σ''). now rewrite eq. +Qed. + +Lemma fresh_global_app kn Σ Σ' : fresh_global kn (Σ ++ Σ') -> + fresh_global kn Σ /\ fresh_global kn Σ'. +Proof. + induction Σ; cbn; intuition auto. + - constructor. + - depelim H. constructor => //. + now eapply Forall_app in H0. + - depelim H. + now eapply Forall_app in H0. +Qed. + +Lemma lookup_global_app_wf Σ Σ' kn : EnvMap.EnvMap.fresh_globals (Σ ++ Σ') -> + forall decl, lookup_global Σ' kn = Some decl -> lookup_global (Σ ++ Σ') kn = Some decl. +Proof. + induction Σ in kn |- *; cbn; auto. + intros fr; depelim fr. + intros decl hd; cbn. + destruct (eqb_spec kn kn0). + eapply PCUICWeakeningEnv.lookup_global_Some_fresh in hd. + subst. now eapply fresh_global_app in H as [H H']. + now eapply IHΣ. +Qed. + +Lemma strictly_extends_decls_lookup {Σ Σ'} : + wf Σ' -> + strictly_extends_decls Σ Σ' -> + forall kn d, lookup_env Σ kn = Some d -> + lookup_env Σ' kn = Some d. +Proof. + intros. + destruct X0 as [? []]. + rewrite /lookup_env in H *. rewrite e0. + erewrite lookup_global_app_wf. reflexivity. 2:eauto. + eapply wf_fresh_globals in X. now rewrite -e0. +Qed. + +Lemma in_deps_in_erase_global_deps X_type X deps decls normalization_in prf kn : + KernameSet.In kn deps -> + KernameSet.In kn (erase_global_deps X_type deps X decls normalization_in prf).2. +Proof. + induction decls in X, deps, normalization_in, prf |- *. + - intros. now cbn. + - intros. cbn. + pose proof (abstract_env_exists X) as [[Σ H']]. + set (Xpop := abstract_pop_decls X). + epose proof (abstract_env_exists Xpop) as [[Σp wfpop]]. + unshelve epose proof (abstract_pop_decls_correct _ decls _ _ _ H' wfpop). intros. rewrite prf => //. now eexists. + destruct H0 as [Hd [Hu Hr]]. + destruct a. destruct g. + + set (eg := erase_global_deps _ _ _ _ _ _); + set (eg' := erase_global_deps _ _ _ _ _ _). + case_eq (KernameSet.mem k deps). + * cbn. intros hm. + destruct (eqb_spec k kn). + ++ subst k. eapply IHdecls. knset. + ++ eapply IHdecls. + knset. + * intros hnin. eapply IHdecls; tea. + + set (eg := erase_global_deps _ _ _ _ _ _); + set (eg' := erase_global_deps _ _ _ _ _ _). + case_eq (KernameSet.mem k deps). + * cbn. intros hm. + destruct (eqb_spec k kn). + ++ subst k. eapply KernameSet.mem_spec in hm. + now eapply IHdecls. + ++ eapply IHdecls. tea. + * intros hnin. eapply IHdecls; tea. +Qed. + +Fixpoint filter_deps deps decls := + match decls with + | [] => [] + | (kn, decl) :: decls => + if KernameSet.mem kn deps then + match decl with + | EAst.ConstantDecl cst => + (kn, decl) :: filter_deps (KernameSet.union deps (constant_decls_deps cst)) decls + | _ => (kn, decl) :: filter_deps deps decls + end + else filter_deps deps decls + end. + +Lemma erase_global_deps_erase_global (X_type : abstract_env_impl) (X:X_type.π1) deps decls {normalization_in} prf : + (@erase_global_deps X_type deps X decls normalization_in prf).1 = + filter_deps deps (@erase_global X_type X decls normalization_in prf). +Proof. + symmetry. + generalize normalization_in at 1. + generalize prf at 1. + intros prf' normalization_in'. + induction decls in normalization_in, normalization_in', prf, prf', X, deps |- *. + - cbn. auto. + - destruct a as [kn []]. + cbn. + + destruct (KernameSet.mem kn deps). + destruct c as [? [] ? ?]. + set (ec := erase_constant_decl _ _ _ _). + set (ec' := erase_constant_decl _ _ _ _). + assert (ec = ec') as <-. + { unfold ec, ec'. eapply erase_constant_decl_irrel. red. cbn. + intros. pose proof (abstract_env_irr _ H H0). subst. red. split; auto. congruence. } + cbn; f_equal. set (d := KernameSet.union _ _). eapply IHdecls. + cbn. f_equal. eapply IHdecls. + eapply IHdecls. + + cbn; destruct (KernameSet.mem kn deps). + cbn. f_equal. eapply IHdecls. + eapply IHdecls. +Qed. + + +Lemma filter_deps_filter {efl : EWellformed.EEnvFlags} deps Σ : + EWellformed.wf_glob Σ -> + (* (forall kn, KernameSet.In kn deps -> In kn (map fst Σ)) -> *) + filter_deps deps Σ = filter (flip KernameSet.mem (global_deps Σ deps) ∘ fst) Σ. +Proof. + induction Σ in deps |- *. + - now cbn. + - destruct a as [kn [[[]]|]]. + + cbn. unfold flip. + destruct (knset_mem_spec kn deps). + * case: (knset_mem_spec kn _). + ** intros hin' hwf. f_equal. depelim hwf. now eapply IHΣ. + ** intros nhin. + elim: nhin. + rewrite global_deps_union. + rewrite KernameSet.union_spec. left. + now eapply in_global_deps. + * intros hwf. + case: (knset_mem_spec kn _); intros hin. + ** elimtype False. + depelim hwf. + eapply in_global_deps_fresh in hin => //. + ** depelim hwf. rewrite -IHΣ => //. + + intros wf; depelim wf. + cbn. unfold flip. + * case: (knset_mem_spec kn _). + ** intros hin'. + case: (knset_mem_spec kn _) => hin''. + ++ f_equal. now eapply IHΣ. + ++ elim: hin''. + rewrite !global_deps_union. + rewrite KernameSet.union_spec. left. + now eapply in_global_deps. + ** intros hwf. + case: (knset_mem_spec kn _); intros hin. + *** elimtype False. + eapply in_global_deps_fresh in hin => //. + *** rewrite -IHΣ => //. + + intros wf; depelim wf. + cbn. unfold flip. + * case: (knset_mem_spec kn _). + ** intros hin'. + case: (knset_mem_spec kn _) => hin''. + ++ f_equal. rewrite IHΣ //. eapply filter_ext. + intros ?. unfold flip. rewrite global_deps_union /=. + rewrite global_deps_empty. + now rewrite KernameSetFact.union_b KernameSetFact.empty_b orb_false_r. + ++ elim: hin''. + rewrite !global_deps_union. + rewrite KernameSet.union_spec. left. + now eapply in_global_deps. + ** intros hwf. + case: (knset_mem_spec kn _); intros hin. + *** elimtype False. + eapply in_global_deps_fresh in hin => //. + *** rewrite -IHΣ => //. +Qed. + +Lemma erase_global_declared_constructor X_type X ind c mind idecl cdecl deps decls normalization_in prf: + forall Σ : global_env, abstract_env_rel X Σ -> declared_constructor Σ (ind, c) mind idecl cdecl -> + KernameSet.In ind.(inductive_mind) deps -> + EGlobalEnv.declared_constructor (erase_global_deps X_type deps X decls normalization_in prf).1 (ind, c) + (erase_mutual_inductive_body mind) (erase_one_inductive_body idecl) + (EAst.mkConstructor cdecl.(cstr_name) cdecl.(cstr_arity)). +Proof. + intros Σ wfΣ [[]] Hin. pose (abstract_env_wf _ wfΣ). sq. + cbn in *. split. split. + - unshelve eapply declared_minductive_to_gen in H; eauto. + red in H |- *. eapply (lookup_env_erase_decl _ _ _ _ _ _ _ (InductiveDecl mind)); eauto. + eapply in_deps_in_erase_global_deps; tea. + + - cbn. now eapply map_nth_error. + - cbn. erewrite map_nth_error; eauto. +Qed. + +Import EGlobalEnv. +Lemma erase_global_deps_fresh X_type kn deps X decls normalization_in heq : + let Σ' := erase_global_deps X_type deps X decls normalization_in heq in + PCUICAst.fresh_global kn decls -> + fresh_global kn Σ'.1. +Proof. + cbn. + revert deps X normalization_in heq. + induction decls; [cbn; auto|]. + - intros. red. constructor. + - destruct a as [kn' d]. intros. depelim H. + cbn in H, H0. + destruct d as []; simpl; destruct KernameSet.mem. + + cbn [EGlobalEnv.closed_env forallb]. cbn. + constructor => //. eapply IHdecls => //. + + eapply IHdecls => //. + + constructor; auto. + eapply IHdecls => //. + + eapply IHdecls => //. +Qed. + +Lemma erase_global_fresh X_type kn X decls normalization_in heq : + let Σ' := @erase_global X_type X decls normalization_in heq in + PCUICAst.fresh_global kn decls -> + fresh_global kn Σ'. +Proof. + cbn. + revert X normalization_in heq. + induction decls; [cbn; auto|]. + - intros. red. constructor. + - destruct a as [kn' d]. intros. depelim H. + cbn in H, H0. cbn. + destruct d as []; simpl. + + cbn [EGlobalEnv.closed_env forallb]. cbn. + constructor => //. eapply IHdecls => //. + + constructor; auto. + eapply IHdecls => //. +Qed. + +From MetaCoq.Erasure Require Import EEtaExpandedFix. + +Lemma erase_brs_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ p ts wt : + erase_brs X_type X Γ p ts wt = + map_All (fun br wt => (erase_context (bcontext br), erase X_type X _ (bbody br) wt)) ts wt. +Proof. + funelim (map_All _ ts wt); cbn; auto. + f_equal => //. f_equal => //. apply erase_irrel. + rewrite -H. eapply erase_irrel. +Qed. + +Lemma erase_fix_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ ts wt : + erase_fix X_type X Γ ts wt = map_All (fun d wt => + let dbody' := erase X_type X _ (dbody d) (fun Σ abs => proj2 (wt Σ abs)) in + let dbody' := if isBox dbody' then + match d.(dbody) with + | tLambda na _ _ => E.tLambda (binder_name na) E.tBox + | _ => dbody' + end else dbody' + in + {| E.dname := d.(dname).(binder_name); E.rarg := d.(rarg); E.dbody := dbody' |}) ts wt. +Proof. + funelim (map_All _ ts wt); cbn; auto. + f_equal => //. f_equal => //. + rewrite (fst (erase_irrel _ _) _ _ _ _ (fun (Σ : global_env_ext) (abs : abstract_env_ext_rel X Σ) => + (map_All_obligation_1 x xs h Σ abs).p2)). + destruct erase => //. + rewrite -H. eapply erase_irrel. +Qed. + +Lemma erase_cofix_eq X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ ts wt : + erase_cofix X_type X Γ ts wt = map_All (fun d wt => + let dbody' := erase X_type X _ (dbody d) wt in + {| E.dname := d.(dname).(binder_name); E.rarg := d.(rarg); E.dbody := dbody' |}) ts wt. +Proof. + funelim (map_All _ ts wt); cbn; auto. + f_equal => //. f_equal => //. + apply erase_irrel. + rewrite -H. eapply erase_irrel. +Qed. + +Lemma isConstruct_erase X_type X {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ t wt : + ~ (PCUICEtaExpand.isConstruct t || PCUICEtaExpand.isFix t || PCUICEtaExpand.isRel t) -> + ~ (isConstruct (erase X_type X Γ t wt) || isFix (erase X_type X Γ t wt) || isRel (erase X_type X Γ t wt)). +Proof. + apply (erase_elim X_type X + (fun Γ t wt e => ~ (PCUICEtaExpand.isConstruct t || PCUICEtaExpand.isFix t || PCUICEtaExpand.isRel t) -> ~ (isConstruct e || isFix e || isRel e)) + (fun Γ l awt e => True) + (fun Γ p l awt e => True) + (fun Γ l awt e => True) + (fun Γ l awt e => True)); intros => //. bang. +Qed. + +Lemma apply_expanded Σ Γ Γ' t t' : + expanded Σ Γ t -> Γ = Γ' -> t = t' -> expanded Σ Γ' t'. +Proof. intros; now subst. Qed. + +Lemma isLambda_inv t : isLambda t -> exists na ty bod, t = tLambda na ty bod. +Proof. destruct t => //; eauto. Qed. + +Lemma erases_deps_erase (cf := config.extraction_checker_flags) + {X_type X} univs + (wfΣ : forall Σ, (abstract_env_rel X Σ) -> ∥ wf_ext (Σ, univs) ∥) decls {normalization_in} prf + (X' := abstract_make_wf_env_ext X univs wfΣ) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} + Γ t + (wt : forall Σ : global_env_ext, abstract_env_ext_rel X' Σ -> welltyped Σ Γ t) : + let et := erase X_type X' Γ t wt in + let deps := EAstUtils.term_global_deps et in + forall Σ, (abstract_env_rel X Σ) -> + erases_deps Σ (erase_global_deps X_type deps X decls normalization_in prf).1 et. +Proof. + intros et deps Σ wf. + pose proof (abstract_env_ext_exists X') as [[Σ' wfΣ']]. + pose proof (wt _ wfΣ'). destruct H. pose proof (wfΣ _ wf) as [w]. + rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ') in X0. + eapply (erase_global_erases_deps w); tea. + eapply (erases_erase (X := X') (Γ := Γ)). + now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). + eapply erase_global_includes. + intros. rewrite (abstract_env_irr _ H0 wf). + eapply term_global_deps_spec in H; eauto. + assumption. + eapply (erases_erase (X := X') (Γ := Γ)). + now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). + eapply KernameSet.subset_spec. reflexivity. + now cbn. +Qed. + +Lemma erases_deps_erase_weaken (cf := config.extraction_checker_flags) + {X_type X} univs + (wfΣ : forall Σ, (abstract_env_rel X Σ) -> ∥ wf_ext (Σ, univs) ∥) decls {normalization_in} prf + (X' := abstract_make_wf_env_ext X univs wfΣ) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} + Γ t + (wt : forall Σ : global_env_ext, abstract_env_ext_rel X' Σ -> welltyped Σ Γ t) + deps : + let et := erase X_type X' Γ t wt in + let tdeps := EAstUtils.term_global_deps et in + forall Σ, (abstract_env_rel X Σ) -> + erases_deps Σ (erase_global_deps X_type (KernameSet.union deps tdeps) X decls normalization_in prf).1 et. +Proof. + intros et tdeps Σ wf. + pose proof (abstract_env_ext_exists X') as [[Σ' wfΣ']]. + pose proof (wt _ wfΣ'). destruct H. pose proof (wfΣ _ wf) as [w]. + rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ') in X0. + eapply (erase_global_erases_deps w); tea. + eapply (erases_erase (X := X') (Γ := Γ)). + now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). + eapply erase_global_includes. + intros. rewrite (abstract_env_irr _ H0 wf). + eapply term_global_deps_spec in H; eauto. + assumption. + eapply (erases_erase (X := X') (Γ := Γ)). + now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). + eapply KernameSet.subset_spec. intros x hin. + eapply KernameSet.union_spec. now right. + now cbn. +Qed. + +Lemma eta_expand_erase {X_type X} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Σ' {Γ t} + (wt : forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> welltyped Σ Γ t) Γ' : + forall Σ : global_env_ext, abstract_env_ext_rel X Σ -> + PCUICEtaExpand.expanded Σ Γ' t -> + erases_global Σ Σ' -> + expanded Σ' Γ' (erase X_type X Γ t wt). +Proof. + intros Σ wfΣ exp deps. + pose proof (abstract_env_ext_wf _ wfΣ) as [wf]. + eapply expanded_erases. apply wf. + eapply erases_erase; eauto. assumption. + pose proof (wt _ wfΣ). destruct H as [T ht]. + eapply erases_global_erases_deps; tea. + eapply erases_erase; eauto. +Qed. + +Lemma erase_global_closed X_type X deps decls {normalization_in} prf : + let X' := erase_global_deps X_type deps X decls normalization_in prf in + EGlobalEnv.closed_env X'.1. +Proof. + revert X normalization_in prf deps. + induction decls; [cbn; auto|]. + intros X normalization_in prf deps. + destruct a as [kn d]. + destruct d as []; simpl; destruct KernameSet.mem; + unfold erase_constant_decl; + set (Xpop := abstract_pop_decls X); + try (set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _); + assert (normalization_in_Xmake : forall Σ : global_env_ext, wf_ext Σ -> Σ ∼_ext Xmake -> NormalizationIn Σ) + by (subst Xmake; unshelve eapply (normalization_in 0); revgoals; try reflexivity; cbn; lia)). + + cbn [EGlobalEnv.closed_env forallb]. cbn. + rewrite [forallb _ _](IHdecls) // andb_true_r. + rewrite /test_snd /EGlobalEnv.closed_decl /=. + destruct c as [ty [] univs]; cbn. + unfold erase_constant_decl. + set (obl := ErasureFunction.erase_constant_body_obligation_1 _ _ _ _ _). + unfold erase_constant_body. cbn. clearbody obl. cbn in obl. + 2:auto. + pose proof (abstract_env_ext_exists Xmake) as [[Σmake wfmake]]. + pose proof (abstract_env_ext_wf _ wfmake) as [[?]]. + unshelve epose proof (erases_erase (X := Xmake) (obl eq_refl) _ _) as H; eauto. + eapply erases_closed in H. + 1: edestruct erase_irrel as [-> _]; eassumption. + cbn. destruct (obl eq_refl _ wfmake). clear H. + now eapply PCUICClosedTyp.subject_closed in X0. + + eapply IHdecls => //. + + cbn [EGlobalEnv.closed_env forallb]. + eapply IHdecls => //. + + eapply IHdecls => //. + Unshelve. eauto. +Qed. + +Import EWellformed. + +Section wffix. + Import EAst. + Fixpoint wf_fixpoints (t : term) : bool := + match t with + | tRel i => true + | tEvar ev args => List.forallb (wf_fixpoints) args + | tLambda N M => wf_fixpoints M + | tApp u v => wf_fixpoints u && wf_fixpoints v + | tLetIn na b b' => wf_fixpoints b && wf_fixpoints b' + | tCase ind c brs => + let brs' := forallb (wf_fixpoints ∘ snd) brs in + wf_fixpoints c && brs' + | tProj p c => wf_fixpoints c + | tFix mfix idx => + (idx isLambda d.(dbody) && wf_fixpoints d.(dbody)) mfix + | tCoFix mfix idx => + (idx true + | tConstruct ind c _ => true + | tVar _ => true + | tBox => true + | tPrim _ => true + end. + +End wffix. + +Lemma erases_deps_wellformed (cf := config.extraction_checker_flags) (efl := all_env_flags) {Σ} {Σ'} et : + erases_deps Σ Σ' et -> + forall n, ELiftSubst.closedn n et -> + wf_fixpoints et -> + wellformed Σ' n et. +Proof. + intros ed. + induction ed using erases_deps_forall_ind; intros => //; + try solve [cbn in *; unfold wf_fix in *; rtoProp; intuition eauto; solve_all]. + - cbn. red in H0. rewrite H0 //. + - cbn -[lookup_constructor]. + cbn. now destruct H0 as [[-> ->] ->]. + - cbn in *. move/andP: H5 => [] cld clbrs. + cbn. apply/andP; split. apply/andP; split. + * now destruct H0 as [-> ->]. + * now move/andP: H6. + * move/andP: H6; solve_all. + - cbn -[lookup_projection] in *. apply/andP; split; eauto. + now rewrite (declared_projection_lookup H0). +Qed. + +Lemma erases_wf_fixpoints Σ Γ t t' : Σ;;; Γ |- t ⇝ℇ t' -> + ErasureProperties.wellformed Σ t -> wf_fixpoints t'. +Proof. + induction 1 using erases_forall_list_ind; cbn; auto; try solve [rtoProp; repeat solve_all]. + - move/andP => []. rewrite (All2_length X) => -> /=. unfold test_def in *. + eapply Forall2_All2 in H. + eapply All2_All2_mix in X; tea. solve_all. + destruct b0; eapply erases_isLambda in H1; tea. + - move/andP => []. rewrite (All2_length X) => -> /=. + unfold test_def in *. solve_all. +Qed. + +Lemma erase_wf_fixpoints (efl := all_env_flags) {X_type X} univs wfΣ {Γ t} wt + (X' := abstract_make_wf_env_ext X univs wfΣ) {normalization_in} : + let t' := erase X_type X' (normalization_in:=normalization_in) Γ t wt in + wf_fixpoints t'. +Proof. + cbn. pose proof (abstract_env_ext_exists X') as [[Σ' wf']]. + pose proof (abstract_env_exists X) as [[Σ wf]]. + eapply erases_wf_fixpoints. + eapply erases_erase; eauto. + specialize (wt _ wf'). destruct (wfΣ _ wf). + unshelve eapply welltyped_wellformed in wt; eauto. + now rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wf'). +Qed. + +Lemma erases_global_erase_global (efl := all_env_flags) {X_type X} Σ decls {normalization_in} prf : + Σ ∼ X -> + erases_global Σ (@erase_global X_type X decls normalization_in prf). +Proof. + intros h. pose proof (prf _ h). red. rewrite H. clear H. + induction decls in X, Σ, prf, normalization_in, h |- *. + - cbn. constructor. + - cbn. + set (Xpop := abstract_pop_decls X). + destruct (abstract_env_exists Xpop) as [[Σpop hpop]]. + unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ h hpop). + intros; eexists; now eapply prf. + destruct H as [? []]. + destruct a as [kn []]. + * constructor. cbn. + unfold erase_constant_decl. + set (Σd := {| universes := _ |}). + assert (Σd = Σpop). + { subst Σd. rewrite -H H0 H1. now destruct Σpop. } + subst Σd. + set (ext := abstract_make_wf_env_ext _ _ _). + pose proof (abstract_env_ext_exists ext) as [[Σ' wf']]. + epose proof (abstract_make_wf_env_ext_correct (abstract_pop_decls X) (cst_universes c) _ _ _ hpop wf'). + subst Σ'. + cbn. pose proof (abstract_env_wf _ hpop) as []. + eapply (erase_constant_body_correct _ _ _ _ (Σpop, cst_universes c)) => //. + now rewrite H2. rewrite H2. split => //. now exists []. + rewrite H0 H1. + now eapply IHdecls. + * pose proof (abstract_env_wf _ h) as [wfΣ]. + destruct wfΣ. rewrite (prf _ h) in o0. + depelim o0. depelim o1. + constructor. + eapply erases_mutual. eapply on_global_decl_d. + rewrite H0 H1. + now eapply IHdecls. +Qed. + +Lemma erase_global_wellformed (efl := all_env_flags) {X_type X} decls {normalization_in} prf univs wfΣ {Γ t} wt + (X' := abstract_make_wf_env_ext X univs wfΣ) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} + (t' := erase X_type X' Γ t wt) : + wellformed (@erase_global X_type X decls normalization_in prf) #|Γ| t'. +Proof. + cbn. + epose proof (@erases_erase X_type X' normalization_in' _ _ wt). + pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. + pose proof (abstract_env_ext_exists X') as [[Σ' wf']]. + pose proof (abstract_env_ext_wf _ wf') as [?]. + eapply (erases_wellformed (wt _ wf')) in H; eauto; tea. + eapply erases_global_all_deps. apply X0. + epose proof (abstract_make_wf_env_ext_correct _ _ wfΣ _ _ wf wf'). subst Σ'. + now eapply erases_global_erase_global. +Qed. + +Lemma erase_wellformed_deps (efl := all_env_flags) {X_type X} decls {normalization_in} prf univs wfΣ {Γ t} wt + (X' := abstract_make_wf_env_ext X univs wfΣ) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} + (t' := erase X_type X' Γ t wt) : + wellformed (erase_global_deps X_type (term_global_deps t') X decls normalization_in prf).1 #|Γ| t'. +Proof. + cbn. + epose proof (@erases_deps_erase X_type X univs wfΣ decls normalization_in prf _ Γ t wt). + pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. + pose proof (abstract_env_ext_exists X') as [[Σ' wf']]. + pose proof (abstract_env_ext_wf _ wf') as [[?]]. + epose proof (erases_deps_wellformed _ H #|Γ|). + apply H0. + eapply (erases_closed _ Γ). + destruct (wt _ wf'). + cbn in X. destruct (wfΣ _ wf). + eapply PCUICClosedTyp.subject_closed in X0. eassumption. + eapply erases_erase; eauto. + eapply erase_wf_fixpoints. Unshelve. eauto. +Qed. + +Lemma erase_wellformed_weaken (efl := all_env_flags) {X_type X} decls {normalization_in} prf univs wfΣ {Γ t} wt +(X' := abstract_make_wf_env_ext X univs wfΣ) +{normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} +deps: +let t' := erase X_type X' Γ t wt in + wellformed (erase_global_deps X_type (KernameSet.union deps (term_global_deps t')) X decls normalization_in prf).1 #|Γ| t'. +Proof. + set (t' := erase _ _ _ _ _). cbn. + epose proof (@erases_deps_erase_weaken _ X univs wfΣ decls _ prf _ Γ t wt deps). + pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. + pose proof (abstract_env_ext_exists X') as [[Σ' wf']]. + pose proof (abstract_env_ext_wf _ wf') as [[?]]. + epose proof (erases_deps_wellformed _ H #|Γ|). + apply H0. + eapply (erases_closed _ Γ). + - destruct (wt _ wf'). + destruct (wfΣ _ wf). + eapply PCUICClosedTyp.subject_closed in X0. eassumption. + - eapply erases_erase; eauto. + - eapply erase_wf_fixpoints. + Unshelve. eauto. +Qed. + +Lemma erase_constant_body_correct'' {X_type X} {cb} {decls normalization_in prf} {univs wfΣ} +(X' := abstract_make_wf_env_ext X univs wfΣ) +{normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} +{onc : forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> ∥ on_constant_decl (lift_typing typing) Σ' cb ∥} {body} deps : + EAst.cst_body (fst (erase_constant_body X_type X' cb onc)) = Some body -> + forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> + ∥ ∑ t T, (Σ' ;;; [] |- t : T) * (Σ' ;;; [] |- t ⇝ℇ body) * + (term_global_deps body = snd (erase_constant_body X_type X' cb onc)) * + wellformed (efl:=all_env_flags) (erase_global_deps X_type (KernameSet.union deps (term_global_deps body)) X decls normalization_in prf).1 0 body ∥. +Proof. + intros ? Σ' wfΣ'. pose proof (abstract_env_exists X) as [[Σ wf]]. + destruct cb as [name [bod|] udecl]; simpl. + simpl in H. noconf H. + set (obl :=(ErasureFunction.erase_constant_body_obligation_1 X_type X' + {| + cst_type := name; + cst_body := Some bod; + cst_universes := udecl |} onc bod eq_refl)). clearbody obl. + destruct (obl _ wfΣ'). sq. + have er : (Σ, univs);;; [] |- bod ⇝ℇ erase X_type X' [] bod obl. + { eapply (erases_erase (X:=X')). + now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). + } + exists bod, A; intuition auto. + now rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). + eapply erase_wellformed_weaken. + now simpl in H. +Qed. + +Lemma erase_global_cst_decl_wf_glob X_type X deps decls normalization_in heq : + forall cb wfΣ hcb, + let X' := abstract_make_wf_env_ext X (cst_universes cb) wfΣ in + forall normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ, + let ecb := erase_constant_body X_type X' cb hcb in + let Σ' := erase_global_deps X_type (KernameSet.union deps ecb.2) X decls normalization_in heq in + (@wf_global_decl all_env_flags Σ'.1 (EAst.ConstantDecl ecb.1) : Prop). +Proof. + intros cb wfΣ hcb X' normalization_in' ecb Σ'. + unfold wf_global_decl. cbn. + pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. + pose proof (abstract_env_ext_exists X') as [[Σmake wfmake]]. + destruct (wfΣ _ wf), (hcb _ wfmake). red in X1. + destruct EAst.cst_body eqn:hb => /= //. + eapply (erase_constant_body_correct'') in hb; eauto. + destruct hb as [[t0 [T [[] ?]]]]. rewrite e in i. exact i. +Qed. + +Lemma erase_global_ind_decl_wf_glob {X_type X} {deps decls normalization_in kn m} heq : + (forall Σ : global_env, abstract_env_rel X Σ -> on_inductive cumulSpec0 (lift_typing typing) (Σ, ind_universes m) kn m) -> + let m' := erase_mutual_inductive_body m in + let Σ' := erase_global_deps X_type deps X decls normalization_in heq in + @wf_global_decl all_env_flags Σ'.1 (EAst.InductiveDecl m'). +Proof. + set (m' := erase_mutual_inductive_body m). + set (Σ' := erase_global_deps _ _ _ _ _ _). simpl. + intros oni. + pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. + pose proof (abstract_env_wf _ wf) as [wfΣ]. + assert (erases_mutual_inductive_body m (erase_mutual_inductive_body m)). + { eapply (erases_mutual (mdecl:=kn)); tea. } + eapply (erases_mutual_inductive_body_wf (univs := Σ.(universes)) (retro := Σ.(retroknowledge)) (Σ := decls) (kn := kn) (Σ' := Σ'.1)) in H; tea. + rewrite -(heq _ wf). now destruct Σ. +Qed. + +Lemma erase_global_deps_wf_glob X_type X deps decls normalization_in heq : + let Σ' := erase_global_deps X_type deps X decls normalization_in heq in + @wf_glob all_env_flags Σ'.1. +Proof. + cbn. + pose proof (abstract_env_exists X) as [[Σ wf]]. + pose proof (abstract_env_wf _ wf) as [wfΣ]. + revert deps X Σ wf wfΣ normalization_in heq. + induction decls; [cbn; auto|]. + { intros. constructor. } + intros. destruct a as [kn []]; simpl; destruct KernameSet.mem; set (Xpop := abstract_pop_decls X); + try set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _); + epose proof (abstract_env_exists Xpop) as [[Σpop wfpop]]; + pose proof (abstract_env_wf _ wfpop) as [wfΣpop]. + + constructor. eapply IHdecls => //; eauto. eapply erase_global_cst_decl_wf_glob; auto. + eapply erase_global_deps_fresh; auto. + destruct wfΣ. destruct wfΣpop. + rewrite (heq _ wf) in o0. depelim o0. now destruct o3. + + cbn. eapply IHdecls; eauto. + + constructor. eapply IHdecls; eauto. + destruct wfΣ as [[onu ond]]. + rewrite (heq _ wf) in o. depelim o. destruct o0. + eapply (erase_global_ind_decl_wf_glob (kn:=kn)); tea. + intros. rewrite (abstract_env_irr _ H wfpop). + unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wf wfpop) as [? ?]. + {intros; now eexists. } + destruct Σpop, Σ; cbn in *. now subst. + eapply erase_global_deps_fresh. + destruct wfΣ as [[onu ond]]. + rewrite (heq _ wf) in o. depelim o. now destruct o0. + + eapply IHdecls; eauto. +Qed. + +Lemma erase_constant_body_correct_glob {X_type X} {cb} {decls normalization_in prf} {univs wfΣ} + (X' := abstract_make_wf_env_ext X univs wfΣ) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} + {onc : forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> ∥ on_constant_decl (lift_typing typing) Σ' cb ∥} {body} : + EAst.cst_body (fst (erase_constant_body X_type X' cb onc)) = Some body -> + forall Σ' : global_env_ext, abstract_env_ext_rel X' Σ' -> + ∥ ∑ t T, (Σ' ;;; [] |- t : T) * (Σ' ;;; [] |- t ⇝ℇ body) * + (term_global_deps body = snd (erase_constant_body X_type X' cb onc)) * + wellformed (efl:=all_env_flags) (@erase_global X_type X decls normalization_in prf) 0 body ∥. +Proof. + intros ? Σ' wfΣ'. pose proof (abstract_env_exists X) as [[Σ wf]]. + destruct cb as [name [bod|] udecl]; simpl. + simpl in H. noconf H. + set (obl :=(ErasureFunction.erase_constant_body_obligation_1 X_type X' + {| + cst_type := name; + cst_body := Some bod; + cst_universes := udecl |} onc bod eq_refl)). clearbody obl. + destruct (obl _ wfΣ'). sq. + have er : (Σ, univs);;; [] |- bod ⇝ℇ erase X_type X' [] bod obl. + { eapply (erases_erase (X:=X')). + now rewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). + } + exists bod, A; intuition auto. + now rewrite (abstract_make_wf_env_ext_correct X univs wfΣ _ _ wf wfΣ'). + eapply (erase_global_wellformed decls _ _ _ (Γ := [])). + now simpl in H. +Qed. + +Lemma erase_global_cst_wf_glob X_type X decls normalization_in heq : + forall cb wfΣ hcb, + let X' := abstract_make_wf_env_ext X (cst_universes cb) wfΣ in + forall normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ, + let ecb := erase_constant_body X_type X' cb hcb in + let Σ' := @erase_global X_type X decls normalization_in heq in + (@wf_global_decl all_env_flags Σ' (EAst.ConstantDecl ecb.1) : Prop). +Proof. + intros cb wfΣ hcb X' normalization_in' ecb Σ'. + unfold wf_global_decl. cbn. + pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. + pose proof (abstract_env_ext_exists X') as [[Σmake wfmake]]. + destruct (wfΣ _ wf), (hcb _ wfmake). red in X1. + destruct EAst.cst_body eqn:hb => /= //. + eapply (erase_constant_body_correct_glob) in hb; eauto. + destruct hb as [[t0 [T [[] ?]]]]. exact i. +Qed. + +Lemma erase_global_ind_decl_wf_glob' {X_type X} {decls normalization_in kn m} heq : + (forall Σ : global_env, abstract_env_rel X Σ -> on_inductive cumulSpec0 (lift_typing typing) (Σ, ind_universes m) kn m) -> + let m' := erase_mutual_inductive_body m in + let Σ' := @erase_global X_type X decls normalization_in heq in + @wf_global_decl all_env_flags Σ' (EAst.InductiveDecl m'). +Proof. + set (m' := erase_mutual_inductive_body m). + set (Σ' := erase_global _ _ _). simpl. + intros oni. + pose proof (abstract_env_exists X) as [[Σ wf]]. specialize_Σ wf. + pose proof (abstract_env_wf _ wf) as [wfΣ]. + assert (erases_mutual_inductive_body m (erase_mutual_inductive_body m)). + { eapply (erases_mutual (mdecl:=kn)); tea. } + eapply (erases_mutual_inductive_body_wf (univs := Σ.(universes)) (retro := Σ.(retroknowledge)) (Σ := decls) (kn := kn) (Σ' := Σ')) in H; tea. + rewrite -(heq _ wf). now destruct Σ. +Qed. + +Lemma erase_global_wf_glob X_type X decls normalization_in heq : + let Σ' := @erase_global X_type X decls normalization_in heq in + @wf_glob all_env_flags Σ'. +Proof. + cbn. + pose proof (abstract_env_exists X) as [[Σ wf]]. + pose proof (abstract_env_wf _ wf) as [wfΣ]. + revert X Σ wf wfΣ normalization_in heq. + induction decls; [cbn; auto|]. + { intros. constructor. } + intros. destruct a as [kn []]; simpl; set (Xpop := abstract_pop_decls X); + try set (Xmake := abstract_make_wf_env_ext Xpop (cst_universes c) _); + epose proof (abstract_env_exists Xpop) as [[Σpop wfpop]]; + pose proof (abstract_env_wf _ wfpop) as [wfΣpop]. + + constructor. eapply IHdecls => //; eauto. + eapply erase_global_cst_wf_glob; auto. + eapply erase_global_fresh; auto. + destruct wfΣ. destruct wfΣpop. + rewrite (heq _ wf) in o0. depelim o0. now destruct o3. + + constructor. eapply IHdecls; eauto. + destruct wfΣ as [[onu ond]]. + rewrite (heq _ wf) in o. depelim o. destruct o0. + eapply (erase_global_ind_decl_wf_glob' (kn:=kn)); tea. + intros. rewrite (abstract_env_irr _ H wfpop). + unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wf wfpop) as [? ?]. + { intros; now eexists. } + destruct Σpop, Σ; cbn in *. now subst. + eapply erase_global_fresh. + destruct wfΣ as [[onu ond]]. + rewrite (heq _ wf) in o. depelim o. now destruct o0. +Qed. + +Lemma extends_cons d Σ Σ' : extends Σ Σ' -> extends (d :: Σ) (d :: Σ'). +Proof. + intros ext kn decl. + cbn. + case: (eqb_spec kn d.1). + - now intros -> [= <-]. + - intros hin. apply ext. +Qed. + +Lemma extends_fresh Σ Σ' kn d : + fresh_global kn Σ' -> + extends Σ Σ' -> + extends Σ ((kn, d) :: Σ'). +Proof. + intros fr ext. + intros kn' decl hl. + cbn. destruct (eqb_spec kn' kn) => //. subst. + eapply ext in hl. now eapply lookup_env_Some_fresh in hl. + now eapply ext. +Qed. + +Lemma Forall_filter {A} {P} f l : @Forall A P l -> Forall P (filter f l). +Proof. + induction 1; try econstructor; auto. + cbn. case E: (f x) => //. now constructor. +Qed. + +Lemma extends_filter_impl {efl : EEnvFlags} (f f' : kername * EAst.global_decl -> bool) Σ : + (forall x, f x -> f' x) -> + wf_glob Σ -> + extends (filter f Σ) (filter f' Σ). +Proof. + intros hf; induction Σ; cbn; auto. + - red; auto. + - intros wf; depelim wf. + case E: (f _); auto. rewrite (hf _ E). + now eapply extends_cons. + case E': (f' _); auto. + eapply extends_fresh. now eapply Forall_filter. + auto. +Qed. + +Lemma extends_filter {efl : EEnvFlags} (f : kername * EAst.global_decl -> bool) Σ : + wf_glob Σ -> + extends (filter f Σ) Σ. +Proof. + intros wf hf; induction Σ; cbn; auto. + intros decl. depelim wf. cbn in *. + case E: (f _); auto; cbn. + case (eqb_spec hf kn) => //. + - intros hd. now eapply IHΣ. + - intros hl; eapply IHΣ in hl => //. rewrite hl. + case (eqb_spec hf kn) => heq. + subst. now eapply lookup_env_Some_fresh in hl. + reflexivity. +Qed. + +Lemma erase_global_deps_eval {efl : EWcbvEval.WcbvFlags} (X_type : abstract_env_impl) (X X':X_type.π1) ext wfX wfX' {normalization_in} {normalization_in'} decls nin prf Γ t wt v wv : + let Xext := abstract_make_wf_env_ext X ext wfX in + let Xext' := abstract_make_wf_env_ext X' ext wfX' in + let t' := erase X_type Xext (normalization_in := normalization_in) Γ t wt in + let v' := erase X_type Xext' (normalization_in := normalization_in') Γ v wv in + let eg := erase_global_deps X_type (term_global_deps t') X decls nin prf in + let eg' := erase_global_deps X_type (term_global_deps v') X decls nin prf in + eg.1 ⊢ t' ⇓ v' -> EGlobalEnv.extends eg'.1 eg.1. +Proof. + intros. + move: (erase_global_deps_wf_glob X_type X (term_global_deps t') decls nin prf). + subst eg eg'. + rewrite !erase_global_deps_erase_global. + assert (wfer : wf_glob (efl := all_env_flags) (@erase_global X_type X decls nin prf)). + { eapply erase_global_wf_glob. } + rewrite !(filter_deps_filter (efl := all_env_flags)) //. + cbn. + intros wf. + eapply extends_filter_impl; tea. unfold flip in *. + intros [kn d]; cbn. clear d. + rewrite /is_true !KernameSet.mem_spec. revert kn. + eapply EWcbvEval.weakening_eval_env in H. 2:exact wfer. + now eapply eval_global_deps in H. + rewrite !erase_global_deps_erase_global (filter_deps_filter (efl := all_env_flags)) //. + now eapply extends_filter. +Qed. + +Lemma lookup_erase_global (cf := config.extraction_checker_flags) {X_type X} {deps deps' decls normalization_in prf} : + KernameSet.Subset deps deps' -> + EExtends.global_subset (erase_global_deps X_type deps X decls normalization_in prf).1 (erase_global_deps X_type deps' X decls normalization_in prf).1. +Proof. + intros sub. + rewrite !erase_global_deps_erase_global. + assert (wfer : wf_glob (efl := all_env_flags) (@erase_global X_type X decls normalization_in prf)). + { eapply erase_global_wf_glob. } + rewrite !(filter_deps_filter (efl := all_env_flags)) //. + eapply extends_filter_impl => //. 2:exact wfer. + unfold flip. + move=> x /KernameSet.mem_spec hin. + apply/KernameSet.mem_spec. + eapply global_deps_subset; tea. +Qed. + +Lemma expanded_weakening_global X_type X deps deps' decls normalization_in prf Γ t : + KernameSet.Subset deps deps' -> + expanded (erase_global_deps X_type deps X decls normalization_in prf).1 Γ t -> + expanded (erase_global_deps X_type deps' X decls normalization_in prf).1 Γ t. +Proof. + intros hs. + eapply expanded_ind; intros; try solve [econstructor; eauto]. + - eapply expanded_tFix; tea. solve_all. + - eapply expanded_tConstruct_app; tea. + destruct H. split; tea. + destruct d; split => //. + cbn in *. red in H. + eapply lookup_erase_global in H; tea. +Qed. + +Lemma expanded_erase (cf := config.extraction_checker_flags) + {X_type X decls normalization_in prf} univs wfΣ t wtp : + forall Σ : global_env, abstract_env_rel X Σ -> PCUICEtaExpand.expanded Σ [] t -> + let X' := abstract_make_wf_env_ext X univs wfΣ in + forall normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ, + let et := (erase X_type X' [] t wtp) in + let deps := EAstUtils.term_global_deps et in + expanded (erase_global_deps X_type deps X decls normalization_in prf).1 [] et. +Proof. + intros Σ wf hexp X' normalization_in'. + pose proof (abstract_env_wf _ wf) as [wf']. + pose proof (abstract_env_ext_exists X') as [[? wfmake]]. + eapply (expanded_erases (Σ := (Σ, univs))); tea. + eapply (erases_erase (X := X')); eauto. + now erewrite <- (abstract_make_wf_env_ext_correct X univs wfΣ). + cbn. + eapply (erases_deps_erase (X := X) univs wfΣ); eauto. +Qed. + +Lemma expanded_erase_global (cf := config.extraction_checker_flags) + deps {X_type X decls normalization_in prf} : + forall Σ : global_env, abstract_env_rel X Σ -> + PCUICEtaExpand.expanded_global_env Σ -> + expanded_global_env (erase_global_deps X_type deps X decls normalization_in prf).1. +Proof. + intros Σ wf etaΣ. pose proof (prf _ wf). + destruct Σ as [univ decls']. + red. revert wf. red in etaΣ. cbn in *. subst. + revert deps X normalization_in prf. + induction etaΣ; intros deps. intros. constructor. intros. + pose proof (abstract_env_exists (abstract_pop_decls X)) as [[Σpop wfpop]]. + unshelve epose proof (abstract_pop_decls_correct X Σ _ _ _ wf wfpop) as [? [? ?]]. + { now eexists. } + destruct Σpop. cbn in H0, H1, H2. subst. + destruct decl as [kn []]; + destruct (KernameSet.mem kn deps) eqn:eqkn; simpl; rewrite eqkn. + constructor; [eapply IHetaΣ; auto|]. + red. cbn. red. cbn in *. + unfold erase_constant_decl. + set (deps' := KernameSet.union _ _). hidebody deps'. + set (wfext := abstract_make_wf_env_ext _ _ _). hidebody wfext. + destruct H. + destruct c as [cst_na [cst_body|] cst_type cst_rel]. + cbn in *. + eapply expanded_weakening_global. + 2:{ eapply expanded_erase; tea. } + set (et := erase _ _ _ _) in *. + unfold deps'. unfold hidebody. intros x hin. + eapply KernameSet.union_spec. right => //. + now cbn. + eapply IHetaΣ; eauto. + constructor. eapply IHetaΣ; eauto. + red. cbn => //. + eapply IHetaΣ; eauto. +Qed. + +(* Sanity checks: the [erase] function maximally erases terms *) +Lemma erasable_tBox_value (wfl := Ee.default_wcbv_flags) (Σ : global_env_ext) (wfΣ : wf_ext Σ) t T v : + forall wt : Σ ;;; [] |- t : T, + Σ |-p t ⇓ v -> erases Σ [] v E.tBox -> ∥ isErasable Σ [] t ∥. +Proof. + intros. + depind H. + eapply Is_type_eval_inv; eauto. eexists; eauto. +Qed. + +Lemma erase_eval_to_box (wfl := Ee.default_wcbv_flags) + {X_type X} {univs wfext t v Σ' t' deps decls normalization_in prf} + (Xext := abstract_make_wf_env_ext X univs wfext) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} + : + forall wt : forall Σ : global_env_ext, abstract_env_ext_rel Xext Σ -> welltyped Σ [] t, + erase X_type Xext [] t wt = t' -> + KernameSet.subset (term_global_deps t') deps -> + erase_global_deps X_type deps X decls normalization_in prf = Σ' -> + forall Σext : global_env_ext, abstract_env_ext_rel Xext Σext -> + (forall Σ : global_env, abstract_env_rel X Σ -> + PCUICWcbvEval.eval Σ t v) -> + @Ee.eval Ee.default_wcbv_flags Σ'.1 t' E.tBox -> ∥ isErasable Σext [] t ∥. +Proof. + intros wt. + intros. + destruct (erase_correct X_type X univs wfext _ _ Σ' _ _ decls prf wt H H0 H1 X0 _ H2) as [ev [eg [eg']]]. + pose proof (Ee.eval_deterministic H3 eg'). subst. + pose proof (abstract_env_exists X) as [[? wf]]. + destruct (wfext _ wf). destruct (wt _ H2) as [T wt']. + pose proof (abstract_env_ext_wf _ H2) as [?]. + eapply erasable_tBox_value; eauto. + pose proof (abstract_make_wf_env_ext_correct X univs wfext _ _ wf H2). subst. + apply X0; eauto. +Qed. + +Lemma erase_eval_to_box_eager (wfl := Ee.default_wcbv_flags) + {X_type X} {univs wfext t v Σ' t' deps decls normalization_in prf} + (Xext := abstract_make_wf_env_ext X univs wfext) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} + : + forall wt : forall Σ : global_env_ext, abstract_env_ext_rel Xext Σ -> welltyped Σ [] t, + erase X_type Xext [] t wt = t' -> + KernameSet.subset (term_global_deps t') deps -> + erase_global_deps X_type deps X decls normalization_in prf = Σ' -> + forall Σext : global_env_ext, abstract_env_ext_rel Xext Σext -> + (forall Σ : global_env, abstract_env_rel X Σ -> + PCUICWcbvEval.eval Σ t v) -> + @Ee.eval Ee.default_wcbv_flags Σ'.1 t' E.tBox -> t' = E.tBox. +Proof. + intros wt. + intros. + destruct (erase_eval_to_box wt H H0 H1 _ H2 X0 H3). + subst t'. + destruct (inspect_bool (is_erasableb X_type Xext [] t wt)) eqn:heq. + - simp erase. rewrite heq. + simp erase => //. + - elimtype False. + pose proof (abstract_env_exists X) as [[? wf]]. + destruct (@is_erasableP X_type Xext _ [] t wt) => //. apply n. + intros. sq. now rewrite (abstract_env_ext_irr _ H H2). +Qed. + +From MetaCoq.PCUIC Require Import PCUICFirstorder. + +From Equations Require Import Equations. + +Inductive firstorder_evalue Σ : EAst.term -> Prop := + | is_fo i n args npars : + EGlobalEnv.lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars -> + Forall (firstorder_evalue Σ) (skipn npars args) -> + firstorder_evalue Σ (EAst.mkApps (EAst.tConstruct i n []) args). + +Lemma list_size_skipn {A} (l : list A) size n : + list_size size (skipn n l) <= list_size size l. +Proof. + induction n in l |- *. + - rewrite skipn_0 //. + - destruct l. rewrite skipn_nil. now cbn. + rewrite skipn_S. cbn. specialize (IHn l); lia. +Qed. + +Section elim. + Context (Σ : E.global_context). + Context (P : EAst.term -> Prop). + Context (ih : (forall i n args npars, + EGlobalEnv.lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars -> + Forall (firstorder_evalue Σ) (skipn npars args) -> + Forall P (skipn npars args) -> + P (EAst.mkApps (EAst.tConstruct i n []) args))). + + Equations? firstorder_evalue_elim (t : EAst.term) (fo : firstorder_evalue Σ t) : P t by wf t (MR lt EInduction.size) := + { | _, is_fo i n args npars hl hf => _ }. + Proof. + eapply ih; tea. + eapply In_Forall. intros x hin. + eapply PCUICWfUniverses.Forall_In in hf; tea. + apply firstorder_evalue_elim => //. red. + rewrite EInduction.size_mkApps. + eapply (In_size id EInduction.size) in hin. + unfold id in *. pose proof (list_size_skipn args EInduction.size npars). + change (fun x => EInduction.size x) with EInduction.size in hin. clear -hin H. + eapply Nat.lt_le_trans; tea. cbn. lia. + Qed. +End elim. + + +Lemma welltyped_mkApps_inv {cf} {Σ : global_env_ext} Γ f args : ∥ wf Σ ∥ -> + welltyped Σ Γ (mkApps f args) -> welltyped Σ Γ f /\ Forall (welltyped Σ Γ) args. +Proof. + intros wf (A & HA). sq. eapply inversion_mkApps in HA as (? & ? & ?). + split. eexists; eauto. + induction t0 in f |- *; econstructor; eauto; econstructor; eauto. +Qed. + +From MetaCoq.PCUIC Require Import PCUICProgress. + +Lemma firstorder_erases_deterministic X_type (X : X_type.π1) + univs wfext {v t' i u args} + (Xext := abstract_make_wf_env_ext X univs wfext) + {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} + : + forall wv : (forall Σ, Σ ∼_ext Xext -> welltyped Σ [] v), + forall Σ, Σ ∼_ext Xext -> + Σ ;;; [] |- v : mkApps (tInd i u) args -> + Σ |-p v ⇓ v -> + @firstorder_ind Σ (firstorder_env Σ) i -> + erases Σ [] v t' -> + t' = erase X_type Xext (normalization_in:=normalization_in) [] v wv. +Proof. + intros wv Σ Hrel Hty Hvalue Hfo Herase. + destruct (firstorder_lookup_inv Hfo) as [mind Hdecl]. + assert (Hext : ∥ wf_ext Σ∥) by now eapply ErasureFunction.heΣ. + sq. eapply firstorder_value_spec in Hty as Hfov; eauto. + clear - Hrel Hext Hfov Herase. + revert t' wv Herase. + pattern v. + revert v Hfov. + eapply firstorder_value_inds. intros. + rewrite erase_mkApps. + - intros Σ0 HΣ0. pose proof (abstract_env_ext_irr _ HΣ0 Hrel). subst. + eapply PCUICValidity.inversion_mkApps in X0 as (? & XX & Happ). + clear XX. revert Happ. clear. generalize (mkApps (tInd i u) pandi). induction 1. + + do 2 econstructor. + + destruct IHHapp. do 2 econstructor. econstructor; eauto. eauto. + - intros. eapply erases_mkApps_inv in Herase as [(? & ? & ? & -> & [Herasable] & ? & ? & ->)|(? & ? & -> & ? & ?)]. all:eauto. + + exfalso. eapply isErasable_Propositional in Herasable; eauto. now rewrite Herasable in H1. + + inv H2. + * cbn. unfold erase_clause_1. destruct (inspect_bool (is_erasableb X_type Xext [] (tConstruct i n ui) Hyp0)). + -- exfalso. sq. destruct (@is_erasableP _ _ _ [] (tConstruct i n ui) Hyp0) => //. + specialize_Σ Hrel. sq. + eapply (isErasable_Propositional (args := [])) in s; eauto. + now rewrite s in H1. + -- f_equal. eapply Forall2_eq. clear X0 H wv. induction H3. + ++ cbn. econstructor. + ++ cbn. econstructor. + ** inv H0. eapply H5. eauto. + ** inv H0. eapply IHForall2. eauto. + * exfalso. eapply (isErasable_Propositional (args := [])) in X1; eauto. + now rewrite X1 in H1. + - eauto. + - intros ? ? H3. assert (Hext_ : ∥ wf_ext Σ0∥) by now eapply ErasureFunction.heΣ. + sq. + specialize_Σ H2. + eapply (isErasable_Propositional) in H3; eauto. + pose proof (abstract_env_ext_irr _ H2 Hrel). subst. + now rewrite H3 in H1. + - intros. assert (Hext__ : ∥ wf_ext Σ0∥) by now eapply ErasureFunction.heΣ. + specialize_Σ H2. eapply welltyped_mkApps_inv in wv; eauto. eapply wv. + now sq. + Unshelve. all: try exact False. + - eapply PCUICWcbvEval.eval_to_value; eauto. +Qed. + +Lemma firstorder_evalue_extends Σ Σ' t : + extends Σ Σ' -> + firstorder_evalue Σ t -> + firstorder_evalue Σ' t. +Proof. + intros ext; move: t. + apply: firstorder_evalue_elim => i n args npars hl hf ih. + econstructor; tea. + red in ext. + move: hl; rewrite /lookup_inductive_pars /lookup_minductive. + move: (ext (inductive_mind i)). + now case: lookup_env => // a; move/(_ _ eq_refl) ->. +Qed. + +Lemma In_map_All {A B C : Type} {Q : C -> Type} {P : C -> A -> Prop} + (fn : forall x : A, (forall y : C, Q y -> P y x) -> B) (l : list A) (Hl : forall y : C, Q y -> ∥ All (P y) l ∥) : + forall x, In x l -> + exists D : forall y : C, Q y -> P y x, + In (fn x D) (map_All fn l Hl). +Proof. + induction l; cbn => //. + intros x []. + - subst. eexists; left; reflexivity. + - destruct (IHl (map_All_obligation_2 a l Hl) _ H). + eexists; now right. +Qed. + +Lemma In_term_global_deps t l : + In t l -> + forall kn, KernameSet.In kn (term_global_deps t) -> KernameSet.In kn (terms_global_deps l). +Proof. + induction l; cbn => //. + intros []. + - subst; intros kn hin; rewrite knset_in_fold_left. + left. eapply KernameSet.union_spec. now left. + - intros kn hin. specialize (IHl H). + rewrite knset_in_fold_left. right. now exists t. +Qed. + +Lemma erase_correct_strong' (wfl := Ee.default_wcbv_flags) X_type (X : X_type.π1) + univs wfext {t v Σ' t' i u args} decls normalization_in prf + (Xext := abstract_make_wf_env_ext X univs wfext) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} + : + forall wt : (forall Σ, Σ ∼_ext Xext -> welltyped Σ [] t), + forall Σ, abstract_env_ext_rel Xext Σ -> + axiom_free Σ -> + Σ ;;; [] |- t : mkApps (tInd i u) args -> + @firstorder_ind Σ (firstorder_env Σ) i -> + erase X_type Xext [] t wt = t' -> + erase_global_deps X_type (term_global_deps t') X decls normalization_in prf = Σ' -> + red Σ [] t v -> + ¬ { t' & Σ ;;; [] |- v ⇝ t'} -> + forall wt', let ev := erase X_type Xext [] v wt' in ∥ Σ'.1 ⊢ t' ⇓ ev ∥ + /\ firstorder_evalue Σ'.1 ev. +Proof. + intros wt Σ Hrel Hax Hty Hfo <- <- Hred Hirred wt'. + destruct (firstorder_lookup_inv Hfo) as [mind Hdecl]. + destruct (abstract_env_exists X) as [[Σg Hg]]. + pose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ Hg Hrel). subst Σ. + pose proof (ErasureFunction.heΣ _ _ _ Hrel) as [Hwf]. eapply wcbv_standardization_fst in Hty as Heval; eauto. + edestruct (erase_correct X_type X univs wfext t v) as [v' [H1 H2]]; eauto. + { eapply KernameSet.subset_spec; reflexivity. } + 1:{ intros ? H_. sq. enough (Σ = Σg) as -> by eauto. + pose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ H_ Hrel). now subst. } + eapply firstorder_erases_deterministic in H1; eauto. + simpl. + destruct H2. split; [sq; tea|]. rewrite -> H1 in e. exact e. + assert (PCUICFirstorder.firstorder_value (Σg, univs) [] v). + { destruct (wfext Σg Hg). eapply PCUICFirstorder.firstorder_value_spec; tea. apply X0. constructor. + eapply PCUICClassification.subject_reduction_eval; tea. + eapply PCUICWcbvEval.eval_to_value; tea. } + { rewrite H1 in e. eapply erase_global_deps_eval in e. + eapply firstorder_evalue_extends; tea. clear e. + clear -Hrel H Hg Hwf. move: v H wt'. + apply: PCUICFirstorder.firstorder_value_inds. + intros i n ui u args pandi hty hf ih isp wt'. + rewrite erase_mkApps //. + { intros ? h; specialize (wt' _ h). eapply welltyped_mkApps_inv in wt' as [? ?] => //. destruct (abstract_env_ext_wf _ h). + now sq; eapply Forall_All. pose proof (abstract_env_ext_wf _ h). now sq. } + 2:{ intros Σ hΣ [ise]. pose proof (abstract_env_ext_irr _ Hrel hΣ). subst Σ. + eapply isErasable_Propositional in ise => //. + cbn in isp, ise. now rewrite ise in isp. } + 2:{ intros Σ hΣ. pose proof (abstract_env_ext_irr _ Hrel hΣ). subst Σ. + pose proof (iswelltyped hty). + now eapply welltyped_mkApps_inv in H. } + intros. simp erase. + destruct inspect_bool; simp erase. + move/is_erasableP: i0. intros hc. + destruct (hc _ Hrel). eapply (isErasable_Propositional (args:=[])) in X0 => //. cbn in X0, isp => //. + rewrite X0 in isp => //. + eapply PCUICInductiveInversion.Construct_Ind_ind_eq' in hty as [mdecl [idecl [cdecl [declc _]]]] => //=. + set (v := E.mkApps _ _). + eapply (erase_global_declared_constructor _ _ _ _ _ _ _ (term_global_deps v) decls normalization_in prf) in declc; tea. + destruct declc as [[hm hi] hc]. + 2:{ rewrite /v /= term_global_deps_mkApps. destruct i => /=. eapply KernameSet.union_spec; left. now eapply KernameSet.singleton_spec. } + 2:{ eapply Hwf. } + econstructor. + { rewrite /lookup_inductive_pars /lookup_minductive hm /= //. } + eapply Forall_skipn. + clear -ih Hrel isp. + { revert v. rewrite erase_terms_eq => v. + eapply Forall_All in ih. + assert (All + (fun t : term => + forall wt' : forall Σ : global_env_ext, Σ ∼_ext Xext -> welltyped Σ [] t, + firstorder_evalue + (erase_global_deps X_type (term_global_deps v) X decls + normalization_in prf).1 (erase X_type Xext [] t wt')) args). + { eapply In_All. intros x hin wt. + eapply All_In in ih; tea. destruct ih as [ih]. + specialize (ih wt). + eapply firstorder_evalue_extends; tea. + subst v. eapply lookup_erase_global. + rewrite term_global_deps_mkApps. + intros kn hin'. eapply KernameSet.union_spec; right. + destruct (In_map_All (erase X_type (abstract_make_wf_env_ext X univs wfext) []) _ Hyp1 _ hin). + eapply In_term_global_deps; tea. + erewrite (erase_irrel _ _).1; tea. } + eapply All_Forall, All_map_All. + { intros. exact X0. } + cbn. intros x y rx h. apply h. exact Hrel. } } + { eapply subject_reduction; eauto. } + { eapply PCUICWcbvEval.value_final. eapply PCUICWcbvEval.eval_to_value. eauto. } + all: eauto. + Unshelve. eauto. +Qed. + +Lemma erase_correct_strong (wfl := Ee.default_wcbv_flags) X_type (X : X_type.π1) + univs wfext {t v i u args} decls normalization_in prf + (Xext := abstract_make_wf_env_ext X univs wfext) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ} + : forall wt : (forall Σ, Σ ∼_ext Xext -> welltyped Σ [] t), + forall Σ, abstract_env_ext_rel Xext Σ -> + axiom_free Σ -> + Σ ;;; [] |- t : mkApps (tInd i u) args -> + let t' := erase X_type Xext [] t wt in + let Σ' := erase_global_deps X_type (term_global_deps t') X decls normalization_in prf in + @firstorder_ind Σ (firstorder_env Σ) i -> + red Σ [] t v -> + ¬ { v' & Σ ;;; [] |- v ⇝ v'} -> + exists wt', let et' := erase X_type Xext [] v wt' in + ∥ Σ'.1 ⊢ t' ⇓ erase X_type Xext [] v wt' ∥ /\ firstorder_evalue Σ'.1 et'. +Proof. + intros wt Σ Hrel Hax Hty Hdecl Hfo Hsub Hred Hirred. + unshelve eexists. + - abstract (intros Σ_ H_; pose proof (ErasureFunction.heΣ _ _ _ H_); sq; + pose proof (abstract_env_ext_irr _ H_ Hrel); subst; eapply red_welltyped; eauto; econstructor; eauto). + - eapply erase_correct_strong'; eauto. + all:reflexivity. +Qed. + +Lemma unique_decls {X_type} {X : X_type.π1} univs wfext +(Xext := abstract_make_wf_env_ext X univs wfext) {Σ'} (wfΣ' : Σ' ∼_ext Xext) : + forall Σ, Σ ∼ X -> declarations Σ = declarations Σ'. +Proof. + intros Σ wfΣ. pose (abstract_make_wf_env_ext_correct X univs wfext Σ Σ' wfΣ wfΣ'). + rewrite e. reflexivity. +Qed. + +Lemma unique_type {X_type : abstract_env_impl} {X : X_type.π2.π1} + {Σ:global_env_ext} + (wfΣ : Σ ∼_ext X) {t T}: + Σ ;;; [] |- t : T -> forall Σ, Σ ∼_ext X -> welltyped Σ [] t. +Proof. + intros wt Σ' wfΣ'. erewrite (abstract_env_ext_irr X wfΣ' wfΣ). + now exists T. +Qed. + +Definition erase_correct_firstorder (wfl := Ee.default_wcbv_flags) + X_type (X : X_type.π1) univs wfext {t v i u args} + (Xext := abstract_make_wf_env_ext X univs wfext) + {normalization_in : forall X Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} : + forall Σ (wfΣ : abstract_env_ext_rel Xext Σ) + (wt : Σ ;;; [] |- t : mkApps (tInd i u) args), + axiom_free Σ -> + let t' := erase X_type Xext [] t (unique_type wfΣ wt) in + let decls := declarations (fst Σ) in + let Σ' := erase_global_deps X_type (term_global_deps t') X decls + (fun _ _ _ _ _ _ _ => normalization_in _ _) (unique_decls univs wfext wfΣ) in + @firstorder_ind Σ (firstorder_env Σ) i -> + red Σ [] t v -> + ¬ { v' & Σ ;;; [] |- v ⇝ v'} -> + exists wt', let et' := erase X_type Xext [] v wt' in ∥ Σ'.1 ⊢ t' ⇓ et' ∥ /\ + firstorder_evalue Σ'.1 et'. +Proof. + intros. + eapply erase_correct_strong; eauto. +Qed. + +(* we use the [match] trick to get typeclass resolution to pick up the right instances without leaving any evidence in the resulting term, and without having to pass them manually everywhere *) +Notation NormalizationIn_erase_global_deps_fast X decls + := (match extraction_checker_flags, extraction_normalizing return _ with + | cf, no + => forall n, + n < List.length decls -> + forall kn cb pf, + List.hd_error (List.skipn n decls) = Some (kn, ConstantDecl cb) -> + let Xext := abstract_make_wf_env_ext X (cst_universes cb) pf in + forall Σ, wf_ext Σ -> Σ ∼_ext Xext -> NormalizationIn Σ + end) + (only parsing). + + +Section EraseGlobalFast. + + Import PCUICEnvironment. + +Definition decls_prefix decls (Σ' : global_env) := + ∑ Σ'', declarations Σ' = Σ'' ++ decls. + +Lemma on_global_decls_prefix {cf} Pcmp P univs retro decls decls' : + PCUICAst.on_global_decls Pcmp P univs retro (decls ++ decls') -> + PCUICAst.on_global_decls Pcmp P univs retro decls'. +Proof using Type. + induction decls => //. + intros ha; depelim ha. + now apply IHdecls. +Qed. + +Lemma decls_prefix_wf {decls Σ} : + decls_prefix decls Σ -> wf Σ -> wf {| universes := Σ.(universes); declarations := decls; retroknowledge := Σ.(retroknowledge) |}. +Proof using Type. + intros [Σ' hd] wfΣ. + split. apply wfΣ. + cbn. destruct wfΣ as [_ ond]. rewrite hd in ond. + now eapply on_global_decls_prefix in ond. +Qed. + +Lemma incl_cs_refl cs : cs ⊂_cs cs. +Proof using Type. + split; [lsets|csets]. +Qed. + +Lemma weaken_prefix {decls Σ kn decl} : + decls_prefix decls Σ -> + wf Σ -> + lookup_env {| universes := Σ; declarations := decls; retroknowledge := Σ.(retroknowledge) |} kn = Some decl -> + on_global_decl cumulSpec0 (lift_typing typing) (Σ, universes_decl_of_decl decl) kn decl. +Proof using Type. + intros prefix wfΣ. + have wfdecls := decls_prefix_wf prefix wfΣ. + epose proof (weakening_env_lookup_on_global_env (lift_typing typing) _ Σ kn decl + weaken_env_prop_typing wfΣ wfdecls). + forward X. apply extends_strictly_on_decls_extends, strictly_extends_decls_extends_strictly_on_decls. red; split => //. + now apply (X wfdecls). +Qed. + + +(* This version of global environment erasure keeps the same global environment throughout the whole + erasure, while folding over the list of declarations. *) + +Local Obligation Tactic := program_simpl. + +Program Fixpoint erase_global_deps_fast (deps : KernameSet.t) + X_type (X:X_type.π1) (decls : global_declarations) + {normalization_in : NormalizationIn_erase_global_deps_fast X decls} + (prop : forall Σ : global_env, abstract_env_rel X Σ -> ∥ decls_prefix decls Σ ∥) : E.global_declarations * KernameSet.t := + match decls with + | [] => ([],deps) + | (kn, ConstantDecl cb) :: decls => + if KernameSet.mem kn deps then + let Xext' := abstract_make_wf_env_ext X (cst_universes cb) _ in + let normalization_in' : @id Prop _ := _ in (* hack to avoid program erroring out on unresolved tc *) + let cb' := @erase_constant_body X_type Xext' normalization_in' cb _ in + let deps := KernameSet.union deps (snd cb') in + let Σ' := erase_global_deps_fast deps X_type X decls _ in + (((kn, E.ConstantDecl (fst cb')) :: Σ'.1), Σ'.2) + else + erase_global_deps_fast deps X_type X decls _ + + | (kn, InductiveDecl mib) :: decls => + if KernameSet.mem kn deps then + let mib' := erase_mutual_inductive_body mib in + let Σ' := erase_global_deps_fast deps X_type X decls _ in + (((kn, E.InductiveDecl mib') :: Σ'.1), Σ'.2) + else erase_global_deps_fast deps X_type X decls _ + end. +Next Obligation. + pose proof (abstract_env_wf _ H) as [?]. + specialize_Σ H. sq. split. cbn. apply X3. cbn. + eapply decls_prefix_wf in X3; tea. + destruct X3 as [onu ond]. cbn in ond. + depelim ond. now destruct o. +Qed. +Next Obligation. + cbv [id]. + unshelve eapply (normalization_in 0); cbn; try reflexivity; try lia. +Defined. +Next Obligation. + pose proof (abstract_env_ext_wf _ H) as [?]. + pose proof (abstract_env_exists X) as [[? wf]]. + pose proof (abstract_env_wf _ wf) as [?]. + pose proof (prop' := prop _ wf). + sq. eapply (weaken_prefix (kn := kn)) in prop'; tea. + 2:{ cbn. rewrite eqb_refl //. } + epose proof (abstract_make_wf_env_ext_correct X (cst_universes cb) _ _ _ wf H). subst. + apply prop'. +Qed. +Next Obligation. + unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. +Defined. +Next Obligation. + specialize_Σ H. sq; red. destruct prop as [Σ' ->]. + eexists (Σ' ++ [(kn, ConstantDecl cb)]); rewrite -app_assoc //. +Qed. +Next Obligation. + unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. +Defined. +Next Obligation. + specialize_Σ H. sq; red. destruct prop as [Σ' ->]. + eexists (Σ' ++ [(kn, ConstantDecl cb)]); rewrite -app_assoc //. +Qed. +Next Obligation. + unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. +Defined. +Next Obligation. + specialize_Σ H. sq; red. destruct prop as [Σ' ->]. + eexists (Σ' ++ [(kn, _)]); rewrite -app_assoc //. +Qed. +Next Obligation. + unshelve eapply (normalization_in (S _)); cbn in *; revgoals; try eassumption; lia. +Defined. +Next Obligation. + specialize_Σ H. sq; red. destruct prop as [Σ' ->]. + eexists (Σ' ++ [(kn, _)]); rewrite -app_assoc //. +Qed. + +Import PCUICAst PCUICEnvironment. + +Lemma wf_lookup Σ kn d suffix g : + wf Σ -> + declarations Σ = suffix ++ (kn, d) :: g -> + lookup_env Σ kn = Some d. +Proof using Type. + unfold lookup_env. + intros [_ ond] eq. rewrite eq in ond |- *. + move: ond; clear. + induction suffix. + - cbn. rewrite eqb_refl //. + - cbn. intros ond. + depelim ond. destruct o as [f ? ? ? ]. cbn. red in f. + eapply Forall_app in f as []. depelim H0. cbn in H0. + destruct (eqb_spec kn kn0); [contradiction|]. + now apply IHsuffix. +Qed. + +Definition add_suffix suffix Σ := set_declarations Σ (suffix ++ Σ.(declarations)). + +Lemma add_suffix_cons d suffix Σ : add_suffix (d :: suffix) Σ = add_global_decl (add_suffix suffix Σ) d. +Proof using Type. reflexivity. Qed. + +Lemma global_erased_with_deps_weaken_prefix suffix Σ Σ' kn : + wf (add_suffix suffix Σ) -> + global_erased_with_deps Σ Σ' kn -> + global_erased_with_deps (add_suffix suffix Σ) Σ' kn. +Proof using Type. + induction suffix. + - unfold add_suffix; cbn. intros wf hg. + now rewrite /set_declarations /= -eta_global_env. + - rewrite add_suffix_cons. intros wf H. + destruct a as [kn' d]. eapply global_erases_with_deps_weaken => //. + apply IHsuffix => //. + destruct wf as [onu ond]. now depelim ond. +Qed. + + +(* Using weakening it is trivial to show that a term found to be erasable in Σ + will be found erasable in any well-formed extension. The converse is not so trivial: + some valid types in the extension are not valid in the restricted global context. + So, we will rather show that the erasure function is invariant under extension. *) + +Lemma isErasable_irrel_global_env {Σ Σ' : global_env_ext} {Γ t} : + wf Σ -> + wf Σ' -> + extends Σ' Σ -> + Σ.2 = Σ'.2 -> + isErasable Σ' Γ t -> isErasable Σ Γ t. +Proof using Type. + unfold isErasable. + intros wfΣ wfΣ' ext eqext [T [ht isar]]. + destruct Σ as [Σ decl], Σ' as [Σ' decl']. cbn in eqext; subst decl'. + exists T; split. + eapply (env_prop_typing weakening_env (Σ', decl)) => //=. + destruct isar as [|s]; [left|right] => //. + destruct s as [u [Hu isp]]. + exists u; split => //. + eapply (env_prop_typing weakening_env (Σ', decl)) => //=. +Qed. + +Definition reduce_stack_eq {cf} {fl} {no:normalizing_flags} {X_type : abstract_env_impl} {X : X_type.π2.π1} {normalization_in : forall Σ, wf_ext Σ -> Σ ∼_ext X -> NormalizationIn Σ} Γ t π wi : reduce_stack fl X_type X Γ t π wi = ` (reduce_stack_full fl X_type X Γ t π wi). +Proof using Type. + unfold reduce_stack. destruct reduce_stack_full => //. +Qed. + +Definition same_principal_type {cf} + {X_type : abstract_env_impl} {X : X_type.π2.π1} + {X_type' : abstract_env_impl} {X' : X_type'.π2.π1} + {Γ : context} {t} (p : PCUICSafeRetyping.principal_type X_type X Γ t) (p' : PCUICSafeRetyping.principal_type X_type' X' Γ t) := + p.π1 = p'.π1. + +(*Lemma erase_global_deps_suffix {deps} {Σ Σ' : wf_env} {decls hprefix hprefix'} : + wf Σ -> wf Σ' -> + universes Σ = universes Σ' -> + erase_global_deps_fast deps Σ decls hprefix = + erase_global_deps_fast deps Σ' decls hprefix'. +Proof using Type. + intros wfΣ wfΣ' equ. + induction decls in deps, hprefix, hprefix' |- *; cbn => //. + destruct a as [kn []]. + set (obl := (erase_global_deps_fast_obligation_1 Σ ((kn, ConstantDecl c) :: decls) + hprefix kn c decls eq_refl)). clearbody obl. + destruct obl. + set (eb := erase_constant_body _ _ _). + set (eb' := erase_constant_body _ _ _). + assert (eb = eb') as <-. + { subst eb eb'. + set (wfΣx := abstract_make_wf_env_ext _ _ _). + set (wfΣ'x := abstract_make_wf_env_ext _ _ _). + set (obl' := erase_global_deps_fast_obligation_2 _ _ _ _ _ _ _). + clearbody obl'. + set (Σd := {| universes := Σ.(universes); declarations := decls |}). + assert (wfΣd : wf Σd). + { eapply decls_prefix_wf; tea. + clear -hprefix. move: hprefix. unfold decls_prefix. + intros [Σ'' eq]. exists (Σ'' ++ [(kn, ConstantDecl c)]). + rewrite -app_assoc //. } + set (wfΣdwfe := build_wf_env_from_env _ (sq wfΣd)). + assert (wfextΣd : wf_ext (Σd, cst_universes c)). + { split => //. cbn. apply w. } + set (wfΣdw := abstract_make_wf_env_ext wfΣdwfe (cst_universes c) (sq wfextΣd)). + assert (ond' : ∥ on_constant_decl (lift_typing typing) (Σd, cst_universes c) c ∥ ). + { pose proof (decls_prefix_wf hprefix wfΣ). + destruct X as [onu ond]. depelim ond. sq. apply o0. } + assert (extd' : extends_decls Σd Σ). + { clear -hprefix. move: hprefix. + intros [? ?]. split => //. eexists (x ++ [(kn, ConstantDecl c)]). cbn. + now rewrite -app_assoc. } + assert (extd'' : extends_decls Σd Σ'). + { clear -equ hprefix'. move: hprefix'. + intros [? ?]. split => //. eexists (x ++ [(kn, ConstantDecl c)]). cbn. + now rewrite -app_assoc. } + rewrite (erase_constant_suffix (Σ := wfΣx) (Σ' := wfΣdw) (ondecl' := ond') wfΣ extd') //. + rewrite (erase_constant_body_irrel (Σ := wfΣ'x) (Σ' := wfΣdw) (ondecl' := ond') wfΣ' extd'') //. } + destruct KernameSet.mem => //. f_equal. + eapply IHdecls. + destruct KernameSet.mem => //. f_equal. + eapply IHdecls. +Qed.*) + +Lemma strictly_extends_decls_extends_global_env Σ Σ' : + wf Σ' -> + strictly_extends_decls Σ Σ' -> + extends_global_env Σ Σ'. +Proof. + intros wfΣ' [uni [Σ'' de] retr]. red. + cbn. unfold lookup_env. rewrite de /primitive_constant retr. + split => //. intros kn decl. + intros hl. + apply lookup_global_Some_iff_In_NoDup. + apply EnvMap.EnvMap.fresh_globals_iff_NoDup. + rewrite -de. now apply wf_fresh_globals. + apply lookup_global_Some_iff_In_NoDup in hl. + apply in_app_iff. now right. + apply EnvMap.EnvMap.fresh_globals_iff_NoDup. + apply wf_fresh_globals in wfΣ'. rewrite de in wfΣ'. + clear -wfΣ'. induction Σ''; auto. apply IHΣ''. + now depelim wfΣ'. +Qed. + +Lemma extends_global_env_equiv_env Σ Σ' : + extends_global_env Σ Σ' -> equiv_env_inter Σ Σ'. +Proof. + intros ext. split. intros kn decl decl' hl hl'. + destruct ext. apply H in hl. congruence. + now destruct ext. +Qed. + +Lemma erase_global_deps_fast_spec_gen {deps} + {X_type X X'} {decls normalization_in normalization_in' hprefix hprefix'} : + (forall Σ Σ', abstract_env_rel X Σ -> abstract_env_rel X' Σ' -> universes Σ = universes Σ' /\ retroknowledge Σ = retroknowledge Σ') -> + erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) hprefix = + erase_global_deps X_type deps X' decls normalization_in' hprefix'. +Proof using Type. + intros equ. + induction decls in X, X', equ, deps, normalization_in, normalization_in', hprefix, hprefix' |- * => //. + pose proof (abstract_env_exists X) as [[Σ wfΣ]]. + pose proof (abstract_env_exists X') as [[Σ' wfΣ']]. + pose proof (abstract_env_wf _ wfΣ) as [wf]. + pose proof (abstract_env_exists (abstract_pop_decls X')) as [[? wfpop]]. + unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' wfpop) as [? [? ?]]. + { now eexists. } + + destruct a as [kn []]. + - cbn. unfold erase_constant_decl. + set (obl := (erase_global_deps_fast_obligation_1 X_type X ((kn, ConstantDecl c) :: decls) + _ hprefix kn c decls eq_refl)). + set (eb := erase_constant_body _ _ _ _). + set (eb' := erase_constant_body _ _ _ _). + set (eg := erase_global_deps _ _ _ _ _ _). + set (eg' := erase_global_deps _ _ _ _ _ _). + + assert (eb = eb') as <-. + { subst eb eb'. + destruct (hprefix _ wfΣ) as [[Σ'' eq]]. + eapply erase_constant_body_irrel; cbn => //. + eapply equiv_env_inter_hlookup. + intros ? ? H2 H3. + epose proof (abstract_make_wf_env_ext_correct X (cst_universes c) _ _ _ wfΣ H2). + epose proof (abstract_make_wf_env_ext_correct (abstract_pop_decls X') (cst_universes c) _ _ _ wfpop H3). + subst. eapply equiv_env_inter_sym. + eapply extends_global_env_equiv_env. + pose proof (abstract_env_wf _ wfΣ) as []. + apply strictly_extends_decls_extends_global_env. + apply X0. red. + rewrite eq. rewrite <- H0, <- H1. split. cbn. symmetry; apply equ; eauto. + cbn. + eexists (Σ'' ++ [(kn, ConstantDecl c)]). cbn. subst. now rewrite -app_assoc. subst. + symmetry. now apply equ. } + destruct KernameSet.mem => //; f_equal; try eapply IHdecls. + f_equal. f_equal. eapply IHdecls. + intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. + { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. + f_equal. eapply IHdecls. + intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. + { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. + intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. + { now eexists. } intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. + + - cbn. + destruct KernameSet.mem => //; f_equal; try eapply IHdecls. + f_equal; f_equal; eapply IHdecls. + intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. + { now eexists. } + intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. + intros. + f_equal; eapply IHdecls. + intros. unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. + { now eexists. } + intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. + intros; unshelve epose proof (abstract_pop_decls_correct X' decls _ _ _ wfΣ' H3) as [? ?]. + { now eexists. } + intuition auto. rewrite <- H6. apply equ; eauto. rewrite <- H7; apply equ; auto. +Qed. + +Lemma erase_global_deps_fast_spec {deps} {X_type X} {decls normalization_in normalization_in' hprefix hprefix'} : + erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) hprefix = + erase_global_deps X_type deps X decls normalization_in' hprefix'. +Proof using Type. + eapply erase_global_deps_fast_spec_gen; intros. + rewrite (abstract_env_irr _ H H0); eauto. +Qed. + +Lemma cored_extends (Σ Σ' : global_env_ext) Γ x y : + wf Σ' -> + extends Σ Σ' -> + cored Σ Γ x y -> cored Σ' Γ x y. +Proof. + intros wf ext. + induction 1. + - constructor. + eapply PCUICWeakeningEnvConv.weakening_env_red1. apply wf. exact ext. + exact X. + - econstructor 2; tea. + eapply PCUICWeakeningEnvConv.weakening_env_red1. apply wf. + exact ext. + exact X. +Qed. + +Lemma normalization_in_extends (Σ Σ' : global_env) ext : + wf Σ -> wf Σ' -> + extends Σ' Σ -> + NormalizationIn (Σ, ext) -> NormalizationIn (Σ', ext). +Proof. + intros wfΣ wfΣ' extH nin. + red in nin. intros Γ t wt. + assert (welltyped (Σ, ext) Γ t). + { destruct wt. exists A. eapply (env_prop_typing weakening_env) in X; tea; cbn. } + specialize (nin _ _ H). clear H. + induction nin. + constructor. intros y r. + eapply H0. eapply cored_extends in r; tea. + eapply cored_welltyped; tea. +Qed. + +Lemma iter_abstract_pop_decls_correct {cf} {abstract_env_impl abstract_env_ext_impl : Type} + {abstract_env_struct0 : abstract_env_struct abstract_env_impl + abstract_env_ext_impl} : + abstract_env_prop abstract_env_impl abstract_env_ext_impl -> + forall (X : abstract_env_impl) (decls : list (kername × global_decl)) n, + (forall Σ : global_env, + Σ ∼ X -> exists decls', #|decls'| = n /\ declarations Σ = decls' ++ decls) -> + let X' := iter abstract_pop_decls n X in + forall Σ Σ' : global_env, + Σ ∼ X -> + Σ' ∼ X' -> + declarations Σ' = decls /\ universes Σ = universes Σ' /\ retroknowledge Σ = retroknowledge Σ'. +Proof. + intros ap X decls. + induction n in X, decls |- *. cbn. + - intros hdecl Σ Σ' HX HX'. + specialize (hdecl _ HX) as [decls' [hlen hde]]. + destruct decls' => //. cbn in hde. + now rewrite -hde (abstract_env_irr _ HX HX'). + - intros hdecl X' Σ Σ' HX HX'. + specialize (hdecl _ HX) as [decls' [hlen hde]]. + destruct decls' => //; cbn in hde. + specialize (IHn (abstract_pop_decls X) decls). + pose proof (abstract_pop_decls_correct X (decls' ++ decls)). + forward H. + { intros ? h. rewrite -(abstract_env_irr _ HX h). now exists p. } + forward IHn. + { intros ? H0. now specialize (H _ _ HX H0). } + pose proof (abstract_env_exists (abstract_pop_decls X)) as [[Σpop hpop]]. + specialize (IHn _ _ hpop HX'). + destruct IHn. split => //. + pose proof (abstract_pop_decls_correct X (decls' ++ decls)). + forward H2. + { intros ? hx. pose proof (abstract_env_irr _ hx HX). subst Σ0. now exists p. } + specialize (H2 _ _ HX hpop). destruct H2 as [? []]. + destruct H1. + split; congruence. +Qed. + +Lemma skipn_app_prefix {A} n (l l' : list A) : skipn (#|l| + n) (l ++ l') = skipn n l'. +Proof. + induction l; cbn; auto. +Qed. + +Lemma erase_global_deps_fast_erase_global_deps {deps} {X_type X} {decls normalization_in hprefix hprefix'} : + exists normalization_in', + erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) hprefix = + erase_global_deps X_type deps X decls normalization_in' hprefix'. +Proof using Type. + unshelve eexists. + intros. + pose proof (abstract_env_exists X) as [[Σ' H']]. + pose proof (abstract_env_wf _ H') as []. + specialize (hprefix _ H') as [[Σ'' ?]]. + pose proof (abstract_env_exists (iter abstract_pop_decls (S n) X)) as [[? ?]]. + pose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ a H1). noconf H2. + epose proof (iter_abstract_pop_decls_correct _ X (skipn (S n) (Σ'' ++ decls)) (S n)). + forward H2. { intros ? h. rewrite (abstract_env_irr _ h H'). exists (firstn (S n) (Σ'' ++ decls)). + rewrite firstn_skipn. split => //. rewrite firstn_length_le //. rewrite app_length. lia. } + specialize (H2 _ _ H' a) as [? []]. + eapply normalization_in_extends. exact X1. exact X0. + { eapply extends_decls_extends, strictly_extends_decls_extends_decls. + red. rewrite H2 H3 H4. split => //. exists (firstn (S n) (Σ'' ++ decls)). + now rewrite firstn_skipn. } + specialize (normalization_in n H kn cb). + assert (wf_ext (Σ', cst_universes cb)). + { split => //. cbn. rewrite H3. apply X0. } + assert (forall Σ : global_env, Σ ∼ X -> ∥ wf_ext (Σ, cst_universes cb) ∥). + { intros. pose proof (abstract_env_irr _ H' H5). now subst Σ'. } + unshelve eapply normalization_in; tea. + set (foo := abstract_make_wf_env_ext X (cst_universes cb) H5). + pose proof (abstract_env_ext_exists foo) as [[Σext hext]]. + epose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ H' hext). now subst Σext. + erewrite erase_global_deps_fast_spec; tea. reflexivity. +Qed. + +Definition erase_global_fast X_type deps X decls {normalization_in} (prf:forall Σ : global_env, abstract_env_rel X Σ -> declarations Σ = decls) := + erase_global_deps_fast deps X_type X decls (normalization_in:=normalization_in) (fun _ H => sq ([] ; prf _ H)). + +Lemma expanded_erase_global_fast (cf := config.extraction_checker_flags) deps + {X_type X decls normalization_in} {normalization_in':NormalizationIn_erase_global_deps X decls} {prf} : + forall Σ : global_env, abstract_env_rel X Σ -> + PCUICEtaExpand.expanded_global_env Σ -> + expanded_global_env (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf).1. +Proof using Type. + unfold erase_global_fast. + erewrite erase_global_deps_fast_spec. + eapply expanded_erase_global. + Unshelve. all: eauto. +Qed. + +Lemma expanded_erase_fast (cf := config.extraction_checker_flags) + {X_type X decls normalization_in prf} {normalization_in'':NormalizationIn_erase_global_deps X decls} univs wfΣ t wtp : + forall Σ : global_env, abstract_env_rel X Σ -> + PCUICEtaExpand.expanded Σ [] t -> + let X' := abstract_make_wf_env_ext X univs wfΣ in + forall (normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ), + let et := (erase X_type X' [] t wtp) in + let deps := EAstUtils.term_global_deps et in + expanded (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf).1 [] et. +Proof using Type. + intros Σ wf hexp X' normalization_in'. pose proof (abstract_env_wf _ wf) as [?]. + eapply (expanded_erases (Σ := (Σ, univs))); tea. + eapply (erases_erase (X := X')). + pose proof (abstract_env_ext_exists X') as [[? wfmake]]. + now rewrite <- (abstract_make_wf_env_ext_correct X univs _ _ _ wf wfmake). + cbn. unfold erase_global_fast. erewrite erase_global_deps_fast_spec => //. + eapply (erases_deps_erase (X := X) univs wfΣ); eauto. + Unshelve. all: eauto. +Qed. + +Lemma erase_global_fast_wf_glob X_type deps X decls normalization_in (normalization_in':NormalizationIn_erase_global_deps X decls) prf : + @wf_glob all_env_flags (erase_global_fast X_type deps X decls (normalization_in:=normalization_in) prf).1. +Proof using Type. + unfold erase_global_fast; erewrite erase_global_deps_fast_spec. + eapply erase_global_deps_wf_glob. + Unshelve. all: eauto. +Qed. + +Lemma erase_wellformed_fast (efl := all_env_flags) + {X_type X decls normalization_in prf} {normalization_in'':NormalizationIn_erase_global_deps X decls} univs wfΣ {Γ t} wt + (X' := abstract_make_wf_env_ext X univs wfΣ) + {normalization_in' : forall Σ, wf_ext Σ -> Σ ∼_ext X' -> NormalizationIn Σ} : + let t' := erase X_type X' Γ t wt in + wellformed (erase_global_fast X_type (term_global_deps t') X decls (normalization_in:=normalization_in) prf).1 #|Γ| t'. +Proof using Type. + intros. + cbn. unfold erase_global_fast. erewrite erase_global_deps_fast_spec. + eapply erase_wellformed_deps. + Unshelve. all: eauto. +Qed. + +End EraseGlobalFast. + +Fixpoint compile_value_erase (t : PCUICAst.term) (acc : list EAst.term) : EAst.term := + match t with + | PCUICAst.tApp f a => compile_value_erase f (compile_value_erase a [] :: acc) + | PCUICAst.tConstruct i n _ => EAst.mkApps (EAst.tConstruct i n []) acc + | _ => EAst.tVar "error" + end. + +Lemma compile_value_erase_mkApps i n ui args acc : + compile_value_erase (mkApps (tConstruct i n ui) args) acc = + EAst.mkApps (EAst.tConstruct i n []) (map (flip compile_value_erase []) args ++ acc). +Proof. + revert acc; induction args using rev_ind. + - cbn. auto. + - intros acc. rewrite PCUICAstUtils.mkApps_app /=. cbn. + now rewrite IHargs map_app /= -app_assoc /=. +Qed. + + +Lemma erases_firstorder Σ Γ t : + PCUICFirstorder.firstorder_value Σ Γ t -> + erases Σ Γ t (compile_value_erase t []). +Proof. + move: t. + apply: (PCUICFirstorder.firstorder_value_inds Σ Γ). + intros i n ui u args pandi hty hfo ih isp. + assert (Forall2 (erases Σ Γ) args (map (flip compile_value_erase []) args)). + { solve_all. eapply All_All2; tea. + cbn. intros x [fo hx]. exact hx. } + unshelve epose proof (erases_mkApps Σ Γ (tConstruct i n ui) (EAst.tConstruct i n []) args _ _ H). + now constructor. + now rewrite compile_value_erase_mkApps app_nil_r. +Qed. + +Lemma erases_firstorder' Σ Γ t : + wf_ext Σ -> + PCUICFirstorder.firstorder_value Σ Γ t -> + forall t', erases Σ Γ t t' -> t' = (compile_value_erase t []). +Proof. + intros wf. move: t. + apply: (PCUICFirstorder.firstorder_value_inds Σ Γ). + intros i n ui u args pandi hty hfo ih isp. + assert (Forall2 (erases Σ Γ) args (map (flip compile_value_erase []) args)). + { solve_all. eapply All_All2; tea. + cbn. intros x [fo hx]. now eapply erases_firstorder. } + unshelve epose proof (erases_mkApps Σ Γ (tConstruct i n ui) (EAst.tConstruct i n []) args _ _ H). + constructor => //. + rewrite compile_value_erase_mkApps app_nil_r. + intros t' ht'. + eapply erases_mkApps_inv in ht'. + destruct ht'. + { destruct H1 as [? [? [? [? [[] ?]]]]]. + eapply isErasable_Propositional in X => //. + now rewrite X in isp. } + destruct H1 as [f' [L' [-> [erc erargs]]]]. + depelim erc. f_equal. + solve_all. clear -erargs. + induction erargs => //. rewrite IHerargs. cbn. f_equal. + destruct r as [[] ?]. eapply e in e0. now unfold flip. + eapply (isErasable_Propositional (args := [])) in X => //. + now rewrite X in isp. +Qed. diff --git a/erasure/theories/Extract.v b/erasure/theories/Extract.v index 39c2ddbac..0d304219a 100644 --- a/erasure/theories/Extract.v +++ b/erasure/theories/Extract.v @@ -1,9 +1,9 @@ (* Distributed under the terms of the MIT license. *) -From Coq Require Import Program. +From Coq Require Import Program ssrbool. From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import config Primitive. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICPrimitive PCUICTyping - PCUICElimination PCUICWcbvEval. + PCUICElimination PCUICWcbvEval PCUICFirstorder. From MetaCoq.Erasure Require EAst EGlobalEnv. Module E := EAst. @@ -13,18 +13,6 @@ Local Existing Instance extraction_checker_flags. Definition isErasable Σ Γ t := ∑ T, Σ ;;; Γ |- t : T × (isArity T + (∑ u, (Σ ;;; Γ |- T : tSort u) * is_propositional u))%type. -Definition isPropositionalArity ar b := - match destArity [] ar with - | Some (_, s) => is_propositional s = b - | None => False - end. - -Definition isPropositional Σ ind b := - match lookup_inductive Σ ind with - | Some (mdecl, idecl) => isPropositionalArity idecl.(ind_type) b - | _ => False - end. - Fixpoint mkAppBox c n := match n with | 0 => c @@ -64,7 +52,7 @@ Inductive erases (Σ : global_env_ext) (Γ : context) : term -> E.term -> Prop : | erases_tConst : forall (kn : kername) (u : Instance.t), Σ;;; Γ |- tConst kn u ⇝ℇ E.tConst kn | erases_tConstruct : forall (kn : inductive) (k : nat) (n : Instance.t), - isPropositional Σ kn false -> + ~~ isPropositional Σ kn -> Σ;;; Γ |- tConstruct kn k n ⇝ℇ E.tConstruct kn k [] | erases_tCase (ci : case_info) (p : predicate term) (c : term) (brs : list (branch term)) (c' : E.term) @@ -124,7 +112,7 @@ Lemma erases_forall_list_ind (Hconst : forall Γ kn u, P Γ (tConst kn u) (E.tConst kn)) (Hconstruct : forall Γ kn k n, - isPropositional Σ kn false -> + ~~ isPropositional Σ kn -> P Γ (tConstruct kn k n) (E.tConstruct kn k [])) (Hcase : forall Γ ci p c brs c' brs', PCUICElimination.Subsingleton Σ ci.(ci_ind) -> @@ -220,7 +208,7 @@ Definition erases_one_inductive_body (oib : one_inductive_body) (oib' : E.one_in Forall2 (fun 'i i' => i.(PCUICEnvironment.proj_name) = i'.(E.proj_name)) oib.(ind_projs) oib'.(E.ind_projs) /\ oib'.(E.ind_name) = oib.(ind_name) /\ oib'.(E.ind_kelim) = oib.(ind_kelim) /\ - isPropositionalArity oib.(ind_type) oib'.(E.ind_propositional). + isPropositionalArity oib.(ind_type) = oib'.(E.ind_propositional). Definition erases_mutual_inductive_body (mib : mutual_inductive_body) (mib' : E.mutual_inductive_body) := let bds := mib.(ind_bodies) in diff --git a/erasure/theories/Typed/ErasureCorrectness.v b/erasure/theories/Typed/ErasureCorrectness.v index 70272cbf9..45d392b0e 100644 --- a/erasure/theories/Typed/ErasureCorrectness.v +++ b/erasure/theories/Typed/ErasureCorrectness.v @@ -5,7 +5,7 @@ From Equations Require Import Equations. From MetaCoq.Erasure Require Import EDeps. From MetaCoq.Erasure Require Import ESubstitution. From MetaCoq.Erasure Require Import Extract. -From MetaCoq.Erasure Require Import ErasureFunction. +From MetaCoq.Erasure Require Import ErasureFunction ErasureFunctionProperties. From MetaCoq.PCUIC Require Import PCUICAstUtils. From MetaCoq.PCUIC Require Import PCUICSR. From MetaCoq.PCUIC Require Import PCUICTyping. @@ -76,10 +76,6 @@ Proof. cbn. constructor; [easy|]. apply IHl. - - destruct wf as [(i & [])]. - unfold isPropositionalArity. - rewrite ind_arity_eq. - now rewrite !destArity_it_mkProd_or_LetIn. Qed. Lemma erase_ind_correct Σ wfΣ kn mib wf : @@ -100,6 +96,7 @@ Qed. End ECorrect. Opaque erase_type flag_of_type ErasureFunction.wf_reduction. + Lemma erase_global_decls_deps_recursive_correct decls univs retro wfΣ include ignore_deps : let Σ := mk_global_env univs decls retro in (forall k, ignore_deps k = false) -> diff --git a/erasure/theories/Typed/ExtractionCorrectness.v b/erasure/theories/Typed/ExtractionCorrectness.v index c4ffc361c..6eb8525f8 100644 --- a/erasure/theories/Typed/ExtractionCorrectness.v +++ b/erasure/theories/Typed/ExtractionCorrectness.v @@ -9,6 +9,7 @@ From MetaCoq.Erasure.Typed Require Import WcbvEvalAux. From Equations Require Import Equations. From MetaCoq.Erasure Require Import ErasureCorrectness. From MetaCoq.Erasure Require Import ErasureFunction. +From MetaCoq.Erasure Require Import ErasureFunctionProperties. From MetaCoq.Erasure Require Import EWcbvEval. From MetaCoq.Erasure Require Import Extract. From MetaCoq.PCUIC Require Import PCUICAstUtils. diff --git a/pcuic/theories/PCUICFirstorder.v b/pcuic/theories/PCUICFirstorder.v index 758a4a2cc..a72c8362e 100644 --- a/pcuic/theories/PCUICFirstorder.v +++ b/pcuic/theories/PCUICFirstorder.v @@ -98,18 +98,16 @@ Section cf. Context {cf : config.checker_flags}. -Definition isPropositional Σ ind b := - match lookup_env Σ (inductive_mind ind) with - | Some (InductiveDecl mdecl) => - match nth_error mdecl.(ind_bodies) (inductive_ind ind) with - | Some idecl => - match destArity [] idecl.(ind_type) with - | Some (_, s) => is_propositional s = b - | None => False - end - | None => False - end - | _ => False +Definition isPropositionalArity ar := + match destArity [] ar with + | Some (_, s) => is_propositional s + | None => false + end. + +Definition isPropositional Σ ind := + match lookup_inductive Σ ind with + | Some (mdecl, idecl) => isPropositionalArity idecl.(ind_type) + | _ => false end. Inductive firstorder_value Σ Γ : term -> Prop := @@ -117,7 +115,7 @@ Inductive firstorder_value Σ Γ : term -> Prop := Σ ;;; Γ |- mkApps (tConstruct i n ui) args : mkApps (tInd i u) pandi -> Forall (firstorder_value Σ Γ) args -> - isPropositional Σ i false -> + ~~ isPropositional Σ i -> firstorder_value Σ Γ (mkApps (tConstruct i n ui) args). Lemma firstorder_value_inds : @@ -127,7 +125,7 @@ Lemma firstorder_value_inds : Σ;;; Γ |- mkApps (tConstruct i n ui) args : mkApps (tInd i u) pandi -> Forall (firstorder_value Σ Γ) args -> Forall P args -> - isPropositional (PCUICEnvironment.fst_ctx Σ) i false -> + ~~ isPropositional (PCUICEnvironment.fst_ctx Σ) i -> P (mkApps (tConstruct i n ui) args)) -> forall t : term, firstorder_value Σ Γ t -> P t. Proof using Type. @@ -139,21 +137,21 @@ Qed. Lemma firstorder_ind_propositional {Σ : global_env_ext} {wfΣ:wf Σ} i mind oind : declared_inductive Σ i mind oind -> @firstorder_ind Σ (firstorder_env Σ) i -> - isPropositional Σ i false. + ~~ isPropositional Σ i. Proof using Type. intros d. unshelve epose proof (d_ := declared_inductive_to_gen d); eauto. pose proof d_ as [d1 d2]. intros H. red in d1. unfold firstorder_ind in H. - red. sq. + unfold isPropositional. unfold PCUICEnvironment.fst_ctx in *. rewrite d1 in H |- *. - solve_all. unfold firstorder_mutind in H. - rewrite d2. move/andP: H => [ind H0]. + rewrite /lookup_inductive /lookup_inductive_gen /lookup_minductive_gen /= d1 d2. + move/andP: H => [ind H0]. eapply forallb_nth_error in H0; tea. erewrite d2 in H0. cbn in H0. unfold firstorder_oneind in H0. solve_all. destruct (ind_sort oind) eqn:E2; inv H0. eapply PCUICInductives.declared_inductive_type in d. - rewrite d. rewrite E2. + rewrite d. rewrite E2. unfold isPropositionalArity. now rewrite destArity_it_mkProd_or_LetIn. Qed. From 915ce93fb1641e227e1318fba70f33e20bbb41ed Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 8 Sep 2023 17:29:53 +0200 Subject: [PATCH 48/61] We're missing preservation of eta-expandedness by reduction in PCUIC... --- erasure-plugin/theories/ErasureCorrectness.v | 23 +++- pcuic/theories/PCUICEtaExpand.v | 11 +- pcuic/theories/PCUICWcbvEval.v | 130 ++++++++++++++++++- 3 files changed, 157 insertions(+), 7 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 5b32b75a6..87e8f948f 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -870,6 +870,19 @@ Admitted. Arguments PCUICFirstorder.firstorder_ind _ _ : clear implicits. +Section PCUICExpandLets. + Import PCUICExpandLets PCUICExpandLetsCorrectness. + + Lemma trans_axiom_free Σ : axiom_free Σ -> axiom_free (trans_global_env Σ). + Proof. + intros ax kn decl. + rewrite /trans_global_env /= /PCUICAst.declared_constant /= /trans_global_decls. + intros h; apply PCUICElimination.In_map in h as [[kn' decl'] [hin heq]]. + noconf heq. destruct decl'; noconf H. + apply ax in hin. + destruct c as [? [] ? ?] => //. + Qed. + Section pipeline_theorem. Instance cf : checker_flags := extraction_checker_flags. @@ -910,7 +923,7 @@ Section pipeline_theorem. - red. cbn. split; eauto. eexists. eapply PCUICClassification.subject_reduction_eval; eauto. - - todo "preservation of eta expandedness". + - todo "preservation of eta expandedness". - cbn. todo "normalization". - todo "normalization". Qed. @@ -1052,8 +1065,12 @@ Section pipeline_theorem. unfold erase_pcuic_program. cbn [fst snd]. exact obs. clear obs b0 ev e w. eapply extends_erase_pcuic_program. cbn. - eapply (PCUICExpandLetsCorrectness.trans_wcbveval (Σ := (Σ.1, Σ.2))). admit. exact X0. - cbn. 2:cbn. 3:cbn. + eapply (PCUICExpandLetsCorrectness.trans_wcbveval (Σ := (Σ.1, Σ.2))). + { clear -HΣ typing. now eapply PCUICClosedTyp.subject_closed in typing. } + cbn. 2:cbn. 3:cbn. exact X0. + + + Admitted. Lemma verified_erasure_pipeline_lambda : diff --git a/pcuic/theories/PCUICEtaExpand.v b/pcuic/theories/PCUICEtaExpand.v index f6c9f6cab..b4db67093 100644 --- a/pcuic/theories/PCUICEtaExpand.v +++ b/pcuic/theories/PCUICEtaExpand.v @@ -1,7 +1,7 @@ From Coq Require Import ssreflect. From Equations Require Import Equations. From MetaCoq.Utils Require Import utils. -From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping PCUICProgram. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping PCUICProgram PCUICCSubst. Definition isConstruct t := match t with tConstruct _ _ _ => true | _ => false end. @@ -341,6 +341,14 @@ Proof. solve_all. Qed. +Lemma csubst_mkApps a k f l : + csubst a k (mkApps f l) = mkApps (csubst a k f) (map (csubst a k) l). +Proof. + induction l in f |- *; cbn; [auto|]. + rewrite IHl. + now cbn. +Qed. + Lemma All_fold_tip {A : Type} (P : list A -> A -> Type) {x} : All_fold P [x] -> P [] x. Proof. intros a; now depelim a. @@ -425,3 +433,4 @@ Proof. solve_all. } all:intros; try solve [econstructor; eauto 1; solve_all; try now rewrite app_assoc]. Qed. + diff --git a/pcuic/theories/PCUICWcbvEval.v b/pcuic/theories/PCUICWcbvEval.v index bb3481dd7..93adfe7fd 100644 --- a/pcuic/theories/PCUICWcbvEval.v +++ b/pcuic/theories/PCUICWcbvEval.v @@ -4,7 +4,7 @@ From MetaCoq.Common Require Import config. From MetaCoq.Utils Require Import utils. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICGlobalEnv PCUICReduction PCUICClosed PCUICCSubst - PCUICClosedTyp. (* Due to reliance on wf Σ instead of closed_env Σ *) + PCUICClosedTyp PCUICEtaExpand. (* Due to reliance on wf Σ instead of closed_env Σ *) Require Import ssreflect ssrbool. From Equations Require Import Equations. @@ -24,14 +24,23 @@ From Equations Require Import Equations. Local Ltac inv H := inversion H; subst. +Lemma nApp_mkApps t f args : + t = mkApps f args -> ~~ isApp t -> t = f /\ args = []. +Proof. + intros -> napp. + destruct args using rev_case; cbn in *; solve_discr; try discriminate => //. split => //. + now rewrite mkApps_app /= in napp. +Qed. + Ltac solve_discr := try progress (prepare_discr; finish_discr; cbn[mkApps] in * ); try match goal with H : mkApps _ _ = mkApps ?f ?l |- _ => eapply mkApps_eq_inj in H as [? ?]; [|easy|easy]; subst; try intuition congruence; try noconf H | H : ?t = mkApps ?f ?l |- _ => - change t with (mkApps t []) in H ; - eapply mkApps_eq_inj in H as [? ?]; [|easy|easy]; subst; try intuition congruence; try noconf H + (change t with (mkApps t []) in H ; + eapply mkApps_eq_inj in H as [? ?]; [|easy|easy]; subst; try intuition congruence; try noconf H) || + (eapply nApp_mkApps in H as [? ?]; [|easy]; subst) | H : mkApps ?f ?l = ?t |- _ => change t with (mkApps t []) in H ; eapply mkApps_eq_inj in H as [? ?]; [|easy|easy]; subst; try intuition congruence; try noconf H @@ -1212,3 +1221,118 @@ Proof. - eapply eval_cofix_case; tea. eapply value_final, value_app. now constructor. auto. Qed. + +Lemma expanded_tApp Σ Γ f a : + ~ isConstructApp f || isFixApp f || isRel (head f) -> + expanded Σ Γ (tApp f a) -> + expanded Σ Γ f /\ expanded Σ Γ a. +Proof. + rewrite /isConstructApp /isFixApp. + intros hf exp; depind exp. + - destruct args using rev_case; solve_discr; try discriminate. + rewrite mkApps_app in H3; noconf H3. + eapply Forall_app in H1 as [? ha]. depelim ha. + destruct args using rev_case; cbn in hf => //. + now rewrite !head_mkApps /= in hf. + - destruct args using rev_case; solve_discr. subst f6. + eapply IHexp => //. + rewrite mkApps_app in H2; noconf H2. + eapply Forall_app in H0 as [? H0]. depelim H0. split => //. + rewrite !head_mkApps in hf. + eapply expanded_mkApps => //. + - destruct args using rev_case; solve_discr. discriminate. + rewrite mkApps_app in H6; noconf H6. + eapply Forall_app in H1 as [? h]. depelim h. split => //. + now rewrite !head_mkApps /= in hf. + - destruct args using rev_case; solve_discr. discriminate. + rewrite mkApps_app in H3; noconf H3. + now rewrite !head_mkApps /= in hf. +Qed. + +Lemma expanded_tLambda_inv Σ Γ na t b : + expanded Σ Γ (tLambda na t b) -> + expanded Σ (0 :: Γ) b. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + +Lemma expanded_tLetIn_inv Σ Γ na t t' b : + expanded Σ Γ (tLetIn na t t' b) -> + expanded Σ Γ t /\ expanded Σ (0 :: Γ) b. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + +Lemma expanded_tCase_inv Σ Γ ci p c brs : + expanded Σ Γ (tCase ci p c brs) -> + Forall (expanded Σ Γ) (pparams p) /\ + Forall + (fun br : branch term => + ∥ All_fold + (fun (Δ : list context_decl) (d : context_decl) + => + ForOption + (fun b : term => + expanded Σ + (repeat 0 #|Δ| ++ + repeat 0 #|pparams p|) b) + (decl_body d)) (bcontext br) ∥ /\ + expanded Σ (repeat 0 #|bcontext br| ++ Γ) (bbody br)) + brs /\ + expanded Σ Γ c. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + +Lemma expanded_tProj_inv Σ Γ p c : + expanded Σ Γ (tProj p c) -> + expanded Σ Γ c. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + +Lemma expanded_red {cf : checker_flags} {Σ : global_env_ext} Γ Γ' t v : wf Σ -> + red1 Σ Γ t v -> + expanded Σ Γ' t -> + expanded Σ Γ' v. +Proof. + intros wf red; induction red using red1_ind_all in Γ' |- *; intros exp. + - eapply expanded_tApp in exp as [] => //. + eapply (expanded_subst _ _ _ _ [] Γ') => //. now constructor. + now eapply expanded_tLambda_inv in H. + - eapply expanded_tLetIn_inv in exp as []. + eapply (expanded_subst _ _ _ _ [] Γ') => //. now constructor. + - admit. + - eapply expanded_tCase_inv in exp as [? []]. + unfold iota_red. admit. + - admit. + - admit. + - eapply expanded_tProj_inv in exp. econstructor. admit. + - admit. + - admit. + - constructor. now eapply expanded_tLambda_inv in exp. + - constructor. eapply IHred. now eapply expanded_tLambda_inv in exp. + - eapply expanded_tLetIn_inv in exp; now constructor. + - eapply expanded_tLetIn_inv in exp. now constructor. + - eapply expanded_tLetIn_inv in exp. now constructor. + - eapply expanded_tCase_inv in exp as [? []]. + constructor; eauto. cbn. + solve_all. eapply OnOne2_impl_All_r; tea. intuition eauto. now apply X0. + now rewrite /PCUICOnOne.set_pparams /= -(OnOne2_length X). + - eapply expanded_tCase_inv in exp as [? []]. + constructor; eauto. + - eapply expanded_tCase_inv in exp as [? []]. + econstructor; eauto. + - eapply expanded_tCase_inv in exp as [? []]. + econstructor; eauto. solve_all. + eapply OnOne2_impl_All_r; tea. + intros x y [? ?]. intros [[] ?]. rewrite -e0. + solve_all. + - eapply expanded_tProj_inv in exp. now econstructor. + - case_eq (isConstructApp (tApp M1 M2) || isFixApp (tApp M1 M2) || isRel (head (tApp M1 M2))). + intros h. + clear red. specialize (IHred Γ'). + depelim exp. + { destruct args using rev_case; cbn in *; noconf H1. rewrite mkApps_app in H1; noconf H1. +(* not the right way to go about this, we need some mkApps induction hypothesis *) +Abort. From f1699162d543998dfafd35a4562b66c2af50db97 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 11 Sep 2023 18:08:45 +0200 Subject: [PATCH 49/61] Prove preservation of expansion by reduction using a custom reduction relation --- erasure/theories/EEtaExpandedFix.v | 2 +- pcuic/theories/PCUICWcbvEval.v | 575 ++++++++++++++++++++++++++++- 2 files changed, 562 insertions(+), 15 deletions(-) diff --git a/erasure/theories/EEtaExpandedFix.v b/erasure/theories/EEtaExpandedFix.v index 8c14b356c..92a783a73 100644 --- a/erasure/theories/EEtaExpandedFix.v +++ b/erasure/theories/EEtaExpandedFix.v @@ -1408,7 +1408,7 @@ Proof. split. eapply All2_app => //. rewrite -H3. eauto. destruct s. * destruct p; solve_discr. noconf H2. - left. split. + left. split.w unfold isStuckFix'; rewrite e1. len. eapply Nat.leb_le. lia. now rewrite -[tApp _ _](mkApps_app _ _ [av]). * right. len. eapply isEtaExp_fixapp_mon; tea. lia. diff --git a/pcuic/theories/PCUICWcbvEval.v b/pcuic/theories/PCUICWcbvEval.v index 93adfe7fd..ddd600473 100644 --- a/pcuic/theories/PCUICWcbvEval.v +++ b/pcuic/theories/PCUICWcbvEval.v @@ -1222,7 +1222,28 @@ Proof. eapply value_final, value_app. now constructor. auto. Qed. -Lemma expanded_tApp Σ Γ f a : +Lemma expanded_tApp (Σ : global_env) (Γ : list nat) (f : term) a : + expanded Σ Γ f -> expanded Σ Γ a -> + expanded Σ Γ (tApp f a). +Proof. + induction 1 using expanded_ind; intros expa. + all:rewrite -?[tApp _ a](mkApps_app _ _ [a]). + all:try (eapply (expanded_mkApps _ _ _ [a]) => //; econstructor; eauto). + + - econstructor; tea. rewrite app_length. lia. eapply app_Forall;eauto. + - econstructor; tea. eapply app_Forall; eauto. + - eapply expanded_tFix; tea. eapply app_Forall; eauto. eauto. rewrite app_length; cbn; eauto. lia. + - eapply expanded_tConstruct_app; tea. rewrite app_length ; lia. eapply app_Forall; eauto. +Qed. + +Lemma expanded_mkApps (Σ : global_env) (Γ : list nat) (f : term) (args : list term) : + expanded Σ Γ f -> Forall (expanded Σ Γ) args -> expanded Σ Γ (mkApps f args). +Proof. + intros expf expa; induction expa in f, expf |- *; eauto. + cbn. eapply IHexpa. now eapply expanded_tApp. +Qed. + +Lemma expanded_tApp_inv Σ Γ f a : ~ isConstructApp f || isFixApp f || isRel (head f) -> expanded Σ Γ (tApp f a) -> expanded Σ Γ f /\ expanded Σ Γ a. @@ -1263,6 +1284,13 @@ Proof. intros exp; depind exp; solve_discr => //; eauto. Qed. +Lemma expanded_tEvar_inv Σ Γ ev l: + expanded Σ Γ (tEvar ev l) -> + Forall (expanded Σ Γ) l. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + Lemma expanded_tCase_inv Σ Γ ci p c brs : expanded Σ Γ (tCase ci p c brs) -> Forall (expanded Σ Γ) (pparams p) /\ @@ -1291,15 +1319,469 @@ Proof. intros exp; depind exp; solve_discr => //; eauto. Qed. +Lemma expanded_tCoFix_inv Σ Γ mfix idx : + expanded Σ Γ (tCoFix mfix idx) -> + Forall (fun d : def term => expanded Σ (repeat 0 #|mfix| ++ Γ) (dbody d)) mfix. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + +Import PCUICOnOne. + +Module Red1Apps. + + + Inductive red1 (Σ : global_env) (Γ : context) : term -> term -> Type := + (** Reductions *) + (** Beta *) + | red_beta na t b a ts : + Σ ;;; Γ |- mkApps (tLambda na t b) (a :: ts) ⇝ mkApps (b {0 := a}) ts + + (** Let *) + | red_zeta na b t b' : + Σ ;;; Γ |- tLetIn na b t b' ⇝ b' {0 := b} + + | red_rel i body : + option_map decl_body (nth_error Γ i) = Some (Some body) -> + Σ ;;; Γ |- tRel i ⇝ lift0 (S i) body + + (** Case *) + | red_iota ci c u args p brs br : + nth_error brs c = Some br -> + #|args| = (ci.(ci_npar) + context_assumptions br.(bcontext))%nat -> + Σ ;;; Γ |- tCase ci p (mkApps (tConstruct ci.(ci_ind) c u) args) brs + ⇝ iota_red ci.(ci_npar) p args br + + (** Fix unfolding, with guard *) + | red_fix mfix idx args narg fn : + unfold_fix mfix idx = Some (narg, fn) -> + is_constructor narg args = true -> + Σ ;;; Γ |- mkApps (tFix mfix idx) args ⇝ mkApps fn args + + (** CoFix-case unfolding *) + | red_cofix_case ip p mfix idx args narg fn brs : + unfold_cofix mfix idx = Some (narg, fn) -> + Σ ;;; Γ |- tCase ip p (mkApps (tCoFix mfix idx) args) brs ⇝ + tCase ip p (mkApps fn args) brs + + (** CoFix-proj unfolding *) + | red_cofix_proj p mfix idx args narg fn : + unfold_cofix mfix idx = Some (narg, fn) -> + Σ ;;; Γ |- tProj p (mkApps (tCoFix mfix idx) args) ⇝ tProj p (mkApps fn args) + + (** Constant unfolding *) + | red_delta c decl body (isdecl : declared_constant Σ c decl) u : + decl.(cst_body) = Some body -> + Σ ;;; Γ |- tConst c u ⇝ subst_instance u body + + (** Proj *) + | red_proj p args u arg: + nth_error args (p.(proj_npars) + p.(proj_arg)) = Some arg -> + Σ ;;; Γ |- tProj p (mkApps (tConstruct p.(proj_ind) 0 u) args) ⇝ arg + + + | abs_red_l na M M' N : Σ ;;; Γ |- M ⇝ M' -> Σ ;;; Γ |- tLambda na M N ⇝ tLambda na M' N + | abs_red_r na M M' N : Σ ;;; Γ ,, vass na N |- M ⇝ M' -> Σ ;;; Γ |- tLambda na N M ⇝ tLambda na N M' + + | letin_red_def na b t b' r : Σ ;;; Γ |- b ⇝ r -> Σ ;;; Γ |- tLetIn na b t b' ⇝ tLetIn na r t b' + | letin_red_ty na b t b' r : Σ ;;; Γ |- t ⇝ r -> Σ ;;; Γ |- tLetIn na b t b' ⇝ tLetIn na b r b' + | letin_red_body na b t b' r : Σ ;;; Γ ,, vdef na b t |- b' ⇝ r -> Σ ;;; Γ |- tLetIn na b t b' ⇝ tLetIn na b t r + + | case_red_param ci p params' c brs : + OnOne2 (fun t u => Σ ;;; Γ |- t ⇝ u) p.(pparams) params' -> + Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci (set_pparams p params') c brs + + | case_red_return ci p preturn' c brs : + Σ ;;; Γ ,,, inst_case_predicate_context p |- p.(preturn) ⇝ preturn' -> + Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci (set_preturn p preturn') c brs + + | case_red_discr ci p c c' brs : + Σ ;;; Γ |- c ⇝ c' -> Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci p c' brs + + | case_red_brs ci p c brs brs' : + OnOne2 (fun br br' => + on_Trel_eq (fun t u => Σ ;;; Γ ,,, inst_case_branch_context p br |- t ⇝ u) bbody bcontext br br') + brs brs' -> + Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci p c brs' + + | proj_red p c c' : Σ ;;; Γ |- c ⇝ c' -> Σ ;;; Γ |- tProj p c ⇝ tProj p c' + + | app_red_l M1 N1 M2 : Σ ;;; Γ |- M1 ⇝ N1 -> ~~ isApp M1 -> ~~ isFix M1 -> M2 <> nil -> + Σ ;;; Γ |- mkApps M1 M2 ⇝ mkApps N1 M2 + + | app_red_r M2 N2 M1 : ~~ isApp M1 -> OnOne2 (fun M2 N2 => Σ ;;; Γ |- M2 ⇝ N2) M2 N2 -> Σ ;;; Γ |- mkApps M1 M2 ⇝ mkApps M1 N2 + + | app_fix_red_ty mfix0 mfix1 idx M2 : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ |- t ⇝ u) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + Σ ;;; Γ |- mkApps (tFix mfix0 idx) M2 ⇝ mkApps (tFix mfix1 idx) M2 + + | app_fix_red_body mfix0 mfix1 idx M2 : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> + Σ ;;; Γ |- mkApps (tFix mfix0 idx) M2 ⇝ mkApps (tFix mfix1 idx) M2 + + | prod_red_l na M1 M2 N1 : Σ ;;; Γ |- M1 ⇝ N1 -> Σ ;;; Γ |- tProd na M1 M2 ⇝ tProd na N1 M2 + | prod_red_r na M2 N2 M1 : Σ ;;; Γ ,, vass na M1 |- M2 ⇝ N2 -> + Σ ;;; Γ |- tProd na M1 M2 ⇝ tProd na M1 N2 + + | evar_red ev l l' : OnOne2 (fun t u => Σ ;;; Γ |- t ⇝ u) l l' -> Σ ;;; Γ |- tEvar ev l ⇝ tEvar ev l' + + | fix_red_ty mfix0 mfix1 idx : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ |- t ⇝ u) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + Σ ;;; Γ |- tFix mfix0 idx ⇝ tFix mfix1 idx + + | fix_red_body mfix0 mfix1 idx : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) dbody (fun x => (dname x, dtype x, rarg x))) + mfix0 mfix1 -> + Σ ;;; Γ |- tFix mfix0 idx ⇝ tFix mfix1 idx + + | cofix_red_ty mfix0 mfix1 idx : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ |- t ⇝ u) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + Σ ;;; Γ |- tCoFix mfix0 idx ⇝ tCoFix mfix1 idx + + | cofix_red_body mfix0 mfix1 idx : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> + Σ ;;; Γ |- tCoFix mfix0 idx ⇝ tCoFix mfix1 idx + where " Σ ;;; Γ |- t ⇝ u " := (red1 Σ Γ t u). + +Lemma red1_ind_all : + forall (Σ : global_env) (P : context -> term -> term -> Type), + + (forall (Γ : context) (na : aname) (t b a : term) ts, + P Γ (mkApps (tLambda na t b) (a :: ts)) (mkApps (b {0 := a}) ts)) -> + + (forall (Γ : context) (na : aname) (b t b' : term), P Γ (tLetIn na b t b') (b' {0 := b})) -> + + (forall (Γ : context) (i : nat) (body : term), + option_map decl_body (nth_error Γ i) = Some (Some body) -> P Γ (tRel i) ((lift0 (S i)) body)) -> + + (forall (Γ : context) (ci : case_info) (c : nat) (u : Instance.t) (args : list term) + (p : predicate term) (brs : list (branch term)) br, + nth_error brs c = Some br -> + #|args| = (ci.(ci_npar) + context_assumptions br.(bcontext))%nat -> + P Γ (tCase ci p (mkApps (tConstruct ci.(ci_ind) c u) args) brs) + (iota_red ci.(ci_npar) p args br)) -> + + (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 Γ (mkApps (tFix mfix idx) args) (mkApps fn args)) -> + + (forall (Γ : context) ci (p : predicate term) (mfix : mfixpoint term) (idx : nat) + (args : list term) (narg : nat) (fn : term) (brs : list (branch term)), + unfold_cofix mfix idx = Some (narg, fn) -> + P Γ (tCase ci p (mkApps (tCoFix mfix idx) args) brs) (tCase ci p (mkApps fn args) brs)) -> + + (forall (Γ : context) (p : projection) (mfix : mfixpoint term) (idx : nat) (args : list term) + (narg : nat) (fn : term), + unfold_cofix mfix idx = Some (narg, fn) -> P Γ (tProj p (mkApps (tCoFix mfix idx) args)) (tProj p (mkApps fn args))) -> + + (forall (Γ : context) c (decl : constant_body) (body : term), + declared_constant Σ c decl -> + forall u : Instance.t, cst_body decl = Some body -> P Γ (tConst c u) (subst_instance u body)) -> + + (forall (Γ : context) p (args : list term) (u : Instance.t) + (arg : term), + nth_error args (p.(proj_npars) + p.(proj_arg)) = Some arg -> + P Γ (tProj p (mkApps (tConstruct p.(proj_ind) 0 u) args)) arg) -> + + (forall (Γ : context) (na : aname) (M M' N : term), + red1 Σ Γ M M' -> P Γ M M' -> P Γ (tLambda na M N) (tLambda na M' N)) -> + + (forall (Γ : context) (na : aname) (M M' N : term), + red1 Σ (Γ,, vass na N) M M' -> P (Γ,, vass na N) M M' -> P Γ (tLambda na N M) (tLambda na N M')) -> + + (forall (Γ : context) (na : aname) (b t b' r : term), + red1 Σ Γ b r -> P Γ b r -> P Γ (tLetIn na b t b') (tLetIn na r t b')) -> + + (forall (Γ : context) (na : aname) (b t b' r : term), + red1 Σ Γ t r -> P Γ t r -> P Γ (tLetIn na b t b') (tLetIn na b r b')) -> + + (forall (Γ : context) (na : aname) (b t b' r : term), + red1 Σ (Γ,, vdef na b t) b' r -> P (Γ,, vdef na b t) b' r -> P Γ (tLetIn na b t b') (tLetIn na b t r)) -> + + (forall (Γ : context) (ci : case_info) p params' c brs, + OnOne2 (Trel_conj (red1 Σ Γ) (P Γ)) p.(pparams) params' -> + P Γ (tCase ci p c brs) + (tCase ci (set_pparams p params') c brs)) -> + + (forall (Γ : context) (ci : case_info) p preturn' c brs, + red1 Σ (Γ ,,, inst_case_predicate_context p) p.(preturn) preturn' -> + P (Γ ,,, inst_case_predicate_context p) p.(preturn) preturn' -> + P Γ (tCase ci p c brs) + (tCase ci (set_preturn p preturn') c brs)) -> + + (forall (Γ : context) (ind : case_info) (p : predicate term) (c c' : term) (brs : list (branch term)), + red1 Σ Γ c c' -> P Γ c c' -> P Γ (tCase ind p c brs) (tCase ind p c' brs)) -> + + (forall (Γ : context) ci p c brs brs', + OnOne2 (fun br br' => + (on_Trel_eq (Trel_conj (red1 Σ (Γ ,,, inst_case_branch_context p br)) + (P (Γ ,,, inst_case_branch_context p br))) + bbody bcontext br br')) brs brs' -> + P Γ (tCase ci p c brs) (tCase ci p c brs')) -> + + (forall (Γ : context) (p : projection) (c c' : term), red1 Σ Γ c c' -> P Γ c c' -> + P Γ (tProj p c) (tProj p c')) -> + + (forall (Γ : context) M1 N1 M2, red1 Σ Γ M1 N1 -> P Γ M1 N1 -> ~~ isApp M1 -> ~~ isFix M1 -> M2 <> [] -> + P Γ (mkApps M1 M2) (mkApps N1 M2)) -> + + (forall (Γ : context) M1 N2 M2, ~~ isApp M1 -> + OnOne2 (Trel_conj (red1 Σ Γ) (P Γ)) M2 N2 -> + P Γ (mkApps M1 M2) (mkApps M1 N2)) -> + + (forall (Γ : context) mfix0 mfix1 idx M2, + OnOne2 (on_Trel_eq (Trel_conj (fun t u => Σ ;;; Γ |- t ⇝ u) (P Γ)) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + P Γ (mkApps (tFix mfix0 idx) M2) (mkApps (tFix mfix1 idx) M2)) -> + + (forall (Γ : context) mfix0 mfix1 idx M2, + OnOne2 (on_Trel_eq (Trel_conj (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) (P (Γ ,,, fix_context mfix0))) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> + P Γ (mkApps (tFix mfix0 idx) M2) (mkApps (tFix mfix1 idx) M2)) -> + + (forall (Γ : context) (na : aname) (M1 M2 N1 : term), + red1 Σ Γ M1 N1 -> P Γ M1 N1 -> P Γ (tProd na M1 M2) (tProd na N1 M2)) -> + + (forall (Γ : context) (na : aname) (M2 N2 M1 : term), + red1 Σ (Γ,, vass na M1) M2 N2 -> P (Γ,, vass na M1) M2 N2 -> P Γ (tProd na M1 M2) (tProd na M1 N2)) -> + + (forall (Γ : context) (ev : nat) (l l' : list term), + OnOne2 (Trel_conj (red1 Σ Γ) (P Γ)) l l' -> P Γ (tEvar ev l) (tEvar ev l')) -> + + (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), + OnOne2 (on_Trel_eq (Trel_conj (red1 Σ Γ) (P Γ)) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + P Γ (tFix mfix0 idx) (tFix mfix1 idx)) -> + + (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), + OnOne2 (on_Trel_eq (Trel_conj (red1 Σ (Γ ,,, fix_context mfix0)) + (P (Γ ,,, fix_context mfix0))) dbody + (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> + P Γ (tFix mfix0 idx) (tFix mfix1 idx)) -> + + (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), + OnOne2 (on_Trel_eq (Trel_conj (red1 Σ Γ) (P Γ)) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + P Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx)) -> + + (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), + OnOne2 (on_Trel_eq (Trel_conj (red1 Σ (Γ ,,, fix_context mfix0)) + (P (Γ ,,, fix_context mfix0))) dbody + (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> + P Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx)) -> + + forall (Γ : context) (t t0 : term), red1 Σ Γ t t0 -> P Γ t t0. +Proof. + intros. rename X29 into Xlast. revert Γ t t0 Xlast. + fix aux 4. intros Γ t T. + move aux at top. + destruct 1; match goal with + | |- P _ (tFix _ _) (tFix _ _) => idtac + | |- P _ (tCoFix _ _) (tCoFix _ _) => idtac + | |- P _ (mkApps (tFix _ _) _) _ => idtac + | |- P _ (tCase _ _ (mkApps (tCoFix _ _) _) _) _ => idtac + | |- P _ (tProj _ (mkApps (tCoFix _ _) _)) _ => idtac + | H : _ |- _ => eapply H; eauto + end. + - eapply X3; eauto. + - eapply X4; eauto. + - eapply X5; eauto. + + - revert params' o. + generalize (pparams p). + fix auxl 3. + intros params params' []. + + constructor. split; auto. + + constructor. auto. + + - revert brs' o. + revert brs. + fix auxl 3. + intros l l' Hl. destruct Hl. + + simpl in *. constructor; intros; intuition auto. + + constructor. eapply auxl. apply Hl. + + - revert M2 N2 o. + fix auxl 3. + intros l l' Hl. destruct Hl. + + constructor. split; auto. + + constructor. auto. + + - eapply X20. + revert mfix0 mfix1 o. + fix auxl 3. + intros l l' Hl. destruct Hl. + + constructor; split; auto; intuition. + + constructor; try split; auto; intuition. + + - eapply X21. revert o. + generalize (fix_context mfix0). + intros c o. revert mfix0 mfix1 o. + fix auxl 3. + intros l l' Hl. destruct Hl. + + constructor. split; auto; intuition. + + constructor; try split; auto; intuition. + + - revert l l' o. + fix auxl 3. + intros l l' Hl. destruct Hl. + + constructor. split; auto. + + constructor. auto. + + - eapply X25. + revert mfix0 mfix1 o; fix auxl 3. + intros l l' Hl; destruct Hl; + constructor; try split; auto; intuition. + + - eapply X26. + revert o. generalize (fix_context mfix0). intros c Xnew. + revert mfix0 mfix1 Xnew; fix auxl 3; intros l l' Hl; + destruct Hl; constructor; try split; auto; intuition. + + - eapply X27. + revert mfix0 mfix1 o. + fix auxl 3; intros l l' Hl; destruct Hl; + constructor; try split; auto; intuition. + + - eapply X28. + revert o. generalize (fix_context mfix0). intros c new. + revert mfix0 mfix1 new; fix auxl 3; intros l l' Hl; destruct Hl; + constructor; try split; auto; intuition. +Defined. + +Lemma red_tApp Σ Γ t v u : + red1 Σ Γ t v -> + red1 Σ Γ (tApp t u) (tApp v u). +Proof. + induction 1. + all:try solve [eapply (app_red_l _ _ _ _ [u]) => //; econstructor; eauto]. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). + eapply red_beta. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). econstructor; eauto. + now apply is_constructor_app_ge. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). econstructor; eauto. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). econstructor; eauto. + now eapply OnOne2_app_r. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). now eapply app_fix_red_ty. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). now eapply app_fix_red_body. + - now eapply (app_fix_red_ty _ _ _ _ _ [_]). + - now eapply (app_fix_red_body _ _ _ _ _ [_]). +Qed. + +Lemma isLambda_red1 Σ Γ b b' : isLambda b -> red1 Σ Γ b b' -> isLambda b'. +Proof. + destruct b; simpl; try discriminate. + intros _ red. + depelim red; solve_discr; eauto. depelim o. +Qed. + +End Red1Apps. + +Lemma red1_red1apps Σ Γ t v : + red1 Σ Γ t v -> Red1Apps.red1 Σ Γ t v. +Proof. + intros red; induction red using red1_ind_all in |- *. + all:try solve [econstructor; eauto; solve_all; eauto]. + - eapply (Red1Apps.red_beta _ _ _ _ _ _ []). + - now eapply Red1Apps.red_tApp. + - remember (decompose_app (tApp M1 M2)). + destruct p as [hd args]. + edestruct (decompose_app_rec_inv' M1 [M2]). rewrite /decompose_app in Heqp. + cbn in Heqp. rewrite -Heqp. reflexivity. + destruct a as [napp [hm2 hm1]]. + symmetry in Heqp; eapply decompose_app_inv in Heqp. rewrite Heqp. + assert (tApp M1 N2 = mkApps hd (firstn x args ++ [N2])) as ->. + { now rewrite mkApps_app -hm1. } + rewrite -{1}(firstn_skipn x args) -hm2. eapply Red1Apps.app_red_r => //. + eapply OnOne2_app. now constructor. +Qed. + +Lemma head_nApp f : + ~~ isApp f -> head f = f. +Proof. + induction f => //. +Qed. + +Lemma expanded_mkApps_inv Σ Γ f v : + ~~ isApp f -> + ~~ (isConstruct f || isFix f || isRel f) -> + expanded Σ Γ (mkApps f v) -> + expanded Σ Γ f /\ Forall (expanded Σ Γ) v. +Proof. + intros happ hc. + induction v using rev_ind; cbn. + - intros exp; split => //. + - rewrite mkApps_app /=. + move/expanded_tApp_inv => e. + forward e. rewrite /isConstructApp /isFixApp head_mkApps. + rewrite head_nApp //. now move/negbTE: hc ->. + intuition auto. eapply app_Forall; eauto. +Qed. + + +Lemma arguments_mkApps f args : + ~~ isApp f -> + arguments (mkApps f args) = args. +Proof. + rewrite /arguments => isa. + rewrite decompose_app_mkApps //. +Qed. + +Lemma arguments_mkApps' f args : + arguments (mkApps f args) = arguments f ++ args. +Proof. + rewrite /arguments. + rewrite /decompose_app decompose_app_rec_mkApps //. + rewrite app_nil_r. + induction f in args |- * => //=. + rewrite IHf1. now rewrite (IHf1 [f2]) -app_assoc. +Qed. + +Lemma expanded_mkApps_inv' Σ Γ f : + expanded Σ Γ f -> + let args := arguments f in + Forall (expanded Σ Γ) args /\ + match head f with + | tRel n => exists m, nth_error Γ n = Some m /\ m <= #|args| + | tConstruct ind c u => exists mind idecl cdecl, + declared_constructor Σ (ind, c) mind idecl cdecl /\ #|args| >= ind_npars mind + context_assumptions (cstr_args cdecl) + | tFix mfix idx => exists d, + Forall + (fun d0 : def term => + isLambda (dbody d0) /\ + (let ctx := + rev_map (fun d1 : def term => (1 + rarg d1)%nat) + mfix in + expanded Σ (ctx ++ Γ) (dbody d0))) mfix /\ + args <> [] /\ + nth_error mfix idx = Some d /\ #|args| > rarg d + | _ => expanded Σ Γ (head f) + end. +Proof. + induction 1 using expanded_ind; cbn. + all: try solve [split; econstructor; eauto]. + all:rewrite head_mkApps /=. + - rewrite arguments_mkApps //. split => //. now exists m. + - destruct IHexpanded. rewrite arguments_mkApps'. split. + eapply app_Forall => //. + rewrite app_length. + destruct (head f6) => //; firstorder (eauto; try lia). + exists x. split => //. firstorder (eauto; try lia). + intros heq; apply H5. now eapply app_eq_nil in heq. + - rewrite arguments_mkApps //. split => //. now exists d. + - rewrite arguments_mkApps //. split => //. + firstorder. +Qed. + Lemma expanded_red {cf : checker_flags} {Σ : global_env_ext} Γ Γ' t v : wf Σ -> + (forall n body, option_map decl_body (nth_error Γ n) = Some (Some body) -> expanded Σ (skipn (S n) Γ') body) -> red1 Σ Γ t v -> expanded Σ Γ' t -> expanded Σ Γ' v. Proof. - intros wf red; induction red using red1_ind_all in Γ' |- *; intros exp. - - eapply expanded_tApp in exp as [] => //. + move=> wf wfΓ /red1_red1apps red. + induction red using Red1Apps.red1_ind_all in wfΓ, Γ' |- *; intros exp. + - eapply expanded_mkApps_inv in exp as [] => //. + eapply expanded_tLambda_inv in H. + depelim H0. + eapply expanded_mkApps => //. eapply (expanded_subst _ _ _ _ [] Γ') => //. now constructor. - now eapply expanded_tLambda_inv in H. - eapply expanded_tLetIn_inv in exp as []. eapply (expanded_subst _ _ _ _ [] Γ') => //. now constructor. - admit. @@ -1311,10 +1793,15 @@ Proof. - admit. - admit. - constructor. now eapply expanded_tLambda_inv in exp. - - constructor. eapply IHred. now eapply expanded_tLambda_inv in exp. + - constructor. eapply expanded_tLambda_inv in exp. + eapply IHred => //. + { intros [] body; cbn => //. rewrite skipn_S. eapply wfΓ. } - eapply expanded_tLetIn_inv in exp; now constructor. - eapply expanded_tLetIn_inv in exp. now constructor. - - eapply expanded_tLetIn_inv in exp. now constructor. + - eapply expanded_tLetIn_inv in exp. constructor; intuition eauto. + eapply IHred => //. + { intros [] ? => //=. intros [=]. subst b. now rewrite skipn_S skipn_0. + rewrite skipn_S. eapply wfΓ. } - eapply expanded_tCase_inv in exp as [? []]. constructor; eauto. cbn. solve_all. eapply OnOne2_impl_All_r; tea. intuition eauto. now apply X0. @@ -1327,12 +1814,72 @@ Proof. econstructor; eauto. solve_all. eapply OnOne2_impl_All_r; tea. intros x y [? ?]. intros [[] ?]. rewrite -e0. - solve_all. + solve_all. eapply e => //. admit. - eapply expanded_tProj_inv in exp. now econstructor. - - case_eq (isConstructApp (tApp M1 M2) || isFixApp (tApp M1 M2) || isRel (head (tApp M1 M2))). - intros h. - clear red. specialize (IHred Γ'). - depelim exp. - { destruct args using rev_case; cbn in *; noconf H1. rewrite mkApps_app in H1; noconf H1. -(* not the right way to go about this, we need some mkApps induction hypothesis *) -Abort. + - eapply expanded_mkApps_inv' in exp. + rewrite head_mkApps head_nApp in exp => //. + rewrite arguments_mkApps // in exp. destruct exp as []. + specialize (IHred Γ' wfΓ). + destruct M1; try solve [eapply expanded_mkApps => //; eauto]. + * depelim red; solve_discr. eapply wfΓ in e. + eapply expanded_mkApps => //. + rewrite -(firstn_skipn (S n) Γ'). eapply (expanded_lift _ _ _ _ []) => //. + destruct H3 as [m [hn ha]]. + rewrite firstn_length_le //. now eapply nth_error_Some_length in hn. + depelim o. + * depelim red; solve_discr. depelim o. + + - eapply expanded_mkApps_inv' in exp. + move: exp. + rewrite head_mkApps head_nApp // arguments_mkApps //. + intros []. + assert(Forall (expanded Σ Γ') N2). + { clear H1. solve_all. eapply OnOne2_impl_All_r; tea. cbn. intuition auto. } + destruct M1; try solve [eapply expanded_mkApps => //; eauto]. + + rewrite (OnOne2_length X) in H1. destruct H1 as [m []]; eapply expanded_tRel; tea. + + rewrite (OnOne2_length X) in H1. destruct H1 as [m [? [? []]]]; eapply expanded_tConstruct_app; tea. + + rewrite (OnOne2_length X) in H1. + destruct H1 as [d [? [? []]]]; eapply expanded_tFix; tea. + destruct N2; cbn in *; eauto. lia. intros eq; discriminate. + + - move/expanded_mkApps_inv': exp. + rewrite head_mkApps head_nApp // arguments_mkApps // => [] [] hargs [d [hf [hm2 [hnth harg]]]]. + eapply OnOne2_nth_error in hnth as [t' [hnth hor]]; tea. + eapply expanded_tFix; tea. + { clear hor. solve_all. eapply OnOne2_impl_All_r; tea. + intros ? ? [] []. noconf e. rewrite -H2. split => //. solve_all. + clear -X H1. + enough (rev_map (fun d1 : def term => S (rarg d1)) mfix0 = (rev_map (fun d1 : def term => S (rarg d1)) mfix1)) as <- => //. + clear -X. rewrite !rev_map_spec. f_equal. induction X. destruct p as []. cbn in *. congruence. cbn. congruence. } + { destruct hor; subst => //. destruct p as [? e]; noconf e. now congruence. } + + - move/expanded_mkApps_inv': exp. + rewrite head_mkApps head_nApp // arguments_mkApps // => [] [] hargs [d [hf [hm2 [hnth harg]]]]. + eapply OnOne2_nth_error in hnth as [t' [hnth hor]]; tea. + eapply expanded_tFix; tea. + { clear hor. solve_all. + assert (rev_map (fun d1 : def term => S (rarg d1)) mfix0 = (rev_map (fun d1 : def term => S (rarg d1)) mfix1)). + { clear -X. rewrite !rev_map_spec. f_equal. induction X. destruct p as []. cbn in *. congruence. cbn. congruence. } + rewrite -H. + eapply OnOne2_impl_All_r; tea. + intros ? ? [] []. noconf e. destruct p. + eapply Red1Apps.isLambda_red1 in r; tea. split => //. + set(Γ'' := rev_map (fun d1 : def term => S (rarg d1)) mfix0). + eapply e => //. + { intros n b hnth'. admit. } } + { destruct hor; subst => //. destruct p as [? e]; noconf e. now congruence. } + - eapply expanded_tProd. + - constructor. + - constructor. + eapply expanded_tEvar_inv in exp. + solve_all; eapply OnOne2_impl_All_r; tea. intuition eauto. + now eapply X0. + - depelim exp; solve_discr. now cbn in H. + - depelim exp; solve_discr. now cbn in H. + - eapply expanded_tCoFix_inv in exp. econstructor. + rewrite -(OnOne2_length X). solve_all; eapply OnOne2_impl_All_r; tea; intuition eauto. + destruct X0. noconf e. now rewrite -H1. + - eapply expanded_tCoFix_inv in exp. econstructor. + rewrite -(OnOne2_length X). solve_all; eapply OnOne2_impl_All_r; tea; intuition eauto. + destruct X0. destruct p. noconf e. eapply e0 => //. admit. +Qed. \ No newline at end of file From 2f5fed3109aaa1f7d8d76328cf2fb2f1084cce2d Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 12 Sep 2023 12:52:42 +0200 Subject: [PATCH 50/61] Finished preservation of expansion proof --- pcuic/theories/PCUICWcbvEval.v | 295 +++++++++++++++++++++++++++++++-- 1 file changed, 284 insertions(+), 11 deletions(-) diff --git a/pcuic/theories/PCUICWcbvEval.v b/pcuic/theories/PCUICWcbvEval.v index ddd600473..6fdc13e16 100644 --- a/pcuic/theories/PCUICWcbvEval.v +++ b/pcuic/theories/PCUICWcbvEval.v @@ -1769,13 +1769,203 @@ Proof. firstorder. Qed. +Lemma All_fold_map_context (P : context -> context_decl -> Type) (f : term -> term) ctx : + All_fold P (map_context f ctx) <~> All_fold (fun Γ d => P (map_context f Γ) (map_decl f d)) ctx. +Proof. + induction ctx. + - split; constructor. + - cbn. split; intros H; depelim H; constructor; auto; now apply IHctx. +Qed. + +Lemma expanded_fix_subst Σ a k b Γ Γs Δ : + #|Γ| = k -> + Forall2 (fun n arg => forall args Γ' k', + Forall (expanded Σ (Γ' ++ Δ)) args -> + n <= #|args| -> + #|Γ'| = k' -> + expanded Σ (Γ' ++ Δ) (mkApps (lift0 k' arg) args)) Γs a -> + expanded Σ (Γ ++ Γs ++ Δ) b -> + expanded Σ (Γ ++ Δ) (subst a k b). +Proof. + intros Hk Hs. + remember (Γ ++ _ ++ Δ)%list as Γ_. + intros exp; revert Γ k Hk HeqΓ_. + induction exp using expanded_ind; intros Γ' k Hk ->. + all:try solve[ cbn; econstructor => // ]. + 2,7:solve[ cbn; econstructor => //; solve_all ]. + - rewrite subst_mkApps /=. + destruct (Nat.leb_spec k n). + destruct (nth_error a _) eqn:hnth. + * pose proof (Forall2_length Hs). + pose proof (Forall2_nth_error_Some_r _ _ _ _ _ _ _ hnth Hs) as [t' [hnth' hargs]]. + eapply hargs => //. + eapply Forall_map. + 2:{ len. rewrite nth_error_app_ge in H. lia. subst k. + rewrite nth_error_app_lt in H. + eapply nth_error_Some_length in hnth'. lia. rewrite hnth' in H. noconf H. exact H0. } + solve_all. + * rewrite nth_error_app_ge in H. lia. + eapply nth_error_None in hnth. + pose proof (Forall2_length Hs). + rewrite nth_error_app_ge in H. lia. + eapply expanded_tRel. rewrite nth_error_app_ge. lia. erewrite <- H. + lia_f_equal. len. solve_all. + * rewrite nth_error_app_lt in H. lia. + eapply expanded_tRel. rewrite nth_error_app_lt. lia. tea. now len. + solve_all. + - cbn. econstructor. + eapply (IHexp (0 :: Γ') (S k)); cbn; auto; try lia. + - cbn. econstructor. apply IHexp1; auto. + eapply (IHexp2 (0 :: Γ') (S k)); cbn; auto; lia. + - rewrite subst_mkApps. + destruct (isConstruct (subst a k f6) || isFix (subst a k f6) || isRel (subst a k f6)) eqn:eqc. + specialize (IHexp _ _ Hk eq_refl). + eapply expanded_mkApps_expanded => //. solve_all. + eapply expanded_mkApps => //. now eapply IHexp. solve_all. + - cbn. econstructor. eauto. cbn. solve_all. solve_all. + specialize (H1 (repeat 0 #|bcontext x| ++ Γ') (#|bcontext x| + k)%nat). + forward H1 by len. + forward H1. now rewrite app_assoc. + rewrite /id. rewrite app_assoc. apply H1. + - rewrite subst_mkApps. cbn. + eapply expanded_tFix. + + solve_all. now eapply isLambda_subst. + specialize (a0 + (rev_map (fun d0 : def term => S (rarg d0)) + (map (map_def (subst a k) (subst a (#|mfix| + k))) mfix) ++ Γ') (#|mfix| + k)%nat). + forward a0 by len. + forward a0. { rewrite app_assoc. f_equal. f_equal. + rewrite !rev_map_spec. f_equal. now rewrite map_map_compose /=. } + rewrite app_assoc. eapply a0. + + solve_all. + + now destruct args. + + rewrite nth_error_map /= H4 //. + + len. + - cbn. constructor. + solve_all. + specialize (a0 (repeat 0 #|mfix| ++ Γ') (#|mfix| + k)%nat). + forward a0 by len. + forward a0. { rewrite app_assoc //. } + rewrite app_assoc. eapply a0. + - rewrite subst_mkApps. cbn. + eapply expanded_tConstruct_app; tea. now len. + solve_all. +Qed. + +Lemma expanded_unfold_fix Σ Γ' mfix idx narg fn : + unfold_fix mfix idx = Some (narg, fn) -> + All (fun d0 : def term => isLambda (dbody d0) /\ expanded Σ (rev_map (fun d1 : def term => S (rarg d1)) mfix ++ Γ') (dbody d0)) mfix -> + expanded Σ Γ' fn. +Proof. + unfold unfold_fix. + destruct nth_error eqn:e => //. + intros [= <- <-] hf. + pose proof (nth_error_all e hf) as [hl hf']. + eapply (expanded_fix_subst _ _ _ _ []) => //; tea. + rewrite rev_map_spec. + eapply Forall2_from_nth_error. len. + intros n rarg f. len. intros hn hrarg hnthf args Γ'' k' hargs hrarg' <-. + eapply PCUICParallelReductionConfluence.nth_error_fix_subst in hnthf. subst f. + move: hrarg. + rewrite nth_error_rev; len. rewrite List.rev_involutive nth_error_map. + intros hrarg. + destruct (nth_error mfix (_ - _)) eqn:e'. cbn in hrarg. noconf hrarg. + eapply expanded_tFix => //. solve_all. + eapply expanded_lift; len. rewrite !rev_map_spec in H1 *. + rewrite map_map => //. destruct args => //. cbn in hrarg'. lia. + rewrite nth_error_map /= e' /= //. cbn. lia. + eapply nth_error_None in e'. lia. +Qed. + +Lemma expanded_unfold_cofix Σ Γ' mfix idx narg fn : + unfold_cofix mfix idx = Some (narg, fn) -> + All (fun d0 : def term => expanded Σ (repeat 0 #|mfix| ++ Γ') (dbody d0)) mfix -> + expanded Σ Γ' fn. +Proof. + unfold unfold_cofix. + destruct nth_error eqn:e => //. + intros [= <- <-] hf. + pose proof (nth_error_all e hf) as hf'. + eapply (expanded_subst _ _ _ _ []) => //; tea. + rewrite /cofix_subst. + generalize #|mfix|. + induction n; repeat constructor; eauto. solve_all. + len. +Qed. + +Lemma expanded_weakening_env {cf : checker_flags} Σ Σ' Γ t : + wf Σ' -> + extends Σ Σ' -> + expanded Σ Γ t -> expanded Σ' Γ t. +Proof. + intros w s. + eapply expanded_ind; intros. + all:try solve [econstructor; eauto]. + - econstructor; eauto. solve_all. sq. eapply All_fold_impl; tea. cbn. + intros. now rewrite repeat_app in H6. + - eapply expanded_tFix; eauto. solve_all. + - eapply expanded_tConstruct_app; eauto. + eapply PCUICWeakeningEnv.weakening_env_declared_constructor; tea. +Qed. + +Lemma expanded_global_env_constant {cf : checker_flags} Σ c decl : + wf Σ -> + expanded_global_env Σ -> + declared_constant Σ c decl -> + ForOption (expanded Σ []) (cst_body decl). +Proof. + intros wf; destruct Σ as [Σ univs] => /=. cbn. + unfold expanded_global_env. cbn. + induction 1; cbn => //. + intros [->|H']. + - depelim H0. + destruct decl as [? [] ?]; cbn in *. + constructor. cbn. + eapply expanded_weakening_env; tea. + eapply extends_strictly_on_decls_extends. + split => //=. eapply incl_cs_refl. 2:eapply Retroknowledge.extends_refl. + set (cb := ConstantDecl _). now exists [(c, cb)]. + constructor. + - forward IHexpanded_global_declarations. + destruct wf. cbn in *. split => //. + cbn. now depelim o0. + eapply IHexpanded_global_declarations in H'. + destruct decl as [? [] ?]; cbn in * => //. 2:constructor. + depelim H'. constructor. + eapply expanded_weakening_env; tea. + eapply extends_strictly_on_decls_extends. + split => //=. eapply incl_cs_refl. 2:eapply Retroknowledge.extends_refl. + now exists [decl0]. +Qed. + +Lemma All_fold_nth_error P (ctx : context) n b : + All_fold P ctx -> nth_error ctx n = Some b -> + P (skipn (S n) ctx) b. +Proof. + induction 1 in n, b |- *. + - rewrite nth_error_nil //. + - destruct n => //=. + * intros [= <-]. now rewrite skipn_S skipn_0. + * intros hnth. now rewrite skipn_S. +Qed. + +Lemma skipn_repeat {A} k n (a : A) : + skipn n (repeat a k) = repeat a (k - n). +Proof. + induction n in k |- *. + - rewrite skipn_0. lia_f_equal. + - destruct k => //=. + rewrite skipn_S IHn. lia_f_equal. +Qed. + Lemma expanded_red {cf : checker_flags} {Σ : global_env_ext} Γ Γ' t v : wf Σ -> + expanded_global_env Σ -> (forall n body, option_map decl_body (nth_error Γ n) = Some (Some body) -> expanded Σ (skipn (S n) Γ') body) -> red1 Σ Γ t v -> expanded Σ Γ' t -> expanded Σ Γ' v. Proof. - move=> wf wfΓ /red1_red1apps red. + move=> wf expΣ wfΓ /red1_red1apps red. induction red using Red1Apps.red1_ind_all in wfΓ, Γ' |- *; intros exp. - eapply expanded_mkApps_inv in exp as [] => //. eapply expanded_tLambda_inv in H. @@ -1784,14 +1974,53 @@ Proof. eapply (expanded_subst _ _ _ _ [] Γ') => //. now constructor. - eapply expanded_tLetIn_inv in exp as []. eapply (expanded_subst _ _ _ _ [] Γ') => //. now constructor. - - admit. + - rewrite -(firstn_skipn (S i) Γ'). + eapply expanded_mkApps_inv' in exp; cbn in exp. + destruct exp as [_ [m []]]. eapply nth_error_Some_length in H0. + eapply (expanded_lift _ _ _ _ []) => //. + rewrite firstn_length_le; try lia. + now eapply wfΓ. + - eapply expanded_tCase_inv in exp as [? []]. + unfold iota_red. + move/expanded_mkApps_inv': H3 => /=. + rewrite arguments_mkApps // head_mkApps //=. + intros [hargs [mind [idecl [cdecl [declc hargs']]]]]. + eapply nth_error_forall in H2; tea. + destruct H2 as [[hbctx] hbod]. + eapply (expanded_subst _ _ _ _ [] _) => //. + + eapply Forall_rev, Forall_skipn => //. + + len. replace #|skipn (ci_npar ci) args| with (context_assumptions (inst_case_branch_context p br)). + eapply expanded_let_expansion; len. red. + sq. rewrite /inst_case_branch_context /inst_case_context /subst_context. + eapply PCUICParallelReduction.All_fold_fold_context_k. + eapply All_fold_map_context, All_fold_impl; tea; cbn. + intros ? ? fo. len. destruct x as [? [] ?] => //; constructor. + cbn in fo. depelim fo. eapply (expanded_subst _ _ _ _ (repeat 0 #|Γ0|) _); len. + eapply Forall_rev; eauto. + eapply expanded_subst_instance. rewrite app_assoc. now eapply expanded_weakening. + rewrite skipn_length. len. + - move/expanded_mkApps_inv': exp. cbn. + rewrite arguments_mkApps // head_mkApps //=. + move=> [hargs [d [hf [hargs' []]]]] hnth hrarg. + eapply expanded_mkApps_expanded => //; solve_all. + eapply expanded_unfold_fix in H; tea. - eapply expanded_tCase_inv in exp as [? []]. - unfold iota_red. admit. - - admit. - - admit. - - eapply expanded_tProj_inv in exp. econstructor. admit. - - admit. - - admit. + constructor => //. + eapply expanded_mkApps_inv' in H2. move: H2; rewrite arguments_mkApps // head_mkApps //=. + intros [hargs' hcof]. cbn in hcof. eapply expanded_tCoFix_inv in hcof. + eapply expanded_unfold_cofix in H; tea. eapply expanded_mkApps; tea => //. solve_all. + - eapply expanded_tProj_inv in exp. econstructor. + eapply expanded_mkApps_inv' in exp. move: exp; rewrite arguments_mkApps // head_mkApps //=. + intros [hargs' hcof]. cbn in hcof. eapply expanded_tCoFix_inv in hcof. + eapply expanded_mkApps => //. + eapply expanded_unfold_cofix in H; tea. solve_all. + - eapply expanded_subst_instance. + eapply expanded_global_env_constant in expΣ; tea. + eapply (expanded_weakening _ []). rewrite H0 in expΣ. now depelim expΣ. + - eapply expanded_tProj_inv in exp. + move/expanded_mkApps_inv': exp. rewrite arguments_mkApps // head_mkApps //=. + intros []. + eapply nth_error_forall in H0; tea. - constructor. now eapply expanded_tLambda_inv in exp. - constructor. eapply expanded_tLambda_inv in exp. eapply IHred => //. @@ -1814,7 +2043,30 @@ Proof. econstructor; eauto. solve_all. eapply OnOne2_impl_All_r; tea. intros x y [? ?]. intros [[] ?]. rewrite -e0. - solve_all. eapply e => //. admit. + solve_all. eapply e => //. + intros n b. + clear -H H2 wfΓ. + destruct nth_error eqn:e' => //. + cbn. intros [=]. destruct c as [? [] ?]. cbn in *. noconf H1. + eapply nth_error_app_inv in e' as [[]|[]]. + { rewrite inst_case_branch_context_length in H0. destruct H2. + rewrite /inst_case_branch_context /inst_case_context in H1. + destruct (nth_error (bcontext x) n) eqn:e. + 2:{ eapply nth_error_None in e. rewrite skipn_app skipn_all2. len. cbn. len. } + rewrite /subst_context in H1. + erewrite nth_error_fold_context_k in H1. 4:{ rewrite nth_error_map e //. } 3:len. 2:exact []. + len in H1. noconf H1. destruct c as [? [] ?]; noconf H1. + rewrite skipn_app. len. eapply All_fold_nth_error in X; tea. cbn in X. depelim X. + rewrite skipn_length in H1. + eapply expanded_subst. rewrite skipn_length. len. + replace (S n - #|bcontext x|) with 0. 2:{ lia. } rewrite skipn_0. eapply Forall_rev. solve_all. + len. rewrite app_assoc. eapply expanded_weakening. eapply expanded_subst_instance. + now rewrite skipn_repeat. } + { rewrite inst_case_branch_context_length in H0, H1. + rewrite skipn_app skipn_all2 /=; len. + replace (S n - #|bcontext x|) with (S (n - #|bcontext x|)). 2:lia. + eapply wfΓ. rewrite H1 //. } + noconf H1. - eapply expanded_tProj_inv in exp. now econstructor. - eapply expanded_mkApps_inv' in exp. rewrite head_mkApps head_nApp in exp => //. @@ -1866,7 +2118,18 @@ Proof. eapply Red1Apps.isLambda_red1 in r; tea. split => //. set(Γ'' := rev_map (fun d1 : def term => S (rarg d1)) mfix0). eapply e => //. - { intros n b hnth'. admit. } } + { intros n b hnth'. + destruct (nth_error (fix_context mfix0) n) eqn:e'. + rewrite nth_error_app_lt in hnth'. now eapply nth_error_Some_length in e'. + rewrite e' in hnth'. noconf hnth'. + pose proof (PCUICParallelReductionConfluence.fix_context_assumption_context mfix0). + eapply PCUICParallelReductionConfluence.nth_error_assumption_context in H6; tea. congruence. + rewrite /Γ''. + eapply nth_error_None in e'. len in e'. + rewrite nth_error_app_ge in hnth' => //. now len. + rewrite skipn_app skipn_all2. len. + cbn. len. replace (S n - #|mfix0|) with (S (n - #|mfix0|)). 2:{ lia. } + eapply wfΓ. now len in hnth'. } } { destruct hor; subst => //. destruct p as [? e]; noconf e. now congruence. } - eapply expanded_tProd. - constructor. @@ -1881,5 +2144,15 @@ Proof. destruct X0. noconf e. now rewrite -H1. - eapply expanded_tCoFix_inv in exp. econstructor. rewrite -(OnOne2_length X). solve_all; eapply OnOne2_impl_All_r; tea; intuition eauto. - destruct X0. destruct p. noconf e. eapply e0 => //. admit. + destruct X0. destruct p. noconf e. eapply e0 => //. + intros n b. + destruct nth_error eqn:e' => //. + cbn. intros [=]. destruct c as [? [] ?]. cbn in *. noconf H4. + eapply nth_error_app_inv in e' as [[]|[]]. + { pose proof (PCUICParallelReductionConfluence.fix_context_assumption_context mfix0). + eapply PCUICParallelReductionConfluence.nth_error_assumption_context in H5; tea. cbn in H5. congruence. } + len in H3. len in H4. + rewrite skipn_app skipn_all2; len. + replace (S n - #|mfix0|) with (S (n - #|mfix0|)) by lia. eapply wfΓ. rewrite H4 /= //. + noconf H4. Qed. \ No newline at end of file From 681a3827e5371dacb705aa594ee1145d39e71fc9 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 12 Sep 2023 17:29:43 +0200 Subject: [PATCH 51/61] Refactor a bit, work on side conditions for let expansion --- erasure-plugin/theories/ErasureCorrectness.v | 88 +- erasure/theories/EEtaExpandedFix.v | 2 +- pcuic/theories/PCUICClassification.v | 6 - pcuic/theories/PCUICConfluence.v | 2 + pcuic/theories/PCUICConversion.v | 4 +- pcuic/theories/PCUICEtaExpand.v | 990 +++++++++++++++++- pcuic/theories/PCUICWcbvEval.v | 937 ----------------- pcuic/theories/PCUICWellScopedCumulativity.v | 2 + .../Typing/PCUICContextConversionTyp.v | 2 + pcuic/theories/utils/PCUICAstUtils.v | 24 + template-coq/theories/TypingWf.v | 1 - template-coq/theories/WcbvEval.v | 4 +- template-coq/theories/WfAst.v | 2 +- utils/theories/MCList.v | 1 + 14 files changed, 1077 insertions(+), 988 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 87e8f948f..dbbaacecf 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -866,11 +866,22 @@ Lemma compile_evalue_strip (Σer : EEnvMap.GlobalContextMap.t) p : firstorder_evalue Σer p -> compile_evalue_box (ERemoveParams.strip Σer p) [] = compile_evalue_box_strip Σer p []. Proof. -Admitted. + move: p. + apply: firstorder_evalue_elim; intros. + rewrite ERemoveParams.strip_mkApps //=. rewrite H. + rewrite compile_evalue_box_mkApps. + rewrite (compile_evalue_box_strip_mkApps (npars:=npars)). + rewrite lookup_inductive_pars_spec //. + rewrite !app_nil_r !skipn_map. f_equal. + rewrite map_map. + ELiftSubst.solve_all. +Qed. Arguments PCUICFirstorder.firstorder_ind _ _ : clear implicits. Section PCUICExpandLets. + Import PCUICAst. + Import PCUICEnvironment. Import PCUICExpandLets PCUICExpandLetsCorrectness. Lemma trans_axiom_free Σ : axiom_free Σ -> axiom_free (trans_global_env Σ). @@ -883,6 +894,29 @@ Section PCUICExpandLets. destruct c as [? [] ? ?] => //. Qed. + Lemma expand_lets_preserves_fo (Σ : global_env_ext) Γ v : + wf Σ -> + PCUICFirstorder.firstorder_value Σ Γ v -> + PCUICFirstorder.firstorder_value (trans_global Σ) (trans_local Γ) (trans v). + Proof. + intros wfΣ; move: v; apply: PCUICFirstorder.firstorder_value_inds. + intros i n ui u args pandi hty hargs ihargs isp. + rewrite trans_mkApps. econstructor. + apply expand_lets_sound in hty. + rewrite !trans_mkApps in hty. cbn in hty. exact hty. + ELiftSubst.solve_all. + move/negbTE: isp. + rewrite /PCUICFirstorder.isPropositional. + rewrite /lookup_inductive /lookup_inductive_gen /lookup_minductive_gen trans_lookup. + destruct lookup_env => //. destruct g => //. cbn. rewrite nth_error_mapi. + destruct nth_error => //=. + rewrite /PCUICFirstorder.isPropositionalArity. + rewrite (trans_destArity []). destruct destArity => //. + destruct p => //. now intros ->. + Qed. +End PCUICExpandLets. + + Section pipeline_theorem. Instance cf : checker_flags := extraction_checker_flags. @@ -894,7 +928,7 @@ Section pipeline_theorem. Variable t : PCUICAst.term. Variable expt : PCUICEtaExpand.expanded Σ.1 [] t. - + Variable axfree : axiom_free Σ. Variable v : PCUICAst.term. Variable i : Kernames.inductive. @@ -923,7 +957,9 @@ Section pipeline_theorem. - red. cbn. split; eauto. eexists. eapply PCUICClassification.subject_reduction_eval; eauto. - - todo "preservation of eta expandedness". + - eapply (PCUICClassification.wcbveval_red (Σ := Σ)) in X; tea. + eapply PCUICEtaExpand.expanded_red in X; tea. apply HΣ. + intros ? ?; rewrite nth_error_nil => //. - cbn. todo "normalization". - todo "normalization". Qed. @@ -952,19 +988,51 @@ Section pipeline_theorem. cbn [fst snd]. intros h. destruct_compose. + destruct Heval. + assert (eqtr : PCUICExpandLets.trans v = v). + { clear -hv. + move: v hv. + eapply PCUICFirstorder.firstorder_value_inds. + intros. + rewrite PCUICExpandLetsCorrectness.trans_mkApps /=. + f_equal. ELiftSubst.solve_all. } assert (PCUICFirstorder.firstorder_value (PCUICExpandLets.trans_global_env Σ.1, Σ.2) [] v). - { todo "expand lets preserves fo values". } + { eapply expand_lets_preserves_fo in hv; eauto. now rewrite eqtr in hv. } + assert (Normalisation': PCUICSN.NormalizationIn (PCUICExpandLets.trans_global Σ)). - { destruct h as [[] ?]. apply H0. cbn. apply X. } + { destruct h as [[] ?]. apply H0. cbn. apply X0. } set (Σ' := build_global_env_map _). set (p := transform erase_transform _ _). - pose proof (@erase_tranform_firstorder _ h v i u args Normalisation'). - forward H0. - { todo "preserves typing of fo values". } + pose proof (@erase_tranform_firstorder _ h v i u (List.map PCUICExpandLets.trans args) Normalisation'). forward H0. - { cbn. todo "preserves axiom freeness". } + { cbn. rewrite -eqtr. + eapply (PCUICClassification.subject_reduction_eval (Σ := Σ)) in X; tea. + eapply PCUICExpandLetsCorrectness.expand_lets_sound in X. + now rewrite PCUICExpandLetsCorrectness.trans_mkApps /= in X. } + forward H0. { cbn. now eapply trans_axiom_free. } forward H0. - { cbn. todo "preserves fo ind". } + { cbn. clear -HΣ fo. + move: fo. + rewrite /PCUICFirstorder.firstorder_ind /= PCUICExpandLetsCorrectness.trans_lookup /=. + destruct PCUICAst.PCUICEnvironment.lookup_env => //. destruct g => //=. + unfold PCUICFirstorder.firstorder_mutind. cbn. + move/andP => [] -> /=. ELiftSubst.solve_all. + eapply forallb_Forall in b. apply forallb_Forall. + ELiftSubst.solve_all. eapply All_mapi. + eapply All_Alli; tea. cbn. + intros n []. rewrite /PCUICFirstorder.firstorder_oneind /=. + move/andP => [] hf ->; rewrite andb_true_r. + eapply forallb_Forall in hf. apply forallb_Forall. + ELiftSubst.solve_all. + rewrite /PCUICFirstorder.firstorder_con in H *. + eapply alli_Alli. eapply alli_Alli in H. + destruct x; cbn in *. + eapply Alli_rev. + ELiftSubst.solve_all. + + re + destruct m; cbn in *. + forward H0. { cbn. todo "preserves values". } specialize (H0 _ eq_refl). diff --git a/erasure/theories/EEtaExpandedFix.v b/erasure/theories/EEtaExpandedFix.v index 92a783a73..8c14b356c 100644 --- a/erasure/theories/EEtaExpandedFix.v +++ b/erasure/theories/EEtaExpandedFix.v @@ -1408,7 +1408,7 @@ Proof. split. eapply All2_app => //. rewrite -H3. eauto. destruct s. * destruct p; solve_discr. noconf H2. - left. split.w + left. split. unfold isStuckFix'; rewrite e1. len. eapply Nat.leb_le. lia. now rewrite -[tApp _ _](mkApps_app _ _ [av]). * right. len. eapply isEtaExp_fixapp_mon; tea. lia. diff --git a/pcuic/theories/PCUICClassification.v b/pcuic/theories/PCUICClassification.v index 47cbe786d..55908fed7 100644 --- a/pcuic/theories/PCUICClassification.v +++ b/pcuic/theories/PCUICClassification.v @@ -126,12 +126,6 @@ Lemma All2_map_left' {A B C} (P : A -> B -> Type) l l' (f : C -> A) : All2 P (map f l) l' -> All2 (fun x y => P (f x) y) l l'. Proof. intros. rewrite - (map_id l') in X. eapply All2_map_inv; eauto. Qed. -Lemma head_mkApps t args : head (mkApps t args) = head t. -Proof. - induction args using rev_ind; simpl; auto. - now rewrite mkApps_app /= head_tapp. -Qed. - Section Spines. Context {cf : checker_flags}. Context {Σ : global_env_ext}. diff --git a/pcuic/theories/PCUICConfluence.v b/pcuic/theories/PCUICConfluence.v index 8f7795231..5427cdec7 100644 --- a/pcuic/theories/PCUICConfluence.v +++ b/pcuic/theories/PCUICConfluence.v @@ -24,6 +24,8 @@ Require Import CRelationClasses CMorphisms. Require Import Equations.Prop.DepElim. Require Import Equations.Type.Relation Equations.Type.Relation_Properties. +Local Ltac intuition_solver ::= auto with *. + #[global] Instance red_Refl Σ Γ : Reflexive (red Σ Γ) := refl_red Σ Γ. diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index 79c797d99..6e7c65626 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -3268,8 +3268,8 @@ Proof. unfold R_global_instance, R_global_instance_gen. destruct global_variance_gen. { induction i in l, i' |- *; destruct l, i'; simpl; auto; try lia; try easy. - * specialize (IHi i' []). simpl in IHi. intuition. - * intros []. intuition. + * specialize (IHi i' []). simpl in IHi. intuition auto with all. + * intros []. intuition eauto with all. } { unfold R_universe_instance. intros H % Forall2_length. now rewrite !map_length in H. } diff --git a/pcuic/theories/PCUICEtaExpand.v b/pcuic/theories/PCUICEtaExpand.v index b4db67093..2b097e2e1 100644 --- a/pcuic/theories/PCUICEtaExpand.v +++ b/pcuic/theories/PCUICEtaExpand.v @@ -1,7 +1,42 @@ -From Coq Require Import ssreflect. +From Coq Require Import ssreflect ssrbool. From Equations Require Import Equations. From MetaCoq.Utils Require Import utils. -From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping PCUICProgram PCUICCSubst. +From MetaCoq.Common Require Import config. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping + PCUICReduction PCUICProgram PCUICLiftSubst PCUICCSubst PCUICUnivSubst. + +(* todo move *) +Lemma All_fold_tip {A : Type} (P : list A -> A -> Type) {x} : All_fold P [x] -> P [] x. +Proof. + intros a; now depelim a. +Qed. + +Lemma All_tip {A} {P : A -> Type} {a : A} : P a <~> All P [a]. +Proof. split; intros. repeat constructor; auto. now depelim X. Qed. + +Lemma nApp_mkApps t f args : + t = mkApps f args -> ~~ isApp t -> t = f /\ args = []. +Proof. + intros -> napp. + destruct args using rev_case; cbn in *; solve_discr; try discriminate => //. split => //. + now rewrite mkApps_app /= in napp. +Qed. + +Ltac solve_discr := + try progress (prepare_discr; finish_discr; cbn[mkApps] in * ); + try match goal with + H : mkApps _ _ = mkApps ?f ?l |- _ => + eapply mkApps_eq_inj in H as [? ?]; [|easy|easy]; subst; try intuition congruence; try noconf H + | H : ?t = mkApps ?f ?l |- _ => + (change t with (mkApps t []) in H ; + eapply mkApps_eq_inj in H as [? ?]; [|easy|easy]; subst; try intuition congruence; try noconf H) || + (eapply nApp_mkApps in H as [? ?]; [|easy]; subst) + | H : mkApps ?f ?l = ?t |- _ => + change t with (mkApps t []) in H ; + eapply mkApps_eq_inj in H as [? ?]; [|easy|easy]; subst; try intuition congruence; try noconf H + end. + +(* end move *) Definition isConstruct t := match t with tConstruct _ _ _ => true | _ => false end. @@ -188,10 +223,6 @@ Definition expanded_global_env (g : global_env) := Definition expanded_pcuic_program (p : pcuic_program) := expanded_global_env p.1 /\ expanded p.1 [] p.2. - -Lemma All_tip {A} {P : A -> Type} {a : A} : P a <~> All P [a]. -Proof. split; intros. repeat constructor; auto. now depelim X. Qed. - Lemma expanded_mkApps_expanded {Σ Γ f args} : expanded Σ Γ f -> All (expanded Σ Γ) args -> expanded Σ Γ (mkApps f args). @@ -200,15 +231,11 @@ Proof. destruct (isConstruct f || isFix f || isRel f) eqn:eqc. destruct f => //. - depelim H; solve_discr. eapply expanded_tRel; tea. cbn in Hle. lia. solve_all. - destruct args0 using rev_case; cbn in *; subst. cbn in H. congruence. - rewrite mkApps_app in H2; noconf H2. + now rewrite eqc in H. - depelim H; solve_discr. - destruct args0 using rev_case; cbn in *; subst. cbn in H. congruence. - rewrite mkApps_app in H2; noconf H2. + now rewrite eqc in H. eapply expanded_tConstruct_app; tea. cbn in H0. lia. solve_all. - - depelim H; solve_discr. - destruct args0 using rev_case; cbn in *; subst. cbn in H. congruence. - rewrite mkApps_app in H2; noconf H2. + - depelim H; solve_discr. now rewrite eqc in H. - eapply expanded_mkApps. now rewrite eqc. auto. solve_all. Qed. @@ -341,19 +368,6 @@ Proof. solve_all. Qed. -Lemma csubst_mkApps a k f l : - csubst a k (mkApps f l) = mkApps (csubst a k f) (map (csubst a k) l). -Proof. - induction l in f |- *; cbn; [auto|]. - rewrite IHl. - now cbn. -Qed. - -Lemma All_fold_tip {A : Type} (P : list A -> A -> Type) {x} : All_fold P [x] -> P [] x. -Proof. - intros a; now depelim a. -Qed. - Lemma expanded_let_expansion Σ (Δ : context) Γ t : expanded_context Σ Γ Δ -> expanded Σ (repeat 0 #|Δ| ++ Γ) t -> @@ -394,8 +408,6 @@ Proof. exact exp. lia. Qed. -Require Import PCUICUnivSubst. - Lemma subst_instance_isConstruct t u : isConstruct t@[u] = isConstruct t. Proof. destruct t => //. Qed. Lemma subst_instance_isRel t u : isRel t@[u] = isRel t. @@ -434,3 +446,925 @@ Proof. all:intros; try solve [econstructor; eauto 1; solve_all; try now rewrite app_assoc]. Qed. + +Lemma expanded_tApp (Σ : global_env) (Γ : list nat) (f : term) a : + expanded Σ Γ f -> expanded Σ Γ a -> + expanded Σ Γ (tApp f a). +Proof. + induction 1 using expanded_ind; intros expa. + all:rewrite -?[tApp _ a](mkApps_app _ _ [a]). + all:try (eapply (expanded_mkApps _ _ _ [a]) => //; econstructor; eauto). + + - econstructor; tea. rewrite app_length. lia. eapply app_Forall;eauto. + - econstructor; tea. eapply app_Forall; eauto. + - eapply expanded_tFix; tea. eapply app_Forall; eauto. eauto. rewrite app_length; cbn; eauto. lia. + - eapply expanded_tConstruct_app; tea. rewrite app_length ; lia. eapply app_Forall; eauto. +Qed. + +Lemma expanded_tApp_inv Σ Γ f a : + ~ isConstruct (head f) || isFix (head f) || isRel (head f) -> + expanded Σ Γ (tApp f a) -> + expanded Σ Γ f /\ expanded Σ Γ a. +Proof. + intros hf exp; depind exp. + - destruct args using rev_case; solve_discr; try discriminate. + rewrite mkApps_app in H3; noconf H3. + eapply Forall_app in H1 as [? ha]. depelim ha. + destruct args using rev_case; cbn in hf => //. + now rewrite !head_mkApps /= in hf. + - destruct args using rev_case; solve_discr. subst f6. + eapply IHexp => //. + rewrite mkApps_app in H2; noconf H2. + eapply Forall_app in H0 as [? H0]. depelim H0. split => //. + rewrite !head_mkApps in hf. + eapply expanded_mkApps => //. + - destruct args using rev_case; solve_discr. discriminate. + rewrite mkApps_app in H6; noconf H6. + eapply Forall_app in H1 as [? h]. depelim h. split => //. + now rewrite !head_mkApps /= in hf. + - destruct args using rev_case; solve_discr. discriminate. + rewrite mkApps_app in H3; noconf H3. + now rewrite !head_mkApps /= in hf. +Qed. + +Lemma expanded_tLambda_inv Σ Γ na t b : + expanded Σ Γ (tLambda na t b) -> + expanded Σ (0 :: Γ) b. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + +Lemma expanded_tLetIn_inv Σ Γ na t t' b : + expanded Σ Γ (tLetIn na t t' b) -> + expanded Σ Γ t /\ expanded Σ (0 :: Γ) b. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + +Lemma expanded_tEvar_inv Σ Γ ev l: + expanded Σ Γ (tEvar ev l) -> + Forall (expanded Σ Γ) l. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + +Lemma expanded_tCase_inv Σ Γ ci p c brs : + expanded Σ Γ (tCase ci p c brs) -> + Forall (expanded Σ Γ) (pparams p) /\ + Forall + (fun br : branch term => + ∥ All_fold + (fun (Δ : list context_decl) (d : context_decl) + => + ForOption + (fun b : term => + expanded Σ + (repeat 0 #|Δ| ++ + repeat 0 #|pparams p|) b) + (decl_body d)) (bcontext br) ∥ /\ + expanded Σ (repeat 0 #|bcontext br| ++ Γ) (bbody br)) + brs /\ + expanded Σ Γ c. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + +Lemma expanded_tProj_inv Σ Γ p c : + expanded Σ Γ (tProj p c) -> + expanded Σ Γ c. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + +Lemma expanded_tCoFix_inv Σ Γ mfix idx : + expanded Σ Γ (tCoFix mfix idx) -> + Forall (fun d : def term => expanded Σ (repeat 0 #|mfix| ++ Γ) (dbody d)) mfix. +Proof. + intros exp; depind exp; solve_discr => //; eauto. +Qed. + +Import PCUICOnOne. + +Module Red1Apps. + + Set Warnings "-notation-overridden". + + Inductive red1 (Σ : global_env) (Γ : context) : term -> term -> Type := + (** Reductions *) + (** Beta *) + | red_beta na t b a ts : + Σ ;;; Γ |- mkApps (tLambda na t b) (a :: ts) ⇝ mkApps (b {0 := a}) ts + + (** Let *) + | red_zeta na b t b' : + Σ ;;; Γ |- tLetIn na b t b' ⇝ b' {0 := b} + + | red_rel i body : + option_map decl_body (nth_error Γ i) = Some (Some body) -> + Σ ;;; Γ |- tRel i ⇝ lift0 (S i) body + + (** Case *) + | red_iota ci c u args p brs br : + nth_error brs c = Some br -> + #|args| = (ci.(ci_npar) + context_assumptions br.(bcontext))%nat -> + Σ ;;; Γ |- tCase ci p (mkApps (tConstruct ci.(ci_ind) c u) args) brs + ⇝ iota_red ci.(ci_npar) p args br + + (** Fix unfolding, with guard *) + | red_fix mfix idx args narg fn : + unfold_fix mfix idx = Some (narg, fn) -> + is_constructor narg args = true -> + Σ ;;; Γ |- mkApps (tFix mfix idx) args ⇝ mkApps fn args + + (** CoFix-case unfolding *) + | red_cofix_case ip p mfix idx args narg fn brs : + unfold_cofix mfix idx = Some (narg, fn) -> + Σ ;;; Γ |- tCase ip p (mkApps (tCoFix mfix idx) args) brs ⇝ + tCase ip p (mkApps fn args) brs + + (** CoFix-proj unfolding *) + | red_cofix_proj p mfix idx args narg fn : + unfold_cofix mfix idx = Some (narg, fn) -> + Σ ;;; Γ |- tProj p (mkApps (tCoFix mfix idx) args) ⇝ tProj p (mkApps fn args) + + (** Constant unfolding *) + | red_delta c decl body (isdecl : declared_constant Σ c decl) u : + decl.(cst_body) = Some body -> + Σ ;;; Γ |- tConst c u ⇝ subst_instance u body + + (** Proj *) + | red_proj p args u arg: + nth_error args (p.(proj_npars) + p.(proj_arg)) = Some arg -> + Σ ;;; Γ |- tProj p (mkApps (tConstruct p.(proj_ind) 0 u) args) ⇝ arg + + + | abs_red_l na M M' N : Σ ;;; Γ |- M ⇝ M' -> Σ ;;; Γ |- tLambda na M N ⇝ tLambda na M' N + | abs_red_r na M M' N : Σ ;;; Γ ,, vass na N |- M ⇝ M' -> Σ ;;; Γ |- tLambda na N M ⇝ tLambda na N M' + + | letin_red_def na b t b' r : Σ ;;; Γ |- b ⇝ r -> Σ ;;; Γ |- tLetIn na b t b' ⇝ tLetIn na r t b' + | letin_red_ty na b t b' r : Σ ;;; Γ |- t ⇝ r -> Σ ;;; Γ |- tLetIn na b t b' ⇝ tLetIn na b r b' + | letin_red_body na b t b' r : Σ ;;; Γ ,, vdef na b t |- b' ⇝ r -> Σ ;;; Γ |- tLetIn na b t b' ⇝ tLetIn na b t r + + | case_red_param ci p params' c brs : + OnOne2 (fun t u => Σ ;;; Γ |- t ⇝ u) p.(pparams) params' -> + Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci (set_pparams p params') c brs + + | case_red_return ci p preturn' c brs : + Σ ;;; Γ ,,, inst_case_predicate_context p |- p.(preturn) ⇝ preturn' -> + Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci (set_preturn p preturn') c brs + + | case_red_discr ci p c c' brs : + Σ ;;; Γ |- c ⇝ c' -> Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci p c' brs + + | case_red_brs ci p c brs brs' : + OnOne2 (fun br br' => + on_Trel_eq (fun t u => Σ ;;; Γ ,,, inst_case_branch_context p br |- t ⇝ u) bbody bcontext br br') + brs brs' -> + Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci p c brs' + + | proj_red p c c' : Σ ;;; Γ |- c ⇝ c' -> Σ ;;; Γ |- tProj p c ⇝ tProj p c' + + | app_red_l M1 N1 M2 : Σ ;;; Γ |- M1 ⇝ N1 -> ~~ isApp M1 -> ~~ isFix M1 -> M2 <> nil -> + Σ ;;; Γ |- mkApps M1 M2 ⇝ mkApps N1 M2 + + | app_red_r M2 N2 M1 : ~~ isApp M1 -> OnOne2 (fun M2 N2 => Σ ;;; Γ |- M2 ⇝ N2) M2 N2 -> Σ ;;; Γ |- mkApps M1 M2 ⇝ mkApps M1 N2 + + | app_fix_red_ty mfix0 mfix1 idx M2 : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ |- t ⇝ u) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + Σ ;;; Γ |- mkApps (tFix mfix0 idx) M2 ⇝ mkApps (tFix mfix1 idx) M2 + + | app_fix_red_body mfix0 mfix1 idx M2 : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> + Σ ;;; Γ |- mkApps (tFix mfix0 idx) M2 ⇝ mkApps (tFix mfix1 idx) M2 + + | prod_red_l na M1 M2 N1 : Σ ;;; Γ |- M1 ⇝ N1 -> Σ ;;; Γ |- tProd na M1 M2 ⇝ tProd na N1 M2 + | prod_red_r na M2 N2 M1 : Σ ;;; Γ ,, vass na M1 |- M2 ⇝ N2 -> + Σ ;;; Γ |- tProd na M1 M2 ⇝ tProd na M1 N2 + + | evar_red ev l l' : OnOne2 (fun t u => Σ ;;; Γ |- t ⇝ u) l l' -> Σ ;;; Γ |- tEvar ev l ⇝ tEvar ev l' + + | fix_red_ty mfix0 mfix1 idx : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ |- t ⇝ u) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + Σ ;;; Γ |- tFix mfix0 idx ⇝ tFix mfix1 idx + + | fix_red_body mfix0 mfix1 idx : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) dbody (fun x => (dname x, dtype x, rarg x))) + mfix0 mfix1 -> + Σ ;;; Γ |- tFix mfix0 idx ⇝ tFix mfix1 idx + + | cofix_red_ty mfix0 mfix1 idx : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ |- t ⇝ u) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + Σ ;;; Γ |- tCoFix mfix0 idx ⇝ tCoFix mfix1 idx + + | cofix_red_body mfix0 mfix1 idx : + OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> + Σ ;;; Γ |- tCoFix mfix0 idx ⇝ tCoFix mfix1 idx + where " Σ ;;; Γ |- t ⇝ u " := (red1 Σ Γ t u). + + Lemma red1_ind_all : + forall (Σ : global_env) (P : context -> term -> term -> Type), + + (forall (Γ : context) (na : aname) (t b a : term) ts, + P Γ (mkApps (tLambda na t b) (a :: ts)) (mkApps (b {0 := a}) ts)) -> + + (forall (Γ : context) (na : aname) (b t b' : term), P Γ (tLetIn na b t b') (b' {0 := b})) -> + + (forall (Γ : context) (i : nat) (body : term), + option_map decl_body (nth_error Γ i) = Some (Some body) -> P Γ (tRel i) ((lift0 (S i)) body)) -> + + (forall (Γ : context) (ci : case_info) (c : nat) (u : Instance.t) (args : list term) + (p : predicate term) (brs : list (branch term)) br, + nth_error brs c = Some br -> + #|args| = (ci.(ci_npar) + context_assumptions br.(bcontext))%nat -> + P Γ (tCase ci p (mkApps (tConstruct ci.(ci_ind) c u) args) brs) + (iota_red ci.(ci_npar) p args br)) -> + + (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 Γ (mkApps (tFix mfix idx) args) (mkApps fn args)) -> + + (forall (Γ : context) ci (p : predicate term) (mfix : mfixpoint term) (idx : nat) + (args : list term) (narg : nat) (fn : term) (brs : list (branch term)), + unfold_cofix mfix idx = Some (narg, fn) -> + P Γ (tCase ci p (mkApps (tCoFix mfix idx) args) brs) (tCase ci p (mkApps fn args) brs)) -> + + (forall (Γ : context) (p : projection) (mfix : mfixpoint term) (idx : nat) (args : list term) + (narg : nat) (fn : term), + unfold_cofix mfix idx = Some (narg, fn) -> P Γ (tProj p (mkApps (tCoFix mfix idx) args)) (tProj p (mkApps fn args))) -> + + (forall (Γ : context) c (decl : constant_body) (body : term), + declared_constant Σ c decl -> + forall u : Instance.t, cst_body decl = Some body -> P Γ (tConst c u) (subst_instance u body)) -> + + (forall (Γ : context) p (args : list term) (u : Instance.t) + (arg : term), + nth_error args (p.(proj_npars) + p.(proj_arg)) = Some arg -> + P Γ (tProj p (mkApps (tConstruct p.(proj_ind) 0 u) args)) arg) -> + + (forall (Γ : context) (na : aname) (M M' N : term), + red1 Σ Γ M M' -> P Γ M M' -> P Γ (tLambda na M N) (tLambda na M' N)) -> + + (forall (Γ : context) (na : aname) (M M' N : term), + red1 Σ (Γ,, vass na N) M M' -> P (Γ,, vass na N) M M' -> P Γ (tLambda na N M) (tLambda na N M')) -> + + (forall (Γ : context) (na : aname) (b t b' r : term), + red1 Σ Γ b r -> P Γ b r -> P Γ (tLetIn na b t b') (tLetIn na r t b')) -> + + (forall (Γ : context) (na : aname) (b t b' r : term), + red1 Σ Γ t r -> P Γ t r -> P Γ (tLetIn na b t b') (tLetIn na b r b')) -> + + (forall (Γ : context) (na : aname) (b t b' r : term), + red1 Σ (Γ,, vdef na b t) b' r -> P (Γ,, vdef na b t) b' r -> P Γ (tLetIn na b t b') (tLetIn na b t r)) -> + + (forall (Γ : context) (ci : case_info) p params' c brs, + OnOne2 (Trel_conj (red1 Σ Γ) (P Γ)) p.(pparams) params' -> + P Γ (tCase ci p c brs) + (tCase ci (set_pparams p params') c brs)) -> + + (forall (Γ : context) (ci : case_info) p preturn' c brs, + red1 Σ (Γ ,,, inst_case_predicate_context p) p.(preturn) preturn' -> + P (Γ ,,, inst_case_predicate_context p) p.(preturn) preturn' -> + P Γ (tCase ci p c brs) + (tCase ci (set_preturn p preturn') c brs)) -> + + (forall (Γ : context) (ind : case_info) (p : predicate term) (c c' : term) (brs : list (branch term)), + red1 Σ Γ c c' -> P Γ c c' -> P Γ (tCase ind p c brs) (tCase ind p c' brs)) -> + + (forall (Γ : context) ci p c brs brs', + OnOne2 (fun br br' => + (on_Trel_eq (Trel_conj (red1 Σ (Γ ,,, inst_case_branch_context p br)) + (P (Γ ,,, inst_case_branch_context p br))) + bbody bcontext br br')) brs brs' -> + P Γ (tCase ci p c brs) (tCase ci p c brs')) -> + + (forall (Γ : context) (p : projection) (c c' : term), red1 Σ Γ c c' -> P Γ c c' -> + P Γ (tProj p c) (tProj p c')) -> + + (forall (Γ : context) M1 N1 M2, red1 Σ Γ M1 N1 -> P Γ M1 N1 -> ~~ isApp M1 -> ~~ isFix M1 -> M2 <> [] -> + P Γ (mkApps M1 M2) (mkApps N1 M2)) -> + + (forall (Γ : context) M1 N2 M2, ~~ isApp M1 -> + OnOne2 (Trel_conj (red1 Σ Γ) (P Γ)) M2 N2 -> + P Γ (mkApps M1 M2) (mkApps M1 N2)) -> + + (forall (Γ : context) mfix0 mfix1 idx M2, + OnOne2 (on_Trel_eq (Trel_conj (fun t u => Σ ;;; Γ |- t ⇝ u) (P Γ)) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + P Γ (mkApps (tFix mfix0 idx) M2) (mkApps (tFix mfix1 idx) M2)) -> + + (forall (Γ : context) mfix0 mfix1 idx M2, + OnOne2 (on_Trel_eq (Trel_conj (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) (P (Γ ,,, fix_context mfix0))) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> + P Γ (mkApps (tFix mfix0 idx) M2) (mkApps (tFix mfix1 idx) M2)) -> + + (forall (Γ : context) (na : aname) (M1 M2 N1 : term), + red1 Σ Γ M1 N1 -> P Γ M1 N1 -> P Γ (tProd na M1 M2) (tProd na N1 M2)) -> + + (forall (Γ : context) (na : aname) (M2 N2 M1 : term), + red1 Σ (Γ,, vass na M1) M2 N2 -> P (Γ,, vass na M1) M2 N2 -> P Γ (tProd na M1 M2) (tProd na M1 N2)) -> + + (forall (Γ : context) (ev : nat) (l l' : list term), + OnOne2 (Trel_conj (red1 Σ Γ) (P Γ)) l l' -> P Γ (tEvar ev l) (tEvar ev l')) -> + + (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), + OnOne2 (on_Trel_eq (Trel_conj (red1 Σ Γ) (P Γ)) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + P Γ (tFix mfix0 idx) (tFix mfix1 idx)) -> + + (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), + OnOne2 (on_Trel_eq (Trel_conj (red1 Σ (Γ ,,, fix_context mfix0)) + (P (Γ ,,, fix_context mfix0))) dbody + (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> + P Γ (tFix mfix0 idx) (tFix mfix1 idx)) -> + + (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), + OnOne2 (on_Trel_eq (Trel_conj (red1 Σ Γ) (P Γ)) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> + P Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx)) -> + + (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), + OnOne2 (on_Trel_eq (Trel_conj (red1 Σ (Γ ,,, fix_context mfix0)) + (P (Γ ,,, fix_context mfix0))) dbody + (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> + P Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx)) -> + + forall (Γ : context) (t t0 : term), red1 Σ Γ t t0 -> P Γ t t0. + Proof. + intros. rename X29 into Xlast. revert Γ t t0 Xlast. + fix aux 4. intros Γ t T. + move aux at top. + destruct 1; match goal with + | |- P _ (tFix _ _) (tFix _ _) => idtac + | |- P _ (tCoFix _ _) (tCoFix _ _) => idtac + | |- P _ (mkApps (tFix _ _) _) _ => idtac + | |- P _ (tCase _ _ (mkApps (tCoFix _ _) _) _) _ => idtac + | |- P _ (tProj _ (mkApps (tCoFix _ _) _)) _ => idtac + | H : _ |- _ => eapply H; eauto + end. + - eapply X3; eauto. + - eapply X4; eauto. + - eapply X5; eauto. + + - revert params' o. + generalize (pparams p). + fix auxl 3. + intros params params' []. + + constructor. split; auto. + + constructor. auto. + + - revert brs' o. + revert brs. + fix auxl 3. + intros l l' Hl. destruct Hl. + + simpl in *. constructor; intros; intuition auto. + + constructor. eapply auxl. apply Hl. + + - revert M2 N2 o. + fix auxl 3. + intros l l' Hl. destruct Hl. + + constructor. split; auto. + + constructor. auto. + + - eapply X20. + revert mfix0 mfix1 o. + fix auxl 3. + intros l l' Hl. destruct Hl. + + constructor; split; auto; intuition. + + constructor; try split; auto; intuition. + + - eapply X21. revert o. + generalize (fix_context mfix0). + intros c o. revert mfix0 mfix1 o. + fix auxl 3. + intros l l' Hl. destruct Hl. + + constructor. split; auto; intuition. + + constructor; try split; auto; intuition. + + - revert l l' o. + fix auxl 3. + intros l l' Hl. destruct Hl. + + constructor. split; auto. + + constructor. auto. + + - eapply X25. + revert mfix0 mfix1 o; fix auxl 3. + intros l l' Hl; destruct Hl; + constructor; try split; auto; intuition. + + - eapply X26. + revert o. generalize (fix_context mfix0). intros c Xnew. + revert mfix0 mfix1 Xnew; fix auxl 3; intros l l' Hl; + destruct Hl; constructor; try split; auto; intuition. + + - eapply X27. + revert mfix0 mfix1 o. + fix auxl 3; intros l l' Hl; destruct Hl; + constructor; try split; auto; intuition. + + - eapply X28. + revert o. generalize (fix_context mfix0). intros c new. + revert mfix0 mfix1 new; fix auxl 3; intros l l' Hl; destruct Hl; + constructor; try split; auto; intuition. + Defined. + + Lemma red_tApp Σ Γ t v u : + red1 Σ Γ t v -> + red1 Σ Γ (tApp t u) (tApp v u). + Proof. + induction 1. + all:try solve [eapply (app_red_l _ _ _ _ [u]) => //; econstructor; eauto]. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). + eapply red_beta. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). econstructor; eauto. + now apply is_constructor_app_ge. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). econstructor; eauto. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). econstructor; eauto. + now eapply OnOne2_app_r. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). now eapply app_fix_red_ty. + - rewrite -![tApp _ u](mkApps_app _ _ [u]). now eapply app_fix_red_body. + - now eapply (app_fix_red_ty _ _ _ _ _ [_]). + - now eapply (app_fix_red_body _ _ _ _ _ [_]). + Qed. + + Lemma isLambda_red1 Σ Γ b b' : isLambda b -> red1 Σ Γ b b' -> isLambda b'. + Proof. + destruct b; simpl; try discriminate. + intros _ red. + depelim red; solve_discr; eauto. depelim o. + Qed. + +End Red1Apps. + +Lemma red1_red1apps Σ Γ t v : + red1 Σ Γ t v -> Red1Apps.red1 Σ Γ t v. +Proof. + intros red; induction red using red1_ind_all in |- *. + all:try solve [econstructor; eauto; solve_all; eauto]. + - eapply (Red1Apps.red_beta _ _ _ _ _ _ []). + - now eapply Red1Apps.red_tApp. + - remember (decompose_app (tApp M1 M2)). + destruct p as [hd args]. + edestruct (decompose_app_rec_inv' M1 [M2]). rewrite /decompose_app in Heqp. + cbn in Heqp. rewrite -Heqp. reflexivity. + destruct a as [napp [hm2 hm1]]. + symmetry in Heqp; eapply decompose_app_inv in Heqp. rewrite Heqp. + assert (tApp M1 N2 = mkApps hd (firstn x args ++ [N2])) as ->. + { now rewrite mkApps_app -hm1. } + rewrite -{1}(firstn_skipn x args) -hm2. eapply Red1Apps.app_red_r => //. + eapply OnOne2_app. now constructor. +Qed. + +Lemma head_nApp f : + ~~ isApp f -> head f = f. +Proof. + induction f => //. +Qed. + +Lemma expanded_mkApps_inv Σ Γ f v : + ~~ isApp f -> + ~~ (isConstruct f || isFix f || isRel f) -> + expanded Σ Γ (mkApps f v) -> + expanded Σ Γ f /\ Forall (expanded Σ Γ) v. +Proof. + intros happ hc. + induction v using rev_ind; cbn. + - intros exp; split => //. + - rewrite mkApps_app /=. + move/expanded_tApp_inv => e. + forward e. rewrite head_mkApps. + rewrite head_nApp //. now move/negbTE: hc ->. + intuition auto. eapply app_Forall; eauto. +Qed. + +Lemma expanded_mkApps_inv' Σ Γ f : + expanded Σ Γ f -> + let args := arguments f in + Forall (expanded Σ Γ) args /\ + match head f with + | tRel n => exists m, nth_error Γ n = Some m /\ m <= #|args| + | tConstruct ind c u => exists mind idecl cdecl, + declared_constructor Σ (ind, c) mind idecl cdecl /\ #|args| >= ind_npars mind + context_assumptions (cstr_args cdecl) + | tFix mfix idx => exists d, + Forall + (fun d0 : def term => + isLambda (dbody d0) /\ + (let ctx := + rev_map (fun d1 : def term => (1 + rarg d1)%nat) + mfix in + expanded Σ (ctx ++ Γ) (dbody d0))) mfix /\ + args <> [] /\ + nth_error mfix idx = Some d /\ #|args| > rarg d + | _ => expanded Σ Γ (head f) + end. +Proof. + induction 1 using expanded_ind; cbn. + all: try solve [split; econstructor; eauto]. + all:rewrite head_mkApps /=. + - rewrite arguments_mkApps_nApp //. split => //. now exists m. + - destruct IHexpanded. rewrite arguments_mkApps. split. + eapply app_Forall => //. + rewrite app_length. + destruct (head f6) => //; firstorder (eauto; try lia). + exists x. split => //. firstorder (eauto; try lia). + intros heq; apply H5. now eapply app_eq_nil in heq. + - rewrite arguments_mkApps_nApp //. split => //. now exists d. + - rewrite arguments_mkApps_nApp //. split => //. + firstorder. +Qed. + +Lemma All_fold_map_context (P : context -> context_decl -> Type) (f : term -> term) ctx : + All_fold P (map_context f ctx) <~> All_fold (fun Γ d => P (map_context f Γ) (map_decl f d)) ctx. +Proof. + induction ctx. + - split; constructor. + - cbn. split; intros H; depelim H; constructor; auto; now apply IHctx. +Qed. + +Lemma expanded_fix_subst Σ a k b Γ Γs Δ : + #|Γ| = k -> + Forall2 (fun n arg => forall args Γ' k', + Forall (expanded Σ (Γ' ++ Δ)) args -> + n <= #|args| -> + #|Γ'| = k' -> + expanded Σ (Γ' ++ Δ) (mkApps (lift0 k' arg) args)) Γs a -> + expanded Σ (Γ ++ Γs ++ Δ) b -> + expanded Σ (Γ ++ Δ) (subst a k b). +Proof. + intros Hk Hs. + remember (Γ ++ _ ++ Δ)%list as Γ_. + intros exp; revert Γ k Hk HeqΓ_. + induction exp using expanded_ind; intros Γ' k Hk ->. + all:try solve[ cbn; econstructor => // ]. + 2,7:solve[ cbn; econstructor => //; solve_all ]. + - rewrite subst_mkApps /=. + destruct (Nat.leb_spec k n). + destruct (nth_error a _) eqn:hnth. + * pose proof (Forall2_length Hs). + pose proof (Forall2_nth_error_Some_r _ _ _ _ _ _ _ hnth Hs) as [t' [hnth' hargs]]. + eapply hargs => //. + eapply Forall_map. + 2:{ len. rewrite nth_error_app_ge in H. lia. subst k. + rewrite nth_error_app_lt in H. + eapply nth_error_Some_length in hnth'. lia. rewrite hnth' in H. noconf H. exact H0. } + solve_all. + * rewrite nth_error_app_ge in H. lia. + eapply nth_error_None in hnth. + pose proof (Forall2_length Hs). + rewrite nth_error_app_ge in H. lia. + eapply expanded_tRel. rewrite nth_error_app_ge. lia. erewrite <- H. + lia_f_equal. len. solve_all. + * rewrite nth_error_app_lt in H. lia. + eapply expanded_tRel. rewrite nth_error_app_lt. lia. tea. now len. + solve_all. + - cbn. econstructor. + eapply (IHexp (0 :: Γ') (S k)); cbn; auto; try lia. + - cbn. econstructor. apply IHexp1; auto. + eapply (IHexp2 (0 :: Γ') (S k)); cbn; auto; lia. + - rewrite subst_mkApps. + destruct (isConstruct (subst a k f6) || isFix (subst a k f6) || isRel (subst a k f6)) eqn:eqc. + specialize (IHexp _ _ Hk eq_refl). + eapply expanded_mkApps_expanded => //. solve_all. + eapply expanded_mkApps => //. rewrite eqc //. now eapply IHexp. solve_all. + - cbn. econstructor. eauto. cbn. solve_all. solve_all. + specialize (H1 (repeat 0 #|bcontext x| ++ Γ') (#|bcontext x| + k)%nat). + forward H1 by len. + forward H1. now rewrite app_assoc. + rewrite /id. rewrite app_assoc. apply H1. + - rewrite subst_mkApps. cbn. + eapply expanded_tFix. + + solve_all. now eapply isLambda_subst. + specialize (a0 + (rev_map (fun d0 : def term => S (rarg d0)) + (map (map_def (subst a k) (subst a (#|mfix| + k))) mfix) ++ Γ') (#|mfix| + k)%nat). + forward a0 by len. + forward a0. { rewrite app_assoc. f_equal. f_equal. + rewrite !rev_map_spec. f_equal. now rewrite map_map_compose /=. } + rewrite app_assoc. eapply a0. + + solve_all. + + now destruct args. + + rewrite nth_error_map /= H4 //. + + len. + - cbn. constructor. + solve_all. + specialize (a0 (repeat 0 #|mfix| ++ Γ') (#|mfix| + k)%nat). + forward a0 by len. + forward a0. { rewrite app_assoc //. } + rewrite app_assoc. eapply a0. + - rewrite subst_mkApps. cbn. + eapply expanded_tConstruct_app; tea. now len. + solve_all. +Qed. + +Lemma expanded_unfold_fix Σ Γ' mfix idx narg fn : + unfold_fix mfix idx = Some (narg, fn) -> + All (fun d0 : def term => isLambda (dbody d0) /\ expanded Σ (rev_map (fun d1 : def term => S (rarg d1)) mfix ++ Γ') (dbody d0)) mfix -> + expanded Σ Γ' fn. +Proof. + unfold unfold_fix. + destruct nth_error eqn:e => //. + intros [= <- <-] hf. + pose proof (nth_error_all e hf) as [hl hf']. + eapply (expanded_fix_subst _ _ _ _ []) => //; tea. + rewrite rev_map_spec. + eapply Forall2_from_nth_error. len. + intros n rarg f. len. intros hn hrarg hnthf args Γ'' k' hargs hrarg' <-. + eapply PCUICParallelReductionConfluence.nth_error_fix_subst in hnthf. subst f. + move: hrarg. + rewrite nth_error_rev; len. rewrite List.rev_involutive nth_error_map. + intros hrarg. + destruct (nth_error mfix (_ - _)) eqn:e'. cbn in hrarg. noconf hrarg. + eapply expanded_tFix => //. solve_all. + eapply expanded_lift; len. rewrite !rev_map_spec in H1 *. + rewrite map_map => //. destruct args => //. cbn in hrarg'. lia. + rewrite nth_error_map /= e' /= //. cbn. lia. + eapply nth_error_None in e'. lia. +Qed. + +Lemma expanded_unfold_cofix Σ Γ' mfix idx narg fn : + unfold_cofix mfix idx = Some (narg, fn) -> + All (fun d0 : def term => expanded Σ (repeat 0 #|mfix| ++ Γ') (dbody d0)) mfix -> + expanded Σ Γ' fn. +Proof. + unfold unfold_cofix. + destruct nth_error eqn:e => //. + intros [= <- <-] hf. + pose proof (nth_error_all e hf) as hf'. + eapply (expanded_subst _ _ _ _ []) => //; tea. + rewrite /cofix_subst. + generalize #|mfix|. + induction n; repeat constructor; eauto. solve_all. + len. +Qed. + +Lemma expanded_weakening_env {cf : checker_flags} Σ Σ' Γ t : + wf Σ' -> + extends Σ Σ' -> + expanded Σ Γ t -> expanded Σ' Γ t. +Proof. + intros w s. + eapply expanded_ind; intros. + all:try solve [econstructor; eauto]. + - econstructor; eauto. solve_all. sq. eapply All_fold_impl; tea. cbn. + intros. now rewrite repeat_app in H6. + - eapply expanded_tFix; eauto. solve_all. + - eapply expanded_tConstruct_app; eauto. + eapply PCUICWeakeningEnv.weakening_env_declared_constructor; tea. +Qed. + +Lemma expanded_global_env_constant {cf : checker_flags} Σ c decl : + wf Σ -> + expanded_global_env Σ -> + declared_constant Σ c decl -> + ForOption (expanded Σ []) (cst_body decl). +Proof. + intros wf; destruct Σ as [Σ univs] => /=. cbn. + unfold expanded_global_env. cbn. + induction 1; cbn => //. + intros [->|H']. + - depelim H0. + destruct decl as [? [] ?]; cbn in *. + constructor. cbn. + eapply expanded_weakening_env; tea. + eapply extends_strictly_on_decls_extends. + split => //=. eapply incl_cs_refl. 2:eapply Retroknowledge.extends_refl. + set (cb := ConstantDecl _). now exists [(c, cb)]. + constructor. + - forward IHexpanded_global_declarations. + destruct wf. cbn in *. split => //. + cbn. now depelim o0. + eapply IHexpanded_global_declarations in H'. + destruct decl as [? [] ?]; cbn in * => //. 2:constructor. + depelim H'. constructor. + eapply expanded_weakening_env; tea. + eapply extends_strictly_on_decls_extends. + split => //=. eapply incl_cs_refl. 2:eapply Retroknowledge.extends_refl. + now exists [decl0]. +Qed. + +Lemma All_fold_nth_error P (ctx : context) n b : + All_fold P ctx -> nth_error ctx n = Some b -> + P (skipn (S n) ctx) b. +Proof. + induction 1 in n, b |- *. + - rewrite nth_error_nil //. + - destruct n => //=. + * intros [= <-]. now rewrite skipn_S skipn_0. + * intros hnth. now rewrite skipn_S. +Qed. + +Lemma skipn_repeat {A} k n (a : A) : + skipn n (repeat a k) = repeat a (k - n). +Proof. + induction n in k |- *. + - rewrite skipn_0. lia_f_equal. + - destruct k => //=. + rewrite skipn_S IHn. lia_f_equal. +Qed. + +Lemma expanded_red1 {cf : checker_flags} {Σ : global_env_ext} Γ Γ' t v : wf Σ -> + expanded_global_env Σ -> + (forall n body, option_map decl_body (nth_error Γ n) = Some (Some body) -> expanded Σ (skipn (S n) Γ') body) -> + red1 Σ Γ t v -> + expanded Σ Γ' t -> + expanded Σ Γ' v. +Proof. + move=> wf expΣ wfΓ /red1_red1apps red. + induction red using Red1Apps.red1_ind_all in wfΓ, Γ' |- *; intros exp. + - eapply expanded_mkApps_inv in exp as [] => //. + eapply expanded_tLambda_inv in H. + depelim H0. + eapply expanded_mkApps_expanded => //. + eapply (expanded_subst _ _ _ _ [] Γ') => //. now constructor. solve_all. + - eapply expanded_tLetIn_inv in exp as []. + eapply (expanded_subst _ _ _ _ [] Γ') => //. now constructor. + - rewrite -(firstn_skipn (S i) Γ'). + eapply expanded_mkApps_inv' in exp; cbn in exp. + destruct exp as [_ [m []]]. eapply nth_error_Some_length in H0. + eapply (expanded_lift _ _ _ _ []) => //. + rewrite firstn_length_le; try lia. + now eapply wfΓ. + - eapply expanded_tCase_inv in exp as [? []]. + unfold iota_red. + move/expanded_mkApps_inv': H3 => /=. + rewrite arguments_mkApps_nApp // head_mkApps //=. + intros [hargs [mind [idecl [cdecl [declc hargs']]]]]. + eapply nth_error_forall in H2; tea. + destruct H2 as [[hbctx] hbod]. + eapply (expanded_subst _ _ _ _ [] _) => //. + + eapply Forall_rev, Forall_skipn => //. + + len. replace #|skipn (ci_npar ci) args| with (context_assumptions (inst_case_branch_context p br)). + eapply expanded_let_expansion; len. red. + sq. rewrite /inst_case_branch_context /inst_case_context /subst_context. + eapply PCUICParallelReduction.All_fold_fold_context_k. + eapply All_fold_map_context, All_fold_impl; tea; cbn. + intros ? ? fo. len. destruct x as [? [] ?] => //; constructor. + cbn in fo. depelim fo. eapply (expanded_subst _ _ _ _ (repeat 0 #|Γ0|) _); len. + eapply Forall_rev; eauto. + eapply expanded_subst_instance. rewrite app_assoc. now eapply expanded_weakening. + rewrite skipn_length. len. + - move/expanded_mkApps_inv': exp. cbn. + rewrite arguments_mkApps_nApp // head_mkApps //=. + move=> [hargs [d [hf [hargs' []]]]] hnth hrarg. + eapply expanded_mkApps_expanded => //; solve_all. + eapply expanded_unfold_fix in H; tea. + - eapply expanded_tCase_inv in exp as [? []]. + constructor => //. + eapply expanded_mkApps_inv' in H2. move: H2; rewrite arguments_mkApps_nApp // head_mkApps //=. + intros [hargs' hcof]. cbn in hcof. eapply expanded_tCoFix_inv in hcof. + eapply expanded_unfold_cofix in H; tea. eapply expanded_mkApps_expanded; tea => //; solve_all. solve_all. + - eapply expanded_tProj_inv in exp. econstructor. + eapply expanded_mkApps_inv' in exp. move: exp; rewrite arguments_mkApps_nApp // head_mkApps //=. + intros [hargs' hcof]. cbn in hcof. eapply expanded_tCoFix_inv in hcof. + eapply expanded_mkApps_expanded => //; solve_all. + eapply expanded_unfold_cofix in H; tea. + - eapply expanded_subst_instance. + eapply expanded_global_env_constant in expΣ; tea. + eapply (expanded_weakening _ []). rewrite H0 in expΣ. now depelim expΣ. + - eapply expanded_tProj_inv in exp. + move/expanded_mkApps_inv': exp. rewrite arguments_mkApps_nApp // head_mkApps //=. + intros []. + eapply nth_error_forall in H0; tea. + - constructor. now eapply expanded_tLambda_inv in exp. + - constructor. eapply expanded_tLambda_inv in exp. + eapply IHred => //. + { intros [] body; cbn => //. rewrite skipn_S. eapply wfΓ. } + - eapply expanded_tLetIn_inv in exp; now constructor. + - eapply expanded_tLetIn_inv in exp. now constructor. + - eapply expanded_tLetIn_inv in exp. constructor; intuition eauto. + eapply IHred => //. + { intros [] ? => //=. intros [=]. subst b. now rewrite skipn_S skipn_0. + rewrite skipn_S. eapply wfΓ. } + - eapply expanded_tCase_inv in exp as [? []]. + constructor; eauto. cbn. + solve_all. eapply OnOne2_impl_All_r; tea. intuition eauto. now apply X0. + now rewrite /PCUICOnOne.set_pparams /= -(OnOne2_length X). + - eapply expanded_tCase_inv in exp as [? []]. + constructor; eauto. + - eapply expanded_tCase_inv in exp as [? []]. + econstructor; eauto. + - eapply expanded_tCase_inv in exp as [? []]. + econstructor; eauto. solve_all. + eapply OnOne2_impl_All_r; tea. + intros x y [? ?]. intros [[] ?]. rewrite -e0. + solve_all. eapply e => //. + intros n b. + clear -H H2 wfΓ. + destruct nth_error eqn:e' => //. + cbn. intros [=]. destruct c as [? [] ?]. cbn in *. noconf H1. + eapply nth_error_app_inv in e' as [[]|[]]. + { rewrite inst_case_branch_context_length in H0. destruct H2. + rewrite /inst_case_branch_context /inst_case_context in H1. + destruct (nth_error (bcontext x) n) eqn:e. + 2:{ eapply nth_error_None in e. rewrite skipn_app skipn_all2. len. cbn. len. } + rewrite /subst_context in H1. + erewrite nth_error_fold_context_k in H1. 4:{ rewrite nth_error_map e //. } 3:len. 2:exact []. + len in H1. noconf H1. destruct c as [? [] ?]; noconf H1. + rewrite skipn_app. len. eapply All_fold_nth_error in X; tea. cbn in X. depelim X. + rewrite skipn_length in H1. + eapply expanded_subst. rewrite skipn_length. len. + replace (S n - #|bcontext x|) with 0. 2:{ lia. } rewrite skipn_0. eapply Forall_rev. solve_all. + len. rewrite app_assoc. eapply expanded_weakening. eapply expanded_subst_instance. + now rewrite skipn_repeat. } + { rewrite inst_case_branch_context_length in H0, H1. + rewrite skipn_app skipn_all2 /=; len. + replace (S n - #|bcontext x|) with (S (n - #|bcontext x|)). 2:lia. + eapply wfΓ. rewrite H1 //. } + noconf H1. + - eapply expanded_tProj_inv in exp. now econstructor. + - eapply expanded_mkApps_inv' in exp. + rewrite head_mkApps head_nApp in exp => //. + rewrite arguments_mkApps_nApp // in exp. destruct exp as []. + specialize (IHred Γ' wfΓ). + destruct M1; try solve [eapply expanded_mkApps_expanded => //; eauto; solve_all]. + * depelim red; solve_discr. eapply wfΓ in e. + eapply expanded_mkApps_expanded => //; solve_all. + rewrite -(firstn_skipn (S n) Γ'). eapply (expanded_lift _ _ _ _ []) => //. + destruct H3 as [m [hn ha]]. + rewrite firstn_length_le //. now eapply nth_error_Some_length in hn. + depelim o. + * depelim red; solve_discr. depelim o. + + - eapply expanded_mkApps_inv' in exp. + move: exp. + rewrite head_mkApps head_nApp // arguments_mkApps_nApp //. + intros []. + assert(Forall (expanded Σ Γ') N2). + { clear H1. solve_all. eapply OnOne2_impl_All_r; tea. cbn. intuition auto. } + destruct M1; try solve [eapply expanded_mkApps => //; eauto]. + + rewrite (OnOne2_length X) in H1. destruct H1 as [m []]; eapply expanded_tRel; tea. + + rewrite (OnOne2_length X) in H1. destruct H1 as [m [? [? []]]]; eapply expanded_tConstruct_app; tea. + + rewrite (OnOne2_length X) in H1. + destruct H1 as [d [? [? []]]]; eapply expanded_tFix; tea. + destruct N2; cbn in *; eauto. lia. intros eq; discriminate. + + - move/expanded_mkApps_inv': exp. + rewrite head_mkApps head_nApp // arguments_mkApps_nApp // => [] [] hargs [d [hf [hm2 [hnth harg]]]]. + eapply OnOne2_nth_error in hnth as [t' [hnth hor]]; tea. + eapply expanded_tFix; tea. + { clear hor. solve_all. eapply OnOne2_impl_All_r; tea. + intros ? ? [] []. noconf e. rewrite -H2. split => //. solve_all. + clear -X H1. + enough (rev_map (fun d1 : def term => S (rarg d1)) mfix0 = (rev_map (fun d1 : def term => S (rarg d1)) mfix1)) as <- => //. + clear -X. rewrite !rev_map_spec. f_equal. induction X. destruct p as []. cbn in *. congruence. cbn. congruence. } + { destruct hor; subst => //. destruct p as [? e]; noconf e. now congruence. } + + - move/expanded_mkApps_inv': exp. + rewrite head_mkApps head_nApp // arguments_mkApps_nApp // => [] [] hargs [d [hf [hm2 [hnth harg]]]]. + eapply OnOne2_nth_error in hnth as [t' [hnth hor]]; tea. + eapply expanded_tFix; tea. + { clear hor. solve_all. + assert (rev_map (fun d1 : def term => S (rarg d1)) mfix0 = (rev_map (fun d1 : def term => S (rarg d1)) mfix1)). + { clear -X. rewrite !rev_map_spec. f_equal. induction X. destruct p as []. cbn in *. congruence. cbn. congruence. } + rewrite -H. + eapply OnOne2_impl_All_r; tea. + intros ? ? [] []. noconf e. destruct p. + eapply Red1Apps.isLambda_red1 in r; tea. split => //. + set(Γ'' := rev_map (fun d1 : def term => S (rarg d1)) mfix0). + eapply e => //. + { intros n b hnth'. + destruct (nth_error (fix_context mfix0) n) eqn:e'. + rewrite nth_error_app_lt in hnth'. now eapply nth_error_Some_length in e'. + rewrite e' in hnth'. noconf hnth'. + pose proof (PCUICParallelReductionConfluence.fix_context_assumption_context mfix0). + eapply PCUICParallelReductionConfluence.nth_error_assumption_context in H6; tea. congruence. + rewrite /Γ''. + eapply nth_error_None in e'. len in e'. + rewrite nth_error_app_ge in hnth' => //. now len. + rewrite skipn_app skipn_all2. len. + cbn. len. replace (S n - #|mfix0|) with (S (n - #|mfix0|)). 2:{ lia. } + eapply wfΓ. now len in hnth'. } } + { destruct hor; subst => //. destruct p as [? e]; noconf e. now congruence. } + - eapply expanded_tProd. + - constructor. + - constructor. + eapply expanded_tEvar_inv in exp. + solve_all; eapply OnOne2_impl_All_r; tea. intuition eauto. + now eapply X0. + - depelim exp; solve_discr. now cbn in H. + - depelim exp; solve_discr. now cbn in H. + - eapply expanded_tCoFix_inv in exp. econstructor. + rewrite -(OnOne2_length X). solve_all; eapply OnOne2_impl_All_r; tea; intuition eauto. + destruct X0. noconf e. now rewrite -H1. + - eapply expanded_tCoFix_inv in exp. econstructor. + rewrite -(OnOne2_length X). solve_all; eapply OnOne2_impl_All_r; tea; intuition eauto. + destruct X0. destruct p. noconf e. eapply e0 => //. + intros n b. + destruct nth_error eqn:e' => //. + cbn. intros [=]. destruct c as [? [] ?]. cbn in *. noconf H4. + eapply nth_error_app_inv in e' as [[]|[]]. + { pose proof (PCUICParallelReductionConfluence.fix_context_assumption_context mfix0). + eapply PCUICParallelReductionConfluence.nth_error_assumption_context in H5; tea. cbn in H5. congruence. } + len in H3. len in H4. + rewrite skipn_app skipn_all2; len. + replace (S n - #|mfix0|) with (S (n - #|mfix0|)) by lia. eapply wfΓ. rewrite H4 /= //. + noconf H4. +Qed. + +Lemma expanded_red {cf : checker_flags} {Σ : global_env_ext} Γ Γ' t v : wf Σ -> + expanded_global_env Σ -> + (forall n body, option_map decl_body (nth_error Γ n) = Some (Some body) -> expanded Σ (skipn (S n) Γ') body) -> + red Σ Γ t v -> + expanded Σ Γ' t -> + expanded Σ Γ' v. +Proof. + intros. + induction X0. now eapply expanded_red1. auto. + eapply IHX0_2. now eapply IHX0_1. +Qed. + diff --git a/pcuic/theories/PCUICWcbvEval.v b/pcuic/theories/PCUICWcbvEval.v index 6fdc13e16..e80e6f7de 100644 --- a/pcuic/theories/PCUICWcbvEval.v +++ b/pcuic/theories/PCUICWcbvEval.v @@ -163,8 +163,6 @@ Proof. simpl. intros. now rewrite mkApps_app in H. Qed. -Global Hint Resolve app_tip_nil : core. - Definition cstr_arity mdecl cdecl := (mdecl.(ind_npars) + context_assumptions cdecl.(cstr_args))%nat. @@ -1220,939 +1218,4 @@ Proof. eapply value_final, value_app. now constructor. auto. - eapply eval_cofix_case; tea. eapply value_final, value_app. now constructor. auto. -Qed. - -Lemma expanded_tApp (Σ : global_env) (Γ : list nat) (f : term) a : - expanded Σ Γ f -> expanded Σ Γ a -> - expanded Σ Γ (tApp f a). -Proof. - induction 1 using expanded_ind; intros expa. - all:rewrite -?[tApp _ a](mkApps_app _ _ [a]). - all:try (eapply (expanded_mkApps _ _ _ [a]) => //; econstructor; eauto). - - - econstructor; tea. rewrite app_length. lia. eapply app_Forall;eauto. - - econstructor; tea. eapply app_Forall; eauto. - - eapply expanded_tFix; tea. eapply app_Forall; eauto. eauto. rewrite app_length; cbn; eauto. lia. - - eapply expanded_tConstruct_app; tea. rewrite app_length ; lia. eapply app_Forall; eauto. -Qed. - -Lemma expanded_mkApps (Σ : global_env) (Γ : list nat) (f : term) (args : list term) : - expanded Σ Γ f -> Forall (expanded Σ Γ) args -> expanded Σ Γ (mkApps f args). -Proof. - intros expf expa; induction expa in f, expf |- *; eauto. - cbn. eapply IHexpa. now eapply expanded_tApp. -Qed. - -Lemma expanded_tApp_inv Σ Γ f a : - ~ isConstructApp f || isFixApp f || isRel (head f) -> - expanded Σ Γ (tApp f a) -> - expanded Σ Γ f /\ expanded Σ Γ a. -Proof. - rewrite /isConstructApp /isFixApp. - intros hf exp; depind exp. - - destruct args using rev_case; solve_discr; try discriminate. - rewrite mkApps_app in H3; noconf H3. - eapply Forall_app in H1 as [? ha]. depelim ha. - destruct args using rev_case; cbn in hf => //. - now rewrite !head_mkApps /= in hf. - - destruct args using rev_case; solve_discr. subst f6. - eapply IHexp => //. - rewrite mkApps_app in H2; noconf H2. - eapply Forall_app in H0 as [? H0]. depelim H0. split => //. - rewrite !head_mkApps in hf. - eapply expanded_mkApps => //. - - destruct args using rev_case; solve_discr. discriminate. - rewrite mkApps_app in H6; noconf H6. - eapply Forall_app in H1 as [? h]. depelim h. split => //. - now rewrite !head_mkApps /= in hf. - - destruct args using rev_case; solve_discr. discriminate. - rewrite mkApps_app in H3; noconf H3. - now rewrite !head_mkApps /= in hf. -Qed. - -Lemma expanded_tLambda_inv Σ Γ na t b : - expanded Σ Γ (tLambda na t b) -> - expanded Σ (0 :: Γ) b. -Proof. - intros exp; depind exp; solve_discr => //; eauto. -Qed. - -Lemma expanded_tLetIn_inv Σ Γ na t t' b : - expanded Σ Γ (tLetIn na t t' b) -> - expanded Σ Γ t /\ expanded Σ (0 :: Γ) b. -Proof. - intros exp; depind exp; solve_discr => //; eauto. -Qed. - -Lemma expanded_tEvar_inv Σ Γ ev l: - expanded Σ Γ (tEvar ev l) -> - Forall (expanded Σ Γ) l. -Proof. - intros exp; depind exp; solve_discr => //; eauto. -Qed. - -Lemma expanded_tCase_inv Σ Γ ci p c brs : - expanded Σ Γ (tCase ci p c brs) -> - Forall (expanded Σ Γ) (pparams p) /\ - Forall - (fun br : branch term => - ∥ All_fold - (fun (Δ : list context_decl) (d : context_decl) - => - ForOption - (fun b : term => - expanded Σ - (repeat 0 #|Δ| ++ - repeat 0 #|pparams p|) b) - (decl_body d)) (bcontext br) ∥ /\ - expanded Σ (repeat 0 #|bcontext br| ++ Γ) (bbody br)) - brs /\ - expanded Σ Γ c. -Proof. - intros exp; depind exp; solve_discr => //; eauto. -Qed. - -Lemma expanded_tProj_inv Σ Γ p c : - expanded Σ Γ (tProj p c) -> - expanded Σ Γ c. -Proof. - intros exp; depind exp; solve_discr => //; eauto. -Qed. - -Lemma expanded_tCoFix_inv Σ Γ mfix idx : - expanded Σ Γ (tCoFix mfix idx) -> - Forall (fun d : def term => expanded Σ (repeat 0 #|mfix| ++ Γ) (dbody d)) mfix. -Proof. - intros exp; depind exp; solve_discr => //; eauto. -Qed. - -Import PCUICOnOne. - -Module Red1Apps. - - - Inductive red1 (Σ : global_env) (Γ : context) : term -> term -> Type := - (** Reductions *) - (** Beta *) - | red_beta na t b a ts : - Σ ;;; Γ |- mkApps (tLambda na t b) (a :: ts) ⇝ mkApps (b {0 := a}) ts - - (** Let *) - | red_zeta na b t b' : - Σ ;;; Γ |- tLetIn na b t b' ⇝ b' {0 := b} - - | red_rel i body : - option_map decl_body (nth_error Γ i) = Some (Some body) -> - Σ ;;; Γ |- tRel i ⇝ lift0 (S i) body - - (** Case *) - | red_iota ci c u args p brs br : - nth_error brs c = Some br -> - #|args| = (ci.(ci_npar) + context_assumptions br.(bcontext))%nat -> - Σ ;;; Γ |- tCase ci p (mkApps (tConstruct ci.(ci_ind) c u) args) brs - ⇝ iota_red ci.(ci_npar) p args br - - (** Fix unfolding, with guard *) - | red_fix mfix idx args narg fn : - unfold_fix mfix idx = Some (narg, fn) -> - is_constructor narg args = true -> - Σ ;;; Γ |- mkApps (tFix mfix idx) args ⇝ mkApps fn args - - (** CoFix-case unfolding *) - | red_cofix_case ip p mfix idx args narg fn brs : - unfold_cofix mfix idx = Some (narg, fn) -> - Σ ;;; Γ |- tCase ip p (mkApps (tCoFix mfix idx) args) brs ⇝ - tCase ip p (mkApps fn args) brs - - (** CoFix-proj unfolding *) - | red_cofix_proj p mfix idx args narg fn : - unfold_cofix mfix idx = Some (narg, fn) -> - Σ ;;; Γ |- tProj p (mkApps (tCoFix mfix idx) args) ⇝ tProj p (mkApps fn args) - - (** Constant unfolding *) - | red_delta c decl body (isdecl : declared_constant Σ c decl) u : - decl.(cst_body) = Some body -> - Σ ;;; Γ |- tConst c u ⇝ subst_instance u body - - (** Proj *) - | red_proj p args u arg: - nth_error args (p.(proj_npars) + p.(proj_arg)) = Some arg -> - Σ ;;; Γ |- tProj p (mkApps (tConstruct p.(proj_ind) 0 u) args) ⇝ arg - - - | abs_red_l na M M' N : Σ ;;; Γ |- M ⇝ M' -> Σ ;;; Γ |- tLambda na M N ⇝ tLambda na M' N - | abs_red_r na M M' N : Σ ;;; Γ ,, vass na N |- M ⇝ M' -> Σ ;;; Γ |- tLambda na N M ⇝ tLambda na N M' - - | letin_red_def na b t b' r : Σ ;;; Γ |- b ⇝ r -> Σ ;;; Γ |- tLetIn na b t b' ⇝ tLetIn na r t b' - | letin_red_ty na b t b' r : Σ ;;; Γ |- t ⇝ r -> Σ ;;; Γ |- tLetIn na b t b' ⇝ tLetIn na b r b' - | letin_red_body na b t b' r : Σ ;;; Γ ,, vdef na b t |- b' ⇝ r -> Σ ;;; Γ |- tLetIn na b t b' ⇝ tLetIn na b t r - - | case_red_param ci p params' c brs : - OnOne2 (fun t u => Σ ;;; Γ |- t ⇝ u) p.(pparams) params' -> - Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci (set_pparams p params') c brs - - | case_red_return ci p preturn' c brs : - Σ ;;; Γ ,,, inst_case_predicate_context p |- p.(preturn) ⇝ preturn' -> - Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci (set_preturn p preturn') c brs - - | case_red_discr ci p c c' brs : - Σ ;;; Γ |- c ⇝ c' -> Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci p c' brs - - | case_red_brs ci p c brs brs' : - OnOne2 (fun br br' => - on_Trel_eq (fun t u => Σ ;;; Γ ,,, inst_case_branch_context p br |- t ⇝ u) bbody bcontext br br') - brs brs' -> - Σ ;;; Γ |- tCase ci p c brs ⇝ tCase ci p c brs' - - | proj_red p c c' : Σ ;;; Γ |- c ⇝ c' -> Σ ;;; Γ |- tProj p c ⇝ tProj p c' - - | app_red_l M1 N1 M2 : Σ ;;; Γ |- M1 ⇝ N1 -> ~~ isApp M1 -> ~~ isFix M1 -> M2 <> nil -> - Σ ;;; Γ |- mkApps M1 M2 ⇝ mkApps N1 M2 - - | app_red_r M2 N2 M1 : ~~ isApp M1 -> OnOne2 (fun M2 N2 => Σ ;;; Γ |- M2 ⇝ N2) M2 N2 -> Σ ;;; Γ |- mkApps M1 M2 ⇝ mkApps M1 N2 - - | app_fix_red_ty mfix0 mfix1 idx M2 : - OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ |- t ⇝ u) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> - Σ ;;; Γ |- mkApps (tFix mfix0 idx) M2 ⇝ mkApps (tFix mfix1 idx) M2 - - | app_fix_red_body mfix0 mfix1 idx M2 : - OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> - Σ ;;; Γ |- mkApps (tFix mfix0 idx) M2 ⇝ mkApps (tFix mfix1 idx) M2 - - | prod_red_l na M1 M2 N1 : Σ ;;; Γ |- M1 ⇝ N1 -> Σ ;;; Γ |- tProd na M1 M2 ⇝ tProd na N1 M2 - | prod_red_r na M2 N2 M1 : Σ ;;; Γ ,, vass na M1 |- M2 ⇝ N2 -> - Σ ;;; Γ |- tProd na M1 M2 ⇝ tProd na M1 N2 - - | evar_red ev l l' : OnOne2 (fun t u => Σ ;;; Γ |- t ⇝ u) l l' -> Σ ;;; Γ |- tEvar ev l ⇝ tEvar ev l' - - | fix_red_ty mfix0 mfix1 idx : - OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ |- t ⇝ u) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> - Σ ;;; Γ |- tFix mfix0 idx ⇝ tFix mfix1 idx - - | fix_red_body mfix0 mfix1 idx : - OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) dbody (fun x => (dname x, dtype x, rarg x))) - mfix0 mfix1 -> - Σ ;;; Γ |- tFix mfix0 idx ⇝ tFix mfix1 idx - - | cofix_red_ty mfix0 mfix1 idx : - OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ |- t ⇝ u) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> - Σ ;;; Γ |- tCoFix mfix0 idx ⇝ tCoFix mfix1 idx - - | cofix_red_body mfix0 mfix1 idx : - OnOne2 (on_Trel_eq (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> - Σ ;;; Γ |- tCoFix mfix0 idx ⇝ tCoFix mfix1 idx - where " Σ ;;; Γ |- t ⇝ u " := (red1 Σ Γ t u). - -Lemma red1_ind_all : - forall (Σ : global_env) (P : context -> term -> term -> Type), - - (forall (Γ : context) (na : aname) (t b a : term) ts, - P Γ (mkApps (tLambda na t b) (a :: ts)) (mkApps (b {0 := a}) ts)) -> - - (forall (Γ : context) (na : aname) (b t b' : term), P Γ (tLetIn na b t b') (b' {0 := b})) -> - - (forall (Γ : context) (i : nat) (body : term), - option_map decl_body (nth_error Γ i) = Some (Some body) -> P Γ (tRel i) ((lift0 (S i)) body)) -> - - (forall (Γ : context) (ci : case_info) (c : nat) (u : Instance.t) (args : list term) - (p : predicate term) (brs : list (branch term)) br, - nth_error brs c = Some br -> - #|args| = (ci.(ci_npar) + context_assumptions br.(bcontext))%nat -> - P Γ (tCase ci p (mkApps (tConstruct ci.(ci_ind) c u) args) brs) - (iota_red ci.(ci_npar) p args br)) -> - - (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 Γ (mkApps (tFix mfix idx) args) (mkApps fn args)) -> - - (forall (Γ : context) ci (p : predicate term) (mfix : mfixpoint term) (idx : nat) - (args : list term) (narg : nat) (fn : term) (brs : list (branch term)), - unfold_cofix mfix idx = Some (narg, fn) -> - P Γ (tCase ci p (mkApps (tCoFix mfix idx) args) brs) (tCase ci p (mkApps fn args) brs)) -> - - (forall (Γ : context) (p : projection) (mfix : mfixpoint term) (idx : nat) (args : list term) - (narg : nat) (fn : term), - unfold_cofix mfix idx = Some (narg, fn) -> P Γ (tProj p (mkApps (tCoFix mfix idx) args)) (tProj p (mkApps fn args))) -> - - (forall (Γ : context) c (decl : constant_body) (body : term), - declared_constant Σ c decl -> - forall u : Instance.t, cst_body decl = Some body -> P Γ (tConst c u) (subst_instance u body)) -> - - (forall (Γ : context) p (args : list term) (u : Instance.t) - (arg : term), - nth_error args (p.(proj_npars) + p.(proj_arg)) = Some arg -> - P Γ (tProj p (mkApps (tConstruct p.(proj_ind) 0 u) args)) arg) -> - - (forall (Γ : context) (na : aname) (M M' N : term), - red1 Σ Γ M M' -> P Γ M M' -> P Γ (tLambda na M N) (tLambda na M' N)) -> - - (forall (Γ : context) (na : aname) (M M' N : term), - red1 Σ (Γ,, vass na N) M M' -> P (Γ,, vass na N) M M' -> P Γ (tLambda na N M) (tLambda na N M')) -> - - (forall (Γ : context) (na : aname) (b t b' r : term), - red1 Σ Γ b r -> P Γ b r -> P Γ (tLetIn na b t b') (tLetIn na r t b')) -> - - (forall (Γ : context) (na : aname) (b t b' r : term), - red1 Σ Γ t r -> P Γ t r -> P Γ (tLetIn na b t b') (tLetIn na b r b')) -> - - (forall (Γ : context) (na : aname) (b t b' r : term), - red1 Σ (Γ,, vdef na b t) b' r -> P (Γ,, vdef na b t) b' r -> P Γ (tLetIn na b t b') (tLetIn na b t r)) -> - - (forall (Γ : context) (ci : case_info) p params' c brs, - OnOne2 (Trel_conj (red1 Σ Γ) (P Γ)) p.(pparams) params' -> - P Γ (tCase ci p c brs) - (tCase ci (set_pparams p params') c brs)) -> - - (forall (Γ : context) (ci : case_info) p preturn' c brs, - red1 Σ (Γ ,,, inst_case_predicate_context p) p.(preturn) preturn' -> - P (Γ ,,, inst_case_predicate_context p) p.(preturn) preturn' -> - P Γ (tCase ci p c brs) - (tCase ci (set_preturn p preturn') c brs)) -> - - (forall (Γ : context) (ind : case_info) (p : predicate term) (c c' : term) (brs : list (branch term)), - red1 Σ Γ c c' -> P Γ c c' -> P Γ (tCase ind p c brs) (tCase ind p c' brs)) -> - - (forall (Γ : context) ci p c brs brs', - OnOne2 (fun br br' => - (on_Trel_eq (Trel_conj (red1 Σ (Γ ,,, inst_case_branch_context p br)) - (P (Γ ,,, inst_case_branch_context p br))) - bbody bcontext br br')) brs brs' -> - P Γ (tCase ci p c brs) (tCase ci p c brs')) -> - - (forall (Γ : context) (p : projection) (c c' : term), red1 Σ Γ c c' -> P Γ c c' -> - P Γ (tProj p c) (tProj p c')) -> - - (forall (Γ : context) M1 N1 M2, red1 Σ Γ M1 N1 -> P Γ M1 N1 -> ~~ isApp M1 -> ~~ isFix M1 -> M2 <> [] -> - P Γ (mkApps M1 M2) (mkApps N1 M2)) -> - - (forall (Γ : context) M1 N2 M2, ~~ isApp M1 -> - OnOne2 (Trel_conj (red1 Σ Γ) (P Γ)) M2 N2 -> - P Γ (mkApps M1 M2) (mkApps M1 N2)) -> - - (forall (Γ : context) mfix0 mfix1 idx M2, - OnOne2 (on_Trel_eq (Trel_conj (fun t u => Σ ;;; Γ |- t ⇝ u) (P Γ)) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> - P Γ (mkApps (tFix mfix0 idx) M2) (mkApps (tFix mfix1 idx) M2)) -> - - (forall (Γ : context) mfix0 mfix1 idx M2, - OnOne2 (on_Trel_eq (Trel_conj (fun t u => Σ ;;; Γ ,,, fix_context mfix0 |- t ⇝ u) (P (Γ ,,, fix_context mfix0))) dbody (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> - P Γ (mkApps (tFix mfix0 idx) M2) (mkApps (tFix mfix1 idx) M2)) -> - - (forall (Γ : context) (na : aname) (M1 M2 N1 : term), - red1 Σ Γ M1 N1 -> P Γ M1 N1 -> P Γ (tProd na M1 M2) (tProd na N1 M2)) -> - - (forall (Γ : context) (na : aname) (M2 N2 M1 : term), - red1 Σ (Γ,, vass na M1) M2 N2 -> P (Γ,, vass na M1) M2 N2 -> P Γ (tProd na M1 M2) (tProd na M1 N2)) -> - - (forall (Γ : context) (ev : nat) (l l' : list term), - OnOne2 (Trel_conj (red1 Σ Γ) (P Γ)) l l' -> P Γ (tEvar ev l) (tEvar ev l')) -> - - (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), - OnOne2 (on_Trel_eq (Trel_conj (red1 Σ Γ) (P Γ)) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> - P Γ (tFix mfix0 idx) (tFix mfix1 idx)) -> - - (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), - OnOne2 (on_Trel_eq (Trel_conj (red1 Σ (Γ ,,, fix_context mfix0)) - (P (Γ ,,, fix_context mfix0))) dbody - (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> - P Γ (tFix mfix0 idx) (tFix mfix1 idx)) -> - - (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), - OnOne2 (on_Trel_eq (Trel_conj (red1 Σ Γ) (P Γ)) dtype (fun x => (dname x, dbody x, rarg x))) mfix0 mfix1 -> - P Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx)) -> - - (forall (Γ : context) (mfix0 mfix1 : list (def term)) (idx : nat), - OnOne2 (on_Trel_eq (Trel_conj (red1 Σ (Γ ,,, fix_context mfix0)) - (P (Γ ,,, fix_context mfix0))) dbody - (fun x => (dname x, dtype x, rarg x))) mfix0 mfix1 -> - P Γ (tCoFix mfix0 idx) (tCoFix mfix1 idx)) -> - - forall (Γ : context) (t t0 : term), red1 Σ Γ t t0 -> P Γ t t0. -Proof. - intros. rename X29 into Xlast. revert Γ t t0 Xlast. - fix aux 4. intros Γ t T. - move aux at top. - destruct 1; match goal with - | |- P _ (tFix _ _) (tFix _ _) => idtac - | |- P _ (tCoFix _ _) (tCoFix _ _) => idtac - | |- P _ (mkApps (tFix _ _) _) _ => idtac - | |- P _ (tCase _ _ (mkApps (tCoFix _ _) _) _) _ => idtac - | |- P _ (tProj _ (mkApps (tCoFix _ _) _)) _ => idtac - | H : _ |- _ => eapply H; eauto - end. - - eapply X3; eauto. - - eapply X4; eauto. - - eapply X5; eauto. - - - revert params' o. - generalize (pparams p). - fix auxl 3. - intros params params' []. - + constructor. split; auto. - + constructor. auto. - - - revert brs' o. - revert brs. - fix auxl 3. - intros l l' Hl. destruct Hl. - + simpl in *. constructor; intros; intuition auto. - + constructor. eapply auxl. apply Hl. - - - revert M2 N2 o. - fix auxl 3. - intros l l' Hl. destruct Hl. - + constructor. split; auto. - + constructor. auto. - - - eapply X20. - revert mfix0 mfix1 o. - fix auxl 3. - intros l l' Hl. destruct Hl. - + constructor; split; auto; intuition. - + constructor; try split; auto; intuition. - - - eapply X21. revert o. - generalize (fix_context mfix0). - intros c o. revert mfix0 mfix1 o. - fix auxl 3. - intros l l' Hl. destruct Hl. - + constructor. split; auto; intuition. - + constructor; try split; auto; intuition. - - - revert l l' o. - fix auxl 3. - intros l l' Hl. destruct Hl. - + constructor. split; auto. - + constructor. auto. - - - eapply X25. - revert mfix0 mfix1 o; fix auxl 3. - intros l l' Hl; destruct Hl; - constructor; try split; auto; intuition. - - - eapply X26. - revert o. generalize (fix_context mfix0). intros c Xnew. - revert mfix0 mfix1 Xnew; fix auxl 3; intros l l' Hl; - destruct Hl; constructor; try split; auto; intuition. - - - eapply X27. - revert mfix0 mfix1 o. - fix auxl 3; intros l l' Hl; destruct Hl; - constructor; try split; auto; intuition. - - - eapply X28. - revert o. generalize (fix_context mfix0). intros c new. - revert mfix0 mfix1 new; fix auxl 3; intros l l' Hl; destruct Hl; - constructor; try split; auto; intuition. -Defined. - -Lemma red_tApp Σ Γ t v u : - red1 Σ Γ t v -> - red1 Σ Γ (tApp t u) (tApp v u). -Proof. - induction 1. - all:try solve [eapply (app_red_l _ _ _ _ [u]) => //; econstructor; eauto]. - - rewrite -![tApp _ u](mkApps_app _ _ [u]). - eapply red_beta. - - rewrite -![tApp _ u](mkApps_app _ _ [u]). econstructor; eauto. - now apply is_constructor_app_ge. - - rewrite -![tApp _ u](mkApps_app _ _ [u]). econstructor; eauto. - - rewrite -![tApp _ u](mkApps_app _ _ [u]). econstructor; eauto. - now eapply OnOne2_app_r. - - rewrite -![tApp _ u](mkApps_app _ _ [u]). now eapply app_fix_red_ty. - - rewrite -![tApp _ u](mkApps_app _ _ [u]). now eapply app_fix_red_body. - - now eapply (app_fix_red_ty _ _ _ _ _ [_]). - - now eapply (app_fix_red_body _ _ _ _ _ [_]). -Qed. - -Lemma isLambda_red1 Σ Γ b b' : isLambda b -> red1 Σ Γ b b' -> isLambda b'. -Proof. - destruct b; simpl; try discriminate. - intros _ red. - depelim red; solve_discr; eauto. depelim o. -Qed. - -End Red1Apps. - -Lemma red1_red1apps Σ Γ t v : - red1 Σ Γ t v -> Red1Apps.red1 Σ Γ t v. -Proof. - intros red; induction red using red1_ind_all in |- *. - all:try solve [econstructor; eauto; solve_all; eauto]. - - eapply (Red1Apps.red_beta _ _ _ _ _ _ []). - - now eapply Red1Apps.red_tApp. - - remember (decompose_app (tApp M1 M2)). - destruct p as [hd args]. - edestruct (decompose_app_rec_inv' M1 [M2]). rewrite /decompose_app in Heqp. - cbn in Heqp. rewrite -Heqp. reflexivity. - destruct a as [napp [hm2 hm1]]. - symmetry in Heqp; eapply decompose_app_inv in Heqp. rewrite Heqp. - assert (tApp M1 N2 = mkApps hd (firstn x args ++ [N2])) as ->. - { now rewrite mkApps_app -hm1. } - rewrite -{1}(firstn_skipn x args) -hm2. eapply Red1Apps.app_red_r => //. - eapply OnOne2_app. now constructor. -Qed. - -Lemma head_nApp f : - ~~ isApp f -> head f = f. -Proof. - induction f => //. -Qed. - -Lemma expanded_mkApps_inv Σ Γ f v : - ~~ isApp f -> - ~~ (isConstruct f || isFix f || isRel f) -> - expanded Σ Γ (mkApps f v) -> - expanded Σ Γ f /\ Forall (expanded Σ Γ) v. -Proof. - intros happ hc. - induction v using rev_ind; cbn. - - intros exp; split => //. - - rewrite mkApps_app /=. - move/expanded_tApp_inv => e. - forward e. rewrite /isConstructApp /isFixApp head_mkApps. - rewrite head_nApp //. now move/negbTE: hc ->. - intuition auto. eapply app_Forall; eauto. -Qed. - - -Lemma arguments_mkApps f args : - ~~ isApp f -> - arguments (mkApps f args) = args. -Proof. - rewrite /arguments => isa. - rewrite decompose_app_mkApps //. -Qed. - -Lemma arguments_mkApps' f args : - arguments (mkApps f args) = arguments f ++ args. -Proof. - rewrite /arguments. - rewrite /decompose_app decompose_app_rec_mkApps //. - rewrite app_nil_r. - induction f in args |- * => //=. - rewrite IHf1. now rewrite (IHf1 [f2]) -app_assoc. -Qed. - -Lemma expanded_mkApps_inv' Σ Γ f : - expanded Σ Γ f -> - let args := arguments f in - Forall (expanded Σ Γ) args /\ - match head f with - | tRel n => exists m, nth_error Γ n = Some m /\ m <= #|args| - | tConstruct ind c u => exists mind idecl cdecl, - declared_constructor Σ (ind, c) mind idecl cdecl /\ #|args| >= ind_npars mind + context_assumptions (cstr_args cdecl) - | tFix mfix idx => exists d, - Forall - (fun d0 : def term => - isLambda (dbody d0) /\ - (let ctx := - rev_map (fun d1 : def term => (1 + rarg d1)%nat) - mfix in - expanded Σ (ctx ++ Γ) (dbody d0))) mfix /\ - args <> [] /\ - nth_error mfix idx = Some d /\ #|args| > rarg d - | _ => expanded Σ Γ (head f) - end. -Proof. - induction 1 using expanded_ind; cbn. - all: try solve [split; econstructor; eauto]. - all:rewrite head_mkApps /=. - - rewrite arguments_mkApps //. split => //. now exists m. - - destruct IHexpanded. rewrite arguments_mkApps'. split. - eapply app_Forall => //. - rewrite app_length. - destruct (head f6) => //; firstorder (eauto; try lia). - exists x. split => //. firstorder (eauto; try lia). - intros heq; apply H5. now eapply app_eq_nil in heq. - - rewrite arguments_mkApps //. split => //. now exists d. - - rewrite arguments_mkApps //. split => //. - firstorder. -Qed. - -Lemma All_fold_map_context (P : context -> context_decl -> Type) (f : term -> term) ctx : - All_fold P (map_context f ctx) <~> All_fold (fun Γ d => P (map_context f Γ) (map_decl f d)) ctx. -Proof. - induction ctx. - - split; constructor. - - cbn. split; intros H; depelim H; constructor; auto; now apply IHctx. -Qed. - -Lemma expanded_fix_subst Σ a k b Γ Γs Δ : - #|Γ| = k -> - Forall2 (fun n arg => forall args Γ' k', - Forall (expanded Σ (Γ' ++ Δ)) args -> - n <= #|args| -> - #|Γ'| = k' -> - expanded Σ (Γ' ++ Δ) (mkApps (lift0 k' arg) args)) Γs a -> - expanded Σ (Γ ++ Γs ++ Δ) b -> - expanded Σ (Γ ++ Δ) (subst a k b). -Proof. - intros Hk Hs. - remember (Γ ++ _ ++ Δ)%list as Γ_. - intros exp; revert Γ k Hk HeqΓ_. - induction exp using expanded_ind; intros Γ' k Hk ->. - all:try solve[ cbn; econstructor => // ]. - 2,7:solve[ cbn; econstructor => //; solve_all ]. - - rewrite subst_mkApps /=. - destruct (Nat.leb_spec k n). - destruct (nth_error a _) eqn:hnth. - * pose proof (Forall2_length Hs). - pose proof (Forall2_nth_error_Some_r _ _ _ _ _ _ _ hnth Hs) as [t' [hnth' hargs]]. - eapply hargs => //. - eapply Forall_map. - 2:{ len. rewrite nth_error_app_ge in H. lia. subst k. - rewrite nth_error_app_lt in H. - eapply nth_error_Some_length in hnth'. lia. rewrite hnth' in H. noconf H. exact H0. } - solve_all. - * rewrite nth_error_app_ge in H. lia. - eapply nth_error_None in hnth. - pose proof (Forall2_length Hs). - rewrite nth_error_app_ge in H. lia. - eapply expanded_tRel. rewrite nth_error_app_ge. lia. erewrite <- H. - lia_f_equal. len. solve_all. - * rewrite nth_error_app_lt in H. lia. - eapply expanded_tRel. rewrite nth_error_app_lt. lia. tea. now len. - solve_all. - - cbn. econstructor. - eapply (IHexp (0 :: Γ') (S k)); cbn; auto; try lia. - - cbn. econstructor. apply IHexp1; auto. - eapply (IHexp2 (0 :: Γ') (S k)); cbn; auto; lia. - - rewrite subst_mkApps. - destruct (isConstruct (subst a k f6) || isFix (subst a k f6) || isRel (subst a k f6)) eqn:eqc. - specialize (IHexp _ _ Hk eq_refl). - eapply expanded_mkApps_expanded => //. solve_all. - eapply expanded_mkApps => //. now eapply IHexp. solve_all. - - cbn. econstructor. eauto. cbn. solve_all. solve_all. - specialize (H1 (repeat 0 #|bcontext x| ++ Γ') (#|bcontext x| + k)%nat). - forward H1 by len. - forward H1. now rewrite app_assoc. - rewrite /id. rewrite app_assoc. apply H1. - - rewrite subst_mkApps. cbn. - eapply expanded_tFix. - + solve_all. now eapply isLambda_subst. - specialize (a0 - (rev_map (fun d0 : def term => S (rarg d0)) - (map (map_def (subst a k) (subst a (#|mfix| + k))) mfix) ++ Γ') (#|mfix| + k)%nat). - forward a0 by len. - forward a0. { rewrite app_assoc. f_equal. f_equal. - rewrite !rev_map_spec. f_equal. now rewrite map_map_compose /=. } - rewrite app_assoc. eapply a0. - + solve_all. - + now destruct args. - + rewrite nth_error_map /= H4 //. - + len. - - cbn. constructor. - solve_all. - specialize (a0 (repeat 0 #|mfix| ++ Γ') (#|mfix| + k)%nat). - forward a0 by len. - forward a0. { rewrite app_assoc //. } - rewrite app_assoc. eapply a0. - - rewrite subst_mkApps. cbn. - eapply expanded_tConstruct_app; tea. now len. - solve_all. -Qed. - -Lemma expanded_unfold_fix Σ Γ' mfix idx narg fn : - unfold_fix mfix idx = Some (narg, fn) -> - All (fun d0 : def term => isLambda (dbody d0) /\ expanded Σ (rev_map (fun d1 : def term => S (rarg d1)) mfix ++ Γ') (dbody d0)) mfix -> - expanded Σ Γ' fn. -Proof. - unfold unfold_fix. - destruct nth_error eqn:e => //. - intros [= <- <-] hf. - pose proof (nth_error_all e hf) as [hl hf']. - eapply (expanded_fix_subst _ _ _ _ []) => //; tea. - rewrite rev_map_spec. - eapply Forall2_from_nth_error. len. - intros n rarg f. len. intros hn hrarg hnthf args Γ'' k' hargs hrarg' <-. - eapply PCUICParallelReductionConfluence.nth_error_fix_subst in hnthf. subst f. - move: hrarg. - rewrite nth_error_rev; len. rewrite List.rev_involutive nth_error_map. - intros hrarg. - destruct (nth_error mfix (_ - _)) eqn:e'. cbn in hrarg. noconf hrarg. - eapply expanded_tFix => //. solve_all. - eapply expanded_lift; len. rewrite !rev_map_spec in H1 *. - rewrite map_map => //. destruct args => //. cbn in hrarg'. lia. - rewrite nth_error_map /= e' /= //. cbn. lia. - eapply nth_error_None in e'. lia. -Qed. - -Lemma expanded_unfold_cofix Σ Γ' mfix idx narg fn : - unfold_cofix mfix idx = Some (narg, fn) -> - All (fun d0 : def term => expanded Σ (repeat 0 #|mfix| ++ Γ') (dbody d0)) mfix -> - expanded Σ Γ' fn. -Proof. - unfold unfold_cofix. - destruct nth_error eqn:e => //. - intros [= <- <-] hf. - pose proof (nth_error_all e hf) as hf'. - eapply (expanded_subst _ _ _ _ []) => //; tea. - rewrite /cofix_subst. - generalize #|mfix|. - induction n; repeat constructor; eauto. solve_all. - len. -Qed. - -Lemma expanded_weakening_env {cf : checker_flags} Σ Σ' Γ t : - wf Σ' -> - extends Σ Σ' -> - expanded Σ Γ t -> expanded Σ' Γ t. -Proof. - intros w s. - eapply expanded_ind; intros. - all:try solve [econstructor; eauto]. - - econstructor; eauto. solve_all. sq. eapply All_fold_impl; tea. cbn. - intros. now rewrite repeat_app in H6. - - eapply expanded_tFix; eauto. solve_all. - - eapply expanded_tConstruct_app; eauto. - eapply PCUICWeakeningEnv.weakening_env_declared_constructor; tea. -Qed. - -Lemma expanded_global_env_constant {cf : checker_flags} Σ c decl : - wf Σ -> - expanded_global_env Σ -> - declared_constant Σ c decl -> - ForOption (expanded Σ []) (cst_body decl). -Proof. - intros wf; destruct Σ as [Σ univs] => /=. cbn. - unfold expanded_global_env. cbn. - induction 1; cbn => //. - intros [->|H']. - - depelim H0. - destruct decl as [? [] ?]; cbn in *. - constructor. cbn. - eapply expanded_weakening_env; tea. - eapply extends_strictly_on_decls_extends. - split => //=. eapply incl_cs_refl. 2:eapply Retroknowledge.extends_refl. - set (cb := ConstantDecl _). now exists [(c, cb)]. - constructor. - - forward IHexpanded_global_declarations. - destruct wf. cbn in *. split => //. - cbn. now depelim o0. - eapply IHexpanded_global_declarations in H'. - destruct decl as [? [] ?]; cbn in * => //. 2:constructor. - depelim H'. constructor. - eapply expanded_weakening_env; tea. - eapply extends_strictly_on_decls_extends. - split => //=. eapply incl_cs_refl. 2:eapply Retroknowledge.extends_refl. - now exists [decl0]. -Qed. - -Lemma All_fold_nth_error P (ctx : context) n b : - All_fold P ctx -> nth_error ctx n = Some b -> - P (skipn (S n) ctx) b. -Proof. - induction 1 in n, b |- *. - - rewrite nth_error_nil //. - - destruct n => //=. - * intros [= <-]. now rewrite skipn_S skipn_0. - * intros hnth. now rewrite skipn_S. -Qed. - -Lemma skipn_repeat {A} k n (a : A) : - skipn n (repeat a k) = repeat a (k - n). -Proof. - induction n in k |- *. - - rewrite skipn_0. lia_f_equal. - - destruct k => //=. - rewrite skipn_S IHn. lia_f_equal. -Qed. - -Lemma expanded_red {cf : checker_flags} {Σ : global_env_ext} Γ Γ' t v : wf Σ -> - expanded_global_env Σ -> - (forall n body, option_map decl_body (nth_error Γ n) = Some (Some body) -> expanded Σ (skipn (S n) Γ') body) -> - red1 Σ Γ t v -> - expanded Σ Γ' t -> - expanded Σ Γ' v. -Proof. - move=> wf expΣ wfΓ /red1_red1apps red. - induction red using Red1Apps.red1_ind_all in wfΓ, Γ' |- *; intros exp. - - eapply expanded_mkApps_inv in exp as [] => //. - eapply expanded_tLambda_inv in H. - depelim H0. - eapply expanded_mkApps => //. - eapply (expanded_subst _ _ _ _ [] Γ') => //. now constructor. - - eapply expanded_tLetIn_inv in exp as []. - eapply (expanded_subst _ _ _ _ [] Γ') => //. now constructor. - - rewrite -(firstn_skipn (S i) Γ'). - eapply expanded_mkApps_inv' in exp; cbn in exp. - destruct exp as [_ [m []]]. eapply nth_error_Some_length in H0. - eapply (expanded_lift _ _ _ _ []) => //. - rewrite firstn_length_le; try lia. - now eapply wfΓ. - - eapply expanded_tCase_inv in exp as [? []]. - unfold iota_red. - move/expanded_mkApps_inv': H3 => /=. - rewrite arguments_mkApps // head_mkApps //=. - intros [hargs [mind [idecl [cdecl [declc hargs']]]]]. - eapply nth_error_forall in H2; tea. - destruct H2 as [[hbctx] hbod]. - eapply (expanded_subst _ _ _ _ [] _) => //. - + eapply Forall_rev, Forall_skipn => //. - + len. replace #|skipn (ci_npar ci) args| with (context_assumptions (inst_case_branch_context p br)). - eapply expanded_let_expansion; len. red. - sq. rewrite /inst_case_branch_context /inst_case_context /subst_context. - eapply PCUICParallelReduction.All_fold_fold_context_k. - eapply All_fold_map_context, All_fold_impl; tea; cbn. - intros ? ? fo. len. destruct x as [? [] ?] => //; constructor. - cbn in fo. depelim fo. eapply (expanded_subst _ _ _ _ (repeat 0 #|Γ0|) _); len. - eapply Forall_rev; eauto. - eapply expanded_subst_instance. rewrite app_assoc. now eapply expanded_weakening. - rewrite skipn_length. len. - - move/expanded_mkApps_inv': exp. cbn. - rewrite arguments_mkApps // head_mkApps //=. - move=> [hargs [d [hf [hargs' []]]]] hnth hrarg. - eapply expanded_mkApps_expanded => //; solve_all. - eapply expanded_unfold_fix in H; tea. - - eapply expanded_tCase_inv in exp as [? []]. - constructor => //. - eapply expanded_mkApps_inv' in H2. move: H2; rewrite arguments_mkApps // head_mkApps //=. - intros [hargs' hcof]. cbn in hcof. eapply expanded_tCoFix_inv in hcof. - eapply expanded_unfold_cofix in H; tea. eapply expanded_mkApps; tea => //. solve_all. - - eapply expanded_tProj_inv in exp. econstructor. - eapply expanded_mkApps_inv' in exp. move: exp; rewrite arguments_mkApps // head_mkApps //=. - intros [hargs' hcof]. cbn in hcof. eapply expanded_tCoFix_inv in hcof. - eapply expanded_mkApps => //. - eapply expanded_unfold_cofix in H; tea. solve_all. - - eapply expanded_subst_instance. - eapply expanded_global_env_constant in expΣ; tea. - eapply (expanded_weakening _ []). rewrite H0 in expΣ. now depelim expΣ. - - eapply expanded_tProj_inv in exp. - move/expanded_mkApps_inv': exp. rewrite arguments_mkApps // head_mkApps //=. - intros []. - eapply nth_error_forall in H0; tea. - - constructor. now eapply expanded_tLambda_inv in exp. - - constructor. eapply expanded_tLambda_inv in exp. - eapply IHred => //. - { intros [] body; cbn => //. rewrite skipn_S. eapply wfΓ. } - - eapply expanded_tLetIn_inv in exp; now constructor. - - eapply expanded_tLetIn_inv in exp. now constructor. - - eapply expanded_tLetIn_inv in exp. constructor; intuition eauto. - eapply IHred => //. - { intros [] ? => //=. intros [=]. subst b. now rewrite skipn_S skipn_0. - rewrite skipn_S. eapply wfΓ. } - - eapply expanded_tCase_inv in exp as [? []]. - constructor; eauto. cbn. - solve_all. eapply OnOne2_impl_All_r; tea. intuition eauto. now apply X0. - now rewrite /PCUICOnOne.set_pparams /= -(OnOne2_length X). - - eapply expanded_tCase_inv in exp as [? []]. - constructor; eauto. - - eapply expanded_tCase_inv in exp as [? []]. - econstructor; eauto. - - eapply expanded_tCase_inv in exp as [? []]. - econstructor; eauto. solve_all. - eapply OnOne2_impl_All_r; tea. - intros x y [? ?]. intros [[] ?]. rewrite -e0. - solve_all. eapply e => //. - intros n b. - clear -H H2 wfΓ. - destruct nth_error eqn:e' => //. - cbn. intros [=]. destruct c as [? [] ?]. cbn in *. noconf H1. - eapply nth_error_app_inv in e' as [[]|[]]. - { rewrite inst_case_branch_context_length in H0. destruct H2. - rewrite /inst_case_branch_context /inst_case_context in H1. - destruct (nth_error (bcontext x) n) eqn:e. - 2:{ eapply nth_error_None in e. rewrite skipn_app skipn_all2. len. cbn. len. } - rewrite /subst_context in H1. - erewrite nth_error_fold_context_k in H1. 4:{ rewrite nth_error_map e //. } 3:len. 2:exact []. - len in H1. noconf H1. destruct c as [? [] ?]; noconf H1. - rewrite skipn_app. len. eapply All_fold_nth_error in X; tea. cbn in X. depelim X. - rewrite skipn_length in H1. - eapply expanded_subst. rewrite skipn_length. len. - replace (S n - #|bcontext x|) with 0. 2:{ lia. } rewrite skipn_0. eapply Forall_rev. solve_all. - len. rewrite app_assoc. eapply expanded_weakening. eapply expanded_subst_instance. - now rewrite skipn_repeat. } - { rewrite inst_case_branch_context_length in H0, H1. - rewrite skipn_app skipn_all2 /=; len. - replace (S n - #|bcontext x|) with (S (n - #|bcontext x|)). 2:lia. - eapply wfΓ. rewrite H1 //. } - noconf H1. - - eapply expanded_tProj_inv in exp. now econstructor. - - eapply expanded_mkApps_inv' in exp. - rewrite head_mkApps head_nApp in exp => //. - rewrite arguments_mkApps // in exp. destruct exp as []. - specialize (IHred Γ' wfΓ). - destruct M1; try solve [eapply expanded_mkApps => //; eauto]. - * depelim red; solve_discr. eapply wfΓ in e. - eapply expanded_mkApps => //. - rewrite -(firstn_skipn (S n) Γ'). eapply (expanded_lift _ _ _ _ []) => //. - destruct H3 as [m [hn ha]]. - rewrite firstn_length_le //. now eapply nth_error_Some_length in hn. - depelim o. - * depelim red; solve_discr. depelim o. - - - eapply expanded_mkApps_inv' in exp. - move: exp. - rewrite head_mkApps head_nApp // arguments_mkApps //. - intros []. - assert(Forall (expanded Σ Γ') N2). - { clear H1. solve_all. eapply OnOne2_impl_All_r; tea. cbn. intuition auto. } - destruct M1; try solve [eapply expanded_mkApps => //; eauto]. - + rewrite (OnOne2_length X) in H1. destruct H1 as [m []]; eapply expanded_tRel; tea. - + rewrite (OnOne2_length X) in H1. destruct H1 as [m [? [? []]]]; eapply expanded_tConstruct_app; tea. - + rewrite (OnOne2_length X) in H1. - destruct H1 as [d [? [? []]]]; eapply expanded_tFix; tea. - destruct N2; cbn in *; eauto. lia. intros eq; discriminate. - - - move/expanded_mkApps_inv': exp. - rewrite head_mkApps head_nApp // arguments_mkApps // => [] [] hargs [d [hf [hm2 [hnth harg]]]]. - eapply OnOne2_nth_error in hnth as [t' [hnth hor]]; tea. - eapply expanded_tFix; tea. - { clear hor. solve_all. eapply OnOne2_impl_All_r; tea. - intros ? ? [] []. noconf e. rewrite -H2. split => //. solve_all. - clear -X H1. - enough (rev_map (fun d1 : def term => S (rarg d1)) mfix0 = (rev_map (fun d1 : def term => S (rarg d1)) mfix1)) as <- => //. - clear -X. rewrite !rev_map_spec. f_equal. induction X. destruct p as []. cbn in *. congruence. cbn. congruence. } - { destruct hor; subst => //. destruct p as [? e]; noconf e. now congruence. } - - - move/expanded_mkApps_inv': exp. - rewrite head_mkApps head_nApp // arguments_mkApps // => [] [] hargs [d [hf [hm2 [hnth harg]]]]. - eapply OnOne2_nth_error in hnth as [t' [hnth hor]]; tea. - eapply expanded_tFix; tea. - { clear hor. solve_all. - assert (rev_map (fun d1 : def term => S (rarg d1)) mfix0 = (rev_map (fun d1 : def term => S (rarg d1)) mfix1)). - { clear -X. rewrite !rev_map_spec. f_equal. induction X. destruct p as []. cbn in *. congruence. cbn. congruence. } - rewrite -H. - eapply OnOne2_impl_All_r; tea. - intros ? ? [] []. noconf e. destruct p. - eapply Red1Apps.isLambda_red1 in r; tea. split => //. - set(Γ'' := rev_map (fun d1 : def term => S (rarg d1)) mfix0). - eapply e => //. - { intros n b hnth'. - destruct (nth_error (fix_context mfix0) n) eqn:e'. - rewrite nth_error_app_lt in hnth'. now eapply nth_error_Some_length in e'. - rewrite e' in hnth'. noconf hnth'. - pose proof (PCUICParallelReductionConfluence.fix_context_assumption_context mfix0). - eapply PCUICParallelReductionConfluence.nth_error_assumption_context in H6; tea. congruence. - rewrite /Γ''. - eapply nth_error_None in e'. len in e'. - rewrite nth_error_app_ge in hnth' => //. now len. - rewrite skipn_app skipn_all2. len. - cbn. len. replace (S n - #|mfix0|) with (S (n - #|mfix0|)). 2:{ lia. } - eapply wfΓ. now len in hnth'. } } - { destruct hor; subst => //. destruct p as [? e]; noconf e. now congruence. } - - eapply expanded_tProd. - - constructor. - - constructor. - eapply expanded_tEvar_inv in exp. - solve_all; eapply OnOne2_impl_All_r; tea. intuition eauto. - now eapply X0. - - depelim exp; solve_discr. now cbn in H. - - depelim exp; solve_discr. now cbn in H. - - eapply expanded_tCoFix_inv in exp. econstructor. - rewrite -(OnOne2_length X). solve_all; eapply OnOne2_impl_All_r; tea; intuition eauto. - destruct X0. noconf e. now rewrite -H1. - - eapply expanded_tCoFix_inv in exp. econstructor. - rewrite -(OnOne2_length X). solve_all; eapply OnOne2_impl_All_r; tea; intuition eauto. - destruct X0. destruct p. noconf e. eapply e0 => //. - intros n b. - destruct nth_error eqn:e' => //. - cbn. intros [=]. destruct c as [? [] ?]. cbn in *. noconf H4. - eapply nth_error_app_inv in e' as [[]|[]]. - { pose proof (PCUICParallelReductionConfluence.fix_context_assumption_context mfix0). - eapply PCUICParallelReductionConfluence.nth_error_assumption_context in H5; tea. cbn in H5. congruence. } - len in H3. len in H4. - rewrite skipn_app skipn_all2; len. - replace (S n - #|mfix0|) with (S (n - #|mfix0|)) by lia. eapply wfΓ. rewrite H4 /= //. - noconf H4. Qed. \ No newline at end of file diff --git a/pcuic/theories/PCUICWellScopedCumulativity.v b/pcuic/theories/PCUICWellScopedCumulativity.v index 1488b87a5..a2fd67f45 100644 --- a/pcuic/theories/PCUICWellScopedCumulativity.v +++ b/pcuic/theories/PCUICWellScopedCumulativity.v @@ -13,6 +13,8 @@ Require Import Equations.Prop.DepElim. Require Import Equations.Type.Relation Equations.Type.Relation_Properties. From Equations Require Import Equations. +Local Ltac intuition_solver ::= auto with *. + (* We show that conversion/cumulativity starting from well-typed terms is transitive. We first use typing to decorate the reductions/comparisons with invariants showing that all the considered contexts/terms are well-scoped. In a second step diff --git a/pcuic/theories/Typing/PCUICContextConversionTyp.v b/pcuic/theories/Typing/PCUICContextConversionTyp.v index e91f60548..9f289f114 100644 --- a/pcuic/theories/Typing/PCUICContextConversionTyp.v +++ b/pcuic/theories/Typing/PCUICContextConversionTyp.v @@ -16,6 +16,8 @@ Arguments red_ctx : clear implicits. Implicit Types (cf : checker_flags) (Σ : global_env_ext). +Local Ltac intuition_solver ::= auto with *. + Lemma lift0_open {cf:checker_flags} {Γ : closed_context} {Γ'' : open_context Γ} {M : open_term Γ} {n} : n = #|Γ''| -> is_open_term (Γ ,,, Γ'') (lift0 n M). diff --git a/pcuic/theories/utils/PCUICAstUtils.v b/pcuic/theories/utils/PCUICAstUtils.v index 1ad4cec57..c83af52f8 100644 --- a/pcuic/theories/utils/PCUICAstUtils.v +++ b/pcuic/theories/utils/PCUICAstUtils.v @@ -595,6 +595,30 @@ Qed. Lemma head_tapp t1 t2 : head (tApp t1 t2) = head t1. Proof. rewrite /head /decompose_app /= fst_decompose_app_rec //. Qed. +Lemma head_mkApps t args : head (mkApps t args) = head t. +Proof. + induction args using rev_ind; simpl; auto. + now rewrite mkApps_app /= head_tapp. +Qed. + +Lemma arguments_mkApps_nApp f args : + ~~ isApp f -> + arguments (mkApps f args) = args. +Proof. + rewrite /arguments => isa. + rewrite decompose_app_mkApps //. +Qed. + +Lemma arguments_mkApps f args : + arguments (mkApps f args) = arguments f ++ args. +Proof. + rewrite /arguments. + rewrite /decompose_app decompose_app_rec_mkApps //. + rewrite app_nil_r. + induction f in args |- * => //=. + rewrite IHf1. now rewrite (IHf1 [f2]) -app_assoc. +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 diff --git a/template-coq/theories/TypingWf.v b/template-coq/theories/TypingWf.v index 9ef53b3c8..125210814 100644 --- a/template-coq/theories/TypingWf.v +++ b/template-coq/theories/TypingWf.v @@ -1117,7 +1117,6 @@ Section TypingWf. (decompose_app (mkApp f u)).2 <> []. Proof using Type. induction f; simpl; auto; try congruence. - destruct args; simpl; congruence. Qed. Lemma mkApps_tApp' f u f' u' : diff --git a/template-coq/theories/WcbvEval.v b/template-coq/theories/WcbvEval.v index 90c22bdef..a41410626 100644 --- a/template-coq/theories/WcbvEval.v +++ b/template-coq/theories/WcbvEval.v @@ -645,10 +645,10 @@ Section Wcbv. cbn in vf'; destruct vf'. - eapply eval_construct; tea. apply app_tip_nil. now eapply All2_app. - rewrite -mkApps_tApp => //. rewrite is_empty_app /= andb_false_r //. - eapply eval_app_cong; auto. eapply app_tip_nil. + eapply eval_app_cong; auto. eapply All2_app; auto. - rewrite -mkApps_tApp => //. rewrite is_empty_app /= andb_false_r //. - eapply eval_app_cong; auto. eapply app_tip_nil. + eapply eval_app_cong; auto. eapply All2_app; auto. - eapply eval_fix_value with (fixargsv := []) (argsv := x0 ++ [y]). 4:tea. all: move=> //. diff --git a/template-coq/theories/WfAst.v b/template-coq/theories/WfAst.v index 0fb54d7b7..a723d762e 100644 --- a/template-coq/theories/WfAst.v +++ b/template-coq/theories/WfAst.v @@ -198,7 +198,7 @@ Lemma wf_mkApp Σ u a : wf Σ u -> wf Σ a -> wf Σ (mkApp u a). Proof. intros H H'. inversion_clear H; try constructor; simpl; auto; try congruence; try constructor; auto. - intro. destruct u0; simpl in *; congruence. solve_all. + solve_all. apply All_app_inv; auto. all:econstructor; eauto. Qed. diff --git a/utils/theories/MCList.v b/utils/theories/MCList.v index b3900805b..d08b8d870 100644 --- a/utils/theories/MCList.v +++ b/utils/theories/MCList.v @@ -1383,6 +1383,7 @@ Lemma app_tip_nil {A} (l : list A) (x : A) : (l ++ [x])%list <> []. Proof. destruct l; cbn; congruence. Qed. +Global Hint Resolve app_tip_nil : core. Definition remove_last {A} (args : list A) := List.firstn (#|args| - 1) args. From b072acf2ec507b1fb3cbce17607d7b973e92afe1 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 13 Sep 2023 10:15:21 +0200 Subject: [PATCH 52/61] Filled all proofs in ErasureCorrectness --- erasure-plugin/theories/ErasureCorrectness.v | 270 ++++++++++--------- 1 file changed, 150 insertions(+), 120 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index dbbaacecf..551ded36b 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -106,91 +106,6 @@ Ltac destruct_compose_no_clear := rewrite H; revert pre'' H end. -(* -Section TransformValue. - Context {program program' : Type}. - Context {value value' : Type}. - Context {eval : program -> value -> Prop}. - Context {eval' : program' -> value' -> Prop}. - Context (t : Transform.t program program' value value' eval eval'). - - Lemma preserves_value p : value p.1 p.2 (transform t p) - - Definition preserves_eval pre (transform : forall p : program, pre p -> program') - (obseq : forall p : program, pre p -> program' -> value -> value' -> Prop) := - forall p v (pr : pre p), - eval p v -> - let p' := transform p pr in - exists v', eval' p' v' /\ obseq p pr p' v v'. - - Record t := - { name : string; - -Lemma transform_value *) - - -Inductive is_construct_app : EAst.term -> Prop := -| is_cstr_app_cstr kn c args : Forall is_construct_app args -> is_construct_app (EAst.tConstruct kn c args) -| is_cstr_app_app f a : is_construct_app f -> is_construct_app a -> is_construct_app (EAst.tApp f a). - -Section lambdabox_theorem. - - Context (Σ Σ' : EEnvMap.GlobalContextMap.t) (v : EAst.term). - - Context (p : pre verified_lambdabox_pipeline (Σ, v)). - Context (p' : pre verified_lambdabox_pipeline (Σ', v)). - Context (is_value : value (wfl := default_wcbv_flags) Σ v). - - Lemma pres : extends_eprogram_env (Σ, v) (Σ', v) -> - extends_eprogram (transform verified_lambdabox_pipeline (Σ, v) p) - (transform verified_lambdabox_pipeline (Σ', v) p'). - Proof. - epose proof (pres := verified_lambdabox_pipeline_extends). - red in pres. specialize (pres _ _ p p'). auto. - Qed. - - (* Final evaluation flags *) - Definition evflags := {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}. - - Lemma pres_firstorder_value : - is_construct_app v -> - is_construct_app (transform verified_lambdabox_pipeline (Σ, v) p).2. - Proof. - intros isapp. - destruct (preservation verified_lambdabox_pipeline (Σ, v) v p) as [v' [[ev] obs]]. - { red. cbn. sq. eapply value_final, is_value. } - set (transp := transform _ _ p) in *. - assert (value (wfl := evflags) transp.1 transp.2). admit. - eapply eval_value in ev => //. subst v'. - clear -obs isapp. - unfold verified_lambdabox_pipeline in obs. - cbn [obseq compose] in obs. - unfold run, time in obs. - decompose [ex and prod] obs. clear obs. subst. - cbn [obseq compose verified_lambdabox_pipeline] in *. - cbn [obseq compose constructors_as_blocks_transformation] in *. - cbn [obseq run compose rebuild_wf_env_transform] in *. - cbn [obseq compose inline_projections_optimization] in *. - cbn [obseq compose remove_match_on_box_trans] in *. - cbn [obseq compose remove_params_optimization] in *. - cbn [obseq compose guarded_to_unguarded_fix] in *. - cbn [obseq compose verified_lambdabox_pipeline] in *. - subst. - cbn [transform rebuild_wf_env_transform] in *. - cbn [transform constructors_as_blocks_transformation] in *. - cbn [transform inline_projections_optimization] in *. - cbn [transform remove_match_on_box_trans] in *. - cbn [transform remove_params_optimization] in *. - cbn [transform guarded_to_unguarded_fix] in *. - clearbody transp. revert b. intros ->. clear transp. - induction isapp. - cbn in *. constructor. constructor. - Admitted. - - -End lambdabox_theorem. - - Lemma rebuild_wf_env_irr {efl : EWellformed.EEnvFlags} p wf p' wf' : p.1 = p'.1 -> (rebuild_wf_env p wf).1 = (rebuild_wf_env p' wf').1. @@ -880,7 +795,7 @@ Qed. Arguments PCUICFirstorder.firstorder_ind _ _ : clear implicits. Section PCUICExpandLets. - Import PCUICAst. + Import PCUICAst PCUICAstUtils PCUICFirstorder. Import PCUICEnvironment. Import PCUICExpandLets PCUICExpandLetsCorrectness. @@ -896,10 +811,10 @@ Section PCUICExpandLets. Lemma expand_lets_preserves_fo (Σ : global_env_ext) Γ v : wf Σ -> - PCUICFirstorder.firstorder_value Σ Γ v -> - PCUICFirstorder.firstorder_value (trans_global Σ) (trans_local Γ) (trans v). + firstorder_value Σ Γ v -> + firstorder_value (trans_global Σ) (trans_local Γ) (trans v). Proof. - intros wfΣ; move: v; apply: PCUICFirstorder.firstorder_value_inds. + intros wfΣ; move: v; apply: firstorder_value_inds. intros i n ui u args pandi hty hargs ihargs isp. rewrite trans_mkApps. econstructor. apply expand_lets_sound in hty. @@ -914,6 +829,122 @@ Section PCUICExpandLets. rewrite (trans_destArity []). destruct destArity => //. destruct p => //. now intros ->. Qed. + + Lemma forallb_mapi {A B} (f f' : nat -> A -> B) (g g' : B -> bool) l : + (forall n x, nth_error l n = Some x -> g (f n x) = g' (f' n x)) -> + forallb g (mapi f l) = forallb g' (mapi f' l). + Proof. + unfold mapi. + intros. + assert + (forall (n : nat) (x : A), nth_error l n = Some x -> g (f (0 + n) x) = g' (f' (0 + n) x)). + { intros n x. now apply H. } + clear H. + revert H0. + generalize 0. + induction l; cbn; auto. + intros n hfg. f_equal. specialize (hfg 0). rewrite Nat.add_0_r in hfg. now apply hfg. + eapply IHl. intros. replace (S n + n0) with (n + S n0) by lia. eapply hfg. + now cbn. + Qed. + + Lemma forallb_mapi_impl {A B} (f f' : nat -> A -> B) (g g' : B -> bool) l : + (forall n x, nth_error l n = Some x -> g (f n x) -> g' (f' n x)) -> + forallb g (mapi f l) -> forallb g' (mapi f' l). + Proof. + unfold mapi. + intros hfg. + assert + (forall (n : nat) (x : A), nth_error l n = Some x -> g (f (0 + n) x) -> g' (f' (0 + n) x)). + { intros n x ?. now apply hfg. } + clear hfg. + revert H. + generalize 0. + induction l; cbn; auto. + intros n hfg. move/andP => [] hg hf. + pose proof (hfg 0). rewrite Nat.add_0_r in H. apply H in hg => //. rewrite hg /=. + eapply IHl. intros. replace (S n + n0) with (n + S n0) by lia. eapply hfg. now cbn. + now rewrite Nat.add_succ_r. assumption. + Qed. + + Lemma expand_firstorder_type Σ n k t : + (forall nm, plookup_env (firstorder_env' Σ) nm = Some true -> + plookup_env (firstorder_env' (trans_global_decls Σ)) nm = Some true) -> + @PCUICFirstorder.firstorder_type (firstorder_env' Σ) n k t -> + @PCUICFirstorder.firstorder_type (firstorder_env' (trans_global_decls Σ)) n k (trans t). + Proof. + rewrite /PCUICFirstorder.firstorder_type /PCUICFirstorder.firstorder_env. + pose proof (trans_decompose_app t). + destruct decompose_app. rewrite H. cbn. + destruct t0 => //. intros hd => /=. + destruct ind. + destruct plookup_env eqn:hp => //. + destruct b => //. + eapply hd in hp. rewrite hp //. + Qed. + + Lemma trans_firstorder_mutind Σ m : + (forall nm, plookup_env (firstorder_env' Σ) nm = Some true -> + plookup_env (firstorder_env' (trans_global_decls Σ)) nm = Some true) -> + @firstorder_mutind (firstorder_env' Σ) m -> + @firstorder_mutind (firstorder_env' (trans_global_decls Σ)) (trans_minductive_body m). + Proof. + intros he. + rewrite /firstorder_mutind. f_equal. + rewrite /trans_minductive_body /=. + rewrite -{1}[ind_bodies m]map_id. rewrite -mapi_cst_map. + move/andP => [] -> /=. + eapply forallb_mapi_impl. + intros n x hnth. + rewrite /firstorder_oneind. destruct x; cbn. + move/andP => [] hc ->; rewrite andb_true_r. + rewrite forallb_map. solve_all. + move: H; rewrite /firstorder_con /=. + rewrite !rev_app_distr !alli_app. + move/andP => [hp ha]. apply/andP; split. + - len. rewrite /trans_local -map_rev alli_map. + eapply alli_impl; tea. intros i []. cbn. + now eapply expand_firstorder_type. + - len. len in ha. clear -ha he. + move: ha; generalize (cstr_args x) as args. + induction args => ha; cbn => //. + destruct a as [na [] ?] => /=. eapply IHargs. cbn in ha. + rewrite alli_app in ha. move/andP: ha => [] //. + cbn. rewrite PCUICSigmaCalculus.smash_context_acc /=. + cbn in ha. rewrite alli_app in ha. move/andP: ha => [] // hc hd. + rewrite alli_app. apply/andP. split. apply IHargs => //. + len. cbn. rewrite andb_true_r. cbn in hd. rewrite andb_true_r in hd. len in hd. + clear -he hd. move: hd. rewrite /firstorder_type. + destruct (decompose_app decl_type) eqn:e. cbn. + destruct t0 => //; apply decompose_app_inv in e; rewrite e. + { rewrite trans_mkApps PCUICLiftSubst.lift_mkApps PCUICLiftSubst.subst_mkApps + decompose_app_mkApps //=. todo "foo". + destruct nth_error eqn:hnth. + move/andP => [] /Nat.leb_le hn /Nat.ltb_lt hn'. + move/nth_error_Some_length: hnth. len. destruct (Nat.leb_spec #|args| n). lia. lia. + len. + move/andP => [] /Nat.leb_le hn /Nat.ltb_lt hn'. + destruct (Nat.leb_spec #|args| n). apply/andP. split. + eapply Nat.leb_le. lia. apply Nat.ltb_lt. lia. + apply/andP; split. apply Nat.leb_le. lia. apply Nat.ltb_lt. lia. } + { rewrite trans_mkApps PCUICLiftSubst.lift_mkApps PCUICLiftSubst.subst_mkApps + decompose_app_mkApps //=. + destruct ind. destruct plookup_env eqn:hp => //. destruct b => //. apply he in hp. rewrite hp //. } + Qed. + + Lemma trans_firstorder_env Σ : + (forall nm, plookup_env (firstorder_env' Σ) nm = Some true -> + plookup_env (firstorder_env' (trans_global_decls Σ)) nm = Some true). + Proof. + induction Σ; cbn => //. + destruct a. destruct g => //=. + intros nm. + destruct eqb => //. now apply IHΣ. + intros nm. destruct eqb => //. + intros [=]. f_equal. + now eapply trans_firstorder_mutind. apply IHΣ. + Qed. + End PCUICExpandLets. @@ -1015,26 +1046,13 @@ Section pipeline_theorem. move: fo. rewrite /PCUICFirstorder.firstorder_ind /= PCUICExpandLetsCorrectness.trans_lookup /=. destruct PCUICAst.PCUICEnvironment.lookup_env => //. destruct g => //=. - unfold PCUICFirstorder.firstorder_mutind. cbn. - move/andP => [] -> /=. ELiftSubst.solve_all. - eapply forallb_Forall in b. apply forallb_Forall. - ELiftSubst.solve_all. eapply All_mapi. - eapply All_Alli; tea. cbn. - intros n []. rewrite /PCUICFirstorder.firstorder_oneind /=. - move/andP => [] hf ->; rewrite andb_true_r. - eapply forallb_Forall in hf. apply forallb_Forall. - ELiftSubst.solve_all. - rewrite /PCUICFirstorder.firstorder_con in H *. - eapply alli_Alli. eapply alli_Alli in H. - destruct x; cbn in *. - eapply Alli_rev. - ELiftSubst.solve_all. - - re - destruct m; cbn in *. - - forward H0. - { cbn. todo "preserves values". } + eapply trans_firstorder_mutind. eapply trans_firstorder_env. } + forward H0. { cbn. rewrite -eqtr. + unshelve eapply (PCUICExpandLetsCorrectness.trans_wcbveval (Σ := Σ)); tea. exact extraction_checker_flags. + apply HΣ. apply PCUICExpandLetsCorrectness.trans_wf, HΣ. + 2:{ eapply PCUICWcbvEval.value_final. now eapply PCUICWcbvEval.eval_to_value in X. } + eapply PCUICWcbvEval.eval_closed; tea. apply HΣ. + unshelve apply (PCUICClosedTyp.subject_closed typing). } specialize (H0 _ eq_refl). rewrite /p. rewrite erase_transform_fo //. @@ -1055,7 +1073,7 @@ Section pipeline_theorem. Import PCUICWfEnv. Lemma verified_erasure_pipeline_theorem : - ∥ eval (wfl := extraction_wcbv_flags) Σ_t t_t v_t∥. + ∥ eval (wfl := extraction_wcbv_flags) Σ_t t_t v_t ∥. Proof. hnf. pose proof (preservation verified_erasure_pipeline (Σ, t)) as Hcorr. @@ -1120,26 +1138,38 @@ Section pipeline_theorem. eapply (ErasureFunctionProperties.firstorder_erases_deterministic optimized_abstract_env_impl wfe Σ.2) in b0. 3:tea. 2:{ cbn. reflexivity. } 2:{ eapply PCUICExpandLetsCorrectness.trans_wcbveval. eapply PCUICWcbvEval.eval_closed; tea. apply HΣ. - admit. + clear -typing HΣ; now eapply PCUICClosedTyp.subject_closed in typing. eapply PCUICWcbvEval.value_final. now eapply PCUICWcbvEval.eval_to_value in X0. } - 2:{ clear -fo. admit. } + 2:{ clear -fo. revert fo. rewrite /PCUICFirstorder.firstorder_ind /=. + rewrite PCUICExpandLetsCorrectness.trans_lookup. + destruct PCUICAst.PCUICEnvironment.lookup_env => //. + destruct g => //=. + eapply trans_firstorder_mutind. eapply trans_firstorder_env. } 2:{ apply HΣ. } 2:{ apply PCUICExpandLetsCorrectness.trans_wf, HΣ. } - rewrite b0. - intros obs. - constructor. + rewrite b0. intros obs. constructor. match goal with [ H1 : eval _ _ ?v1 |- eval _ _ ?v2 ] => enough (v2 = v1) as -> by exact ev end. eapply obseq_lambdabox; revgoals. unfold erase_pcuic_program. cbn [fst snd]. exact obs. + Unshelve. all:eauto. + 2:{ eapply PCUICExpandLetsCorrectness.trans_wf, HΣ. } clear obs b0 ev e w. eapply extends_erase_pcuic_program. cbn. - eapply (PCUICExpandLetsCorrectness.trans_wcbveval (Σ := (Σ.1, Σ.2))). + unshelve eapply (PCUICExpandLetsCorrectness.trans_wcbveval (cf := extraction_checker_flags) (Σ := (Σ.1, Σ.2))). + { now eapply PCUICExpandLetsCorrectness.trans_wf. } { clear -HΣ typing. now eapply PCUICClosedTyp.subject_closed in typing. } cbn. 2:cbn. 3:cbn. exact X0. - - - - Admitted. + now eapply trans_axiom_free. + pose proof (PCUICExpandLetsCorrectness.expand_lets_sound typing). + rewrite PCUICExpandLetsCorrectness.trans_mkApps in X1. eapply X1. + cbn. + move: fo. clear. + { rewrite /PCUICFirstorder.firstorder_ind /=. + rewrite PCUICExpandLetsCorrectness.trans_lookup. + destruct PCUICAst.PCUICEnvironment.lookup_env => //. + destruct g => //=. + eapply trans_firstorder_mutind. eapply trans_firstorder_env. } + Qed. Lemma verified_erasure_pipeline_lambda : PCUICAst.isLambda t -> EAst.isLambda t_t. From 708302f5901363a86bc522ba9bbeed91a1aa9366 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 13 Sep 2023 11:29:35 +0200 Subject: [PATCH 53/61] Cleanup a bit, remain only normalization side-conditions --- erasure-plugin/theories/ErasureCorrectness.v | 144 +++++++++++++------ erasure/theories/ErasureFunctionProperties.v | 33 +++-- erasure/theories/Typed/Erasure.v | 2 - erasure/theories/Typed/Extraction.v | 2 +- 4 files changed, 120 insertions(+), 61 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 551ded36b..bcf1c9600 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -581,6 +581,41 @@ Proof. Qed. Derive Signature for firstorder_evalue. +Section PCUICenv. + Import PCUICAst. + Import PCUICEnvironment. + + Lemma pres_inductives_lookup Σ Σ' i n mdecl idecl cdecl : + wf Σ -> + erase_preserves_inductives Σ Σ' -> + declared_constructor Σ (i, n) mdecl idecl cdecl -> + forall npars nargs, EGlobalEnv.lookup_constructor_pars_args Σ' i n = Some (npars, nargs) -> + npars = mdecl.(ind_npars) /\ nargs = cdecl.(cstr_arity). + Proof. + intros wf he. + rewrite /declared_constructor /declared_inductive. + intros [[declm decli] declc]. + unshelve eapply declared_minductive_to_gen in declm. 3:exact wf. red in declm. + intros npars nargs. rewrite /EGlobalEnv.lookup_constructor_pars_args. + rewrite /EGlobalEnv.lookup_constructor /EGlobalEnv.lookup_inductive /EGlobalEnv.lookup_minductive. + destruct EGlobalEnv.lookup_env eqn:e => //=. + destruct g => //. + eapply he in declm; tea. subst m. + rewrite nth_error_map decli /=. + rewrite nth_error_map declc /=. intuition congruence. + Qed. +End PCUICenv. + +Lemma lookup_constructor_pars_args_lookup_inductive_pars Σ i n npars nargs : + EGlobalEnv.lookup_constructor_pars_args Σ i n = Some (npars, nargs) -> + EGlobalEnv.lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars. +Proof. + rewrite /EGlobalEnv.lookup_constructor_pars_args /EGlobalEnv.lookup_inductive_pars. + rewrite /EGlobalEnv.lookup_constructor /EGlobalEnv.lookup_inductive. + destruct EGlobalEnv.lookup_minductive => //=. + destruct nth_error => //. destruct nth_error => //. congruence. +Qed. + Lemma compile_evalue_erase (Σ : PCUICAst.PCUICEnvironment.global_env_ext) (Σ' : EEnvMap.GlobalContextMap.t) v : wf Σ.1 -> PCUICFirstorder.firstorder_value Σ [] v -> @@ -593,16 +628,9 @@ Proof. intros i n ui u args pandi hty hargs ih isp. eapply PCUICInductiveInversion.Construct_Ind_ind_eq' in hty as [mdecl [idecl [cdecl [declc _]]]] => //. rewrite compile_value_erase_mkApps. - intros fo'. depelim fo'. EAstUtils.solve_discr. noconf H1. + intros fo'. depelim fo'. EAstUtils.solve_discr. noconf H1. noconf H2. assert (npars = PCUICAst.PCUICEnvironment.ind_npars mdecl). - { destruct declc as [[declm decli] declc]. - unshelve eapply declared_minductive_to_gen in declm. 3:exact wf. red in declm. - rewrite /EGlobalEnv.lookup_inductive_pars /EGlobalEnv.lookup_minductive in H. - destruct (PCUICAst.PCUICEnvironment.lookup_env) eqn:hl => //. - noconf declm. - destruct (EGlobalEnv.lookup_env) eqn:hl' => //. destruct g => //. - red in hΣ. - eapply hΣ in hl'; tea. cbn in H. noconf H. subst m. reflexivity. } + { now eapply pres_inductives_lookup in declc; tea. } subst npars. rewrite (compile_value_box_mkApps (npars := PCUICAst.PCUICEnvironment.ind_npars mdecl)). { destruct declc as [[declm decli] declc]. @@ -611,6 +639,7 @@ Proof. rewrite /pcuic_lookup_inductive_pars // declm //. } rewrite (compile_evalue_box_strip_mkApps (npars := PCUICAst.PCUICEnvironment.ind_npars mdecl)) //. rewrite lookup_inductive_pars_spec //. + eapply lookup_constructor_pars_args_lookup_inductive_pars; tea. rewrite !app_nil_r. f_equal. rewrite app_nil_r skipn_map in H0. eapply Forall_map_inv in H0. @@ -619,6 +648,19 @@ Proof. ELiftSubst.solve_all. Qed. + +Lemma lookup_constructor_pars_args_nopars {efl : EEnvFlags} Σ ind c npars nargs : + wf_glob Σ -> + has_cstr_params = false -> + EGlobalEnv.lookup_constructor_pars_args Σ ind c = Some (npars, nargs) -> npars = 0. +Proof. + intros wf h. + rewrite /EGlobalEnv.lookup_constructor_pars_args. + destruct EGlobalEnv.lookup_constructor eqn:e => //=. + destruct p as [[m i] cb]. intros [= <- <-]. + now eapply wellformed_lookup_constructor_pars in e. +Qed. + Lemma compile_evalue_box_firstorder {efl : EEnvFlags} {Σ : EEnvMap.GlobalContextMap.t} v : has_cstr_params = false -> EWellformed.wf_glob Σ -> @@ -628,10 +670,8 @@ Proof. move: v; apply: firstorder_evalue_elim. intros. rewrite /flip (compile_evalue_box_mkApps) // ?app_nil_r. - rewrite /EGlobalEnv.lookup_inductive_pars /= in H. - destruct EGlobalEnv.lookup_minductive eqn:e => //. - eapply wellformed_lookup_inductive_pars in hpars; tea => //. - noconf H. rewrite hpars in H1. rewrite skipn_0 in H1. + eapply lookup_constructor_pars_args_nopars in H; tea. subst npars. + rewrite skipn_0 in H1. constructor. ELiftSubst.solve_all. Qed. @@ -667,11 +707,13 @@ Proof. split => //. clear wf; move: t1 fo. unfold fo_evalue, fo_evalue_map. cbn. apply: firstorder_evalue_elim; intros. econstructor. - rewrite EInlineProjections.lookup_inductive_pars_optimize in H => //; tea. auto. + move: H. rewrite /EGlobalEnv.lookup_constructor_pars_args. + rewrite EInlineProjections.lookup_constructor_optimize //. intros h; exact h. auto. auto. - rewrite /fo_evalue_map. intros [] pr fo. cbn in *. unfold EInlineProjections.optimize_program. cbn. f_equal. destruct pr as [[pr _] _]. cbn in *. move: t1 fo. apply: firstorder_evalue_elim; intros. - eapply wf_glob_lookup_inductive_pars in H => //. subst npars; rewrite skipn_0 in H0 H1. + eapply lookup_constructor_pars_args_nopars in H; tea => //. subst npars. + rewrite skipn_0 in H0 H1. rewrite EInlineProjections.optimize_mkApps /=. f_equal. ELiftSubst.solve_all. Qed. @@ -688,17 +730,26 @@ Proof. split => //. clear wf; move: t1 fo. apply: firstorder_evalue_elim; intros. econstructor; tea. - rewrite EOptimizePropDiscr.lookup_inductive_pars_optimize in H0 => //; tea. + move: H0. + rewrite /EGlobalEnv.lookup_constructor_pars_args EOptimizePropDiscr.lookup_constructor_remove_match_on_box //. - intros [] pr fo. cbn in *. unfold EOptimizePropDiscr.remove_match_on_box_program; cbn. f_equal. destruct pr as [[pr _] _]; cbn in *; move: t1 fo. apply: firstorder_evalue_elim; intros. - eapply wf_glob_lookup_inductive_pars in H0 => //. subst npars; rewrite skipn_0 in H2. - rewrite EOptimizePropDiscr.remove_match_on_box_mkApps /=. f_equal. + eapply lookup_constructor_pars_args_nopars in H0; tea => //. subst npars. + rewrite skipn_0 in H2. rewrite EOptimizePropDiscr.remove_match_on_box_mkApps /=. f_equal. ELiftSubst.solve_all. Qed. +Lemma lookup_constructor_pars_args_strip (Σ : t) i n npars nargs : + EGlobalEnv.lookup_constructor_pars_args Σ i n = Some (npars, nargs) -> + EGlobalEnv.lookup_constructor_pars_args (ERemoveParams.strip_env Σ) i n = Some (0, nargs). +Proof. + rewrite /EGlobalEnv.lookup_constructor_pars_args. rewrite ERemoveParams.lookup_constructor_strip //=. + destruct EGlobalEnv.lookup_constructor => //. destruct p as [[] ?] => //=. now intros [= <- <-]. +Qed. + #[global] Instance remove_params_optimization_pres {fl : WcbvFlags} {wcon : with_constructor_as_block = false} : ETransformPresFO.t (remove_params_optimization (wcon := wcon)) @@ -708,12 +759,12 @@ Proof. split => //. cbn [transform remove_params_optimization] in *. destruct pr as [[pr _] _]; cbn -[ERemoveParams.strip] in *; move: t1 fo. apply: firstorder_evalue_elim; intros. - rewrite ERemoveParams.strip_mkApps //. cbn -[EGlobalEnv.lookup_inductive_pars]. rewrite H. - econstructor. cbn -[EGlobalEnv.lookup_inductive_pars]. - now eapply ERemoveParams.lookup_inductive_pars_strip in H; tea. - rewrite skipn_0 /=. - rewrite skipn_map. - ELiftSubst.solve_all. + rewrite ERemoveParams.strip_mkApps //. cbn -[EGlobalEnv.lookup_inductive_pars]. + rewrite (lookup_constructor_pars_args_lookup_inductive_pars _ _ _ _ _ H). + eapply lookup_constructor_pars_args_strip in H. + econstructor; tea. rewrite skipn_0 /= skipn_map. + ELiftSubst.solve_all. len. + rewrite skipn_map. len. rewrite skipn_length. lia. Qed. #[global] Instance constructors_as_blocks_transformation_pres {efl : EWellformed.EEnvFlags} @@ -726,16 +777,18 @@ Proof. split. - intros v pr fo; eapply compile_evalue_box_firstorder; tea. apply pr. - move=> [Σ v] /= pr fo. rewrite /flip. - clear pr. move: v fo. + destruct pr as [[wf _] _]. cbn in wf. + move: v fo. apply: firstorder_evalue_elim; intros. rewrite /transform_blocks_program /=. f_equal. rewrite EConstructorsAsBlocks.transform_blocks_decompose. rewrite EAstUtils.decompose_app_mkApps // /=. rewrite compile_evalue_box_mkApps // ?app_nil_r. - (* rewrite lookup_inductive_pars_spec //. *) - admit. -Admitted. - + rewrite H. + eapply lookup_constructor_pars_args_nopars in H => //. subst npars. + rewrite chop_all. len. cbn. f_equal. rewrite skipn_0 in H1 H0. + ELiftSubst.solve_all. unfold transform_blocks_program in a. now noconf a. +Qed. #[global] Instance guarded_to_unguarded_fix_pres {efl : EWellformed.EEnvFlags} {has_guard : with_guarded_fix} {has_cstrblocks : with_constructor_as_block = false} : @@ -783,9 +836,11 @@ Lemma compile_evalue_strip (Σer : EEnvMap.GlobalContextMap.t) p : Proof. move: p. apply: firstorder_evalue_elim; intros. - rewrite ERemoveParams.strip_mkApps //=. rewrite H. + rewrite ERemoveParams.strip_mkApps //=. + rewrite (lookup_constructor_pars_args_lookup_inductive_pars _ _ _ _ _ H). rewrite compile_evalue_box_mkApps. rewrite (compile_evalue_box_strip_mkApps (npars:=npars)). + pose proof (lookup_constructor_pars_args_lookup_inductive_pars _ _ _ _ _ H). rewrite lookup_inductive_pars_spec //. rewrite !app_nil_r !skipn_map. f_equal. rewrite map_map. @@ -917,16 +972,18 @@ Section PCUICExpandLets. clear -he hd. move: hd. rewrite /firstorder_type. destruct (decompose_app decl_type) eqn:e. cbn. destruct t0 => //; apply decompose_app_inv in e; rewrite e. - { rewrite trans_mkApps PCUICLiftSubst.lift_mkApps PCUICLiftSubst.subst_mkApps - decompose_app_mkApps //=. todo "foo". + { move/andP => [] /Nat.leb_le hn /Nat.ltb_lt hn'. + rewrite trans_mkApps PCUICLiftSubst.lift_mkApps PCUICLiftSubst.subst_mkApps + decompose_app_mkApps //=. + { destruct nth_error eqn:hnth. + pose proof (nth_error_Some_length hnth). + move: hnth H. + destruct (Nat.leb_spec #|args| n). len. lia. len. lia. now cbn. } destruct nth_error eqn:hnth. - move/andP => [] /Nat.leb_le hn /Nat.ltb_lt hn'. - move/nth_error_Some_length: hnth. len. destruct (Nat.leb_spec #|args| n). lia. lia. - len. - move/andP => [] /Nat.leb_le hn /Nat.ltb_lt hn'. - destruct (Nat.leb_spec #|args| n). apply/andP. split. - eapply Nat.leb_le. lia. apply Nat.ltb_lt. lia. - apply/andP; split. apply Nat.leb_le. lia. apply Nat.ltb_lt. lia. } + move/nth_error_Some_length: hnth. len. destruct (Nat.leb_spec #|args| n). lia. lia. len. + destruct (Nat.leb_spec #|args| n). apply/andP. split. + eapply Nat.leb_le. lia. apply Nat.ltb_lt. lia. + apply/andP; split. apply Nat.leb_le. lia. apply Nat.ltb_lt. lia. } { rewrite trans_mkApps PCUICLiftSubst.lift_mkApps PCUICLiftSubst.subst_mkApps decompose_app_mkApps //=. destruct ind. destruct plookup_env eqn:hp => //. destruct b => //. apply he in hp. rewrite hp //. } @@ -1080,10 +1137,6 @@ Section pipeline_theorem. unshelve eapply Hcorr in Heval as Hev. eapply precond. destruct Hev as [v' [[H1] H2]]. move: H2. - - (* repeat match goal with - [ H : obseq _ _ _ _ _ |- _ ] => hnf in H ; decompose [ex and prod] H ; subst - end. *) rewrite v_t_spec. subst v_t Σ_t t_t. revert H1. @@ -1171,10 +1224,11 @@ Section pipeline_theorem. eapply trans_firstorder_mutind. eapply trans_firstorder_env. } Qed. - Lemma verified_erasure_pipeline_lambda : + (*Lemma verified_erasure_pipeline_lambda : PCUICAst.isLambda t -> EAst.isLambda t_t. Proof. unfold t_t. clear. - Admitted. + Qed.*) End pipeline_theorem. + diff --git a/erasure/theories/ErasureFunctionProperties.v b/erasure/theories/ErasureFunctionProperties.v index b1c0a710e..ca453b065 100644 --- a/erasure/theories/ErasureFunctionProperties.v +++ b/erasure/theories/ErasureFunctionProperties.v @@ -1989,9 +1989,10 @@ From MetaCoq.PCUIC Require Import PCUICFirstorder. From Equations Require Import Equations. Inductive firstorder_evalue Σ : EAst.term -> Prop := - | is_fo i n args npars : - EGlobalEnv.lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars -> + | is_fo i n args npars nargs : + EGlobalEnv.lookup_constructor_pars_args Σ i n = Some (npars, nargs) -> Forall (firstorder_evalue Σ) (skipn npars args) -> + #|args| = npars + nargs -> firstorder_evalue Σ (EAst.mkApps (EAst.tConstruct i n []) args). Lemma list_size_skipn {A} (l : list A) size n : @@ -2006,14 +2007,15 @@ Qed. Section elim. Context (Σ : E.global_context). Context (P : EAst.term -> Prop). - Context (ih : (forall i n args npars, - EGlobalEnv.lookup_inductive_pars Σ (Kernames.inductive_mind i) = Some npars -> + Context (ih : (forall i n args npars nargs, + EGlobalEnv.lookup_constructor_pars_args Σ i n = Some (npars, nargs) -> Forall (firstorder_evalue Σ) (skipn npars args) -> Forall P (skipn npars args) -> + #|args| = npars + nargs -> P (EAst.mkApps (EAst.tConstruct i n []) args))). Equations? firstorder_evalue_elim (t : EAst.term) (fo : firstorder_evalue Σ t) : P t by wf t (MR lt EInduction.size) := - { | _, is_fo i n args npars hl hf => _ }. + { | _, is_fo i n args npars nargs hl hf hargs => _ }. Proof. eapply ih; tea. eapply In_Forall. intros x hin. @@ -2101,10 +2103,10 @@ Lemma firstorder_evalue_extends Σ Σ' t : firstorder_evalue Σ' t. Proof. intros ext; move: t. - apply: firstorder_evalue_elim => i n args npars hl hf ih. + apply: firstorder_evalue_elim => i n args npars nargs hl hf hargs ih. econstructor; tea. red in ext. - move: hl; rewrite /lookup_inductive_pars /lookup_minductive. + move: hl; rewrite /lookup_constructor_pars_args /lookup_constructor /lookup_inductive /lookup_minductive. move: (ext (inductive_mind i)). now case: lookup_env => // a; move/(_ _ eq_refl) ->. Qed. @@ -2167,6 +2169,7 @@ Proof. { destruct (wfext Σg Hg). eapply PCUICFirstorder.firstorder_value_spec; tea. apply X0. constructor. eapply PCUICClassification.subject_reduction_eval; tea. eapply PCUICWcbvEval.eval_to_value; tea. } + 2:{ eapply subject_reduction; eauto. } { rewrite H1 in e. eapply erase_global_deps_eval in e. eapply firstorder_evalue_extends; tea. clear e. clear -Hrel H Hg Hwf. move: v H wt'. @@ -2186,16 +2189,18 @@ Proof. move/is_erasableP: i0. intros hc. destruct (hc _ Hrel). eapply (isErasable_Propositional (args:=[])) in X0 => //. cbn in X0, isp => //. rewrite X0 in isp => //. - eapply PCUICInductiveInversion.Construct_Ind_ind_eq' in hty as [mdecl [idecl [cdecl [declc _]]]] => //=. + eapply PCUICInductiveInversion.Construct_Ind_ind_eq' in hty as [mdecl [idecl [cdecl [declc [hcstr _]]]]] => //=. + destruct hcstr as [_ hcstr]. set (v := E.mkApps _ _). - eapply (erase_global_declared_constructor _ _ _ _ _ _ _ (term_global_deps v) decls normalization_in prf) in declc; tea. + pose proof declc as declc'. + eapply (erase_global_declared_constructor _ _ _ _ _ _ _ (term_global_deps v) decls normalization_in prf _ Hg) in declc; tea. destruct declc as [[hm hi] hc]. 2:{ rewrite /v /= term_global_deps_mkApps. destruct i => /=. eapply KernameSet.union_spec; left. now eapply KernameSet.singleton_spec. } 2:{ eapply Hwf. } econstructor. - { rewrite /lookup_inductive_pars /lookup_minductive hm /= //. } + { rewrite /lookup_constructor_pars_args /lookup_constructor /lookup_inductive /lookup_minductive hm /= hi //= hc //. } eapply Forall_skipn. - clear -ih Hrel isp. + clear -ih Hwf Hrel isp hcstr declc'. { revert v. rewrite erase_terms_eq => v. eapply Forall_All in ih. assert (All @@ -2216,8 +2221,10 @@ Proof. erewrite (erase_irrel _ _).1; tea. } eapply All_Forall, All_map_All. { intros. exact X0. } - cbn. intros x y rx h. apply h. exact Hrel. } } - { eapply subject_reduction; eauto. } + cbn. intros x y rx h. apply h. exact Hrel. } + rewrite erase_terms_eq map_All_length. cbn. + clear -wt' Hwf declc' Hrel hcstr. specialize (wt' _ Hrel). + now rewrite -(declared_constructor_arity declc'). } { eapply PCUICWcbvEval.value_final. eapply PCUICWcbvEval.eval_to_value. eauto. } all: eauto. Unshelve. eauto. diff --git a/erasure/theories/Typed/Erasure.v b/erasure/theories/Typed/Erasure.v index 4e1e9640e..8092b1566 100644 --- a/erasure/theories/Typed/Erasure.v +++ b/erasure/theories/Typed/Erasure.v @@ -572,8 +572,6 @@ Record type_flag {Σ Γ T} := Global Arguments type_flag : clear implicits. -Import PCUICSN. - Hint Resolve abstract_env_wf : erase. Definition isTT Γ T := diff --git a/erasure/theories/Typed/Extraction.v b/erasure/theories/Typed/Extraction.v index 2e7c75a9d..9ff35f659 100644 --- a/erasure/theories/Typed/Extraction.v +++ b/erasure/theories/Typed/Extraction.v @@ -185,7 +185,7 @@ Definition get_projections (env : ExAst.global_env) : list (ident * ExAst.one_in | [oib] => match oib.(ExAst.ind_ctors), oib.(ExAst.ind_projs) with (* case 1-ind with primitive projections *) - | [ctor],_::_ => map (fun '(na, _) => (na, oib)) oib.(ExAst.ind_projs) + | [_],_::_ => map (fun '(na, _) => (na, oib)) oib.(ExAst.ind_projs) (* case 1-ind without primitive projections *) | [(_,ctor_args,_)],[] => (* let is_named '(nm,_) := match nm with nNamed _ => true | _ => false end in *) From a0dc628f1b4566c2b36dda823ee4f73ab164dff5 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 13 Sep 2023 11:31:59 +0200 Subject: [PATCH 54/61] Fix erasure plugin compilation --- erasure-plugin/_PluginProject.in | 2 ++ erasure-plugin/src/metacoq_erasure_plugin.mlpack | 1 + 2 files changed, 3 insertions(+) diff --git a/erasure-plugin/_PluginProject.in b/erasure-plugin/_PluginProject.in index 1b344fb8a..dc5b48e22 100644 --- a/erasure-plugin/_PluginProject.in +++ b/erasure-plugin/_PluginProject.in @@ -79,6 +79,8 @@ src/eRemoveParams.ml src/erasureFunction.mli src/erasureFunction.ml +src/erasureFunctionProperties.mli +src/erasureFunctionProperties.ml src/ePretty.mli src/ePretty.ml src/eOptimizePropDiscr.mli diff --git a/erasure-plugin/src/metacoq_erasure_plugin.mlpack b/erasure-plugin/src/metacoq_erasure_plugin.mlpack index 662a2d49a..29a6ec0d0 100644 --- a/erasure-plugin/src/metacoq_erasure_plugin.mlpack +++ b/erasure-plugin/src/metacoq_erasure_plugin.mlpack @@ -54,6 +54,7 @@ Extract EEtaExpanded ERemoveParams ErasureFunction +ErasureFunctionProperties EOptimizePropDiscr EInlineProjections EConstructorsAsBlocks From 487cdf52b8d8edbc76a5fa3133bb141fae0dbedc Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 15 Sep 2023 14:12:38 +0200 Subject: [PATCH 55/61] Cleanup --- erasure-plugin/theories/ErasureCorrectness.v | 29 ++------------------ 1 file changed, 3 insertions(+), 26 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index bcf1c9600..5a2e5ed3e 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -9,8 +9,6 @@ From MetaCoq Require Import ETransform EConstructorsAsBlocks. From MetaCoq.Erasure Require Import EWcbvEvalNamed ErasureFunction ErasureFunctionProperties. From MetaCoq.ErasurePlugin Require Import Erasure. Import PCUICProgram. -(* Import TemplateProgram (template_eta_expand). - *) Import PCUICTransform (template_to_pcuic_transform, pcuic_expand_lets_transform). (* This is the total erasure function + @@ -300,16 +298,6 @@ Import EWcbvEval. Arguments erase_global_deps _ _ _ _ _ : clear implicits. Arguments erase_global_deps_fast _ _ _ _ _ _ : clear implicits. -(*Lemma erase_pcuic_program_spec {guard : abstract_guard_impl} - (p : pcuic_program) - (nin : (wf_ext p.1 -> PCUICSN.NormalizationIn p.1)) - (nin' : (wf_ext p.1 -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p.1)) - (wfext : ∥ wf_ext p.1 ∥) - (wt : ∥ ∑ T : PCUICAst.term, p.1;;; [] |- p.2 : T ∥) : - erase_pcuic_program p nin nin'wfext wt = - let et' := @erase optimized_abstract_env_impl - @erase_global_deps optimized_abstract_env_impl*) - Section PCUICProof. Import PCUICAst.PCUICEnvironment. @@ -442,15 +430,6 @@ Proof. eapply erase_transform_fo_gen; tea. reflexivity. Qed. - -(* Import PCUICAst. - -Lemma compile_fo_value (Σ : global_env_ext) Σ' t : - PCUICFirstorder.firstorder_value Σ [] t -> - erases_global - firstorder_evalue Σ (compile_value_erase t []). -Proof. Admitted. *) - Import MetaCoq.Common.Transform. From Coq Require Import Morphisms. @@ -1027,14 +1006,13 @@ Section pipeline_theorem. Variable fo : @PCUICFirstorder.firstorder_ind Σ (PCUICFirstorder.firstorder_env Σ) i. - Variable Normalisation : PCUICSN.NormalizationIn Σ. + Variable Normalisation : (forall Σ, wf_ext Σ -> PCUICSN.NormalizationIn Σ). Lemma precond : pre verified_erasure_pipeline (Σ, t). Proof. hnf. repeat eapply conj; sq; cbn; eauto. - red. cbn. eauto. - - todo "normalization". - - todo "normalization". + - intros. red. intros. now eapply Normalisation. Qed. Variable Heval : ∥PCUICWcbvEval.eval Σ t v∥. @@ -1048,8 +1026,7 @@ Section pipeline_theorem. - eapply (PCUICClassification.wcbveval_red (Σ := Σ)) in X; tea. eapply PCUICEtaExpand.expanded_red in X; tea. apply HΣ. intros ? ?; rewrite nth_error_nil => //. - - cbn. todo "normalization". - - todo "normalization". + - cbn. intros wf ? ? ? ? ? ?. now eapply Normalisation. Qed. Let Σ_t := (transform verified_erasure_pipeline (Σ, t) precond).1. From dfb03d1f3dd68613fc123c50683e25027b078594 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 2 Oct 2023 13:34:43 +0200 Subject: [PATCH 56/61] Minor fixes --- common/theories/Transform.v | 22 ++++++++-------- erasure-plugin/Makefile.plugin.local | 1 + erasure-plugin/theories/ETransform.v | 18 ++++++------- erasure-plugin/theories/Erasure.v | 16 ++++++------ erasure-plugin/theories/ErasureCorrectness.v | 27 ++++++++++---------- erasure/theories/EProgram.v | 20 ++++++++------- pcuic/Makefile.plugin.local | 2 ++ pcuic/theories/PCUICEtaExpand.v | 2 +- safechecker-plugin/Makefile.plugin.local | 1 + template-coq/Makefile.plugin.local | 6 +++-- template-coq/theories/TemplateProgram.v | 2 +- template-pcuic/theories/PCUICTransform.v | 2 +- 12 files changed, 64 insertions(+), 55 deletions(-) diff --git a/common/theories/Transform.v b/common/theories/Transform.v index d1678fb27..55785e023 100644 --- a/common/theories/Transform.v +++ b/common/theories/Transform.v @@ -23,13 +23,14 @@ Module Transform. Section Opt. Context {env env' : Type}. Context {term term' : Type}. + Context {value value' : Type}. Notation program' := (program env' term'). Notation program := (program env term). - Context {eval : program -> term -> Prop}. - Context {eval' : program' -> term' -> Prop}. + Context {eval : program -> value -> Prop}. + Context {eval' : program' -> value' -> Prop}. Definition preserves_eval pre (transform : forall p : program, pre p -> program') - (obseq : forall p : program, pre p -> program' -> term -> term' -> Prop) := + (obseq : forall p : program, pre p -> program' -> value -> value' -> Prop) := forall p v (pr : pre p), eval p v -> let p' := transform p pr in @@ -41,7 +42,7 @@ Module Transform. transform : forall p : program, pre p -> program'; post : program' -> Prop; correctness : forall input (p : pre input), post (transform input p); - obseq : forall p : program, pre p -> program' -> term -> term' -> Prop; + obseq : forall p : program, pre p -> program' -> value -> value' -> Prop; preservation : preserves_eval pre transform obseq }. Definition run (x : t) (p : program) (pr : pre x p) : program' := @@ -50,19 +51,20 @@ Module Transform. End Opt. Arguments t : clear implicits. - Definition self_transform env term eval eval' := t env env term term eval eval'. + Definition self_transform env term eval eval' := t env env term term term term eval eval'. Section Comp. Context {env env' env'' : Type}. Context {term term' term'' : Type}. - Context {eval : program env term -> term -> Prop}. - Context {eval' : program env' term' -> term' -> Prop}. - Context {eval'' : program env'' term'' -> term'' -> Prop}. + Context {value value' value'' : Type}. + Context {eval : program env term -> value -> Prop}. + Context {eval' : program env' term' -> value' -> Prop}. + Context {eval'' : program env'' term'' -> value'' -> Prop}. Local Obligation Tactic := idtac. - Program Definition compose (o : t env env' term term' eval eval') (o' : t env' env'' term' term'' eval' eval'') + Program Definition compose (o : t env env' term term' _ _ eval eval') (o' : t env' env'' term' term'' _ _ eval' eval'') (hpp : (forall p, o.(post) p -> o'.(pre) p)) - : t env env'' term term'' eval eval'' := + : t env env'' term term'' _ _ eval eval'' := {| name := (o.(name) ^ " -> " ^ o'.(name))%bs; transform p hp := run o' (run o p hp) (hpp _ (o.(correctness) _ hp)); diff --git a/erasure-plugin/Makefile.plugin.local b/erasure-plugin/Makefile.plugin.local index 2514f50e5..62cfbc2d2 100644 --- a/erasure-plugin/Makefile.plugin.local +++ b/erasure-plugin/Makefile.plugin.local @@ -1,3 +1,4 @@ +CAMLFLAGS :=-thread -bin-annot -strict-sequence -w -a+1..3-4+5..8-9+10..26-27+28..39-40-41-42+43-44-45+46..47-48+49..57-58+59..66-67-68+69-70 CAMLFLAGS+=-open Metacoq_template_plugin CAMLFLAGS+=-w -8 # Non-exhaustive matches due to translation of comparison to int CAMLFLAGS+=-w -20 # Unused arguments diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index f108396d6..fe2ab20b3 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -195,7 +195,7 @@ Qed. Local Obligation Tactic := try solve [ eauto ]. -Program Definition erase_transform {guard : abstract_guard_impl} : Transform.t _ _ PCUICAst.term EAst.term +Program Definition erase_transform {guard : abstract_guard_impl} : Transform.t _ _ PCUICAst.term EAst.term PCUICAst.term EAst.term eval_pcuic_program (eval_eprogram_env EWcbvEval.default_wcbv_flags) := {| name := "erasure"; pre p := @@ -327,7 +327,7 @@ End PCUICEnv. Import EWcbvEval (WcbvFlags, with_prop_case, with_guarded_fix). Program Definition guarded_to_unguarded_fix {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} {efl : EEnvFlags} (wguard : with_guarded_fix) : - Transform.t _ _ EAst.term EAst.term + Transform.t _ _ EAst.term EAst.term _ _ (eval_eprogram_env fl) (eval_eprogram_env (EWcbvEval.switch_unguarded_fix fl)) := {| name := "switching to unguarded fixpoints"; transform p pre := p; @@ -356,7 +356,7 @@ Definition rebuild_wf_env {efl} (p : eprogram) (hwf : wf_eprogram efl p): eprogr (GlobalContextMap.make p.1 (wf_glob_fresh p.1 (proj1 hwf)), p.2). Program Definition rebuild_wf_env_transform {fl : EWcbvEval.WcbvFlags} {efl} (with_exp : bool) : - Transform.t _ _ EAst.term EAst.term (eval_eprogram fl) (eval_eprogram_env fl) := + Transform.t _ _ EAst.term EAst.term _ _ (eval_eprogram fl) (eval_eprogram_env fl) := {| name := "rebuilding environment lookup table"; pre p := wf_eprogram efl p /\ (with_exp ==> EEtaExpanded.expanded_eprogram_cstrs p); transform p pre := rebuild_wf_env p (proj1 pre); @@ -378,7 +378,7 @@ Qed. Program Definition remove_params_optimization {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := all_env_flags): - Transform.t _ _ EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram fl) := + Transform.t _ _ EAst.term EAst.term _ _ (eval_eprogram_env fl) (eval_eprogram fl) := {| name := "stripping constructor parameters"; transform p pre := ERemoveParams.strip_program p; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; @@ -414,7 +414,7 @@ Qed. Program Definition remove_params_fast_optimization {fl : EWcbvEval.WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := all_env_flags) : - Transform.t _ _ EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram fl) := + Transform.t _ _ EAst.term EAst.term _ _ (eval_eprogram_env fl) (eval_eprogram fl) := {| name := "stripping constructor parameters (faster?)"; transform p _ := (ERemoveParams.Fast.strip_env p.1, ERemoveParams.Fast.strip p.1 [] p.2); pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; @@ -456,7 +456,7 @@ Qed. Import EOptimizePropDiscr EWcbvEval. Program Definition remove_match_on_box_trans {fl : WcbvFlags} {wcon : with_constructor_as_block = false} {efl : EEnvFlags} {hastrel : has_tRel} {hastbox : has_tBox} : - Transform.t _ _ EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram (disable_prop_cases fl)) := + Transform.t _ _ EAst.term EAst.term _ _ (eval_eprogram_env fl) (eval_eprogram (disable_prop_cases fl)) := {| name := "optimize_prop_discr"; transform p _ := remove_match_on_box_program p ; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; @@ -492,7 +492,7 @@ From MetaCoq.Erasure Require Import EInlineProjections. Program Definition inline_projections_optimization {fl : WcbvFlags} {wcon : EWcbvEval.with_constructor_as_block = false} (efl := switch_no_params all_env_flags) {hastrel : has_tRel} {hastbox : has_tBox} : - Transform.t _ _ EAst.term EAst.term (eval_eprogram_env fl) (eval_eprogram fl) := + Transform.t _ _ EAst.term EAst.term _ _ (eval_eprogram_env fl) (eval_eprogram fl) := {| name := "primitive projection inlining"; transform p _ := EInlineProjections.optimize_program p ; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; @@ -528,7 +528,7 @@ From MetaCoq.Erasure Require Import EConstructorsAsBlocks. Program Definition constructors_as_blocks_transformation {efl : EEnvFlags} {has_app : has_tApp} {has_rel : has_tRel} {hasbox : has_tBox} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = false} : - Transform.t _ _ EAst.term EAst.term (eval_eprogram_env target_wcbv_flags) (eval_eprogram block_wcbv_flags) := + Transform.t _ _ EAst.term EAst.term _ _ (eval_eprogram_env target_wcbv_flags) (eval_eprogram block_wcbv_flags) := {| name := "transforming to constuctors as blocks"; transform p _ := EConstructorsAsBlocks.transform_blocks_program p ; pre p := wf_eprogram_env efl p /\ EEtaExpanded.expanded_eprogram_env_cstrs p; @@ -568,7 +568,7 @@ From MetaCoq.Erasure Require Import EImplementBox. Program Definition implement_box_transformation {efl : EEnvFlags} {has_app : has_tApp} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = true} : - Transform.t _ _ EAst.term EAst.term (eval_eprogram block_wcbv_flags) (eval_eprogram block_wcbv_flags) := + Transform.t _ _ EAst.term EAst.term _ _ (eval_eprogram block_wcbv_flags) (eval_eprogram block_wcbv_flags) := {| name := "transforming to constuctors as blocks"; transform p _ := EImplementBox.implement_box_program p ; pre p := wf_eprogram efl p ; diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 175f9210f..f04c84f23 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -42,7 +42,7 @@ Axiom assume_preservation_template_program_env_expansion : (** We kludge the normalization assumptions by parameterizing over a continuation of "what will be done to the program later" as well as what properties we'll need of it *) -Program Definition eta_expand K : Transform.t _ _ Ast.term Ast.term +Program Definition eta_expand K : Transform.t _ _ Ast.term Ast.term _ _ eval_template_program_env eval_template_program := {| name := "eta expand cstrs and fixpoints"; pre := fun p => ∥ wt_template_program_env p ∥ /\ K (eta_expand_global_env p.1) ; @@ -59,7 +59,7 @@ Next Obligation. Qed. Program Definition verified_lambdabox_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - Transform.t _ _ EAst.term EAst.term + Transform.t _ _ EAst.term EAst.term _ _ (EProgram.eval_eprogram_env {| with_prop_case := true; with_guarded_fix := true; with_constructor_as_block := false |}) (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := (* Simulation of the guarded fixpoint rules with a single unguarded one: @@ -99,7 +99,7 @@ Qed. Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : Transform.t _ _ - PCUICAst.term EAst.term + PCUICAst.term EAst.term _ _ PCUICTransform.eval_pcuic_program (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := (* a bunch of nonsense for normalization preconditions *) @@ -117,10 +117,10 @@ Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl Import EGlobalEnv EWellformed. Definition transform_compose - {env env' env'' term term' term'' : Type} + {env env' env'' term term' term'' value value' value'' : Type} {eval eval' eval''} - (o : t env env' term term' eval eval') - (o' : t env' env'' term' term'' eval' eval'') + (o : t env env' term term' value value' eval eval') + (o' : t env' env'' term' term'' value' value'' eval' eval'') (pre : forall p, post o p -> pre o' p) : forall x p1, exists p3, transform (compose o o' pre) x p1 = transform o' (transform o x p1) p3. @@ -138,7 +138,7 @@ Qed. Program Definition pre_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : Transform.t _ _ - Ast.term PCUICAst.term + Ast.term PCUICAst.term _ _ TemplateProgram.eval_template_program PCUICTransform.eval_pcuic_program := (* a bunch of nonsense for normalization preconditions *) @@ -159,7 +159,7 @@ Program Definition pre_erasure_pipeline {guard : abstract_guard_impl} (efl := EW Program Definition erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : Transform.t _ _ - Ast.term EAst.term + Ast.term EAst.term _ _ TemplateProgram.eval_template_program (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := pre_erasure_pipeline ▷ diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 5a2e5ed3e..d85c2cb9a 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -25,11 +25,11 @@ Import Transform. Import EWcbvEval. Lemma transform_compose_assoc - {env env' env'' env''' term term' term'' term''' : Type} + {env env' env'' env''' term term' term'' term''' value value' value'' value''' : Type} {eval eval' eval'' eval'''} - (o : t env env' term term' eval eval') - (o' : t env' env'' term' term'' eval' eval'') - (o'' : t env'' env''' term'' term''' eval'' eval''') + (o : t env env' term term' value value' eval eval') + (o' : t env' env'' term' term'' value' value'' eval' eval'') + (o'' : t env'' env''' term'' term''' value'' value''' eval'' eval''') (prec : forall p, post o p -> pre o' p) (prec' : forall p, post o' p -> pre o'' p) : forall x p1, @@ -43,11 +43,11 @@ Proof. Qed. Lemma obseq_compose_assoc - {env env' env'' env''' term term' term'' term''' : Type} + {env env' env'' env''' term term' term'' term''' value value' value'' value''' : Type} {eval eval' eval'' eval'''} - (o : t env env' term term' eval eval') - (o' : t env' env'' term' term'' eval' eval'') - (o'' : t env'' env''' term'' term''' eval'' eval''') + (o : t env env' term term' value value' eval eval') + (o' : t env' env'' term' term'' value' value'' eval' eval'') + (o'' : t env'' env''' term'' term''' value'' value''' eval'' eval''') (prec : forall p, post o p -> pre o' p) (prec' : forall p, post o' p -> pre o'' p) : forall x p1 p2 v1 v2, obseq (compose o (compose o' o'' prec') prec) x p1 p2 v1 v2 <-> @@ -438,7 +438,7 @@ Module ETransformPresFO. Context {env env' : Type}. Context {eval : program env EAst.term -> EAst.term -> Prop}. Context {eval' : program env' EAst.term -> EAst.term -> Prop}. - Context (o : Transform.t _ _ _ _ eval eval'). + Context (o : Transform.t _ _ _ _ _ _ eval eval'). Context (firstorder_value : program env EAst.term -> Prop). Context (firstorder_value' : program env' EAst.term -> Prop). Context (compile_fo_value : forall p : program env EAst.term, o.(pre) p -> @@ -453,7 +453,7 @@ Module ETransformPresFO. Context {env env' : Type}. Context {eval : program env EAst.term -> EAst.term -> Prop}. Context {eval' : program env' EAst.term -> EAst.term -> Prop}. - Context (o : Transform.t _ _ _ _ eval eval'). + Context (o : Transform.t _ _ _ _ _ _ eval eval'). Context (firstorder_value : program env EAst.term -> Prop). Context (firstorder_value' : program env' EAst.term -> Prop). @@ -478,7 +478,7 @@ Module ETransformPresFO. Context (firstorder_value : program env EAst.term -> Prop). Context (firstorder_value' : program env' EAst.term -> Prop). Context (firstorder_value'' : program env'' EAst.term -> Prop). - Context (o : Transform.t _ _ _ _ eval eval') (o' : Transform.t _ _ _ _ eval' eval''). + Context (o : Transform.t _ _ _ _ _ _ eval eval') (o' : Transform.t _ _ _ _ _ _ eval' eval''). Context compile_fo_value compile_fo_value' (oext : t o firstorder_value firstorder_value' compile_fo_value) (o'ext : t o' firstorder_value' firstorder_value'' compile_fo_value') @@ -909,9 +909,8 @@ Section PCUICExpandLets. Proof. rewrite /PCUICFirstorder.firstorder_type /PCUICFirstorder.firstorder_env. pose proof (trans_decompose_app t). - destruct decompose_app. rewrite H. cbn. - destruct t0 => //. intros hd => /=. - destruct ind. + destruct decompose_app. rewrite {}H. cbn. + case: t0 => //; case => /= kn _ _ hd. destruct plookup_env eqn:hp => //. destruct b => //. eapply hd in hp. rewrite hp //. diff --git a/erasure/theories/EProgram.v b/erasure/theories/EProgram.v index 15cfe1e16..f069d6c25 100644 --- a/erasure/theories/EProgram.v +++ b/erasure/theories/EProgram.v @@ -55,10 +55,11 @@ Module TransformExt. Section Opt. Context {env env' env'' : Type}. Context {term term' term'' : Type}. - Context {eval : program env term -> term -> Prop}. - Context {eval' : program env' term' -> term' -> Prop}. - Context {eval'' : program env'' term'' -> term'' -> Prop}. - Context (o : Transform.t env env' term term' eval eval'). + Context {value value' value'' : Type}. + Context {eval : program env term -> value -> Prop}. + Context {eval' : program env' term' -> value' -> Prop}. + Context {eval'' : program env'' term'' -> value'' -> Prop}. + Context (o : Transform.t env env' term term' value value' eval eval'). Context (extends : program env term -> program env term -> Prop). Context (extends' : program env' term' -> program env' term' -> Prop). @@ -70,17 +71,18 @@ Module TransformExt. Section Comp. Context {env env' env'' : Type}. Context {term term' term'' : Type}. - Context {eval : program env term -> term -> Prop}. - Context {eval' : program env' term' -> term' -> Prop}. - Context {eval'' : program env'' term'' -> term'' -> Prop}. + Context {value value' value'' : Type}. + Context {eval : program env term -> value -> Prop}. + Context {eval' : program env' term' -> value' -> Prop}. + Context {eval'' : program env'' term'' -> value'' -> Prop}. Context {extends : program env term -> program env term -> Prop}. Context {extends' : program env' term' -> program env' term' -> Prop}. Context {extends'' : program env'' term'' -> program env'' term'' -> Prop}. Local Obligation Tactic := idtac. #[global] - Instance compose (o : Transform.t env env' term term' eval eval') - (o' : Transform.t env' env'' term' term'' eval' eval'') + Instance compose (o : Transform.t env env' term term' value value' eval eval') + (o' : Transform.t env' env'' term' term'' value' value'' eval' eval'') (oext : t o extends extends') (o'ext : t o' extends' extends'') (hpp : (forall p, o.(post) p -> o'.(pre) p)) diff --git a/pcuic/Makefile.plugin.local b/pcuic/Makefile.plugin.local index f8a95036e..e30465def 100644 --- a/pcuic/Makefile.plugin.local +++ b/pcuic/Makefile.plugin.local @@ -1,2 +1,4 @@ + +CAMLFLAGS :=-thread -bin-annot -strict-sequence -w -a+1..3-4+5..8-9+10..26-27+28..39-40-41-42+43-44-45+46..47-48+49..57-58+59..66-67-68+69-70 CAMLFLAGS+=-open Metacoq_template_plugin -open Metacoq_checker_plugin CAMLFLAGS+=-w -33 # Unused opens diff --git a/pcuic/theories/PCUICEtaExpand.v b/pcuic/theories/PCUICEtaExpand.v index 2b097e2e1..ef7fca054 100644 --- a/pcuic/theories/PCUICEtaExpand.v +++ b/pcuic/theories/PCUICEtaExpand.v @@ -5,7 +5,7 @@ From MetaCoq.Common Require Import config. From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping PCUICReduction PCUICProgram PCUICLiftSubst PCUICCSubst PCUICUnivSubst. -(* todo move *) +(* move *) Lemma All_fold_tip {A : Type} (P : list A -> A -> Type) {x} : All_fold P [x] -> P [] x. Proof. intros a; now depelim a. diff --git a/safechecker-plugin/Makefile.plugin.local b/safechecker-plugin/Makefile.plugin.local index 2514f50e5..62cfbc2d2 100644 --- a/safechecker-plugin/Makefile.plugin.local +++ b/safechecker-plugin/Makefile.plugin.local @@ -1,3 +1,4 @@ +CAMLFLAGS :=-thread -bin-annot -strict-sequence -w -a+1..3-4+5..8-9+10..26-27+28..39-40-41-42+43-44-45+46..47-48+49..57-58+59..66-67-68+69-70 CAMLFLAGS+=-open Metacoq_template_plugin CAMLFLAGS+=-w -8 # Non-exhaustive matches due to translation of comparison to int CAMLFLAGS+=-w -20 # Unused arguments diff --git a/template-coq/Makefile.plugin.local b/template-coq/Makefile.plugin.local index 3d1b9c40a..f52f9c737 100644 --- a/template-coq/Makefile.plugin.local +++ b/template-coq/Makefile.plugin.local @@ -1,3 +1,7 @@ +CAMLPKGS+=-package stdlib-shims +INSTALLDEFAULTROOT=MetaCoq/Template + +CAMLFLAGS :=-thread -bin-annot -strict-sequence -w -a+1..3-4+5..8-9+10..26-27+28..39-40-41-42+43-44-45+46..47-48+49..57-58+59..66-67-68+69-70 CAMLFLAGS+=-w -20 # Unused argument (produced by extraction) CAMLFLAGS+=-w -33 # Unused opens CAMLFLAGS+=-w -32 # Unused value @@ -7,7 +11,5 @@ CAMLFLAGS+=-w -34 # Unused type CAMLFLAGS+=-w -60 # Unused module # CAMLFLAGS+=-w -8 # Non-exhaustive pattern-matchings (BEWARE, just for extracted code) CAMLFLAGS+=-bin-annot # For merlin -CAMLPKGS+=-package stdlib-shims -INSTALLDEFAULTROOT=MetaCoq/Template .PHONY: META \ No newline at end of file diff --git a/template-coq/theories/TemplateProgram.v b/template-coq/theories/TemplateProgram.v index 6db0926ba..482e5ed03 100644 --- a/template-coq/theories/TemplateProgram.v +++ b/template-coq/theories/TemplateProgram.v @@ -47,7 +47,7 @@ Definition make_template_program_env {cf : checker_flags} (p : template_program) (** We kludge the normalization assumptions by parameterizing over a continuation of "what will be done to the program later" as well as what properties we'll need of it *) Program Definition build_template_program_env {cf : checker_flags} K : - Transform.t global_env GlobalEnvMap.t Ast.term Ast.term eval_template_program eval_template_program_env := + Transform.t global_env GlobalEnvMap.t Ast.term Ast.term Ast.term Ast.term eval_template_program eval_template_program_env := {| name := "rebuilding environment lookup table"; pre p := ∥ wt_template_program p ∥ /\ forall pf, K (GlobalEnvMap.make p.1 (wt_template_program_fresh p pf)) : Prop ; transform p pre := make_template_program_env p (proj1 pre); diff --git a/template-pcuic/theories/PCUICTransform.v b/template-pcuic/theories/PCUICTransform.v index ef840ddaf..8bafdc8c5 100644 --- a/template-pcuic/theories/PCUICTransform.v +++ b/template-pcuic/theories/PCUICTransform.v @@ -34,7 +34,7 @@ Local Obligation Tactic := idtac. (** We kludge the normalization assumptions by parameterizing over a continuation of "what will be done to the program later" as well as what properties we'll need of it *) Program Definition template_to_pcuic_transform {cf : checker_flags} K : - Transform.t Ast.Env.global_env global_env_ext_map Ast.term term + Transform.t Ast.Env.global_env global_env_ext_map Ast.term term Ast.term term eval_template_program eval_pcuic_program := {| name := "template to pcuic"; pre p := ∥ wt_template_program p ∥ /\ EtaExpand.expanded_program p /\ K (trans_global (Ast.Env.empty_ext p.1)) ; From f2b1cc0783dd745d9dd52e562bac30221599dd63 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 3 Oct 2023 10:30:52 +0200 Subject: [PATCH 57/61] Fix after rebase --- erasure-plugin/theories/ErasureCorrectness.v | 14 +++++++++----- 1 file changed, 9 insertions(+), 5 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index d85c2cb9a..dea582694 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -910,10 +910,12 @@ Section PCUICExpandLets. rewrite /PCUICFirstorder.firstorder_type /PCUICFirstorder.firstorder_env. pose proof (trans_decompose_app t). destruct decompose_app. rewrite {}H. cbn. - case: t0 => //; case => /= kn _ _ hd. - destruct plookup_env eqn:hp => //. - destruct b => //. - eapply hd in hp. rewrite hp //. + case: t0 => //=. + - intros n' hn. destruct l => //. + - case => /= kn _ _ hd. destruct l => //. + destruct plookup_env eqn:hp => //. + destruct b => //. + eapply hd in hp. rewrite hp //. Qed. Lemma trans_firstorder_mutind Σ m : @@ -950,6 +952,7 @@ Section PCUICExpandLets. clear -he hd. move: hd. rewrite /firstorder_type. destruct (decompose_app decl_type) eqn:e. cbn. destruct t0 => //; apply decompose_app_inv in e; rewrite e. + destruct l => //. { move/andP => [] /Nat.leb_le hn /Nat.ltb_lt hn'. rewrite trans_mkApps PCUICLiftSubst.lift_mkApps PCUICLiftSubst.subst_mkApps decompose_app_mkApps //=. @@ -964,7 +967,8 @@ Section PCUICExpandLets. apply/andP; split. apply Nat.leb_le. lia. apply Nat.ltb_lt. lia. } { rewrite trans_mkApps PCUICLiftSubst.lift_mkApps PCUICLiftSubst.subst_mkApps decompose_app_mkApps //=. - destruct ind. destruct plookup_env eqn:hp => //. destruct b => //. apply he in hp. rewrite hp //. } + destruct ind, l => //=. destruct plookup_env eqn:hp => //. + destruct b => //. apply he in hp. rewrite hp //. } Qed. Lemma trans_firstorder_env Σ : From 357e8a7ddd140f4396c18eb9bd95e822f9363d25 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 3 Oct 2023 11:27:01 +0200 Subject: [PATCH 58/61] Fix test suite --- test-suite/erasure_live_test.v | 2 ++ 1 file changed, 2 insertions(+) diff --git a/test-suite/erasure_live_test.v b/test-suite/erasure_live_test.v index c619b7d5e..b3f89a2fb 100644 --- a/test-suite/erasure_live_test.v +++ b/test-suite/erasure_live_test.v @@ -323,8 +323,10 @@ Require Import Coq.Arith.Compare_dec. Require Import Coq.Arith.PeanoNat. Require Import Coq.Arith.Peano_dec. Require Import Arith Init.Wf. +Require Import Program. Program Fixpoint provedCopy (n:nat) {wf lt n} : nat := match n with 0 => 0 | S k => S (provedCopy k) end. + Print Assumptions provedCopy. (* MetaCoq Quote Recursively Definition pCopy := provedCopy. program *) From 1dc79610c4994d1b0f03150bee2d6ea65617d321 Mon Sep 17 00:00:00 2001 From: nicolas tabareau Date: Tue, 10 Oct 2023 13:38:07 +0200 Subject: [PATCH 59/61] add not_isErasable lemma in EArities --- erasure/theories/EArities.v | 21 +++++++++++++++++++++ 1 file changed, 21 insertions(+) diff --git a/erasure/theories/EArities.v b/erasure/theories/EArities.v index 3b11eeb08..056133d5b 100644 --- a/erasure/theories/EArities.v +++ b/erasure/theories/EArities.v @@ -1031,3 +1031,24 @@ Proof. intros [Tty []] hred. exists Tty. split => //. eapply subject_reduction; tea. Qed. + +Lemma not_isErasable Σ Γ f A u : + wf_ext Σ -> wf_local Σ Γ -> + ∥Σ;;; Γ |- f : A∥ -> + (forall B, ∥Σ;;; Γ ⊢ A ⇝ B∥ -> A = B) -> + (forall B, ∥Σ ;;; Γ |- f : B∥ -> ∥Σ ;;; Γ ⊢ A ≤ B∥) -> + ~ ∥ isArity A ∥ -> + ∥ Σ;;; Γ |- A : tSort u ∥ -> + ~ is_propositional u -> + ~ ∥ Extract.isErasable Σ Γ f ∥. +Proof. + intros wfΣ Hlocal Hf Hnf Hprinc Harity Hfu Hu [[T [HT []]]]; sq. + - eapply Harity; sq. + eapply EArities.arity_type_inv in i as [T' [? ?]]; eauto. + eapply Hnf in H. subst; eauto. + - destruct s as [s [? ?]]. eapply Hu. + specialize (Hprinc _ (sq HT)). + pose proof (Hs := i). sq. + eapply PCUICElimination.unique_sorting_equality_propositional in Hprinc; eauto. + rewrite Hprinc; eauto. +Qed. From 34c831e73af5661e5b3ae2bd6c24cd056fd66eba Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Mon, 16 Oct 2023 11:01:24 +0200 Subject: [PATCH 60/61] fix Makefile --- Makefile | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/Makefile b/Makefile index 31e6a96fc..cae917e41 100644 --- a/Makefile +++ b/Makefile @@ -1,5 +1,5 @@ -all: printconf template-coq pcuic safechecker erasure erasure-plugin +all: printconf template-coq pcuic safechecker erasure erasure-plugin -include Makefile.conf @@ -26,7 +26,7 @@ else endif endif -install: all translations +install: all $(MAKE) -C utils install $(MAKE) -C common install $(MAKE) -C template-coq install @@ -37,7 +37,6 @@ install: all translations $(MAKE) -C safechecker-plugin install $(MAKE) -C erasure install $(MAKE) -C erasure-plugin install - $(MAKE) -C translations install uninstall: $(MAKE) -C utils uninstall @@ -177,7 +176,7 @@ cleanplugins: ci-local-noclean: ./configure.sh local - $(MAKE) all test-suite TIMED=pretty-timed + $(MAKE) all translations test-suite TIMED=pretty-timed ci-local: ci-local-noclean $(MAKE) clean From 39e31f6e0b56cd5ad5ea557aaa65a5a801a2c241 Mon Sep 17 00:00:00 2001 From: Yannick Forster Date: Mon, 16 Oct 2023 17:10:12 +0200 Subject: [PATCH 61/61] Remove nix --- .../workflows/nix-action-coq-dev-macos.yml | 187 ------------------ .../workflows/nix-action-coq-dev-ubuntu.yml | 187 ------------------ 2 files changed, 374 deletions(-) delete mode 100644 .github/workflows/nix-action-coq-dev-macos.yml delete mode 100644 .github/workflows/nix-action-coq-dev-ubuntu.yml diff --git a/.github/workflows/nix-action-coq-dev-macos.yml b/.github/workflows/nix-action-coq-dev-macos.yml deleted file mode 100644 index ed523bc63..000000000 --- a/.github/workflows/nix-action-coq-dev-macos.yml +++ /dev/null @@ -1,187 +0,0 @@ -jobs: - coq: - needs: [] - runs-on: macos-latest - concurrency: - group: ${{ github.workflow }}-MacOS-coq-${{ github.event_name }}-${{ github.head_ref || github.run_id }} - cancel-in-progress: true - steps: - - name: Determine which commit to test - run: "if [ ${{ github.event_name }} = \"push\" ]; then\n echo \"tested_commit=${{\ - \ github.sha }}\" >> $GITHUB_ENV\nelse\n merge_commit=$(git ls-remote ${{\ - \ github.event.repository.html_url }} refs/pull/${{ github.event.number }}/merge\ - \ | cut -f1)\n if [ -z \"$merge_commit\" ]; then\n echo \"tested_commit=${{\ - \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ - \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - - name: Git checkout - uses: actions/checkout@v4 - with: - fetch-depth: 0 - ref: ${{ env.tested_commit }} - - name: Cachix install - uses: cachix/install-nix-action@v23 - with: - nix_path: nixpkgs=channel:nixpkgs-unstable - - name: Cachix setup metacoq - uses: cachix/cachix-action@v12 - with: - authToken: ${{ secrets.CACHIX_AUTH_TOKEN }} - extraPullNames: coq, coq-community - name: metacoq - - id: stepCheck - name: Checking presence of CI target coq - run: "nb_dry_run=$(NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link \\\n --argstr\ - \ bundle \"coq-dev\" --argstr job \"coq\" \\\n --dry-run 2>&1 > /dev/null)\n\ - echo $nb_dry_run\necho ::set-output name=status::$(echo $nb_dry_run | grep\ - \ \"built:\" | sed \"s/.*/built/\")\n" - - if: steps.stepCheck.outputs.status == 'built' - name: Building/fetching current CI target - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "coq" - equations: - needs: - - coq - runs-on: macos-latest - concurrency: - group: ${{ github.workflow }}-MacOS-equations-${{ github.event_name }}-${{ github.head_ref || github.run_id }} - cancel-in-progress: true - steps: - - name: Determine which commit to test - run: "if [ ${{ github.event_name }} = \"push\" ]; then\n echo \"tested_commit=${{\ - \ github.sha }}\" >> $GITHUB_ENV\nelse\n merge_commit=$(git ls-remote ${{\ - \ github.event.repository.html_url }} refs/pull/${{ github.event.number }}/merge\ - \ | cut -f1)\n if [ -z \"$merge_commit\" ]; then\n echo \"tested_commit=${{\ - \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ - \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - - name: Git checkout - uses: actions/checkout@v4 - with: - fetch-depth: 0 - ref: ${{ env.tested_commit }} - - name: Cachix install - uses: cachix/install-nix-action@v23 - with: - nix_path: nixpkgs=channel:nixpkgs-unstable - - name: Cachix setup metacoq - uses: cachix/cachix-action@v12 - with: - authToken: ${{ secrets.CACHIX_AUTH_TOKEN }} - extraPullNames: coq, coq-community - name: metacoq - - id: stepCheck - name: Checking presence of CI target equations - run: "nb_dry_run=$(NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link \\\n --argstr\ - \ bundle \"coq-dev\" --argstr job \"equations\" \\\n --dry-run 2>&1 > /dev/null)\n\ - echo $nb_dry_run\necho ::set-output name=status::$(echo $nb_dry_run | grep\ - \ \"built:\" | sed \"s/.*/built/\")\n" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: coq' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "coq" - - if: steps.stepCheck.outputs.status == 'built' - name: Building/fetching current CI target - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "equations" - metacoq: - needs: - - coq - - equations - runs-on: macos-latest - concurrency: - group: ${{ github.workflow }}-MacOS-metacoq-${{ github.event_name }}-${{ github.head_ref || github.run_id }} - cancel-in-progress: true - steps: - - name: Determine which commit to test - run: "if [ ${{ github.event_name }} = \"push\" ]; then\n echo \"tested_commit=${{\ - \ github.sha }}\" >> $GITHUB_ENV\nelse\n merge_commit=$(git ls-remote ${{\ - \ github.event.repository.html_url }} refs/pull/${{ github.event.number }}/merge\ - \ | cut -f1)\n if [ -z \"$merge_commit\" ]; then\n echo \"tested_commit=${{\ - \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ - \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - - name: Git checkout - uses: actions/checkout@v4 - with: - fetch-depth: 0 - ref: ${{ env.tested_commit }} - - name: Cachix install - uses: cachix/install-nix-action@v23 - with: - nix_path: nixpkgs=channel:nixpkgs-unstable - - name: Cachix setup metacoq - uses: cachix/cachix-action@v12 - with: - authToken: ${{ secrets.CACHIX_AUTH_TOKEN }} - extraPullNames: coq, coq-community - name: metacoq - - id: stepCheck - name: Checking presence of CI target metacoq - run: "nb_dry_run=$(NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link \\\n --argstr\ - \ bundle \"coq-dev\" --argstr job \"metacoq\" \\\n --dry-run 2>&1 > /dev/null)\n\ - echo $nb_dry_run\necho ::set-output name=status::$(echo $nb_dry_run | grep\ - \ \"built:\" | sed \"s/.*/built/\")\n" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: coq' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "coq" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: equations' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "equations" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-utils' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-utils" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-common' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-common" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-template-coq' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-template-coq" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-pcuic' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-pcuic" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-safechecker' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-safechecker" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-template-pcuic' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-template-pcuic" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-erasure' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-erasure" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-quotation' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-quotation" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-safechecker-plugin' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-safechecker-plugin" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-erasure-plugin' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-erasure-plugin" - - if: steps.stepCheck.outputs.status == 'built' - name: Building/fetching current CI target - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq" -name: Nix CI for bundle coq-dev -'on': - pull_request: - paths: - - .github/workflows/** - pull_request_target: - types: - - opened - - synchronize - - reopened - push: - branches: - - main - - coq-dev diff --git a/.github/workflows/nix-action-coq-dev-ubuntu.yml b/.github/workflows/nix-action-coq-dev-ubuntu.yml deleted file mode 100644 index 324b24726..000000000 --- a/.github/workflows/nix-action-coq-dev-ubuntu.yml +++ /dev/null @@ -1,187 +0,0 @@ -jobs: - coq: - needs: [] - runs-on: ubuntu-latest - concurrency: - group: ${{ github.workflow }}-Ubuntu-coq-${{ github.event_name }}-${{ github.head_ref || github.run_id }} - cancel-in-progress: true - steps: - - name: Determine which commit to test - run: "if [ ${{ github.event_name }} = \"push\" ]; then\n echo \"tested_commit=${{\ - \ github.sha }}\" >> $GITHUB_ENV\nelse\n merge_commit=$(git ls-remote ${{\ - \ github.event.repository.html_url }} refs/pull/${{ github.event.number }}/merge\ - \ | cut -f1)\n if [ -z \"$merge_commit\" ]; then\n echo \"tested_commit=${{\ - \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ - \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - - name: Git checkout - uses: actions/checkout@v4 - with: - fetch-depth: 0 - ref: ${{ env.tested_commit }} - - name: Cachix install - uses: cachix/install-nix-action@v23 - with: - nix_path: nixpkgs=channel:nixpkgs-unstable - - name: Cachix setup metacoq - uses: cachix/cachix-action@v12 - with: - authToken: ${{ secrets.CACHIX_AUTH_TOKEN }} - extraPullNames: coq, coq-community - name: metacoq - - id: stepCheck - name: Checking presence of CI target coq - run: "nb_dry_run=$(NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link \\\n --argstr\ - \ bundle \"coq-dev\" --argstr job \"coq\" \\\n --dry-run 2>&1 > /dev/null)\n\ - echo $nb_dry_run\necho ::set-output name=status::$(echo $nb_dry_run | grep\ - \ \"built:\" | sed \"s/.*/built/\")\n" - - if: steps.stepCheck.outputs.status == 'built' - name: Building/fetching current CI target - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "coq" - equations: - needs: - - coq - runs-on: ubuntu-latest - concurrency: - group: ${{ github.workflow }}-Ubuntu-equations-${{ github.event_name }}-${{ github.head_ref || github.run_id }} - cancel-in-progress: true - steps: - - name: Determine which commit to test - run: "if [ ${{ github.event_name }} = \"push\" ]; then\n echo \"tested_commit=${{\ - \ github.sha }}\" >> $GITHUB_ENV\nelse\n merge_commit=$(git ls-remote ${{\ - \ github.event.repository.html_url }} refs/pull/${{ github.event.number }}/merge\ - \ | cut -f1)\n if [ -z \"$merge_commit\" ]; then\n echo \"tested_commit=${{\ - \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ - \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - - name: Git checkout - uses: actions/checkout@v4 - with: - fetch-depth: 0 - ref: ${{ env.tested_commit }} - - name: Cachix install - uses: cachix/install-nix-action@v23 - with: - nix_path: nixpkgs=channel:nixpkgs-unstable - - name: Cachix setup metacoq - uses: cachix/cachix-action@v12 - with: - authToken: ${{ secrets.CACHIX_AUTH_TOKEN }} - extraPullNames: coq, coq-community - name: metacoq - - id: stepCheck - name: Checking presence of CI target equations - run: "nb_dry_run=$(NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link \\\n --argstr\ - \ bundle \"coq-dev\" --argstr job \"equations\" \\\n --dry-run 2>&1 > /dev/null)\n\ - echo $nb_dry_run\necho ::set-output name=status::$(echo $nb_dry_run | grep\ - \ \"built:\" | sed \"s/.*/built/\")\n" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: coq' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "coq" - - if: steps.stepCheck.outputs.status == 'built' - name: Building/fetching current CI target - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "equations" - metacoq: - needs: - - coq - - equations - runs-on: ubuntu-latest - concurrency: - group: ${{ github.workflow }}-Ubuntu-metacoq-${{ github.event_name }}-${{ github.head_ref || github.run_id }} - cancel-in-progress: true - steps: - - name: Determine which commit to test - run: "if [ ${{ github.event_name }} = \"push\" ]; then\n echo \"tested_commit=${{\ - \ github.sha }}\" >> $GITHUB_ENV\nelse\n merge_commit=$(git ls-remote ${{\ - \ github.event.repository.html_url }} refs/pull/${{ github.event.number }}/merge\ - \ | cut -f1)\n if [ -z \"$merge_commit\" ]; then\n echo \"tested_commit=${{\ - \ github.event.pull_request.head.sha }}\" >> $GITHUB_ENV\n else\n echo\ - \ \"tested_commit=$merge_commit\" >> $GITHUB_ENV\n fi\nfi\n" - - name: Git checkout - uses: actions/checkout@v4 - with: - fetch-depth: 0 - ref: ${{ env.tested_commit }} - - name: Cachix install - uses: cachix/install-nix-action@v23 - with: - nix_path: nixpkgs=channel:nixpkgs-unstable - - name: Cachix setup metacoq - uses: cachix/cachix-action@v12 - with: - authToken: ${{ secrets.CACHIX_AUTH_TOKEN }} - extraPullNames: coq, coq-community - name: metacoq - - id: stepCheck - name: Checking presence of CI target metacoq - run: "nb_dry_run=$(NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link \\\n --argstr\ - \ bundle \"coq-dev\" --argstr job \"metacoq\" \\\n --dry-run 2>&1 > /dev/null)\n\ - echo $nb_dry_run\necho ::set-output name=status::$(echo $nb_dry_run | grep\ - \ \"built:\" | sed \"s/.*/built/\")\n" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: coq' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "coq" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: equations' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "equations" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-utils' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-utils" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-common' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-common" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-template-coq' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-template-coq" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-pcuic' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-pcuic" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-safechecker' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-safechecker" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-template-pcuic' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-template-pcuic" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-erasure' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-erasure" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-quotation' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-quotation" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-safechecker-plugin' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-safechecker-plugin" - - if: steps.stepCheck.outputs.status == 'built' - name: 'Building/fetching previous CI target: metacoq-erasure-plugin' - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq-erasure-plugin" - - if: steps.stepCheck.outputs.status == 'built' - name: Building/fetching current CI target - run: NIXPKGS_ALLOW_UNFREE=1 nix-build --no-out-link --argstr bundle "coq-dev" - --argstr job "metacoq" -name: Nix CI for bundle coq-dev -'on': - pull_request: - paths: - - .github/workflows/** - pull_request_target: - types: - - opened - - synchronize - - reopened - push: - branches: - - main - - coq-dev