From 35b62250f812cbb9c46d34540c82d914fb3c55cf Mon Sep 17 00:00:00 2001 From: nicolas tabareau Date: Tue, 30 Jan 2024 09:32:50 +0100 Subject: [PATCH 01/28] erasure_pipeline_extends_app with todo irrel --- erasure-plugin/theories/ErasureCorrectness.v | 65 +++++++++++++++++--- 1 file changed, 58 insertions(+), 7 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index fbeac012c..83424da23 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -763,7 +763,7 @@ Definition fo_evalue (p : program E.global_context EAst.term) : Prop := firstord 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_fo {fl : WcbvFlags} {efl : EEnvFlags} we wf : - ETransformPresFO.t + ETransformPresFO.t (rebuild_wf_env_transform we wf) fo_evalue fo_evalue_map (fun p pr fo => rebuild_wf_env p pr.p1). Proof. split => //. Qed. @@ -1432,16 +1432,17 @@ Section PCUICErase. Lemma extends_eq Σ Σ0 Σ' : EGlobalEnv.extends Σ Σ' -> Σ = Σ0 -> EGlobalEnv.extends Σ0 Σ'. Proof. now intros ext ->. Qed. - Lemma erasure_pipeline_eta_app (Σ : global_env_ext_map) t u pre : + Lemma erasure_pipeline_extends_app (Σ : global_env_ext_map) t u pre : ∥ nisErasable Σ [] (tApp t u) ∥ -> PCUICEtaExpand.expanded Σ [] t -> exists pre' pre'', - transform verified_erasure_pipeline (Σ, tApp t u) pre = - ((transform verified_erasure_pipeline (Σ, tApp t u) pre).1, - EAst.tApp (transform verified_erasure_pipeline (Σ, t) pre').2 - (transform verified_erasure_pipeline (Σ, u) pre'').2). + let trapp := transform verified_erasure_pipeline (Σ, PCUICAst.tApp t u) pre in + let trt := transform verified_erasure_pipeline (Σ, t) pre' in + let tru := transform verified_erasure_pipeline (Σ, u) pre'' in + (EGlobalEnv.extends trt.1 trapp.1 /\ EGlobalEnv.extends tru.1 trapp.1) /\ + trapp = (trapp.1, EAst.tApp trt.2 tru.2). Proof. - intros ner exp. + intros ner exp. unfold verified_erasure_pipeline. destruct_compose. set (K:= (fun p : global_env_ext_map => (wf_ext p -> PCUICSN.NormalizationIn p) /\ (wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p))). @@ -1450,6 +1451,56 @@ Section PCUICErase. { destruct ner as [ner]. destruct pre, s. eapply nisErasable_spec in ner => //. eapply w. } destruct (expand_lets_eta_app _ _ _ K pre ner' exp) as [pre' [pre'' eq]]. exists pre', pre''. + split. + set (tr := transform _ _ _). + destruct tr eqn:heq. cbn -[transform]. + replace t0 with tr.2. assert (heq_env:tr.1=g) by now rewrite heq. subst tr. + 2:{ now rewrite heq. } + clear heq. revert H. + destruct_compose_no_clear. rewrite eq. intros pre3 eq2 pre4. + epose proof (erase_eta_app _ _ _ pre3) as H0. + pose proof (correctness (pcuic_expand_lets_transform K) (Σ, tApp t u) pre). + destruct H as [[wtapp] [expapp Kapp]]. + pose proof (correctness (pcuic_expand_lets_transform K) (Σ, t) pre'). + destruct H as [[wtt] [expt Kt]]. + forward H0. + { clear -wtapp ner eq. apply (f_equal snd) in eq. cbn [snd] in eq. rewrite -eq. + destruct pre as [[wtp] rest]. + destruct ner as [ner]. eapply (nisErasable_lets (Σ, tApp t u)) in ner. + eapply nisErasable_spec in ner => //. cbn. + apply wtapp. apply wtp. } + forward H0 by apply expt. + destruct H0 as [pre'0 [pre''0 [eta [extapp [extapp' heq]]]]]. + split. + { rewrite <- heq_env. cbn -[transform]. unfold verified_lambdabox_pipeline. + repeat (destruct_compose; intros). eapply constructors_as_blocks_extends. + repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. + repeat (destruct_compose; intros). eapply inline_projections_optimization_extends. + repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. + repeat (destruct_compose; intros). eapply remove_match_on_box_extends. + repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. + repeat (destruct_compose; intros). eapply remove_params_extends. + repeat (destruct_compose; intros). eapply guarded_to_unguarded_fix_extends. + repeat (destruct_compose; intros). cbn - [transform]. + generalize dependent pre3. rewrite <- eq; intros. revert extapp. intro. + cbn; cbn in extapp. unfold expand_lets_program; cbn. + revert H14; intro. cbn in H14. red in H14. cbn in H14. + todo "irrel". + } + { cbn -[transform]. unfold verified_lambdabox_pipeline. + repeat (destruct_compose; intros). eapply constructors_as_blocks_extends. + repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. + repeat (destruct_compose; intros). eapply inline_projections_optimization_extends. + repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. + repeat (destruct_compose; intros). eapply remove_match_on_box_extends. + repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. + repeat (destruct_compose; intros). eapply remove_params_extends. + repeat (destruct_compose; intros). eapply guarded_to_unguarded_fix_extends. + repeat (destruct_compose; intros). cbn - [transform]. + generalize dependent pre3. rewrite <- eq; intros. revert extapp'. intro. + cbn; cbn in extapp'. unfold expand_lets_program; cbn. + todo "irrel". + } set (tr := transform _ _ _). destruct tr eqn:heq. cbn -[transform]. f_equal. replace t0 with tr.2. subst tr. From 78937906b8f15e10fe4f0b0b558cf9fddfbe1f14 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 1 Feb 2024 12:07:19 +0100 Subject: [PATCH 02/28] Fill proofs --- erasure-plugin/theories/ErasureCorrectness.v | 104 ++++++++----------- erasure/theories/ErasureFunction.v | 67 ++++++++++++ 2 files changed, 108 insertions(+), 63 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 83424da23..c2576d635 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -293,16 +293,16 @@ Section PCUICProof. decl = erase_mutual_inductive_body decl'). Definition erase_decl_equal cmp decl decl' := - match decl, decl' with + match decl, decl' with EAst.InductiveDecl decl , InductiveDecl decl' => decl = cmp decl' | EAst.ConstantDecl _ , ConstantDecl _ => True - | _ , _ => False - end. + | _ , _ => False + end. 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 decl -> - exists decl', lookup_global decls kn = Some decl' /\ + exists decl', lookup_global decls kn = Some decl' /\ erase_decl_equal erase_mutual_inductive_body decl decl'. Proof. induction decls in deps, X, normalization_in, prf |- *; cbn [erase_global_deps] => //. @@ -310,8 +310,8 @@ Section PCUICProof. - case: (knset_mem_spec k deps) => // hdeps. cbn [EGlobalEnv.lookup_env fst lookup_env lookup_global]. { destruct (eqb_spec kn k) => //. - intro hl; depelim hl; cbn. intro e; noconf e. - eexists; split; eauto. cbn; eauto. + intro hl; depelim hl; cbn. intro e; noconf e. + eexists; split; eauto. cbn; eauto. intros hl. cbn. eapply IHdecls. now depelim hl. } { intros hl. depelim hl. intros hl'. @@ -321,12 +321,12 @@ Section PCUICProof. destruct (eqb_spec kn k) => //. subst k. destruct H0. now eapply PCUICWeakeningEnv.lookup_global_Some_fresh in H0. - exact hl'. } + 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) => //; cbn. - intros hl. noconf hl. subst k. eexists; split; cbn; eauto. cbn; eauto. + intros hl. noconf hl. subst k. eexists; split; cbn; eauto. cbn; eauto. intros hl'. eapply IHdecls => //. tea. } { intros hl'. eapply IHdecls in hf; tea. destruct hf. exists x. @@ -763,7 +763,7 @@ Definition fo_evalue (p : program E.global_context EAst.term) : Prop := firstord 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_fo {fl : WcbvFlags} {efl : EEnvFlags} we wf : - ETransformPresFO.t + ETransformPresFO.t (rebuild_wf_env_transform we wf) fo_evalue fo_evalue_map (fun p pr fo => rebuild_wf_env p pr.p1). Proof. split => //. Qed. @@ -1439,10 +1439,10 @@ Section PCUICErase. let trapp := transform verified_erasure_pipeline (Σ, PCUICAst.tApp t u) pre in let trt := transform verified_erasure_pipeline (Σ, t) pre' in let tru := transform verified_erasure_pipeline (Σ, u) pre'' in - (EGlobalEnv.extends trt.1 trapp.1 /\ EGlobalEnv.extends tru.1 trapp.1) /\ + (EGlobalEnv.extends trt.1 trapp.1 /\ EGlobalEnv.extends tru.1 trapp.1) /\ trapp = (trapp.1, EAst.tApp trt.2 tru.2). Proof. - intros ner exp. + intros ner exp. unfold verified_erasure_pipeline. destruct_compose. set (K:= (fun p : global_env_ext_map => (wf_ext p -> PCUICSN.NormalizationIn p) /\ (wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p))). @@ -1451,9 +1451,8 @@ Section PCUICErase. { destruct ner as [ner]. destruct pre, s. eapply nisErasable_spec in ner => //. eapply w. } destruct (expand_lets_eta_app _ _ _ K pre ner' exp) as [pre' [pre'' eq]]. exists pre', pre''. - split. - set (tr := transform _ _ _). - destruct tr eqn:heq. cbn -[transform]. + set (tr := transform _ _ _). + destruct tr eqn:heq. cbn -[transform]. replace t0 with tr.2. assert (heq_env:tr.1=g) by now rewrite heq. subst tr. 2:{ now rewrite heq. } clear heq. revert H. @@ -1472,54 +1471,31 @@ Section PCUICErase. forward H0 by apply expt. destruct H0 as [pre'0 [pre''0 [eta [extapp [extapp' heq]]]]]. split. - { rewrite <- heq_env. cbn -[transform]. unfold verified_lambdabox_pipeline. - repeat (destruct_compose; intros). eapply constructors_as_blocks_extends. - repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. - repeat (destruct_compose; intros). eapply inline_projections_optimization_extends. - repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. - repeat (destruct_compose; intros). eapply remove_match_on_box_extends. - repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. - repeat (destruct_compose; intros). eapply remove_params_extends. - repeat (destruct_compose; intros). eapply guarded_to_unguarded_fix_extends. - repeat (destruct_compose; intros). cbn - [transform]. - generalize dependent pre3. rewrite <- eq; intros. revert extapp. intro. - cbn; cbn in extapp. unfold expand_lets_program; cbn. - revert H14; intro. cbn in H14. red in H14. cbn in H14. - todo "irrel". - } - { cbn -[transform]. unfold verified_lambdabox_pipeline. - repeat (destruct_compose; intros). eapply constructors_as_blocks_extends. - repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. - repeat (destruct_compose; intros). eapply inline_projections_optimization_extends. - repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. - repeat (destruct_compose; intros). eapply remove_match_on_box_extends. - repeat (destruct_compose; intros). eapply rebuild_wf_env_extends. - repeat (destruct_compose; intros). eapply remove_params_extends. - repeat (destruct_compose; intros). eapply guarded_to_unguarded_fix_extends. - repeat (destruct_compose; intros). cbn - [transform]. - generalize dependent pre3. rewrite <- eq; intros. revert extapp'. intro. - cbn; cbn in extapp'. unfold expand_lets_program; cbn. - todo "irrel". - } + { rewrite <- heq_env. cbn -[transform]. + pose proof (EProgram.TransformExt.preserves_obs _ _ _ (t:=verified_lambdabox_pipeline_extends')). + unfold extends_eprogram in H. + split. + { repeat (destruct_compose; intros). eapply verified_lambdabox_pipeline_extends. + repeat (destruct_compose; intros). cbn - [transform]. + generalize dependent pre3. rewrite <- eq. + cbn [transform pcuic_expand_lets_transform expand_lets_program]. + unfold expand_lets_program. cbn [fst snd]. + intros pre3. cbn in pre3. intros <-. intros. + assert (pre'0 = H1). apply proof_irrelevance. subst H1. + exact extapp. } + { repeat (destruct_compose; intros). eapply verified_lambdabox_pipeline_extends. + repeat (destruct_compose; intros). cbn - [transform]. + generalize dependent pre3. rewrite <- eq. + cbn [transform pcuic_expand_lets_transform expand_lets_program]. + unfold expand_lets_program. cbn [fst snd]. + intros pre3. cbn in pre3. intros <-. intros. + assert (pre''0 = H1). apply proof_irrelevance. subst H1. + exact extapp'. } } set (tr := transform _ _ _). - destruct tr eqn:heq. cbn -[transform]. f_equal. - replace t0 with tr.2. subst tr. - 2:{ now rewrite heq. } - clear heq. revert H. - destruct_compose_no_clear. rewrite eq. intros pre3 eq2 pre4. - epose proof (erase_eta_app _ _ _ pre3) as H0. - pose proof (correctness (pcuic_expand_lets_transform K) (Σ, tApp t u) pre). - destruct H as [[wtapp] [expapp Kapp]]. - pose proof (correctness (pcuic_expand_lets_transform K) (Σ, t) pre'). - destruct H as [[wtt] [expt Kt]]. - forward H0. - { clear -wtapp ner eq. apply (f_equal snd) in eq. cbn [snd] in eq. rewrite -eq. - destruct pre as [[wtp] rest]. - destruct ner as [ner]. eapply (nisErasable_lets (Σ, tApp t u)) in ner. - eapply nisErasable_spec in ner => //. cbn. - apply wtapp. apply wtp. } - forward H0 by apply expt. - destruct H0 as [pre'0 [pre''0 [eta [extapp [extapp' heq]]]]]. + destruct tr eqn:heqtr. cbn -[transform]. f_equal. + replace t1 with tr.2. subst tr. + 2:{ now rewrite heqtr; cbn. } + clear heqtr. move: pre4. rewrite heq. intros h. epose proof (transform_lambda_box_eta_app _ _ _ h). @@ -1559,6 +1535,8 @@ Section PCUICErase. { red. cbn. split => //. } reflexivity. Qed. + Transparent erase_transform. + End PCUICErase. Lemma compile_evalue_strip (Σer : EEnvMap.GlobalContextMap.t) p : @@ -1622,6 +1600,7 @@ Section pipeline_cond. End pipeline_cond. + Section pipeline_theorem. Instance cf_ : checker_flags := extraction_checker_flags. @@ -1784,7 +1763,7 @@ Section pipeline_theorem. revert H8. cbn. set (Σ.1). induction 1; econstructor; eauto. cbn. clear -H. induction H; econstructor; eauto. } destruct Hlookup as [decl'' [? ?]]. exists decl''; split ; eauto. - cbn in H10. inversion H10. + cbn in H10. inversion H10. now destruct decl' , decl''. Qed. @@ -1963,4 +1942,3 @@ Section pipeline_theorem. Qed. End pipeline_theorem. - diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index e1b500218..283724d96 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -1799,6 +1799,73 @@ Proof. destruct KernameSet.mem; cbn. f_equal. f_equal. 1-2:f_equal. all:now eapply IHdecls. Qed. + +Lemma erase_global_irr X_type (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 X decls (normalization_in:=normalization_in) prf = +erase_global X' decls (normalization_in:=normalization_in') prf'. +Proof. +revert X X' 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. + red. intros. + 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' _ _ H0). + rewrite -(abstract_env_lookup_correct' _ _ H1). + move: H4 H5. rewrite /lookup_env H2 H3 /= Hd Hd'. congruence. } + intros. + { 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')). + f_equal; f_equal. f_equal; now eapply 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. } + f_equal. f_equal. all:now eapply IHdecls. +Qed. + (* TODO: Figure out if this (and [erases]) should use [strictly_extends_decls] or [extends] -Jason Gross *) Lemma erases_weakening_env {Σ Σ' : global_env_ext} {Γ t t' T} : wf Σ -> wf Σ' -> strictly_extends_decls Σ Σ' -> From f9d5fa69ef9a3a2524d1df2ec5843aaeeda412d9 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 13 Feb 2024 07:24:18 +0100 Subject: [PATCH 03/28] Fix typed erasure calls to _not_ trim inductive masks --- erasure-plugin/theories/ETransform.v | 36 +++-- erasure-plugin/theories/Erasure.v | 28 +++- erasure/theories/EPretty.v | 4 +- .../theories/Typed/ExtractionCorrectness.v | 12 +- test-suite/erasure_live_test.v | 133 ++++++++++++++++++ test-suite/erasure_test.v | 6 + 6 files changed, 192 insertions(+), 27 deletions(-) diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index ab431bf18..717e7320f 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -205,7 +205,7 @@ Program Definition erase_transform {guard : abstract_guard_impl} : Transform.t _ /\ 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]; + post p := [/\ wf_eprogram_env all_env_flags p & EEtaExpandedFix.expanded_eprogram_env p]; obseq p hp p' v v' := let Σ := p.1 in Σ ;;; [] |- v ⇝ℇ v' |}. Next Obligation. @@ -474,17 +474,22 @@ Next Obligation. Unshelve. eauto. Qed. -Definition check_dearging_precond Σ t := +Record dearging_config := + { overridden_masks : kername -> option bitmask; + do_trim_const_masks : bool; + do_trim_ctor_masks : bool }. + +Definition check_dearging_precond (cf : dearging_config) Σ t := if closed_env (trans_env Σ) then - match ExtractionCorrectness.compute_masks (fun _ : kername => None) true true Σ with + match ExtractionCorrectness.compute_masks cf.(overridden_masks) cf.(do_trim_const_masks) cf.(do_trim_ctor_masks) Σ with | ResultMonad.Ok masks => if valid_cases (ind_masks masks) t && is_expanded (ind_masks masks) (const_masks masks) t then Some masks else None | _ => None end else None. -Lemma check_dearging_precond_spec env t d : - check_dearging_precond env t = Some d -> +Lemma check_dearging_precond_spec cf env t d : + check_dearging_precond cf env t = Some d -> closed_env (trans_env env) /\ valid_masks_env (ind_masks d) (const_masks d) env /\ is_expanded_env (ind_masks d) (const_masks d) env /\ @@ -502,22 +507,22 @@ Proof. intros [= <-]; cbn; intuition. Qed. -Definition check_dearging_trans p := - ((check_dearging_precond p.1 p.2, p.1), p.2). +Definition check_dearging_trans cf p := + ((check_dearging_precond cf p.1 p.2, p.1), p.2). Definition eval_typed_eprogram_masks fl := (fun (p : (option dearg_set * global_env) * term) v => ∥ @EWcbvEval.eval fl (ExAst.trans_env p.1.2) p.2 v ∥). -Definition post_dearging_checks efl (p : (option dearg_set × global_env) × term) := - (p.1.1 = check_dearging_precond p.1.2 p.2) /\ +Definition post_dearging_checks efl cf (p : (option dearg_set × global_env) × term) := + (p.1.1 = check_dearging_precond cf p.1.2 p.2) /\ wf_eprogram efl (trans_env p.1.2, p.2) /\ EEtaExpandedFix.expanded_eprogram (trans_env p.1.2, p.2). -Program Definition dearging_checks_transform {efl : EEnvFlags} {hastrel : has_tRel} {hastbox : has_tBox} : +Program Definition dearging_checks_transform {efl : EEnvFlags} cf {hastrel : has_tRel} {hastbox : has_tBox} : Transform.t _ _ EAst.term EAst.term _ _ (eval_typed_eprogram opt_wcbv_flags) (eval_typed_eprogram_masks opt_wcbv_flags) := {| name := "dearging"; - transform p _ := check_dearging_trans p ; + transform p _ := check_dearging_trans cf p ; pre p := pre_remove_match_on_box_typed efl p; - post := post_dearging_checks efl; + post := post_dearging_checks efl cf; obseq p hp p' v v' := v' = v |}. Next Obligation. cbn. intros. split => //. @@ -535,11 +540,11 @@ Section Dearging. | None => (trans_env p.1.2, p.2) end. - Program Definition dearging_transform (efl := all_env_flags) : + Program Definition dearging_transform (efl := all_env_flags) cf : Transform.t _ _ EAst.term EAst.term _ _ (eval_typed_eprogram_masks opt_wcbv_flags) (eval_eprogram opt_wcbv_flags) := {| name := "dearging"; transform p _ := dearg p ; - pre p := post_dearging_checks efl p; + pre p := post_dearging_checks efl cf p; post p := wf_eprogram efl p /\ EEtaExpandedFix.expanded_eprogram p ; obseq p hp p' v v' := v' = (dearg (p.1, v)).2 |}. @@ -572,7 +577,8 @@ Section Dearging. intros [ev]. unfold dearg. rewrite he. destruct (check_dearging_precond) as [masks|] eqn:cpre. - * eapply (extract_correct_gen' _ _ _ masks) in ev as ev'. destruct ev' as [ev']. + * eapply (extract_correct_gen' _ _ _ masks cf.(overridden_masks) cf.(do_trim_const_masks) cf.(do_trim_ctor_masks)) in ev as ev'. + destruct ev' as [ev']. eexists. split. red. sq. cbn. exact ev'. auto. - destruct wfe. now eapply wellformed_closed_env. - destruct wfe. now eapply wellformed_closed. diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 5b55c8016..05e3a084e 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -208,7 +208,9 @@ Program Definition verified_lambdabox_typed_pipeline {guard : abstract_guard_imp Local Obligation Tactic := intros; eauto. -Program Definition verified_typed_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : +Program Definition verified_typed_erasure_pipeline {guard : abstract_guard_impl} + (efl := EWellformed.all_env_flags) + (cf : dearging_config) : Transform.t _ _ PCUICAst.term EAst.term _ _ PCUICTransform.eval_pcuic_program @@ -226,8 +228,8 @@ Program Definition verified_typed_erasure_pipeline {guard : abstract_guard_impl} (* Remove match on box early for dearging *) remove_match_on_box_typed_transform (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ (* Check if the preconditions for dearging are valid, otherwise dearging will be the identity *) - dearging_checks_transform (hastrel := eq_refl) (hastbox := eq_refl) ▷ - dearging_transform ▷ + dearging_checks_transform cf (hastrel := eq_refl) (hastbox := eq_refl) ▷ + dearging_transform cf ▷ rebuild_wf_env_transform true true ▷ verified_lambdabox_typed_pipeline. @@ -238,13 +240,15 @@ Program Definition verified_typed_erasure_pipeline {guard : abstract_guard_impl} cbn in H |- *; intuition eauto. Qed. -Program Definition typed_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : +Program Definition typed_erasure_pipeline {guard : abstract_guard_impl} + (efl := EWellformed.all_env_flags) + cf : 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 |}) := pre_erasure_pipeline ▷ - verified_typed_erasure_pipeline. + verified_typed_erasure_pipeline cf. (* 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 @@ -355,9 +359,10 @@ Next Obligation. split; typeclasses eauto. Qed. -Program Definition typed_erase_and_print_template_program (p : Ast.Env.program) +(* Parameterized by a configuration for dearging, allowing to, e.g., override masks. *) +Program Definition typed_erase_and_print_template_program_gen (p : Ast.Env.program) cf : string := - let p' := run typed_erasure_pipeline p _ in + let p' := run (typed_erasure_pipeline cf) p _ in time "Pretty printing" EPretty.print_program p'. Next Obligation. split. @@ -367,3 +372,12 @@ Next Obligation. split; typeclasses eauto. Qed. +Definition default_dearging_config := + {| overridden_masks := fun _ => None; + do_trim_const_masks := true; + do_trim_ctor_masks := false |}. + +Definition typed_erase_and_print_template_program (p : Ast.Env.program) + : string := + typed_erase_and_print_template_program_gen p default_dearging_config. + diff --git a/erasure/theories/EPretty.v b/erasure/theories/EPretty.v index 9afba35b0..47324cd2d 100644 --- a/erasure/theories/EPretty.v +++ b/erasure/theories/EPretty.v @@ -225,7 +225,9 @@ Module PrintTermTree. Notation print_env := print_global_context. Definition print_program (p : program) : t := - pr p.1 p.2 ^ nl ^ "in" ^ print_env p.1. + "Environment: " ^ nl ^ + print_env p.1 ^ nl ^ + "Program: " ^ pr p.1 p.2. End PrintTermTree. diff --git a/erasure/theories/Typed/ExtractionCorrectness.v b/erasure/theories/Typed/ExtractionCorrectness.v index 1a9db7311..95b0fd34a 100644 --- a/erasure/theories/Typed/ExtractionCorrectness.v +++ b/erasure/theories/Typed/ExtractionCorrectness.v @@ -641,7 +641,10 @@ Theorem extract_correct_gen (H := EWellformed.all_env_flags) (wfl := default_wcbv_flags) (Σ : P.global_env_ext) (wfΣ : ∥wf_ext Σ∥) - t v ignored masks : + t v ignored masks + overridden_masks + do_trim_const_masks do_trim_ctor_masks + : axiom_free Σ -> forall wt : welltyped Σ [] t, Σ p⊢ t ⇓ v -> @@ -653,7 +656,7 @@ Theorem extract_correct_gen let gerΣ := EEnvMap.GlobalContextMap.make (trans_env erΣ) (trans_env_fresh_globals erΣ (remove_match_on_box_env_lemma (map_squash fst wfΣ) deps ignored)) in let erΣ := remove_match_on_box_env erΣ (remove_match_on_box_env_lemma _ deps ignored) in - compute_masks (fun _ => None) true true erΣ = Ok masks -> + compute_masks overridden_masks do_trim_const_masks do_trim_ctor_masks erΣ = Ok masks -> let t'' := EOptimizePropDiscr.remove_match_on_box gerΣ t' in valid_cases (ind_masks masks) t'' -> is_expanded (ind_masks masks) (const_masks masks) t'' -> @@ -712,11 +715,12 @@ Qed. Theorem extract_correct_gen' (H := EWellformed.all_env_flags) (wfl := opt_wcbv_flags) - (Σ : ExAst.global_env) t v masks : + (Σ : ExAst.global_env) t v masks + overridden_masks do_trim_const_masks do_trim_ctor_masks : EGlobalEnv.closed_env (trans_env Σ) -> ELiftSubst.closedn 0 t -> trans_env Σ e⊢ t ⇓ v -> - compute_masks (fun _ => None) true true Σ = Ok masks -> + compute_masks overridden_masks do_trim_const_masks do_trim_ctor_masks Σ = Ok masks -> valid_cases (ind_masks masks) t -> is_expanded (ind_masks masks) (const_masks masks) t -> ∥ EWcbvEval.eval (trans_env (dearg_env masks Σ)) (dearg_term masks t) (dearg_term masks v) ∥. diff --git a/test-suite/erasure_live_test.v b/test-suite/erasure_live_test.v index fd4d37a63..277658dc9 100644 --- a/test-suite/erasure_live_test.v +++ b/test-suite/erasure_live_test.v @@ -351,6 +351,15 @@ From MetaCoq.ErasurePlugin Require Import Loader. MetaCoq Erase provedCopyx. MetaCoq Typed Erase provedCopyx. + +(* From MetaCoq.Erasure.Typed Require Import CertifyingEta. +MetaCoq Run (eta_expand_def +(fun _ => None) +true true +provedCopy). *) + +Print P_provedCopyx. + (* 0.2s purely in the bytecode VM *) (*Time Definition P_provedCopyxvm' := Eval vm_compute in (test p_provedCopyx). *) (* Goal @@ -422,3 +431,127 @@ Proof. show_match. *) + + +From MetaCoq.Common Require Import Transform. +From MetaCoq.Erasure.Typed Require Import ExtractionCorrectness. +From MetaCoq.PCUIC Require Import PCUICAst PCUICProgram. +From MetaCoq.ErasurePlugin Require Import Erasure. +Import Transform. + +Program Definition verified_typed_erasure_pipeline + (cf : checker_flags := extraction_checker_flags) + {guard : PCUICWfEnvImpl.abstract_guard_impl} + (efl := EWellformed.all_env_flags) : + Transform.t _ _ + PCUICAst.term EAst.term _ _ + PCUICTransform.eval_pcuic_program _ := + (* a bunch of nonsense for normalization preconditions *) + let K ty (T : ty -> global_env_ext) p + := let p := T p in + (PCUICTyping.wf_ext p -> @PCUICSN.NormalizationIn _ _ p) /\ + (PCUICTyping.wf_ext p -> @PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn _ _ p) in + let T1 (p:PCUICProgram.global_env_ext_map) := p in + (* Branches of cases are expanded to bind only variables, constructor types are expanded accordingly *) + PCUICTransform.pcuic_expand_lets_transform (K _ T1) ▷ + (* Erasure of proofs terms in Prop and types *) + ETransform.typed_erase_transform ▷ + (* Remove match on box early for dearging *) + ETransform.remove_match_on_box_typed_transform (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl). + + (* ▷ + (* Check if the preconditions for dearging are valid, otherwise dearging will be the identity *) + ETransform.dearging_checks_transform (hastrel := eq_refl) (hastbox := eq_refl). *) + +#[local] Existing Instance extraction_checker_flags. +Next Obligation. exact extraction_checker_flags. Defined. +Next Obligation. exact extraction_normalizing. Defined. + +Next Obligation. exact extraction_checker_flags. Defined. +Next Obligation. exact extraction_normalizing. Defined. +Next Obligation. now split. Qed. + +Program Definition pipeline_upto (cf : checker_flags := extraction_checker_flags) + (guard := fake_guard_impl) + (efl := EWellformed.all_env_flags) := + Transform.compose pre_erasure_pipeline + verified_typed_erasure_pipeline _. + +Program Definition exintro_typed_er p := Transform.Transform.run pipeline_upto p _. +Next Obligation. cbn. todo "assum wt". Qed. + +MetaCoq Quote Recursively Definition exintro_proj := + (proj1_sig (@exist _ (fun x => x = 0) 0 (@eq_refl _ 0))). + +Eval lazy in testty cbv_provedCopyx. + +Definition exintro_before_checks := Eval compute in exintro_typed_er exintro_proj. + +MetaCoq Quote Recursively Definition quoted_provedCopyx := (provedCopy x). + +Definition provedCop_before_checks' := Eval lazy in exintro_typed_er cbv_provedCopyx. + +Definition provedCop_before_checks := Eval lazy in exintro_typed_er quoted_provedCopyx. + +Import ETransform Optimize. + +Definition masks := +Eval compute in Optimize.analyze_env (fun _ : kername => None) exintro_before_checks.1. + +Eval compute in (masks).(Optimize.ind_masks). + +Eval compute in Optimize.trim_ind_masks (masks).(Optimize.ind_masks). +Eval compute in Optimize.valid_masks_env ( Optimize.trim_ind_masks masks.(ind_masks)) masks.(const_masks) (exintro_before_checks.1). + + +Eval compute in compute_masks (fun _ : kername => None) false false exintro_before_checks.1. + +Definition masksprovd := analyze_env (fun _ : kername => None) provedCop_before_checks.1. + +Lemma forallb_cons {A} (p : A -> bool) hd tl : p hd -> forallb p tl -> forallb p (cons hd tl) . +Proof. Admitted. +Lemma and_andb (b b' : bool) : b -> b' -> b && b'. +Admitted. +Goal is_expanded_env (Optimize.trim_ind_masks masksprovd.(ind_masks)) (Optimize.trim_const_masks masksprovd.(const_masks)) provedCop_before_checks.1 = true. +Proof. now lazy. Qed. + +Goal valid_masks_env masksprovd.(ind_masks) (Optimize.trim_const_masks masksprovd.(const_masks)) provedCop_before_checks.1 = true. +Proof. + unfold valid_masks_env. lazy. + eapply forallb_cons. now lazy. + eapply forallb_cons. now lazy. + eapply forallb_cons. + set (trm := Optimize.trim_ind_masks (ind_masks masksprovd)). lazy in trm. + unfold valid_masks_decl. + lazy -[check_oib_masks]. + unfold check_oib_masks. + cbn -[] + + unfold is_expanded at 1. + cbn [is_expanded_aux]. + cbn [andb]. + eapply and_andb. + 2:{ lazy. } + + unfold get_ctor_mask. cbn. + + lazy. + eapply forallb_cons. now lazy. + eapply forallb_cons. now lazy. + eapply forallb_cons. now lazy. + + unfold provedCop_before_checks. + + cbn. + lazy. + + +Eval compute in compute_masks (fun _ : kername => None) false false provedCop_before_checks.1. + +Eval compute in check_dearging_precond default_dearging_config provedCop_before_checks.1 provedCop_before_checks.2. + +Eval compute in check_dearging_precond default_dearging_config provedCop_before_checks.1 provedCop_before_checks.2. + +MetaCoq Typed Erase provedCopyx. +Eval lazy in testty cbv_provedCopyx. + diff --git a/test-suite/erasure_test.v b/test-suite/erasure_test.v index e9cd59265..601963af6 100644 --- a/test-suite/erasure_test.v +++ b/test-suite/erasure_test.v @@ -18,6 +18,12 @@ Construct(Coq.Init.Datatypes.bool,0,0) *) MetaCoq Erase (exist _ 0 (eq_refl) : {x : nat | x = 0}). + +Definition test := (proj1_sig (exist (fun x => x = 0) 0 (eq_refl))). + +MetaCoq Typed Erase test. + + (* (* *) (* Environment is well-formed and exist nat (fun x : nat => eq nat x O) O (eq_refl nat O):sig nat (fun x : nat => eq nat x O) erases to: *) (* (fun f => f) (exist ∎ ∎ O ∎) *) From 3fbfc8e5935d0298c56d8b2aa34c175175ca29d7 Mon Sep 17 00:00:00 2001 From: nicolas tabareau Date: Tue, 13 Feb 2024 07:33:26 +0100 Subject: [PATCH 04/28] generalize verified_erasure_pipeline_lookup_env_in --- erasure-plugin/theories/ErasureCorrectness.v | 193 ++++++++++--------- 1 file changed, 100 insertions(+), 93 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index c2576d635..c820a968d 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -1619,12 +1619,81 @@ Section pipeline_theorem. Variable u : Universes.Instance.t. Variable args : list PCUICAst.term. + Variable Normalisation : (forall Σ, wf_ext Σ -> PCUICSN.NormalizationIn Σ). + + Lemma verified_erasure_pipeline_lookup_env_in kn decl (efl := EInlineProjections.switch_no_params all_env_flags) + {has_rel : has_tRel} {has_box : has_tBox} + T (typing: ∥PCUICTyping.typing Σ [] t T∥): + let Σ_t := (transform verified_erasure_pipeline (Σ, t) (precond _ _ _ _ expΣ expt typing _)).1 in + EGlobalEnv.lookup_env Σ_t kn = Some decl -> + exists decl', + PCUICAst.PCUICEnvironment.lookup_global (PCUICExpandLets.trans_global_decls + (PCUICAst.PCUICEnvironment.declarations Σ.1)) kn = Some decl' + /\ erase_decl_equal (fun decl => ERemoveParams.strip_inductive_decl (erase_mutual_inductive_body decl)) + decl decl'. + Proof. + Opaque compose. + unfold verified_erasure_pipeline. + repeat rewrite -transform_compose_assoc. + destruct_compose; intro. cbn. + destruct_compose; intro. cbn. + set (erase_program _ _). + unfold verified_lambdabox_pipeline. + repeat rewrite -transform_compose_assoc. + repeat (destruct_compose; intro). + unfold transform at 1. cbn -[transform]. + rewrite EConstructorsAsBlocks.lookup_env_transform_blocks. + set (EConstructorsAsBlocks.transform_blocks_decl _). + unfold transform at 1. cbn -[transform]. + unfold transform at 1. cbn -[transform]. + erewrite EInlineProjections.lookup_env_optimize. + 2: { + eapply EOptimizePropDiscr.remove_match_on_box_env_wf; eauto. + apply ERemoveParams.strip_env_wf. + unfold transform at 1; cbn -[transform]. + rewrite erase_global_deps_fast_spec. + eapply erase_global_deps_wf_glob. + intros ? He; now rewrite He. } + set (EInlineProjections.optimize_decl _). + unfold transform at 1. cbn -[transform]. + unfold transform at 1. cbn -[transform]. + erewrite EOptimizePropDiscr.lookup_env_remove_match_on_box. + 2: { + apply ERemoveParams.strip_env_wf. + unfold transform at 1. cbn -[transform]. + rewrite erase_global_deps_fast_spec. + eapply erase_global_deps_wf_glob. + intros ? He; now rewrite He. } + set (EOptimizePropDiscr.remove_match_on_box_decl _). + unfold transform at 1. cbn -[transform]. + unfold transform at 1. cbn -[transform]. + erewrite ERemoveParams.lookup_env_strip. + set (ERemoveParams.strip_decl _). + unfold transform at 1. cbn -[transform]. + rewrite erase_global_deps_fast_spec. + 2: { cbn. intros ? He. rewrite He. eauto. } + intro. + set (EAstUtils.term_global_deps _). + set (build_wf_env_from_env _ _). + set (EGlobalEnv.lookup_env _ _). + case_eq o. 2: { intros ?. inversion 1. } + intros decl' Heq. + unshelve epose proof + (Hlookup := lookup_env_in_erase_global_deps optimized_abstract_env_impl w t0 + _ kn _ Hyp0 decl' _ Heq). + { epose proof (wf_fresh_globals _ HΣ). clear - H8. + revert H8. cbn. set (Σ.1). induction 1; econstructor; eauto. + cbn. clear -H. induction H; econstructor; eauto. } + destruct Hlookup as [decl'' [? ?]]. exists decl''; split ; eauto. + cbn in H10. inversion H10. + now destruct decl' , decl''. + Qed. + + Variable typing : ∥PCUICTyping.typing Σ [] t (PCUICAst.mkApps (PCUICAst.tInd i u) args)∥. Variable fo : @PCUICFirstorder.firstorder_ind Σ (PCUICFirstorder.firstorder_env Σ) i. - Variable Normalisation : (forall Σ, wf_ext Σ -> PCUICSN.NormalizationIn Σ). - Variable Heval : ∥PCUICWcbvEval.eval Σ t v∥. Let Σ_t := (transform verified_erasure_pipeline (Σ, t) (precond _ _ _ _ expΣ expt typing _)).1. @@ -1632,6 +1701,35 @@ Section pipeline_theorem. Let Σ_v := (transform verified_erasure_pipeline (Σ, v) (precond2 _ _ _ _ expΣ expt typing _ _ Heval)).1. Let v_t := compile_value_box (PCUICExpandLets.trans_global_env Σ) v []. + + Lemma verified_erasure_pipeline_extends (efl := EInlineProjections.switch_no_params all_env_flags) + {has_rel : has_tRel} {has_box : has_tBox} : + EGlobalEnv.extends Σ_v Σ_t. + Proof. + unfold Σ_v, Σ_t. unfold verified_erasure_pipeline. + repeat (destruct_compose; intro). destruct typing as [typing'], Heval. + cbn [transform compose pcuic_expand_lets_transform] in *. + unfold run, time. + cbn [transform erase_transform] in *. + set (erase_program _ _). set (erase_program _ _). + eapply verified_lambdabox_pipeline_extends. + eapply extends_erase_pcuic_program; eauto. + 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'. } + assumption. + now eapply PCUICExpandLetsCorrectness.trans_axiom_free. + pose proof (PCUICExpandLetsCorrectness.expand_lets_sound typing'). + rewrite PCUICExpandLetsCorrectness.trans_mkApps in X. eapply X. + move: fo. clear. + { rewrite /PCUICFirstorder.firstorder_ind /=. + rewrite PCUICExpandLetsCorrectness.trans_lookup. + destruct PCUICAst.PCUICEnvironment.lookup_env => //. + destruct g => //=. + eapply PCUICExpandLetsCorrectness.trans_firstorder_mutind. + eapply PCUICExpandLetsCorrectness.trans_firstorder_env. } + Qed. + Lemma fo_v : PCUICFirstorder.firstorder_value Σ [] v. Proof. destruct typing, Heval. sq. @@ -1702,70 +1800,6 @@ Section pipeline_theorem. Import PCUICWfEnv. - Lemma verified_erasure_pipeline_lookup_env_in kn decl (efl := EInlineProjections.switch_no_params all_env_flags) {has_rel : has_tRel} {has_box : has_tBox} : - EGlobalEnv.lookup_env Σ_t kn = Some decl -> - exists decl', - PCUICAst.PCUICEnvironment.lookup_global (PCUICExpandLets.trans_global_decls - (PCUICAst.PCUICEnvironment.declarations Σ.1)) kn = Some decl' - /\ erase_decl_equal (fun decl => ERemoveParams.strip_inductive_decl (erase_mutual_inductive_body decl)) - decl decl'. - Proof. - Opaque compose. - unfold Σ_t, verified_erasure_pipeline. - repeat rewrite -transform_compose_assoc. - destruct_compose; intro. cbn. - destruct_compose; intro. cbn. - set (erase_program _ _). - unfold verified_lambdabox_pipeline. - repeat rewrite -transform_compose_assoc. - repeat (destruct_compose; intro). - unfold transform at 1. cbn -[transform]. - rewrite EConstructorsAsBlocks.lookup_env_transform_blocks. - set (EConstructorsAsBlocks.transform_blocks_decl _). - unfold transform at 1. cbn -[transform]. - unfold transform at 1. cbn -[transform]. - erewrite EInlineProjections.lookup_env_optimize. - 2: { - eapply EOptimizePropDiscr.remove_match_on_box_env_wf; eauto. - apply ERemoveParams.strip_env_wf. - unfold transform at 1; cbn -[transform]. - rewrite erase_global_deps_fast_spec. - eapply erase_global_deps_wf_glob. - intros ? He; now rewrite He. } - set (EInlineProjections.optimize_decl _). - unfold transform at 1. cbn -[transform]. - unfold transform at 1. cbn -[transform]. - erewrite EOptimizePropDiscr.lookup_env_remove_match_on_box. - 2: { - apply ERemoveParams.strip_env_wf. - unfold transform at 1. cbn -[transform]. - rewrite erase_global_deps_fast_spec. - eapply erase_global_deps_wf_glob. - intros ? He; now rewrite He. } - set (EOptimizePropDiscr.remove_match_on_box_decl _). - unfold transform at 1. cbn -[transform]. - unfold transform at 1. cbn -[transform]. - erewrite ERemoveParams.lookup_env_strip. - set (ERemoveParams.strip_decl _). - unfold transform at 1. cbn -[transform]. - rewrite erase_global_deps_fast_spec. - 2: { cbn. intros ? He. rewrite He. eauto. } - intro. - set (EAstUtils.term_global_deps _). - set (build_wf_env_from_env _ _). - set (EGlobalEnv.lookup_env _ _). - case_eq o. 2: { intros ?. inversion 1. } - intros decl' Heq. - unshelve epose proof - (Hlookup := lookup_env_in_erase_global_deps optimized_abstract_env_impl w t0 - _ kn _ Hyp0 decl' _ Heq). - { epose proof (wf_fresh_globals _ HΣ). clear - H8. - revert H8. cbn. set (Σ.1). induction 1; econstructor; eauto. - cbn. clear -H. induction H; econstructor; eauto. } - destruct Hlookup as [decl'' [? ?]]. exists decl''; split ; eauto. - cbn in H10. inversion H10. - now destruct decl' , decl''. - Qed. Lemma verified_erasure_pipeline_firstorder_evalue_block : firstorder_evalue_block Σ_v v_t. @@ -1822,33 +1856,6 @@ Section pipeline_theorem. simpl. unfold fo_evalue_map. rewrite eqtr. exact H1. Qed. - Lemma verified_erasure_pipeline_extends (efl := EInlineProjections.switch_no_params all_env_flags) {has_rel : has_tRel} {has_box : has_tBox} : - EGlobalEnv.extends Σ_v Σ_t. - Proof. - unfold Σ_v, Σ_t. unfold verified_erasure_pipeline. - repeat (destruct_compose; intro). destruct typing as [typing'], Heval. - cbn [transform compose pcuic_expand_lets_transform] in *. - unfold run, time. - cbn [transform erase_transform] in *. - set (erase_program _ _). set (erase_program _ _). - eapply verified_lambdabox_pipeline_extends. - eapply extends_erase_pcuic_program; eauto. - 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'. } - assumption. - now eapply PCUICExpandLetsCorrectness.trans_axiom_free. - pose proof (PCUICExpandLetsCorrectness.expand_lets_sound typing'). - rewrite PCUICExpandLetsCorrectness.trans_mkApps in X. eapply X. - move: fo. clear. - { rewrite /PCUICFirstorder.firstorder_ind /=. - rewrite PCUICExpandLetsCorrectness.trans_lookup. - destruct PCUICAst.PCUICEnvironment.lookup_env => //. - destruct g => //=. - eapply PCUICExpandLetsCorrectness.trans_firstorder_mutind. - eapply PCUICExpandLetsCorrectness.trans_firstorder_env. } - Qed. - Lemma verified_erasure_pipeline_theorem : ∥ eval (wfl := extraction_wcbv_flags) Σ_t t_t v_t ∥. Proof. From 53b5abf2acb385534a0ec6c1f301805bae315ead Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 13 Feb 2024 09:21:59 +0100 Subject: [PATCH 05/28] Fix test-suite file --- test-suite/erasure_live_test.v | 58 ---------------------------------- 1 file changed, 58 deletions(-) diff --git a/test-suite/erasure_live_test.v b/test-suite/erasure_live_test.v index 277658dc9..57c9e17b3 100644 --- a/test-suite/erasure_live_test.v +++ b/test-suite/erasure_live_test.v @@ -495,63 +495,5 @@ Definition provedCop_before_checks := Eval lazy in exintro_typed_er quoted_prove Import ETransform Optimize. -Definition masks := -Eval compute in Optimize.analyze_env (fun _ : kername => None) exintro_before_checks.1. - -Eval compute in (masks).(Optimize.ind_masks). - -Eval compute in Optimize.trim_ind_masks (masks).(Optimize.ind_masks). -Eval compute in Optimize.valid_masks_env ( Optimize.trim_ind_masks masks.(ind_masks)) masks.(const_masks) (exintro_before_checks.1). - - -Eval compute in compute_masks (fun _ : kername => None) false false exintro_before_checks.1. - -Definition masksprovd := analyze_env (fun _ : kername => None) provedCop_before_checks.1. - -Lemma forallb_cons {A} (p : A -> bool) hd tl : p hd -> forallb p tl -> forallb p (cons hd tl) . -Proof. Admitted. -Lemma and_andb (b b' : bool) : b -> b' -> b && b'. -Admitted. -Goal is_expanded_env (Optimize.trim_ind_masks masksprovd.(ind_masks)) (Optimize.trim_const_masks masksprovd.(const_masks)) provedCop_before_checks.1 = true. -Proof. now lazy. Qed. - -Goal valid_masks_env masksprovd.(ind_masks) (Optimize.trim_const_masks masksprovd.(const_masks)) provedCop_before_checks.1 = true. -Proof. - unfold valid_masks_env. lazy. - eapply forallb_cons. now lazy. - eapply forallb_cons. now lazy. - eapply forallb_cons. - set (trm := Optimize.trim_ind_masks (ind_masks masksprovd)). lazy in trm. - unfold valid_masks_decl. - lazy -[check_oib_masks]. - unfold check_oib_masks. - cbn -[] - - unfold is_expanded at 1. - cbn [is_expanded_aux]. - cbn [andb]. - eapply and_andb. - 2:{ lazy. } - - unfold get_ctor_mask. cbn. - - lazy. - eapply forallb_cons. now lazy. - eapply forallb_cons. now lazy. - eapply forallb_cons. now lazy. - - unfold provedCop_before_checks. - - cbn. - lazy. - - -Eval compute in compute_masks (fun _ : kername => None) false false provedCop_before_checks.1. - -Eval compute in check_dearging_precond default_dearging_config provedCop_before_checks.1 provedCop_before_checks.2. - -Eval compute in check_dearging_precond default_dearging_config provedCop_before_checks.1 provedCop_before_checks.2. - MetaCoq Typed Erase provedCopyx. Eval lazy in testty cbv_provedCopyx. - From 99c6e851a8f7f2cafc6ca8cccdd5433b7df8b464 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 13 Feb 2024 15:43:42 +0100 Subject: [PATCH 06/28] Erase function lemma (#1053) * Fix typed erasure calls to _not_ trim inductive masks * Comment WIP app_construct proof * Comment WIP erase_function proof * Prove that functions are preserved by compilation * Show that inhabitants of product types evaluate to functions through the erasure pipeline --- erasure-plugin/theories/ErasureCorrectness.v | 680 ++++++++++++++++++- 1 file changed, 667 insertions(+), 13 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index c820a968d..62d1f21a9 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -515,8 +515,354 @@ Definition compile_app f t := | EAst.tApp fn a => EAst.tApp (f fn) (f a) | _ => f t end. + Fixpoint forallbi {A} (f : nat -> A -> bool) n l := + match l with + | [] => true + | hd :: tl => f n hd && forallbi f (S n) tl + end. + +(* TODO Move, in optimizeCorrectness as well *) +Lemma forallbi_Alli {A} (f : nat -> A -> bool) n l : + Alli f n l <~> forallbi f n l. +Proof. + split. + - induction 1; cbn; auto. + - induction l in n |- *; cbn; auto. + * constructor. + * move/andP => [] hf hl. constructor; eauto. +Qed. + +Lemma forallb_mapi_forallbi {A B} p (f : nat -> A -> B) n l : + forallb p (mapi_rec f l n) <~> forallbi (fun i x => p (f i x)) n l. +Proof. + split. + - induction l in n |- *; cbn; auto. move/andP => [] ->; eauto. + - induction l in n |- *; cbn; auto. move/andP => [] ->; eauto. +Qed. + +Lemma forallb_forallbi {A} (f : A -> bool) l : + forallb f l = forallbi (fun i => f) 0 l. +Proof. + generalize 0. + induction l; simpl; auto => n. bool_congr. eauto. +Qed. + +Lemma forallbi_impl {A} (f g : nat -> A -> bool) n l : + (forall n x, f n x -> g n x) -> + forallbi f n l -> forallbi g n l. +Proof. + intros hfg. + induction l in n |- *; simpl; auto; move/andP => [] /hfg -> //=; eauto. +Qed. + + +Section PCUICInv. + Import PCUICAst PCUICAstUtils PCUICTyping PCUICInversion PCUICWcbvEval PCUICWellScopedCumulativity PCUICEtaExpand. + Import PCUIC.PCUICConversion PCUICSigmaCalculus PCUICExpandLetsCorrectness. + + Definition is_vdef (d : context_decl) := match d with {| decl_body := None |} => false | _ => true end. + + Lemma red_vdefs_tProd {Σ Γ ctx concl na dom codom} : + wf_ext Σ -> + forallb is_vdef ctx -> + Σ ;;; Γ ⊢ it_mkProd_or_LetIn ctx concl ⇝ tProd na dom codom -> + Σ ;;; Γ ⊢ expand_lets ctx concl ⇝ tProd na dom codom. + Proof. + intros wfΣ. revert ctx concl. apply: PCUICInduction.ctx_length_rev_ind => [concl|d ctx IH concl]. + - rewrite expand_lets_nil //. + - rewrite forallb_app /= andb_true_r => /andP[] hctx hdef. + rewrite it_mkProd_or_LetIn_app. destruct d as [na' [b|] ty] => //=. + cbn. rewrite expand_lets_vdef. + move/(invert_red_letin (cf := extraction_checker_flags) (wfΣ := wfΣ)) => [[d' [ty' [b' []]]]|] //. + rewrite [T.subst10 _ _]PCUICLiftSubst.subst_it_mkProd_or_LetIn Nat.add_0_r. + eapply IH. now len. clear -hctx. + rewrite subst_context_alt. + { rewrite forallb_mapi_forallbi. rewrite forallb_forallbi in hctx. + eapply forallbi_impl; tea. + intros n [na [b'|] ty]; cbn; eauto. auto. } + Qed. -Module ETransformPresApp. + (* Constructors are impossible as they are eta expanded. Primitives can't have type Prod. + Fix and CoFix must be unreducible (that's the value hypothesis). *) + Lemma closed_prod_type {Σ : global_env_ext} {f na A B} {wfΣ : wf_ext Σ} + (wf : Σ ;;; [] |- f : PCUICAst.tProd na A B) : + value Σ f -> + expanded Σ [] f -> + match decompose_app f with + | (tLambda na ty b, []) => True + | (tInd i u, args) => True + | (tFix mfix idx, args) => True + | (tCoFix mfix idx, args) => True + | _ => False + end. + Proof. + induction 1 using value_values_ind. + - destruct t0 => //. + + eapply inversion_Sort in wf as [? []]; eauto. + now eapply ws_cumul_pb_Sort_Prod_inv in w0. + + eapply inversion_Prod in wf as [? [? [? []]]]; eauto. + now eapply ws_cumul_pb_Sort_Prod_inv in w. + +(* Constructor case: due to eta expansion *) + cbn. intros exp. depelim exp. + now rewrite nth_error_nil in H0. + destruct f; solve_discr. noconf H2. now cbn in H0. + solve_discr. solve_discr. + cbn in H1. assert (context_assumptions (cstr_args cdecl) = 0). lia. + assert (ind_npars mind = 0). lia. + eapply inversion_Construct in wf as [mdecl [idecl' [cdecl' [wf [hd [cu hty]]]]]]; eauto. + unshelve eapply declared_constructor_to_gen in hd, H0; tea; eauto. 3,6:eapply wfΣ. + eapply PCUICGlobalEnv.declared_constructor_inj in H0; tea. destruct H0 as [? []]; subst. + unfold type_of_constructor in hty. + eapply declared_constructor_from_gen in hd. + destruct (PCUICWeakeningEnvTyp.on_declared_constructor hd) as [[onmib onind] [univs [[]]]]. + destruct onmib, o. rewrite H4 in onNpars. + cbn in on_lets_in_type. + move: hty. rewrite cstr_eq. + rewrite PCUICUnivSubst.subst_instance_it_mkProd_or_LetIn. + rewrite PCUICLiftSubst.subst_it_mkProd_or_LetIn. + move/ws_cumul_pb_Prod_r_inv => [na' [dom' [codom' []]]] hred eqb eqdom cumcod; move: hred. + set (pars := subst_context _ _ _). + rewrite PCUICUnivSubst.subst_instance_it_mkProd_or_LetIn. + rewrite PCUICLiftSubst.subst_it_mkProd_or_LetIn. + set (args := subst_context _ _ _). + assert (forallb is_vdef pars). + { clear -onNpars. subst pars. + generalize 0. + induction (ind_params mdecl) using rev_ind; cbn; intros n'. now rewrite subst_context_nil. + rewrite PCUICUnivSubstitutionConv.subst_instance_app. + rewrite context_assumptions_app /= in onNpars. + destruct x as [na [d|] ty]; cbn in * => //; try lia. cbn. + rewrite PCUICSubstitution.subst_context_app' /= forallb_app. + rewrite IHc //. lia. } + assert (forallb is_vdef args). + { clear -H3. subst args. + generalize 0. + induction (cstr_args cdecl) using rev_ind; cbn; intros n'. now rewrite subst_context_nil. + rewrite PCUICUnivSubstitutionConv.subst_instance_app. + rewrite context_assumptions_app /= in H3. + destruct x as [na [d|] ty]; cbn in * => //; try lia. cbn. + rewrite PCUICSubstitution.subst_context_app' /= forallb_app. cbn. + rewrite -Nat.add_assoc. + rewrite IHc //. lia. } + rewrite -it_mkProd_or_LetIn_app. + move/red_vdefs_tProd. rewrite forallb_app H0 H5 //=. + rewrite /cstr_concl. len. rewrite PCUICSubstitution.subst_cstr_concl_head. + destruct hd. destruct H6. now eapply nth_error_Some_length in H8. + rewrite expand_lets_mkApps //=. cbn. + intros h. specialize (h eq_refl). + eapply invert_red_mkApps_tInd in h as [args' []]. solve_discr. + - eapply inversion_Prim in wf as [prim_ty [cdecl []]]; eauto. + now eapply invert_cumul_prim_type_prod in w; tea. + - clear X1. depelim X. + move/expanded_mkApps_inv'. + rewrite arguments_mkApps head_mkApps /=. intros [expargs [mind' [idecl' [cdecl' []]]]]. + unshelve eapply declared_constructor_to_gen in H0. 3:eapply wfΣ. + eapply PCUICGlobalEnv.declared_constructor_inj in d; tea. destruct d as [? []]; subst. + rewrite /cstr_arity in l. + all:rewrite decompose_app_mkApps //. + eapply PCUICValidity.inversion_mkApps in wf as (? & ? & ?). + eapply inversion_Construct in t0 as [mdecl [idecl'' [cdecl' [wf [hd [cu hty]]]]]]; eauto. + eapply PCUICSpine.typing_spine_strengthen in t1. 3:tea. + 2:{ eapply PCUICValidity.validity. econstructor; eauto. } + clear hty. + unfold type_of_constructor in t1. + unshelve eapply declared_constructor_to_gen in hd. 3:apply wfΣ. + eapply (PCUICGlobalEnv.declared_constructor_inj H0) in hd as [? []]; subst. + eapply declared_constructor_from_gen in H0. + destruct (PCUICWeakeningEnvTyp.on_declared_constructor H0) as [[onmib onind] [univs [[]]]]. + destruct onmib, o. + cbn in on_lets_in_type. + move: t1. rewrite cstr_eq. + rewrite PCUICUnivSubst.subst_instance_it_mkProd_or_LetIn. + rewrite PCUICLiftSubst.subst_it_mkProd_or_LetIn. + set (pars := subst_context _ _ _). + rewrite PCUICUnivSubst.subst_instance_it_mkProd_or_LetIn. + rewrite PCUICLiftSubst.subst_it_mkProd_or_LetIn. + set (args' := subst_context _ _ _). + (* move/ws_cumul_pb_Prod_r_inv => [na' [dom' [codom' []]]] hred eqb eqdom cumcod; move: hred. *) + rewrite /cstr_concl. len. rewrite PCUICSubstitution.subst_cstr_concl_head. + destruct H0. destruct H0. now eapply nth_error_Some_length in H3. + rewrite -it_mkProd_or_LetIn_app. + rewrite -onNpars in l, H1. + move/PCUICClassification.typing_spine_smash. + rewrite expand_lets_mkApps. + have lenargs : #|args| = context_assumptions (ind_params mind') + context_assumptions (cstr_args cdecl') by lia. + move/(PCUICClassification.typing_spine_all_inv _). + intros h; forward h. rewrite context_assumptions_smash_context /= context_assumptions_app. + rewrite /args' /pars. now len. + rewrite expand_lets_mkApps PCUICLiftSubst.subst_mkApps /= in h. + move: h => [] h _; move: h. + move/ws_cumul_pb_Prod_r_inv => [na' [dom' [codom' []]]] hred eqb eqdom cumcod; move: hred. + move/invert_red_mkApps_tInd => [args'' [heq]]. solve_discr. + Qed. + + Lemma pcuic_eval_function `{WcbvFlags} (cf:=config.extraction_checker_flags) (Σ:global_env_ext_map) f v na A B + (wfΣ : wf_ext Σ) (wf : Σ ;;; [] |- f : PCUICAst.tProd na A B) : + expanded_global_env Σ -> + expanded Σ [] f -> + eval Σ f v -> + match decompose_app v with + | (tLambda na ty b, []) => True + | (tInd i u, args) => True + | (tFix mfix idx, args) => True + | (tCoFix mfix idx, args) => True + | _ => False + end. + Proof. + intros expΣ exp ev. + assert (expanded Σ [] v). + { eapply (expanded_red (Σ := Σ)); tea. exact wfΣ. + 2:eapply (PCUICClassification.wcbveval_red (Σ := Σ)); tea. + intros ??; now rewrite nth_error_nil. } + eapply PCUICClassification.subject_reduction_eval in wf; tea. + eapply eval_to_value in ev. + eapply closed_prod_type; tea. + Qed. + + Lemma nisErasable_eval {Σ : global_env_ext} {t t'} {wfΣ : wf Σ} : + nisErasable Σ [] t -> + eval Σ t t' -> + nisErasable Σ [] t'. + Proof. + intros [nf [s [hty norm nisa hs nonprop]]] ev. + eapply PCUICClassification.subject_reduction_eval in hty. 2:eassumption. + exists nf, s; split => //. + Qed. + +End PCUICInv. + +Section ErasureFunction. + Import PCUICAst PCUICAst.PCUICEnvironment PCUIC.PCUICConversion PCUICArities PCUICSpine PCUICOnFreeVars PCUICWellScopedCumulativity. + Import EAst EAstUtils EWcbvEval EArities. + + Definition isFunction (t : EAst.term) := EAst.isLambda t || isFixApp t || isCoFix (head t). + + Lemma typing_spine_isArity {Σ : global_env_ext} {Γ T args T'} {wfΣ : wf Σ} : + wf_local Σ Γ -> + typing_spine Σ Γ T args T' -> + isArity T -> + exists ar, ∥ typing_spine Σ Γ T args ar ∥ /\ isArity ar. + Proof. + intros wfΓ sp isa. + unshelve epose proof (PCUICClassification.isArity_typing_spine _ sp); eauto. + forward H. + { exists T. split; eauto. sq. + pose proof (typing_spine_isType_dom sp). + eapply PCUICContextConversion.closed_red_refl; fvs. } + destruct H as [ar [[redar] isa']]. exists ar; split => //. sq. + eapply typing_spine_weaken_concl; tea. now eapply red_ws_cumul_pb. + eapply typing_spine_isType_codom in sp. eapply isType_red; tea. + eapply redar. + Qed. + + Lemma erase_function (wfl := default_wcbv_flags) + {guard_impl : abstract_guard_impl} + (cf:=config.extraction_checker_flags) (Σ:global_env_ext_map) f v v' na A B + (wf : ∥ Σ ;;; [] |- f : PCUICAst.tProd na A B ∥) pr : + ∥ nisErasable Σ [] f ∥ -> + PCUICWcbvEval.eval Σ f v -> + let (Σ', f') := transform erase_transform (Σ, f) pr in + eval Σ' f' v' -> isFunction v' = true. + Proof. + intros [ne]. destruct wf as [hty]. + intros evf. + unfold transform, erase_transform. + destruct pr as [pre [[expΣ expt] [norm norm']]]. destruct pre as [[wfΣ wft]]. + eapply nisErasable_eval in ne; tea. + eapply nisErasable_spec in ne; eauto. + eapply PCUICClassification.subject_reduction_eval in hty as hty'; tea. + unfold erase_program, erase_pcuic_program. + set (Σ' := build_wf_env_from_env _ _). + set (env' := PCUICWfEnv.abstract_make_wf_env_ext _ _ _). + set (obl := (fun Σ (wfΣ : _) => _)) at 5. + set (obl' := (fun Σ => _)) at 12. + cbn -[env']. + set (wtf := (fun Σ => _)) at 13. change obl' with wtf. clear obl'. clearbody wtf. + set (eqdecls := (fun Σ => _)) at 9. clearbody eqdecls. + set (deps := term_global_deps _). + set (nin := (fun (n : nat) => _)). clearbody nin. + epose proof (@erase_global_deps_fast_erase_global_deps deps optimized_abstract_env_impl Σ' (PCUICAst.PCUICEnvironment.declarations Σ) nin) as [nin2 eq]. + rewrite /erase_global_fast. erewrite eq. clear eq nin. + set (eg := erase_global_deps _ _ _ _ _ _). + + unshelve epose proof (erase_correct optimized_abstract_env_impl Σ' Σ.2 _ f v _ _ deps _ _ _ eq_refl _ eq_refl _ Σ eq_refl); eauto. + { eapply Kernames.KernameSet.subset_spec. rewrite /deps -/env'. cbn [fst snd]. apply Kernames.KernameSetProp.subset_refl. } + { cbn => ? -> //. } + destruct H as [v'' [ervv'' [ev]]]. + eset (eg' := erase_global_deps _ _ _ _ _ _) in ev. + replace eg with eg'. 2:{ rewrite /eg /eg'. eapply reflexivity. } + intros ev'. + assert (v' = v''). { epose proof (eval_deterministic ev). symmetry. eapply H. + set(er := erase _ _ _ _ _) in ev'. + set(er' := erase _ _ _ _ _). + replace er' with er. exact ev'. + rewrite /er /er' /env'. apply reflexivity. } + subst v''. cbn in ev. + eapply pcuic_eval_function in evf; tea. + destruct (PCUICAstUtils.decompose_app v) eqn:da. + eapply PCUICAstUtils.decompose_app_inv in da. cbn in da. + move: evf. destruct t0 => //; cbn in da; subst v. 1:destruct l => //. 1-4:intros _. + - clear -ne ervv''. depelim ervv''. cbn => //. elim ne. sq. exact X. + - clear -wfΣ hty' ne. elim ne. + assert (exists mdecl idecl, PCUICAst.declared_inductive Σ ind mdecl idecl) as [mdecl [idecl decli]]. + { eapply PCUICValidity.inversion_mkApps in hty' as [? [hty _]]. clear -hty. + depind hty; eauto. } + eapply PCUICInductives.invert_type_mkApps_ind in hty' as [sp cu]; eauto. + pose proof (typing_spine_isArity ltac:(constructor) sp). + forward H. + { apply (PCUICUnivSubstitutionTyp.isArity_subst_instance ui). + now eapply isArity_ind_type. } + destruct H as [ar [[spar] isa']]. + sq. + eexists; cbn. split. + eapply PCUICSpine.type_mkApps. econstructor; eauto with pcuic. exact spar. + now left. + - clear -wfΣ hty' ne ervv''. + assert (exists mfix' l', v' = EAst.mkApps (EAst.tFix mfix' idx) l') as [mfix' [l' ->]]. + { eapply ErasureProperties.erases_mkApps_inv in ervv'' as [[L1 [L2 [L2' [? [? [? []]]]]]]|[f'[L'[eq [erf erargs]]]]]. + - subst l v'. elim ne. destruct H0. rewrite PCUICAstUtils.mkApps_app. eapply Is_type_app; [eauto|eauto| |eauto]. + now rewrite -PCUICAstUtils.mkApps_app. + - depelim erf. do 2 eexists; trea. subst v'. + elim ne. eapply Is_type_app; eauto. } + now rewrite /isFunction /isFixApp !head_mkApps //= orb_true_r. + - clear -wfΣ hty' ne ervv''. + assert (exists mfix' l', v' = EAst.mkApps (EAst.tCoFix mfix' idx) l') as [mfix' [l' ->]]. + { eapply ErasureProperties.erases_mkApps_inv in ervv'' as [[L1 [L2 [L2' [? [? [? []]]]]]]|[f'[L'[eq [erf erargs]]]]]. + - subst l v'. elim ne. destruct H0. rewrite PCUICAstUtils.mkApps_app. eapply Is_type_app; [eauto|eauto| |eauto]. + now rewrite -PCUICAstUtils.mkApps_app. + - depelim erf. do 2 eexists; trea. subst v'. + elim ne. eapply Is_type_app; eauto. } + now rewrite /isFunction /isFixApp !head_mkApps //= orb_true_r. + Qed. + + Lemma pcuic_function_value (wfl := default_wcbv_flags) + {guard_impl : abstract_guard_impl} + (cf:=config.extraction_checker_flags) {Σ : global_env_ext} {f na A B} + (wfΣ : wf_ext Σ) (axfree : axiom_free Σ) (wf : Σ ;;; [] |- f : PCUICAst.tProd na A B) : { v & PCUICWcbvEval.eval Σ f v }. + Proof. + eapply (PCUICNormalization.wcbv_normalization wfΣ axfree wf). Unshelve. + exact PCUICSN.extraction_normalizing. + now eapply PCUICSN.normalization. + Qed. + + Lemma erase_function_to_function (wfl := default_wcbv_flags) + {guard_impl : abstract_guard_impl} + (cf:=config.extraction_checker_flags) (Σ:global_env_ext_map) f v' na A B + (wf : ∥ Σ ;;; [] |- f : PCUICAst.tProd na A B ∥) pr : + axiom_free Σ -> + ∥ nisErasable Σ [] f ∥ -> + let (Σ', f') := transform erase_transform (Σ, f) pr in + eval Σ' f' v' -> isFunction v' = true. + Proof. + intros axfree nise. + destruct pr as [[]]. destruct wf. + epose proof (pcuic_function_value w.1 axfree X) as [v hv]. + eapply erase_function; tea. now sq. + Qed. + +End ErasureFunction. + +Module ETransformPresAppLam. Section Opt. Context {env env' : Type}. Context {eval : program env EAst.term -> EAst.term -> Prop}. @@ -534,7 +880,8 @@ Module ETransformPresApp. exists prt pru, o.(transform) (Σ, EAst.tApp t u) pr = ((o.(transform) (Σ, EAst.tApp t u) pr).1, - EAst.tApp (o.(transform) (Σ, t) prt).2 (o.(transform) (Σ, u) pru).2) }. + EAst.tApp (o.(transform) (Σ, t) prt).2 (o.(transform) (Σ, u) pru).2); + transform_lam : forall p (pr : o.(pre) p), isFunction p.2 -> isFunction (o.(transform) p pr).2 }. End Opt. Section Comp. @@ -603,10 +950,11 @@ Module ETransformPresApp. now apply oext.(transform_env_irrel _ _ _). destruct (transform o _ pr'') => //. cbn. intros pru. now eapply transform_pre_irrel. + - intros. cbn; unfold run, time. now eapply o'ext.(transform_lam _ _ _), oext.(transform_lam _ _ _). Qed. End Comp. -End ETransformPresApp. +End ETransformPresAppLam. Import EWellformed. @@ -774,7 +1122,7 @@ Definition is_eta_fix_app (p : program E.global_context EAst.term) : Prop := EEt Definition is_eta_fix_app_map (p : program EEnvMap.GlobalContextMap.t EAst.term) : Prop := EEtaExpandedFix.isEtaExp p.1 [] p.2. #[global] Instance rebuild_wf_env_transform_pres_app {fl : WcbvFlags} {efl : EEnvFlags} we : - ETransformPresApp.t + ETransformPresAppLam.t (rebuild_wf_env_transform we false) is_eta_app is_eta_app_map. Proof. split => //. - intros. unfold transform, rebuild_wf_env_transform. @@ -832,7 +1180,7 @@ Import EAstUtils. #[global] Instance inline_projections_optimization_pres_app {fl : WcbvFlags} (efl := EInlineProjections.switch_no_params all_env_flags) {wcon : with_constructor_as_block = false} {has_rel : has_tRel} {has_box : has_tBox} : - ETransformPresApp.t + ETransformPresAppLam.t (inline_projections_optimization (wcon := wcon) (hastrel := has_rel) (hastbox := has_box)) is_eta_app_map is_eta_app. Proof. @@ -853,6 +1201,14 @@ Proof. now move/andP: H0 => []. move/andP: H1 => [] etactx etaapp. apply/andP => //. split => //. now apply EEtaExpanded.isEtaExp_tApp_arg in etaapp. reflexivity. + - intros [Σ t] pr; cbn. clear pr. move: t. + apply: EInduction.MkAppsInd.rec => //= t u. + rewrite /isFunction /isFixApp head_mkApps EInlineProjections.optimize_mkApps !head_mkApps; rtoProp; intuition auto. + destruct u using rev_case => //. rewrite ?map_app !mkApps_app /= in H2 *. + rewrite -!orb_assoc in H1. + forward H1. apply/or3P. move/orP in H2; intuition auto. now constructor 2. + now constructor 3. + apply/orP. move/or3P: H1 => []; intuition auto. destruct t => //. Qed. #[global] Instance remove_match_on_box_pres {fl : WcbvFlags} {efl : EEnvFlags} {wcon : with_constructor_as_block = false} @@ -882,7 +1238,7 @@ Qed. #[global] Instance remove_match_on_box_pres_app {fl : WcbvFlags} {efl : EEnvFlags} {wcon : with_constructor_as_block = false} {has_rel : has_tRel} {has_box : has_tBox} : has_cstr_params = false -> - ETransformPresApp.t + ETransformPresAppLam.t (remove_match_on_box_trans (wcon := wcon) (hastrel := has_rel) (hastbox := has_box)) is_eta_app_map is_eta_app. Proof. @@ -904,6 +1260,15 @@ Proof. { destruct pr as [[] pr']; move/andP: pr' => [] etactx; split => //. split => //. cbn in H0. now move/andP: H0 => [] /andP []. apply/andP. split => //. now apply EEtaExpanded.isEtaExp_tApp_arg in b. } now rewrite /EOptimizePropDiscr.remove_match_on_box_program /=. + - intros [Σ t] pr; cbn. + clear -t. revert t. + apply: EInduction.MkAppsInd.rec => //= t u. + rewrite /isFunction /isFixApp head_mkApps EOptimizePropDiscr.remove_match_on_box_mkApps !head_mkApps; rtoProp; intuition auto. + destruct u using rev_case => //. rewrite ?map_app !mkApps_app /= in H2 *. + rewrite -!orb_assoc in H1. + forward H1. apply/or3P. move/orP in H2; intuition auto. now constructor 2. + now constructor 3. + apply/orP. move/or3P: H1 => []; intuition auto. destruct t => //. Qed. #[global] Instance remove_params_optimization_pres {fl : WcbvFlags} {wcon : with_constructor_as_block = false} : @@ -924,7 +1289,7 @@ Proof. split => //. Qed. #[global] Instance remove_params_optimization_pres_app {fl : WcbvFlags} {wcon : with_constructor_as_block = false} : - ETransformPresApp.t + ETransformPresAppLam.t (remove_params_optimization (wcon := wcon)) is_eta_app_map is_eta_app. Proof. @@ -945,6 +1310,18 @@ Proof. apply/andP. split => //. now apply EEtaExpanded.isEtaExp_tApp_arg in b. } rewrite /ERemoveParams.strip_program /=. f_equal. rewrite (ERemoveParams.strip_mkApps_etaexp _ [u]) //. + - intros [Σ t] pr. cbn [fst snd transform remove_params_optimization]. + clear -t. revert t. + apply: EInduction.MkAppsInd.rec => //= t u. + rewrite /isFunction /isFixApp !head_mkApps. + intros napp nnil. + intros IH IHargs isl. + rewrite ERemoveParams.strip_mkApps //. + destruct EEtaExpanded.construct_viewc => //=; cbn in isl; + rewrite OptimizeCorrectness.isLambda_mkApps //= in isl *. + rewrite OptimizeCorrectness.isLambda_mkApps. len. + rewrite !head_mkApps. + destruct t0 => //=. Qed. #[global] Instance constructors_as_blocks_transformation_pres {efl : EWellformed.EEnvFlags} @@ -972,7 +1349,7 @@ Qed. #[global] Instance constructors_as_blocks_transformation_pres_app {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} : - ETransformPresApp.t + ETransformPresAppLam.t (@constructors_as_blocks_transformation efl has_app has_rel hasbox has_pars has_cstrblocks) is_eta_app_map (fun _ => True). Proof. @@ -988,6 +1365,20 @@ Proof. apply/andP. split => //. now apply EEtaExpanded.isEtaExp_tApp_arg in exp'. } simpl. rewrite /EConstructorsAsBlocks.transform_blocks_program /=. f_equal. rewrite (EConstructorsAsBlocks.transform_blocks_mkApps_eta_fn _ _ [u]) //. + - intros [Σ t] pr; cbn [fst snd transform constructors_as_blocks_transformation]. + destruct pr as [_ h]. move/andP: h => [] _ /=. + clear -t. + destruct (decompose_app t) as [f l] eqn:da. + pose proof (decompose_app_notApp _ _ _ da). + eapply decompose_app_inv in da. subst t. + rewrite /isFunction /isFixApp !head_mkApps. + rewrite OptimizeCorrectness.isLambda_mkApps //=. + rewrite EEtaExpanded.isEtaExp_mkApps_napp //=. + destruct EEtaExpanded.construct_viewc => //. + move/andP => [etat etal]. + rewrite !(EConstructorsAsBlocks.transform_blocks_mkApps_eta_fn _ _ l) // !head_mkApps /=. + destruct t0 => //=. rewrite !orb_false_r. move/Nat.eqb_eq. + destruct l => //=. all:now rewrite !orb_false_r !orb_true_r. Qed. #[global] Instance guarded_to_unguarded_fix_pres {efl : EWellformed.EEnvFlags} @@ -1002,7 +1393,7 @@ Qed. #[global] Instance guarded_to_unguarded_fix_pres_app {efl : EWellformed.EEnvFlags} {has_guard : with_guarded_fix} {has_cstrblocks : with_constructor_as_block = false} : - ETransformPresApp.t + ETransformPresAppLam.t (@guarded_to_unguarded_fix default_wcbv_flags has_cstrblocks efl has_guard) is_eta_fix_app_map is_eta_app_map. Proof. @@ -1044,15 +1435,30 @@ Proof. Qed. #[local] Instance lambdabox_pres_app : - ETransformPresApp.t verified_lambdabox_pipeline is_eta_fix_app_map (fun _ => True). + ETransformPresAppLam.t verified_lambdabox_pipeline is_eta_fix_app_map (fun _ => True). Proof. unfold verified_lambdabox_pipeline. - do 5 (unshelve eapply ETransformPresApp.compose; [shelve| |tc]). + do 5 (unshelve eapply ETransformPresAppLam.compose; [shelve| |tc]). 2:{ eapply remove_match_on_box_pres_app => //. } - do 2 (unshelve eapply ETransformPresApp.compose; [shelve| |tc]). + do 2 (unshelve eapply ETransformPresAppLam.compose; [shelve| |tc]). tc. Qed. + +Lemma expand_lets_function (wfl := default_wcbv_flags) + {guard_impl : abstract_guard_impl} + (cf:=config.extraction_checker_flags) (Σ:global_env_ext_map) K f na A B + (wf : ∥ Σ ;;; [] |- f : PCUICAst.tProd na A B ∥) pr : + let (Σ', f') := transform (pcuic_expand_lets_transform K) (Σ, f) pr in + ∥ Σ' ;;; [] |- f' : PCUICExpandLets.trans (PCUICAst.tProd na A B) ∥. +Proof. + unfold transform, pcuic_expand_lets_transform. cbn. + destruct pr as [[[]] ?]. + sq. + eapply PCUICExpandLetsCorrectness.pcuic_expand_lets in wf; eauto. + now eapply PCUICExpandLetsCorrectness.trans_wf. +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) []). @@ -1072,7 +1478,7 @@ Lemma transform_lambda_box_eta_app (Σer : EEnvMap.GlobalContextMap.t) t u pre : (transform verified_lambdabox_pipeline (Σer, u) pre'').2). Proof. intros etat. - epose proof (ETransformPresApp.transform_app verified_lambdabox_pipeline is_eta_fix_app_map (fun _ => True) Σer t u pre etat). + epose proof (ETransformPresAppLam.transform_app verified_lambdabox_pipeline is_eta_fix_app_map (fun _ => True) Σer t u pre etat). exact H. Qed. @@ -1419,6 +1825,119 @@ Section PCUICErase. now eapply (PCUICExpandLetsCorrectness.expand_lets_sound (cf := extraction_checker_flags)) in Hsort. Qed. + #[local] Existing Instance PCUICSN.extraction_normalizing. + + (* Beware, this internally uses preservation of observations and determinism of evaluation + from the canonical evaluation of [f] in the source to the evaluation in the target. + *) + Lemma transform_erasure_pipeline_function + (wfl := default_wcbv_flags) + {guard_impl : abstract_guard_impl} + (cf:=config.extraction_checker_flags) (Σ:global_env_ext_map) + {f v' na A B} + (wf : ∥ Σ ;;; [] |- f : PCUICAst.tProd na A B ∥) pr : + axiom_free Σ -> + ∥ nisErasable Σ [] f ∥ -> + let tr := transform verified_erasure_pipeline (Σ, f) pr in + eval (wfl := extraction_wcbv_flags) tr.1 tr.2 v' -> isFunction v' = true. + Proof. + intros axfree nise. + unfold verified_erasure_pipeline. + rewrite -!transform_compose_assoc. + pose proof (expand_lets_function Σ (fun p : global_env_ext_map => + (wf_ext p -> PCUICSN.NormalizationIn p) /\ + (wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p)) f na A B wf pr). + destruct_compose. intros pre. + set (trexp := transform (pcuic_expand_lets_transform _) _ _) in *. + eapply (PCUICExpandLetsCorrectness.trans_axiom_free Σ) in axfree. + have nise' : ∥ nisErasable trexp.1 [] trexp.2 ∥. + destruct pr as [[[]] ?], nise. sq; now eapply nisErasable_lets. + change (trans_global_env _) with (global_env_ext_map_global_env_ext trexp.1).1 in axfree. + clearbody trexp. clear nise pr wf Σ f. destruct trexp as [Σ f]. + pose proof pre as pre'; destruct pre' as [[[wf _]] _]. + pose proof (map_squash (pcuic_function_value wf axfree) H) as [[v ev]]. + epose proof (Transform.preservation erase_transform). + specialize (H0 _ v pre (sq ev)). + revert H0. + destruct_compose. intros pre' htr. + destruct htr as [v'' [ev' _]]. + epose proof (erase_function_to_function _ f v'' _ _ _ H pre axfree nise'). + set (tre := transform erase_transform _ _) in *. clearbody tre. + cbn -[transform obseq]. + intros ev2. red in ev'. destruct ev'. + epose proof (Transform.preservation verified_lambdabox_pipeline). + destruct tre as [Σ' f']. + specialize (H2 _ v'' pre' (sq H1)) as [finalv [[evfinal] obseq]]. + pose proof (eval_deterministic evfinal ev2). subst v'. + have prev : Transform.pre verified_lambdabox_pipeline (Σ', v''). + { clear -wfl pre' H1. cbn in H1. + destruct pre' as [[] []]. split; split => //=. + eapply Prelim.Ee.eval_wellformed; eauto. + eapply EEtaExpandedFix.isEtaExp_expanded. + eapply (@EEtaExpandedFix.eval_etaexp wfl); eauto. + now eapply EEtaExpandedFix.expanded_global_env_isEtaExp_env. + now eapply EEtaExpandedFix.expanded_isEtaExp. } + specialize (H0 H1). + eapply (obseq_lambdabox (Σ', f') (Σ', v'')) in obseq. + epose proof (ETransformPresAppLam.transform_lam _ _ _ (t := lambdabox_pres_app) (Σ', v'') prev H0). + rewrite -obseq. exact H2. cbn. red; tauto. + Qed. + + (* This version provides the evaluation proof. *) + Lemma transform_erasure_pipeline_function' + (wfl := default_wcbv_flags) + {guard_impl : abstract_guard_impl} + (cf:=config.extraction_checker_flags) (Σ:global_env_ext_map) + {f na A B} + (wf : ∥ Σ ;;; [] |- f : PCUICAst.tProd na A B ∥) pr : + axiom_free Σ -> + ∥ nisErasable Σ [] f ∥ -> + let tr := transform verified_erasure_pipeline (Σ, f) pr in + exists v, ∥ eval (wfl := extraction_wcbv_flags) tr.1 tr.2 v ∥ /\ isFunction v = true. + Proof. + intros axfree nise. + unfold verified_erasure_pipeline. + rewrite -!transform_compose_assoc. + pose proof (expand_lets_function Σ (fun p : global_env_ext_map => + (wf_ext p -> PCUICSN.NormalizationIn p) /\ + (wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p)) f na A B wf pr). + destruct_compose. intros pre. + set (trexp := transform (pcuic_expand_lets_transform _) _ _) in *. + eapply (PCUICExpandLetsCorrectness.trans_axiom_free Σ) in axfree. + have nise' : ∥ nisErasable trexp.1 [] trexp.2 ∥. + destruct pr as [[[]] ?], nise. sq; now eapply nisErasable_lets. + change (trans_global_env _) with (global_env_ext_map_global_env_ext trexp.1).1 in axfree. + clearbody trexp. clear nise pr wf Σ f. destruct trexp as [Σ f]. + pose proof pre as pre'; destruct pre' as [[[wf _]] _]. + pose proof (map_squash (pcuic_function_value wf axfree) H) as [[v ev]]. + epose proof (Transform.preservation erase_transform). + specialize (H0 _ v pre (sq ev)). + revert H0. + destruct_compose. intros pre' htr. + destruct htr as [v'' [ev' _]]. + epose proof (erase_function_to_function _ f v'' _ _ _ H pre axfree nise'). + set (tre := transform erase_transform _ _) in *. clearbody tre. + cbn -[transform obseq]. + red in ev'. destruct ev'. + epose proof (Transform.preservation verified_lambdabox_pipeline). + destruct tre as [Σ' f']. + specialize (H2 _ v'' pre' (sq H1)) as [finalv [[evfinal] obseq]]. + exists finalv. + split. now sq. + have prev : Transform.pre verified_lambdabox_pipeline (Σ', v''). + { clear -wfl pre' H1. cbn in H1. + destruct pre' as [[] []]. split; split => //=. + eapply Prelim.Ee.eval_wellformed; eauto. + eapply EEtaExpandedFix.isEtaExp_expanded. + eapply (@EEtaExpandedFix.eval_etaexp wfl); eauto. + now eapply EEtaExpandedFix.expanded_global_env_isEtaExp_env. + now eapply EEtaExpandedFix.expanded_isEtaExp. } + specialize (H0 H1). + eapply (obseq_lambdabox (Σ', f') (Σ', v'')) in obseq. + epose proof (ETransformPresAppLam.transform_lam _ _ _ (t := lambdabox_pres_app) (Σ', v'') prev H0). + rewrite -obseq. exact H2. cbn. red; tauto. + Qed. + Lemma expand_lets_transform_env K p p' pre pre' : p.1 = p'.1 -> (transform (pcuic_expand_lets_transform K) p pre).1 = @@ -1535,6 +2054,141 @@ Section PCUICErase. { red. cbn. split => //. } reflexivity. Qed. +Transparent erase_transform. + + + (* Lemma erase_transform_extends_app_construct {X_type} {X} {normalization_in} ind i u args wt : + let t := (mkApps (tConstruct ind i u) args) in + let trt := @erase_global_fast X_type X normalization_in [] t wt in + Forall (fun arg => exists wtarg, + let trarg := @erase X_type X normalization_in [] t wtarg in + EGlobalEnv.extends trarg.1 trt.1) args. + Proof. + intros t. + cbn. rewrite /erase_transform /=. + set (deps := term_global_deps _). + subst t. *) + + + (* Lemma erase_transform_extends_app_construct (Σ : global_env_ext_map) ind i u args pre : + let t := (mkApps (tConstruct ind i u) args) in + let trt := transform erase_transform (Σ, t) pre in + Forall (fun arg => exists pre', + let trarg := transform erase_transform (Σ, arg) pre' in + EGlobalEnv.extends trarg.1 trt.1) args. + Proof. + intros t. + cbn. rewrite /erase_transform /=. + set (deps := term_global_deps _). + subst t. + + induction args. constructor. constructor; eauto. + intros H. + assert (ner' : ~ ∥ isErasable Σ [] t ∥). + { destruct ner as [ner]. destruct pre, s. eapply nisErasable_spec in ner => //. eapply w. } + set (tr := transform _ _ _). + destruct tr eqn:heq. cbn -[transform]. + + Lemma erasure_pipeline_extends_app_construct (Σ : global_env_ext_map) ind i u args pre : + let t := (mkApps (tConstruct ind i u) args) in + ∥ nisErasable Σ [] t ∥ -> + PCUICEtaExpand.expanded Σ [] t -> + let trt := transform verified_erasure_pipeline (Σ, t) pre in + Forall (fun arg => exists pre', + let trarg := transform verified_erasure_pipeline (Σ, arg) pre' in + EGlobalEnv.extends trarg.1 trt.1) args. + Proof. + intros t ner exp. + unfold verified_erasure_pipeline. + destruct_compose. + set (K:= (fun p : global_env_ext_map => (wf_ext p -> PCUICSN.NormalizationIn p) /\ (wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p))). + intros H. + assert (ner' : ~ ∥ isErasable Σ [] t ∥). + { destruct ner as [ner]. destruct pre, s. eapply nisErasable_spec in ner => //. eapply w. } + set (tr := transform _ _ _). + destruct tr eqn:heq. cbn -[transform]. + replace t0 with tr.2. assert (heq_env:tr.1=g) by now rewrite heq. subst tr. + 2:{ now rewrite heq. } + clear heq. + epose proof ( + erase_eta_app _ _ _ pre3) as H0. + pose proof (correctness (pcuic_expand_lets_transform K) (Σ, tApp t u) pre). + destruct H as [[wtapp] [expapp Kapp]]. + pose proof (correctness (pcuic_expand_lets_transform K) (Σ, t) pre'). + destruct H as [[wtt] [expt Kt]]. + forward H0. + { clear -wtapp ner eq. apply (f_equal snd) in eq. cbn [snd] in eq. rewrite -eq. + destruct pre as [[wtp] rest]. + destruct ner as [ner]. eapply (nisErasable_lets (Σ, tApp t u)) in ner. + eapply nisErasable_spec in ner => //. cbn. + apply wtapp. apply wtp. } + forward H0 by apply expt. + destruct H0 as [pre'0 [pre''0 [eta [extapp [extapp' heq]]]]]. + split. + { rewrite <- heq_env. cbn -[transform]. + pose proof (EProgram.TransformExt.preserves_obs _ _ _ (t:=verified_lambdabox_pipeline_extends')). + unfold extends_eprogram in H. + split. + { repeat (destruct_compose; intros). eapply verified_lambdabox_pipeline_extends. + repeat (destruct_compose; intros). cbn - [transform]. + generalize dependent pre3. rewrite <- eq. + cbn [transform pcuic_expand_lets_transform expand_lets_program]. + unfold expand_lets_program. cbn [fst snd]. + intros pre3. cbn in pre3. intros <-. intros. + assert (pre'0 = H1). apply proof_irrelevance. subst H1. + exact extapp. } + { repeat (destruct_compose; intros). eapply verified_lambdabox_pipeline_extends. + repeat (destruct_compose; intros). cbn - [transform]. + generalize dependent pre3. rewrite <- eq. + cbn [transform pcuic_expand_lets_transform expand_lets_program]. + unfold expand_lets_program. cbn [fst snd]. + intros pre3. cbn in pre3. intros <-. intros. + assert (pre''0 = H1). apply proof_irrelevance. subst H1. + exact extapp'. } } + set (tr := transform _ _ _). + destruct tr eqn:heqtr. cbn -[transform]. f_equal. + replace t1 with tr.2. subst tr. + 2:{ now rewrite heqtr; cbn. } + clear heqtr. + move: pre4. + rewrite heq. intros h. + epose proof (transform_lambda_box_eta_app _ _ _ h). + forward H. { cbn [fst snd]. + clear -eq eta extapp. revert pre3 extapp. + rewrite -eq. pose proof (correctness _ _ pre'0). + destruct H as [? []]. cbn [fst snd] in eta |- *. revert pre'0 H H0 H1 eta. rewrite eq. + intros. cbn -[transform] in H1. cbn -[transform]. + eapply EEtaExpandedFix.expanded_isEtaExp in H1. + eapply EEtaExpandedFix.isEtaExp_extends; tea. + pose proof (correctness _ _ pre3). apply H2. } + destruct H as [prelam [prelam' eqlam]]. rewrite eqlam. + rewrite snd_pair. clear eqlam. + destruct_compose_no_clear. + intros hlt heqlt. symmetry. + apply f_equal2. + eapply transform_lambdabox_pres_term. + split. rewrite fst_pair. + { destruct_compose_no_clear. intros H eq'. clear -extapp. + eapply extends_eq; tea. do 2 f_equal. clear extapp. + change (transform (pcuic_expand_lets_transform K) (Σ, tApp t u) pre).1 with + (transform (pcuic_expand_lets_transform K) (Σ, t) pre').1 in pre'0 |- *. + revert pre'0. + rewrite -surjective_pairing. intros pre'0. f_equal. apply proof_irrelevance. } + rewrite snd_pair. + destruct_compose_no_clear. intros ? ?. + eapply transform_erase_pres_term. + rewrite fst_pair. + { red. cbn. split => //. } reflexivity. + eapply transform_lambdabox_pres_term. + split. rewrite fst_pair. + { unfold run, time. destruct_compose_no_clear. intros H eq'. clear -extapp'. + assert (pre''0 = H). apply proof_irrelevance. subst H. apply extapp'. } + cbn [snd run]. unfold run, time. + destruct_compose_no_clear. intros ? ?. + eapply transform_erase_pres_term. cbn [fst]. + { red. cbn. split => //. } reflexivity. + Qed.*) + Transparent erase_transform. End PCUICErase. From 6a26f385ce19b3204d93a81cd4e30994dbcd6ccf Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 14 Feb 2024 07:47:52 +0100 Subject: [PATCH 07/28] =?UTF-8?q?Resurrect=20the=20cofix=20transform,=20ad?= =?UTF-8?q?ding=20a=20new=20axiom=20for=20the=20admitted=20pr=E2=80=A6=20(?= =?UTF-8?q?#1056)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Resurrect the cofix transform, adding a new axiom for the admitted proofs. Generalize `MetaCoq Erase` to take options allowing to optionally run this pass * Minor fixes * Fix metacoq_tour * Fix quoting of cofix to make translation correct --- erasure-plugin/_PluginProject.in | 2 + erasure-plugin/src/g_metacoq_erasure.mlg | 135 +- .../src/metacoq_erasure_plugin.mlpack | 1 + erasure-plugin/theories/ETransform.v | 108 +- erasure-plugin/theories/Erasure.v | 177 ++- erasure/_CoqProject.in | 1 + erasure/theories/ECoInductiveToInductive.v | 1158 +++++++++++++++++ examples/metacoq_tour.v | 2 +- template-coq/src/ast_quoter.ml | 6 +- template-coq/src/constr_quoter.ml | 6 +- template-coq/src/quoter.ml | 14 +- test-suite/erasure_live_test.v | 26 +- test-suite/erasure_test.v | 20 +- 13 files changed, 1532 insertions(+), 124 deletions(-) create mode 100644 erasure/theories/ECoInductiveToInductive.v diff --git a/erasure-plugin/_PluginProject.in b/erasure-plugin/_PluginProject.in index 089e2bba1..c4015c525 100644 --- a/erasure-plugin/_PluginProject.in +++ b/erasure-plugin/_PluginProject.in @@ -118,6 +118,8 @@ src/eInlineProjections.mli src/eInlineProjections.ml src/eConstructorsAsBlocks.mli src/eConstructorsAsBlocks.ml +src/eCoInductiveToInductive.mli +src/eCoInductiveToInductive.ml src/eTransform.mli src/eTransform.ml src/erasure0.mli diff --git a/erasure-plugin/src/g_metacoq_erasure.mlg b/erasure-plugin/src/g_metacoq_erasure.mlg index be1defda1..5e61fde14 100644 --- a/erasure-plugin/src/g_metacoq_erasure.mlg +++ b/erasure-plugin/src/g_metacoq_erasure.mlg @@ -9,6 +9,23 @@ open PeanoNat.Nat open Datatypes open Vernacextend open Tm_util +open Erasure0 + +open Stdarg +open Pcoq.Prim +open Ltac_plugin +open Tacexpr +open Tacinterp +open Stdarg +open Tacarg +open Genredexpr + +type erasure_command_args = + | Unsafe + | Time + | Typed + | BypassQeds + | Fast let pr_char c = str (Char.escaped c) @@ -25,49 +42,97 @@ let pr_char_list l = (* We allow utf8 encoding *) str (Caml_bytestring.caml_string_of_bytestring l) -let check ~bypass ~fast ?(with_types=false) env evm c = +type erasure_config = + { unsafe : bool; + time : bool; + typed : bool; + bypass_qeds : bool; + fast : bool; + } + +let default_config = + { unsafe = false; + time = false; + typed = false; + bypass_qeds = false; + fast = false } + +let make_erasure_config config = + let open Erasure0 in + { enable_cofix_to_fix = config.unsafe; + enable_typed_erasure = config.typed; + enable_fast_remove_params = config.fast; + dearging_config = default_dearging_config; + } + +let time_opt config str fn arg = + if config.time then + time str fn arg + else fn arg + +let check config env evm c = debug (fun () -> str"Quoting"); - let term = time (str"Quoting") (Ast_quoter.quote_term_rec ~bypass env evm) (EConstr.to_constr evm c) in - let erase = time (str"Erasure") - (if fast then Erasure0.erase_fast_and_print_template_program - else - if with_types then - Erasure0.typed_erase_and_print_template_program - else Erasure0.erase_and_print_template_program) - term + let time str f arg = time_opt config str f arg in + let term = time (str"Quoting") (Ast_quoter.quote_term_rec ~bypass:config.bypass_qeds env evm) (EConstr.to_constr evm c) in + let erasure_config = make_erasure_config config in + let erase = + time (str"Erasure") (Erasure0.erase_and_print_template_program erasure_config) term in - Feedback.msg_info (pr_char_list erase) + Feedback.msg_notice (pr_char_list erase) + +let interp_erase args env evm c = + let open Erasure0 in + let config = + let rec aux config = function + | [] -> config + | arg :: args -> + match arg with + | Unsafe -> aux {config with unsafe = true} args + | Time -> aux {config with time = true} args + | Typed -> aux {config with typed = true} args + | BypassQeds -> aux {config with bypass_qeds = true} args + | Fast -> aux {config with fast = true} args + in aux default_config args + in + check config env evm c + +let help_msg : string = + "Usage:\n\ +To erase a Gallina definition named type:\n\ + MetaCoq Erase .\n\n\ +To show this help message type:\n\ + MetaCoq Erase -help.\n\n\ +Valid options:\n\ +-typed : Run the typed erasure pipeline including a dearging phase. By default we run the pipeline without this phase.\n\ +-unsafe : Run also partially verified passes of the pipeline. This includes the cofix to fix translation.\n\ +-time : Time each compilation phase\n\ +-bypass-qeds : Bypass the use of Qed and quote opaque proofs. Beware, this can result in large memory\n\ + consumption due to reification of large proof terms.\n\ + By default, we use the (trusted) Template-Coq quoting optimization that quotes every opaque term as an axiom.\n\ + All these axioms should live in Prop so that erasure is not affected by the absence of their bodies.\n\ +-fast : Enables an alternative implementation of the parameter stripping phase that uses accumulators\n\ + instead of a view (see Erasure.ERemoveParams).\n\ +\n\ +See https://metacoq.github.io for more detailed information." + } +ARGUMENT EXTEND erase_args +| [ "-unsafe" ] -> { Unsafe } +| [ "-time" ] -> { Time } +| [ "-typed" ] -> { Typed } +| [ "-bypass-qeds" ] -> { BypassQeds } +| [ "-fast" ] -> { Fast } +END + VERNAC COMMAND EXTEND MetaCoqErase CLASSIFIED AS QUERY -| [ "MetaCoq" "Bypass" "Erase" constr(c) ] -> { +| [ "MetaCoq" "Erase" erase_args_list(l) constr(c) ] -> { let env = Global.env () in let evm = Evd.from_env env in let (c, _) = Constrintern.interp_constr env evm c in - check ~bypass:true ~fast:false env evm c + interp_erase l env evm c } -| [ "MetaCoq" "Bypass" "Typed" "Erase" constr(c) ] -> { - let env = Global.env () in - let evm = Evd.from_env env in - let (c, _) = Constrintern.interp_constr env evm c in - check ~bypass:true ~fast:false ~with_types:true env evm c - } -| [ "MetaCoq" "Typed" "Erase" constr(c) ] -> { - let env = Global.env () in - let evm = Evd.from_env env in - let (c, _) = Constrintern.interp_constr env evm c in - check ~bypass:false ~fast:false ~with_types:true env evm c - } -| [ "MetaCoq" "Erase" constr(c) ] -> { - let env = Global.env () in - let evm = Evd.from_env env in - let (c, _) = Constrintern.interp_constr env evm c in - check ~bypass:false ~fast:false env evm c - } -| [ "MetaCoq" "Fast" "Erase" constr(c) ] -> { - let env = Global.env () in - let evm = Evd.from_env env in - let (c, _) = Constrintern.interp_constr env evm c in - check ~bypass:false ~fast:true env evm c +| [ "MetaCoq" "Erase" "-help" ] -> { + Feedback.msg_notice (str help_msg) } END diff --git a/erasure-plugin/src/metacoq_erasure_plugin.mlpack b/erasure-plugin/src/metacoq_erasure_plugin.mlpack index a366f9564..485dad598 100644 --- a/erasure-plugin/src/metacoq_erasure_plugin.mlpack +++ b/erasure-plugin/src/metacoq_erasure_plugin.mlpack @@ -64,6 +64,7 @@ ErasureFunctionProperties EOptimizePropDiscr EInlineProjections EConstructorsAsBlocks +ECoInductiveToInductive EProgram OptimizePropDiscr diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index 717e7320f..1dfc50628 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -744,7 +744,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) := - {| name := "stripping constructor parameters"; + {| name := "stripping constructor parameters (using a view)"; 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; @@ -789,7 +789,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) := - {| name := "stripping constructor parameters (faster?)"; + {| name := "stripping constructor parameters (using accumulators)"; 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; @@ -927,7 +927,7 @@ Qed. 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} : + {has_app : has_tApp} {has_rel : has_tRel} {has_box : 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) := {| name := "transforming to constuctors as blocks"; transform p _ := EConstructorsAsBlocks.transform_blocks_program p ; @@ -953,8 +953,8 @@ 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)) + {has_app : has_tApp} {has_rel : has_tRel} {has_box : 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) (has_box := has_box) (has_pars := has_pars) (has_cstrblocks := has_cstrblocks)) (fun p p' => extends p.1 p'.1) (fun p p' => extends p.1 p'.1). Proof. red. intros p p' pr pr' ext. rewrite /transform /=. @@ -963,8 +963,8 @@ 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)) + {has_app : has_tApp} {has_rel : has_tRel} {has_box : 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) (has_box := has_box) (has_pars := has_pars) (has_cstrblocks := has_cstrblocks)) extends_eprogram_env extends_eprogram. Proof. red. intros p p' pr pr' [ext eq]. rewrite /transform /=. split. @@ -974,3 +974,97 @@ Proof. eapply transform_blocks_extends; eauto. apply pr. apply pr'. Qed. +From MetaCoq.Erasure Require ECoInductiveToInductive. + +Program Definition coinductive_to_inductive_transformation (efl : EEnvFlags) + {has_app : has_tApp} {has_box : has_tBox} {has_rel : has_tRel} {has_pars : has_cstr_params = false} + {has_cstrblocks : cstr_as_blocks = true} : + Transform.t _ _ EAst.term EAst.term _ _ + (eval_eprogram_env block_wcbv_flags) (eval_eprogram block_wcbv_flags) := + {| name := "transforming co-inductive to inductive types"; + transform p _ := ECoInductiveToInductive.trans_program p ; + pre p := wf_eprogram_env efl p ; + post p := wf_eprogram efl p ; + obseq p hp p' v v' := v' = ECoInductiveToInductive.trans p.1 v |}. + +Next Obligation. + move=> efl hasapp hasbox hasrel haspars hascstrs [Σ t] [wftp wft]. + cbn in *. eapply ECoInductiveToInductive.trans_program_wf; eauto. split => //. +Qed. +Next Obligation. + red. move=> efl hasapp hasbox hasrel haspars hascstrs [Σ t] /= v [wfe1 wfe2] [ev]. + eexists. split; [ | eauto]. + econstructor. + cbn -[transform_blocks]. + eapply ECoInductiveToInductive.trans_correct; cbn; eauto. + eapply wellformed_closed_env, wfe1. + Unshelve. all:eauto. +Qed. + +#[global] +Instance coinductive_to_inductive_transformation_ext (efl : EEnvFlags) + {has_app : has_tApp} {has_rel : has_tRel} {has_box : has_tBox} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = true} : + TransformExt.t (coinductive_to_inductive_transformation efl (has_app := has_app) (has_rel := has_rel) + (has_box := has_box) (has_pars := has_pars) (has_cstrblocks := has_cstrblocks)) + (fun p p' => extends p.1 p'.1) (fun p p' => extends p.1 p'.1). +Proof. + red. intros p p' pr pr' ext. rewrite /transform /=. + eapply ECoInductiveToInductive.trust_cofix. +Qed. + +#[global] +Instance coinductive_to_inductive_transformation_ext' (efl : EEnvFlags) + {has_app : has_tApp} {has_rel : has_tRel} {has_box : has_tBox} {has_pars : has_cstr_params = false} {has_cstrblocks : cstr_as_blocks = true} : + TransformExt.t (coinductive_to_inductive_transformation efl (has_app := has_app) (has_rel := has_rel) + (has_box := has_box) (has_pars := has_pars) (has_cstrblocks := has_cstrblocks)) + extends_eprogram_env extends_eprogram. +Proof. + red. intros p p' pr pr' ext. rewrite /transform /=. + eapply ECoInductiveToInductive.trust_cofix. +Qed. + +Program Definition optional_transform {env env' term term' value value' eval eval'} (activate : bool) + (tr : Transform.t env env' term term' value value' eval eval') : + if activate then Transform.t env env' term term' value value' eval eval' + else Transform.t env env term term value value eval eval := + if activate return if activate then Transform.t env env' term term' value value' eval eval' + else Transform.t env env term term value value eval eval + then tr + else + {| name := ("skipped " ^ tr.(name)); + transform p pr := p; + pre := tr.(pre) ; + post := tr.(pre) ; + obseq p hp p' v v' := v' = v |}. +Next Obligation. + intros. cbn. exact p. +Defined. +Next Obligation. + cbn. intros. red. intros. exists v. split => //. +Defined. + +Program Definition optional_self_transform {env term eval} (activate : bool) + (tr : Transform.self_transform env term eval eval) : + Transform.self_transform env term eval eval := + if activate return Transform.self_transform env term eval eval + then tr + else + {| name := ("skipped " ^ tr.(name)); + transform p pr := p; + pre := tr.(pre) ; + post := tr.(pre) ; + obseq p hp p' v v' := v' = v |}. +Next Obligation. + intros. cbn. exact p. +Defined. +Next Obligation. + cbn. intros. red. intros. exists v. split => //. +Defined. + +#[global] +Instance optional_self_transformation_ext {env term eval} activate tr extends : + TransformExt.t tr extends extends -> + TransformExt.t (@optional_self_transform env term eval activate tr) extends extends. +Proof. + red; intros. destruct activate; cbn in * => //. now apply H. +Qed. \ No newline at end of file diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 05e3a084e..9f8688745 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -30,6 +30,32 @@ Import EWcbvEval. Local Obligation Tactic := program_simpl. +Record erasure_configuration := { + enable_cofix_to_fix : bool; + enable_typed_erasure : bool; + enable_fast_remove_params : bool; + dearging_config : dearging_config + }. + +Definition default_dearging_config := + {| overridden_masks := fun _ => None; + do_trim_const_masks := true; + do_trim_ctor_masks := false |}. + +(* This runs the cofix -> fix translation which is not entirely verified yet *) +Definition default_erasure_config := + {| enable_cofix_to_fix := true; + dearging_config := default_dearging_config; + enable_typed_erasure := true; + enable_fast_remove_params := true |}. + +(* This runs only the verified phases without the typed erasure and "fast" remove params *) +Definition safe_erasure_config := + {| enable_cofix_to_fix := false; + enable_typed_erasure := false; + enable_fast_remove_params := false; + dearging_config := default_dearging_config |}. + Axiom assume_welltyped_template_program_expansion : forall p (wtp : ∥ wt_template_program_env p ∥), let p' := EtaExpand.eta_expand_program p in @@ -58,14 +84,35 @@ Next Obligation. apply assume_preservation_template_program_env_expansion in ev as [ev']; eauto. Qed. -Program Definition verified_lambdabox_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - 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 |}) := +Definition final_wcbv_flags := {| + with_prop_case := false; + with_guarded_fix := false; + with_constructor_as_block := true |}. + +Program Definition optional_cofix_to_fix_transform econf := + ETransform.optional_self_transform econf.(enable_cofix_to_fix) + ((* Rebuild the efficient lookup table *) + rebuild_wf_env_transform (efl := EConstructorsAsBlocks.switch_cstr_as_blocks + (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags))) false false ▷ + (* Coinductives & cofixpoints are translated to inductive types and thunked fixpoints *) + let efl := EConstructorsAsBlocks.switch_cstr_as_blocks + (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags)) in + coinductive_to_inductive_transformation efl + (has_app := eq_refl) (has_box := eq_refl) (has_rel := eq_refl) (has_pars := eq_refl) (has_cstrblocks := eq_refl)) + . + +Program Definition verified_lambdabox_pipeline {guard : abstract_guard_impl} + (efl := EWellformed.all_env_flags) + : Transform.t _ _ EAst.term EAst.term _ _ + (* Standard evaluation, with cases on prop, guarded fixpoints, applied constructors *) + (EProgram.eval_eprogram_env default_wcbv_flags) + (* Target evaluation, with no more cases on prop, unguarded fixpoints, constructors as block *) + (EProgram.eval_eprogram final_wcbv_flags) := + (* 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 (fl := {| with_prop_case := true; with_guarded_fix := true; with_constructor_as_block := false |}) (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 *) @@ -81,7 +128,7 @@ Program Definition verified_lambdabox_pipeline {guard : abstract_guard_impl} (ef (* First-order constructor representation *) 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). + (has_app := eq_refl) (has_pars := eq_refl) (has_rel := eq_refl) (has_box := 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 @@ -96,12 +143,11 @@ Next Obligation. now eapply ETransform.expanded_eprogram_env_expanded_eprogram_cstrs. Qed. - Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : 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 |}) := + (EProgram.eval_eprogram final_wcbv_flags) := (* a bunch of nonsense for normalization preconditions *) let K ty (T : ty -> _) p := let p := T p in @@ -131,14 +177,14 @@ Proof. Qed. Lemma verified_lambdabox_pipeline_extends {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - TransformExt.t (verified_lambdabox_pipeline) (fun p p' => extends (EEnvMap.GlobalContextMap.global_decls p.1) + TransformExt.t verified_lambdabox_pipeline (fun p p' => extends (EEnvMap.GlobalContextMap.global_decls p.1) (EEnvMap.GlobalContextMap.global_decls p'.1)) (fun p p' => extends p.1 p'.1). Proof. unfold verified_lambdabox_pipeline. tc. Qed. Lemma verified_lambdabox_pipeline_extends' {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - TransformExt.t (verified_lambdabox_pipeline) extends_eprogram_env extends_eprogram. + TransformExt.t verified_lambdabox_pipeline extends_eprogram_env extends_eprogram. Proof. unfold verified_lambdabox_pipeline. tc. Qed. @@ -167,15 +213,14 @@ 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 _ _ TemplateProgram.eval_template_program - (EProgram.eval_eprogram {| with_prop_case := false; with_guarded_fix := false; with_constructor_as_block := true |}) := + (EProgram.eval_eprogram final_wcbv_flags) := pre_erasure_pipeline ▷ verified_erasure_pipeline. - -Program Definition verified_lambdabox_typed_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : +Program Definition verified_lambdabox_typed_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) econf : Transform.t _ _ EAst.term EAst.term _ _ (EProgram.eval_eprogram_env {| with_prop_case := false; 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 |}) := + (EProgram.eval_eprogram final_wcbv_flags) := (* 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. *) @@ -191,7 +236,17 @@ Program Definition verified_lambdabox_typed_pipeline {guard : abstract_guard_imp (* First-order constructor representation *) 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). + (has_app := eq_refl) (has_pars := eq_refl) (has_rel := eq_refl) (has_box := eq_refl) (has_cstrblocks := eq_refl) ▷ + ETransform.optional_self_transform econf.(enable_cofix_to_fix) + ((* Rebuild the efficient lookup table *) + rebuild_wf_env_transform (efl := EConstructorsAsBlocks.switch_cstr_as_blocks + (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags))) false false ▷ + (* Coinductives & cofixpoints are translated to inductive types and thunked fixpoints *) + let efl := EConstructorsAsBlocks.switch_cstr_as_blocks + (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags)) in + coinductive_to_inductive_transformation efl + (has_app := eq_refl) (has_box := eq_refl) (has_rel := eq_refl) (has_pars := 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 @@ -205,16 +260,19 @@ Program Definition verified_lambdabox_typed_pipeline {guard : abstract_guard_imp destruct H. split => //. sq. now eapply ETransform.expanded_eprogram_env_expanded_eprogram_cstrs. Qed. + Next Obligation. + destruct H. destruct enable_cofix_to_fix => //. + Qed. Local Obligation Tactic := intros; eauto. Program Definition verified_typed_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) - (cf : dearging_config) : + econf : 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 |}) := + (EProgram.eval_eprogram final_wcbv_flags) := (* a bunch of nonsense for normalization preconditions *) let K ty (T : ty -> _) p := let p := T p in @@ -228,10 +286,10 @@ Program Definition verified_typed_erasure_pipeline {guard : abstract_guard_impl} (* Remove match on box early for dearging *) remove_match_on_box_typed_transform (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ (* Check if the preconditions for dearging are valid, otherwise dearging will be the identity *) - dearging_checks_transform cf (hastrel := eq_refl) (hastbox := eq_refl) ▷ - dearging_transform cf ▷ + dearging_checks_transform econf.(dearging_config) (hastrel := eq_refl) (hastbox := eq_refl) ▷ + dearging_transform econf.(dearging_config) ▷ rebuild_wf_env_transform true true ▷ - verified_lambdabox_typed_pipeline. + verified_lambdabox_typed_pipeline econf. Next Obligation. cbn in H. split; cbn; intuition eauto. @@ -242,13 +300,13 @@ Program Definition verified_typed_erasure_pipeline {guard : abstract_guard_impl} Program Definition typed_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) - cf : + econf : 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 |}) := + (EProgram.eval_eprogram final_wcbv_flags) := pre_erasure_pipeline ▷ - verified_typed_erasure_pipeline cf. + verified_typed_erasure_pipeline econf. (* 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 @@ -258,9 +316,7 @@ Program Definition typed_erasure_pipeline {guard : abstract_guard_impl} 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) := +Program Definition erasure_pipeline_fast {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) econf := (* a bunch of nonsense for normalization preconditions *) let K ty (T : ty -> _) p := let p := T p in @@ -283,7 +339,8 @@ Program Definition erasure_pipeline_fast {guard : abstract_guard_impl} (efl := E 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 false ▷ - 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 := efl) (has_app := eq_refl) (has_pars := eq_refl) (has_rel := eq_refl) (has_box := eq_refl) (has_cstrblocks := eq_refl) ▷ + optional_cofix_to_fix_transform econf. Next Obligation. destruct H; split => //. now eapply ETransform.expanded_eprogram_env_expanded_eprogram_cstrs. Qed. @@ -302,11 +359,15 @@ Qed. Next Obligation. cbn in H. split; cbn; intuition eauto. Qed. +Next Obligation. + cbn in H. unfold optional_cofix_to_fix_transform. destruct enable_cofix_to_fix => //. +Qed. Next Obligation. cbn in H. split; cbn; intuition eauto. Qed. -Definition run_erase_program_fast {guard : abstract_guard_impl} := run erasure_pipeline_fast. +Definition run_erase_program_fast {guard : abstract_guard_impl} (econf : erasure_configuration) := + run (erasure_pipeline_fast econf). Local Open Scope string_scope. @@ -335,34 +396,20 @@ Global Existing Instance fake_normalization. Axiom assume_that_we_only_erase_on_welltyped_programs : forall {cf : checker_flags}, forall (p : Ast.Env.program), squash (TemplateProgram.wt_template_program p). -Program Definition erase_and_print_template_program (p : Ast.Env.program) - : string := - let p' := run_erase_program p _ in - time "Pretty printing" EPretty.print_program p'. -Next Obligation. - split. - now eapply assume_that_we_only_erase_on_welltyped_programs. - cbv [PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn]. - pose proof @PCUICSN.normalization. - split; typeclasses eauto. -Qed. - -Program Definition erase_fast_and_print_template_program (p : Ast.Env.program) - : string := - let p' := run_erase_program_fast p _ in - time "pretty-printing" EPretty.print_program p'. +(* This also optionally runs the cofix to fix translation *) +Program Definition run_erase_program {guard : abstract_guard_impl} econf := + if econf.(enable_typed_erasure) then run (typed_erasure_pipeline econf) + else if econf.(enable_fast_remove_params) then + run (erasure_pipeline_fast econf) + else run (erasure_pipeline ▷ (optional_cofix_to_fix_transform econf)). Next Obligation. - split. - now eapply assume_that_we_only_erase_on_welltyped_programs. - cbv [PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn]. - pose proof @PCUICSN.normalization. - split; typeclasses eauto. +Proof. + unfold optional_cofix_to_fix_transform. + destruct enable_cofix_to_fix => //. Qed. -(* Parameterized by a configuration for dearging, allowing to, e.g., override masks. *) -Program Definition typed_erase_and_print_template_program_gen (p : Ast.Env.program) cf - : string := - let p' := run (typed_erasure_pipeline cf) p _ in +Program Definition erase_and_print_template_program econf (p : Ast.Env.program) : string := + let p' := run_erase_program econf p _ in time "Pretty printing" EPretty.print_program p'. Next Obligation. split. @@ -372,12 +419,22 @@ Next Obligation. split; typeclasses eauto. Qed. -Definition default_dearging_config := - {| overridden_masks := fun _ => None; - do_trim_const_masks := true; - do_trim_ctor_masks := false |}. +Definition erasure_fast_config := + {| enable_cofix_to_fix := false; + dearging_config := default_dearging_config; + enable_typed_erasure := false; + enable_fast_remove_params := true |}. -Definition typed_erase_and_print_template_program (p : Ast.Env.program) - : string := - typed_erase_and_print_template_program_gen p default_dearging_config. +Program Definition erase_fast_and_print_template_program (p : Ast.Env.program) : string := + erase_and_print_template_program erasure_fast_config p. + +Definition typed_erasure_config := + {| enable_cofix_to_fix := false; + dearging_config := default_dearging_config; + enable_typed_erasure := true; + enable_fast_remove_params := true |}. +(* Parameterized by a configuration for dearging, allowing to, e.g., override masks. *) +Program Definition typed_erase_and_print_template_program (p : Ast.Env.program) + : string := + erase_and_print_template_program typed_erasure_config p. diff --git a/erasure/_CoqProject.in b/erasure/_CoqProject.in index a4a94c374..e0be67b79 100644 --- a/erasure/_CoqProject.in +++ b/erasure/_CoqProject.in @@ -38,6 +38,7 @@ theories/EInlineProjections.v theories/EConstructorsAsBlocks.v theories/EWcbvEvalCstrsAsBlocksInd.v theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v +theories/ECoInductiveToInductive.v theories/EImplementBox.v theories/Typed/Annotations.v diff --git a/erasure/theories/ECoInductiveToInductive.v b/erasure/theories/ECoInductiveToInductive.v new file mode 100644 index 000000000..3132dd66c --- /dev/null +++ b/erasure/theories/ECoInductiveToInductive.v @@ -0,0 +1,1158 @@ +(* 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. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils + PCUICReflect PCUICWeakeningEnvConv PCUICWeakeningEnvTyp + PCUICTyping PCUICInversion + PCUICSafeLemmata. (* for welltyped *) +From MetaCoq.SafeChecker Require Import PCUICWfEnvImpl. +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils EDeps EExtends + EEtaExpanded + ELiftSubst ECSubst ESpineView EGlobalEnv EInduction EWellformed EWcbvEval Extract Prelim + EEnvMap EArities EProgram. + +Import MCList (map_InP, map_InP_elim, map_InP_spec). + +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. + +(** We assumes [Prop ... -> Tn -> C here + and we want to produce an expansion: + λ x0 .. xn (), t x0 xn () *) + Equations make_n_aux (n : nat) (t : term) (acc : list term) : term := + make_n_aux 0 t acc => tLambda + (nNamed "thunk") + (mkApps (lift0 1 t) (rev (tRel 0 :: map (lift0 1) acc))); + make_n_aux (S n) t acc => + tLambda + (nNamed (make_name "x" (S n))) + (make_n_aux n (lift0 1 t) (tRel 0 :: map (lift0 1) acc)). + + Definition make_n (n : nat) (t : term) := make_n_aux n t []. + + (* Eval compute in make_n 2 (tRel 0). *) + +End Thunk. + +Section trans. + Context (Σ : GlobalContextMap.t). + + Definition trans_cofix (d : def term) := + {| dname := d.(dname); + dbody := Thunk.make_n d.(rarg) d.(dbody); + rarg := d.(rarg) |}. + + Fixpoint trans (t : term) : term := + match t with + | tRel i => tRel i + | tEvar ev args => tEvar ev (List.map trans args) + | tLambda na M => tLambda na (trans M) + | tApp u v => tApp (trans u) (trans v) + | tLetIn na b b' => tLetIn na (trans b) (trans b') + | tCase ind c brs => + let brs' := List.map (on_snd trans) brs in + match GlobalContextMap.lookup_inductive_kind Σ (fst ind).(inductive_mind) with + | Some CoFinite => + tCase ind (Thunk.force (trans c)) brs' + | _ => tCase ind (trans c) brs' + end + | tProj p c => + match GlobalContextMap.lookup_inductive_kind Σ p.(proj_ind).(inductive_mind) with + | Some CoFinite => tProj p (Thunk.force (trans c)) + | _ => tProj p (trans c) + end + | tFix mfix idx => + let mfix' := List.map (map_def trans) mfix in + tFix mfix' idx + | tCoFix mfix idx => + let mfix' := List.map (map_def trans) mfix in + let mfix' := List.map trans_cofix mfix' in + match nth_error mfix' idx with + | Some d => Thunk.make_n d.(rarg) (tFix mfix' idx) + | None => tCoFix mfix' idx + end + | tBox => t + | tVar _ => t + | tConst _ => t + | tConstruct ind i args => + match GlobalContextMap.lookup_inductive_kind Σ ind.(inductive_mind) with + | Some CoFinite => Thunk.make (tConstruct ind i (map trans args)) + | _ => tConstruct ind i (map trans args) + end + | tPrim p => tPrim (map_prim trans p) + end. + + (* cofix succ x := match x with Stream x xs => Stream (x + 1) (succ xs) -> + + fix succ x := match x () with Stream x xs => fun () => Stream (x + 1) (succ xs) + + cofix ones := Stream 1 ones -> + fix ones := fun () => Stream 1 ones + *) + + Lemma trans_mkApps f l : trans (mkApps f l) = mkApps (trans f) (map trans l). + Proof using Type. + induction l using rev_ind; simpl; auto. + now rewrite mkApps_app /= IHl map_app /= mkApps_app /=. + Qed. + + Lemma map_repeat {A B} (f : A -> B) x n : map f (repeat x n) = repeat (f x) n. + Proof using Type. + now induction n; simpl; auto; rewrite IHn. + Qed. + + Lemma map_trans_repeat_box n : map trans (repeat tBox n) = repeat tBox n. + Proof using Type. by rewrite map_repeat. Qed. + + Import ECSubst. + + Lemma csubst_mkApps {a k f l} : csubst a k (mkApps f l) = mkApps (csubst a k f) (map (csubst a k) l). + Proof using Type. + induction l using rev_ind; simpl; auto. + rewrite mkApps_app /= IHl. + now rewrite -[EAst.tApp _ _](mkApps_app _ _ [_]) map_app. + Qed. + + Lemma csubst_closed t k x : closedn k x -> csubst t k x = x. + Proof using Type. + induction x in k |- * using EInduction.term_forall_list_ind; simpl; auto. + all:try solve [intros; f_equal; solve_all; eauto]. + intros Hn. eapply Nat.ltb_lt in Hn. + - destruct (Nat.compare_spec k n); try lia. reflexivity. + - move/andP => []. intros. f_equal; solve_all; eauto. + - move/andP => []. intros. f_equal; solve_all; eauto. + - move/andP => []. intros. f_equal; solve_all; eauto. + Qed. + + Lemma closed_trans t k : closedn k t -> closedn k (trans t). + Proof using Type. + induction t in k |- * using EInduction.term_forall_list_ind; simpl; auto; + 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]; try easy. + - destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //; solve_all. + rewrite -Nat.add_1_r. now eapply closedn_lift. + - move/andP: H => [] clt clargs. + destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //; rtoProp; solve_all; solve_all. + - destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //; rtoProp; solve_all; solve_all. + - solve_all. destruct nth_error eqn:hnth. + * apply trust_cofix. + * cbn. unfold trans_cofix. len. solve_all. + unfold test_def. cbn. apply trust_cofix. + - primProp. solve_all_k 6. + Qed. + + Lemma subst_csubst_comm l t k b : + forallb (closedn 0) l -> closed t -> + subst l 0 (csubst t (#|l| + k) b) = + csubst t k (subst l 0 b). + Proof using Type. + intros hl cl. + rewrite !closed_subst //. + rewrite distr_subst. f_equal. + symmetry. solve_all. + rewrite subst_closed //. + eapply closed_upwards; tea. lia. + Qed. + + Lemma substl_subst s t : + forallb (closedn 0) s -> + substl s t = subst s 0 t. + Proof using Type. + induction s in t |- *; cbn; auto. + intros _. now rewrite subst_empty. + move/andP=> []cla cls. + rewrite (subst_app_decomp [_]). + cbn. rewrite lift_closed //. + rewrite closed_subst //. now eapply IHs. + Qed. + + Lemma substl_csubst_comm l t k b : + forallb (closedn 0) l -> closed t -> + substl l (csubst t (#|l| + k) b) = + csubst t k (substl l b). + Proof using Type. + intros hl cl. + rewrite substl_subst //. + rewrite substl_subst //. + apply subst_csubst_comm => //. + Qed. + + Lemma trans_csubst a k b : + closed a -> + trans (ECSubst.csubst a k b) = ECSubst.csubst (trans a) k (trans b). + Proof using Type. + induction b in k |- * using EInduction.term_forall_list_ind; simpl; auto; + intros cl; 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]; try easy. + - destruct (k ?= n)%nat; auto. + - destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //. + 1,3,4:f_equal; rewrite map_map_compose; solve_all. + unfold Thunk.make. f_equal. cbn. + rewrite !map_map_compose. f_equal; solve_all. + specialize (H k cl). rewrite H. + rewrite closed_subst. now apply closed_trans. + rewrite closed_subst. now apply closed_trans. + now rewrite commut_lift_subst_rec. + - destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //. + all:f_equal; eauto; try (rewrite /on_snd map_map_compose; solve_all). + unfold Thunk.force. f_equal; eauto. + - destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //. + all:f_equal; eauto; try (rewrite /on_snd map_map_compose; solve_all). + unfold Thunk.force. f_equal; eauto. + - f_equal. solve_all. + rewrite !nth_error_map. destruct nth_error eqn:hnth => //=. + 2:{ f_equal. rewrite map_map_compose. eapply map_ext_in => x hin. + rewrite /trans_cofix /map_def //=. f_equal. len. + rewrite /Thunk.make_n. apply trust_cofix. + } + apply trust_cofix. + Qed. + + Lemma trans_substl s t : + forallb (closedn 0) s -> + trans (substl s t) = substl (map trans s) (trans t). + Proof using Type. + induction s in t |- *; simpl; auto. + move/andP => [] cla cls. + rewrite IHs //. f_equal. + now rewrite trans_csubst. + Qed. + + Lemma trans_iota_red pars args br : + forallb (closedn 0) args -> + trans (EGlobalEnv.iota_red pars args br) = EGlobalEnv.iota_red pars (map trans args) (on_snd trans br). + Proof using Type. + intros cl. + unfold EGlobalEnv.iota_red. + rewrite trans_substl //. + rewrite forallb_rev forallb_skipn //. + now rewrite map_rev map_skipn. + Qed. + + Lemma trans_fix_subst mfix : EGlobalEnv.fix_subst (map (map_def trans) mfix) = map trans (EGlobalEnv.fix_subst mfix). + Proof using Type. + unfold EGlobalEnv.fix_subst. + rewrite map_length. + generalize #|mfix|. + induction n; simpl; auto. + f_equal; auto. + Qed. + + Lemma trans_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 trans) mfix) idx = Some (n, trans f). + Proof using Type. + intros hfix. + unfold cunfold_fix. + rewrite nth_error_map. + destruct nth_error. + intros [= <- <-] => /=. f_equal. + now rewrite trans_substl // trans_fix_subst. + discriminate. + Qed. + + Lemma trans_cunfold_cofix mfix idx n f : + forallb (closedn 0) (EGlobalEnv.cofix_subst mfix) -> + cunfold_cofix mfix idx = Some (n, f) -> + exists d, nth_error mfix idx = Some d /\ + cunfold_fix (map trans_cofix mfix) idx = Some (n, substl (fix_subst (map trans_cofix mfix)) (Thunk.make_n (rarg d) (dbody d))). + Proof using Type. + intros hcofix. + unfold cunfold_cofix, cunfold_fix. + rewrite nth_error_map. + destruct nth_error. + intros [= <- <-] => /=. f_equal. exists d. split => //. + discriminate. + Qed. + + Lemma trans_nth {n l d} : + trans (nth n l d) = nth n (map trans l) (trans d). + Proof using Type. + induction l in n |- *; destruct n; simpl; auto. + Qed. + +End trans. + +Lemma is_box_inv b : is_box b -> ∑ args, b = mkApps tBox args. +Proof. + unfold is_box, EAstUtils.head. + destruct decompose_app eqn:da. + simpl. destruct t => //. + eapply decompose_app_inv in da. subst. + eexists; eauto. +Qed. + +Import EWcbvEval. + +Notation "Σ ⊢ s ⇓ t" := (eval Σ s t) (at level 50, s, t at next level) : type_scope. + +Lemma eval_is_box {wfl:WcbvFlags} Σ t u : Σ ⊢ t ⇓ u -> is_box t -> u = EAst.tBox. +Proof. + intros ev; induction ev => //. + - rewrite is_box_tApp. + intros isb. intuition congruence. + - rewrite is_box_tApp. move/IHev1 => ?; solve_discr. + - rewrite is_box_tApp. move/IHev1 => ?; solve_discr. + - rewrite is_box_tApp. move/IHev1 => ?. subst => //. + - rewrite is_box_tApp. move/IHev1 => ?. subst. solve_discr. + - rewrite is_box_tApp. move/IHev1 => ?. subst. cbn in i. + destruct EWcbvEval.with_guarded_fix => //. + - destruct t => //. +Qed. + +Lemma isType_tSort {cf:checker_flags} {Σ : global_env_ext} {Γ l A} {wfΣ : wf Σ} : + Σ ;;; Γ |- tSort (sType (Universe.make l)) : A -> isType Σ Γ (tSort (sType (Universe.make l))). +Proof. + intros HT. + eapply inversion_Sort in HT as [l' [wfΓ Hs]]; auto. + eexists; econstructor; eauto. cbn. split; eauto. econstructor; eauto. +Qed. + +Lemma isType_it_mkProd {cf:checker_flags} {Σ : global_env_ext} {Γ na dom codom A} {wfΣ : wf Σ} : + Σ ;;; Γ |- tProd na dom codom : A -> + isType Σ Γ (tProd na dom codom). +Proof. + intros HT. + eapply inversion_Prod in HT as (? & ? & ? & ? & ?); auto. + eexists; cbn; econstructor; split; eauto. econstructor; eauto. +Qed. + +Definition trans_constant_decl Σ cb := + {| cst_body := option_map (trans Σ) cb.(cst_body) |}. + +Definition trans_decl Σ d := + match d with + | ConstantDecl cb => ConstantDecl (trans_constant_decl Σ cb) + | InductiveDecl idecl => d + end. + +Definition trans_env Σ := + map (on_snd (trans_decl Σ)) Σ.(GlobalContextMap.global_decls). + +Import EnvMap. + +Program Fixpoint trans_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 (trans_decl Σg) hd :: trans_env' tl (fresh_globals_cons_inv HΣ) + end. + +Import EGlobalEnv EExtends. + +(* Lemma extends_is_propositional {Σ Σ'} : + wf_glob Σ' -> extends Σ Σ' -> + forall ind, + match inductive_isprop_and_pars Σ ind with + | Some b => inductive_isprop_and_pars Σ' ind = Some b + | None => inductive_isprop_and_pars Σ' ind = None + end. +Proof. + intros wf ex ind. + rewrite /inductive_isprop_and_pars. + destruct lookup_env eqn:lookup => //. + now rewrite (extends_lookup wf ex lookup). + +Qed. *) + +Lemma extends_inductive_isprop_and_pars {efl : EEnvFlags} {Σ Σ' ind} : extends Σ Σ' -> wf_glob Σ' -> + isSome (lookup_inductive Σ ind) -> + inductive_isprop_and_pars Σ ind = inductive_isprop_and_pars Σ' ind. +Proof. + intros ext wf; cbn. + unfold inductive_isprop_and_pars. cbn. + destruct lookup_env as [[]|] eqn:hl => //. + rewrite (extends_lookup wf ext hl). + destruct nth_error => //. +Qed. + +Lemma extends_lookup_inductive_kind {efl : EEnvFlags} {Σ Σ' ind} : extends Σ Σ' -> wf_glob Σ' -> + isSome (lookup_minductive Σ ind) -> + lookup_inductive_kind Σ ind = lookup_inductive_kind Σ' ind. +Proof. + intros ext wf. + unfold lookup_inductive_kind. cbn. + destruct lookup_env as [[]|] eqn:hl => //. + now rewrite (extends_lookup wf ext hl). +Qed. + +Lemma wellformed_trans_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : + forall n, EWellformed.wellformed Σ n t -> + forall {Σ' : GlobalContextMap.t}, extends Σ Σ' -> wf_glob Σ' -> + trans Σ t = trans Σ' t. +Proof. + induction t using EInduction.term_forall_list_ind; cbn -[lookup_constant lookup_inductive + lookup_projection + lookup_constructor + GlobalContextMap.lookup_inductive_kind]; intros => //. + all:unfold wf_fix_gen in *; rtoProp; intuition auto. + all:try now f_equal; eauto; solve_all. + - rewrite !GlobalContextMap.lookup_inductive_kind_spec. + destruct lookup_constructor as [[[mdecl idecl] cdecl]|] eqn:hl => //. + destruct cstr_as_blocks. + { move/andP: H2 => [] hpars hargs. + assert (map (trans Σ) args = map (trans Σ') args) as -> by solve_all. + rewrite (extends_lookup_inductive_kind H0 H1) //. + apply lookup_constructor_lookup_inductive in hl. + unfold lookup_inductive in hl. + destruct lookup_minductive => //. } + { destruct args => // /=. + rewrite (extends_lookup_inductive_kind H0 H1) //. + apply lookup_constructor_lookup_inductive in hl. + unfold lookup_inductive in hl. + destruct lookup_minductive => //. } + - rewrite !GlobalContextMap.lookup_inductive_kind_spec. + destruct lookup_inductive as [[mdecl idecl]|] eqn:hl => //. + assert (map (on_snd (trans Σ)) l = map (on_snd (trans Σ')) l) as -> by solve_all. + rewrite (extends_lookup_inductive_kind H0 H1) //. + unfold lookup_inductive in hl. + destruct lookup_minductive => //. + rewrite (IHt _ H4 _ H0 H1) //. + - rewrite !GlobalContextMap.lookup_inductive_kind_spec. + destruct (lookup_projection) as [[[[mdecl idecl] cdecl] pdecl]|] eqn:hl => //. + eapply lookup_projection_lookup_constructor in hl. + eapply lookup_constructor_lookup_inductive in hl. + rewrite (extends_lookup_inductive_kind H0 H1) //. + unfold lookup_inductive in hl. + destruct lookup_minductive => //. + rewrite (IHt _ H2 _ H0 H1) //. + - apply trust_cofix. +Qed. + +Lemma wellformed_trans_decl_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : + wf_global_decl Σ t -> + forall {Σ' : GlobalContextMap.t}, extends Σ Σ' -> wf_glob Σ' -> + trans_decl Σ t = trans_decl Σ' t. +Proof. + destruct t => /= //. + intros wf Σ' ext wf'. f_equal. unfold trans_constant_decl. f_equal. + destruct (cst_body c) => /= //. f_equal. + now eapply wellformed_trans_extends. +Qed. + +Lemma lookup_env_trans_env_Some {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn d : + wf_glob Σ -> + GlobalContextMap.lookup_env Σ kn = Some d -> + ∑ Σ' : GlobalContextMap.t, + [× extends_prefix Σ' Σ, wf_global_decl Σ' d & + lookup_env (trans_env Σ) kn = Some (trans_decl Σ' d)]. +Proof. + rewrite GlobalContextMap.lookup_env_spec. + destruct Σ as [Σ map repr wf]. + induction Σ in map, repr, wf |- *; simpl; auto => //. + intros wfg. + case: eqb_specT => //. + - intros ->. cbn. intros [= <-]. + exists (GlobalContextMap.make Σ (fresh_globals_cons_inv wf)). split. + cbn. now exists [a]. now depelim wfg. + depelim wfg; eauto. + f_equal. symmetry. eapply wellformed_trans_decl_extends. cbn; auto. cbn. + now eapply extends_fresh. cbn. now constructor. + - intros _. + set (Σ' := GlobalContextMap.make Σ (fresh_globals_cons_inv wf)). + specialize (IHΣ (GlobalContextMap.map Σ') (GlobalContextMap.repr Σ') (GlobalContextMap.wf Σ')). + cbn in IHΣ. 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. + eapply map_ext_in => kn hin. unfold on_snd. f_equal. + symmetry. eapply wellformed_trans_decl_extends => //. cbn. + eapply lookup_env_In in hin. 2:now depelim wfg. + depelim wfg. eapply lookup_env_wellformed; tea. + cbn. destruct a. eapply extends_fresh. 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_trans_env_None {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn : + GlobalContextMap.lookup_env Σ kn = None -> + lookup_env (trans_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_trans {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn : + wf_glob Σ -> + lookup_env (trans_env Σ) kn = option_map (trans_decl Σ) (lookup_env Σ kn). +Proof. + intros wf. + rewrite -GlobalContextMap.lookup_env_spec. + destruct (GlobalContextMap.lookup_env Σ kn) eqn:hl. + - eapply lookup_env_trans_env_Some in hl as [Σ' [ext wf' hl']] => /= //. + rewrite hl'. f_equal. + eapply wellformed_trans_decl_extends; eauto. + now eapply extends_prefix_extends. + - cbn. now eapply lookup_env_trans_env_None in hl. +Qed. + +Lemma is_propositional_trans {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : + wf_glob Σ -> + inductive_isprop_and_pars Σ ind = inductive_isprop_and_pars (trans_env Σ) ind. +Proof. + rewrite /inductive_isprop_and_pars => wf. + rewrite /lookup_inductive /lookup_minductive. + rewrite (lookup_env_trans (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_trans {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind c : + wf_glob Σ -> + constructor_isprop_pars_decl Σ ind c = constructor_isprop_pars_decl (trans_env Σ) ind c. +Proof. + rewrite /constructor_isprop_pars_decl => wf. + rewrite /lookup_constructor /lookup_inductive /lookup_minductive. + rewrite (lookup_env_trans (inductive_mind ind) wf). + rewrite /GlobalContextMap.inductive_isprop_and_pars /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 -> + nth_error brs c = Some br -> + #|skipn pars args| = #|br.1| -> + closedn #|br.1| br.2 -> + closed (iota_red pars args br). +Proof. + intros clargs hnth hskip clbr. + rewrite /iota_red. + eapply ECSubst.closed_substl => //. + now rewrite forallb_rev forallb_skipn. + now rewrite List.rev_length hskip Nat.add_0_r. +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_trans {efl : EEnvFlags} {Σ : GlobalContextMap.t} {ind c} : + wf_glob Σ -> + lookup_constructor Σ ind c = lookup_constructor (trans_env Σ) ind c. +Proof. + intros wfΣ. rewrite /lookup_constructor /lookup_inductive /lookup_minductive. + rewrite lookup_env_trans // /=. destruct lookup_env => // /=. + destruct g => //. +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 value_constructor_as_block {wfl : WcbvFlags} {Σ ind c args} : value Σ (tConstruct ind c args) -> + All (value Σ) args. +Proof. + intros h; depelim h. + - depelim a. cbn in i. destruct args => //. + - eauto. + - depelim v; solve_discr. +Qed. + +Lemma constructor_isprop_pars_decl_constructor {Σ ind c prop npars cdecl} : + constructor_isprop_pars_decl Σ ind c = Some (prop, npars, cdecl) -> + ∑ mdecl idecl, lookup_constructor Σ ind c = Some (mdecl, idecl, cdecl) /\ npars = ind_npars mdecl. +Proof. + rewrite /constructor_isprop_pars_decl. + destruct lookup_constructor as [[[mdecl idecl] cdecl']|] eqn:hl => //=. + intros [= <- <- <-]. + exists mdecl, idecl; split => //. +Qed. + +Lemma isLambda_make_n n t : isLambda (Thunk.make_n n t). +Proof. + induction n; cbn => //. +Qed. + +Lemma value_trans {efl : EEnvFlags} {fl : WcbvFlags} {hasc : cstr_as_blocks = true} {wcon : with_constructor_as_block = true} {Σ : GlobalContextMap.t} {c} : + has_tApp -> wf_glob Σ -> + wellformed Σ 0 c -> + value Σ c -> + value (trans_env Σ) (trans Σ c). +Proof. + intros hasapp wfg wf h. + revert c h wf. apply: Ee.value_values_ind. + - intros t; destruct t => //; cbn -[lookup_constructor GlobalContextMap.lookup_inductive_kind]. + all:try solve [intros; repeat constructor => //]. + destruct args => //. + move=> /andP[] wc. now rewrite wcon in wc. + move=> _ /andP [] hascof /andP[] /Nat.ltb_lt /nth_error_Some hnth hm. + destruct nth_error => //. + pose proof (isLambda_make_n (rarg d) (tFix (map trans_cofix mfix) idx)). + destruct Thunk.make_n => //. apply trust_cofix. + (* do 3 constructor. *) + - intros p pv IH wf. cbn. constructor. constructor 2. + cbn in wf. move/andP: wf => [hasp tp]. + primProp. depelim tp; depelim pv; depelim IH; constructor; cbn in *; rtoProp; intuition auto; solve_all. + - intros ind c mdecl idecl cdecl args wc hl harglen hargs ihargs. + simpl. rewrite hasc. move/andP => [] /andP[] hascstr _ /andP[] hpargs wfargs. + cbn -[GlobalContextMap.lookup_inductive_kind]. + destruct GlobalContextMap.lookup_inductive_kind as [[]|] eqn:hl' => //. + 1,3,4:eapply value_constructor; tea; [erewrite <-lookup_constructor_trans; tea|now len|solve_all]. + now do 2 constructor. + - intros f args vh harglen hargs ihargs. + rewrite wellformed_mkApps // => /andP[] wff wfargs. + rewrite trans_mkApps. + depelim vh. congruence. + cbn. + simpl in wff; move/andP: wff => [] hascof /andP[] /Nat.ltb_lt wfidx wffix. + apply nth_error_Some in wfidx. + destruct nth_error eqn:heq => //. + all: apply trust_cofix. +Qed. + +Ltac destruct_times := + match goal with + | [ H : pair _ _ |- _ ] => destruct H + | [ H : MCProd.and3 _ _ _ |- _ ] => destruct H + | [ H : MCProd.and4 _ _ _ _ |- _ ] => destruct H + | [ H : MCProd.and5 _ _ _ _ _ |- _ ] => destruct H + | [ H : MCProd.and6 _ _ _ _ _ _ |- _ ] => destruct H + | [ H : MCProd.and7 _ _ _ _ _ _ _ |- _ ] => destruct H + | [ H : MCProd.and8 _ _ _ _ _ _ _ _ |- _ ] => destruct H + | [ H : MCProd.and9 _ _ _ _ _ _ _ _ _ |- _ ] => destruct H + | [ H : MCProd.and10 _ _ _ _ _ _ _ _ _ _ |- _ ] => destruct H + end. + +From MetaCoq.Erasure Require Import EWcbvEvalCstrsAsBlocksInd. +Lemma trans_correct {efl : EEnvFlags} {fl} {wcon : with_constructor_as_block = true} + {wcb : cstr_as_blocks = true} + {Σ : GlobalContextMap.t} t v : + has_tApp -> + wf_glob Σ -> + closed_env Σ -> + @Ee.eval fl Σ t v -> + wellformed Σ 0 t -> + @Ee.eval fl (trans_env Σ) (trans Σ t) (trans Σ v). +Proof. + intros hasapp wfΣ clΣ ev wf. + revert t v wf ev. + eapply + (eval_preserve_mkApps_ind fl wcon (efl := efl) Σ _ + (wellformed Σ) (Qpres := Qpreserves_wellformed efl _ wfΣ)) => //; eauto; + intros; repeat destruct_times; try solve [econstructor; eauto 3]. + + - intros. eapply eval_wellformed in H; tea. + + - econstructor; eauto. + rewrite trans_csubst // in e. now eapply wellformed_closed. + + - rewrite trans_csubst // in e. now eapply wellformed_closed. + cbn. econstructor; eauto. + + - assert (forallb (wellformed Σ 0) args). + cbn -[lookup_constructor lookup_constructor_pars_args] in i2. + rewrite wcb in i2. move/and3P: i2 => [] _ _ hargs. + solve_all. + rewrite trans_iota_red // in e. + { solve_all. now eapply wellformed_closed. } + cbn -[lookup_constructor lookup_constructor_pars_args + GlobalContextMap.lookup_inductive_kind] in e0 |- *. + eapply eval_closed in H as evc => //. + 2:{ now eapply wellformed_closed. } + rewrite GlobalContextMap.lookup_inductive_kind_spec in e0 *. + destruct lookup_inductive_kind as [[]|] eqn:hl => //. + 1,3,4:eapply eval_iota_block; eauto; + [now rewrite -is_propositional_cstr_trans| + rewrite nth_error_map H2 //|now len| + try (cbn; rewrite -H4 !skipn_length map_length //)]. + eapply eval_iota_block. + 1,3,4: eauto. + + now rewrite -is_propositional_cstr_trans. + + rewrite nth_error_map H2 //. + + eapply eval_beta. eapply e0; eauto. + constructor; eauto. + rewrite closed_subst // simpl_subst_k //. + eapply Ee.eval_to_value in H. + eapply value_constructor_as_block in H. + eapply constructor_isprop_pars_decl_constructor in H1 as [mdecl [idecl [hl' hp]]]. + econstructor; eauto. + now erewrite <-lookup_constructor_trans. len. + now rewrite /cstr_arity. + instantiate (1 := map (trans Σ) args). + eapply All2_All2_Set. + eapply values_final. solve_all. + unshelve eapply value_trans; tea. + + now len. + + now len. + + exact e. + + - subst brs. + cbn -[lookup_constructor lookup_constructor_pars_args + GlobalContextMap.lookup_inductive_kind] in e0 |- *. + rewrite GlobalContextMap.lookup_inductive_kind_spec. + rewrite trans_substl ?map_repeat /= in e. + { now apply forallb_repeat. } + destruct lookup_inductive_kind as [[]|] eqn:hl => //. + 1,3,4:eapply eval_iota_sing; [eauto|eauto| + now rewrite -is_propositional_trans|reflexivity| + rewrite /= ?trans_substl //; simpl; eauto]. + eapply eval_iota_sing; eauto. + eapply eval_box; eauto. + rewrite -is_propositional_trans //. + reflexivity. + + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. +Qed. +(* + - rewrite trans_mkApps in e1. + 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 trans_cunfold_fix; tea. + eapply closed_fix_subst. tea. + rewrite trans_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. + + - move/andP => [] clf cla. + eapply eval_closed in ev1 => //. + rewrite closedn_mkApps in ev1. + move: ev1 => /andP [] clfix clargs. + eapply eval_closed in ev2; tas. + rewrite trans_mkApps in IHev1 |- *. + simpl in *. eapply Ee.eval_fix_value. auto. auto. auto. + eapply trans_cunfold_fix; eauto. + eapply closed_fix_subst => //. + now rewrite map_length. + + - move/andP => [] clf cla. + eapply eval_closed in ev1 => //. + eapply eval_closed in ev2; tas. + simpl in *. eapply Ee.eval_fix'. auto. auto. + eapply trans_cunfold_fix; eauto. + eapply closed_fix_subst => //. + eapply IHev2; tea. eapply IHev3. + apply/andP; split => //. + eapply closed_cunfold_fix; tea. + + - move/andP => [] cd clbrs. specialize (IHev1 cd). + rewrite closedn_mkApps in IHev2. + move: (eval_closed _ clΣ _ _ cd ev1). + rewrite closedn_mkApps. + move/andP => [] clfix clargs. + forward IHev2. + { rewrite clargs clbrs !andb_true_r. + eapply closed_cunfold_cofix; tea. } + rewrite -> trans_mkApps in IHev1, IHev2. simpl. + rewrite GlobalContextMap.lookup_inductive_kind_spec in IHev2 |- *. + destruct EGlobalEnv.lookup_inductive_kind as [[]|] eqn:isp => //. + simpl in IHev1. + eapply Ee.eval_cofix_case. tea. + apply trans_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + apply IHev2. + eapply Ee.eval_cofix_case; tea. + apply trans_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + simpl in *. + eapply Ee.eval_cofix_case; tea. + apply trans_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + eapply Ee.eval_cofix_case; tea. + apply trans_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + + - intros cd. specialize (IHev1 cd). + move: (eval_closed _ clΣ _ _ cd ev1). + rewrite closedn_mkApps; move/andP => [] clfix clargs. forward IHev2. + { 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 -> trans_mkApps in IHev1, IHev2. simpl in *. + econstructor; eauto. + apply trans_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + rewrite -> trans_mkApps in IHev1, IHev2. simpl in *. + econstructor; eauto. + apply trans_cunfold_cofix; tea. eapply closed_cofix_subst; tea. + + - rewrite /declared_constant in isdecl. + move: (lookup_env_trans c wfΣ). + rewrite isdecl /= //. + intros hl. + econstructor; tea. cbn. rewrite e //. + apply IHev. + eapply lookup_env_closed in clΣ; tea. + move: clΣ. rewrite /closed_decl e //. + + - move=> cld. + eapply eval_closed in ev1; tea. + move: ev1; rewrite closedn_mkApps /= => clargs. + rewrite GlobalContextMap.inductive_isprop_and_pars_spec. + rewrite (constructor_isprop_pars_decl_inductive e1). + rewrite trans_mkApps in IHev1. + specialize (IHev1 cld). + eapply Ee.eval_proj; tea. + now rewrite -is_propositional_cstr_trans. + now len. rewrite nth_error_map e3 //. + eapply IHev2. + eapply nth_error_forallb in e3; tea. + + - congruence. + + - rewrite GlobalContextMap.inductive_isprop_and_pars_spec. + now rewrite e0. + + - move/andP=> [] clf cla. + rewrite trans_mkApps. + eapply eval_construct; tea. + rewrite -lookup_constructor_trans //. exact e0. + rewrite trans_mkApps in IHev1. now eapply IHev1. + now len. + now eapply IHev2. + + - congruence. + + - move/andP => [] clf cla. + specialize (IHev1 clf). specialize (IHev2 cla). + eapply Ee.eval_app_cong; eauto. + eapply Ee.eval_to_value in ev1. + destruct ev1; simpl in *; eauto. + * destruct t => //; rewrite trans_mkApps /=. + * destruct with_guarded_fix. + + move: i. + rewrite !negb_or. + rewrite trans_mkApps !isFixApp_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. + rewrite !andb_true_r. + rtoProp; intuition auto. + destruct v => /= //. + destruct v => /= //. + destruct v => /= //. + + move: i. + rewrite !negb_or. + rewrite trans_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. + destruct v => /= //. + - destruct t => //. + all:constructor; eauto. cbn [atom trans] in i |- *. + rewrite -lookup_constructor_trans //. destruct l => //. +Qed. +*) +From MetaCoq.Erasure Require Import EEtaExpanded. + +Lemma isLambda_trans Σ t : isLambda t -> isLambda (trans Σ t). +Proof. destruct t => //. Qed. +Lemma isBox_trans Σ t : isBox t -> isBox (trans Σ t). +Proof. destruct t => //. Qed. + +Lemma trans_expanded {Σ : GlobalContextMap.t} t : expanded Σ t -> expanded Σ (trans Σ t). +Proof. + induction 1 using expanded_ind. + all:try solve[constructor; eauto; solve_all]. + all:rewrite ?trans_mkApps. + - eapply expanded_mkApps_expanded => //. solve_all. + - cbn -[GlobalContextMap.lookup_inductive_kind]. + rewrite GlobalContextMap.lookup_inductive_kind_spec. + destruct lookup_inductive_kind as [[]|] => /= //. + 2-3:constructor; eauto; solve_all. + constructor; eauto; solve_all. cbn. + unfold Thunk.force. + eapply isEtaExp_expanded. + all:apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. + - apply trust_cofix. +Qed. + (*cbn. + eapply isEtaExp_substl. eapply forallb_repeat => //. + destruct branches as [|[]]; cbn in heq; noconf heq. + cbn -[isEtaExp] in *. depelim H1. cbn in H1. + now eapply expanded_isEtaExp. + constructor; eauto; solve_all. + depelim H1. depelim H1. do 2 (constructor; intuition auto). + solve_all. + - cbn -[GlobalContextMap.inductive_isprop_and_pars]. + rewrite GlobalContextMap.inductive_isprop_and_pars_spec. + destruct inductive_isprop_and_pars as [[[|] _]|] => /= //. + constructor. all:constructor; auto. + - cbn. eapply expanded_tFix. solve_all. + rewrite isLambda_trans //. + - eapply expanded_tConstruct_app; tea. + now len. solve_all. +Qed. +*) +Lemma trans_expanded_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded Σ t -> expanded (trans_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_trans // /= H //. 1-2:eauto. auto. solve_all. +Qed. + +Lemma trans_expanded_decl {Σ : GlobalContextMap.t} t : expanded_decl Σ t -> expanded_decl Σ (trans_decl Σ t). +Proof. + destruct t as [[[]]|] => /= //. + unfold expanded_constant_decl => /=. + apply trans_expanded. +Qed. + +Lemma trans_expanded_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded_decl Σ t -> expanded_decl (trans_env Σ) t. +Proof. + destruct t as [[[]]|] => /= //. + unfold expanded_constant_decl => /=. + apply trans_expanded_irrel. +Qed. + +Lemma trans_env_extends' {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : + extends_prefix Σ Σ' -> + wf_glob Σ' -> + List.map (on_snd (trans_decl Σ)) Σ.(GlobalContextMap.global_decls) = + List.map (on_snd (trans_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_trans_decl_extends => //. cbn. + eapply extends_wf_global_decl. 3:tea. + eapply extends_wf_glob; tea. eapply extends_prefix_extends; tea. + 2:{ eapply extends_wf_glob; tea. } + destruct hin. exists (x ++ [(kn, d)]). rewrite -app_assoc /= //. + cbn. eapply extends_prefix_extends; tea. +Qed. + +Lemma trans_env_eq {efl : EEnvFlags} (Σ : GlobalContextMap.t) : + wf_glob Σ -> trans_env Σ = trans_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). +Proof. + intros wf. + unfold trans_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_trans_decl_extends => //. cbn. now depelim wf. cbn. + eapply extends_prefix_extends; tea. 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 (trans_env_extends' (Σ := Σg') (Σ' := Σg)) => //. + cbn. now exists [a]. +Qed. + +Lemma trans_env_expanded {efl : EEnvFlags} {Σ : GlobalContextMap.t} : + wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (trans_env Σ). +Proof. + unfold expanded_global_env; move=> wfg. + rewrite trans_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 -(trans_env_eq Σ'). cbn. now depelim wfg. + eapply (trans_expanded_decl_irrel (Σ := Σ')). now depelim wfg. + now unshelve eapply (trans_expanded_decl (Σ:=Σ')). +Qed. + +Lemma trans_wellformed {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : + has_tBox -> has_tRel -> + wf_glob Σ -> EWellformed.wellformed Σ n t -> EWellformed.wellformed Σ n (trans Σ t). +Proof. + intros wfΣ hbox hrel. + induction t in n |- * using EInduction.term_forall_list_ind => //. + all:try solve [cbn; rtoProp; intuition auto; solve_all]. + all:apply trust_cofix. + (*- + cbn -[lookup_constructor]. intros. destruct cstr_as_blocks; rtoProp; repeat split; eauto. 2:solve_all. + 2: now destruct args; inv H0. len. eauto. + - cbn -[GlobalContextMap.inductive_isprop_and_pars lookup_inductive]. move/and3P => [] hasc /andP[]hs ht hbrs. + destruct GlobalContextMap.inductive_isprop_and_pars as [[[|] _]|] => /= //. + destruct l as [|[br n'] [|l']] eqn:eql; simpl. + all:rewrite ?hasc ?hs /= ?andb_true_r. + rewrite IHt //. + depelim X. cbn in hbrs. + rewrite andb_true_r in hbrs. + specialize (i _ hbrs). + eapply wellformed_substl => //. solve_all. eapply All_repeat => //. + now rewrite repeat_length. + cbn in hbrs; rtoProp; solve_all. depelim X; depelim X. solve_all. + do 2 depelim X. solve_all. + do 2 depelim X. solve_all. + rtoProp; solve_all. solve_all. + rtoProp; solve_all. solve_all. + - 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_trans. 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. *) +Qed. + +Import EWellformed. + +Lemma trans_wellformed_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : + wf_glob Σ -> + forall n, wellformed Σ n t -> wellformed (trans_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_trans //. + destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. subst g. + destruct (cst_body c) => //. + - rewrite lookup_env_trans //. + 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_trans //. + destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. subst g. + destruct nth_error => /= //. + intros; rtoProp; intuition auto; solve_all. + - rewrite lookup_env_trans //. + destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. + rewrite andb_false_r => //. + destruct nth_error => /= //. + all:intros; rtoProp; intuition auto; solve_all. +Qed. + +Lemma trans_wellformed_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} d : + wf_glob Σ -> + wf_global_decl Σ d -> wf_global_decl (trans_env Σ) d. +Proof. + intros wf; destruct d => /= //. + destruct (cst_body c) => /= //. + now eapply trans_wellformed_irrel. +Qed. + +Lemma trans_decl_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : + has_tBox -> has_tRel -> wf_glob Σ -> + forall d, wf_global_decl Σ d -> wf_global_decl (trans_env Σ) (trans_decl Σ d). +Proof. + intros hasb hasr wf d. + intros hd. + eapply trans_wellformed_decl_irrel; tea. + move: hd. + destruct d => /= //. + destruct (cst_body c) => /= //. + now eapply trans_wellformed => //. +Qed. + +Lemma fresh_global_trans_env {Σ : GlobalContextMap.t} kn : + fresh_global kn Σ -> + fresh_global kn (trans_env Σ). +Proof. + destruct Σ as [Σ map repr wf]; cbn in *. + induction 1; cbn; constructor; auto. + now eapply Forall_map; cbn. +Qed. + +Lemma trans_env_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : + has_tBox -> has_tRel -> + wf_glob Σ -> wf_glob (trans_env Σ). +Proof. + intros hasb hasrel. + intros wfg. rewrite trans_env_eq //. + destruct Σ as [Σ map repr wf]; cbn in *. + clear map repr. + induction wfg; cbn; constructor; auto. + - rewrite /= -(trans_env_eq (GlobalContextMap.make Σ (fresh_globals_cons_inv wf))) //. + eapply trans_decl_wf => //. + - rewrite /= -(trans_env_eq (GlobalContextMap.make Σ (fresh_globals_cons_inv wf))) //. + now eapply fresh_global_trans_env. +Qed. + +Definition trans_program (p : eprogram_env) := + (trans_env p.1, trans p.1 p.2). + +Definition trans_program_wf {efl} (p : eprogram_env) {hastbox : has_tBox} {hastrel : has_tRel} : + wf_eprogram_env efl p -> wf_eprogram efl (trans_program p). +Proof. + intros []; split. + now eapply trans_env_wf. + cbn. eapply trans_wellformed_irrel => //. now eapply trans_wellformed. +Qed. + +Definition trans_program_expanded {efl} (p : eprogram_env) : + wf_eprogram_env efl p -> + expanded_eprogram_env_cstrs p -> expanded_eprogram_cstrs (trans_program p). +Proof. + unfold expanded_eprogram_env_cstrs. + move=> [wfe wft] /andP[] etae etat. + apply/andP; split. + cbn. eapply expanded_global_env_isEtaExp_env, trans_env_expanded => //. + now eapply isEtaExp_env_expanded_global_env. + eapply expanded_isEtaExp. + eapply trans_expanded_irrel => //. + now apply trans_expanded, isEtaExp_expanded. +Qed. diff --git a/examples/metacoq_tour.v b/examples/metacoq_tour.v index 2f56ce470..3830aa3de 100644 --- a/examples/metacoq_tour.v +++ b/examples/metacoq_tour.v @@ -99,7 +99,7 @@ From MetaCoq.ErasurePlugin Require Import Erasure Loader. (** Running erasure live in Coq *) Definition test (p : Ast.Env.program) : string := - erase_and_print_template_program p. + erase_and_print_template_program default_erasure_config p. MetaCoq Quote Recursively Definition zero := 0. diff --git a/template-coq/src/ast_quoter.ml b/template-coq/src/ast_quoter.ml index 2a987dfb5..af1168407 100644 --- a/template-coq/src/ast_quoter.ml +++ b/template-coq/src/ast_quoter.ml @@ -243,16 +243,16 @@ struct let block = List.rev defs in Coq_tFix (block, b) - let mkCoFix (a,(ns,ts,ds)) = + let mkCoFix ((a,b),(ns,ts,ds)) = let mk_fun xs i = { dname = Array.get ns i ; dtype = Array.get ts i ; dbody = Array.get ds i ; - rarg = Datatypes.O } :: xs + rarg = Array.get a i } :: xs in let defs = List.fold_left mk_fun [] (seq 0 (Array.length ns)) in let block = List.rev defs in - Coq_tCoFix (block, a) + Coq_tCoFix (block, b) let mkCase (ind, npar, r) (univs, pars, pctx, pret) c brs = let info = { ci_ind = ind; ci_npar = npar; ci_relevance = r } in diff --git a/template-coq/src/constr_quoter.ml b/template-coq/src/constr_quoter.ml index 1bab1577f..4d447b739 100644 --- a/template-coq/src/constr_quoter.ml +++ b/template-coq/src/constr_quoter.ml @@ -59,14 +59,14 @@ struct let mkConstruct (ind, i) u = constr_mkApp (tConstructor, [| ind ; i ; u |]) - let mkCoFix (a,(ns,ts,ds)) = + let mkCoFix ((a,b),(ns,ts,ds)) = let mk_fun xs i = constr_mkApp (tmkdef, [| Lazy.force tTerm ; Array.get ns i ; - Array.get ts i ; Array.get ds i ; Lazy.force tO |]) :: xs + Array.get ts i ; Array.get ds i ; Array.get a i |]) :: xs in let defs = List.fold_left mk_fun [] (seq 0 (Array.length ns)) in let block = to_coq_list (constr_mkAppl (tdef, [| tTerm |])) (List.rev defs) in - constr_mkApp (tCoFix, [| block ; a |]) + constr_mkApp (tCoFix, [| block ; b |]) let mkInd i u = constr_mkApp (tInd, [| i ; u |]) diff --git a/template-coq/src/quoter.ml b/template-coq/src/quoter.ml index 3f93489a9..dffaa164f 100644 --- a/template-coq/src/quoter.ml +++ b/template-coq/src/quoter.ml @@ -66,7 +66,7 @@ sig t -> (quoted_aname array * t) list (* branches *) -> t val mkProj : quoted_proj -> t -> t val mkFix : (quoted_int array * quoted_int) * (quoted_aname array * t array * t array) -> t - val mkCoFix : quoted_int * (quoted_aname array * t array * t array) -> t + val mkCoFix : (quoted_int array * quoted_int) * (quoted_aname array * t array * t array) -> t val mkInt : quoted_int63 -> t val mkFloat : quoted_float64 -> t val mkArray : quoted_univ_level -> t array -> default:t -> ty:t -> t @@ -250,6 +250,12 @@ struct CArray.fold_left_map (fun acc t -> let (x, acc) = quote_term acc env sigma t in acc, x) acc ts in ts, acc + let cofixpoint_arities ts = + let cofix_arity t = + let ctx, _concl = Term.decompose_prod_assum t in + Context.Rel.nhyps ctx + in Array.map cofix_arity ts + let quote_term_remember (add_constant : KerName.t -> 'a -> 'a) (add_inductive : Names.inductive -> Declarations.mutual_inductive_body -> 'a -> 'a) = @@ -359,9 +365,11 @@ struct let a' = Array.map Q.quote_int a in let (b',decl'),acc = quote_recdecl acc env sigma b decl in (Q.mkFix ((a',b'),decl'), acc) - and quote_cofixpoint acc env sigma (a,decl) = + and quote_cofixpoint acc env sigma (a,(_, ts, _ as decl)) = + let arities = cofixpoint_arities ts in + let qarities = Array.map Q.quote_int arities in let (a',decl'),acc = quote_recdecl acc env sigma a decl in - (Q.mkCoFix (a',decl'), acc) + (Q.mkCoFix ((qarities,a'),decl'), acc) and quote_minductive_type (acc : 'a) env sigma (t : MutInd.t) mib = let uctx = get_abstract_inductive_universes mib.Declarations.mind_universes in let inst = Univ.UContext.instance uctx in diff --git a/test-suite/erasure_live_test.v b/test-suite/erasure_live_test.v index 57c9e17b3..c780c2996 100644 --- a/test-suite/erasure_live_test.v +++ b/test-suite/erasure_live_test.v @@ -18,7 +18,7 @@ Unset MetaCoq Debug. #[local] Existing Instance config.extraction_checker_flags. Definition test (p : Ast.Env.program) : string := - erase_and_print_template_program p. + erase_and_print_template_program default_erasure_config p. Definition testty (p : Ast.Env.program) : string := typed_erase_and_print_template_program p. @@ -51,7 +51,7 @@ Definition singlelim := ((fun (X : Set) (x : X) (e : x = x) => Definition erase {A} (a : A) : TemplateMonad unit := aq <- tmQuoteRec a ;; - s <- tmEval lazy (erase_and_print_template_program aq) ;; + s <- tmEval lazy (erase_and_print_template_program default_erasure_config aq) ;; tmMsg s. Definition erase_fast {A} (a : A) : TemplateMonad unit := @@ -80,6 +80,13 @@ Set MetaCoq Timing. Time MetaCoq Run (tmEval hnf vplus0123 >>= erase). Time MetaCoq Run (tmEval hnf vplus0123 >>= erase_fast). +(** Cofix *) +From Coq Require Import StreamMemo. + +MetaCoq Quote Recursively Definition memo := memo_make. + +Definition testmemo := Eval lazy in test memo. + (** Ackermann **) Fixpoint ack (n m:nat) {struct n} : nat := match n with @@ -350,7 +357,7 @@ Time Definition P_provedCopyx := Eval lazy in (test_fast cbv_provedCopyx). From MetaCoq.ErasurePlugin Require Import Loader. MetaCoq Erase provedCopyx. -MetaCoq Typed Erase provedCopyx. +MetaCoq Erase -time -typed -unsafe provedCopyx. (* From MetaCoq.Erasure.Typed Require Import CertifyingEta. MetaCoq Run (eta_expand_def @@ -360,6 +367,16 @@ provedCopy). *) Print P_provedCopyx. +From Coq Require Import Streams. + +CoFixpoint ones : Stream nat := Cons 1 ones. + +MetaCoq Erase ones. +MetaCoq Erase -unsafe ones. + +MetaCoq Erase -typed -time -unsafe (map S ones). + + (* 0.2s purely in the bytecode VM *) (*Time Definition P_provedCopyxvm' := Eval vm_compute in (test p_provedCopyx). *) (* Goal @@ -431,7 +448,7 @@ Proof. show_match. *) - +(* Debuggging From MetaCoq.Common Require Import Transform. From MetaCoq.Erasure.Typed Require Import ExtractionCorrectness. @@ -497,3 +514,4 @@ Import ETransform Optimize. MetaCoq Typed Erase provedCopyx. Eval lazy in testty cbv_provedCopyx. +*) \ No newline at end of file diff --git a/test-suite/erasure_test.v b/test-suite/erasure_test.v index 601963af6..ddea7abab 100644 --- a/test-suite/erasure_test.v +++ b/test-suite/erasure_test.v @@ -3,6 +3,8 @@ From MetaCoq.Template Require Import Loader. Set MetaCoq Timing. Local Open Scope string_scope. +MetaCoq Erase -help. + MetaCoq Erase nat. (* Environment is well-formed and Ind(Coq.Init.Datatypes.nat,0,[]) has type: ⧆ @@ -17,24 +19,26 @@ Environment is well-formed and Construct(Coq.Init.Datatypes.bool,0,0,[]) erases Construct(Coq.Init.Datatypes.bool,0,0) *) -MetaCoq Erase (exist _ 0 (eq_refl) : {x : nat | x = 0}). +MetaCoq Erase (exist (fun x => x = 0) 0 (eq_refl)). Definition test := (proj1_sig (exist (fun x => x = 0) 0 (eq_refl))). -MetaCoq Typed Erase test. +MetaCoq Erase -typed test. + +(** Cofix *) +From Coq Require Import StreamMemo. + +MetaCoq Quote Recursively Definition memo := memo_make. +MetaCoq Erase -typed -unsafe memo_make. -(* (* *) -(* Environment is well-formed and exist nat (fun x : nat => eq nat x O) O (eq_refl nat O):sig nat (fun x : nat => eq nat x O) erases to: *) -(* (fun f => f) (exist ∎ ∎ O ∎) *) -(* *) *) MetaCoq Erase (3 + 1). Universe i. -MetaCoq Fast Erase ((fun (X : Set) (x : X) => x) nat). +MetaCoq Erase ((fun (X : Set) (x : X) => x) nat). (** Check that optimization of singleton pattern-matchings work *) -MetaCoq Erase ((fun (X : Set) (x : X) (e : x = x) => +MetaCoq Erase ((fun (X : Set) (x : X) (e : x = x) => match e in eq _ x' return bool with | eq_refl => true end)). From 07c0a825a2a971704d5f09f9e828435aee96b83e Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Sat, 17 Feb 2024 10:52:24 +0100 Subject: [PATCH 08/28] Avoid Ee := EWcbvEval module aliases which result in ugly extraction issues (#1057) --- erasure-plugin/theories/ErasureCorrectness.v | 10 ++--- erasure/theories/ECoInductiveToInductive.v | 28 ++++++------- erasure/theories/EOptimizePropDiscr.v | 24 +++++------ erasure/theories/ErasureCorrectness.v | 42 +++++++++---------- erasure/theories/ErasureFunctionProperties.v | 20 ++++----- erasure/theories/ErasureProperties.v | 32 +++++++------- erasure/theories/Prelim.v | 2 - .../theories/Typed/ExtractionCorrectness.v | 4 +- erasure/theories/Typed/OptimizePropDiscr.v | 8 ++-- 9 files changed, 83 insertions(+), 87 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 62d1f21a9..5e020479a 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -1872,7 +1872,7 @@ Section PCUICErase. have prev : Transform.pre verified_lambdabox_pipeline (Σ', v''). { clear -wfl pre' H1. cbn in H1. destruct pre' as [[] []]. split; split => //=. - eapply Prelim.Ee.eval_wellformed; eauto. + eapply EWcbvEval.eval_wellformed; eauto. eapply EEtaExpandedFix.isEtaExp_expanded. eapply (@EEtaExpandedFix.eval_etaexp wfl); eauto. now eapply EEtaExpandedFix.expanded_global_env_isEtaExp_env. @@ -1927,7 +1927,7 @@ Section PCUICErase. have prev : Transform.pre verified_lambdabox_pipeline (Σ', v''). { clear -wfl pre' H1. cbn in H1. destruct pre' as [[] []]. split; split => //=. - eapply Prelim.Ee.eval_wellformed; eauto. + eapply EWcbvEval.eval_wellformed; eauto. eapply EEtaExpandedFix.isEtaExp_expanded. eapply (@EEtaExpandedFix.eval_etaexp wfl); eauto. now eapply EEtaExpandedFix.expanded_global_env_isEtaExp_env. @@ -2275,8 +2275,8 @@ Section pipeline_theorem. Variable Normalisation : (forall Σ, wf_ext Σ -> PCUICSN.NormalizationIn Σ). - Lemma verified_erasure_pipeline_lookup_env_in kn decl (efl := EInlineProjections.switch_no_params all_env_flags) - {has_rel : has_tRel} {has_box : has_tBox} + Lemma verified_erasure_pipeline_lookup_env_in kn decl (efl := EInlineProjections.switch_no_params all_env_flags) + {has_rel : has_tRel} {has_box : has_tBox} T (typing: ∥PCUICTyping.typing Σ [] t T∥): let Σ_t := (transform verified_erasure_pipeline (Σ, t) (precond _ _ _ _ expΣ expt typing _)).1 in EGlobalEnv.lookup_env Σ_t kn = Some decl -> @@ -2356,7 +2356,7 @@ Section pipeline_theorem. Let v_t := compile_value_box (PCUICExpandLets.trans_global_env Σ) v []. - Lemma verified_erasure_pipeline_extends (efl := EInlineProjections.switch_no_params all_env_flags) + Lemma verified_erasure_pipeline_extends (efl := EInlineProjections.switch_no_params all_env_flags) {has_rel : has_tRel} {has_box : has_tBox} : EGlobalEnv.extends Σ_v Σ_t. Proof. diff --git a/erasure/theories/ECoInductiveToInductive.v b/erasure/theories/ECoInductiveToInductive.v index 3132dd66c..5aff66bce 100644 --- a/erasure/theories/ECoInductiveToInductive.v +++ b/erasure/theories/ECoInductiveToInductive.v @@ -623,7 +623,7 @@ Lemma value_trans {efl : EEnvFlags} {fl : WcbvFlags} {hasc : cstr_as_blocks = tr value (trans_env Σ) (trans Σ c). Proof. intros hasapp wfg wf h. - revert c h wf. apply: Ee.value_values_ind. + revert c h wf. apply: EWcbvEval.value_values_ind. - intros t; destruct t => //; cbn -[lookup_constructor GlobalContextMap.lookup_inductive_kind]. all:try solve [intros; repeat constructor => //]. destruct args => //. @@ -673,9 +673,9 @@ Lemma trans_correct {efl : EEnvFlags} {fl} {wcon : with_constructor_as_block = t has_tApp -> wf_glob Σ -> closed_env Σ -> - @Ee.eval fl Σ t v -> + @EWcbvEval.eval fl Σ t v -> wellformed Σ 0 t -> - @Ee.eval fl (trans_env Σ) (trans Σ t) (trans Σ v). + @EWcbvEval.eval fl (trans_env Σ) (trans Σ t) (trans Σ v). Proof. intros hasapp wfΣ clΣ ev wf. revert t v wf ev. @@ -715,7 +715,7 @@ Proof. + eapply eval_beta. eapply e0; eauto. constructor; eauto. rewrite closed_subst // simpl_subst_k //. - eapply Ee.eval_to_value in H. + eapply EWcbvEval.eval_to_value in H. eapply value_constructor_as_block in H. eapply constructor_isprop_pars_decl_constructor in H1 as [mdecl [idecl [hl' hp]]]. econstructor; eauto. @@ -763,7 +763,7 @@ Qed. eapply eval_closed in ev1 => //. rewrite closedn_mkApps in ev1. move: ev1 => /andP [] clfix clargs. - eapply Ee.eval_fix; eauto. + eapply EWcbvEval.eval_fix; eauto. rewrite map_length. eapply trans_cunfold_fix; tea. eapply closed_fix_subst. tea. @@ -778,7 +778,7 @@ Qed. move: ev1 => /andP [] clfix clargs. eapply eval_closed in ev2; tas. rewrite trans_mkApps in IHev1 |- *. - simpl in *. eapply Ee.eval_fix_value. auto. auto. auto. + simpl in *. eapply EWcbvEval.eval_fix_value. auto. auto. auto. eapply trans_cunfold_fix; eauto. eapply closed_fix_subst => //. now rewrite map_length. @@ -786,7 +786,7 @@ Qed. - move/andP => [] clf cla. eapply eval_closed in ev1 => //. eapply eval_closed in ev2; tas. - simpl in *. eapply Ee.eval_fix'. auto. auto. + simpl in *. eapply EWcbvEval.eval_fix'. auto. auto. eapply trans_cunfold_fix; eauto. eapply closed_fix_subst => //. eapply IHev2; tea. eapply IHev3. @@ -805,15 +805,15 @@ Qed. rewrite GlobalContextMap.lookup_inductive_kind_spec in IHev2 |- *. destruct EGlobalEnv.lookup_inductive_kind as [[]|] eqn:isp => //. simpl in IHev1. - eapply Ee.eval_cofix_case. tea. + eapply EWcbvEval.eval_cofix_case. tea. apply trans_cunfold_cofix; tea. eapply closed_cofix_subst; tea. apply IHev2. - eapply Ee.eval_cofix_case; tea. + eapply EWcbvEval.eval_cofix_case; tea. apply trans_cunfold_cofix; tea. eapply closed_cofix_subst; tea. simpl in *. - eapply Ee.eval_cofix_case; tea. + eapply EWcbvEval.eval_cofix_case; tea. apply trans_cunfold_cofix; tea. eapply closed_cofix_subst; tea. - eapply Ee.eval_cofix_case; tea. + eapply EWcbvEval.eval_cofix_case; tea. apply trans_cunfold_cofix; tea. eapply closed_cofix_subst; tea. - intros cd. specialize (IHev1 cd). @@ -845,7 +845,7 @@ Qed. rewrite (constructor_isprop_pars_decl_inductive e1). rewrite trans_mkApps in IHev1. specialize (IHev1 cld). - eapply Ee.eval_proj; tea. + eapply EWcbvEval.eval_proj; tea. now rewrite -is_propositional_cstr_trans. now len. rewrite nth_error_map e3 //. eapply IHev2. @@ -868,8 +868,8 @@ Qed. - move/andP => [] clf cla. specialize (IHev1 clf). specialize (IHev2 cla). - eapply Ee.eval_app_cong; eauto. - eapply Ee.eval_to_value in ev1. + eapply EWcbvEval.eval_app_cong; eauto. + eapply EWcbvEval.eval_to_value in ev1. destruct ev1; simpl in *; eauto. * destruct t => //; rewrite trans_mkApps /=. * destruct with_guarded_fix. diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 7257c75df..673b9223e 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -544,9 +544,9 @@ Qed. 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 -> + @EWcbvEval.eval fl Σ t v -> closed t -> - @Ee.eval (disable_prop_cases fl) (remove_match_on_box_env Σ) (remove_match_on_box Σ t) (remove_match_on_box Σ v). + @EWcbvEval.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 *. @@ -599,7 +599,7 @@ Proof. eapply eval_closed in ev1 => //. rewrite closedn_mkApps in ev1. move: ev1 => /andP [] clfix clargs. - eapply Ee.eval_fix; eauto. + eapply EWcbvEval.eval_fix; eauto. rewrite map_length. eapply remove_match_on_box_cunfold_fix; tea. eapply closed_fix_subst. tea. @@ -614,7 +614,7 @@ Proof. move: ev1 => /andP [] clfix clargs. eapply eval_closed in ev2; tas. rewrite remove_match_on_box_mkApps in IHev1 |- *. - simpl in *. eapply Ee.eval_fix_value. auto. auto. auto. + simpl in *. eapply EWcbvEval.eval_fix_value. auto. auto. auto. eapply remove_match_on_box_cunfold_fix; eauto. eapply closed_fix_subst => //. now rewrite map_length. @@ -622,7 +622,7 @@ Proof. - move/andP => [] clf cla. eapply eval_closed in ev1 => //. eapply eval_closed in ev2; tas. - simpl in *. eapply Ee.eval_fix'. auto. auto. + simpl in *. eapply EWcbvEval.eval_fix'. auto. auto. eapply remove_match_on_box_cunfold_fix; eauto. eapply closed_fix_subst => //. eapply IHev2; tea. eapply IHev3. @@ -642,15 +642,15 @@ Proof. 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. + eapply EWcbvEval.eval_cofix_case. tea. apply remove_match_on_box_cunfold_cofix; tea. eapply closed_cofix_subst; tea. apply IHev2. - eapply Ee.eval_cofix_case; tea. + eapply EWcbvEval.eval_cofix_case; tea. apply remove_match_on_box_cunfold_cofix; tea. eapply closed_cofix_subst; tea. simpl in *. - eapply Ee.eval_cofix_case; tea. + eapply EWcbvEval.eval_cofix_case; tea. apply remove_match_on_box_cunfold_cofix; tea. eapply closed_cofix_subst; tea. - eapply Ee.eval_cofix_case; tea. + eapply EWcbvEval.eval_cofix_case; tea. apply remove_match_on_box_cunfold_cofix; tea. eapply closed_cofix_subst; tea. - intros cd. specialize (IHev1 cd). @@ -682,7 +682,7 @@ Proof. rewrite (constructor_isprop_pars_decl_inductive e1). rewrite remove_match_on_box_mkApps in IHev1. specialize (IHev1 cld). - eapply Ee.eval_proj; tea. + eapply EWcbvEval.eval_proj; tea. now rewrite -is_propositional_cstr_remove_match_on_box. now len. rewrite nth_error_map e3 //. eapply IHev2. @@ -705,8 +705,8 @@ Proof. - move/andP => [] clf cla. specialize (IHev1 clf). specialize (IHev2 cla). - eapply Ee.eval_app_cong; eauto. - eapply Ee.eval_to_value in ev1. + eapply EWcbvEval.eval_app_cong; eauto. + eapply EWcbvEval.eval_to_value in ev1. destruct ev1; simpl in *; eauto. * depelim a0. destruct t => //; rewrite ?remove_match_on_box_mkApps /=. cbn in i. now rewrite !orb_false_r orb_true_r in i. diff --git a/erasure/theories/ErasureCorrectness.v b/erasure/theories/ErasureCorrectness.v index fa16aa49f..4a12486ef 100644 --- a/erasure/theories/ErasureCorrectness.v +++ b/erasure/theories/ErasureCorrectness.v @@ -265,8 +265,8 @@ Proof. now eapply isErasable_Proof. } 2:{ exists x2. split; eauto. constructor. eapply eval_iota_sing => //. - pose proof (Ee.eval_to_value _ _ He_v'). - let X0 := match goal with H : Ee.value _ (EAst.mkApps _ _) |- _ => H end in + pose proof (EWcbvEval.eval_to_value _ _ He_v'). + let X0 := match goal with H : EWcbvEval.value _ (EAst.mkApps _ _) |- _ => H end in eapply value_app_inv in X0. subst. eassumption. depelim H2. eapply isErasable_Propositional in X0; eauto. @@ -446,7 +446,7 @@ Proof. constructor. destruct x1 as [n br']. eapply eval_iota_sing => //. - pose proof (Ee.eval_to_value _ _ He_v') as X0. + pose proof (EWcbvEval.eval_to_value _ _ He_v') as X0. apply value_app_inv in X0; subst x0. apply He_v'. now rewrite -eq_npars. @@ -517,7 +517,7 @@ Proof. eapply isErasable_Proof. constructor. eauto. eapply eval_proj_prop => //. - pose proof (Ee.eval_to_value _ _ Hty_vc') as X0. + pose proof (EWcbvEval.eval_to_value _ _ Hty_vc') as X0. eapply value_app_inv in X0. subst. eassumption. now rewrite -eqpars. * rename H3 into Hinf. @@ -534,7 +534,7 @@ Proof. destruct (declared_projection_inj d d0) as [? [? []]]; subst mdecl0 idecl0 pdecl0 cdecl. invs H2. -- exists x9. split; eauto. constructor. - eapply Ee.eval_proj; eauto. rewrite -eqpars. + eapply EWcbvEval.eval_proj; eauto. rewrite -eqpars. erewrite isPropositional_propositional_cstr; eauto. move/negbTE: H9 => ->. reflexivity. eapply declared_constructor_from_gen. apply d0. cbn. eapply decli'. cbn. rewrite -lenx5 //. @@ -558,7 +558,7 @@ Proof. eassumption. eapply isErasable_Proof. constructor. eapply eval_proj_prop => //. - pose proof (Ee.eval_to_value _ _ Hty_vc') as X0. + pose proof (EWcbvEval.eval_to_value _ _ Hty_vc') as X0. eapply value_app_inv in X0. subst. eassumption. now rewrite -eqpars. -- eapply erases_deps_eval in Hty_vc'; [|now eauto]. @@ -675,11 +675,11 @@ Proof. - now eapply erases_deps_eval in ev_arg; eauto. } exists x3. split. eauto. - constructor. eapply Ee.eval_fix. + constructor. eapply EWcbvEval.eval_fix. ++ eauto. ++ eauto. ++ eauto. - ++ rewrite <- Ee.closed_unfold_fix_cunfold_eq. + ++ rewrite <- EWcbvEval.closed_unfold_fix_cunfold_eq. { unfold EGlobalEnv.unfold_fix. rewrite e. eapply All2_nth_error in Hmfix' as []; tea. cbn in *. rewrite -eargsv -(Forall2_length H4); trea. } @@ -727,7 +727,7 @@ Proof. eapply eval_fix; eauto. 1-2:eapply value_final, eval_to_value; eauto. rewrite /cunfold_fix e0 //. congruence. - ++ constructor. eapply Ee.eval_box; [|now eauto]. + ++ constructor. eapply EWcbvEval.eval_box; [|now eauto]. apply eval_to_mkApps_tBox_inv in ev_stuck as ?; subst. eauto. -- eauto. @@ -744,7 +744,7 @@ Proof. eapply erases_App in He as [(-> & [])|(? & ? & -> & ? & ?)]. + exists E.tBox. - split; [|now constructor; eauto using @Ee.eval]. + split; [|now constructor; eauto using @EWcbvEval.eval]. constructor. eapply Is_type_red. * eauto. @@ -765,7 +765,7 @@ Proof. rewrite -> !app_nil_r in *. cbn in *. exists E.tBox. - split; [|now constructor; eauto using @Ee.eval]. + split; [|now constructor; eauto using @EWcbvEval.eval]. eapply (Is_type_app _ _ _ [av]) in X as []. -- constructor. apply X. @@ -779,7 +779,7 @@ Proof. unfold cunfold_fix in *. destruct (nth_error _ _) eqn:nth; [|congruence]. eapply All2_nth_error_Some in X as (?&?&[? ? ? ? ?]); [|eauto]. - constructor. eapply Ee.eval_fix_value. + constructor. eapply EWcbvEval.eval_fix_value. ++ eauto. ++ eauto. ++ eauto. @@ -788,7 +788,7 @@ Proof. -- exists E.tBox. apply eval_to_mkApps_tBox_inv in H3 as ?; subst. - split; [|now constructor; eauto using @Ee.eval]. + split; [|now constructor; eauto using @EWcbvEval.eval]. eapply Is_type_app in X as []. ++ constructor. rewrite <- mkApps_snoc. @@ -897,13 +897,13 @@ Proof. - now apply Forall_All, Forall_erases_deps_cofix_subst; eauto. - now eapply nth_error_forall in H2; eauto. } exists v'. split => //. split. - eapply Ee.eval_cofix_case; tea. + eapply EWcbvEval.eval_cofix_case; tea. rewrite /EGlobalEnv.cunfold_cofix nth' //. f_equal. f_equal. - rewrite -(Ee.closed_cofix_substl_subst_eq (idx:=idx)) //. } + rewrite -(EWcbvEval.closed_cofix_substl_subst_eq (idx:=idx)) //. } { eapply eval_to_mkApps_tBox_inv in H1 as H'; subst L'; cbn in *. depelim hl'. edestruct IHeval1 as (? & ? & [?]); tea. - pose proof (Ee.eval_deterministic H1 H3). subst x0. + pose proof (EWcbvEval.eval_deterministic H1 H3). subst x0. eapply erases_deps_eval in Hed; tea. specialize (IHeval2 (E.tCase (ip, ci_npar ip) E.tBox brs')). forward IHeval2. @@ -1017,13 +1017,13 @@ Proof. - now apply Forall_All, Forall_erases_deps_cofix_subst; eauto. - now eapply nth_error_forall in H1; eauto. } exists v'. split => //. split. - eapply Ee.eval_cofix_proj; tea. + eapply EWcbvEval.eval_cofix_proj; tea. rewrite /EGlobalEnv.cunfold_cofix nth' //. f_equal. f_equal. - rewrite -(Ee.closed_cofix_substl_subst_eq (idx:=idx)) //. } + rewrite -(EWcbvEval.closed_cofix_substl_subst_eq (idx:=idx)) //. } { eapply eval_to_mkApps_tBox_inv in H2 as H'; subst L'; cbn in *. depelim hl'. edestruct IHeval1 as (? & ? & [?]); tea. - pose proof (Ee.eval_deterministic H3 H2). subst x0. + pose proof (EWcbvEval.eval_deterministic H3 H2). subst x0. eapply erases_deps_eval in Hed; tea. specialize (IHeval2 (E.tProj p E.tBox)). forward IHeval2. @@ -1083,7 +1083,7 @@ Proof. eapply erases_deps_eval in Hed1; tea. eapply erases_deps_mkApps_inv in Hed1 as []. depelim H8. - constructor. eapply Ee.eval_construct; tea. eauto. + constructor. eapply EWcbvEval.eval_construct; tea. eauto. eapply (EGlobalEnv.declared_constructor_lookup H9). rewrite -(Forall2_length H7). rewrite /EAst.cstr_arity. @@ -1129,7 +1129,7 @@ Proof. eapply wcbveval_red; eauto. * exists (E.tApp x2 x3). split; [econstructor; eauto|]. - constructor; eapply Ee.eval_app_cong; eauto. + constructor; eapply EWcbvEval.eval_app_cong; eauto. eapply ssrbool.negbT. repeat eapply orb_false_intro. -- destruct x2; try reflexivity. diff --git a/erasure/theories/ErasureFunctionProperties.v b/erasure/theories/ErasureFunctionProperties.v index 8c1b5163d..8b6d3bf82 100644 --- a/erasure/theories/ErasureFunctionProperties.v +++ b/erasure/theories/ErasureFunctionProperties.v @@ -650,7 +650,7 @@ Proof. 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) +Lemma erase_correct (wfl := EWcbvEval.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 Σ} @@ -1951,7 +1951,7 @@ Proof. 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 : +Lemma erasable_tBox_value (wfl := EWcbvEval.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. @@ -1960,7 +1960,7 @@ Proof. eapply Is_type_eval_inv; eauto. eexists; eauto. Qed. -Lemma erase_eval_to_box (wfl := Ee.default_wcbv_flags) +Lemma erase_eval_to_box (wfl := EWcbvEval.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 Σ} @@ -1972,12 +1972,12 @@ 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 Σ'.1 t' E.tBox -> ∥ isErasable Σext [] t ∥. + @EWcbvEval.eval EWcbvEval.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 (EWcbvEval.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 [?]. @@ -1986,7 +1986,7 @@ Proof. apply X0; eauto. Qed. -Lemma erase_eval_to_box_eager (wfl := Ee.default_wcbv_flags) +Lemma erase_eval_to_box_eager (wfl := EWcbvEval.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 Σ} @@ -1998,7 +1998,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 Σ'.1 t' E.tBox -> t' = E.tBox. + @EWcbvEval.eval EWcbvEval.default_wcbv_flags Σ'.1 t' E.tBox -> t' = E.tBox. Proof. intros wt. intros. @@ -2166,7 +2166,7 @@ Proof. 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) +Lemma erase_correct_strong' (wfl := EWcbvEval.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 Σ} @@ -2260,7 +2260,7 @@ Proof. Unshelve. eauto. Qed. -Lemma erase_correct_strong (wfl := Ee.default_wcbv_flags) X_type (X : X_type.π1) +Lemma erase_correct_strong (wfl := EWcbvEval.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 Σ} @@ -2301,7 +2301,7 @@ Proof. now exists T. Qed. -Definition erase_correct_firstorder (wfl := Ee.default_wcbv_flags) +Definition erase_correct_firstorder (wfl := EWcbvEval.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 Σ} : diff --git a/erasure/theories/ErasureProperties.v b/erasure/theories/ErasureProperties.v index 0edf1a6ec..cce1e4af8 100644 --- a/erasure/theories/ErasureProperties.v +++ b/erasure/theories/ErasureProperties.v @@ -671,7 +671,7 @@ Proof. depelim X0; depelim X1; repeat constructor; cbn; intuition eauto. solve_all. Qed. -Lemma eval_empty_brs {wfl : Ee.WcbvFlags} Σ ci p e : Σ ⊢ E.tCase ci p [] ⇓ e -> False. +Lemma eval_empty_brs {wfl : EWcbvEval.WcbvFlags} Σ ci p e : Σ ⊢ E.tCase ci p [] ⇓ e -> False. Proof. intros He. depind He. @@ -682,7 +682,7 @@ Proof. - cbn in i. discriminate. Qed. -Lemma eval_case_tBox_inv {wfl : Ee.WcbvFlags} {Σ ci e brs} : +Lemma eval_case_tBox_inv {wfl : EWcbvEval.WcbvFlags} {Σ ci e brs} : Σ ⊢ E.tCase ci EAst.tBox brs ⇓ e -> ∑ n br, brs = [(n, br)] × inductive_isprop_and_pars Σ ci.1 = Some (true, ci.2) × Σ ⊢ ECSubst.substl (repeat EAst.tBox #|n|) br ⇓ e. @@ -700,26 +700,26 @@ Proof. - cbn in i. discriminate. Qed. -Lemma eval_case_eval_discr {wfl : Ee.WcbvFlags} {Σ ci c c' e brs} : +Lemma eval_case_eval_discr {wfl : EWcbvEval.WcbvFlags} {Σ ci c c' e brs} : Σ ⊢ E.tCase ci c brs ⇓ e -> Σ ⊢ c ⇓ c' -> Σ ⊢ E.tCase ci c' brs ⇓ e. Proof. intros He Hc. depind He. - - pose proof (Ee.eval_deterministic He1 Hc). subst c'. - econstructor; eauto. now eapply Ee.value_final, Ee.eval_to_value. - - pose proof (Ee.eval_deterministic He1 Hc). subst c'. - eapply Ee.eval_iota_block; eauto. now eapply Ee.value_final, Ee.eval_to_value. - - pose proof (Ee.eval_deterministic He1 Hc). subst c'. - eapply Ee.eval_iota_sing; tea. now constructor. - - pose proof (Ee.eval_deterministic He1 Hc). subst c'. - eapply Ee.eval_cofix_case; tea. - now eapply Ee.value_final, Ee.eval_to_value. + - pose proof (EWcbvEval.eval_deterministic He1 Hc). subst c'. + econstructor; eauto. now eapply EWcbvEval.value_final, EWcbvEval.eval_to_value. + - pose proof (EWcbvEval.eval_deterministic He1 Hc). subst c'. + eapply EWcbvEval.eval_iota_block; eauto. now eapply EWcbvEval.value_final, EWcbvEval.eval_to_value. + - pose proof (EWcbvEval.eval_deterministic He1 Hc). subst c'. + eapply EWcbvEval.eval_iota_sing; tea. now constructor. + - pose proof (EWcbvEval.eval_deterministic He1 Hc). subst c'. + eapply EWcbvEval.eval_cofix_case; tea. + now eapply EWcbvEval.value_final, EWcbvEval.eval_to_value. - cbn in i. discriminate. Qed. -Lemma eval_case_eval_inv_discr {wfl : Ee.WcbvFlags} {Σ ci c c' e brs} : +Lemma eval_case_eval_inv_discr {wfl : EWcbvEval.WcbvFlags} {Σ ci c c' e brs} : Σ ⊢ E.tCase ci c brs ⇓ e -> Σ ⊢ c' ⇓ c -> Σ ⊢ E.tCase ci c' brs ⇓ e. @@ -731,13 +731,13 @@ Proof. - pose proof (eval_trans' Hc He1); subst discr. now econstructor; eauto. - pose proof (eval_trans' Hc He1); subst discr. - eapply Ee.eval_iota_sing; tea. + eapply EWcbvEval.eval_iota_sing; tea. - pose proof (eval_trans' Hc He1); subst discr. - eapply Ee.eval_cofix_case; tea. + eapply EWcbvEval.eval_cofix_case; tea. - cbn in i. discriminate. Qed. -Lemma eval_proj_eval_inv_discr {wfl : Ee.WcbvFlags} {Σ p c c' e} : +Lemma eval_proj_eval_inv_discr {wfl : EWcbvEval.WcbvFlags} {Σ p c c' e} : Σ ⊢ E.tProj p c ⇓ e -> Σ ⊢ c' ⇓ c -> Σ ⊢ E.tProj p c' ⇓ e. diff --git a/erasure/theories/Prelim.v b/erasure/theories/Prelim.v index 49d362695..02bbd94d9 100644 --- a/erasure/theories/Prelim.v +++ b/erasure/theories/Prelim.v @@ -93,8 +93,6 @@ Proof. simpl. apply IHt1. Qed. -Module Ee := EWcbvEval. - Lemma fst_decompose_app_rec t l : fst (EAstUtils.decompose_app_rec t l) = fst (EAstUtils.decompose_app t). Proof. induction t in l |- *; simpl; auto. rewrite IHt1. diff --git a/erasure/theories/Typed/ExtractionCorrectness.v b/erasure/theories/Typed/ExtractionCorrectness.v index 95b0fd34a..8f13f506d 100644 --- a/erasure/theories/Typed/ExtractionCorrectness.v +++ b/erasure/theories/Typed/ExtractionCorrectness.v @@ -507,8 +507,8 @@ Lemma optimize_correct `{EWellformed.EEnvFlags} Σ fgΣ t v : ELiftSubst.closed t = true -> EGlobalEnv.closed_env (trans_env Σ) = true -> EWellformed.wf_glob (trans_env Σ) -> - @Prelim.Ee.eval default_wcbv_flags (trans_env Σ) t v -> - @Prelim.Ee.eval + @EWcbvEval.eval default_wcbv_flags (trans_env Σ) t v -> + @EWcbvEval.eval (EWcbvEval.disable_prop_cases opt_wcbv_flags) (trans_env (map (on_snd (remove_match_on_box_decl (EEnvMap.GlobalContextMap.make (trans_env Σ) (OptimizePropDiscr.trans_env_fresh_globals _ fgΣ)))) Σ)) (EOptimizePropDiscr.remove_match_on_box (EEnvMap.GlobalContextMap.make (trans_env Σ) (OptimizePropDiscr.trans_env_fresh_globals _ fgΣ)) t) diff --git a/erasure/theories/Typed/OptimizePropDiscr.v b/erasure/theories/Typed/OptimizePropDiscr.v index 947efe2c0..3bbdbdef8 100644 --- a/erasure/theories/Typed/OptimizePropDiscr.v +++ b/erasure/theories/Typed/OptimizePropDiscr.v @@ -44,8 +44,6 @@ Next Obligation. now apply trans_env_fresh_globals. Qed. -Module Ee := EWcbvEval. - 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Σ)). @@ -64,13 +62,13 @@ Proof. now destruct o. Qed. -Lemma remove_match_on_box_correct `{EWellformed.EEnvFlags} `{Ee.WcbvFlags} Σ fgΣ t v : +Lemma remove_match_on_box_correct `{EWellformed.EEnvFlags} `{EWcbvEval.WcbvFlags} Σ fgΣ t v : EWcbvEval.with_constructor_as_block = false -> ELiftSubst.closed t = true -> EGlobalEnv.closed_env (trans_env Σ) = true -> EWellformed.wf_glob (trans_env Σ) -> - @Prelim.Ee.eval _ (trans_env Σ) t v -> - @Prelim.Ee.eval + @EWcbvEval.eval _ (trans_env Σ) t v -> + @EWcbvEval.eval (EWcbvEval.disable_prop_cases _) (trans_env (remove_match_on_box_env Σ fgΣ)) (remove_match_on_box (EEnvMap.GlobalContextMap.make (trans_env Σ) (trans_env_fresh_globals _ fgΣ)) t) From f23a18c853a5809fb3c5e5c923ecd64b48145edc Mon Sep 17 00:00:00 2001 From: nicolas tabareau Date: Mon, 19 Feb 2024 12:12:30 +0100 Subject: [PATCH 09/28] clean up in ErasureCorrectness --- erasure-plugin/theories/ErasureCorrectness.v | 258 +++++-------------- 1 file changed, 59 insertions(+), 199 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index 5e020479a..d7cee5569 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -21,6 +21,7 @@ Import Common.Transform.Transform. #[local] Obligation Tactic := program_simpl. #[local] Existing Instance extraction_checker_flags. +#[local] Existing Instance PCUICSN.extraction_normalizing. Import EWcbvEval. @@ -336,7 +337,7 @@ Section PCUICProof. now eapply PCUICWeakeningEnv.lookup_global_Some_fresh in H0. } Qed. - Lemma erase_tranform_firstorder (no := PCUICSN.extraction_normalizing) (wfl := default_wcbv_flags) + Lemma erase_tranform_firstorder (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), @@ -1694,7 +1695,7 @@ Section PCUICErase. match goal with [ |- context [ @erase ?X_type ?X ?nin ?G (tApp _ _) ?wt ] ] => unshelve epose proof (@erase_mkApps X_type X nin G t [u] wt (wt'_erase_pcuic_program (Σ, t) prf prf0)) - end. exact PCUICSN.extraction_normalizing. + end. assert (hargs : forall Σ : global_env_ext, Σ ∼_ext env -> ∥ All (welltyped Σ []) [u] ∥). { cbn; intros ? ->. do 2 constructor; auto. destruct prf. destruct prf2 as [[T HT]]. eapply PCUICInversion.inversion_App in HT as HT'. destruct HT' as [na [A [B [Hp []]]]]. now eexists. eapply w. } @@ -1741,7 +1742,7 @@ Section PCUICErase. match goal with [ |- context [ @erase ?X_type ?X ?nin ?G (tApp _ _) ?wt ] ] => unshelve epose proof (@erase_mkApps X_type X nin G t [u] wt (wt'_erase_pcuic_program (Σ, t) prf3 prf0)) - end. exact PCUICSN.extraction_normalizing. + end. assert (hargs : forall Σ : global_env_ext, Σ ∼_ext env -> ∥ All (welltyped Σ []) [u] ∥). { cbn; intros ? ->. do 2 constructor; auto. destruct prf4 as [[T HT]]. eexists; eapply HT. } specialize (H hargs). @@ -1825,8 +1826,6 @@ Section PCUICErase. now eapply (PCUICExpandLetsCorrectness.expand_lets_sound (cf := extraction_checker_flags)) in Hsort. Qed. - #[local] Existing Instance PCUICSN.extraction_normalizing. - (* Beware, this internally uses preservation of observations and determinism of evaluation from the canonical evaluation of [f] in the source to the evaluation in the target. *) @@ -2056,141 +2055,6 @@ Section PCUICErase. Transparent erase_transform. - - (* Lemma erase_transform_extends_app_construct {X_type} {X} {normalization_in} ind i u args wt : - let t := (mkApps (tConstruct ind i u) args) in - let trt := @erase_global_fast X_type X normalization_in [] t wt in - Forall (fun arg => exists wtarg, - let trarg := @erase X_type X normalization_in [] t wtarg in - EGlobalEnv.extends trarg.1 trt.1) args. - Proof. - intros t. - cbn. rewrite /erase_transform /=. - set (deps := term_global_deps _). - subst t. *) - - - (* Lemma erase_transform_extends_app_construct (Σ : global_env_ext_map) ind i u args pre : - let t := (mkApps (tConstruct ind i u) args) in - let trt := transform erase_transform (Σ, t) pre in - Forall (fun arg => exists pre', - let trarg := transform erase_transform (Σ, arg) pre' in - EGlobalEnv.extends trarg.1 trt.1) args. - Proof. - intros t. - cbn. rewrite /erase_transform /=. - set (deps := term_global_deps _). - subst t. - - induction args. constructor. constructor; eauto. - intros H. - assert (ner' : ~ ∥ isErasable Σ [] t ∥). - { destruct ner as [ner]. destruct pre, s. eapply nisErasable_spec in ner => //. eapply w. } - set (tr := transform _ _ _). - destruct tr eqn:heq. cbn -[transform]. - - Lemma erasure_pipeline_extends_app_construct (Σ : global_env_ext_map) ind i u args pre : - let t := (mkApps (tConstruct ind i u) args) in - ∥ nisErasable Σ [] t ∥ -> - PCUICEtaExpand.expanded Σ [] t -> - let trt := transform verified_erasure_pipeline (Σ, t) pre in - Forall (fun arg => exists pre', - let trarg := transform verified_erasure_pipeline (Σ, arg) pre' in - EGlobalEnv.extends trarg.1 trt.1) args. - Proof. - intros t ner exp. - unfold verified_erasure_pipeline. - destruct_compose. - set (K:= (fun p : global_env_ext_map => (wf_ext p -> PCUICSN.NormalizationIn p) /\ (wf_ext p -> PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn p))). - intros H. - assert (ner' : ~ ∥ isErasable Σ [] t ∥). - { destruct ner as [ner]. destruct pre, s. eapply nisErasable_spec in ner => //. eapply w. } - set (tr := transform _ _ _). - destruct tr eqn:heq. cbn -[transform]. - replace t0 with tr.2. assert (heq_env:tr.1=g) by now rewrite heq. subst tr. - 2:{ now rewrite heq. } - clear heq. - epose proof ( - erase_eta_app _ _ _ pre3) as H0. - pose proof (correctness (pcuic_expand_lets_transform K) (Σ, tApp t u) pre). - destruct H as [[wtapp] [expapp Kapp]]. - pose proof (correctness (pcuic_expand_lets_transform K) (Σ, t) pre'). - destruct H as [[wtt] [expt Kt]]. - forward H0. - { clear -wtapp ner eq. apply (f_equal snd) in eq. cbn [snd] in eq. rewrite -eq. - destruct pre as [[wtp] rest]. - destruct ner as [ner]. eapply (nisErasable_lets (Σ, tApp t u)) in ner. - eapply nisErasable_spec in ner => //. cbn. - apply wtapp. apply wtp. } - forward H0 by apply expt. - destruct H0 as [pre'0 [pre''0 [eta [extapp [extapp' heq]]]]]. - split. - { rewrite <- heq_env. cbn -[transform]. - pose proof (EProgram.TransformExt.preserves_obs _ _ _ (t:=verified_lambdabox_pipeline_extends')). - unfold extends_eprogram in H. - split. - { repeat (destruct_compose; intros). eapply verified_lambdabox_pipeline_extends. - repeat (destruct_compose; intros). cbn - [transform]. - generalize dependent pre3. rewrite <- eq. - cbn [transform pcuic_expand_lets_transform expand_lets_program]. - unfold expand_lets_program. cbn [fst snd]. - intros pre3. cbn in pre3. intros <-. intros. - assert (pre'0 = H1). apply proof_irrelevance. subst H1. - exact extapp. } - { repeat (destruct_compose; intros). eapply verified_lambdabox_pipeline_extends. - repeat (destruct_compose; intros). cbn - [transform]. - generalize dependent pre3. rewrite <- eq. - cbn [transform pcuic_expand_lets_transform expand_lets_program]. - unfold expand_lets_program. cbn [fst snd]. - intros pre3. cbn in pre3. intros <-. intros. - assert (pre''0 = H1). apply proof_irrelevance. subst H1. - exact extapp'. } } - set (tr := transform _ _ _). - destruct tr eqn:heqtr. cbn -[transform]. f_equal. - replace t1 with tr.2. subst tr. - 2:{ now rewrite heqtr; cbn. } - clear heqtr. - move: pre4. - rewrite heq. intros h. - epose proof (transform_lambda_box_eta_app _ _ _ h). - forward H. { cbn [fst snd]. - clear -eq eta extapp. revert pre3 extapp. - rewrite -eq. pose proof (correctness _ _ pre'0). - destruct H as [? []]. cbn [fst snd] in eta |- *. revert pre'0 H H0 H1 eta. rewrite eq. - intros. cbn -[transform] in H1. cbn -[transform]. - eapply EEtaExpandedFix.expanded_isEtaExp in H1. - eapply EEtaExpandedFix.isEtaExp_extends; tea. - pose proof (correctness _ _ pre3). apply H2. } - destruct H as [prelam [prelam' eqlam]]. rewrite eqlam. - rewrite snd_pair. clear eqlam. - destruct_compose_no_clear. - intros hlt heqlt. symmetry. - apply f_equal2. - eapply transform_lambdabox_pres_term. - split. rewrite fst_pair. - { destruct_compose_no_clear. intros H eq'. clear -extapp. - eapply extends_eq; tea. do 2 f_equal. clear extapp. - change (transform (pcuic_expand_lets_transform K) (Σ, tApp t u) pre).1 with - (transform (pcuic_expand_lets_transform K) (Σ, t) pre').1 in pre'0 |- *. - revert pre'0. - rewrite -surjective_pairing. intros pre'0. f_equal. apply proof_irrelevance. } - rewrite snd_pair. - destruct_compose_no_clear. intros ? ?. - eapply transform_erase_pres_term. - rewrite fst_pair. - { red. cbn. split => //. } reflexivity. - eapply transform_lambdabox_pres_term. - split. rewrite fst_pair. - { unfold run, time. destruct_compose_no_clear. intros H eq'. clear -extapp'. - assert (pre''0 = H). apply proof_irrelevance. subst H. apply extapp'. } - cbn [snd run]. unfold run, time. - destruct_compose_no_clear. intros ? ?. - eapply transform_erase_pres_term. cbn [fst]. - { red. cbn. split => //. } reflexivity. - Qed.*) - - Transparent erase_transform. - End PCUICErase. Lemma compile_evalue_strip (Σer : EEnvMap.GlobalContextMap.t) p : @@ -2214,9 +2078,6 @@ Arguments PCUICFirstorder.firstorder_ind _ _ : clear implicits. Section pipeline_cond. - Instance cf : checker_flags := extraction_checker_flags. - Instance nf : PCUICSN.normalizing_flags := PCUICSN.extraction_normalizing. - Variable Σ : global_env_ext_map. Variable t : PCUICAst.term. Variable T : PCUICAst.term. @@ -2252,33 +2113,15 @@ Section pipeline_cond. - cbn. intros wf ? ? ? ? ? ?. now eapply Normalisation. Qed. -End pipeline_cond. - - -Section pipeline_theorem. - - 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 axfree : axiom_free Σ. - Variable v : PCUICAst.term. - - Variable i : Kernames.inductive. - Variable u : Universes.Instance.t. - Variable args : list PCUICAst.term. + Let Σ_t := (transform verified_erasure_pipeline (Σ, t) precond).1. + Let t_t := (transform verified_erasure_pipeline (Σ, t) precond).2. + Let Σ_v := (transform verified_erasure_pipeline (Σ, v) precond2).1. + Let v_t := compile_value_box (PCUICExpandLets.trans_global_env Σ) v []. - Variable Normalisation : (forall Σ, wf_ext Σ -> PCUICSN.NormalizationIn Σ). + Opaque compose. - Lemma verified_erasure_pipeline_lookup_env_in kn decl (efl := EInlineProjections.switch_no_params all_env_flags) - {has_rel : has_tRel} {has_box : has_tBox} - T (typing: ∥PCUICTyping.typing Σ [] t T∥): - let Σ_t := (transform verified_erasure_pipeline (Σ, t) (precond _ _ _ _ expΣ expt typing _)).1 in + Lemma verified_erasure_pipeline_lookup_env_in kn decl (efl := EInlineProjections.switch_no_params all_env_flags) + {has_rel : has_tRel} {has_box : has_tBox} : EGlobalEnv.lookup_env Σ_t kn = Some decl -> exists decl', PCUICAst.PCUICEnvironment.lookup_global (PCUICExpandLets.trans_global_decls @@ -2286,8 +2129,7 @@ Section pipeline_theorem. /\ erase_decl_equal (fun decl => ERemoveParams.strip_inductive_decl (erase_mutual_inductive_body decl)) decl decl'. Proof. - Opaque compose. - unfold verified_erasure_pipeline. + unfold Σ_t, verified_erasure_pipeline. repeat rewrite -transform_compose_assoc. destruct_compose; intro. cbn. destruct_compose; intro. cbn. @@ -2343,6 +2185,25 @@ Section pipeline_theorem. now destruct decl' , decl''. Qed. +End pipeline_cond. + + +Section pipeline_theorem. + + 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 axfree : axiom_free Σ. + Variable v : PCUICAst.term. + + Variable i : Kernames.inductive. + Variable u : Universes.Instance.t. + Variable args : list PCUICAst.term. + + Variable Normalisation : (forall Σ, wf_ext Σ -> PCUICSN.NormalizationIn Σ). Variable typing : ∥PCUICTyping.typing Σ [] t (PCUICAst.mkApps (PCUICAst.tInd i u) args)∥. @@ -2355,41 +2216,40 @@ Section pipeline_theorem. Let Σ_v := (transform verified_erasure_pipeline (Σ, v) (precond2 _ _ _ _ expΣ expt typing _ _ Heval)).1. Let v_t := compile_value_box (PCUICExpandLets.trans_global_env Σ) v []. - - Lemma verified_erasure_pipeline_extends (efl := EInlineProjections.switch_no_params all_env_flags) - {has_rel : has_tRel} {has_box : has_tBox} : + Lemma verified_erasure_pipeline_extends (efl := EInlineProjections.switch_no_params all_env_flags) + {has_rel : has_tRel} {has_box : has_tBox} : EGlobalEnv.extends Σ_v Σ_t. Proof. - unfold Σ_v, Σ_t. unfold verified_erasure_pipeline. - repeat (destruct_compose; intro). destruct typing as [typing'], Heval. - cbn [transform compose pcuic_expand_lets_transform] in *. - unfold run, time. - cbn [transform erase_transform] in *. - set (erase_program _ _). set (erase_program _ _). - eapply verified_lambdabox_pipeline_extends. - eapply extends_erase_pcuic_program; eauto. - 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'. } - assumption. - now eapply PCUICExpandLetsCorrectness.trans_axiom_free. - pose proof (PCUICExpandLetsCorrectness.expand_lets_sound typing'). - rewrite PCUICExpandLetsCorrectness.trans_mkApps in X. eapply X. - move: fo. clear. - { rewrite /PCUICFirstorder.firstorder_ind /=. - rewrite PCUICExpandLetsCorrectness.trans_lookup. - destruct PCUICAst.PCUICEnvironment.lookup_env => //. - destruct g => //=. - eapply PCUICExpandLetsCorrectness.trans_firstorder_mutind. - eapply PCUICExpandLetsCorrectness.trans_firstorder_env. } + unfold Σ_v, Σ_t. unfold verified_erasure_pipeline. + repeat (destruct_compose; intro). destruct typing as [typing'], Heval. + cbn [transform compose pcuic_expand_lets_transform] in *. + unfold run, time. + cbn [transform erase_transform] in *. + set (erase_program _ _). set (erase_program _ _). + eapply verified_lambdabox_pipeline_extends. + eapply extends_erase_pcuic_program; eauto. + 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'. } + assumption. + now eapply PCUICExpandLetsCorrectness.trans_axiom_free. + pose proof (PCUICExpandLetsCorrectness.expand_lets_sound typing'). + rewrite PCUICExpandLetsCorrectness.trans_mkApps in X. eapply X. + move: fo. clear. + { rewrite /PCUICFirstorder.firstorder_ind /=. + rewrite PCUICExpandLetsCorrectness.trans_lookup. + destruct PCUICAst.PCUICEnvironment.lookup_env => //. + destruct g => //=. + eapply PCUICExpandLetsCorrectness.trans_firstorder_mutind. + eapply PCUICExpandLetsCorrectness.trans_firstorder_env. } Qed. Lemma fo_v : PCUICFirstorder.firstorder_value Σ [] v. Proof. - destruct typing, Heval. sq. - eapply PCUICFirstorder.firstorder_value_spec; eauto. - - eapply PCUICClassification.subject_reduction_eval; eauto. - - eapply PCUICWcbvEval.eval_to_value; eauto. + destruct typing, Heval. 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 _ _ _ _ expΣ expt typing _ _ Heval)).2. From 9e3e264be33e248918b553150f1f2c4accb74392 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 20 Feb 2024 13:50:43 +0100 Subject: [PATCH 10/28] Implement tLazy and tForce in EAst (#1058) * Add tLazy and tForce constructors With no evaluation semantics yet. Use them in ECoInductiveToInductive to allow efficient (unverified) implementation of cofixpoints in target languages. * Install archive file for static linking --- erasure/theories/EAst.v | 4 +- erasure/theories/EAstUtils.v | 4 + erasure/theories/ECSubst.v | 2 + erasure/theories/ECoInductiveToInductive.v | 142 +++++------------- erasure/theories/EConstructorsAsBlocks.v | 4 +- erasure/theories/EDeps.v | 6 + erasure/theories/EEtaExpanded.v | 22 ++- erasure/theories/EEtaExpandedFix.v | 8 +- erasure/theories/EImplementBox.v | 5 +- erasure/theories/EInduction.v | 20 ++- erasure/theories/EInlineProjections.v | 3 + erasure/theories/ELiftSubst.v | 12 +- erasure/theories/EOptimizePropDiscr.v | 2 + erasure/theories/EPretty.v | 2 + erasure/theories/EReflect.v | 2 + erasure/theories/ERemoveParams.v | 6 +- erasure/theories/ESpineView.v | 8 +- erasure/theories/EWcbvEval.v | 6 + .../EWcbvEvalCstrsAsBlocksFixLambdaInd.v | 6 +- erasure/theories/EWcbvEvalCstrsAsBlocksInd.v | 6 +- erasure/theories/EWcbvEvalEtaInd.v | 62 +++++--- erasure/theories/EWcbvEvalNamed.v | 62 +++++++- erasure/theories/EWellformed.v | 4 + erasure/theories/ErasureFunctionProperties.v | 2 + erasure/theories/Typed/Optimize.v | 10 ++ erasure/theories/Typed/OptimizeCorrectness.v | 20 +++ erasure/theories/Typed/TypeAnnotations.v | 2 + template-coq/Makefile.plugin.local-late | 1 + 28 files changed, 279 insertions(+), 154 deletions(-) create mode 100644 template-coq/Makefile.plugin.local-late diff --git a/erasure/theories/EAst.v b/erasure/theories/EAst.v index cd1eec17f..7620ba662 100644 --- a/erasure/theories/EAst.v +++ b/erasure/theories/EAst.v @@ -40,7 +40,9 @@ Inductive term : Set := | tProj (p : projection) (c : term) | tFix (mfix : mfixpoint term) (idx : nat) | tCoFix (mfix : mfixpoint term) (idx : nat) -| tPrim (prim : prim_val term). +| tPrim (prim : prim_val term) +| tLazy (t : term) +| tForce (t : term). Derive NoConfusion for term. diff --git a/erasure/theories/EAstUtils.v b/erasure/theories/EAstUtils.v index ec9068a60..de7c2655c 100644 --- a/erasure/theories/EAstUtils.v +++ b/erasure/theories/EAstUtils.v @@ -374,6 +374,8 @@ Fixpoint string_of_term (t : term) : string := | tFix l n => "Fix(" ^ (string_of_list (string_of_def string_of_term) l) ^ "," ^ string_of_nat n ^ ")" | tCoFix l n => "CoFix(" ^ (string_of_list (string_of_def string_of_term) l) ^ "," ^ string_of_nat n ^ ")" | tPrim p => "Prim(" ^ EPrimitive.string_of_prim string_of_term p ^ ")" + | tLazy t => "Lazy(" ^ string_of_term t ^ ")" + | tForce t => "Force(" ^ string_of_term t ^ ")" end. (** Compute all the global environment dependencies of the term *) @@ -409,5 +411,7 @@ Fixpoint term_global_deps (t : term) := KernameSet.union (KernameSet.singleton (inductive_mind p.(proj_ind))) (term_global_deps c) | tPrim p => prim_global_deps term_global_deps p + | tLazy t => term_global_deps t + | tForce t => term_global_deps t | _ => KernameSet.empty end. \ No newline at end of file diff --git a/erasure/theories/ECSubst.v b/erasure/theories/ECSubst.v index e05deb65d..3cc2d2881 100644 --- a/erasure/theories/ECSubst.v +++ b/erasure/theories/ECSubst.v @@ -38,6 +38,8 @@ Fixpoint csubst t k u := tCoFix mfix' idx | tConstruct ind n args => tConstruct ind n (map (csubst t k) args) | tPrim p => tPrim (map_prim (csubst t k) p) + | tLazy u => tLazy (csubst t k u) + | tForce u => tForce (csubst t k u) | x => x end. diff --git a/erasure/theories/ECoInductiveToInductive.v b/erasure/theories/ECoInductiveToInductive.v index 5aff66bce..485284539 100644 --- a/erasure/theories/ECoInductiveToInductive.v +++ b/erasure/theories/ECoInductiveToInductive.v @@ -33,43 +33,9 @@ Axiom trust_cofix : forall {A}, A. #[global] Hint Constructors eval : core. -Module Thunk. - Definition make (t : term) : term := - tLambda (nNamed "thunk") (lift 1 0 t). - - Definition force (t : term) := - tApp t tBox. - - Definition make_name (x : string) (n : nat) : string := - (x ++ string_of_nat n)%bs. - - (* Thunk an n-ary function: - [t] is supposed to be of type T0 -> ... -> Tn -> C here - and we want to produce an expansion: - λ x0 .. xn (), t x0 xn () *) - Equations make_n_aux (n : nat) (t : term) (acc : list term) : term := - make_n_aux 0 t acc => tLambda - (nNamed "thunk") - (mkApps (lift0 1 t) (rev (tRel 0 :: map (lift0 1) acc))); - make_n_aux (S n) t acc => - tLambda - (nNamed (make_name "x" (S n))) - (make_n_aux n (lift0 1 t) (tRel 0 :: map (lift0 1) acc)). - - Definition make_n (n : nat) (t : term) := make_n_aux n t []. - - (* Eval compute in make_n 2 (tRel 0). *) - -End Thunk. - Section trans. Context (Σ : GlobalContextMap.t). - Definition trans_cofix (d : def term) := - {| dname := d.(dname); - dbody := Thunk.make_n d.(rarg) d.(dbody); - rarg := d.(rarg) |}. - Fixpoint trans (t : term) : term := match t with | tRel i => tRel i @@ -81,12 +47,12 @@ Section trans. let brs' := List.map (on_snd trans) brs in match GlobalContextMap.lookup_inductive_kind Σ (fst ind).(inductive_mind) with | Some CoFinite => - tCase ind (Thunk.force (trans c)) brs' + tCase ind (tForce (trans c)) brs' | _ => tCase ind (trans c) brs' end | tProj p c => match GlobalContextMap.lookup_inductive_kind Σ p.(proj_ind).(inductive_mind) with - | Some CoFinite => tProj p (Thunk.force (trans c)) + | Some CoFinite => tProj p (tForce (trans c)) | _ => tProj p (trans c) end | tFix mfix idx => @@ -94,20 +60,18 @@ Section trans. tFix mfix' idx | tCoFix mfix idx => let mfix' := List.map (map_def trans) mfix in - let mfix' := List.map trans_cofix mfix' in - match nth_error mfix' idx with - | Some d => Thunk.make_n d.(rarg) (tFix mfix' idx) - | None => tCoFix mfix' idx - end + tFix mfix' idx | tBox => t | tVar _ => t | tConst _ => t | tConstruct ind i args => match GlobalContextMap.lookup_inductive_kind Σ ind.(inductive_mind) with - | Some CoFinite => Thunk.make (tConstruct ind i (map trans args)) + | Some CoFinite => tLazy (tConstruct ind i (map trans args)) | _ => tConstruct ind i (map trans args) end | tPrim p => tPrim (map_prim trans p) + | tLazy t => tLazy (trans t) + | tForce t => tForce (trans t) end. (* cofix succ x := match x with Stream x xs => Stream (x + 1) (succ xs) -> @@ -160,14 +124,9 @@ Section trans. unfold test_def in *; simpl closed in *; try solve [simpl subst; simpl closed; f_equal; auto; rtoProp; solve_all]; try easy. - destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //; solve_all. - rewrite -Nat.add_1_r. now eapply closedn_lift. - move/andP: H => [] clt clargs. destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //; rtoProp; solve_all; solve_all. - destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //; rtoProp; solve_all; solve_all. - - solve_all. destruct nth_error eqn:hnth. - * apply trust_cofix. - * cbn. unfold trans_cofix. len. solve_all. - unfold test_def. cbn. apply trust_cofix. - primProp. solve_all_k 6. Qed. @@ -219,25 +178,13 @@ Section trans. - destruct (k ?= n)%nat; auto. - destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //. 1,3,4:f_equal; rewrite map_map_compose; solve_all. - unfold Thunk.make. f_equal. cbn. - rewrite !map_map_compose. f_equal; solve_all. - specialize (H k cl). rewrite H. - rewrite closed_subst. now apply closed_trans. - rewrite closed_subst. now apply closed_trans. - now rewrite commut_lift_subst_rec. + do 2 f_equal; solve_all. - destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //. all:f_equal; eauto; try (rewrite /on_snd map_map_compose; solve_all). - unfold Thunk.force. f_equal; eauto. + f_equal. eauto. - destruct GlobalContextMap.lookup_inductive_kind as [[]|] => /= //. all:f_equal; eauto; try (rewrite /on_snd map_map_compose; solve_all). - unfold Thunk.force. f_equal; eauto. - - f_equal. solve_all. - rewrite !nth_error_map. destruct nth_error eqn:hnth => //=. - 2:{ f_equal. rewrite map_map_compose. eapply map_ext_in => x hin. - rewrite /trans_cofix /map_def //=. f_equal. len. - rewrite /Thunk.make_n. apply trust_cofix. - } - apply trust_cofix. + f_equal; eauto. Qed. Lemma trans_substl s t : @@ -284,11 +231,11 @@ Section trans. discriminate. Qed. - Lemma trans_cunfold_cofix mfix idx n f : + (* Lemma trans_cunfold_cofix mfix idx n f : forallb (closedn 0) (EGlobalEnv.cofix_subst mfix) -> cunfold_cofix mfix idx = Some (n, f) -> exists d, nth_error mfix idx = Some d /\ - cunfold_fix (map trans_cofix mfix) idx = Some (n, substl (fix_subst (map trans_cofix mfix)) (Thunk.make_n (rarg d) (dbody d))). + cunfold_fix mfix idx = Some (n, substl (fix_subst mfix) (dbody d)). Proof using Type. intros hcofix. unfold cunfold_cofix, cunfold_fix. @@ -296,7 +243,7 @@ Section trans. destruct nth_error. intros [= <- <-] => /=. f_equal. exists d. split => //. discriminate. - Qed. + Qed. *) Lemma trans_nth {n l d} : trans (nth n l d) = nth n (map trans l) (trans d). @@ -450,7 +397,6 @@ Proof. unfold lookup_inductive in hl. destruct lookup_minductive => //. rewrite (IHt _ H2 _ H0 H1) //. - - apply trust_cofix. Qed. Lemma wellformed_trans_decl_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : @@ -611,11 +557,6 @@ Proof. exists mdecl, idecl; split => //. Qed. -Lemma isLambda_make_n n t : isLambda (Thunk.make_n n t). -Proof. - induction n; cbn => //. -Qed. - Lemma value_trans {efl : EEnvFlags} {fl : WcbvFlags} {hasc : cstr_as_blocks = true} {wcon : with_constructor_as_block = true} {Σ : GlobalContextMap.t} {c} : has_tApp -> wf_glob Σ -> wellformed Σ 0 c -> @@ -628,11 +569,6 @@ Proof. all:try solve [intros; repeat constructor => //]. destruct args => //. move=> /andP[] wc. now rewrite wcon in wc. - move=> _ /andP [] hascof /andP[] /Nat.ltb_lt /nth_error_Some hnth hm. - destruct nth_error => //. - pose proof (isLambda_make_n (rarg d) (tFix (map trans_cofix mfix) idx)). - destruct Thunk.make_n => //. apply trust_cofix. - (* do 3 constructor. *) - intros p pv IH wf. cbn. constructor. constructor 2. cbn in wf. move/andP: wf => [hasp tp]. primProp. depelim tp; depelim pv; depelim IH; constructor; cbn in *; rtoProp; intuition auto; solve_all. @@ -641,7 +577,8 @@ Proof. cbn -[GlobalContextMap.lookup_inductive_kind]. destruct GlobalContextMap.lookup_inductive_kind as [[]|] eqn:hl' => //. 1,3,4:eapply value_constructor; tea; [erewrite <-lookup_constructor_trans; tea|now len|solve_all]. - now do 2 constructor. + apply trust_cofix. + (* now do 2 constructor. *) - intros f args vh harglen hargs ihargs. rewrite wellformed_mkApps // => /andP[] wff wfargs. rewrite trans_mkApps. @@ -669,6 +606,7 @@ Ltac destruct_times := From MetaCoq.Erasure Require Import EWcbvEvalCstrsAsBlocksInd. Lemma trans_correct {efl : EEnvFlags} {fl} {wcon : with_constructor_as_block = true} {wcb : cstr_as_blocks = true} + {wpc : with_prop_case = false} (* Avoid tLazy tBox values *) {Σ : GlobalContextMap.t} t v : has_tApp -> wf_glob Σ -> @@ -712,7 +650,8 @@ Proof. 1,3,4: eauto. + now rewrite -is_propositional_cstr_trans. + rewrite nth_error_map H2 //. - + eapply eval_beta. eapply e0; eauto. + + eapply trust_cofix. + (* eapply eval_beta. eapply e0; eauto. constructor; eauto. rewrite closed_subst // simpl_subst_k //. eapply EWcbvEval.eval_to_value in H. @@ -724,27 +663,19 @@ Proof. instantiate (1 := map (trans Σ) args). eapply All2_All2_Set. eapply values_final. solve_all. - unshelve eapply value_trans; tea. + unshelve eapply value_trans; tea.*) + now len. + now len. - + exact e. + + apply trust_cofix. - - subst brs. - cbn -[lookup_constructor lookup_constructor_pars_args - GlobalContextMap.lookup_inductive_kind] in e0 |- *. - rewrite GlobalContextMap.lookup_inductive_kind_spec. - rewrite trans_substl ?map_repeat /= in e. - { now apply forallb_repeat. } - destruct lookup_inductive_kind as [[]|] eqn:hl => //. - 1,3,4:eapply eval_iota_sing; [eauto|eauto| - now rewrite -is_propositional_trans|reflexivity| - rewrite /= ?trans_substl //; simpl; eauto]. - eapply eval_iota_sing; eauto. - eapply eval_box; eauto. - rewrite -is_propositional_trans //. - reflexivity. + - now rewrite H in wpc. - apply trust_cofix. + (*rewrite trans_mkApps /= in e1. + cbn; eapply eval_fix => //; tea. + len. apply trust_cofix*) + + - apply trust_cofix. - apply trust_cofix. - apply trust_cofix. @@ -910,13 +841,22 @@ Proof. destruct lookup_inductive_kind as [[]|] => /= //. 2-3:constructor; eauto; solve_all. constructor; eauto; solve_all. cbn. - unfold Thunk.force. - eapply isEtaExp_expanded. - all:apply trust_cofix. - - apply trust_cofix. - - apply trust_cofix. - - apply trust_cofix. - - apply trust_cofix. + now constructor. + constructor; eauto; solve_all. + - cbn -[GlobalContextMap.lookup_inductive_kind]. + rewrite GlobalContextMap.lookup_inductive_kind_spec. + destruct lookup_inductive_kind as [[]|] => /= //. + 2-3:constructor; eauto; solve_all. + constructor; eauto; solve_all. cbn. + now constructor. + constructor; eauto; solve_all. + - cbn. econstructor; eauto. solve_all. now eapply isLambda_trans. + - cbn. econstructor; eauto; solve_all. apply trust_cofix. + - cbn -[GlobalContextMap.lookup_inductive_kind]. + rewrite GlobalContextMap.lookup_inductive_kind_spec. + destruct lookup_inductive_kind as [[]|] => /= //. + 1,3,4:eapply expanded_tConstruct_app; eauto; solve_all. + apply trust_cofix. (* Needs constructor_as_blocks = true invariant *) Qed. (*cbn. eapply isEtaExp_substl. eapply forallb_repeat => //. diff --git a/erasure/theories/EConstructorsAsBlocks.v b/erasure/theories/EConstructorsAsBlocks.v index ce67f4197..0032bef87 100644 --- a/erasure/theories/EConstructorsAsBlocks.v +++ b/erasure/theories/EConstructorsAsBlocks.v @@ -62,7 +62,9 @@ Section transform_blocks. | tVar n => EAst.tVar n | tConst n => EAst.tConst n | tConstruct ind i block_args => EAst.tConstruct ind i [] - | tPrim p => EAst.tPrim (map_primIn p (fun x H => transform_blocks x)) }. + | tPrim p => EAst.tPrim (map_primIn p (fun x H => transform_blocks x)) + | tLazy t => EAst.tLazy (transform_blocks t) + | tForce t => EAst.tForce (transform_blocks t) }. Proof. all:try lia. all:try apply (In_size); tea. diff --git a/erasure/theories/EDeps.v b/erasure/theories/EDeps.v index 88a9d7716..b1180b5cc 100644 --- a/erasure/theories/EDeps.v +++ b/erasure/theories/EDeps.v @@ -86,6 +86,8 @@ Proof. now apply e. - depelim X; depelim er; constructor; cbn. solve_all. destruct p. solve_all. + - depelim er. + - depelim er. Qed. Lemma erases_deps_subst Σ Σ' s k t : @@ -137,6 +139,8 @@ Proof. constructor; [|easy]. now apply e. - depelim X; depelim er; constructor; cbn; intuition auto; solve_all. + - depelim er. + - depelim er. Qed. Lemma erases_deps_subst1 Σ Σ' t k u : @@ -195,6 +199,8 @@ Proof. constructor; [|easy]. now apply e. - depelim X; depelim er; constructor; cbn; intuition auto; solve_all. + - depelim er. + - depelim er. Qed. Lemma erases_deps_substl Σ Σ' s t : diff --git a/erasure/theories/EEtaExpanded.v b/erasure/theories/EEtaExpanded.v index a2e6dfa25..503908928 100644 --- a/erasure/theories/EEtaExpanded.v +++ b/erasure/theories/EEtaExpanded.v @@ -76,6 +76,8 @@ Section isEtaExp. | tVar _ => true | tConst _ => true | tPrim p => test_primIn p (fun x H => isEtaExp x) + | tLazy t => isEtaExp t + | tForce t => isEtaExp t | tConstruct ind i block_args => isEtaExp_app ind i 0 && is_nil block_args }. Proof. all:try lia. @@ -473,6 +475,8 @@ Inductive expanded : term -> Prop := Forall expanded args -> expanded (mkApps (tConstruct ind idx []) args) | expanded_tPrim p : primProp expanded p -> expanded (tPrim p) +| expanded_tLazy t : expanded t -> expanded (tLazy t) +| expanded_tForce t : expanded t -> expanded (tForce t) | expanded_tBox : expanded tBox. End expanded. @@ -512,19 +516,21 @@ forall (Σ : global_declarations) (P : term -> Prop), declared_constructor Σ (ind, idx) mind idecl cdecl -> #|args| >= cstr_arity mind cdecl -> Forall (expanded Σ) args -> Forall P args -> P (mkApps (tConstruct ind idx []) args)) -> (forall p, primProp (expanded Σ) p -> primProp P p -> P (tPrim p)) -> +(forall t, expanded Σ t -> P t -> P (tLazy t)) -> +(forall t, expanded Σ t -> P t -> P (tForce t)) -> (P tBox) -> forall t : term, expanded Σ t -> P t. Proof. - intros. revert t H13. + intros Σ P. intros hrel hvar hevar hlam hletin happ hconst hcase hproj hfix hcofix hcapp hprim hlazy hforce hbox. fix f 2. intros t Hexp. destruct Hexp; eauto. - - eapply H1; eauto. induction H13; econstructor; cbn in *; eauto. - - eapply H4; eauto. clear H14. induction H15; econstructor; cbn in *; eauto. - - eapply H6; eauto. induction H13; econstructor; cbn in *; eauto. - - eapply H8; eauto. induction H13; econstructor; cbn in *; intuition eauto. - - eapply H9; eauto. induction H13; econstructor; cbn in *; eauto. - - eapply H10; eauto. clear - H15 f. induction H15; econstructor; cbn in *; eauto. - - eapply H11; eauto. + - eapply hevar; eauto. induction H; econstructor; cbn in *; eauto. + - eapply happ; eauto. clear H0. induction H1; econstructor; cbn in *; eauto. + - eapply hcase; eauto. induction H; econstructor; cbn in *; eauto. + - eapply hfix; eauto. induction H; econstructor; cbn in *; intuition eauto. + - eapply hcofix; eauto. induction H; econstructor; cbn in *; eauto. + - eapply hcapp; eauto. clear - H1 f. induction H1; econstructor; cbn in *; eauto. + - eapply hprim; eauto. depelim X; constructor. destruct p; split; eauto. eapply (make_All_All f a0). Qed. diff --git a/erasure/theories/EEtaExpandedFix.v b/erasure/theories/EEtaExpandedFix.v index 9d5eeafb6..c70e37487 100644 --- a/erasure/theories/EEtaExpandedFix.v +++ b/erasure/theories/EEtaExpandedFix.v @@ -61,6 +61,8 @@ Inductive expanded (Γ : list nat): term -> Prop := Forall (expanded Γ) args -> expanded Γ (mkApps (tConstruct ind idx []) args) | expanded_tPrim p : primProp (expanded Γ) p -> expanded Γ (tPrim p) +| expanded_tLazy t : expanded Γ t -> expanded Γ (tLazy t) +| expanded_tForce t : expanded Γ t -> expanded Γ (tForce t) | expanded_tBox : expanded Γ tBox. End expanded. @@ -139,10 +141,12 @@ Lemma expanded_ind : → Forall (P Γ) args → P Γ (mkApps (tConstruct ind idx []) args)) → (∀ Γ p, primProp (expanded Σ Γ) p -> primProp (P Γ) p -> P Γ (tPrim p)) + → (∀ Γ t, expanded Σ Γ t -> P Γ t -> P Γ (tLazy t)) + → (∀ Γ t, expanded Σ Γ t -> P Γ t -> P Γ (tForce t)) → (∀ Γ : list nat, P Γ tBox) → ∀ (Γ : list nat) (t : term), expanded Σ Γ t → P Γ t. Proof. - intros Σ P HRel_app HVar HEvar HLamdba HLetIn HmkApps HConst HCase HProj HFix HCoFix HConstruct HPrim HBox. + intros Σ P HRel_app HVar HEvar HLamdba HLetIn HmkApps HConst HCase HProj HFix HCoFix HConstruct HPrim HLazy HForce HBox. fix f 3. intros Γ t Hexp. destruct Hexp; eauto. - eapply HRel_app; eauto. clear - f H0. induction H0; econstructor; eauto. @@ -297,6 +301,8 @@ Section isEtaExp. | tVar _ => true | tConst _ => true | tPrim p => test_primIn p (fun x H => isEtaExp Γ x) + | tLazy t => isEtaExp Γ t + | tForce t => isEtaExp Γ t | tConstruct ind i block_args => isEtaExp_app ind i 0 && is_nil block_args }. Proof using Σ. all:try lia. diff --git a/erasure/theories/EImplementBox.v b/erasure/theories/EImplementBox.v index f6db6a82c..99deb0a5f 100644 --- a/erasure/theories/EImplementBox.v +++ b/erasure/theories/EImplementBox.v @@ -54,7 +54,9 @@ Section implement_box. | 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 (map_primIn p (fun x H => implement_box x)). + | tPrim p => EAst.tPrim (map_primIn p (fun x H => implement_box x)) + | tLazy t => EAst.tLazy (implement_box t) + | tForce t => EAst.tForce (implement_box t). Proof. all:try lia. all:try apply (In_size); tea. @@ -445,6 +447,7 @@ Definition disable_box_term_flags (et : ETermFlags) := ; has_tFix := true ; has_tCoFix := has_tCoFix ; has_tPrim := has_tPrim + ; has_tLazy_Force := has_tLazy_Force |}. Definition switch_off_box (efl : EEnvFlags) := diff --git a/erasure/theories/EInduction.v b/erasure/theories/EInduction.v index 306b6d25c..53d4d9bc0 100644 --- a/erasure/theories/EInduction.v +++ b/erasure/theories/EInduction.v @@ -35,6 +35,8 @@ Lemma term_forall_list_ind : (forall (m : mfixpoint term) (n : nat), All (fun x => P (dbody x)) m -> P (tFix m n)) -> (forall (m : mfixpoint term) (n : nat), All (fun x => P (dbody x)) m -> P (tCoFix m n)) -> (forall p, primProp P p -> P (tPrim p)) -> + (forall t, P t -> P (tLazy t)) -> + (forall t, P t -> P (tForce t)) -> forall t : term, P t. Proof. intros until t. revert t. @@ -115,6 +117,8 @@ Fixpoint size (t : term) : nat := | tCoFix mfix idx => S (list_size (fun x => size (dbody x)) mfix) | tConstruct _ _ ignore_args => S (list_size size ignore_args) | tPrim p => S (prim_size size p) + | tLazy t => S (size t) + | tForce t => S (size t) | _ => 1 end. @@ -239,7 +243,9 @@ Section MkApps_rec. (pproj : forall (s : projection) (t : term), P t -> P (tProj s t)) (pfix : forall (m : mfixpoint term) (n : nat), All (fun x => P (dbody x)) m -> P (tFix m n)) (pcofix : forall (m : mfixpoint term) (n : nat), All (fun x => P (dbody x)) m -> P (tCoFix m n)) - (pprim : forall p, primProp P p -> P (tPrim p)). + (pprim : forall p, primProp P p -> P (tPrim p)) + (plazy : forall t, P t -> P (tLazy t)) + (pforce : forall t, P t -> P (tForce t)). Definition inspect {A} (x : A) : { y : A | x = y } := exist _ x eq_refl. @@ -265,7 +271,9 @@ Section MkApps_rec. | tProj p c => pproj p c (rec c) | tFix mfix idx => pfix mfix idx (All_rec P dbody mfix (fun x H => rec x)) | tCoFix mfix idx => pcofix mfix idx (All_rec P dbody mfix (fun x H => rec x)) - | tPrim p => pprim p _. + | tPrim p => pprim p _ + | tLazy t => plazy t (rec t) + | tForce t => pforce t (rec t). Proof. all:unfold MR; cbn; auto with arith. 4:lia. - clear -napp nonnil da rec. @@ -303,7 +311,9 @@ Section MkApps_rec. (pproj : forall (s : projection) (t : term), P (tProj s t)) (pfix : forall (m : mfixpoint term) (n : nat), P (tFix m n)) (pcofix : forall (m : mfixpoint term) (n : nat), P (tCoFix m n)) - (pprim : forall p, P (tPrim p)). + (pprim : forall p, P (tPrim p)) + (plazy : forall t, P (tLazy t)) + (pforce : forall t, P (tForce t)). Import EqNotations. @@ -325,7 +335,9 @@ Section MkApps_rec. | tProj p c => pproj p c | tFix mfix idx => pfix mfix idx | tCoFix mfix idx => pcofix mfix idx - | tPrim p => pprim p. + | tPrim p => pprim p + | tLazy t => plazy t + | tForce t => pforce t. End MkApps_case. diff --git a/erasure/theories/EInlineProjections.v b/erasure/theories/EInlineProjections.v index 1035e146a..a37b97ae7 100644 --- a/erasure/theories/EInlineProjections.v +++ b/erasure/theories/EInlineProjections.v @@ -49,6 +49,7 @@ Definition disable_projections_term_flags (et : ETermFlags) := ; has_tFix := has_tFix ; has_tCoFix := has_tCoFix ; has_tPrim := has_tPrim + ; has_tLazy_Force := has_tLazy_Force |}. Definition disable_projections_env_flag (efl : EEnvFlags) := @@ -124,6 +125,8 @@ Section optimize. | tConst _ => t | tConstruct ind n args => tConstruct ind n (map optimize args) | tPrim p => tPrim (map_prim optimize p) + | tLazy t => tLazy (optimize t) + | tForce t => tForce (optimize t) end. Lemma optimize_mkApps f l : optimize (mkApps f l) = mkApps (optimize f) (map optimize l). diff --git a/erasure/theories/ELiftSubst.v b/erasure/theories/ELiftSubst.v index a03c221eb..40632fd9a 100644 --- a/erasure/theories/ELiftSubst.v +++ b/erasure/theories/ELiftSubst.v @@ -37,6 +37,8 @@ Fixpoint lift n k t : term := | tConst _ => t | tConstruct ind i args => tConstruct ind i (map (lift n k) args) | tPrim p => tPrim (map_prim (lift n k) p) + | tLazy t => tLazy (lift n k t) + | tForce t => tForce (lift n k t) end. Notation lift0 n := (lift n 0). @@ -72,7 +74,11 @@ Fixpoint subst s k u := tCoFix mfix' idx | tConstruct ind i args => tConstruct ind i (map (subst s k) args) | tPrim p => tPrim (map_prim (subst s k) p) - | x => x + | tLazy t => tLazy (subst s k t) + | tForce t => tForce (subst s k t) + | tBox => tBox + | tVar n => tVar n + | tConst c => tConst c end. (** Substitutes [t1 ; .. ; tn] in u for [Rel 0; .. Rel (n-1)] *in parallel* *) @@ -100,6 +106,8 @@ Fixpoint closedn k (t : term) : bool := List.forallb (test_def (closedn k')) mfix | tConstruct ind i args => forallb (closedn k) args | tPrim p => test_prim (closedn k) p + | tLazy t => closedn k t + | tForce t => closedn k t | _ => true end. @@ -634,6 +642,8 @@ Proof. intros. specialize (H (#|m| + k')). now rewrite !Nat.add_assoc !(Nat.add_comm k) in H |- *. - solve_all. + - solve_all. + - solve_all. Qed. Lemma closedn_subst s k t : diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 673b9223e..211562f18 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -69,6 +69,8 @@ Section remove_match_on_box. | tConst _ => t | tConstruct ind i args => tConstruct ind i (map remove_match_on_box args) | tPrim p => tPrim (map_prim remove_match_on_box p) + | tLazy t => tLazy (remove_match_on_box t) + | tForce t => tForce (remove_match_on_box t) end. 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). diff --git a/erasure/theories/EPretty.v b/erasure/theories/EPretty.v index 47324cd2d..97e6a2081 100644 --- a/erasure/theories/EPretty.v +++ b/erasure/theories/EPretty.v @@ -170,6 +170,8 @@ Module PrintTermTree. parens top ("let cofix " ^ print_defs print_term Γ l ^ nl ^ " in " ^ List.nth_default (string_of_nat n) (map (string_of_name ∘ dname) l) n) | tPrim p => parens top (print_prim (print_term Γ false false) p) + | tLazy t => parens top ("lazy " ^ print_term Γ false false t) + | tForce t => parens top ("force " ^ print_term Γ false false t) end. End print_term. diff --git a/erasure/theories/EReflect.v b/erasure/theories/EReflect.v index 71284edc2..70bb89dc6 100644 --- a/erasure/theories/EReflect.v +++ b/erasure/theories/EReflect.v @@ -121,6 +121,8 @@ Proof. revert array_value; induction hv; intros []; eauto; nodec. destruct (p t); subst; nodec. destruct (IHhv l0); nodec. noconf e; eauto. + - destruct (IHx t); subst; nodec. now left. + - destruct (IHx t); subst; nodec. now left. Defined. #[global] diff --git a/erasure/theories/ERemoveParams.v b/erasure/theories/ERemoveParams.v index f574e2893..c58d9e5d9 100644 --- a/erasure/theories/ERemoveParams.v +++ b/erasure/theories/ERemoveParams.v @@ -59,7 +59,9 @@ Section strip. | tVar n => EAst.tVar n | tConst n => EAst.tConst n | tConstruct ind i block_args => EAst.tConstruct ind i block_args - | tPrim p => EAst.tPrim (map_primIn p (fun x H => strip x)) }. + | tPrim p => EAst.tPrim (map_primIn p (fun x H => strip x)) + | tLazy t => EAst.tLazy (strip t) + | tForce t => EAst.tForce (strip t) }. Proof. all:try lia. all:try apply (In_size); tea. @@ -569,6 +571,8 @@ Module Fast. | app, tPrim (primFloat; primFloatModel f) => mkApps (tPrim (primFloat; primFloatModel f)) app | app, tPrim (primArray; primArrayModel a) => mkApps (tPrim (primArray; primArrayModel {| array_default := strip [] a.(array_default); array_value := strip_args a.(array_value) |})) app + | app, tLazy t => mkApps (tLazy (strip [] t)) app + | app, tForce t => mkApps (tForce (strip [] t)) app | app, x => mkApps x app } where strip_args (t : list term) : list term := diff --git a/erasure/theories/ESpineView.v b/erasure/theories/ESpineView.v index 6b5a58903..6278254eb 100644 --- a/erasure/theories/ESpineView.v +++ b/erasure/theories/ESpineView.v @@ -23,7 +23,9 @@ Inductive t : term -> Set := | tProj p c : t (tProj p c) | tFix mfix idx : t (tFix mfix idx) | tCoFix mfix idx : t (tCoFix mfix idx) -| tPrim p : t (tPrim p). +| tPrim p : t (tPrim p) +| tLazy p : t (tLazy p) +| tForce p : t (tForce p). Derive Signature for t. Definition view : forall x : term, t x := @@ -39,7 +41,9 @@ Definition view : forall x : term, t x := (fun p t => tProj p t) (fun mfix n => tFix mfix n) (fun mfix n => tCoFix mfix n) - (fun p => tPrim p). + (fun p => tPrim p) + tLazy + tForce. Lemma view_mkApps {f v} (vi : t (mkApps f v)) : ~~ isApp f -> v <> [] -> exists hf vn, vi = tApp f v hf vn. diff --git a/erasure/theories/EWcbvEval.v b/erasure/theories/EWcbvEval.v index d6ca6e787..7cd775a2a 100644 --- a/erasure/theories/EWcbvEval.v +++ b/erasure/theories/EWcbvEval.v @@ -272,6 +272,12 @@ Section Wcbv. eval_primitive eval p p' -> eval (tPrim p) (tPrim p') + (* + | eval_lazy : eval (tLazy t) (tLazy t) + | eval_force t v v' : eval t (tLazy v) -> + eval v v' -> + eval (tForce t) v' *) + (** Atoms are values (includes abstractions, cofixpoints and constructors) *) | eval_atom t : atom Σ t -> eval t t. diff --git a/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v b/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v index e380b4cd0..dbaed8575 100644 --- a/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v +++ b/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v @@ -42,7 +42,9 @@ Section OnSubterm. | on_proj p c : has_tProj -> Q n c -> on_subterms Q n (tProj p c) | on_fix mfix idx : has_tFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tFix mfix idx) | on_cofix mfix idx : has_tCoFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tCoFix mfix idx) - | on_prim p : has_prim p -> primProp (Q n) p -> on_subterms Q n (tPrim p). + | on_prim p : has_prim p -> primProp (Q n) p -> on_subterms Q n (tPrim p) + | on_lazy t : has_tLazy_Force -> Q n t -> on_subterms Q n (tLazy t) + | on_force t : has_tLazy_Force -> Q n t -> on_subterms Q n (tForce t). Derive Signature for on_subterms. End OnSubterm. @@ -593,6 +595,8 @@ Proof. rtoProp; intuition auto. eapply on_cofix => //. move/andP: H0 => [] _ ha. solve_all. move/andP: H => [] hp ha. eapply on_prim => //. solve_all. + eapply on_lazy; rtoProp; intuition auto. + eapply on_force; rtoProp; intuition auto. - red. intros kn decl. move/(lookup_env_wellformed clΣ). unfold wf_global_decl. destruct cst_body => //. diff --git a/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v b/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v index 86161f502..158684f81 100644 --- a/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v +++ b/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v @@ -42,7 +42,9 @@ Section OnSubterm. | on_proj p c : has_tProj -> Q n c -> on_subterms Q n (tProj p c) | on_fix mfix idx : has_tFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tFix mfix idx) | on_cofix mfix idx : has_tCoFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tCoFix mfix idx) - | on_prim p : has_prim p -> primProp (Q n) p -> on_subterms Q n (tPrim p). + | on_prim p : has_prim p -> primProp (Q n) p -> on_subterms Q n (tPrim p) + | on_lazy t : has_tLazy_Force -> Q n t -> on_subterms Q n (tLazy t) + | on_force t : has_tLazy_Force -> Q n t -> on_subterms Q n (tForce t). Derive Signature for on_subterms. End OnSubterm. @@ -492,6 +494,8 @@ Proof. rtoProp; intuition auto. eapply on_cofix => //. move/andP: H0 => [] _ ha. solve_all. move/andP: H => [] hasp ht. eapply on_prim => //. solve_all. + eapply on_lazy; rtoProp; intuition auto. + eapply on_force; rtoProp; intuition auto. - red. intros kn decl. move/(lookup_env_wellformed clΣ). unfold wf_global_decl. destruct cst_body => //. diff --git a/erasure/theories/EWcbvEvalEtaInd.v b/erasure/theories/EWcbvEvalEtaInd.v index c92f882a1..36ac102b6 100644 --- a/erasure/theories/EWcbvEvalEtaInd.v +++ b/erasure/theories/EWcbvEvalEtaInd.v @@ -42,7 +42,9 @@ Section OnSubterm. | on_proj p c : has_tProj -> Q n c -> on_subterms Q n (tProj p c) | on_fix mfix idx : has_tFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tFix mfix idx) | on_cofix mfix idx : has_tCoFix -> All (fun d => Q (#|mfix| + n) d.(dbody)) mfix -> on_subterms Q n (tCoFix mfix idx) - | on_prim p : has_prim p -> primProp (Q n) p -> on_subterms Q n (tPrim p). + | on_prim p : has_prim p -> primProp (Q n) p -> on_subterms Q n (tPrim p) + | on_lazy t : has_tLazy_Force -> Q n t -> on_subterms Q n (tLazy t) + | on_force t : has_tLazy_Force -> Q n t -> on_subterms Q n (tForce t). Derive Signature for on_subterms. End OnSubterm. @@ -324,6 +326,12 @@ Lemma eval_preserve_mkApps_ind : eval_primitive_ind _ (fun x y _ => P x y) _ _ ev -> P' (tPrim p) (tPrim p')) -> + (* (forall t t', eval Σ t t' -> Q 0 t -> Q 0 t' -> P t t' -> + P' (tLazy t) (tLazy t')) -> + + (forall t t', eval Σ t t' -> Q 0 t -> Q 0 t' -> P t t' -> + P' (tForce t) (tForce t')) -> *) + (∀ t : term, atom Σ t → Q 0 t -> isEtaExp Σ t -> P' t t) -> ∀ (t t0 : term), Q 0 t -> isEtaExp Σ t -> eval Σ t t0 → P' t t0. Proof. @@ -712,30 +720,6 @@ Proof. - intros ise. split => //. eapply Qatom; tea. Qed. -Definition term_flags := - {| - has_tBox := true; - has_tRel := true; - has_tVar := false; - 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 := all_primitive_flags; - |}. - -Definition env_flags := - {| has_axioms := false; - has_cstr_params := false; - term_switches := term_flags ; - cstr_as_blocks := false - |}. From MetaCoq.Erasure Require Import ELiftSubst. Lemma Qpreserves_wellformed (efl : EEnvFlags) Σ : @@ -759,6 +743,8 @@ Proof. eapply on_fix; eauto. move/andP: H0 => [] _ wf. solve_all. eapply on_cofix; eauto. move/andP: H0 => [] _ wf. solve_all. eapply on_prim; eauto. solve_all. + eapply on_lazy; eauto. + eapply on_force; eauto. - red. intros kn decl. move/(lookup_env_wellformed clΣ). unfold wf_global_decl. destruct cst_body => //. @@ -782,6 +768,32 @@ Ltac destruct_nary_times := | [ H : [× _, _, _, _ & _] |- _ ] => destruct H end. +Definition term_flags := + {| + has_tBox := true; + has_tRel := true; + has_tVar := false; + 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 := all_primitive_flags; + has_tLazy_Force := false; + |}. + +Definition env_flags := + {| has_axioms := false; + has_cstr_params := false; + term_switches := term_flags ; + cstr_as_blocks := false + |}. + Lemma eval_etaexp {fl : WcbvFlags} (efl := env_flags) {Σ a a'} : with_constructor_as_block = false -> isEtaExp_env Σ -> diff --git a/erasure/theories/EWcbvEvalNamed.v b/erasure/theories/EWcbvEvalNamed.v index 48d2a598a..ffef00c29 100644 --- a/erasure/theories/EWcbvEvalNamed.v +++ b/erasure/theories/EWcbvEvalNamed.v @@ -294,6 +294,13 @@ Inductive represents : list ident -> environment -> term -> term -> Set := | represents_tPrim Γ E p p' : onPrims (represents Γ E) p p' -> Γ ;;; E ⊩ tPrim p ~ tPrim p' +| represents_tLazy Γ E t t' : + Γ ;;; E ⊩ t ~ t' -> + Γ ;;; E ⊩ tLazy t ~ tLazy t' +| represents_tForce Γ E t t' : + Γ ;;; E ⊩ t ~ t' -> + Γ ;;; E ⊩ tForce t ~ tForce t' + with represents_value : value -> term -> Set := | 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) @@ -371,6 +378,17 @@ Program Definition represents_ind := (represents_tFix Γ E mfix mfix' idx nms Hbodies Hnodup a a0)) (f9 : forall Γ E p p' (o : onPrims (represents Γ E) p p'), onPrims_dep _ (P Γ E) _ _ o -> P Γ E (tPrim p) (tPrim p') (represents_tPrim Γ E p p' o)) + + (flazy : forall Γ E t t' + (he : Γ ;;; E ⊩ t ~ t'), + P Γ E t t' he -> + P Γ E _ _ (represents_tLazy Γ E t t' he)) + + (fforce : forall Γ E t t' + (he : Γ ;;; E ⊩ t ~ t'), + P Γ E t t' he -> + P Γ E _ _ (represents_tForce Γ E t t' he)) + (f10 : ∀ (na : ident) (E : environment) @@ -419,7 +437,11 @@ Program Definition represents_ind := f8 Γ E mfix mfix' idx nms Hbodies Hnodup a a0 _ | represents_tPrim Γ E p p' r => f9 Γ E p p' r (map_onPrims (F Γ E) r) - end + | represents_tLazy Γ E t t' e => + flazy _ _ _ _ e (F Γ E _ _ e) + | represents_tForce Γ E t t' e => + fforce _ _ _ _ e (F Γ E _ _ e) + end 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 @@ -520,6 +542,16 @@ Program Definition represents_value_ind := (f9 : forall Γ E p p' (o : onPrims (represents Γ E) p p'), onPrims_dep _ (P Γ E) _ _ o -> P Γ E (tPrim p) (tPrim p') (represents_tPrim Γ E p p' o)) + (flazy : forall Γ E t t' + (he : Γ ;;; E ⊩ t ~ t'), + P Γ E t t' he -> + P Γ E _ _ (represents_tLazy Γ E t t' he)) + + (fforce : forall Γ E t t' + (he : Γ ;;; E ⊩ t ~ t'), + P Γ E t t' he -> + P Γ E _ _ (represents_tForce Γ E t t' he)) + (f10 : ∀ (na : ident) (E : environment) @@ -571,6 +603,10 @@ Program Definition represents_value_ind := f8 Γ E mfix mfix' idx nms Hbodies Hnodup a a0 _ | represents_tPrim Γ E p p' r => f9 Γ E p p' r (map_onPrims (F Γ E) r) + | represents_tLazy Γ E t t' e => + flazy _ _ _ _ e (F Γ E _ _ e) + | represents_tForce Γ E t t' e => + fforce _ _ _ _ e (F Γ E _ _ e) end with F0 (v : value) (t : term) (r : represents_value v t) {struct r} : @@ -672,8 +708,16 @@ Definition rep_ind := (represents_tFix Γ E mfix mfix' idx nms Hbodies Hnodup a a0)) (f9 : forall Γ E p p' (o : onPrims (represents Γ E) p p'), onPrims_dep _ (P Γ E) _ _ o -> P Γ E (tPrim p) (tPrim p') (represents_tPrim Γ E p p' o)) - - (f10 : + (flazy : forall Γ E t t' + (he : Γ ;;; E ⊩ t ~ t'), + P Γ E t t' he -> + P Γ E _ _ (represents_tLazy Γ E t t' he)) + + (fforce : forall Γ E t t' + (he : Γ ;;; E ⊩ t ~ t'), + P Γ E t t' he -> + P Γ E _ _ (represents_tForce Γ E t t' he)) + (f10 : ∀ (na : ident) (E : environment) (s t : term) @@ -703,8 +747,8 @@ Definition rep_ind := (f13 : forall p p' (o : onPrims represents_value p p'), onPrims_dep _ P0 _ _ o -> P0 (vPrim p) (tPrim p') (represents_value_tPrim p p' o)) , - (represents_ind P P0 f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13, - represents_value_ind P P0 f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13)). + (represents_ind P P0 f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 flazy fforce f10 f11 f12 f13, + represents_value_ind P P0 f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 flazy fforce f10 f11 f12 f13)). Local Notation "'⊩' v ~ s" := (represents_value v s) (at level 50). Local Hint Constructors represents : core. @@ -893,6 +937,8 @@ Fixpoint annotate (s : list ident) (u : term) {struct u} : term := 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 | tPrim p => tPrim (map_prim (annotate s) p) + | tLazy t => tLazy (annotate s t) + | tForce t => tForce (annotate s t) | _ => u end. @@ -918,6 +964,7 @@ Definition extraction_term_flags := ; has_tFix := true ; has_tCoFix := false ; has_tPrim := all_primitive_flags + ; has_tLazy_Force := true |}. Definition extraction_env_flags := @@ -963,6 +1010,7 @@ Definition named_extraction_term_flags := ; has_tFix := true ; has_tCoFix := false ; has_tPrim := all_primitive_flags + ; has_tLazy_Force := true |}. Definition named_extraction_env_flags := @@ -1101,7 +1149,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. @@ -1376,6 +1424,8 @@ Fixpoint sunny Γ (t : term) : bool := forallb (test_def (sunny (names ++ Γ))) mfix | tConstruct _ _ args => forallb (sunny Γ) args | tPrim p => test_prim (sunny Γ) p + | tLazy t => sunny Γ t + | tForce t => sunny Γ t | _ => true end. diff --git a/erasure/theories/EWellformed.v b/erasure/theories/EWellformed.v index d50b7bbf5..25037ce0e 100644 --- a/erasure/theories/EWellformed.v +++ b/erasure/theories/EWellformed.v @@ -47,6 +47,7 @@ Class ETermFlags := ; has_tFix : bool ; has_tCoFix : bool ; has_tPrim :: EPrimitiveFlags + ; has_tLazy_Force : bool }. Class EEnvFlags := { @@ -80,6 +81,7 @@ Definition all_term_flags := ; has_tFix := true ; has_tCoFix := true ; has_tPrim := all_primitive_flags + ; has_tLazy_Force := true |}. Definition all_env_flags := @@ -139,6 +141,8 @@ Section wf. && forallb (wellformed k) block_args else is_nil block_args | tVar _ => has_tVar | tPrim p => has_prim p && test_prim (wellformed k) p + | tLazy t => has_tLazy_Force && wellformed k t + | tForce t => has_tLazy_Force && wellformed k t end. End wf. diff --git a/erasure/theories/ErasureFunctionProperties.v b/erasure/theories/ErasureFunctionProperties.v index 8b6d3bf82..7a6e35739 100644 --- a/erasure/theories/ErasureFunctionProperties.v +++ b/erasure/theories/ErasureFunctionProperties.v @@ -1433,6 +1433,8 @@ Section wffix. | tVar _ => true | tBox => true | tPrim p => test_prim wf_fixpoints p + | tLazy t => wf_fixpoints t + | tForce t => wf_fixpoints t end. End wffix. diff --git a/erasure/theories/Typed/Optimize.v b/erasure/theories/Typed/Optimize.v index a659f1272..0309d62c9 100644 --- a/erasure/theories/Typed/Optimize.v +++ b/erasure/theories/Typed/Optimize.v @@ -21,6 +21,8 @@ Definition map_subterms (f : term -> term) (t : term) : term := | tFix def i => tFix (map (map_def f) def) i | tCoFix def i => tCoFix (map (map_def f) def) i | tPrim p => tPrim (map_prim f p) + | tLazy t => tLazy (f t) + | tForce t => tForce (f t) | t => t end. @@ -313,6 +315,8 @@ Fixpoint is_dead (rel : nat) (t : term) : bool := | tCoFix defs _ => forallb (is_dead (#|defs| + rel) ∘ EAst.dbody) defs | tConstruct _ _ args => forallb (is_dead rel) args | tPrim p => test_prim (is_dead rel) p + | tLazy t => is_dead rel t + | tForce t => is_dead rel t | _ => true end. @@ -371,6 +375,8 @@ Fixpoint valid_cases (t : term) : bool := | tCoFix defs _ => forallb (valid_cases ∘ EAst.dbody) defs | tConstruct _ _ (_ :: _) => false (* check whether constructors are not blocks*) | tPrim p => test_prim valid_cases p + | tLazy t => valid_cases t + | tForce t => valid_cases t | _ => true end. @@ -429,6 +435,8 @@ Fixpoint is_expanded_aux (nargs : nat) (t : term) : bool := | tFix defs _ | tCoFix defs _ => forallb (is_expanded_aux 0 ∘ EAst.dbody) defs | tPrim p => test_prim (is_expanded_aux 0) p + | tLazy t => is_expanded_aux 0 t + | tForce t => is_expanded_aux 0 t end. (** Check if all applications are applied enough to be deboxed without eta expansion *) @@ -666,6 +674,8 @@ Fixpoint analyze (state : analyze_state) (t : term) {struct t} : analyze_state : let state := fold_left (fun state d => analyze state (dbody d)) defs state in remove_vars state #|defs| | tPrim p => fold_prim analyze p state + | tLazy t => analyze state t + | tForce t => analyze state t end. Fixpoint decompose_TArr (bt : box_type) : list box_type × box_type := diff --git a/erasure/theories/Typed/OptimizeCorrectness.v b/erasure/theories/Typed/OptimizeCorrectness.v index f771c1f06..7a375830e 100644 --- a/erasure/theories/Typed/OptimizeCorrectness.v +++ b/erasure/theories/Typed/OptimizeCorrectness.v @@ -287,6 +287,8 @@ Proof. rewrite <- Nat.add_succ_r in *. now eapply IHX. - solve_all. + - solve_all. + - solve_all. Qed. Lemma is_dead_csubst k t u k' : @@ -522,6 +524,8 @@ Proof. + rewrite <- !Nat.add_succ_r in *. now apply IHX. - f_equal; solve_all. + - f_equal; solve_all. + - f_equal; solve_all. Qed. Lemma masked_nil {X} mask : @@ -871,6 +875,8 @@ Proof. f_equal. now rewrite p. - rewrite lift_mkApps. f_equal. simpl lift. f_equal. solve_all. + - rewrite lift_mkApps. f_equal. simpl lift. f_equal. solve_all. + - rewrite lift_mkApps. f_equal. simpl lift. f_equal. solve_all. Qed. Lemma lift_dearg n k t : @@ -922,6 +928,8 @@ Proof. rewrite <- !Nat.add_succ_r. now apply IHX. - solve_all. + - solve_all. + - solve_all. Qed. Lemma is_dead_lift_all k k' n t : @@ -1486,6 +1494,8 @@ Proof. split; [easy|]. now apply IHX. - solve_all. rtoProp; intuition solve_all. + - solve_all. rtoProp; intuition solve_all. + - solve_all. rtoProp; intuition solve_all. Qed. Lemma valid_dearg_mask_dearg mask t : @@ -1645,6 +1655,10 @@ Proof. - rewrite subst_mkApps, map_map; cbn; f_equal. f_equal. solve_all. eapply map_prim_eq_prop; tea; cbn; intuition eauto. specialize (a s k []). eauto. + - rewrite subst_mkApps, map_map; cbn; f_equal. + f_equal. specialize (IHt s k []); cbn in IHt. eauto. + - rewrite subst_mkApps, map_map; cbn; f_equal. + f_equal. specialize (IHt s k []); cbn in IHt. eauto. Qed. Lemma dearg_subst s k t : @@ -1761,6 +1775,8 @@ Proof. rewrite <- !Nat.add_succ_r. now rewrite p, IHX. - solve_all_k 6. + - solve_all. + - solve_all. Qed. Lemma is_expanded_aux_subst s n t k : @@ -1805,6 +1821,8 @@ Proof. rewrite <- !Nat.add_succ_r. now rewrite p, IHX. - solve_all_k 6. + - solve_all. + - solve_all. Qed. Lemma is_expanded_substl s n t : @@ -2416,6 +2434,8 @@ Proof. rewrite <- !Nat.add_succ_r in *. now apply IHX. - rewrite closedn_mkApps; cbn; rtoProp; intuition solve_all. solve_all_k 6. + - rewrite closedn_mkApps; cbn; rtoProp; intuition solve_all. + - rewrite closedn_mkApps; cbn; rtoProp; intuition solve_all. Qed. Lemma Alli_map {A B P n} {f : A -> B} l : diff --git a/erasure/theories/Typed/TypeAnnotations.v b/erasure/theories/Typed/TypeAnnotations.v index 55c322504..d56451529 100644 --- a/erasure/theories/Typed/TypeAnnotations.v +++ b/erasure/theories/Typed/TypeAnnotations.v @@ -586,6 +586,8 @@ Proof. intros. exact (f _ All_nil _ X). - refine (annot_mkApps _ argsa). cbn. cbn in ta. exact ta. + - refine (annot_mkApps _ argsa). cbn. cbn in ta. exact ta. + - refine (annot_mkApps _ argsa). cbn. cbn in ta. exact ta. Defined. Definition annot_dearg im cm {t : term} (ta : annots box_type t) : annots box_type (dearg im cm t) := diff --git a/template-coq/Makefile.plugin.local-late b/template-coq/Makefile.plugin.local-late new file mode 100644 index 000000000..eb053443d --- /dev/null +++ b/template-coq/Makefile.plugin.local-late @@ -0,0 +1 @@ +FINDLIBFILESTOINSTALL += gen-src/metacoq_template_plugin.a \ No newline at end of file From e896e06b7911c56020b4db3f9455c0f16169e735 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 21 Feb 2024 11:32:24 +0100 Subject: [PATCH 11/28] Implement reordering of constructors (for Extract Inductive) --- erasure-plugin/_PluginProject.in | 2 + erasure-plugin/src/g_metacoq_erasure.mlg | 2 +- .../src/metacoq_erasure_plugin.mlpack | 1 + erasure-plugin/theories/ETransform.v | 42 +++++++- erasure-plugin/theories/Erasure.v | 59 ++++++------ erasure/_CoqProject.in | 1 + erasure/theories/EReorderCstrs.v | 96 +++++++++++++++++++ 7 files changed, 171 insertions(+), 32 deletions(-) create mode 100644 erasure/theories/EReorderCstrs.v diff --git a/erasure-plugin/_PluginProject.in b/erasure-plugin/_PluginProject.in index c4015c525..d4a2f6132 100644 --- a/erasure-plugin/_PluginProject.in +++ b/erasure-plugin/_PluginProject.in @@ -120,6 +120,8 @@ src/eConstructorsAsBlocks.mli src/eConstructorsAsBlocks.ml src/eCoInductiveToInductive.mli src/eCoInductiveToInductive.ml +src/eReorderCstrs.mli +src/eReorderCstrs.ml src/eTransform.mli src/eTransform.ml src/erasure0.mli diff --git a/erasure-plugin/src/g_metacoq_erasure.mlg b/erasure-plugin/src/g_metacoq_erasure.mlg index 5e61fde14..ba9dce88e 100644 --- a/erasure-plugin/src/g_metacoq_erasure.mlg +++ b/erasure-plugin/src/g_metacoq_erasure.mlg @@ -63,7 +63,7 @@ let make_erasure_config config = enable_typed_erasure = config.typed; enable_fast_remove_params = config.fast; dearging_config = default_dearging_config; - } + inductives_mapping = [] } let time_opt config str fn arg = if config.time then diff --git a/erasure-plugin/src/metacoq_erasure_plugin.mlpack b/erasure-plugin/src/metacoq_erasure_plugin.mlpack index 485dad598..5471ffbee 100644 --- a/erasure-plugin/src/metacoq_erasure_plugin.mlpack +++ b/erasure-plugin/src/metacoq_erasure_plugin.mlpack @@ -65,6 +65,7 @@ EOptimizePropDiscr EInlineProjections EConstructorsAsBlocks ECoInductiveToInductive +EReorderCstrs EProgram OptimizePropDiscr diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index 1dfc50628..71d8f2118 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -8,7 +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 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. @@ -1023,6 +1024,45 @@ Proof. eapply ECoInductiveToInductive.trust_cofix. Qed. +From MetaCoq.Erasure Require Import EReorderCstrs. + +Axiom trust_reorder_cstrs_wf : + forall efl : EEnvFlags, + WcbvFlags -> + forall (m : inductives_mapping) (input : Transform.program E.global_context term), + wf_eprogram efl input -> wf_eprogram efl (reorder_program m input). +Axiom trust_reorder_cstrs_pres : + forall (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping) (p : Transform.program E.global_context term) + (v : term), + wf_eprogram efl p -> + eval_eprogram wfl p v -> exists v' : term, eval_eprogram wfl (reorder_program m p) v' /\ v' = reorder m v. + +Program Definition reorder_cstrs_transformation (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping) : + Transform.t _ _ EAst.term EAst.term _ _ + (eval_eprogram wfl) (eval_eprogram wfl) := + {| name := "reoder inductive constructors "; + transform p _ := EReorderCstrs.reorder_program m p ; + pre p := wf_eprogram efl p ; + post p := wf_eprogram efl p ; + obseq p hp p' v v' := v' = EReorderCstrs.reorder m v |}. + +Next Obligation. + move=> efl wfl m. cbn. now apply trust_reorder_cstrs_wf. +Qed. +Next Obligation. + red. eapply trust_reorder_cstrs_pres. +Qed. + +#[global] +Axiom trust_reorder_cstrs_transformation_ext : forall (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping), + TransformExt.t (reorder_cstrs_transformation efl wfl m) + (fun p p' => extends p.1 p'.1) (fun p p' => extends p.1 p'.1). + +#[global] +Axiom trust_reorder_cstrs_transformation_ext' : forall (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping), + TransformExt.t (reorder_cstrs_transformation efl wfl m) + extends_eprogram extends_eprogram. + Program Definition optional_transform {env env' term term' value value' eval eval'} (activate : bool) (tr : Transform.t env env' term term' value value' eval eval') : if activate then Transform.t env env' term term' value value' eval eval' diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 9f8688745..1459e58d4 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -34,7 +34,8 @@ Record erasure_configuration := { enable_cofix_to_fix : bool; enable_typed_erasure : bool; enable_fast_remove_params : bool; - dearging_config : dearging_config + dearging_config : dearging_config; + inductives_mapping : EReorderCstrs.inductives_mapping }. Definition default_dearging_config := @@ -47,14 +48,16 @@ Definition default_erasure_config := {| enable_cofix_to_fix := true; dearging_config := default_dearging_config; enable_typed_erasure := true; - enable_fast_remove_params := true |}. + enable_fast_remove_params := true; + inductives_mapping := [] |}. (* This runs only the verified phases without the typed erasure and "fast" remove params *) Definition safe_erasure_config := {| enable_cofix_to_fix := false; enable_typed_erasure := false; enable_fast_remove_params := false; - dearging_config := default_dearging_config |}. + dearging_config := default_dearging_config; + inductives_mapping := [] |}. Axiom assume_welltyped_template_program_expansion : forall p (wtp : ∥ wt_template_program_env p ∥), @@ -89,7 +92,7 @@ Definition final_wcbv_flags := {| with_guarded_fix := false; with_constructor_as_block := true |}. -Program Definition optional_cofix_to_fix_transform econf := +Program Definition optional_unsafe_transforms econf := ETransform.optional_self_transform econf.(enable_cofix_to_fix) ((* Rebuild the efficient lookup table *) rebuild_wf_env_transform (efl := EConstructorsAsBlocks.switch_cstr_as_blocks @@ -98,8 +101,8 @@ Program Definition optional_cofix_to_fix_transform econf := let efl := EConstructorsAsBlocks.switch_cstr_as_blocks (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags)) in coinductive_to_inductive_transformation efl - (has_app := eq_refl) (has_box := eq_refl) (has_rel := eq_refl) (has_pars := eq_refl) (has_cstrblocks := eq_refl)) - . + (has_app := eq_refl) (has_box := eq_refl) (has_rel := eq_refl) (has_pars := eq_refl) (has_cstrblocks := eq_refl) ▷ + reorder_cstrs_transformation efl final_wcbv_flags econf.(inductives_mapping)). Program Definition verified_lambdabox_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) @@ -237,16 +240,7 @@ Program Definition verified_lambdabox_typed_pipeline {guard : abstract_guard_imp 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) (has_box := eq_refl) (has_cstrblocks := eq_refl) ▷ - ETransform.optional_self_transform econf.(enable_cofix_to_fix) - ((* Rebuild the efficient lookup table *) - rebuild_wf_env_transform (efl := EConstructorsAsBlocks.switch_cstr_as_blocks - (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags))) false false ▷ - (* Coinductives & cofixpoints are translated to inductive types and thunked fixpoints *) - let efl := EConstructorsAsBlocks.switch_cstr_as_blocks - (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags)) in - coinductive_to_inductive_transformation efl - (has_app := eq_refl) (has_box := eq_refl) (has_rel := eq_refl) (has_pars := eq_refl) (has_cstrblocks := eq_refl)) -. + optional_unsafe_transforms econf. (* 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 @@ -254,15 +248,18 @@ Program Definition verified_lambdabox_typed_pipeline {guard : abstract_guard_imp 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. +Import EGlobalEnv EWellformed. + +Next Obligation. + destruct H. split => //. sq. + now eapply ETransform.expanded_eprogram_env_expanded_eprogram_cstrs. +Qed. - Next Obligation. - destruct H. split => //. sq. - now eapply ETransform.expanded_eprogram_env_expanded_eprogram_cstrs. - Qed. - Next Obligation. - destruct H. destruct enable_cofix_to_fix => //. - Qed. +Next Obligation. + unfold optional_unsafe_transforms. cbn. + unfold optional_self_transform. cbn. + destruct enable_cofix_to_fix => //. +Qed. Local Obligation Tactic := intros; eauto. @@ -340,7 +337,7 @@ Program Definition erasure_pipeline_fast {guard : abstract_guard_impl} (efl := E let efl := EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags) in rebuild_wf_env_transform (efl := efl) true false ▷ constructors_as_blocks_transformation (efl := efl) (has_app := eq_refl) (has_pars := eq_refl) (has_rel := eq_refl) (has_box := eq_refl) (has_cstrblocks := eq_refl) ▷ - optional_cofix_to_fix_transform econf. + optional_unsafe_transforms econf. Next Obligation. destruct H; split => //. now eapply ETransform.expanded_eprogram_env_expanded_eprogram_cstrs. Qed. @@ -360,7 +357,7 @@ Next Obligation. cbn in H. split; cbn; intuition eauto. Qed. Next Obligation. - cbn in H. unfold optional_cofix_to_fix_transform. destruct enable_cofix_to_fix => //. + cbn in H. unfold optional_unsafe_transforms. destruct enable_cofix_to_fix => //. Qed. Next Obligation. cbn in H. split; cbn; intuition eauto. @@ -401,10 +398,10 @@ Program Definition run_erase_program {guard : abstract_guard_impl} econf := if econf.(enable_typed_erasure) then run (typed_erasure_pipeline econf) else if econf.(enable_fast_remove_params) then run (erasure_pipeline_fast econf) - else run (erasure_pipeline ▷ (optional_cofix_to_fix_transform econf)). + else run (erasure_pipeline ▷ (optional_unsafe_transforms econf)). Next Obligation. Proof. - unfold optional_cofix_to_fix_transform. + unfold optional_unsafe_transforms. destruct enable_cofix_to_fix => //. Qed. @@ -423,7 +420,8 @@ Definition erasure_fast_config := {| enable_cofix_to_fix := false; dearging_config := default_dearging_config; enable_typed_erasure := false; - enable_fast_remove_params := true |}. + enable_fast_remove_params := true; + inductives_mapping := [] |}. Program Definition erase_fast_and_print_template_program (p : Ast.Env.program) : string := erase_and_print_template_program erasure_fast_config p. @@ -432,7 +430,8 @@ Definition typed_erasure_config := {| enable_cofix_to_fix := false; dearging_config := default_dearging_config; enable_typed_erasure := true; - enable_fast_remove_params := true |}. + enable_fast_remove_params := true; + inductives_mapping := [] |}. (* Parameterized by a configuration for dearging, allowing to, e.g., override masks. *) Program Definition typed_erase_and_print_template_program (p : Ast.Env.program) diff --git a/erasure/_CoqProject.in b/erasure/_CoqProject.in index e0be67b79..6fd7c9cfc 100644 --- a/erasure/_CoqProject.in +++ b/erasure/_CoqProject.in @@ -39,6 +39,7 @@ theories/EConstructorsAsBlocks.v theories/EWcbvEvalCstrsAsBlocksInd.v theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v theories/ECoInductiveToInductive.v +theories/EReorderCstrs.v theories/EImplementBox.v theories/Typed/Annotations.v diff --git a/erasure/theories/EReorderCstrs.v b/erasure/theories/EReorderCstrs.v new file mode 100644 index 000000000..16c2ded40 --- /dev/null +++ b/erasure/theories/EReorderCstrs.v @@ -0,0 +1,96 @@ +From Coq Require Import List String Arith Lia. +Import ListNotations. +From Equations Require Import Equations. + +From MetaCoq.PCUIC Require Import PCUICAstUtils. +From MetaCoq.Utils Require Import MCList bytestring utils monad_utils. +From MetaCoq.Erasure Require Import EPrimitive EAst ESpineView EEtaExpanded EInduction ERemoveParams Erasure EGlobalEnv. + +Import Kernames. +Import MCMonadNotation. + +Definition inductive_mapping : Set := Kernames.inductive * (bytestring.string * list nat). +Definition inductives_mapping := list inductive_mapping. + +Fixpoint lookup_inductive {A} (Σ : list (inductive × A)) (kn : inductive) {struct Σ} : option A := + match Σ with + | [] => None + | d :: tl => if kn == d.1 then Some d.2 else lookup_inductive tl kn + end. + +Section Reorder. + Context (Σ : global_declarations). + Context (mapping : inductives_mapping). + + Definition lookup_constructor_mapping i m := + '(tyna, tags) <- lookup_inductive mapping i ;; + List.nth_error tags m. + + Definition lookup_constructor_ordinal i m := + match lookup_constructor_mapping i m with + | None => m + | Some m' => m' + end. + + Definition reorder_list_opt {A} tags (brs : list A) : option (list A) := + mapopt (nth_error brs) tags. + + Definition reorder_list {A} tags (l : list A) := + option_get l (reorder_list_opt tags l). + + Definition reorder_branches (i : inductive) (brs : list (list BasicAst.name × term)) : list (list BasicAst.name × term) := + match lookup_inductive mapping i with + | None => brs + | Some (tyna, tags) => reorder_list tags brs + end. + + Equations reorder (t : term) : term := + | tVar na => tVar na + | tLambda nm bod => tLambda nm (reorder bod) + | tLetIn nm dfn bod => tLetIn nm dfn bod + | tApp fn arg => tApp (reorder fn) (reorder arg) + | tConst nm => tConst nm + | tConstruct i m args => tConstruct i (lookup_constructor_ordinal i m) (map reorder args) + | tCase i mch brs => tCase i mch (reorder_branches i.1 (map (on_snd reorder) brs)) + | tFix mfix idx => tFix (map (map_def reorder) mfix) idx + | tCoFix mfix idx => tCoFix (map (map_def reorder) mfix) idx + | tProj (Kernames.mkProjection ind i arg) bod => + tProj (Kernames.mkProjection ind i (lookup_constructor_ordinal ind arg)) (reorder bod) + | tPrim p => tPrim (map_prim reorder p) + | tLazy t => tLazy (reorder t) + | tForce t => tForce (reorder t) + | tRel n => tRel n + | tBox => tBox + | tEvar ev args => tEvar ev (map reorder args). + + Definition reorder_constant_decl cb := + {| cst_body := option_map reorder cb.(cst_body) |}. + + Definition reorder_one_ind kn i (oib : one_inductive_body) : one_inductive_body := + match lookup_inductive mapping {| inductive_mind := kn; inductive_ind := i |} with + | None => oib + | Some (tyna, tags) => + {| ind_name := oib.(ind_name); + ind_propositional := oib.(ind_propositional); + ind_kelim := oib.(ind_kelim); + ind_ctors := reorder_list tags oib.(ind_ctors); + ind_projs := reorder_list tags oib.(ind_projs) |} + end. + + Definition reorder_inductive_decl kn idecl := + {| ind_finite := idecl.(ind_finite); ind_npars := 0; + ind_bodies := mapi (reorder_one_ind kn) idecl.(ind_bodies) |}. + + Definition reorder_decl d := + match d.2 with + | ConstantDecl cb => (d.1, ConstantDecl (reorder_constant_decl cb)) + | InductiveDecl idecl => (d.1, InductiveDecl (reorder_inductive_decl d.1 idecl)) + end. + + Definition reorder_env Σ := + map (reorder_decl) Σ. + +End Reorder. + +Definition reorder_program mapping (p : program) : program := + (reorder_env mapping p.1, reorder mapping p.2). From e622b4308c8f03886a839ea0d994b8b2a598efe7 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 21 Feb 2024 17:58:38 +0100 Subject: [PATCH 12/28] Make reordering function transparent so it can compute in Coq --- erasure/theories/EReorderCstrs.v | 1 + 1 file changed, 1 insertion(+) diff --git a/erasure/theories/EReorderCstrs.v b/erasure/theories/EReorderCstrs.v index 16c2ded40..071d6b57a 100644 --- a/erasure/theories/EReorderCstrs.v +++ b/erasure/theories/EReorderCstrs.v @@ -1,6 +1,7 @@ From Coq Require Import List String Arith Lia. Import ListNotations. From Equations Require Import Equations. +Set Equations Transparent. From MetaCoq.PCUIC Require Import PCUICAstUtils. From MetaCoq.Utils Require Import MCList bytestring utils monad_utils. From 5ed7a02771637982aa2e4affbd2cc865181c23bc Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 12 Feb 2024 07:26:09 +0100 Subject: [PATCH 13/28] WIP unboxing transform --- erasure/theories/Unboxing.v | 2272 +++++++++++++++++++++++++++++++++++ 1 file changed, 2272 insertions(+) create mode 100644 erasure/theories/Unboxing.v diff --git a/erasure/theories/Unboxing.v b/erasure/theories/Unboxing.v new file mode 100644 index 000000000..f4695489a --- /dev/null +++ b/erasure/theories/Unboxing.v @@ -0,0 +1,2272 @@ +(* 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 Primitive BasicAst EnvMap. +From MetaCoq.Erasure Require Import EPrimitive 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 nil) : A := +| [] | nnil := False_rect _ (nnil eq_refl) +| hd :: _ | _nnil := hd. + +Section optimize. + Context (Σ : GlobalContextMap.t). + + Section Def. + Import TermSpineView. + + + Definition unboxable (idecl : one_inductive_body) (cdecl : constructor_body) := + (#|idecl.(ind_ctors)| == 1) && (cdecl.(cstr_nargs) == 1). + + Equations is_unboxable (kn : inductive) (c : nat) : bool := + | kn | 0 with GlobalContextMap.lookup_constructor Σ kn 0 := { + | Some (mdecl, idecl, cdecl) := unboxable idecl cdecl + | None := false } + | kn | S _ := false. + + + Notation " t 'eqn:' h " := (exist t h) (only parsing, at level 12). + Import Equations.Prop.Logic. + + Opaque is_unboxable. + + Equations get_unboxable_case_branch (ind : inductive) (brs : list (list name * term)) : option (name * term) := + get_unboxable_case_branch ind [([brna], bbody)] := Some (brna, bbody); + get_unboxable_case_branch ind _ := None. + + Lemma get_unboxable_case_branch_spec {ind : inductive} {brs : list (list name * term)} {brna bbody} : + get_unboxable_case_branch ind brs = Some (brna, bbody) -> + brs = [([brna], bbody)]. + Proof. + funelim (get_unboxable_case_branch ind brs) => //. + now intros [= <- <-]. + Qed. + + Equations? optimize (t : term) : term + by wf t (fun x y : EAst.term => size x < size y) := + | e with TermSpineView.view e := { + | tRel i => EAst.tRel i + | tEvar ev args => EAst.tEvar ev (map_InP args (fun x H => optimize x)) + | tLambda na M => EAst.tLambda na (optimize M) + | tApp u v napp nnil with construct_viewc u := { + | view_construct kn c block_args with is_unboxable kn c := + { | true => optimize (safe_hd v nnil) + | false => mkApps (EAst.tConstruct kn c block_args) (map_InP v (fun x H => optimize x)) } + | view_other u nconstr => + mkApps (optimize u) (map_InP v (fun x H => optimize x)) + } + | tLetIn na b b' => EAst.tLetIn na (optimize b) (optimize b') + | tCase ind c brs with inspect (is_unboxable ind.1 0) => + { | true eqn:unb with inspect (get_unboxable_case_branch ind.1 brs ) := { + | Some (brna, bbody) eqn:heqbr => EAst.tLetIn brna (optimize c) (optimize bbody) + | None eqn:heqbr := + let brs' := map_InP brs (fun x H => (x.1, optimize x.2)) in + EAst.tCase ind (optimize c) brs' } + | false eqn:unb := + let brs' := map_InP brs (fun x H => (x.1, optimize x.2)) in + EAst.tCase ind (optimize c) brs' } + | tProj p c with inspect (is_unboxable p.(proj_ind) 0) := { + | true eqn:unb => optimize c + | false eqn:unb => EAst.tProj p (optimize c) } + | tConstruct ind n args => EAst.tConstruct ind n args + | tFix mfix idx => + let mfix' := map_InP mfix (fun d H => {| dname := dname d; dbody := optimize 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 := optimize d.(dbody); rarg := d.(rarg) |}) in + EAst.tCoFix mfix' idx + | tBox => EAst.tBox + | tVar n => EAst.tVar n + | tConst n => EAst.tConst n + | tPrim p => EAst.tPrim (map_primIn p (fun x H => optimize x)) }. + Proof. + all:try lia. + all:try apply (In_size); tea. + - eapply (In_size id size) in H. unfold id in H. + change (fun x => size x) with size in *. lia. + - rewrite size_mkApps. destruct v; cbn => //. lia. + - rewrite size_mkApps. + eapply (In_size id size) in H. change (fun x => size (id x)) with size in H. unfold id in H. cbn. lia. + - rewrite size_mkApps. destruct v => //=. lia. + - rewrite size_mkApps. + eapply (In_size id size) in H. change (fun x => size (id x)) with size in H. unfold id in H. cbn. lia. + - move/get_unboxable_case_branch_spec: heqbr => -> //=. lia. + - eapply (In_size snd size) in H; cbn in H; lia. + - eapply (In_size snd size) in H; cbn in H; lia. + - eapply (In_size dbody size) in H; cbn in H; lia. + - eapply (In_size dbody size) in H; cbn in H; lia. + - destruct p as [? []]; cbn in H; eauto. + intuition auto; subst; cbn; try lia. + eapply (In_size id size) in H0. unfold id in H0. + change (fun x => size x) with size in H0. lia. + Qed. + End Def. + + #[universes(polymorphic)] Hint Rewrite @map_primIn_spec @map_InP_spec : optimize. + + Lemma map_repeat {A B} (f : A -> B) x n : map f (repeat x n) = repeat (f x) n. + Proof using Type. + now induction n; simpl; auto; rewrite IHn. + Qed. + + Lemma map_optimize_repeat_box n : map optimize (repeat tBox n) = repeat tBox n. + Proof using Type. now rewrite map_repeat. Qed. + + Arguments eqb : simpl never. + + Opaque optimize_unfold_clause_1. + Opaque optimize. + Opaque isEtaExp. + Opaque isEtaExp_unfold_clause_1. + + Lemma closedn_mkApps k f l : closedn k (mkApps f l) = closedn k f && forallb (closedn k) l. + Proof using Type. + induction l in f |- *; cbn; auto. + - now rewrite andb_true_r. + - now rewrite IHl /= andb_assoc. + Qed. + + Lemma closed_optimize t k : closedn k t -> closedn k (optimize t). + Proof using Type Σ. + funelim (optimize t); simp optimize; rewrite -?optimize_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. + - solve_all_k 6. + - rewrite !closedn_mkApps in H1 *. + rtoProp; intuition auto. + solve_all. + - rewrite !closedn_mkApps /= in H0 *. rtoProp. + destruct v => //=. cbn in H1. rtoProp; intuition auto. + - rewrite !closedn_mkApps /= in H0 *. rtoProp. repeat solve_all. + - rtoProp; intuition auto. clear H1 H2 Heq. + move: H4; move/get_unboxable_case_branch_spec: heqbr => -> //=. + now rewrite andb_true_r. + Qed. + + Hint Rewrite @forallb_InP_spec : isEtaExp. + Transparent isEtaExp_unfold_clause_1. + + Local Lemma optimize_mkApps_nonnil f v (nnil : v <> []): + ~~ isApp f -> + optimize (mkApps f v) = match construct_viewc f with + | view_construct kn c block_args => + if is_unboxable kn c then optimize (safe_hd v nnil) + else mkApps (EAst.tConstruct kn c block_args) (map optimize v) + | view_other u nconstr => mkApps (optimize f) (map optimize v) + end. + Proof using Type. + intros napp. rewrite optimize_equation_1. + destruct (TermSpineView.view_mkApps (TermSpineView.view (mkApps f v)) napp nnil) as [hna [hv' ->]]. + simp optimize; rewrite -optimize_equation_1. + destruct (construct_viewc f). + 2:cbn; simp optimize => //. + simp optimize. + destruct (is_unboxable ind n) => //=. simp optimize. + replace (safe_hd v hv') with (safe_hd v nnil) => //. + destruct v; cbn => //. + f_equal. now simp optimize. + Qed. + + Lemma optimize_mkApps f v : ~~ isApp f -> + optimize (mkApps f v) = match construct_viewc f with + | view_construct kn c block_args => + if is_unboxable kn c then + match v with + | hd :: _ => optimize hd + | _ => mkApps (EAst.tConstruct kn c block_args) (map optimize v) + end + else mkApps (EAst.tConstruct kn c block_args) (map optimize v) + | view_other u nconstr => mkApps (optimize f) (map optimize v) + end. + Proof using Type. + intros napp. + destruct v using rev_case; simpl. + - destruct construct_viewc => //. simp optimize. + destruct is_unboxable => //. + - rewrite (optimize_mkApps_nonnil f (v ++ [x])) => //. + destruct construct_viewc => //. + destruct is_unboxable eqn:unb => //. + destruct v eqn:hc => //=. + Qed. + + Lemma lookup_inductive_pars_constructor_pars_args {ind n pars args} : + lookup_constructor_pars_args Σ ind n = Some (pars, args) -> + lookup_inductive_pars Σ (inductive_mind ind) = Some pars. + Proof using Type. + rewrite /lookup_constructor_pars_args /lookup_inductive_pars. + rewrite /lookup_constructor /lookup_inductive. destruct lookup_minductive => //. + cbn. do 2 destruct nth_error => //. congruence. + Qed. + + Lemma optimize_csubst a k b : + closed a -> + isEtaExp Σ a -> + isEtaExp Σ b -> + optimize (ECSubst.csubst a k b) = ECSubst.csubst (optimize a) k (optimize b). + Proof using Type. + intros cla etaa; move cla before a. move etaa before a. + funelim (optimize b); cbn; simp optimize isEtaExp; rewrite -?isEtaExp_equation_1 -?optimize_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]; try easy. + + - destruct Nat.compare => //. + - f_equal. rtoProp. solve_all. destruct args; inv H0. eauto. + - f_equal. solve_all. move/andP: b => [] _ he. solve_all. + - specialize (H a etaa cla k). + rewrite !csubst_mkApps in H1 *. + rewrite isEtaExp_mkApps_napp // in H1. + destruct construct_viewc. + * cbn. rewrite optimize_mkApps //. + * move/andP: H1 => [] et ev. + rewrite -H //. + assert (map (csubst a k) v <> []). + { destruct v; cbn; congruence. } + pose proof (etaExp_csubst Σ _ k _ etaa et). + destruct (isApp (csubst a k t)) eqn:eqa. + { destruct (decompose_app (csubst a k t)) eqn:eqk. + rewrite (decompose_app_inv eqk) in H2 *. + pose proof (decompose_app_notApp _ _ _ eqk). + assert (l <> []). + { intros ->. rewrite (decompose_app_inv eqk) in eqa. now rewrite eqa in H3. } + rewrite isEtaExp_mkApps_napp // in H2. + assert ((l ++ map (csubst a k) v)%list <> []). + { destruct l; cbn; congruence. } + + destruct (construct_viewc t0) eqn:hc. + { rewrite -mkApps_app /=. + rewrite optimize_mkApps // optimize_mkApps //. + cbn -[lookup_inductive_pars]. + move/andP: H2 => [] ise hl. + unfold isEtaExp_app in ise. + destruct is_unboxable eqn:isunb => //. + destruct l => //=. + rewrite (lookup_inductive_pars_constructor_pars_args eqpars). + rewrite -mkApps_app /= !skipn_map. f_equal. + rewrite skipn_app map_app. f_equal. + assert (pars - #|l| = 0). rtoProp. rename H2 into ise. eapply Nat.leb_le in ise; lia. + rewrite H2 skipn_0. + rewrite !map_map_compose. + clear -etaa cla ev H0. solve_all. } + { rewrite -mkApps_app. + rewrite optimize_mkApps //. rewrite hc. + rewrite optimize_mkApps // hc -mkApps_app map_app //. + f_equal. f_equal. + rewrite !map_map_compose. + clear -etaa cla ev H0. solve_all. } } + { rewrite optimize_mkApps ?eqa //. + destruct (construct_viewc (csubst a k t)) eqn:eqc. + 2:{ f_equal. rewrite !map_map_compose. clear -etaa cla ev H0. solve_all. } + simp isEtaExp in H2. + rewrite /isEtaExp_app in H2. + destruct lookup_constructor_pars_args as [[pars args]|] eqn:eqpars => // /=. + rewrite (lookup_inductive_pars_constructor_pars_args eqpars). + assert (pars = 0). rtoProp. eapply Nat.leb_le in H2. lia. + subst pars. rewrite skipn_0. + simp optimize; rewrite -optimize_equation_1. + { f_equal. rewrite !map_map_compose. clear -etaa cla ev H0. solve_all. } } + - pose proof (etaExp_csubst _ _ k _ etaa H0). + rewrite !csubst_mkApps /= in H1 *. + assert (map (csubst a k) v <> []). + { destruct v; cbn; congruence. } + rewrite optimize_mkApps //. + rewrite isEtaExp_Constructor // in H1. + move/andP: H1. rewrite map_length. move=> [] etaapp etav. + cbn -[lookup_inductive_pars]. + unfold isEtaExp_app in etaapp. + rewrite GlobalContextMap.lookup_inductive_pars_spec in Heq. + rewrite Heq in etaapp *. + f_equal. + now destruct block_args; inv etav. + rewrite map_skipn. f_equal. + rewrite !map_map_compose. + rewrite isEtaExp_Constructor // in H0. rtoProp. solve_all. + - pose proof (etaExp_csubst _ _ k _ etaa H0). + rewrite !csubst_mkApps /= in H1 *. + assert (map (csubst a k) v <> []). + { destruct v; cbn; congruence. } + rewrite optimize_mkApps //. + rewrite isEtaExp_Constructor // in H1. + rewrite GlobalContextMap.lookup_inductive_pars_spec in Heq. + move/andP: H1. rewrite map_length. move=> [] etaapp etav. + cbn -[lookup_inductive_pars]. + unfold isEtaExp_app in etaapp. + destruct lookup_constructor_pars_args as [[pars args]|] eqn:eqpars => //. + now rewrite (lookup_inductive_pars_constructor_pars_args eqpars) in Heq. + Qed. + + Lemma optimize_substl s t : + forallb (closedn 0) s -> + forallb (isEtaExp Σ) s -> + isEtaExp Σ t -> + optimize (substl s t) = substl (map optimize s) (optimize t). + Proof using Type. + induction s in t |- *; simpl; auto. + move=> /andP[] cla cls /andP[] etaa etas etat. + rewrite IHs //. now eapply etaExp_csubst. f_equal. + now rewrite optimize_csubst. + Qed. + + Lemma optimize_iota_red pars args br : + forallb (closedn 0) args -> + forallb (isEtaExp Σ) args -> + isEtaExp Σ br.2 -> + optimize (EGlobalEnv.iota_red pars args br) = EGlobalEnv.iota_red pars (map optimize args) (on_snd optimize br). + Proof using Type. + intros cl etaargs etabr. + unfold EGlobalEnv.iota_red. + rewrite optimize_substl //. + rewrite forallb_rev forallb_skipn //. + 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). + Proof using Type. + unfold EGlobalEnv.fix_subst. + rewrite map_length. + generalize #|mfix|. + induction n; simpl; auto. + f_equal; auto. now simp optimize. + Qed. + + Lemma optimize_cofix_subst mfix : EGlobalEnv.cofix_subst (map (map_def optimize) mfix) = map optimize (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 optimize. + Qed. + + Lemma optimize_cunfold_fix mfix idx n f : + forallb (closedn 0) (fix_subst mfix) -> + forallb (fun d => isLambda (dbody d) && isEtaExp Σ (dbody d)) mfix -> + cunfold_fix mfix idx = Some (n, f) -> + cunfold_fix (map (map_def optimize) mfix) idx = Some (n, optimize f). + Proof using Type. + intros hfix heta. + unfold cunfold_fix. + rewrite nth_error_map. + destruct nth_error eqn:heq. + intros [= <- <-] => /=. f_equal. f_equal. + rewrite optimize_substl //. + now apply isEtaExp_fix_subst. + solve_all. eapply nth_error_all in heta; tea. cbn in heta. + rtoProp; intuition auto. + f_equal. f_equal. apply optimize_fix_subst. + discriminate. + Qed. + + + Lemma optimize_cunfold_cofix mfix idx n f : + forallb (closedn 0) (cofix_subst mfix) -> + forallb (isEtaExp Σ ∘ dbody) mfix -> + cunfold_cofix mfix idx = Some (n, f) -> + cunfold_cofix (map (map_def optimize) mfix) idx = Some (n, optimize f). + Proof using Type. + intros hcofix heta. + unfold cunfold_cofix. + rewrite nth_error_map. + destruct nth_error eqn:heq. + intros [= <- <-] => /=. f_equal. + rewrite optimize_substl //. + now apply isEtaExp_cofix_subst. + solve_all. now eapply nth_error_all in heta; tea. + f_equal. f_equal. apply optimize_cofix_subst. + discriminate. + Qed. + + Lemma optimize_nth {n l d} : + optimize (nth n l d) = nth n (map optimize l) (optimize d). + Proof using Type. + induction l in n |- *; destruct n; simpl; auto. + Qed. + +End optimize. + +#[universes(polymorphic)] Global Hint Rewrite @map_primIn_spec @map_InP_spec : optimize. +Tactic Notation "simp_eta" "in" hyp(H) := simp isEtaExp in H; rewrite -?isEtaExp_equation_1 in H. +Ltac simp_eta := simp isEtaExp; rewrite -?isEtaExp_equation_1. +Tactic Notation "simp_optimize" "in" hyp(H) := simp optimize in H; rewrite -?optimize_equation_1 in H. +Ltac simp_optimize := simp optimize; rewrite -?optimize_equation_1. + +Definition optimize_constant_decl Σ cb := + {| cst_body := option_map (optimize Σ) cb.(cst_body) |}. + +Definition optimize_inductive_decl idecl := + {| ind_finite := idecl.(ind_finite); ind_npars := 0; ind_bodies := idecl.(ind_bodies) |}. + +Definition optimize_decl Σ d := + match d with + | ConstantDecl cb => ConstantDecl (optimize_constant_decl Σ cb) + | InductiveDecl idecl => InductiveDecl (optimize_inductive_decl idecl) + end. + +Definition optimize_env Σ := + map (on_snd (optimize_decl Σ)) Σ.(GlobalContextMap.global_decls). + +Definition optimize_program (p : eprogram_env) : eprogram := + (ERemoveParams.optimize_env p.1, ERemoveParams.optimize p.1 p.2). + +Import EGlobalEnv. + +Lemma lookup_env_optimize Σ kn : + lookup_env (optimize_env Σ) kn = + option_map (optimize_decl Σ) (lookup_env Σ.(GlobalContextMap.global_decls) kn). +Proof. + unfold optimize_env. + destruct Σ as [Σ map repr wf]; cbn. + induction Σ at 2 4; simpl; auto. + case: eqb_spec => //. +Qed. + +Lemma lookup_constructor_optimize {Σ kn c} : + lookup_constructor (optimize_env Σ) kn c = + match lookup_constructor Σ.(GlobalContextMap.global_decls) kn c with + | Some (mdecl, idecl, cdecl) => Some (optimize_inductive_decl mdecl, idecl, cdecl) + | None => None + end. +Proof. + rewrite /lookup_constructor /lookup_inductive /lookup_minductive. + rewrite lookup_env_optimize. + destruct lookup_env as [[]|] => /= //. + do 2 destruct nth_error => //. +Qed. + +Lemma lookup_constructor_pars_args_optimize (Σ : GlobalContextMap.t) i n npars nargs : + EGlobalEnv.lookup_constructor_pars_args Σ i n = Some (npars, nargs) -> + EGlobalEnv.lookup_constructor_pars_args (optimize_env Σ) i n = Some (0, nargs). +Proof. + rewrite /EGlobalEnv.lookup_constructor_pars_args. rewrite lookup_constructor_optimize //=. + destruct EGlobalEnv.lookup_constructor => //. destruct p as [[] ?] => //=. now intros [= <- <-]. +Qed. + +Lemma is_propositional_optimize (Σ : GlobalContextMap.t) ind : + match inductive_isprop_and_pars Σ.(GlobalContextMap.global_decls) ind with + | Some (prop, npars) => + inductive_isprop_and_pars (optimize_env Σ) ind = Some (prop, 0) + | None => + inductive_isprop_and_pars (optimize_env Σ) ind = None + end. +Proof. + rewrite /inductive_isprop_and_pars /lookup_inductive /lookup_minductive. + rewrite lookup_env_optimize. + destruct lookup_env; simpl; auto. + destruct g; simpl; auto. destruct nth_error => //. +Qed. + +Lemma is_propositional_cstr_optimize {Σ : GlobalContextMap.t} {ind c} : + match constructor_isprop_pars_decl Σ.(GlobalContextMap.global_decls) ind c with + | Some (prop, npars, cdecl) => + constructor_isprop_pars_decl (optimize_env Σ) ind c = Some (prop, 0, cdecl) + | None => + constructor_isprop_pars_decl (optimize_env Σ) ind c = None + end. +Proof. + rewrite /constructor_isprop_pars_decl /lookup_constructor /lookup_inductive /lookup_minductive. + rewrite lookup_env_optimize. + destruct lookup_env; simpl; auto. + destruct g; simpl; auto. do 2 destruct nth_error => //. +Qed. + +Lemma lookup_inductive_pars_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : + wf_glob Σ -> + forall pars, lookup_inductive_pars Σ ind = Some pars -> + EGlobalEnv.lookup_inductive_pars (optimize_env Σ) ind = Some 0. +Proof. + rewrite /lookup_inductive_pars => wf pars. + rewrite /lookup_inductive /lookup_minductive. + rewrite (lookup_env_optimize _ ind). + destruct lookup_env as [[decl|]|] => //=. +Qed. + +Arguments eval {wfl}. + +Arguments isEtaExp : simpl never. + +Lemma isEtaExp_mkApps {Σ} {f u} : isEtaExp Σ (tApp f u) -> + let (hd, args) := decompose_app (tApp f u) in + match construct_viewc hd with + | view_construct kn c block_args => + args <> [] /\ f = mkApps hd (remove_last args) /\ u = last args u /\ + isEtaExp_app Σ kn c #|args| && forallb (isEtaExp Σ) args && is_nil block_args + | view_other _ discr => + [&& isEtaExp Σ hd, forallb (isEtaExp Σ) args, isEtaExp Σ f & isEtaExp Σ u] + end. +Proof. + move/(isEtaExp_mkApps Σ f [u]). + cbn -[decompose_app]. destruct decompose_app eqn:da. + destruct construct_viewc eqn:cv => //. + intros ->. + pose proof (decompose_app_inv da). + pose proof (decompose_app_notApp _ _ _ da). + destruct l using rev_case. cbn. intuition auto. solve_discr. noconf H. + rewrite mkApps_app in H. noconf H. + rewrite remove_last_app last_last. intuition auto. + destruct l; cbn in *; congruence. + pose proof (decompose_app_inv da). + pose proof (decompose_app_notApp _ _ _ da). + destruct l using rev_case. cbn. intuition auto. destruct t => //. + rewrite mkApps_app in H. noconf H. + move=> /andP[] etat. rewrite forallb_app => /andP[] etal /=. + rewrite andb_true_r => etaa. rewrite etaa andb_true_r. + rewrite etat etal. cbn. rewrite andb_true_r. + eapply isEtaExp_mkApps_intro; auto; solve_all. +Qed. + +Lemma decompose_app_tApp_split f a hd args : + decompose_app (tApp f a) = (hd, args) -> f = mkApps hd (remove_last args) /\ a = last args a. +Proof. + unfold decompose_app. cbn. + move/decompose_app_rec_inv' => [n [napp [hskip heq]]]. + rewrite -(firstn_skipn n args). + rewrite -hskip. rewrite last_last; split => //. + rewrite heq. f_equal. + now rewrite remove_last_app. +Qed. + +Arguments lookup_inductive_pars_constructor_pars_args {Σ ind n pars args}. + +Lemma optimize_tApp {Σ : GlobalContextMap.t} f a : isEtaExp Σ f -> isEtaExp Σ a -> optimize Σ (EAst.tApp f a) = EAst.tApp (optimize Σ f) (optimize Σ a). +Proof. + move=> etaf etaa. + pose proof (isEtaExp_mkApps_intro _ _ [a] etaf). + forward H by eauto. + move/isEtaExp_mkApps: H. + destruct decompose_app eqn:da. + destruct construct_viewc eqn:cv => //. + { intros [? [? []]]. rewrite H0 /=. + rewrite -[EAst.tApp _ _ ](mkApps_app _ _ [a]). + move/andP: H2 => []. rewrite /isEtaExp_app. + rewrite !optimize_mkApps // cv. + destruct lookup_constructor_pars_args as [[pars args]|] eqn:hl => //. + rewrite (lookup_inductive_pars_constructor_pars_args hl). + intros hpars etal. + rewrite -[EAst.tApp _ _ ](mkApps_app _ _ [optimize Σ a]). + f_equal. rewrite !skipn_map !skipn_app map_app. f_equal. + assert (pars - #|remove_last l| = 0) as ->. + 2:{ rewrite skipn_0 //. } + rewrite H0 in etaf. + rewrite isEtaExp_mkApps_napp // in etaf. + simp construct_viewc in etaf. + move/andP: etaf => []. rewrite /isEtaExp_app hl. + move => /andP[] /Nat.leb_le. lia. } + { move/and4P=> [] iset isel _ _. rewrite (decompose_app_inv da). + pose proof (decompose_app_notApp _ _ _ da). + rewrite optimize_mkApps //. + destruct (decompose_app_tApp_split _ _ _ _ da). + rewrite cv. rewrite H0. + rewrite optimize_mkApps // cv. + rewrite -[EAst.tApp _ _ ](mkApps_app _ _ [optimize Σ a]). f_equal. + rewrite -[(_ ++ _)%list](map_app (optimize Σ) _ [a]). + f_equal. + assert (l <> []). + { destruct l; try congruence. eapply decompose_app_inv in da. cbn in *. now subst t. } + rewrite H1. + now apply remove_last_last. } +Qed. + +Module Fast. + Section fastoptimize. + Context (Σ : GlobalContextMap.t). + + Equations optimize (app : list term) (t : term) : term := { + | app, tEvar ev args => mkApps (EAst.tEvar ev (optimize_args args)) app + | app, tLambda na M => mkApps (EAst.tLambda na (optimize [] M)) app + | app, tApp u v => optimize (optimize [] v :: app) u + | app, tLetIn na b b' => mkApps (EAst.tLetIn na (optimize [] b) (optimize [] b')) app + | app, tCase ind c brs => + let brs' := optimize_brs brs in + mkApps (EAst.tCase (ind.1, 0) (optimize [] c) brs') app + | app, tProj p c => mkApps (EAst.tProj {| proj_ind := p.(proj_ind); proj_npars := 0; proj_arg := p.(proj_arg) |} (optimize [] c)) app + | app, tFix mfix idx => + let mfix' := optimize_defs mfix in + mkApps (EAst.tFix mfix' idx) app + | app, tCoFix mfix idx => + let mfix' := optimize_defs mfix in + mkApps (EAst.tCoFix mfix' idx) app + | app, tConstruct kn c block_args with GlobalContextMap.lookup_inductive_pars Σ (inductive_mind kn) := { + | Some npars => mkApps (EAst.tConstruct kn c block_args) (List.skipn npars app) + | None => mkApps (EAst.tConstruct kn c block_args) app } + | app, tPrim (primInt; primIntModel i) => mkApps (tPrim (primInt; primIntModel i)) app + | app, tPrim (primFloat; primFloatModel f) => mkApps (tPrim (primFloat; primFloatModel f)) app + | app, tPrim (primArray; primArrayModel a) => + mkApps (tPrim (primArray; primArrayModel {| array_default := optimize [] a.(array_default); array_value := optimize_args a.(array_value) |})) app + | app, x => mkApps x app } + + where optimize_args (t : list term) : list term := + { | [] := [] + | a :: args := (optimize [] a) :: optimize_args args + } + + where optimize_brs (t : list (list BasicAst.name × term)) : list (list BasicAst.name × term) := + { | [] := [] + | a :: args := (a.1, (optimize [] a.2)) :: optimize_brs args } + + where optimize_defs (t : mfixpoint term) : mfixpoint term := + { | [] := [] + | d :: defs := {| dname := dname d; dbody := optimize [] d.(dbody); rarg := d.(rarg) |} :: optimize_defs defs + }. + + Local Ltac specIH := + match goal with + | [ H : (forall args : list term, _) |- _ ] => specialize (H [] eq_refl) + end. + + Lemma optimize_acc_opt t : + forall args, ERemoveParams.optimize Σ (mkApps t args) = optimize (map (ERemoveParams.optimize Σ) args) t. + Proof using Type. + intros args. + remember (map (ERemoveParams.optimize Σ) args) as args'. + revert args Heqargs'. + set (P:= fun args' t fs => forall args, args' = map (ERemoveParams.optimize Σ) args -> ERemoveParams.optimize Σ (mkApps t args) = fs). + apply (optimize_elim P + (fun l l' => map (ERemoveParams.optimize Σ) l = l') + (fun l l' => map (on_snd (ERemoveParams.optimize Σ)) l = l') + (fun l l' => map (map_def (ERemoveParams.optimize Σ)) l = l')); subst P; cbn -[GlobalContextMap.lookup_inductive_pars isEtaExp ERemoveParams.optimize]; clear. + all: try reflexivity. + all: intros *; simp_eta; simp_optimize. + all: try solve [intros; subst; rtoProp; rewrite optimize_mkApps // /=; simp_optimize; repeat specIH; repeat (f_equal; intuition eauto)]. + all: try solve [rewrite optimize_mkApps //]. + - intros IHv IHu. + specialize (IHv [] eq_refl). simpl in IHv. + intros args ->. specialize (IHu (v :: args)). + forward IHu. now rewrite -IHv. exact IHu. + - intros Hl hargs args ->. + rewrite optimize_mkApps //=. simp_optimize. + f_equal. f_equal. cbn. + do 2 f_equal. rewrite /map_array_model. + specialize (Hl [] eq_refl). f_equal; eauto. + - intros Hl args ->. + rewrite optimize_mkApps // /=. + rewrite GlobalContextMap.lookup_inductive_pars_spec in Hl. now rewrite Hl. + - intros Hl args ->. + rewrite GlobalContextMap.lookup_inductive_pars_spec in Hl. + now rewrite optimize_mkApps // /= Hl. + - intros IHa heq. + specIH. now rewrite IHa. + - intros IHa heq; specIH. f_equal; eauto. unfold on_snd. now rewrite IHa. + - intros IHa heq; specIH. unfold map_def. f_equal; eauto. now rewrite IHa. + Qed. + + Lemma optimize_fast t : ERemoveParams.optimize Σ t = optimize [] t. + Proof using Type. now apply (optimize_acc_opt t []). Qed. + + End fastoptimize. + + Notation optimize' Σ := (optimize Σ []). + + Definition optimize_constant_decl Σ cb := + {| cst_body := option_map (optimize' Σ) cb.(cst_body) |}. + + Definition optimize_inductive_decl idecl := + {| ind_finite := idecl.(ind_finite); ind_npars := 0; ind_bodies := idecl.(ind_bodies) |}. + + Definition optimize_decl Σ d := + match d with + | ConstantDecl cb => ConstantDecl (optimize_constant_decl Σ cb) + | InductiveDecl idecl => InductiveDecl (optimize_inductive_decl idecl) + end. + + Definition optimize_env Σ := + map (on_snd (optimize_decl Σ)) Σ.(GlobalContextMap.global_decls). + + Lemma optimize_env_fast Σ : ERemoveParams.optimize_env Σ = optimize_env Σ. + Proof. + unfold ERemoveParams.optimize_env, optimize_env. + destruct Σ as [Σ map repr wf]. cbn. + induction Σ at 2 4; cbn; auto. + f_equal; eauto. + destruct a as [kn []]; cbn; auto. + destruct c as [[b|]]; cbn; auto. unfold on_snd; cbn. + do 2 f_equal. unfold ERemoveParams.optimize_constant_decl, optimize_constant_decl. + simpl. f_equal. f_equal. apply optimize_fast. + Qed. + +End Fast. + +Lemma isLambda_mkApps' f l : + l <> nil -> + ~~ EAst.isLambda (EAst.mkApps f l). +Proof. + induction l using rev_ind; try congruence. + rewrite mkApps_app /= //. +Qed. + +Lemma isBox_mkApps' f l : + l <> nil -> + ~~ isBox (EAst.mkApps f l). +Proof. + induction l using rev_ind; try congruence. + rewrite mkApps_app /= //. +Qed. + +Lemma isFix_mkApps' f l : + l <> nil -> + ~~ isFix (EAst.mkApps f l). +Proof. + induction l using rev_ind; try congruence. + rewrite mkApps_app /= //. +Qed. + +Lemma isLambda_mkApps_Construct ind n block_args l : + ~~ EAst.isLambda (EAst.mkApps (EAst.tConstruct ind n block_args) l). +Proof. + induction l using rev_ind; cbn; try congruence. + rewrite mkApps_app /= //. +Qed. + +Lemma isBox_mkApps_Construct ind n block_args l : + ~~ isBox (EAst.mkApps (EAst.tConstruct ind n block_args) l). +Proof. + induction l using rev_ind; cbn; try congruence. + rewrite mkApps_app /= //. +Qed. + +Lemma isFix_mkApps_Construct ind n block_args l : + ~~ isFix (EAst.mkApps (EAst.tConstruct ind n block_args) l). +Proof. + induction l using rev_ind; cbn; try congruence. + rewrite mkApps_app /= //. +Qed. + +Lemma optimize_isLambda Σ f : + EAst.isLambda f = EAst.isLambda (optimize Σ f). +Proof. + funelim (optimize Σ f); cbn -[optimize]; (try simp_optimize) => //. + rewrite (negbTE (isLambda_mkApps' _ _ _)) //. + rewrite (negbTE (isLambda_mkApps' _ _ _)) //; try apply map_nil => //. + all:rewrite !(negbTE (isLambda_mkApps_Construct _ _ _ _)) //. +Qed. + +Lemma optimize_isBox Σ f : + isBox f = isBox (optimize Σ f). +Proof. + funelim (optimize Σ f); cbn -[optimize] => //. + all:rewrite map_InP_spec. + rewrite (negbTE (isBox_mkApps' _ _ _)) //. + rewrite (negbTE (isBox_mkApps' _ _ _)) //; try apply map_nil => //. + all:rewrite !(negbTE (isBox_mkApps_Construct _ _ _ _)) //. +Qed. + +Lemma isApp_mkApps u v : v <> nil -> isApp (mkApps u v). +Proof. + destruct v using rev_case; try congruence. + rewrite mkApps_app /= //. +Qed. + +Lemma optimize_isApp Σ f : + ~~ EAst.isApp f -> + ~~ EAst.isApp (optimize Σ f). +Proof. + funelim (optimize Σ f); cbn -[optimize] => //. + all:rewrite map_InP_spec. + all:rewrite isApp_mkApps //. +Qed. + +Lemma optimize_isFix Σ f : + isFix f = isFix (optimize Σ f). +Proof. + funelim (optimize Σ f); cbn -[optimize] => //. + all:rewrite map_InP_spec. + rewrite (negbTE (isFix_mkApps' _ _ _)) //. + rewrite (negbTE (isFix_mkApps' _ _ _)) //; try apply map_nil => //. + all:rewrite !(negbTE (isFix_mkApps_Construct _ _ _ _)) //. +Qed. + +Lemma optimize_isFixApp Σ f : + isFixApp f = isFixApp (optimize Σ f). +Proof. + funelim (optimize Σ f); cbn -[optimize] => //. + all:rewrite map_InP_spec. + all:rewrite isFixApp_mkApps isFixApp_mkApps //. +Qed. + +Lemma optimize_isConstructApp Σ f : + isConstructApp f = isConstructApp (optimize Σ f). +Proof. + funelim (optimize Σ f); cbn -[optimize] => //. + all:rewrite map_InP_spec. + all:rewrite isConstructApp_mkApps isConstructApp_mkApps //. +Qed. + +Lemma optimize_isPrimApp Σ f : + isPrimApp f = isPrimApp (optimize Σ f). +Proof. + funelim (optimize Σ f); cbn -[optimize] => //. + all:rewrite map_InP_spec. + all:rewrite !isPrimApp_mkApps //. +Qed. + +Lemma lookup_inductive_pars_is_prop_and_pars {Σ ind b pars} : + inductive_isprop_and_pars Σ ind = Some (b, pars) -> + lookup_inductive_pars Σ (inductive_mind ind) = Some pars. +Proof. + rewrite /inductive_isprop_and_pars /lookup_inductive_pars /lookup_inductive /lookup_minductive. + destruct lookup_env => //. + destruct g => /= //. + destruct nth_error => //. congruence. +Qed. + +Lemma constructor_isprop_pars_decl_lookup {Σ ind c b pars decl} : + constructor_isprop_pars_decl Σ ind c = Some (b, pars, decl) -> + lookup_inductive_pars Σ (inductive_mind ind) = Some pars. +Proof. + rewrite /constructor_isprop_pars_decl /lookup_constructor /lookup_inductive_pars /lookup_inductive /lookup_minductive. + destruct lookup_env => //. + destruct g => /= //. + destruct nth_error => //. destruct nth_error => //. congruence. +Qed. + +Lemma constructor_isprop_pars_inductive_pars {Σ ind c b pars decl} : + constructor_isprop_pars_decl Σ ind c = Some (b, pars, decl) -> + inductive_isprop_and_pars Σ ind = Some (b, 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 lookup_constructor_lookup_inductive_pars {Σ ind c mdecl idecl cdecl} : + lookup_constructor Σ ind c = Some (mdecl, idecl, cdecl) -> + lookup_inductive_pars Σ (inductive_mind ind) = Some mdecl.(ind_npars). +Proof. + rewrite /lookup_constructor /lookup_inductive_pars /lookup_inductive /lookup_minductive. + destruct lookup_env => //. + destruct g => /= //. + destruct nth_error => //. destruct nth_error => //. congruence. +Qed. + +Lemma optimize_mkApps_etaexp {Σ : GlobalContextMap.t} fn args : + isEtaExp Σ fn -> + optimize Σ (EAst.mkApps fn args) = EAst.mkApps (optimize Σ fn) (map (optimize Σ) args). +Proof. + destruct (decompose_app fn) eqn:da. + rewrite (decompose_app_inv da). + rewrite isEtaExp_mkApps_napp. now eapply decompose_app_notApp. + destruct construct_viewc eqn:vc. + + move=> /andP[] hl0 etal0. + rewrite -mkApps_app. + rewrite (optimize_mkApps Σ (tConstruct ind n block_args)) // /=. + rewrite optimize_mkApps // /=. + unfold isEtaExp_app in hl0. + destruct lookup_constructor_pars_args as [[pars args']|] eqn:hl => //. + rtoProp. + eapply Nat.leb_le in H. + rewrite (lookup_inductive_pars_constructor_pars_args hl). + rewrite -mkApps_app. f_equal. rewrite map_app. + rewrite skipn_app. len. assert (pars - #|l| = 0) by lia. + now rewrite H1 skipn_0. + + move=> /andP[] etat0 etal0. + rewrite -mkApps_app !optimize_mkApps; try now eapply decompose_app_notApp. + rewrite vc. rewrite -mkApps_app !map_app //. +Qed. + + #[export] Instance Qpreserves_closedn (efl := all_env_flags) Σ : closed_env Σ -> + Qpreserves (fun n x => closedn n x) Σ. +Proof. + intros clΣ. + split. + - red. move=> n t. + destruct t; cbn; intuition auto; try solve [constructor; auto]. + eapply on_evar; solve_all. + eapply on_letin; rtoProp; intuition auto. + eapply on_app; rtoProp; intuition auto. + eapply on_case; rtoProp; intuition auto. solve_all. + eapply on_fix. solve_all. solve_all. + eapply on_cofix; solve_all. + eapply on_prim; solve_all. + - red. intros kn decl. + move/(lookup_env_closed clΣ). + unfold closed_decl. destruct EAst.cst_body => //. + - red. move=> hasapp n t args. rewrite closedn_mkApps. + split; intros; rtoProp; intuition auto; solve_all. + - red. move=> hascase n ci discr brs. simpl. + intros; rtoProp; intuition auto; solve_all. + - red. move=> hasproj n p discr. simpl. + intros; rtoProp; intuition auto; solve_all. + - red. move=> t args clt cll. + eapply closed_substl. solve_all. now rewrite Nat.add_0_r. + - red. move=> n mfix idx. cbn. + intros; rtoProp; intuition auto; solve_all. + - red. move=> n mfix idx. cbn. + intros; rtoProp; intuition auto; solve_all. +Qed. + +Lemma optimize_eval (efl := all_env_flags) {wfl:WcbvFlags} {wcon : with_constructor_as_block = false} {Σ : GlobalContextMap.t} t v : + isEtaExp_env Σ -> + closed_env Σ -> + wf_glob Σ -> + closedn 0 t -> + isEtaExp Σ t -> + eval Σ t v -> + eval (optimize_env Σ) (optimize Σ t) (optimize Σ v). +Proof. + intros etaΣ clΣ wfΣ. + revert t v. + unshelve eapply (eval_preserve_mkApps_ind wfl wcon Σ (fun x y => eval (optimize_env Σ) (optimize Σ x) (optimize Σ y)) + (fun n x => closedn n x) (Qpres := Qpreserves_closedn Σ clΣ)) => //. + { intros. eapply eval_closed; tea. } + all:intros; simpl in *. + (* 1-15: try solve [econstructor; eauto]. *) + all:repeat destruct_nary_times => //. + - rewrite optimize_tApp //. + econstructor; eauto. + + - rewrite optimize_tApp //. simp_optimize in e1. + econstructor; eauto. + rewrite optimize_csubst // in e. now simp_eta in i10. + + - simp_optimize. + rewrite optimize_csubst // in e. + econstructor; eauto. + + - simp_optimize. + set (brs' := map _ brs). cbn -[optimize]. + cbn in H3. + rewrite optimize_mkApps // /= in e0. + apply constructor_isprop_pars_decl_lookup in H2 as H1'. + rewrite H1' in e0. + pose proof (@is_propositional_cstr_optimize Σ ind c). + rewrite H2 in H1. + econstructor; eauto. + * rewrite nth_error_map H3 //. + * len. cbn in H4, H5. rewrite skipn_length. lia. + * cbn -[optimize]. rewrite skipn_0. len. + * cbn -[optimize]. + have etaargs : forallb (isEtaExp Σ) args. + { rewrite isEtaExp_Constructor in i6. + now move/andP: i6 => [] /andP[]. } + rewrite optimize_iota_red // in e. + rewrite closedn_mkApps in i4. now move/andP: i4. + cbn. now eapply nth_error_forallb in H; tea. + + - subst brs. cbn in H4. + rewrite andb_true_r in H4. + rewrite optimize_substl // in e. + eapply All_forallb, All_repeat => //. + eapply All_forallb, All_repeat => //. + rewrite map_optimize_repeat_box in e. + simp_optimize. set (brs' := map _ _). + cbn -[optimize]. eapply eval_iota_sing => //. + now move: (is_propositional_optimize Σ ind); rewrite H2. + + - rewrite optimize_tApp //. + rewrite optimize_mkApps // /= in e1. + simp_optimize in e1. + eapply eval_fix; tea. + * rewrite map_length. + eapply optimize_cunfold_fix; tea. + eapply closed_fix_subst. tea. + move: i8; rewrite closedn_mkApps => /andP[] //. + move: i10; rewrite isEtaExp_mkApps_napp // /= => /andP[] //. simp_eta. + * move: e. + rewrite -[tApp _ _](mkApps_app _ _ [av]). + rewrite -[tApp _ _](mkApps_app _ _ [optimize _ av]). + rewrite optimize_mkApps_etaexp // map_app //. + + - rewrite optimize_tApp //. + rewrite optimize_tApp //. + rewrite optimize_mkApps //. simpl. + simp_optimize. set (mfix' := map _ _). simpl. + rewrite optimize_mkApps // /= in e0. simp_optimize in e0. + eapply eval_fix_value; tea. + eapply optimize_cunfold_fix; eauto. + eapply closed_fix_subst => //. + { move: i4. + rewrite closedn_mkApps. now move/andP => []. } + { move: i6. rewrite isEtaExp_mkApps_napp // /= => /andP[] //. now simp isEtaExp. } + now rewrite map_length. + + - rewrite optimize_tApp //. simp_optimize in e0. + simp_optimize in e1. + eapply eval_fix'; tea. + eapply optimize_cunfold_fix; tea. + { eapply closed_fix_subst => //. } + { simp isEtaExp in i10. } + rewrite optimize_tApp // in e. + + - simp_optimize. + rewrite optimize_mkApps // /= in e0. + simp_optimize in e. + simp_optimize in e0. + set (brs' := map _ _) in *; simpl. + eapply eval_cofix_case; tea. + eapply optimize_cunfold_cofix; tea => //. + { eapply closed_cofix_subst; tea. + move: i5; rewrite closedn_mkApps => /andP[] //. } + { move: i7. rewrite isEtaExp_mkApps_napp // /= => /andP[] //. now simp isEtaExp. } + rewrite optimize_mkApps_etaexp // in e. + + - destruct p as [[ind pars] arg]. + simp_optimize. + simp_optimize in e. + rewrite optimize_mkApps // /= in e0. + simp_optimize in e0. + rewrite optimize_mkApps_etaexp // in e. + simp_optimize in e0. + eapply eval_cofix_proj; tea. + eapply optimize_cunfold_cofix; tea. + { eapply closed_cofix_subst; tea. + move: i4; rewrite closedn_mkApps => /andP[] //. } + { move: i6. rewrite isEtaExp_mkApps_napp // /= => /andP[] //. now simp isEtaExp. } + + - econstructor. red in H |- *. + rewrite lookup_env_optimize H //. + now rewrite /optimize_constant_decl H0. + exact e. + + - simp_optimize. + rewrite optimize_mkApps // /= in e0. + rewrite (constructor_isprop_pars_decl_lookup H2) in e0. + eapply eval_proj; eauto. + move: (@is_propositional_cstr_optimize Σ p.(proj_ind) 0). now rewrite H2. simpl. + len. rewrite skipn_length. cbn in H3. lia. + rewrite nth_error_skipn nth_error_map /= H4 //. + + - simp_optimize. eapply eval_proj_prop => //. + move: (is_propositional_optimize Σ p.(proj_ind)); now rewrite H3. + + - rewrite !optimize_tApp //. + eapply eval_app_cong; tea. + move: H1. eapply contraNN. + rewrite -optimize_isLambda -optimize_isConstructApp -optimize_isFixApp -optimize_isBox -optimize_isPrimApp //. + rewrite -optimize_isFix //. + + - rewrite !optimize_mkApps // /=. + rewrite (lookup_constructor_lookup_inductive_pars H). + eapply eval_mkApps_Construct; tea. + + rewrite lookup_constructor_optimize H //. + + constructor. cbn [atom]. rewrite wcon lookup_constructor_optimize H //. + + rewrite /cstr_arity /=. + move: H0; rewrite /cstr_arity /=. + rewrite skipn_length map_length => ->. lia. + + cbn in H0. eapply All2_skipn, All2_map. + eapply All2_impl; tea; cbn -[optimize]. + intros x y []; auto. + + - depelim X; simp optimize; repeat constructor. + eapply All2_over_undep in a. eapply All2_Set_All2 in ev. eapply All2_All2_Set. solve_all. now destruct b. + now destruct a0. + + - destruct t => //. + all:constructor; eauto. simp optimize. + cbn [atom optimize] in H |- *. + rewrite lookup_constructor_optimize. + destruct lookup_constructor eqn:hl => //. destruct p as [[] ?] => //. +Qed. + +From MetaCoq.Erasure Require Import EEtaExpanded. + +Lemma optimize_declared_constructor {Σ : GlobalContextMap.t} {k mdecl idecl cdecl} : + declared_constructor Σ.(GlobalContextMap.global_decls) k mdecl idecl cdecl -> + declared_constructor (optimize_env Σ) k (optimize_inductive_decl mdecl) idecl cdecl. +Proof. + intros [[] ?]; do 2 split => //. + red in H; red. + rewrite lookup_env_optimize H //. +Qed. + +Lemma lookup_inductive_pars_spec {Σ} {mind} {mdecl} : + declared_minductive Σ mind mdecl -> + lookup_inductive_pars Σ mind = Some (ind_npars mdecl). +Proof. + rewrite /declared_minductive /lookup_inductive_pars /lookup_minductive. + 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 optimize_wellformed {efl} {Σ : GlobalContextMap.t} n t : + cstr_as_blocks = false -> + has_tApp -> + @wf_glob efl Σ -> + @wellformed efl Σ n t -> + @wellformed (switch_no_params efl) (optimize_env Σ) n (optimize Σ t). +Proof. + intros cab hasapp wfΣ. + revert n. + funelim (optimize Σ t); try intros n. + all:cbn -[optimize lookup_constructor lookup_inductive]; simp_optimize; intros. + all:try solve[unfold wf_fix_gen in *; rtoProp; intuition auto; toAll; solve_all]. + - cbn -[optimize]; simp_optimize. intros; rtoProp; intuition auto. + rewrite lookup_env_optimize. destruct lookup_env eqn:hl => // /=. + destruct g => /= //. destruct (cst_body c) => //. + - rewrite cab in H |- *. cbn -[optimize] in *. + rewrite lookup_env_optimize. destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. subst g. + destruct nth_error => //. destruct nth_error => //. + - cbn -[optimize]. + rewrite lookup_env_optimize. cbn in H1. destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. subst g. + destruct nth_error => //. rtoProp; intuition auto. + simp_optimize. toAll; solve_all. + toAll. solve_all. + - cbn -[optimize] in H0 |- *. + rewrite lookup_env_optimize. destruct lookup_env eqn:hl => // /=. + 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 -optimize_isLambda. toAll; solve_all. + - primProp. rtoProp; intuition eauto; solve_all_k 6. + - move:H1; rewrite !wellformed_mkApps //. rtoProp; intuition auto. + toAll; solve_all. toAll; solve_all. + - move:H0; rewrite !wellformed_mkApps //. rtoProp; intuition auto. + move: H1. cbn. rewrite cab. + rewrite lookup_env_optimize. 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. rewrite cab. + rewrite lookup_env_optimize. destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. subst g. + destruct nth_error => //. destruct nth_error => //. + toAll; solve_all. +Qed. + +Lemma optimize_wellformed_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : + cstr_as_blocks = false -> + wf_glob Σ -> + forall n, wellformed Σ n t -> wellformed (optimize_env Σ) n t. +Proof. + intros hcstrs wfΣ. induction t using EInduction.term_forall_list_ind; cbn => //. + all:try solve [intros; unfold wf_fix in *; rtoProp; intuition eauto; solve_all]. + - rewrite lookup_env_optimize //. + destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. subst g. + destruct (cst_body c) => //. + - rewrite lookup_env_optimize //. + destruct lookup_env eqn:hl => // /=; intros; rtoProp; eauto. + destruct g eqn:hg => /= //; intros; rtoProp; eauto. + destruct cstr_as_blocks => //; repeat split; eauto. + destruct nth_error => /= //. + destruct nth_error => /= //. + - rewrite lookup_env_optimize //. + destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. subst g. + destruct nth_error => /= //. + intros; rtoProp; intuition auto; solve_all. + - rewrite lookup_env_optimize //. + destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. + rewrite andb_false_r => //. + destruct nth_error => /= //. + destruct EAst.ind_ctors => /= //. + all:intros; rtoProp; intuition auto; solve_all. + destruct nth_error => //. +Qed. + +Lemma optimize_wellformed_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} d : + cstr_as_blocks = false -> + wf_glob Σ -> + wf_global_decl Σ d -> wf_global_decl (optimize_env Σ) d. +Proof. + intros hcstrs wf; destruct d => /= //. + destruct (cst_body c) => /= //. + now eapply optimize_wellformed_irrel. +Qed. + +Lemma optimize_decl_wf (efl := all_env_flags) {Σ : GlobalContextMap.t} : + wf_glob Σ -> + forall d, wf_global_decl Σ d -> + wf_global_decl (efl := switch_no_params efl) (optimize_env Σ) (optimize_decl Σ d). +Proof. + intros wf d. + destruct d => /= //. + destruct (cst_body c) => /= //. + now apply (optimize_wellformed (Σ := Σ) 0 t). +Qed. + +Lemma fresh_global_optimize_env {Σ : GlobalContextMap.t} kn : + fresh_global kn Σ -> + fresh_global kn (optimize_env Σ). +Proof. + unfold fresh_global. cbn. unfold optimize_env. + induction (GlobalContextMap.global_decls Σ); cbn; constructor; auto. cbn. + now depelim H. depelim H. eauto. +Qed. + +From MetaCoq.Erasure Require Import EProgram. + +Program Fixpoint optimize_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Σ) + end. + +Lemma lookup_minductive_declared_minductive Σ ind mdecl : + lookup_minductive Σ ind = Some mdecl <-> declared_minductive Σ ind mdecl. +Proof. + unfold declared_minductive, lookup_minductive. + destruct lookup_env => /= //. destruct g => /= //; split => //; congruence. +Qed. + +Lemma lookup_minductive_declared_inductive Σ ind mdecl idecl : + lookup_inductive Σ ind = Some (mdecl, idecl) <-> declared_inductive Σ ind mdecl idecl. +Proof. + unfold declared_inductive, lookup_inductive. + rewrite <- lookup_minductive_declared_minductive. + destruct lookup_minductive => /= //. + destruct nth_error eqn:hnth => //. + split. intros [=]. subst. split => //. + intros [[=] hnth']. subst. congruence. + split => //. intros [[=] hnth']. congruence. + split => //. intros [[=] hnth']. +Qed. + +Lemma extends_lookup_inductive_pars {efl : EEnvFlags} {Σ Σ'} : + extends Σ Σ' -> wf_glob Σ' -> + forall ind t, lookup_inductive_pars Σ ind = Some t -> + lookup_inductive_pars Σ' ind = Some t. +Proof. + intros ext wf ind t. + rewrite /lookup_inductive_pars. + destruct (lookup_minductive Σ ind) eqn:hl => /= //. + intros [= <-]. + apply lookup_minductive_declared_minductive in hl. + eapply EExtends.weakening_env_declared_minductive in hl; tea. + apply lookup_minductive_declared_minductive in hl. now rewrite hl. +Qed. + +Lemma optimize_extends {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : + has_tApp -> + extends Σ Σ' -> wf_glob Σ' -> + forall n t, wellformed Σ n t -> optimize Σ t = optimize Σ' t. +Proof. + intros hasapp hext hwf n t. revert n. + funelim (optimize Σ t); intros ?n; simp_optimize => /= //. + all:try solve [intros; unfold wf_fix in *; rtoProp; intuition auto; f_equal; auto; toAll; solve_all]. + - intros; rtoProp; intuition auto. + move: H1; rewrite wellformed_mkApps // => /andP[] wfu wfargs. + rewrite optimize_mkApps // Heq /=. f_equal; eauto. toAll; solve_all. + - rewrite wellformed_mkApps // => /andP[] wfc wfv. + rewrite optimize_mkApps // /=. + rewrite GlobalContextMap.lookup_inductive_pars_spec in Heq. + eapply (extends_lookup_inductive_pars hext hwf) in Heq. + rewrite Heq. f_equal. f_equal. + toAll; solve_all. + - rewrite wellformed_mkApps // => /andP[] wfc wfv. + rewrite optimize_mkApps // /=. + move: Heq. + rewrite GlobalContextMap.lookup_inductive_pars_spec. + unfold wellformed in wfc. move/andP: wfc => [] /andP[] hacc hc bargs. + unfold lookup_inductive_pars. destruct lookup_minductive eqn:heq => //. + unfold lookup_constructor, lookup_inductive in hc. rewrite heq /= // in hc. +Qed. + +Lemma optimize_decl_extends {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : + has_tApp -> + extends Σ Σ' -> wf_glob Σ' -> + forall d, wf_global_decl Σ d -> optimize_decl Σ d = optimize_decl Σ' d. +Proof. + intros hast ext wf []; cbn. + unfold optimize_constant_decl. + destruct (EAst.cst_body c) eqn:hc => /= //. + intros hwf. f_equal. f_equal. f_equal. + eapply optimize_extends => //. exact hwf. + intros _ => //. +Qed. + +From MetaCoq.Erasure Require Import EGenericGlobalMap. + +#[local] +Instance GT : GenTransform := { gen_transform := optimize; gen_transform_inductive_decl := optimize_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 optimize_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 optimize_wellformed => //. +Defined. + +Lemma optimize_extends' {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). +Proof. + intros hast ext wf. + now apply (gen_transform_env_extends' (gt := GTExt efl hast) ext). +Qed. + +Lemma optimize_extends_env {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : + has_tApp -> + extends Σ Σ' -> + wf_glob Σ -> + wf_glob Σ' -> + extends (optimize_env Σ) (optimize_env Σ'). +Proof. + intros hast ext wf. + now apply (gen_transform_extends (gt := GTExt efl hast) ext). +Qed. + +Lemma optimize_env_eq (efl := all_env_flags) (Σ : 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 optimize_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 (optimize_extends' (Σ := Σg') (Σ' := Σg)) => //. + cbn. eapply EExtends.extends_prefix_extends => //. now exists [a]. + cbn. now depelim wf. +Qed. + +Lemma optimize_env_wf (efl := all_env_flags) {Σ : GlobalContextMap.t} : + wf_glob Σ -> wf_glob (efl := switch_no_params efl) (optimize_env Σ). +Proof. + intros wf. + rewrite (optimize_env_eq _ wf). + destruct Σ as [Σ map repr fr]; cbn in *. + induction Σ in map, repr, fr, wf |- *. + - cbn. constructor. + - cbn. + set (Σg := GlobalContextMap.make Σ (fresh_globals_cons_inv fr)). + constructor; eauto. + eapply (IHΣ (GlobalContextMap.map Σg) (GlobalContextMap.repr Σg)). now depelim wf. + depelim wf. cbn. + rewrite -(optimize_env_eq Σg). now cbn. cbn. + now eapply (optimize_decl_wf (Σ:=Σg)). + rewrite -(optimize_env_eq Σg). now depelim wf. + eapply fresh_global_optimize_env. now depelim fr. +Qed. + +Lemma optimize_program_wf (efl := all_env_flags) (p : eprogram_env) : + wf_eprogram_env efl p -> + wf_eprogram (switch_no_params efl) (optimize_program p). +Proof. + intros []; split => //. + eapply (optimize_env_wf H). + now eapply optimize_wellformed. +Qed. + +Lemma optimize_expanded {Σ : GlobalContextMap.t} {t} : expanded Σ t -> expanded (optimize_env Σ) (optimize Σ t). +Proof. + induction 1 using expanded_ind. + all:try solve[simp_optimize; constructor; eauto; solve_all]. + - rewrite optimize_mkApps_etaexp. now eapply expanded_isEtaExp. + eapply expanded_mkApps_expanded => //. solve_all. + - simp_optimize; constructor; eauto. solve_all. + rewrite -optimize_isLambda //. + - rewrite optimize_mkApps // /=. + rewrite (lookup_inductive_pars_spec (proj1 (proj1 H))). + eapply expanded_tConstruct_app. + eapply optimize_declared_constructor; tea. + len. rewrite skipn_length /= /EAst.cstr_arity /=. + rewrite /EAst.cstr_arity in H0. lia. + solve_all. eapply All_skipn. solve_all. +Qed. + +Lemma optimize_expanded_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded Σ t -> expanded (optimize_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. + cbn. rewrite /EAst.cstr_arity in H0. lia. solve_all. +Qed. + +Lemma optimize_expanded_decl {Σ : GlobalContextMap.t} t : expanded_decl Σ t -> expanded_decl (optimize_env Σ) (optimize_decl Σ t). +Proof. + destruct t as [[[]]|] => /= //. + unfold expanded_constant_decl => /=. + apply optimize_expanded. +Qed. + +Lemma optimize_env_expanded (efl := all_env_flags) {Σ : GlobalContextMap.t} : + wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (optimize_env Σ). +Proof. + unfold expanded_global_env. + intros wf. rewrite optimize_env_eq //. + destruct Σ as [Σ map repr wf']; cbn in *. + intros exp; induction exp in map, repr, wf', wf |- *; cbn. + - constructor; auto. + - set (Σ' := GlobalContextMap.make Σ (fresh_globals_cons_inv wf')). + constructor; auto. + eapply IHexp. eapply Σ'. now depelim wf. cbn. + eapply (optimize_expanded_decl (Σ := Σ')) in H. + rewrite -(optimize_env_eq Σ'). cbn. now depelim wf. + exact H. +Qed. + +Import EProgram. + +Lemma optimize_program_expanded (efl := all_env_flags) (p : eprogram_env) : + wf_eprogram_env efl p -> + expanded_eprogram_env_cstrs p -> + expanded_eprogram_cstrs (optimize_program p). +Proof. + unfold expanded_eprogram_env_cstrs, expanded_eprogram_cstrs. + move=> [] wfe wft. move/andP => [] etaenv etap. + apply/andP; split. + eapply expanded_global_env_isEtaExp_env, optimize_env_expanded => //. + now eapply isEtaExp_env_expanded_global_env. + now eapply expanded_isEtaExp, optimize_expanded, isEtaExp_expanded. +Qed. + + + tConstruct ind n (map optimize args) + | tPrim p => tPrim (map_prim optimize p) + end. + + Lemma optimize_mkApps f l : optimize (mkApps f l) = mkApps (optimize f) (map optimize l). + Proof using Type. + induction l using rev_ind; simpl; auto. + now rewrite mkApps_app /= IHl map_app /= mkApps_app /=. + Qed. + + Lemma map_optimize_repeat_box n : map optimize (repeat tBox n) = repeat tBox n. + Proof using Type. by rewrite map_repeat. Qed. + + (* move to globalenv *) + + Lemma isLambda_optimize t : isLambda t -> isLambda (optimize t). + Proof. destruct t => //. Qed. + Lemma isBox_optimize t : isBox t -> isBox (optimize t). + Proof. destruct t => //. Qed. + + Lemma wf_optimize t k : + wf_glob Σ -> + wellformed Σ k t -> wellformed Σ k (optimize t). + Proof using Type. + intros wfΣ. + induction t in k |- * using EInduction.term_forall_list_ind; simpl; auto; + intros; try easy; + rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + unfold wf_fix_gen, test_def in *; + simpl closed in *; try solve [simpl subst; simpl closed; f_equal; auto; rtoProp; solve_all]; try easy. + - rtoProp. split; eauto. destruct args; eauto. + - move/andP: H => [] /andP[] -> clt cll /=. + rewrite IHt //=. solve_all. + - rewrite GlobalContextMap.lookup_projection_spec. + destruct lookup_projection as [[[[mdecl idecl] cdecl] pdecl]|] eqn:hl; auto => //. + simpl. + have arglen := wellformed_projection_args wfΣ hl. + apply lookup_projection_lookup_constructor, lookup_constructor_lookup_inductive in hl. + rewrite hl /= andb_true_r. + rewrite IHt //=; len. apply Nat.ltb_lt. + lia. + - len. rtoProp; solve_all. solve_all. + now eapply isLambda_optimize. solve_all. + - len. rtoProp; repeat solve_all. + - rewrite test_prim_map. rtoProp; intuition eauto; solve_all. + Qed. + + Lemma optimize_csubst {a k b} n : + wf_glob Σ -> + wellformed Σ (k + n) b -> + optimize (ECSubst.csubst a k b) = ECSubst.csubst (optimize a) k (optimize b). + Proof using Type. + intros wfΣ. + induction b in k |- * using EInduction.term_forall_list_ind; simpl; auto; + intros wft; try easy; + rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + unfold wf_fix, test_def in *; + simpl closed in *; try solve [simpl subst; simpl closed; f_equal; auto; rtoProp; solve_all]; try easy. + - destruct (k ?= n0)%nat; auto. + - f_equal. rtoProp. now destruct args; inv H0. + - move/andP: wft => [] /andP[] hi hb hl. rewrite IHb. f_equal. unfold on_snd; solve_all. + repeat toAll. f_equal. solve_all. unfold on_snd; cbn. f_equal. + rewrite a0 //. now rewrite -Nat.add_assoc. + - move/andP: wft => [] hp hb. + rewrite GlobalContextMap.lookup_projection_spec. + destruct lookup_projection as [[[[mdecl idecl] cdecl] pdecl]|] eqn:hl => /= //. + f_equal; eauto. f_equal. len. f_equal. + have arglen := wellformed_projection_args wfΣ hl. + case: Nat.compare_spec. lia. lia. + auto. + - f_equal. move/andP: wft => [hlam /andP[] hidx hb]. + solve_all. unfold map_def. f_equal. + eapply a0. now rewrite -Nat.add_assoc. + - f_equal. move/andP: wft => [hidx hb]. + solve_all. unfold map_def. f_equal. + eapply a0. now rewrite -Nat.add_assoc. + Qed. + + Lemma optimize_substl s t : + wf_glob Σ -> + forallb (wellformed Σ 0) s -> + wellformed Σ #|s| t -> + optimize (substl s t) = substl (map optimize s) (optimize t). + Proof using Type. + intros wfΣ. induction s in t |- *; simpl; auto. + move/andP => [] cla cls wft. + rewrite IHs //. eapply wellformed_csubst => //. + f_equal. rewrite (optimize_csubst (S #|s|)) //. + Qed. + + Lemma optimize_iota_red pars args br : + wf_glob Σ -> + forallb (wellformed Σ 0) args -> + wellformed Σ #|skipn pars args| br.2 -> + optimize (EGlobalEnv.iota_red pars args br) = EGlobalEnv.iota_red pars (map optimize args) (on_snd optimize br). + Proof using Type. + intros wfΣ wfa wfbr. + unfold EGlobalEnv.iota_red. + rewrite optimize_substl //. + rewrite forallb_rev forallb_skipn //. + now rewrite List.rev_length. + 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). + Proof using Type. + unfold EGlobalEnv.fix_subst. + rewrite map_length. + generalize #|mfix|. + induction n; simpl; auto. + f_equal; auto. + Qed. + + Lemma optimize_cofix_subst mfix : EGlobalEnv.cofix_subst (map (map_def optimize) mfix) = map optimize (EGlobalEnv.cofix_subst mfix). + Proof using Type. + unfold EGlobalEnv.cofix_subst. + rewrite map_length. + generalize #|mfix|. + induction n; simpl; auto. + f_equal; auto. + Qed. + + Lemma optimize_cunfold_fix mfix idx n f : + wf_glob Σ -> + wellformed Σ 0 (tFix mfix idx) -> + cunfold_fix mfix idx = Some (n, f) -> + cunfold_fix (map (map_def optimize) mfix) idx = Some (n, optimize f). + Proof using Type. + intros wfΣ hfix. + unfold cunfold_fix. + rewrite nth_error_map. + cbn in hfix. move/andP: hfix => [] hlam /andP[] hidx hfix. + destruct nth_error eqn:hnth => //. + intros [= <- <-] => /=. f_equal. + rewrite optimize_substl //. eapply wellformed_fix_subst => //. + rewrite fix_subst_length. + eapply nth_error_forallb in hfix; tea. now rewrite Nat.add_0_r in hfix. + now rewrite optimize_fix_subst. + Qed. + + Lemma optimize_cunfold_cofix mfix idx n f : + wf_glob Σ -> + wellformed Σ 0 (tCoFix mfix idx) -> + cunfold_cofix mfix idx = Some (n, f) -> + cunfold_cofix (map (map_def optimize) mfix) idx = Some (n, optimize f). + Proof using Type. + intros wfΣ hfix. + unfold cunfold_cofix. + rewrite nth_error_map. + cbn in hfix. move/andP: hfix => [] hidx hfix. + destruct nth_error eqn:hnth => //. + intros [= <- <-] => /=. f_equal. + rewrite optimize_substl //. eapply wellformed_cofix_subst => //. + rewrite cofix_subst_length. + eapply nth_error_forallb in hfix; tea. now rewrite Nat.add_0_r in hfix. + now rewrite optimize_cofix_subst. + Qed. + +End optimize. + +Definition optimize_constant_decl Σ cb := + {| cst_body := option_map (optimize Σ) cb.(cst_body) |}. + +Definition optimize_decl Σ d := + match d with + | ConstantDecl cb => ConstantDecl (optimize_constant_decl Σ cb) + | InductiveDecl idecl => InductiveDecl idecl + end. + +Definition optimize_env Σ := + map (on_snd (optimize_decl Σ)) Σ.(GlobalContextMap.global_decls). + +Import EnvMap. + +Program Fixpoint optimize_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Σ) + end. + +Import EGlobalEnv EExtends. + +Lemma extends_lookup_projection {efl : EEnvFlags} {Σ Σ' p} : extends Σ Σ' -> wf_glob Σ' -> + isSome (lookup_projection Σ p) -> + lookup_projection Σ p = lookup_projection Σ' p. +Proof. + intros ext wf; cbn. + unfold lookup_projection. + destruct lookup_constructor as [[[mdecl idecl] cdecl]|] eqn:hl => //. + simpl. + rewrite (extends_lookup_constructor wf ext _ _ _ hl) //. +Qed. + +Lemma wellformed_optimize_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : + forall n, EWellformed.wellformed Σ n t -> + forall {Σ' : GlobalContextMap.t}, extends Σ Σ' -> wf_glob Σ' -> + optimize Σ t = optimize Σ' t. +Proof. + induction t using EInduction.term_forall_list_ind; cbn -[lookup_constant lookup_inductive + GlobalContextMap.lookup_projection]; intros => //. + all:unfold wf_fix_gen in *; rtoProp; intuition auto. + 5:{ destruct cstr_as_blocks; rtoProp. f_equal; eauto; solve_all. destruct args; cbn in *; eauto. } + all:f_equal; eauto; solve_all. + - rewrite !GlobalContextMap.lookup_projection_spec. + rewrite -(extends_lookup_projection H0 H1 H3). + destruct lookup_projection as [[[[]]]|]. f_equal; eauto. + now cbn in H3. +Qed. + +Lemma wellformed_optimize_decl_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : + wf_global_decl Σ t -> + forall {Σ' : GlobalContextMap.t}, extends Σ Σ' -> wf_glob Σ' -> + optimize_decl Σ t = optimize_decl Σ' t. +Proof. + destruct t => /= //. + intros wf Σ' ext wf'. f_equal. unfold optimize_constant_decl. f_equal. + destruct (cst_body c) => /= //. f_equal. + now eapply wellformed_optimize_extends. +Qed. + +Lemma lookup_env_optimize_env_Some {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn d : + wf_glob Σ -> + GlobalContextMap.lookup_env Σ kn = Some d -> + ∑ Σ' : GlobalContextMap.t, + [× extends_prefix Σ' Σ, wf_global_decl Σ' d & + lookup_env (optimize_env Σ) kn = Some (optimize_decl Σ' d)]. +Proof. + rewrite GlobalContextMap.lookup_env_spec. + destruct Σ as [Σ map repr wf]. + induction Σ in map, repr, wf |- *; simpl; auto => //. + intros wfg. + case: eqb_specT => //. + - intros ->. cbn. intros [= <-]. + 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. + 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 Σ')). + cbn in IHΣ. 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. + eapply map_ext_in => kn hin. unfold on_snd. f_equal. + symmetry. eapply wellformed_optimize_decl_extends => //. cbn. + eapply lookup_env_In in hin. 2:now depelim wfg. + depelim wfg. eapply lookup_env_wellformed; tea. + 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). +Proof. + induction Σ; cbn; auto. + case: eqb_spec => //. +Qed. + +Lemma lookup_env_optimize_env_None {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn : + GlobalContextMap.lookup_env Σ kn = None -> + lookup_env (optimize_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 : + wf_glob Σ -> + lookup_env (optimize_env Σ) kn = option_map (optimize_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']] => /=. + rewrite hl'. f_equal. + eapply wellformed_optimize_decl_extends; eauto. + now eapply extends_prefix_extends. auto. + + - cbn. now eapply lookup_env_optimize_env_None in hl. +Qed. + +Lemma is_propositional_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : + wf_glob Σ -> + inductive_isprop_and_pars Σ ind = inductive_isprop_and_pars (optimize_env Σ) ind. +Proof. + rewrite /inductive_isprop_and_pars => wf. + rewrite /lookup_inductive /lookup_minductive. + rewrite (lookup_env_optimize (inductive_mind ind) wf). + rewrite /GlobalContextMap.inductive_isprop_and_pars /GlobalContextMap.lookup_inductive + /GlobalContextMap.lookup_minductive. + 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. +Proof. + rewrite /constructor_isprop_pars_decl => wf. + rewrite /lookup_constructor /lookup_inductive /lookup_minductive. + rewrite (lookup_env_optimize (inductive_mind ind) wf). + rewrite /GlobalContextMap.inductive_isprop_and_pars /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 -> + nth_error brs c = Some br -> + #|skipn pars args| = #|br.1| -> + closedn #|br.1| br.2 -> + closed (iota_red pars args br). +Proof. + intros clargs hnth hskip clbr. + rewrite /iota_red. + eapply ECSubst.closed_substl => //. + now rewrite forallb_rev forallb_skipn. + now rewrite List.rev_length hskip Nat.add_0_r. +Qed. + +Definition disable_prop_cases fl := {| with_prop_case := false; with_guarded_fix := fl.(@with_guarded_fix) ; with_constructor_as_block := false |}. + +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_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} {ind c} : + wf_glob Σ -> + lookup_constructor Σ ind c = lookup_constructor (optimize_env Σ) ind c. +Proof. + intros wfΣ. rewrite /lookup_constructor /lookup_inductive /lookup_minductive. + rewrite lookup_env_optimize // /=. destruct lookup_env => // /=. + destruct g => //. +Qed. + +Lemma lookup_projection_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} {p} : + wf_glob Σ -> + lookup_projection Σ p = lookup_projection (optimize_env Σ) p. +Proof. + intros wfΣ. rewrite /lookup_projection. + rewrite -lookup_constructor_optimize //. +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 Σ k f args : reflect (wellformed Σ k f /\ forallb (wellformed Σ k) args) (wellformed Σ k (mkApps f args)). +Proof. + rewrite wellformed_mkApps //. eapply andP. +Qed. + +Lemma substl_closed s t : closed t -> substl s t = t. +Proof. + induction s in t |- *; cbn => //. + intros clt. rewrite csubst_closed //. now apply IHs. +Qed. + +Lemma substl_rel s k a : + closed a -> + nth_error s k = Some a -> + substl s (tRel k) = a. +Proof. + intros cla. + induction s in k |- *. + - rewrite nth_error_nil //. + - destruct k => //=. + * intros [= ->]. rewrite substl_closed //. + * intros hnth. now apply IHs. +Qed. + + +Lemma optimize_correct {fl} {wcon : with_constructor_as_block = false} { Σ : GlobalContextMap.t} t v : + wf_glob Σ -> + @eval fl Σ t v -> + wellformed Σ 0 t -> + @eval fl (optimize_env Σ) (optimize Σ t) (optimize Σ v). +Proof. + intros wfΣ ev. + induction ev; simpl in *. + + - move/andP => [] cla clt. econstructor; eauto. + - move/andP => [] clf cla. + eapply eval_wellformed in ev2; tea => //. + eapply eval_wellformed in ev1; tea => //. + econstructor; eauto. + rewrite -(optimize_csubst _ 1) //. + apply IHev3. eapply wellformed_csubst => //. + + - move/andP => [] clb0 clb1. + intuition auto. + eapply eval_wellformed in ev1; tea => //. + forward IHev2 by eapply wellformed_csubst => //. + econstructor; eauto. rewrite -(optimize_csubst _ 1) //. + + - move/andP => [] /andP[] hl wfd wfbrs. rewrite optimize_mkApps in IHev1. + eapply eval_wellformed in ev1 => //. + move/wf_mkApps: ev1 => [] wfc' wfargs. + eapply nth_error_forallb in wfbrs; tea. + rewrite Nat.add_0_r in wfbrs. + forward IHev2. eapply wellformed_iota_red; tea => //. + rewrite optimize_iota_red in IHev2 => //. now rewrite e4. + econstructor; eauto. + rewrite -is_propositional_cstr_optimize //. tea. + rewrite nth_error_map e2 //. len. len. + + - congruence. + + - move/andP => [] /andP[] hl wfd wfbrs. + forward IHev2. eapply wellformed_substl; tea => //. + rewrite forallb_repeat //. len. + rewrite e1 /= Nat.add_0_r in wfbrs. now move/andP: wfbrs. + rewrite optimize_substl in IHev2 => //. + rewrite forallb_repeat //. len. + rewrite e1 /= Nat.add_0_r in wfbrs. now move/andP: wfbrs. + eapply eval_iota_sing => //; eauto. + rewrite -is_propositional_optimize //. + rewrite e1 //. simpl. + rewrite map_repeat in IHev2 => //. + + - move/andP => [] clf cla. rewrite optimize_mkApps in IHev1. + simpl in *. + eapply eval_wellformed in ev1 => //. + move/wf_mkApps: ev1 => [] wff wfargs. + eapply eval_fix; eauto. + rewrite map_length. + eapply optimize_cunfold_fix; tea. + rewrite optimize_mkApps in IHev3. apply IHev3. + rewrite wellformed_mkApps // wfargs. + eapply eval_wellformed in ev2; tas => //. rewrite ev2 /= !andb_true_r. + eapply wellformed_cunfold_fix; tea. + + - move/andP => [] clf cla. + eapply eval_wellformed in ev1 => //. + move/wf_mkApps: ev1 => [] clfix clargs. + eapply eval_wellformed in ev2; tas => //. + rewrite optimize_mkApps in IHev1 |- *. + simpl in *. eapply eval_fix_value. auto. auto. auto. + eapply optimize_cunfold_fix; eauto. + now rewrite map_length. + + - move/andP => [] clf cla. + eapply eval_wellformed in ev1 => //. + eapply eval_wellformed in ev2; tas => //. + simpl in *. eapply eval_fix'. auto. auto. + eapply optimize_cunfold_fix; eauto. + eapply IHev2; tea. eapply IHev3. + apply/andP; split => //. + eapply wellformed_cunfold_fix; tea. now cbn. + + - move/andP => [] /andP[] hl cd clbrs. specialize (IHev1 cd). + eapply eval_wellformed in ev1; tea => //. + move/wf_mkApps: ev1 => [] wfcof wfargs. + forward IHev2. + rewrite hl wellformed_mkApps // /= wfargs clbrs !andb_true_r. + eapply wellformed_cunfold_cofix; tea => //. + rewrite !optimize_mkApps /= in IHev1, IHev2. + eapply eval_cofix_case. tea. + eapply optimize_cunfold_cofix; tea. + exact IHev2. + + - move/andP => [] hl hd. + rewrite GlobalContextMap.lookup_projection_spec in IHev2 |- *. + destruct lookup_projection as [[[[mdecl idecl] cdecl] pdecl]|] eqn:hl' => //. + eapply eval_wellformed in ev1 => //. + move/wf_mkApps: ev1 => [] wfcof wfargs. + forward IHev2. + { rewrite /= wellformed_mkApps // wfargs andb_true_r. + eapply wellformed_cunfold_cofix; tea. } + rewrite optimize_mkApps /= in IHev1. + eapply eval_cofix_case. eauto. + eapply optimize_cunfold_cofix; tea. + rewrite optimize_mkApps in IHev2 => //. + + - rewrite /declared_constant in isdecl. + move: (lookup_env_optimize c wfΣ). + rewrite isdecl /= //. + intros hl. + econstructor; tea. cbn. rewrite e //. + apply IHev. + eapply lookup_env_wellformed in wfΣ; tea. + move: wfΣ. rewrite /wf_global_decl /= e //. + + - move=> /andP[] iss cld. + rewrite GlobalContextMap.lookup_projection_spec. + eapply eval_wellformed in ev1; tea => //. + move/wf_mkApps: ev1 => [] wfc wfargs. + destruct lookup_projection as [[[[mdecl idecl] cdecl'] pdecl]|] eqn:hl' => //. + pose proof (lookup_projection_lookup_constructor hl'). + rewrite (constructor_isprop_pars_decl_constructor H) in e1. noconf e1. + forward IHev1 by auto. + forward IHev2. eapply nth_error_forallb in wfargs; tea. + rewrite optimize_mkApps /= in IHev1. + eapply eval_iota; tea. + rewrite /constructor_isprop_pars_decl -lookup_constructor_optimize // H //= //. + rewrite H0 H1; reflexivity. cbn. reflexivity. len. len. + rewrite skipn_length. lia. + unfold iota_red. cbn. + rewrite (substl_rel _ _ (optimize Σ a)) => //. + { eapply nth_error_forallb in wfargs; tea. + eapply wf_optimize in wfargs => //. + now eapply wellformed_closed in wfargs. } + pose proof (wellformed_projection_args wfΣ hl'). cbn in H1. + rewrite nth_error_rev. len. rewrite skipn_length. lia. + rewrite List.rev_involutive. len. rewrite skipn_length. + rewrite nth_error_skipn nth_error_map. + rewrite e2 -H1. + assert((ind_npars mdecl + cstr_nargs cdecl - ind_npars mdecl) = cstr_nargs cdecl) by lia. + rewrite H3. + eapply (f_equal (option_map (optimize Σ))) in e3. + cbn in e3. rewrite -e3. f_equal. f_equal. lia. + + - congruence. + + - move=> /andP[] iss cld. + rewrite GlobalContextMap.lookup_projection_spec. + destruct lookup_projection as [[[[mdecl idecl] cdecl'] pdecl]|] eqn:hl' => //. + pose proof (lookup_projection_lookup_constructor hl'). + simpl in H. + move: e0. rewrite /inductive_isprop_and_pars. + rewrite (lookup_constructor_lookup_inductive H) /=. + intros [= eq <-]. + eapply eval_iota_sing => //; eauto. + rewrite -is_propositional_optimize // /inductive_isprop_and_pars + (lookup_constructor_lookup_inductive H) //=. congruence. + have lenarg := wellformed_projection_args wfΣ hl'. + rewrite (substl_rel _ _ tBox) => //. + { rewrite nth_error_repeat //. len. } + now constructor. + + - move/andP=> [] clf cla. + rewrite optimize_mkApps. + eapply eval_construct; tea. + rewrite -lookup_constructor_optimize //. exact e0. + rewrite optimize_mkApps in IHev1. now eapply IHev1. + now len. + now eapply IHev2. + + - congruence. + + - move/andP => [] clf cla. + specialize (IHev1 clf). specialize (IHev2 cla). + eapply eval_app_cong; eauto. + eapply eval_to_value in ev1. + destruct ev1; simpl in *; eauto. + * depelim a0. + + destruct t => //; rewrite optimize_mkApps /=. + + now rewrite /= !orb_false_r orb_true_r in i. + * destruct with_guarded_fix. + + move: i. + rewrite !negb_or. + rewrite optimize_mkApps !isFixApp_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. + rewrite !andb_true_r. + rtoProp; intuition auto; destruct v => /= //. + + move: i. + rewrite !negb_or. + rewrite optimize_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. + destruct v => /= //. + - intros; rtoProp; intuition eauto. + depelim X; repeat constructor. + eapply All2_over_undep in a. + eapply All2_Set_All2 in ev. eapply All2_All2_Set. primProp. + subst a0 a'; cbn in *. depelim H0; cbn in *. intuition auto; solve_all. + primProp; depelim H0; intuition eauto. + - destruct t => //. + all:constructor; eauto. + cbn [atom optimize] in i |- *. + rewrite -lookup_constructor_optimize //. + destruct args; cbn in *; eauto. +Qed. + +From MetaCoq.Erasure Require Import EEtaExpanded. + +Lemma optimize_expanded {Σ : GlobalContextMap.t} t : expanded Σ t -> expanded Σ (optimize Σ t). +Proof. + induction 1 using expanded_ind. + all:try solve[constructor; eauto; solve_all]. + all:rewrite ?optimize_mkApps. + - eapply expanded_mkApps_expanded => //. solve_all. + - cbn -[GlobalContextMap.inductive_isprop_and_pars]. + rewrite GlobalContextMap.lookup_projection_spec. + destruct lookup_projection as [[[[mdecl idecl] cdecl] pdecl]|] => /= //. + 2:constructor; eauto; solve_all. + destruct proj as [[ind npars] arg]. + econstructor; eauto. repeat constructor. + - cbn. eapply expanded_tFix. solve_all. + rewrite isLambda_optimize //. + - 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. +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. +Qed. + +Lemma optimize_expanded_decl {Σ : GlobalContextMap.t} t : expanded_decl Σ t -> expanded_decl Σ (optimize_decl Σ t). +Proof. + destruct t as [[[]]|] => /= //. + unfold expanded_constant_decl => /=. + apply optimize_expanded. +Qed. + +Lemma optimize_expanded_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded_decl Σ t -> expanded_decl (optimize_env Σ) t. +Proof. + destruct t as [[[]]|] => /= //. + unfold expanded_constant_decl => /=. + apply optimize_expanded_irrel. +Qed. + +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). +Proof. + intros hbox hrel wfΣ. + induction t in n |- * using EInduction.term_forall_list_ind => //. + all:try solve [cbn; rtoProp; intuition auto; solve_all]. + - cbn -[lookup_constructor_pars_args]. intros. rtoProp. repeat split; eauto. + destruct cstr_as_blocks; rtoProp; eauto. + destruct lookup_constructor_pars_args as [ [] | ]; eauto. split; len. solve_all. split; eauto. + solve_all. now destruct args; invs H0. + - cbn. move/andP => [] /andP[] hast hl wft. + rewrite GlobalContextMap.lookup_projection_spec. + destruct lookup_projection as [[[[mdecl idecl] cdecl] pdecl]|] eqn:hl'; auto => //. + simpl. + rewrite (lookup_constructor_lookup_inductive (lookup_projection_lookup_constructor hl')) /=. + rewrite hrel IHt //= andb_true_r. + have hargs' := wellformed_projection_args wfΣ hl'. + apply Nat.ltb_lt. len. + - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now eapply isLambda_optimize. + - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. + - cbn. rtoProp; intuition eauto; solve_all_k 6. +Qed. + +Import EWellformed. + +Lemma optimize_wellformed_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : + wf_glob Σ -> + forall n, wellformed (efl := disable_projections_env_flag efl) Σ n t -> + wellformed (efl := disable_projections_env_flag efl) (optimize_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 //. + destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. + repeat (rtoProp; intuition auto). + destruct has_axioms => //. cbn in *. + destruct (cst_body c) => //. + - rewrite lookup_env_optimize //. + 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_optimize //. + destruct lookup_env eqn:hl => // /=. + destruct g eqn:hg => /= //. subst g. + destruct nth_error => /= //. + intros; rtoProp; intuition auto; solve_all. +Qed. + +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; 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. + 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 _ _ => has_tBox /\ has_tRel |}; auto. + intros Σ n t [] wfΣ wft. + now apply optimize_wellformed. +Defined. + +Lemma optimize_extends_env {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : + has_tApp -> + extends Σ Σ' -> + wf_glob Σ -> + wf_glob Σ' -> + extends (optimize_env Σ) (optimize_env Σ'). +Proof. + 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). +Proof. + intros wf. + now apply (gen_transform_env_eq (gt := GTExt efl)). +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. + +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 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 hasre wfg. + eapply (gen_transform_env_wf (gt := GTExt efl)) => //. + now apply Pre_glob. +Qed. + +Definition optimize_program (p : eprogram_env) := + (optimize_env p.1, optimize p.1 p.2). + +Definition optimize_program_wf {efl : EEnvFlags} (p : eprogram_env) {hastbox : has_tBox} {hastrel : has_tRel} : + wf_eprogram_env efl p -> wf_eprogram (disable_projections_env_flag efl) (optimize_program p). +Proof. + intros []; split. + now eapply optimize_env_wf. + cbn. now eapply optimize_wellformed. +Qed. + +Definition optimize_program_expanded {efl} (p : eprogram_env) : + wf_eprogram_env efl p -> + expanded_eprogram_env_cstrs p -> expanded_eprogram_cstrs (optimize_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 => //. + now eapply isEtaExp_env_expanded_global_env. + eapply expanded_isEtaExp. + eapply optimize_expanded_irrel => //. + now apply optimize_expanded, isEtaExp_expanded. +Qed. From 3be71de60284a8693365e287137845aa4256cf67 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 22 Feb 2024 18:12:03 +0100 Subject: [PATCH 14/28] Implement unboxing structurally (cstrs as blocks version) --- erasure-plugin/_PluginProject.in | 2 + .../src/metacoq_erasure_plugin.mlpack | 1 + erasure-plugin/theories/ETransform.v | 41 + erasure-plugin/theories/Erasure.v | 11 +- erasure/_CoqProject.in | 1 + erasure/theories/{Unboxing.v => EUnboxing.v} | 1043 ++++++++--------- test-suite/erasure_live_test.v | 4 +- 7 files changed, 561 insertions(+), 542 deletions(-) rename erasure/theories/{Unboxing.v => EUnboxing.v} (65%) diff --git a/erasure-plugin/_PluginProject.in b/erasure-plugin/_PluginProject.in index d4a2f6132..dca62f1cd 100644 --- a/erasure-plugin/_PluginProject.in +++ b/erasure-plugin/_PluginProject.in @@ -122,6 +122,8 @@ src/eCoInductiveToInductive.mli src/eCoInductiveToInductive.ml src/eReorderCstrs.mli src/eReorderCstrs.ml +src/eUnboxing.mli +src/eUnboxing.ml src/eTransform.mli src/eTransform.ml src/erasure0.mli diff --git a/erasure-plugin/src/metacoq_erasure_plugin.mlpack b/erasure-plugin/src/metacoq_erasure_plugin.mlpack index 5471ffbee..c689fdf8b 100644 --- a/erasure-plugin/src/metacoq_erasure_plugin.mlpack +++ b/erasure-plugin/src/metacoq_erasure_plugin.mlpack @@ -66,6 +66,7 @@ EInlineProjections EConstructorsAsBlocks ECoInductiveToInductive EReorderCstrs +EUnboxing EProgram OptimizePropDiscr diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index 71d8f2118..b38027537 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -1063,6 +1063,47 @@ Axiom trust_reorder_cstrs_transformation_ext' : forall (efl : EEnvFlags) (wfl : TransformExt.t (reorder_cstrs_transformation efl wfl m) extends_eprogram extends_eprogram. +From MetaCoq.Erasure Require Import EUnboxing. + +Axiom trust_unboxing_wf : + forall efl : EEnvFlags, + WcbvFlags -> + forall (input : Transform.program _ term), + wf_eprogram_env efl input -> wf_eprogram efl (unbox_program input). +Axiom trust_unboxing_pres : + forall (efl : EEnvFlags) (wfl : WcbvFlags) (p : Transform.program _ term) + (v : term), + wf_eprogram_env efl p -> + eval_eprogram_env wfl p v -> exists v' : term, eval_eprogram wfl (unbox_program p) v' /\ v' = unbox p.1 v. + +Program Definition unbox_transformation (efl : EEnvFlags) (wfl : WcbvFlags) : + Transform.t _ _ EAst.term EAst.term _ _ + (eval_eprogram_env wfl) (eval_eprogram wfl) := + {| name := "unbox singleton constructors "; + transform p _ := EUnboxing.unbox_program p ; + pre p := wf_eprogram_env efl p ; + post p := wf_eprogram efl p ; + obseq p hp p' v v' := v' = unbox p.1 v |}. + +Next Obligation. + move=> efl wfl m. cbn. now apply trust_unboxing_wf. +Qed. +Next Obligation. + red. eapply trust_unboxing_pres. +Qed. + +#[global] +Axiom trust_unbox_transformation_ext : + forall (efl : EEnvFlags) (wfl : WcbvFlags), + TransformExt.t (unbox_transformation efl wfl) + (fun p p' => extends p.1 p'.1) (fun p p' => extends p.1 p'.1). + +#[global] +Axiom trust_unbox_transformation_ext' : + forall (efl : EEnvFlags) (wfl : WcbvFlags), + TransformExt.t (unbox_transformation efl wfl) + extends_eprogram_env extends_eprogram. + Program Definition optional_transform {env env' term term' value value' eval eval'} (activate : bool) (tr : Transform.t env env' term term' value value' eval eval') : if activate then Transform.t env env' term term' value value' eval eval' diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 1459e58d4..1eb37df4d 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -93,16 +93,17 @@ Definition final_wcbv_flags := {| with_constructor_as_block := true |}. Program Definition optional_unsafe_transforms econf := + let efl := EConstructorsAsBlocks.switch_cstr_as_blocks + (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags)) in ETransform.optional_self_transform econf.(enable_cofix_to_fix) ((* Rebuild the efficient lookup table *) - rebuild_wf_env_transform (efl := EConstructorsAsBlocks.switch_cstr_as_blocks - (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags))) false false ▷ + rebuild_wf_env_transform (efl := efl) false false ▷ (* Coinductives & cofixpoints are translated to inductive types and thunked fixpoints *) - let efl := EConstructorsAsBlocks.switch_cstr_as_blocks - (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags)) in coinductive_to_inductive_transformation efl (has_app := eq_refl) (has_box := eq_refl) (has_rel := eq_refl) (has_pars := eq_refl) (has_cstrblocks := eq_refl) ▷ - reorder_cstrs_transformation efl final_wcbv_flags econf.(inductives_mapping)). + reorder_cstrs_transformation efl final_wcbv_flags econf.(inductives_mapping) ▷ + rebuild_wf_env_transform (efl := efl) false false ▷ + unbox_transformation efl final_wcbv_flags). Program Definition verified_lambdabox_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) diff --git a/erasure/_CoqProject.in b/erasure/_CoqProject.in index 6fd7c9cfc..4d0ec4f9b 100644 --- a/erasure/_CoqProject.in +++ b/erasure/_CoqProject.in @@ -39,6 +39,7 @@ theories/EConstructorsAsBlocks.v theories/EWcbvEvalCstrsAsBlocksInd.v theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v theories/ECoInductiveToInductive.v +theories/EUnboxing.v theories/EReorderCstrs.v theories/EImplementBox.v diff --git a/erasure/theories/Unboxing.v b/erasure/theories/EUnboxing.v similarity index 65% rename from erasure/theories/Unboxing.v rename to erasure/theories/EUnboxing.v index f4695489a..4bd0d9a73 100644 --- a/erasure/theories/Unboxing.v +++ b/erasure/theories/EUnboxing.v @@ -29,12 +29,13 @@ Equations safe_hd {A} (l : list A) (nnil : l <> nil) : A := | [] | nnil := False_rect _ (nnil eq_refl) | hd :: _ | _nnil := hd. -Section optimize. +Definition inspect {A : Type} (x : A) : {y : A | x = y} := + @exist _ (fun y => x = y) x eq_refl. + +Section unbox. Context (Σ : GlobalContextMap.t). Section Def. - Import TermSpineView. - Definition unboxable (idecl : one_inductive_body) (cdecl : constructor_body) := (#|idecl.(ind_ctors)| == 1) && (cdecl.(cstr_nargs) == 1). @@ -47,7 +48,6 @@ Section optimize. Notation " t 'eqn:' h " := (exist t h) (only parsing, at level 12). - Import Equations.Prop.Logic. Opaque is_unboxable. @@ -63,82 +63,49 @@ Section optimize. now intros [= <- <-]. Qed. - Equations? optimize (t : term) : term - by wf t (fun x y : EAst.term => size x < size y) := - | e with TermSpineView.view e := { + Equations unbox (t : term) : term := | tRel i => EAst.tRel i - | tEvar ev args => EAst.tEvar ev (map_InP args (fun x H => optimize x)) - | tLambda na M => EAst.tLambda na (optimize M) - | tApp u v napp nnil with construct_viewc u := { - | view_construct kn c block_args with is_unboxable kn c := - { | true => optimize (safe_hd v nnil) - | false => mkApps (EAst.tConstruct kn c block_args) (map_InP v (fun x H => optimize x)) } - | view_other u nconstr => - mkApps (optimize u) (map_InP v (fun x H => optimize x)) - } - | tLetIn na b b' => EAst.tLetIn na (optimize b) (optimize b') + | tEvar ev args => EAst.tEvar ev (map unbox args) + | tLambda na M => EAst.tLambda na (unbox M) + | tApp u v => EAst.tApp (unbox u) (unbox v) + | tLetIn na b b' => EAst.tLetIn na (unbox b) (unbox b') | tCase ind c brs with inspect (is_unboxable ind.1 0) => - { | true eqn:unb with inspect (get_unboxable_case_branch ind.1 brs ) := { - | Some (brna, bbody) eqn:heqbr => EAst.tLetIn brna (optimize c) (optimize bbody) - | None eqn:heqbr := - let brs' := map_InP brs (fun x H => (x.1, optimize x.2)) in - EAst.tCase ind (optimize c) brs' } - | false eqn:unb := - let brs' := map_InP brs (fun x H => (x.1, optimize x.2)) in - EAst.tCase ind (optimize c) brs' } + { | true eqn:unb with inspect (get_unboxable_case_branch ind.1 (map (on_snd unbox) brs)) := { + | Some (brna, bbody) eqn:heqbr => EAst.tLetIn brna (unbox c) bbody + | None eqn:heqbr := EAst.tCase ind (unbox c) (map (on_snd unbox) brs) } + | false eqn:unb := EAst.tCase ind (unbox c) (map (on_snd unbox) brs) } | tProj p c with inspect (is_unboxable p.(proj_ind) 0) := { - | true eqn:unb => optimize c - | false eqn:unb => EAst.tProj p (optimize c) } - | tConstruct ind n args => EAst.tConstruct ind n args + | true eqn:unb => unbox c + | false eqn:unb => EAst.tProj p (unbox c) } + | tConstruct ind n args with inspect (is_unboxable ind n) => { + unbox (tConstruct ind n [a]) (true eqn:unb) => unbox a ; + unbox (tConstruct ind n args) isunb => tConstruct ind n (map unbox args) } | tFix mfix idx => - let mfix' := map_InP mfix (fun d H => {| dname := dname d; dbody := optimize d.(dbody); rarg := d.(rarg) |}) in + let mfix' := map (map_def unbox) mfix in EAst.tFix mfix' idx | tCoFix mfix idx => - let mfix' := map_InP mfix (fun d H => {| dname := dname d; dbody := optimize d.(dbody); rarg := d.(rarg) |}) in + let mfix' := map (map_def unbox) mfix in EAst.tCoFix mfix' idx | tBox => EAst.tBox | tVar n => EAst.tVar n | tConst n => EAst.tConst n - | tPrim p => EAst.tPrim (map_primIn p (fun x H => optimize x)) }. - Proof. - all:try lia. - all:try apply (In_size); tea. - - eapply (In_size id size) in H. unfold id in H. - change (fun x => size x) with size in *. lia. - - rewrite size_mkApps. destruct v; cbn => //. lia. - - rewrite size_mkApps. - eapply (In_size id size) in H. change (fun x => size (id x)) with size in H. unfold id in H. cbn. lia. - - rewrite size_mkApps. destruct v => //=. lia. - - rewrite size_mkApps. - eapply (In_size id size) in H. change (fun x => size (id x)) with size in H. unfold id in H. cbn. lia. - - move/get_unboxable_case_branch_spec: heqbr => -> //=. lia. - - eapply (In_size snd size) in H; cbn in H; lia. - - eapply (In_size snd size) in H; cbn in H; lia. - - eapply (In_size dbody size) in H; cbn in H; lia. - - eapply (In_size dbody size) in H; cbn in H; lia. - - destruct p as [? []]; cbn in H; eauto. - intuition auto; subst; cbn; try lia. - eapply (In_size id size) in H0. unfold id in H0. - change (fun x => size x) with size in H0. lia. - Qed. - End Def. + | tPrim p => EAst.tPrim (map_prim unbox p) + | tLazy t => EAst.tLazy (unbox t) + | tForce t => EAst.tForce (unbox t). - #[universes(polymorphic)] Hint Rewrite @map_primIn_spec @map_InP_spec : optimize. + End Def. Lemma map_repeat {A B} (f : A -> B) x n : map f (repeat x n) = repeat (f x) n. Proof using Type. now induction n; simpl; auto; rewrite IHn. Qed. - Lemma map_optimize_repeat_box n : map optimize (repeat tBox n) = repeat tBox n. + Lemma map_unbox_repeat_box n : map unbox (repeat tBox n) = repeat tBox n. Proof using Type. now rewrite map_repeat. Qed. Arguments eqb : simpl never. - Opaque optimize_unfold_clause_1. - Opaque optimize. - Opaque isEtaExp. - Opaque isEtaExp_unfold_clause_1. + (* Lemma closedn_mkApps k f l : closedn k (mkApps f l) = closedn k f && forallb (closedn k) l. Proof using Type. @@ -147,9 +114,9 @@ Section optimize. - now rewrite IHl /= andb_assoc. Qed. - Lemma closed_optimize t k : closedn k t -> closedn k (optimize t). + Lemma closed_unbox t k : closedn k t -> closedn k (unbox t). Proof using Type Σ. - funelim (optimize t); simp optimize; rewrite -?optimize_equation_1; toAll; simpl; + revert k. induction t using EInduction.term_forall_list_ind; simp unbox; rewrite -?unbox_equation_1; toAll; simpl; intros; try easy; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; unfold test_def in *; @@ -170,44 +137,45 @@ Section optimize. Hint Rewrite @forallb_InP_spec : isEtaExp. Transparent isEtaExp_unfold_clause_1. - Local Lemma optimize_mkApps_nonnil f v (nnil : v <> []): + Local Lemma unbox_mkApps_nonnil f v (nnil : v <> []): ~~ isApp f -> - optimize (mkApps f v) = match construct_viewc f with + unbox (mkApps f v) = match construct_viewc f with | view_construct kn c block_args => - if is_unboxable kn c then optimize (safe_hd v nnil) - else mkApps (EAst.tConstruct kn c block_args) (map optimize v) - | view_other u nconstr => mkApps (optimize f) (map optimize v) + if is_unboxable kn c then unbox (safe_hd v nnil) + else mkApps (EAst.tConstruct kn c block_args) (map unbox v) + | view_other u nconstr => mkApps (unbox f) (map unbox v) end. Proof using Type. - intros napp. rewrite optimize_equation_1. + Admitted. + intros napp. rewrite unbox_equation_1. destruct (TermSpineView.view_mkApps (TermSpineView.view (mkApps f v)) napp nnil) as [hna [hv' ->]]. - simp optimize; rewrite -optimize_equation_1. + simp unbox; rewrite -unbox_equation_1. destruct (construct_viewc f). - 2:cbn; simp optimize => //. - simp optimize. - destruct (is_unboxable ind n) => //=. simp optimize. + 2:cbn; simp unbox => //. + simp unbox. + destruct (is_unboxable ind n) => //=. simp unbox. replace (safe_hd v hv') with (safe_hd v nnil) => //. destruct v; cbn => //. - f_equal. now simp optimize. + f_equal. now simp unbox. Qed. - Lemma optimize_mkApps f v : ~~ isApp f -> - optimize (mkApps f v) = match construct_viewc f with + Lemma unbox_mkApps f v : ~~ isApp f -> + unbox (mkApps f v) = match construct_viewc f with | view_construct kn c block_args => if is_unboxable kn c then match v with - | hd :: _ => optimize hd - | _ => mkApps (EAst.tConstruct kn c block_args) (map optimize v) + | hd :: _ => unbox hd + | _ => mkApps (EAst.tConstruct kn c block_args) (map unbox v) end - else mkApps (EAst.tConstruct kn c block_args) (map optimize v) - | view_other u nconstr => mkApps (optimize f) (map optimize v) + else mkApps (EAst.tConstruct kn c block_args) (map unbox v) + | view_other u nconstr => mkApps (unbox f) (map unbox v) end. Proof using Type. intros napp. destruct v using rev_case; simpl. - - destruct construct_viewc => //. simp optimize. + - destruct construct_viewc => //. simp unbox. destruct is_unboxable => //. - - rewrite (optimize_mkApps_nonnil f (v ++ [x])) => //. + - rewrite (unbox_mkApps_nonnil f (v ++ [x])) => //. destruct construct_viewc => //. destruct is_unboxable eqn:unb => //. destruct v eqn:hc => //=. @@ -222,14 +190,14 @@ Section optimize. cbn. do 2 destruct nth_error => //. congruence. Qed. - Lemma optimize_csubst a k b : + Lemma unbox_csubst a k b : closed a -> isEtaExp Σ a -> isEtaExp Σ b -> - optimize (ECSubst.csubst a k b) = ECSubst.csubst (optimize a) k (optimize b). + unbox (ECSubst.csubst a k b) = ECSubst.csubst (unbox a) k (unbox b). Proof using Type. intros cla etaa; move cla before a. move etaa before a. - funelim (optimize b); cbn; simp optimize isEtaExp; rewrite -?isEtaExp_equation_1 -?optimize_equation_1; toAll; simpl; + funelim (unbox b); cbn; simp unbox isEtaExp; rewrite -?isEtaExp_equation_1 -?unbox_equation_1; toAll; simpl; intros; try easy; rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; unfold test_def in *; @@ -242,7 +210,7 @@ Section optimize. rewrite !csubst_mkApps in H1 *. rewrite isEtaExp_mkApps_napp // in H1. destruct construct_viewc. - * cbn. rewrite optimize_mkApps //. + * cbn. rewrite unbox_mkApps //. * move/andP: H1 => [] et ev. rewrite -H //. assert (map (csubst a k) v <> []). @@ -260,7 +228,7 @@ Section optimize. destruct (construct_viewc t0) eqn:hc. { rewrite -mkApps_app /=. - rewrite optimize_mkApps // optimize_mkApps //. + rewrite unbox_mkApps // unbox_mkApps //. cbn -[lookup_inductive_pars]. move/andP: H2 => [] ise hl. unfold isEtaExp_app in ise. @@ -274,12 +242,12 @@ Section optimize. rewrite !map_map_compose. clear -etaa cla ev H0. solve_all. } { rewrite -mkApps_app. - rewrite optimize_mkApps //. rewrite hc. - rewrite optimize_mkApps // hc -mkApps_app map_app //. + rewrite unbox_mkApps //. rewrite hc. + rewrite unbox_mkApps // hc -mkApps_app map_app //. f_equal. f_equal. rewrite !map_map_compose. clear -etaa cla ev H0. solve_all. } } - { rewrite optimize_mkApps ?eqa //. + { rewrite unbox_mkApps ?eqa //. destruct (construct_viewc (csubst a k t)) eqn:eqc. 2:{ f_equal. rewrite !map_map_compose. clear -etaa cla ev H0. solve_all. } simp isEtaExp in H2. @@ -288,13 +256,13 @@ Section optimize. rewrite (lookup_inductive_pars_constructor_pars_args eqpars). assert (pars = 0). rtoProp. eapply Nat.leb_le in H2. lia. subst pars. rewrite skipn_0. - simp optimize; rewrite -optimize_equation_1. + simp unbox; rewrite -unbox_equation_1. { f_equal. rewrite !map_map_compose. clear -etaa cla ev H0. solve_all. } } - pose proof (etaExp_csubst _ _ k _ etaa H0). rewrite !csubst_mkApps /= in H1 *. assert (map (csubst a k) v <> []). { destruct v; cbn; congruence. } - rewrite optimize_mkApps //. + rewrite unbox_mkApps //. rewrite isEtaExp_Constructor // in H1. move/andP: H1. rewrite map_length. move=> [] etaapp etav. cbn -[lookup_inductive_pars]. @@ -310,7 +278,7 @@ Section optimize. rewrite !csubst_mkApps /= in H1 *. assert (map (csubst a k) v <> []). { destruct v; cbn; congruence. } - rewrite optimize_mkApps //. + rewrite unbox_mkApps //. rewrite isEtaExp_Constructor // in H1. rewrite GlobalContextMap.lookup_inductive_pars_spec in Heq. move/andP: H1. rewrite map_length. move=> [] etaapp etav. @@ -320,189 +288,191 @@ Section optimize. now rewrite (lookup_inductive_pars_constructor_pars_args eqpars) in Heq. Qed. - Lemma optimize_substl s t : + Lemma unbox_substl s t : forallb (closedn 0) s -> forallb (isEtaExp Σ) s -> isEtaExp Σ t -> - optimize (substl s t) = substl (map optimize s) (optimize t). + unbox (substl s t) = substl (map unbox s) (unbox t). Proof using Type. induction s in t |- *; simpl; auto. move=> /andP[] cla cls /andP[] etaa etas etat. rewrite IHs //. now eapply etaExp_csubst. f_equal. - now rewrite optimize_csubst. + now rewrite unbox_csubst. Qed. - Lemma optimize_iota_red pars args br : + Lemma unbox_iota_red pars args br : forallb (closedn 0) args -> forallb (isEtaExp Σ) args -> isEtaExp Σ br.2 -> - optimize (EGlobalEnv.iota_red pars args br) = EGlobalEnv.iota_red pars (map optimize args) (on_snd optimize br). + unbox (EGlobalEnv.iota_red pars args br) = EGlobalEnv.iota_red pars (map unbox args) (on_snd unbox br). Proof using Type. intros cl etaargs etabr. unfold EGlobalEnv.iota_red. - rewrite optimize_substl //. + rewrite unbox_substl //. rewrite forallb_rev forallb_skipn //. 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 unbox_fix_subst mfix : EGlobalEnv.fix_subst (map (map_def unbox) mfix) = map unbox (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 optimize. + f_equal; auto. now simp unbox. Qed. - Lemma optimize_cofix_subst mfix : EGlobalEnv.cofix_subst (map (map_def optimize) mfix) = map optimize (EGlobalEnv.cofix_subst mfix). + Lemma unbox_cofix_subst mfix : EGlobalEnv.cofix_subst (map (map_def unbox) mfix) = map unbox (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 optimize. + f_equal; auto. now simp unbox. Qed. - Lemma optimize_cunfold_fix mfix idx n f : + Lemma unbox_cunfold_fix mfix idx n f : forallb (closedn 0) (fix_subst mfix) -> forallb (fun d => isLambda (dbody d) && isEtaExp Σ (dbody d)) 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 unbox) mfix) idx = Some (n, unbox f). Proof using Type. intros hfix heta. unfold cunfold_fix. rewrite nth_error_map. destruct nth_error eqn:heq. intros [= <- <-] => /=. f_equal. f_equal. - rewrite optimize_substl //. + rewrite unbox_substl //. now apply isEtaExp_fix_subst. solve_all. eapply nth_error_all in heta; tea. cbn in heta. rtoProp; intuition auto. - f_equal. f_equal. apply optimize_fix_subst. + f_equal. f_equal. apply unbox_fix_subst. discriminate. Qed. - Lemma optimize_cunfold_cofix mfix idx n f : + Lemma unbox_cunfold_cofix mfix idx n f : forallb (closedn 0) (cofix_subst mfix) -> forallb (isEtaExp Σ ∘ dbody) 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 unbox) mfix) idx = Some (n, unbox f). Proof using Type. intros hcofix heta. unfold cunfold_cofix. rewrite nth_error_map. destruct nth_error eqn:heq. intros [= <- <-] => /=. f_equal. - rewrite optimize_substl //. + rewrite unbox_substl //. now apply isEtaExp_cofix_subst. solve_all. now eapply nth_error_all in heta; tea. - f_equal. f_equal. apply optimize_cofix_subst. + f_equal. f_equal. apply unbox_cofix_subst. discriminate. Qed. - Lemma optimize_nth {n l d} : - optimize (nth n l d) = nth n (map optimize l) (optimize d). + Lemma unbox_nth {n l d} : + unbox (nth n l d) = nth n (map unbox l) (unbox d). Proof using Type. induction l in n |- *; destruct n; simpl; auto. Qed. -End optimize. + *) + +End unbox. -#[universes(polymorphic)] Global Hint Rewrite @map_primIn_spec @map_InP_spec : optimize. +#[universes(polymorphic)] Global Hint Rewrite @map_primIn_spec @map_InP_spec : unbox. Tactic Notation "simp_eta" "in" hyp(H) := simp isEtaExp in H; rewrite -?isEtaExp_equation_1 in H. Ltac simp_eta := simp isEtaExp; rewrite -?isEtaExp_equation_1. -Tactic Notation "simp_optimize" "in" hyp(H) := simp optimize in H; rewrite -?optimize_equation_1 in H. -Ltac simp_optimize := simp optimize; rewrite -?optimize_equation_1. +Tactic Notation "simp_unbox" "in" hyp(H) := simp unbox in H; rewrite -?unbox_equation_1 in H. +Ltac simp_unbox := simp unbox; rewrite -?unbox_equation_1. -Definition optimize_constant_decl Σ cb := - {| cst_body := option_map (optimize Σ) cb.(cst_body) |}. +Definition unbox_constant_decl Σ cb := + {| cst_body := option_map (unbox Σ) cb.(cst_body) |}. -Definition optimize_inductive_decl idecl := +Definition unbox_inductive_decl idecl := {| ind_finite := idecl.(ind_finite); ind_npars := 0; ind_bodies := idecl.(ind_bodies) |}. -Definition optimize_decl Σ d := +Definition unbox_decl Σ d := match d with - | ConstantDecl cb => ConstantDecl (optimize_constant_decl Σ cb) - | InductiveDecl idecl => InductiveDecl (optimize_inductive_decl idecl) + | ConstantDecl cb => ConstantDecl (unbox_constant_decl Σ cb) + | InductiveDecl idecl => InductiveDecl (unbox_inductive_decl idecl) end. -Definition optimize_env Σ := - map (on_snd (optimize_decl Σ)) Σ.(GlobalContextMap.global_decls). +Definition unbox_env Σ := + map (on_snd (unbox_decl Σ)) Σ.(GlobalContextMap.global_decls). -Definition optimize_program (p : eprogram_env) : eprogram := - (ERemoveParams.optimize_env p.1, ERemoveParams.optimize p.1 p.2). +Definition unbox_program (p : eprogram_env) : eprogram := + (unbox_env p.1, unbox p.1 p.2). Import EGlobalEnv. -Lemma lookup_env_optimize Σ kn : - lookup_env (optimize_env Σ) kn = - option_map (optimize_decl Σ) (lookup_env Σ.(GlobalContextMap.global_decls) kn). +Lemma lookup_env_unbox Σ kn : + lookup_env (unbox_env Σ) kn = + option_map (unbox_decl Σ) (lookup_env Σ.(GlobalContextMap.global_decls) kn). Proof. - unfold optimize_env. + unfold unbox_env. destruct Σ as [Σ map repr wf]; cbn. induction Σ at 2 4; simpl; auto. case: eqb_spec => //. Qed. -Lemma lookup_constructor_optimize {Σ kn c} : - lookup_constructor (optimize_env Σ) kn c = +Lemma lookup_constructor_unbox {Σ kn c} : + lookup_constructor (unbox_env Σ) kn c = match lookup_constructor Σ.(GlobalContextMap.global_decls) kn c with - | Some (mdecl, idecl, cdecl) => Some (optimize_inductive_decl mdecl, idecl, cdecl) + | Some (mdecl, idecl, cdecl) => Some (unbox_inductive_decl mdecl, idecl, cdecl) | None => None end. Proof. rewrite /lookup_constructor /lookup_inductive /lookup_minductive. - rewrite lookup_env_optimize. + rewrite lookup_env_unbox. destruct lookup_env as [[]|] => /= //. do 2 destruct nth_error => //. Qed. -Lemma lookup_constructor_pars_args_optimize (Σ : GlobalContextMap.t) i n npars nargs : +Lemma lookup_constructor_pars_args_unbox (Σ : GlobalContextMap.t) i n npars nargs : EGlobalEnv.lookup_constructor_pars_args Σ i n = Some (npars, nargs) -> - EGlobalEnv.lookup_constructor_pars_args (optimize_env Σ) i n = Some (0, nargs). + EGlobalEnv.lookup_constructor_pars_args (unbox_env Σ) i n = Some (0, nargs). Proof. - rewrite /EGlobalEnv.lookup_constructor_pars_args. rewrite lookup_constructor_optimize //=. + rewrite /EGlobalEnv.lookup_constructor_pars_args. rewrite lookup_constructor_unbox //=. destruct EGlobalEnv.lookup_constructor => //. destruct p as [[] ?] => //=. now intros [= <- <-]. Qed. -Lemma is_propositional_optimize (Σ : GlobalContextMap.t) ind : +Lemma is_propositional_unbox (Σ : GlobalContextMap.t) ind : match inductive_isprop_and_pars Σ.(GlobalContextMap.global_decls) ind with | Some (prop, npars) => - inductive_isprop_and_pars (optimize_env Σ) ind = Some (prop, 0) + inductive_isprop_and_pars (unbox_env Σ) ind = Some (prop, 0) | None => - inductive_isprop_and_pars (optimize_env Σ) ind = None + inductive_isprop_and_pars (unbox_env Σ) ind = None end. Proof. rewrite /inductive_isprop_and_pars /lookup_inductive /lookup_minductive. - rewrite lookup_env_optimize. + rewrite lookup_env_unbox. destruct lookup_env; simpl; auto. destruct g; simpl; auto. destruct nth_error => //. Qed. -Lemma is_propositional_cstr_optimize {Σ : GlobalContextMap.t} {ind c} : +Lemma is_propositional_cstr_unbox {Σ : GlobalContextMap.t} {ind c} : match constructor_isprop_pars_decl Σ.(GlobalContextMap.global_decls) ind c with | Some (prop, npars, cdecl) => - constructor_isprop_pars_decl (optimize_env Σ) ind c = Some (prop, 0, cdecl) + constructor_isprop_pars_decl (unbox_env Σ) ind c = Some (prop, 0, cdecl) | None => - constructor_isprop_pars_decl (optimize_env Σ) ind c = None + constructor_isprop_pars_decl (unbox_env Σ) ind c = None end. Proof. rewrite /constructor_isprop_pars_decl /lookup_constructor /lookup_inductive /lookup_minductive. - rewrite lookup_env_optimize. + rewrite lookup_env_unbox. destruct lookup_env; simpl; auto. destruct g; simpl; auto. do 2 destruct nth_error => //. Qed. -Lemma lookup_inductive_pars_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : +Lemma lookup_inductive_pars_unbox {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : wf_glob Σ -> forall pars, lookup_inductive_pars Σ ind = Some pars -> - EGlobalEnv.lookup_inductive_pars (optimize_env Σ) ind = Some 0. + EGlobalEnv.lookup_inductive_pars (unbox_env Σ) ind = Some 0. Proof. rewrite /lookup_inductive_pars => wf pars. rewrite /lookup_inductive /lookup_minductive. - rewrite (lookup_env_optimize _ ind). + rewrite (lookup_env_unbox _ ind). destruct lookup_env as [[decl|]|] => //=. Qed. @@ -551,9 +521,10 @@ Proof. now rewrite remove_last_app. Qed. +(* Arguments lookup_inductive_pars_constructor_pars_args {Σ ind n pars args}. -Lemma optimize_tApp {Σ : GlobalContextMap.t} f a : isEtaExp Σ f -> isEtaExp Σ a -> optimize Σ (EAst.tApp f a) = EAst.tApp (optimize Σ f) (optimize Σ a). +Lemma unbox_tApp {Σ : GlobalContextMap.t} f a : isEtaExp Σ f -> isEtaExp Σ a -> unbox Σ (EAst.tApp f a) = EAst.tApp (unbox Σ f) (unbox Σ a). Proof. move=> etaf etaa. pose proof (isEtaExp_mkApps_intro _ _ [a] etaf). @@ -564,11 +535,11 @@ Proof. { intros [? [? []]]. rewrite H0 /=. rewrite -[EAst.tApp _ _ ](mkApps_app _ _ [a]). move/andP: H2 => []. rewrite /isEtaExp_app. - rewrite !optimize_mkApps // cv. + rewrite !unbox_mkApps // cv. destruct lookup_constructor_pars_args as [[pars args]|] eqn:hl => //. rewrite (lookup_inductive_pars_constructor_pars_args hl). intros hpars etal. - rewrite -[EAst.tApp _ _ ](mkApps_app _ _ [optimize Σ a]). + rewrite -[EAst.tApp _ _ ](mkApps_app _ _ [unbox Σ a]). f_equal. rewrite !skipn_map !skipn_app map_app. f_equal. assert (pars - #|remove_last l| = 0) as ->. 2:{ rewrite skipn_0 //. } @@ -579,12 +550,12 @@ Proof. move => /andP[] /Nat.leb_le. lia. } { move/and4P=> [] iset isel _ _. rewrite (decompose_app_inv da). pose proof (decompose_app_notApp _ _ _ da). - rewrite optimize_mkApps //. + rewrite unbox_mkApps //. destruct (decompose_app_tApp_split _ _ _ _ da). rewrite cv. rewrite H0. - rewrite optimize_mkApps // cv. - rewrite -[EAst.tApp _ _ ](mkApps_app _ _ [optimize Σ a]). f_equal. - rewrite -[(_ ++ _)%list](map_app (optimize Σ) _ [a]). + rewrite unbox_mkApps // cv. + rewrite -[EAst.tApp _ _ ](mkApps_app _ _ [unbox Σ a]). f_equal. + rewrite -[(_ ++ _)%list](map_app (unbox Σ) _ [a]). f_equal. assert (l <> []). { destruct l; try congruence. eapply decompose_app_inv in da. cbn in *. now subst t. } @@ -593,23 +564,23 @@ Proof. Qed. Module Fast. - Section fastoptimize. + Section fastunbox. Context (Σ : GlobalContextMap.t). - Equations optimize (app : list term) (t : term) : term := { - | app, tEvar ev args => mkApps (EAst.tEvar ev (optimize_args args)) app - | app, tLambda na M => mkApps (EAst.tLambda na (optimize [] M)) app - | app, tApp u v => optimize (optimize [] v :: app) u - | app, tLetIn na b b' => mkApps (EAst.tLetIn na (optimize [] b) (optimize [] b')) app + Equations unbox (app : list term) (t : term) : term := { + | app, tEvar ev args => mkApps (EAst.tEvar ev (unbox_args args)) app + | app, tLambda na M => mkApps (EAst.tLambda na (unbox [] M)) app + | app, tApp u v => unbox (unbox [] v :: app) u + | app, tLetIn na b b' => mkApps (EAst.tLetIn na (unbox [] b) (unbox [] b')) app | app, tCase ind c brs => - let brs' := optimize_brs brs in - mkApps (EAst.tCase (ind.1, 0) (optimize [] c) brs') app - | app, tProj p c => mkApps (EAst.tProj {| proj_ind := p.(proj_ind); proj_npars := 0; proj_arg := p.(proj_arg) |} (optimize [] c)) app + let brs' := unbox_brs brs in + mkApps (EAst.tCase (ind.1, 0) (unbox [] c) brs') app + | app, tProj p c => mkApps (EAst.tProj {| proj_ind := p.(proj_ind); proj_npars := 0; proj_arg := p.(proj_arg) |} (unbox [] c)) app | app, tFix mfix idx => - let mfix' := optimize_defs mfix in + let mfix' := unbox_defs mfix in mkApps (EAst.tFix mfix' idx) app | app, tCoFix mfix idx => - let mfix' := optimize_defs mfix in + let mfix' := unbox_defs mfix in mkApps (EAst.tCoFix mfix' idx) app | app, tConstruct kn c block_args with GlobalContextMap.lookup_inductive_pars Σ (inductive_mind kn) := { | Some npars => mkApps (EAst.tConstruct kn c block_args) (List.skipn npars app) @@ -617,21 +588,21 @@ Module Fast. | app, tPrim (primInt; primIntModel i) => mkApps (tPrim (primInt; primIntModel i)) app | app, tPrim (primFloat; primFloatModel f) => mkApps (tPrim (primFloat; primFloatModel f)) app | app, tPrim (primArray; primArrayModel a) => - mkApps (tPrim (primArray; primArrayModel {| array_default := optimize [] a.(array_default); array_value := optimize_args a.(array_value) |})) app + mkApps (tPrim (primArray; primArrayModel {| array_default := unbox [] a.(array_default); array_value := unbox_args a.(array_value) |})) app | app, x => mkApps x app } - where optimize_args (t : list term) : list term := + where unbox_args (t : list term) : list term := { | [] := [] - | a :: args := (optimize [] a) :: optimize_args args + | a :: args := (unbox [] a) :: unbox_args args } - where optimize_brs (t : list (list BasicAst.name × term)) : list (list BasicAst.name × term) := + where unbox_brs (t : list (list BasicAst.name × term)) : list (list BasicAst.name × term) := { | [] := [] - | a :: args := (a.1, (optimize [] a.2)) :: optimize_brs args } + | a :: args := (a.1, (unbox [] a.2)) :: unbox_brs args } - where optimize_defs (t : mfixpoint term) : mfixpoint term := + where unbox_defs (t : mfixpoint term) : mfixpoint term := { | [] := [] - | d :: defs := {| dname := dname d; dbody := optimize [] d.(dbody); rarg := d.(rarg) |} :: optimize_defs defs + | d :: defs := {| dname := dname d; dbody := unbox [] d.(dbody); rarg := d.(rarg) |} :: unbox_defs defs }. Local Ltac specIH := @@ -639,74 +610,74 @@ Module Fast. | [ H : (forall args : list term, _) |- _ ] => specialize (H [] eq_refl) end. - Lemma optimize_acc_opt t : - forall args, ERemoveParams.optimize Σ (mkApps t args) = optimize (map (ERemoveParams.optimize Σ) args) t. + Lemma unbox_acc_opt t : + forall args, ERemoveParams.unbox Σ (mkApps t args) = unbox (map (ERemoveParams.unbox Σ) args) t. Proof using Type. intros args. - remember (map (ERemoveParams.optimize Σ) args) as args'. + remember (map (ERemoveParams.unbox Σ) args) as args'. revert args Heqargs'. - set (P:= fun args' t fs => forall args, args' = map (ERemoveParams.optimize Σ) args -> ERemoveParams.optimize Σ (mkApps t args) = fs). - apply (optimize_elim P - (fun l l' => map (ERemoveParams.optimize Σ) l = l') - (fun l l' => map (on_snd (ERemoveParams.optimize Σ)) l = l') - (fun l l' => map (map_def (ERemoveParams.optimize Σ)) l = l')); subst P; cbn -[GlobalContextMap.lookup_inductive_pars isEtaExp ERemoveParams.optimize]; clear. + set (P:= fun args' t fs => forall args, args' = map (ERemoveParams.unbox Σ) args -> ERemoveParams.unbox Σ (mkApps t args) = fs). + apply (unbox_elim P + (fun l l' => map (ERemoveParams.unbox Σ) l = l') + (fun l l' => map (on_snd (ERemoveParams.unbox Σ)) l = l') + (fun l l' => map (map_def (ERemoveParams.unbox Σ)) l = l')); subst P; cbn -[GlobalContextMap.lookup_inductive_pars isEtaExp ERemoveParams.unbox]; clear. all: try reflexivity. - all: intros *; simp_eta; simp_optimize. - all: try solve [intros; subst; rtoProp; rewrite optimize_mkApps // /=; simp_optimize; repeat specIH; repeat (f_equal; intuition eauto)]. - all: try solve [rewrite optimize_mkApps //]. + all: intros *; simp_eta; simp_unbox. + all: try solve [intros; subst; rtoProp; rewrite unbox_mkApps // /=; simp_unbox; repeat specIH; repeat (f_equal; intuition eauto)]. + all: try solve [rewrite unbox_mkApps //]. - intros IHv IHu. specialize (IHv [] eq_refl). simpl in IHv. intros args ->. specialize (IHu (v :: args)). forward IHu. now rewrite -IHv. exact IHu. - intros Hl hargs args ->. - rewrite optimize_mkApps //=. simp_optimize. + rewrite unbox_mkApps //=. simp_unbox. f_equal. f_equal. cbn. do 2 f_equal. rewrite /map_array_model. specialize (Hl [] eq_refl). f_equal; eauto. - intros Hl args ->. - rewrite optimize_mkApps // /=. + rewrite unbox_mkApps // /=. rewrite GlobalContextMap.lookup_inductive_pars_spec in Hl. now rewrite Hl. - intros Hl args ->. rewrite GlobalContextMap.lookup_inductive_pars_spec in Hl. - now rewrite optimize_mkApps // /= Hl. + now rewrite unbox_mkApps // /= Hl. - intros IHa heq. specIH. now rewrite IHa. - intros IHa heq; specIH. f_equal; eauto. unfold on_snd. now rewrite IHa. - intros IHa heq; specIH. unfold map_def. f_equal; eauto. now rewrite IHa. Qed. - Lemma optimize_fast t : ERemoveParams.optimize Σ t = optimize [] t. - Proof using Type. now apply (optimize_acc_opt t []). Qed. + Lemma unbox_fast t : ERemoveParams.unbox Σ t = unbox [] t. + Proof using Type. now apply (unbox_acc_opt t []). Qed. - End fastoptimize. + End fastunbox. - Notation optimize' Σ := (optimize Σ []). + Notation unbox' Σ := (unbox Σ []). - Definition optimize_constant_decl Σ cb := - {| cst_body := option_map (optimize' Σ) cb.(cst_body) |}. + Definition unbox_constant_decl Σ cb := + {| cst_body := option_map (unbox' Σ) cb.(cst_body) |}. - Definition optimize_inductive_decl idecl := + Definition unbox_inductive_decl idecl := {| ind_finite := idecl.(ind_finite); ind_npars := 0; ind_bodies := idecl.(ind_bodies) |}. - Definition optimize_decl Σ d := + Definition unbox_decl Σ d := match d with - | ConstantDecl cb => ConstantDecl (optimize_constant_decl Σ cb) - | InductiveDecl idecl => InductiveDecl (optimize_inductive_decl idecl) + | ConstantDecl cb => ConstantDecl (unbox_constant_decl Σ cb) + | InductiveDecl idecl => InductiveDecl (unbox_inductive_decl idecl) end. - Definition optimize_env Σ := - map (on_snd (optimize_decl Σ)) Σ.(GlobalContextMap.global_decls). + Definition unbox_env Σ := + map (on_snd (unbox_decl Σ)) Σ.(GlobalContextMap.global_decls). - Lemma optimize_env_fast Σ : ERemoveParams.optimize_env Σ = optimize_env Σ. + Lemma unbox_env_fast Σ : ERemoveParams.unbox_env Σ = unbox_env Σ. Proof. - unfold ERemoveParams.optimize_env, optimize_env. + unfold ERemoveParams.unbox_env, unbox_env. destruct Σ as [Σ map repr wf]. cbn. induction Σ at 2 4; cbn; auto. f_equal; eauto. destruct a as [kn []]; cbn; auto. destruct c as [[b|]]; cbn; auto. unfold on_snd; cbn. - do 2 f_equal. unfold ERemoveParams.optimize_constant_decl, optimize_constant_decl. - simpl. f_equal. f_equal. apply optimize_fast. + do 2 f_equal. unfold ERemoveParams.unbox_constant_decl, unbox_constant_decl. + simpl. f_equal. f_equal. apply unbox_fast. Qed. End Fast. @@ -756,19 +727,19 @@ Proof. rewrite mkApps_app /= //. Qed. -Lemma optimize_isLambda Σ f : - EAst.isLambda f = EAst.isLambda (optimize Σ f). +Lemma unbox_isLambda Σ f : + EAst.isLambda f = EAst.isLambda (unbox Σ f). Proof. - funelim (optimize Σ f); cbn -[optimize]; (try simp_optimize) => //. + funelim (unbox Σ f); cbn -[unbox]; (try simp_unbox) => //. rewrite (negbTE (isLambda_mkApps' _ _ _)) //. rewrite (negbTE (isLambda_mkApps' _ _ _)) //; try apply map_nil => //. all:rewrite !(negbTE (isLambda_mkApps_Construct _ _ _ _)) //. Qed. -Lemma optimize_isBox Σ f : - isBox f = isBox (optimize Σ f). +Lemma unbox_isBox Σ f : + isBox f = isBox (unbox Σ f). Proof. - funelim (optimize Σ f); cbn -[optimize] => //. + funelim (unbox Σ f); cbn -[unbox] => //. all:rewrite map_InP_spec. rewrite (negbTE (isBox_mkApps' _ _ _)) //. rewrite (negbTE (isBox_mkApps' _ _ _)) //; try apply map_nil => //. @@ -781,45 +752,45 @@ Proof. rewrite mkApps_app /= //. Qed. -Lemma optimize_isApp Σ f : +Lemma unbox_isApp Σ f : ~~ EAst.isApp f -> - ~~ EAst.isApp (optimize Σ f). + ~~ EAst.isApp (unbox Σ f). Proof. - funelim (optimize Σ f); cbn -[optimize] => //. + funelim (unbox Σ f); cbn -[unbox] => //. all:rewrite map_InP_spec. all:rewrite isApp_mkApps //. Qed. -Lemma optimize_isFix Σ f : - isFix f = isFix (optimize Σ f). +Lemma unbox_isFix Σ f : + isFix f = isFix (unbox Σ f). Proof. - funelim (optimize Σ f); cbn -[optimize] => //. + funelim (unbox Σ f); cbn -[unbox] => //. all:rewrite map_InP_spec. rewrite (negbTE (isFix_mkApps' _ _ _)) //. rewrite (negbTE (isFix_mkApps' _ _ _)) //; try apply map_nil => //. all:rewrite !(negbTE (isFix_mkApps_Construct _ _ _ _)) //. Qed. -Lemma optimize_isFixApp Σ f : - isFixApp f = isFixApp (optimize Σ f). +Lemma unbox_isFixApp Σ f : + isFixApp f = isFixApp (unbox Σ f). Proof. - funelim (optimize Σ f); cbn -[optimize] => //. + funelim (unbox Σ f); cbn -[unbox] => //. all:rewrite map_InP_spec. all:rewrite isFixApp_mkApps isFixApp_mkApps //. Qed. -Lemma optimize_isConstructApp Σ f : - isConstructApp f = isConstructApp (optimize Σ f). +Lemma unbox_isConstructApp Σ f : + isConstructApp f = isConstructApp (unbox Σ f). Proof. - funelim (optimize Σ f); cbn -[optimize] => //. + funelim (unbox Σ f); cbn -[unbox] => //. all:rewrite map_InP_spec. all:rewrite isConstructApp_mkApps isConstructApp_mkApps //. Qed. -Lemma optimize_isPrimApp Σ f : - isPrimApp f = isPrimApp (optimize Σ f). +Lemma unbox_isPrimApp Σ f : + isPrimApp f = isPrimApp (unbox Σ f). Proof. - funelim (optimize Σ f); cbn -[optimize] => //. + funelim (unbox Σ f); cbn -[unbox] => //. all:rewrite map_InP_spec. all:rewrite !isPrimApp_mkApps //. Qed. @@ -863,9 +834,9 @@ Proof. destruct nth_error => //. destruct nth_error => //. congruence. Qed. -Lemma optimize_mkApps_etaexp {Σ : GlobalContextMap.t} fn args : +Lemma unbox_mkApps_etaexp {Σ : GlobalContextMap.t} fn args : isEtaExp Σ fn -> - optimize Σ (EAst.mkApps fn args) = EAst.mkApps (optimize Σ fn) (map (optimize Σ) args). + unbox Σ (EAst.mkApps fn args) = EAst.mkApps (unbox Σ fn) (map (unbox Σ) args). Proof. destruct (decompose_app fn) eqn:da. rewrite (decompose_app_inv da). @@ -873,8 +844,8 @@ Proof. destruct construct_viewc eqn:vc. + move=> /andP[] hl0 etal0. rewrite -mkApps_app. - rewrite (optimize_mkApps Σ (tConstruct ind n block_args)) // /=. - rewrite optimize_mkApps // /=. + rewrite (unbox_mkApps Σ (tConstruct ind n block_args)) // /=. + rewrite unbox_mkApps // /=. unfold isEtaExp_app in hl0. destruct lookup_constructor_pars_args as [[pars args']|] eqn:hl => //. rtoProp. @@ -884,7 +855,7 @@ Proof. rewrite skipn_app. len. assert (pars - #|l| = 0) by lia. now rewrite H1 skipn_0. + move=> /andP[] etat0 etal0. - rewrite -mkApps_app !optimize_mkApps; try now eapply decompose_app_notApp. + rewrite -mkApps_app !unbox_mkApps; try now eapply decompose_app_notApp. rewrite vc. rewrite -mkApps_app !map_app //. Qed. @@ -919,178 +890,177 @@ Proof. intros; rtoProp; intuition auto; solve_all. Qed. -Lemma optimize_eval (efl := all_env_flags) {wfl:WcbvFlags} {wcon : with_constructor_as_block = false} {Σ : GlobalContextMap.t} t v : +Lemma unbox_eval (efl := all_env_flags) {wfl:WcbvFlags} {wcon : with_constructor_as_block = false} {Σ : GlobalContextMap.t} t v : isEtaExp_env Σ -> closed_env Σ -> wf_glob Σ -> closedn 0 t -> isEtaExp Σ t -> eval Σ t v -> - eval (optimize_env Σ) (optimize Σ t) (optimize Σ v). + eval (unbox_env Σ) (unbox Σ t) (unbox Σ v). Proof. intros etaΣ clΣ wfΣ. revert t v. - unshelve eapply (eval_preserve_mkApps_ind wfl wcon Σ (fun x y => eval (optimize_env Σ) (optimize Σ x) (optimize Σ y)) + unshelve eapply (eval_preserve_mkApps_ind wfl wcon Σ (fun x y => eval (unbox_env Σ) (unbox Σ x) (unbox Σ y)) (fun n x => closedn n x) (Qpres := Qpreserves_closedn Σ clΣ)) => //. { intros. eapply eval_closed; tea. } all:intros; simpl in *. - (* 1-15: try solve [econstructor; eauto]. *) all:repeat destruct_nary_times => //. - - rewrite optimize_tApp //. + - rewrite unbox_tApp //. econstructor; eauto. - - rewrite optimize_tApp //. simp_optimize in e1. + - rewrite unbox_tApp //. simp_unbox in e1. econstructor; eauto. - rewrite optimize_csubst // in e. now simp_eta in i10. + rewrite unbox_csubst // in e. now simp_eta in i10. - - simp_optimize. - rewrite optimize_csubst // in e. + - simp_unbox. + rewrite unbox_csubst // in e. econstructor; eauto. - - simp_optimize. - set (brs' := map _ brs). cbn -[optimize]. + - simp_unbox. + set (brs' := map _ brs). cbn -[unbox]. cbn in H3. - rewrite optimize_mkApps // /= in e0. + rewrite unbox_mkApps // /= in e0. apply constructor_isprop_pars_decl_lookup in H2 as H1'. rewrite H1' in e0. - pose proof (@is_propositional_cstr_optimize Σ ind c). + pose proof (@is_propositional_cstr_unbox Σ ind c). rewrite H2 in H1. econstructor; eauto. * rewrite nth_error_map H3 //. * len. cbn in H4, H5. rewrite skipn_length. lia. - * cbn -[optimize]. rewrite skipn_0. len. - * cbn -[optimize]. + * cbn -[unbox]. rewrite skipn_0. len. + * cbn -[unbox]. have etaargs : forallb (isEtaExp Σ) args. { rewrite isEtaExp_Constructor in i6. now move/andP: i6 => [] /andP[]. } - rewrite optimize_iota_red // in e. + rewrite unbox_iota_red // in e. rewrite closedn_mkApps in i4. now move/andP: i4. cbn. now eapply nth_error_forallb in H; tea. - subst brs. cbn in H4. rewrite andb_true_r in H4. - rewrite optimize_substl // in e. + rewrite unbox_substl // in e. eapply All_forallb, All_repeat => //. eapply All_forallb, All_repeat => //. - rewrite map_optimize_repeat_box in e. - simp_optimize. set (brs' := map _ _). - cbn -[optimize]. eapply eval_iota_sing => //. - now move: (is_propositional_optimize Σ ind); rewrite H2. - - - rewrite optimize_tApp //. - rewrite optimize_mkApps // /= in e1. - simp_optimize in e1. + rewrite map_unbox_repeat_box in e. + simp_unbox. set (brs' := map _ _). + cbn -[unbox]. eapply eval_iota_sing => //. + now move: (is_propositional_unbox Σ ind); rewrite H2. + + - rewrite unbox_tApp //. + rewrite unbox_mkApps // /= in e1. + simp_unbox in e1. eapply eval_fix; tea. * rewrite map_length. - eapply optimize_cunfold_fix; tea. + eapply unbox_cunfold_fix; tea. eapply closed_fix_subst. tea. move: i8; rewrite closedn_mkApps => /andP[] //. move: i10; rewrite isEtaExp_mkApps_napp // /= => /andP[] //. simp_eta. * move: e. rewrite -[tApp _ _](mkApps_app _ _ [av]). - rewrite -[tApp _ _](mkApps_app _ _ [optimize _ av]). - rewrite optimize_mkApps_etaexp // map_app //. - - - rewrite optimize_tApp //. - rewrite optimize_tApp //. - rewrite optimize_mkApps //. simpl. - simp_optimize. set (mfix' := map _ _). simpl. - rewrite optimize_mkApps // /= in e0. simp_optimize in e0. + rewrite -[tApp _ _](mkApps_app _ _ [unbox _ av]). + rewrite unbox_mkApps_etaexp // map_app //. + + - rewrite unbox_tApp //. + rewrite unbox_tApp //. + rewrite unbox_mkApps //. simpl. + simp_unbox. set (mfix' := map _ _). simpl. + rewrite unbox_mkApps // /= in e0. simp_unbox in e0. eapply eval_fix_value; tea. - eapply optimize_cunfold_fix; eauto. + eapply unbox_cunfold_fix; eauto. eapply closed_fix_subst => //. { move: i4. rewrite closedn_mkApps. now move/andP => []. } { move: i6. rewrite isEtaExp_mkApps_napp // /= => /andP[] //. now simp isEtaExp. } now rewrite map_length. - - rewrite optimize_tApp //. simp_optimize in e0. - simp_optimize in e1. + - rewrite unbox_tApp //. simp_unbox in e0. + simp_unbox in e1. eapply eval_fix'; tea. - eapply optimize_cunfold_fix; tea. + eapply unbox_cunfold_fix; tea. { eapply closed_fix_subst => //. } { simp isEtaExp in i10. } - rewrite optimize_tApp // in e. + rewrite unbox_tApp // in e. - - simp_optimize. - rewrite optimize_mkApps // /= in e0. - simp_optimize in e. - simp_optimize in e0. + - simp_unbox. + rewrite unbox_mkApps // /= in e0. + simp_unbox in e. + simp_unbox in e0. set (brs' := map _ _) in *; simpl. eapply eval_cofix_case; tea. - eapply optimize_cunfold_cofix; tea => //. + eapply unbox_cunfold_cofix; tea => //. { eapply closed_cofix_subst; tea. move: i5; rewrite closedn_mkApps => /andP[] //. } { move: i7. rewrite isEtaExp_mkApps_napp // /= => /andP[] //. now simp isEtaExp. } - rewrite optimize_mkApps_etaexp // in e. + rewrite unbox_mkApps_etaexp // in e. - destruct p as [[ind pars] arg]. - simp_optimize. - simp_optimize in e. - rewrite optimize_mkApps // /= in e0. - simp_optimize in e0. - rewrite optimize_mkApps_etaexp // in e. - simp_optimize in e0. + simp_unbox. + simp_unbox in e. + rewrite unbox_mkApps // /= in e0. + simp_unbox in e0. + rewrite unbox_mkApps_etaexp // in e. + simp_unbox in e0. eapply eval_cofix_proj; tea. - eapply optimize_cunfold_cofix; tea. + eapply unbox_cunfold_cofix; tea. { eapply closed_cofix_subst; tea. move: i4; rewrite closedn_mkApps => /andP[] //. } { move: i6. rewrite isEtaExp_mkApps_napp // /= => /andP[] //. now simp isEtaExp. } - econstructor. red in H |- *. - rewrite lookup_env_optimize H //. - now rewrite /optimize_constant_decl H0. + rewrite lookup_env_unbox H //. + now rewrite /unbox_constant_decl H0. exact e. - - simp_optimize. - rewrite optimize_mkApps // /= in e0. + - simp_unbox. + rewrite unbox_mkApps // /= in e0. rewrite (constructor_isprop_pars_decl_lookup H2) in e0. eapply eval_proj; eauto. - move: (@is_propositional_cstr_optimize Σ p.(proj_ind) 0). now rewrite H2. simpl. + move: (@is_propositional_cstr_unbox Σ p.(proj_ind) 0). now rewrite H2. simpl. len. rewrite skipn_length. cbn in H3. lia. rewrite nth_error_skipn nth_error_map /= H4 //. - - simp_optimize. eapply eval_proj_prop => //. - move: (is_propositional_optimize Σ p.(proj_ind)); now rewrite H3. + - simp_unbox. eapply eval_proj_prop => //. + move: (is_propositional_unbox Σ p.(proj_ind)); now rewrite H3. - - rewrite !optimize_tApp //. + - rewrite !unbox_tApp //. eapply eval_app_cong; tea. move: H1. eapply contraNN. - rewrite -optimize_isLambda -optimize_isConstructApp -optimize_isFixApp -optimize_isBox -optimize_isPrimApp //. - rewrite -optimize_isFix //. + rewrite -unbox_isLambda -unbox_isConstructApp -unbox_isFixApp -unbox_isBox -unbox_isPrimApp //. + rewrite -unbox_isFix //. - - rewrite !optimize_mkApps // /=. + - rewrite !unbox_mkApps // /=. rewrite (lookup_constructor_lookup_inductive_pars H). eapply eval_mkApps_Construct; tea. - + rewrite lookup_constructor_optimize H //. - + constructor. cbn [atom]. rewrite wcon lookup_constructor_optimize H //. + + rewrite lookup_constructor_unbox H //. + + constructor. cbn [atom]. rewrite wcon lookup_constructor_unbox H //. + rewrite /cstr_arity /=. move: H0; rewrite /cstr_arity /=. rewrite skipn_length map_length => ->. lia. + cbn in H0. eapply All2_skipn, All2_map. - eapply All2_impl; tea; cbn -[optimize]. + eapply All2_impl; tea; cbn -[unbox]. intros x y []; auto. - - depelim X; simp optimize; repeat constructor. + - depelim X; simp unbox; repeat constructor. eapply All2_over_undep in a. eapply All2_Set_All2 in ev. eapply All2_All2_Set. solve_all. now destruct b. now destruct a0. - destruct t => //. - all:constructor; eauto. simp optimize. - cbn [atom optimize] in H |- *. - rewrite lookup_constructor_optimize. + all:constructor; eauto. simp unbox. + cbn [atom unbox] in H |- *. + rewrite lookup_constructor_unbox. destruct lookup_constructor eqn:hl => //. destruct p as [[] ?] => //. Qed. From MetaCoq.Erasure Require Import EEtaExpanded. -Lemma optimize_declared_constructor {Σ : GlobalContextMap.t} {k mdecl idecl cdecl} : +Lemma unbox_declared_constructor {Σ : GlobalContextMap.t} {k mdecl idecl cdecl} : declared_constructor Σ.(GlobalContextMap.global_decls) k mdecl idecl cdecl -> - declared_constructor (optimize_env Σ) k (optimize_inductive_decl mdecl) idecl cdecl. + declared_constructor (unbox_env Σ) k (unbox_inductive_decl mdecl) idecl cdecl. Proof. intros [[] ?]; do 2 split => //. red in H; red. - rewrite lookup_env_optimize H //. + rewrite lookup_env_unbox H //. Qed. Lemma lookup_inductive_pars_spec {Σ} {mind} {mdecl} : @@ -1109,80 +1079,80 @@ Definition switch_no_params (efl : EEnvFlags) := |}. (* Stripping preserves well-formedness directly, not caring about eta-expansion *) -Lemma optimize_wellformed {efl} {Σ : GlobalContextMap.t} n t : +Lemma unbox_wellformed {efl} {Σ : GlobalContextMap.t} n t : cstr_as_blocks = false -> has_tApp -> @wf_glob efl Σ -> @wellformed efl Σ n t -> - @wellformed (switch_no_params efl) (optimize_env Σ) n (optimize Σ t). + @wellformed (switch_no_params efl) (unbox_env Σ) n (unbox Σ t). Proof. intros cab hasapp wfΣ. revert n. - funelim (optimize Σ t); try intros n. - all:cbn -[optimize lookup_constructor lookup_inductive]; simp_optimize; intros. + funelim (unbox Σ t); try intros n. + all:cbn -[unbox lookup_constructor lookup_inductive]; simp_unbox; intros. all:try solve[unfold wf_fix_gen in *; rtoProp; intuition auto; toAll; solve_all]. - - cbn -[optimize]; simp_optimize. intros; rtoProp; intuition auto. - rewrite lookup_env_optimize. destruct lookup_env eqn:hl => // /=. + - cbn -[unbox]; simp_unbox. intros; rtoProp; intuition auto. + rewrite lookup_env_unbox. destruct lookup_env eqn:hl => // /=. destruct g => /= //. destruct (cst_body c) => //. - - rewrite cab in H |- *. cbn -[optimize] in *. - rewrite lookup_env_optimize. destruct lookup_env eqn:hl => // /=. + - rewrite cab in H |- *. cbn -[unbox] in *. + rewrite lookup_env_unbox. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => //. destruct nth_error => //. - - cbn -[optimize]. - rewrite lookup_env_optimize. cbn in H1. destruct lookup_env eqn:hl => // /=. + - cbn -[unbox]. + rewrite lookup_env_unbox. cbn in H1. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => //. rtoProp; intuition auto. - simp_optimize. toAll; solve_all. + simp_unbox. toAll; solve_all. toAll. solve_all. - - cbn -[optimize] in H0 |- *. - rewrite lookup_env_optimize. destruct lookup_env eqn:hl => // /=. + - cbn -[unbox] in H0 |- *. + rewrite lookup_env_unbox. destruct lookup_env eqn:hl => // /=. 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 -optimize_isLambda. toAll; solve_all. + now rewrite -unbox_isLambda. toAll; solve_all. - primProp. rtoProp; intuition eauto; solve_all_k 6. - move:H1; rewrite !wellformed_mkApps //. rtoProp; intuition auto. toAll; solve_all. toAll; solve_all. - move:H0; rewrite !wellformed_mkApps //. rtoProp; intuition auto. move: H1. cbn. rewrite cab. - rewrite lookup_env_optimize. destruct lookup_env eqn:hl => // /=. + rewrite lookup_env_unbox. 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. rewrite cab. - rewrite lookup_env_optimize. destruct lookup_env eqn:hl => // /=. + rewrite lookup_env_unbox. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => //. destruct nth_error => //. toAll; solve_all. Qed. -Lemma optimize_wellformed_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : +Lemma unbox_wellformed_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : cstr_as_blocks = false -> wf_glob Σ -> - forall n, wellformed Σ n t -> wellformed (optimize_env Σ) n t. + forall n, wellformed Σ n t -> wellformed (unbox_env Σ) n t. Proof. intros hcstrs wfΣ. induction t using EInduction.term_forall_list_ind; cbn => //. all:try solve [intros; unfold wf_fix in *; rtoProp; intuition eauto; solve_all]. - - rewrite lookup_env_optimize //. + - rewrite lookup_env_unbox //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct (cst_body c) => //. - - rewrite lookup_env_optimize //. + - rewrite lookup_env_unbox //. destruct lookup_env eqn:hl => // /=; intros; rtoProp; eauto. destruct g eqn:hg => /= //; intros; rtoProp; eauto. destruct cstr_as_blocks => //; repeat split; eauto. destruct nth_error => /= //. destruct nth_error => /= //. - - rewrite lookup_env_optimize //. + - rewrite lookup_env_unbox //. 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_unbox //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. rewrite andb_false_r => //. @@ -1192,44 +1162,44 @@ Proof. destruct nth_error => //. Qed. -Lemma optimize_wellformed_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} d : +Lemma unbox_wellformed_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} d : cstr_as_blocks = false -> wf_glob Σ -> - wf_global_decl Σ d -> wf_global_decl (optimize_env Σ) d. + wf_global_decl Σ d -> wf_global_decl (unbox_env Σ) d. Proof. intros hcstrs wf; destruct d => /= //. destruct (cst_body c) => /= //. - now eapply optimize_wellformed_irrel. + now eapply unbox_wellformed_irrel. Qed. -Lemma optimize_decl_wf (efl := all_env_flags) {Σ : GlobalContextMap.t} : +Lemma unbox_decl_wf (efl := all_env_flags) {Σ : GlobalContextMap.t} : wf_glob Σ -> forall d, wf_global_decl Σ d -> - wf_global_decl (efl := switch_no_params efl) (optimize_env Σ) (optimize_decl Σ d). + wf_global_decl (efl := switch_no_params efl) (unbox_env Σ) (unbox_decl Σ d). Proof. intros wf d. destruct d => /= //. destruct (cst_body c) => /= //. - now apply (optimize_wellformed (Σ := Σ) 0 t). + now apply (unbox_wellformed (Σ := Σ) 0 t). Qed. -Lemma fresh_global_optimize_env {Σ : GlobalContextMap.t} kn : +Lemma fresh_global_unbox_env {Σ : GlobalContextMap.t} kn : fresh_global kn Σ -> - fresh_global kn (optimize_env Σ). + fresh_global kn (unbox_env Σ). Proof. - unfold fresh_global. cbn. unfold optimize_env. + unfold fresh_global. cbn. unfold unbox_env. induction (GlobalContextMap.global_decls Σ); cbn; constructor; auto. cbn. now depelim H. depelim H. eauto. Qed. From MetaCoq.Erasure Require Import EProgram. -Program Fixpoint optimize_env' Σ : EnvMap.fresh_globals Σ -> global_context := +Program Fixpoint unbox_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 (unbox_decl Σg) hd :: unbox_env' tl (fresh_globals_cons_inv HΣ) end. Lemma lookup_minductive_declared_minductive Σ ind mdecl : @@ -1266,25 +1236,25 @@ Proof. apply lookup_minductive_declared_minductive in hl. now rewrite hl. Qed. -Lemma optimize_extends {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : +Lemma unbox_extends {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : has_tApp -> extends Σ Σ' -> wf_glob Σ' -> - forall n t, wellformed Σ n t -> optimize Σ t = optimize Σ' t. + forall n t, wellformed Σ n t -> unbox Σ t = unbox Σ' t. Proof. intros hasapp hext hwf n t. revert n. - funelim (optimize Σ t); intros ?n; simp_optimize => /= //. + funelim (unbox Σ t); intros ?n; simp_unbox => /= //. all:try solve [intros; unfold wf_fix in *; rtoProp; intuition auto; f_equal; auto; toAll; solve_all]. - intros; rtoProp; intuition auto. move: H1; rewrite wellformed_mkApps // => /andP[] wfu wfargs. - rewrite optimize_mkApps // Heq /=. f_equal; eauto. toAll; solve_all. + rewrite unbox_mkApps // Heq /=. f_equal; eauto. toAll; solve_all. - rewrite wellformed_mkApps // => /andP[] wfc wfv. - rewrite optimize_mkApps // /=. + rewrite unbox_mkApps // /=. rewrite GlobalContextMap.lookup_inductive_pars_spec in Heq. eapply (extends_lookup_inductive_pars hext hwf) in Heq. rewrite Heq. f_equal. f_equal. toAll; solve_all. - rewrite wellformed_mkApps // => /andP[] wfc wfv. - rewrite optimize_mkApps // /=. + rewrite unbox_mkApps // /=. move: Heq. rewrite GlobalContextMap.lookup_inductive_pars_spec. unfold wellformed in wfc. move/andP: wfc => [] /andP[] hacc hc bargs. @@ -1292,29 +1262,29 @@ Proof. unfold lookup_constructor, lookup_inductive in hc. rewrite heq /= // in hc. Qed. -Lemma optimize_decl_extends {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : +Lemma unbox_decl_extends {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : has_tApp -> extends Σ Σ' -> wf_glob Σ' -> - forall d, wf_global_decl Σ d -> optimize_decl Σ d = optimize_decl Σ' d. + forall d, wf_global_decl Σ d -> unbox_decl Σ d = unbox_decl Σ' d. Proof. intros hast ext wf []; cbn. - unfold optimize_constant_decl. + unfold unbox_constant_decl. destruct (EAst.cst_body c) eqn:hc => /= //. intros hwf. f_equal. f_equal. f_equal. - eapply optimize_extends => //. exact hwf. + eapply unbox_extends => //. exact hwf. intros _ => //. Qed. From MetaCoq.Erasure Require Import EGenericGlobalMap. #[local] -Instance GT : GenTransform := { gen_transform := optimize; gen_transform_inductive_decl := optimize_inductive_decl }. +Instance GT : GenTransform := { gen_transform := unbox; gen_transform_inductive_decl := unbox_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 optimize_extends; tea. + eapply unbox_extends; tea. Qed. #[local] Instance GTWf efl : GenTransformWf efl (switch_no_params efl) GT. @@ -1322,55 +1292,55 @@ 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 optimize_wellformed => //. + eapply unbox_wellformed => //. Defined. -Lemma optimize_extends' {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : +Lemma unbox_extends' {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). + List.map (on_snd (unbox_decl Σ)) Σ.(GlobalContextMap.global_decls) = + List.map (on_snd (unbox_decl Σ')) Σ.(GlobalContextMap.global_decls). Proof. intros hast ext wf. now apply (gen_transform_env_extends' (gt := GTExt efl hast) ext). Qed. -Lemma optimize_extends_env {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : +Lemma unbox_extends_env {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : has_tApp -> extends Σ Σ' -> wf_glob Σ -> wf_glob Σ' -> - extends (optimize_env Σ) (optimize_env Σ'). + extends (unbox_env Σ) (unbox_env Σ'). Proof. intros hast ext wf. now apply (gen_transform_extends (gt := GTExt efl hast) ext). Qed. -Lemma optimize_env_eq (efl := all_env_flags) (Σ : GlobalContextMap.t) : wf_glob Σ -> optimize_env Σ = optimize_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). +Lemma unbox_env_eq (efl := all_env_flags) (Σ : GlobalContextMap.t) : wf_glob Σ -> unbox_env Σ = unbox_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). Proof. intros wf. - unfold optimize_env. + unfold unbox_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 optimize_decl_extends => //. cbn. cbn. eapply EExtends.extends_prefix_extends. now exists [(kn, d)]. auto. cbn. now depelim wf. + eapply unbox_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 (optimize_extends' (Σ := Σg') (Σ' := Σg)) => //. + symmetry. eapply (unbox_extends' (Σ := Σg') (Σ' := Σg)) => //. cbn. eapply EExtends.extends_prefix_extends => //. now exists [a]. cbn. now depelim wf. Qed. -Lemma optimize_env_wf (efl := all_env_flags) {Σ : GlobalContextMap.t} : - wf_glob Σ -> wf_glob (efl := switch_no_params efl) (optimize_env Σ). +Lemma unbox_env_wf (efl := all_env_flags) {Σ : GlobalContextMap.t} : + wf_glob Σ -> wf_glob (efl := switch_no_params efl) (unbox_env Σ). Proof. intros wf. - rewrite (optimize_env_eq _ wf). + rewrite (unbox_env_eq _ wf). destruct Σ as [Σ map repr fr]; cbn in *. induction Σ in map, repr, fr, wf |- *. - cbn. constructor. @@ -1379,111 +1349,111 @@ Proof. constructor; eauto. eapply (IHΣ (GlobalContextMap.map Σg) (GlobalContextMap.repr Σg)). now depelim wf. depelim wf. cbn. - rewrite -(optimize_env_eq Σg). now cbn. cbn. - now eapply (optimize_decl_wf (Σ:=Σg)). - rewrite -(optimize_env_eq Σg). now depelim wf. - eapply fresh_global_optimize_env. now depelim fr. + rewrite -(unbox_env_eq Σg). now cbn. cbn. + now eapply (unbox_decl_wf (Σ:=Σg)). + rewrite -(unbox_env_eq Σg). now depelim wf. + eapply fresh_global_unbox_env. now depelim fr. Qed. -Lemma optimize_program_wf (efl := all_env_flags) (p : eprogram_env) : +Lemma unbox_program_wf (efl := all_env_flags) (p : eprogram_env) : wf_eprogram_env efl p -> - wf_eprogram (switch_no_params efl) (optimize_program p). + wf_eprogram (switch_no_params efl) (unbox_program p). Proof. intros []; split => //. - eapply (optimize_env_wf H). - now eapply optimize_wellformed. + eapply (unbox_env_wf H). + now eapply unbox_wellformed. Qed. -Lemma optimize_expanded {Σ : GlobalContextMap.t} {t} : expanded Σ t -> expanded (optimize_env Σ) (optimize Σ t). +Lemma unbox_expanded {Σ : GlobalContextMap.t} {t} : expanded Σ t -> expanded (unbox_env Σ) (unbox Σ t). Proof. induction 1 using expanded_ind. - all:try solve[simp_optimize; constructor; eauto; solve_all]. - - rewrite optimize_mkApps_etaexp. now eapply expanded_isEtaExp. + all:try solve[simp_unbox; constructor; eauto; solve_all]. + - rewrite unbox_mkApps_etaexp. now eapply expanded_isEtaExp. eapply expanded_mkApps_expanded => //. solve_all. - - simp_optimize; constructor; eauto. solve_all. - rewrite -optimize_isLambda //. - - rewrite optimize_mkApps // /=. + - simp_unbox; constructor; eauto. solve_all. + rewrite -unbox_isLambda //. + - rewrite unbox_mkApps // /=. rewrite (lookup_inductive_pars_spec (proj1 (proj1 H))). eapply expanded_tConstruct_app. - eapply optimize_declared_constructor; tea. + eapply unbox_declared_constructor; tea. len. rewrite skipn_length /= /EAst.cstr_arity /=. rewrite /EAst.cstr_arity in H0. lia. solve_all. eapply All_skipn. solve_all. Qed. -Lemma optimize_expanded_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded Σ t -> expanded (optimize_env Σ) t. +Lemma unbox_expanded_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded Σ t -> expanded (unbox_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. + red in H. rewrite lookup_env_unbox // /= H //. 1-2:eauto. cbn. rewrite /EAst.cstr_arity in H0. lia. solve_all. Qed. -Lemma optimize_expanded_decl {Σ : GlobalContextMap.t} t : expanded_decl Σ t -> expanded_decl (optimize_env Σ) (optimize_decl Σ t). +Lemma unbox_expanded_decl {Σ : GlobalContextMap.t} t : expanded_decl Σ t -> expanded_decl (unbox_env Σ) (unbox_decl Σ t). Proof. destruct t as [[[]]|] => /= //. unfold expanded_constant_decl => /=. - apply optimize_expanded. + apply unbox_expanded. Qed. -Lemma optimize_env_expanded (efl := all_env_flags) {Σ : GlobalContextMap.t} : - wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (optimize_env Σ). +Lemma unbox_env_expanded (efl := all_env_flags) {Σ : GlobalContextMap.t} : + wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (unbox_env Σ). Proof. unfold expanded_global_env. - intros wf. rewrite optimize_env_eq //. + intros wf. rewrite unbox_env_eq //. destruct Σ as [Σ map repr wf']; cbn in *. intros exp; induction exp in map, repr, wf', wf |- *; cbn. - constructor; auto. - set (Σ' := GlobalContextMap.make Σ (fresh_globals_cons_inv wf')). constructor; auto. eapply IHexp. eapply Σ'. now depelim wf. cbn. - eapply (optimize_expanded_decl (Σ := Σ')) in H. - rewrite -(optimize_env_eq Σ'). cbn. now depelim wf. + eapply (unbox_expanded_decl (Σ := Σ')) in H. + rewrite -(unbox_env_eq Σ'). cbn. now depelim wf. exact H. Qed. Import EProgram. -Lemma optimize_program_expanded (efl := all_env_flags) (p : eprogram_env) : +Lemma unbox_program_expanded (efl := all_env_flags) (p : eprogram_env) : wf_eprogram_env efl p -> expanded_eprogram_env_cstrs p -> - expanded_eprogram_cstrs (optimize_program p). + expanded_eprogram_cstrs (unbox_program p). Proof. unfold expanded_eprogram_env_cstrs, expanded_eprogram_cstrs. move=> [] wfe wft. move/andP => [] etaenv etap. apply/andP; split. - eapply expanded_global_env_isEtaExp_env, optimize_env_expanded => //. + eapply expanded_global_env_isEtaExp_env, unbox_env_expanded => //. now eapply isEtaExp_env_expanded_global_env. - now eapply expanded_isEtaExp, optimize_expanded, isEtaExp_expanded. + now eapply expanded_isEtaExp, unbox_expanded, isEtaExp_expanded. Qed. - tConstruct ind n (map optimize args) - | tPrim p => tPrim (map_prim optimize p) + tConstruct ind n (map unbox args) + | tPrim p => tPrim (map_prim unbox p) end. - Lemma optimize_mkApps f l : optimize (mkApps f l) = mkApps (optimize f) (map optimize l). + Lemma unbox_mkApps f l : unbox (mkApps f l) = mkApps (unbox f) (map unbox l). Proof using Type. induction l using rev_ind; simpl; auto. now rewrite mkApps_app /= IHl map_app /= mkApps_app /=. Qed. - Lemma map_optimize_repeat_box n : map optimize (repeat tBox n) = repeat tBox n. + Lemma map_unbox_repeat_box n : map unbox (repeat tBox n) = repeat tBox n. Proof using Type. by rewrite map_repeat. Qed. (* move to globalenv *) - Lemma isLambda_optimize t : isLambda t -> isLambda (optimize t). + Lemma isLambda_unbox t : isLambda t -> isLambda (unbox t). Proof. destruct t => //. Qed. - Lemma isBox_optimize t : isBox t -> isBox (optimize t). + Lemma isBox_unbox t : isBox t -> isBox (unbox t). Proof. destruct t => //. Qed. - Lemma wf_optimize t k : + Lemma wf_unbox t k : wf_glob Σ -> - wellformed Σ k t -> wellformed Σ k (optimize t). + wellformed Σ k t -> wellformed Σ k (unbox t). Proof using Type. intros wfΣ. induction t in k |- * using EInduction.term_forall_list_ind; simpl; auto; @@ -1503,15 +1473,15 @@ Qed. rewrite IHt //=; len. apply Nat.ltb_lt. lia. - len. rtoProp; solve_all. solve_all. - now eapply isLambda_optimize. solve_all. + now eapply isLambda_unbox. solve_all. - len. rtoProp; repeat solve_all. - rewrite test_prim_map. rtoProp; intuition eauto; solve_all. Qed. - Lemma optimize_csubst {a k b} n : + Lemma unbox_csubst {a k b} n : wf_glob Σ -> wellformed Σ (k + n) b -> - optimize (ECSubst.csubst a k b) = ECSubst.csubst (optimize a) k (optimize b). + unbox (ECSubst.csubst a k b) = ECSubst.csubst (unbox a) k (unbox b). Proof using Type. intros wfΣ. induction b in k |- * using EInduction.term_forall_list_ind; simpl; auto; @@ -1539,33 +1509,33 @@ Qed. eapply a0. now rewrite -Nat.add_assoc. Qed. - Lemma optimize_substl s t : + Lemma unbox_substl s t : wf_glob Σ -> forallb (wellformed Σ 0) s -> wellformed Σ #|s| t -> - optimize (substl s t) = substl (map optimize s) (optimize t). + unbox (substl s t) = substl (map unbox s) (unbox t). Proof using Type. intros wfΣ. induction s in t |- *; simpl; auto. move/andP => [] cla cls wft. rewrite IHs //. eapply wellformed_csubst => //. - f_equal. rewrite (optimize_csubst (S #|s|)) //. + f_equal. rewrite (unbox_csubst (S #|s|)) //. Qed. - Lemma optimize_iota_red pars args br : + Lemma unbox_iota_red pars args br : wf_glob Σ -> forallb (wellformed Σ 0) args -> wellformed Σ #|skipn pars args| br.2 -> - optimize (EGlobalEnv.iota_red pars args br) = EGlobalEnv.iota_red pars (map optimize args) (on_snd optimize br). + unbox (EGlobalEnv.iota_red pars args br) = EGlobalEnv.iota_red pars (map unbox args) (on_snd unbox br). Proof using Type. intros wfΣ wfa wfbr. unfold EGlobalEnv.iota_red. - rewrite optimize_substl //. + rewrite unbox_substl //. rewrite forallb_rev forallb_skipn //. now rewrite List.rev_length. 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 unbox_fix_subst mfix : EGlobalEnv.fix_subst (map (map_def unbox) mfix) = map unbox (EGlobalEnv.fix_subst mfix). Proof using Type. unfold EGlobalEnv.fix_subst. rewrite map_length. @@ -1574,7 +1544,7 @@ Qed. f_equal; auto. Qed. - Lemma optimize_cofix_subst mfix : EGlobalEnv.cofix_subst (map (map_def optimize) mfix) = map optimize (EGlobalEnv.cofix_subst mfix). + Lemma unbox_cofix_subst mfix : EGlobalEnv.cofix_subst (map (map_def unbox) mfix) = map unbox (EGlobalEnv.cofix_subst mfix). Proof using Type. unfold EGlobalEnv.cofix_subst. rewrite map_length. @@ -1583,11 +1553,11 @@ Qed. f_equal; auto. Qed. - Lemma optimize_cunfold_fix mfix idx n f : + Lemma unbox_cunfold_fix mfix idx n f : wf_glob Σ -> wellformed Σ 0 (tFix mfix idx) -> cunfold_fix mfix idx = Some (n, f) -> - cunfold_fix (map (map_def optimize) mfix) idx = Some (n, optimize f). + cunfold_fix (map (map_def unbox) mfix) idx = Some (n, unbox f). Proof using Type. intros wfΣ hfix. unfold cunfold_fix. @@ -1595,17 +1565,17 @@ Qed. cbn in hfix. move/andP: hfix => [] hlam /andP[] hidx hfix. destruct nth_error eqn:hnth => //. intros [= <- <-] => /=. f_equal. - rewrite optimize_substl //. eapply wellformed_fix_subst => //. + rewrite unbox_substl //. eapply wellformed_fix_subst => //. rewrite fix_subst_length. eapply nth_error_forallb in hfix; tea. now rewrite Nat.add_0_r in hfix. - now rewrite optimize_fix_subst. + now rewrite unbox_fix_subst. Qed. - Lemma optimize_cunfold_cofix mfix idx n f : + Lemma unbox_cunfold_cofix mfix idx n f : wf_glob Σ -> wellformed Σ 0 (tCoFix mfix idx) -> cunfold_cofix mfix idx = Some (n, f) -> - cunfold_cofix (map (map_def optimize) mfix) idx = Some (n, optimize f). + cunfold_cofix (map (map_def unbox) mfix) idx = Some (n, unbox f). Proof using Type. intros wfΣ hfix. unfold cunfold_cofix. @@ -1613,34 +1583,34 @@ Qed. cbn in hfix. move/andP: hfix => [] hidx hfix. destruct nth_error eqn:hnth => //. intros [= <- <-] => /=. f_equal. - rewrite optimize_substl //. eapply wellformed_cofix_subst => //. + rewrite unbox_substl //. eapply wellformed_cofix_subst => //. rewrite cofix_subst_length. eapply nth_error_forallb in hfix; tea. now rewrite Nat.add_0_r in hfix. - now rewrite optimize_cofix_subst. + now rewrite unbox_cofix_subst. Qed. -End optimize. +End unbox. -Definition optimize_constant_decl Σ cb := - {| cst_body := option_map (optimize Σ) cb.(cst_body) |}. +Definition unbox_constant_decl Σ cb := + {| cst_body := option_map (unbox Σ) cb.(cst_body) |}. -Definition optimize_decl Σ d := +Definition unbox_decl Σ d := match d with - | ConstantDecl cb => ConstantDecl (optimize_constant_decl Σ cb) + | ConstantDecl cb => ConstantDecl (unbox_constant_decl Σ cb) | InductiveDecl idecl => InductiveDecl idecl end. -Definition optimize_env Σ := - map (on_snd (optimize_decl Σ)) Σ.(GlobalContextMap.global_decls). +Definition unbox_env Σ := + map (on_snd (unbox_decl Σ)) Σ.(GlobalContextMap.global_decls). Import EnvMap. -Program Fixpoint optimize_env' Σ : EnvMap.fresh_globals Σ -> global_context := +Program Fixpoint unbox_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 (unbox_decl Σg) hd :: unbox_env' tl (fresh_globals_cons_inv HΣ) end. Import EGlobalEnv EExtends. @@ -1656,10 +1626,10 @@ Proof. rewrite (extends_lookup_constructor wf ext _ _ _ hl) //. Qed. -Lemma wellformed_optimize_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : +Lemma wellformed_unbox_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : forall n, EWellformed.wellformed Σ n t -> forall {Σ' : GlobalContextMap.t}, extends Σ Σ' -> wf_glob Σ' -> - optimize Σ t = optimize Σ' t. + unbox Σ t = unbox Σ' t. Proof. induction t using EInduction.term_forall_list_ind; cbn -[lookup_constant lookup_inductive GlobalContextMap.lookup_projection]; intros => //. @@ -1672,23 +1642,23 @@ Proof. now cbn in H3. Qed. -Lemma wellformed_optimize_decl_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : +Lemma wellformed_unbox_decl_extends {wfl: EEnvFlags} {Σ : GlobalContextMap.t} t : wf_global_decl Σ t -> forall {Σ' : GlobalContextMap.t}, extends Σ Σ' -> wf_glob Σ' -> - optimize_decl Σ t = optimize_decl Σ' t. + unbox_decl Σ t = unbox_decl Σ' t. Proof. destruct t => /= //. - intros wf Σ' ext wf'. f_equal. unfold optimize_constant_decl. f_equal. + intros wf Σ' ext wf'. f_equal. unfold unbox_constant_decl. f_equal. destruct (cst_body c) => /= //. f_equal. - now eapply wellformed_optimize_extends. + now eapply wellformed_unbox_extends. Qed. -Lemma lookup_env_optimize_env_Some {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn d : +Lemma lookup_env_unbox_env_Some {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn d : wf_glob Σ -> GlobalContextMap.lookup_env Σ kn = Some d -> ∑ Σ' : GlobalContextMap.t, [× extends_prefix Σ' Σ, wf_global_decl Σ' d & - lookup_env (optimize_env Σ) kn = Some (optimize_decl Σ' d)]. + lookup_env (unbox_env Σ) kn = Some (unbox_decl Σ' d)]. Proof. rewrite GlobalContextMap.lookup_env_spec. destruct Σ as [Σ map repr wf]. @@ -1699,7 +1669,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_unbox_decl_extends. cbn. now depelim wfg. eapply extends_prefix_extends. cbn. now exists [a]. now cbn. now cbn. - intros _. @@ -1713,7 +1683,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_unbox_decl_extends => //. cbn. eapply lookup_env_In in hin. 2:now depelim wfg. depelim wfg. eapply lookup_env_wellformed; tea. eapply extends_prefix_extends. @@ -1726,60 +1696,60 @@ Proof. case: eqb_spec => //. Qed. -Lemma lookup_env_optimize_env_None {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn : +Lemma lookup_env_unbox_env_None {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn : GlobalContextMap.lookup_env Σ kn = None -> - lookup_env (optimize_env Σ) kn = None. + lookup_env (unbox_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_unbox {efl : EEnvFlags} {Σ : GlobalContextMap.t} kn : wf_glob Σ -> - lookup_env (optimize_env Σ) kn = option_map (optimize_decl Σ) (lookup_env Σ kn). + lookup_env (unbox_env Σ) kn = option_map (unbox_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_unbox_env_Some in hl as [Σ' [ext wf' hl']] => /=. rewrite hl'. f_equal. - eapply wellformed_optimize_decl_extends; eauto. + eapply wellformed_unbox_decl_extends; eauto. now eapply extends_prefix_extends. auto. - - cbn. now eapply lookup_env_optimize_env_None in hl. + - cbn. now eapply lookup_env_unbox_env_None in hl. Qed. -Lemma is_propositional_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : +Lemma is_propositional_unbox {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 (unbox_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_unbox (inductive_mind ind) wf). rewrite /GlobalContextMap.inductive_isprop_and_pars /GlobalContextMap.lookup_inductive /GlobalContextMap.lookup_minductive. destruct lookup_env as [[decl|]|] => //. Qed. -Lemma lookup_inductive_pars_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : +Lemma lookup_inductive_pars_unbox {efl : EEnvFlags} {Σ : GlobalContextMap.t} ind : wf_glob Σ -> - EGlobalEnv.lookup_inductive_pars Σ ind = EGlobalEnv.lookup_inductive_pars (optimize_env Σ) ind. + EGlobalEnv.lookup_inductive_pars Σ ind = EGlobalEnv.lookup_inductive_pars (unbox_env Σ) ind. Proof. rewrite /lookup_inductive_pars => wf. rewrite /lookup_inductive /lookup_minductive. - rewrite (lookup_env_optimize ind wf). + rewrite (lookup_env_unbox 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 : +Lemma is_propositional_cstr_unbox {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 (unbox_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_unbox (inductive_mind ind) wf). rewrite /GlobalContextMap.inductive_isprop_and_pars /GlobalContextMap.lookup_inductive /GlobalContextMap.lookup_minductive. destruct lookup_env as [[decl|]|] => //. @@ -1809,21 +1779,21 @@ 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_unbox {efl : EEnvFlags} {Σ : GlobalContextMap.t} {ind c} : wf_glob Σ -> - lookup_constructor Σ ind c = lookup_constructor (optimize_env Σ) ind c. + lookup_constructor Σ ind c = lookup_constructor (unbox_env Σ) ind c. Proof. intros wfΣ. rewrite /lookup_constructor /lookup_inductive /lookup_minductive. - rewrite lookup_env_optimize // /=. destruct lookup_env => // /=. + rewrite lookup_env_unbox // /=. destruct lookup_env => // /=. destruct g => //. Qed. -Lemma lookup_projection_optimize {efl : EEnvFlags} {Σ : GlobalContextMap.t} {p} : +Lemma lookup_projection_unbox {efl : EEnvFlags} {Σ : GlobalContextMap.t} {p} : wf_glob Σ -> - lookup_projection Σ p = lookup_projection (optimize_env Σ) p. + lookup_projection Σ p = lookup_projection (unbox_env Σ) p. Proof. intros wfΣ. rewrite /lookup_projection. - rewrite -lookup_constructor_optimize //. + rewrite -lookup_constructor_unbox //. Qed. Lemma constructor_isprop_pars_decl_inductive {Σ ind c} {prop pars cdecl} : @@ -1867,11 +1837,11 @@ Proof. Qed. -Lemma optimize_correct {fl} {wcon : with_constructor_as_block = false} { Σ : GlobalContextMap.t} t v : +Lemma unbox_correct {fl} {wcon : with_constructor_as_block = false} { Σ : GlobalContextMap.t} t v : wf_glob Σ -> @eval fl Σ t v -> wellformed Σ 0 t -> - @eval fl (optimize_env Σ) (optimize Σ t) (optimize Σ v). + @eval fl (unbox_env Σ) (unbox Σ t) (unbox Σ v). Proof. intros wfΣ ev. induction ev; simpl in *. @@ -1881,24 +1851,24 @@ Proof. eapply eval_wellformed in ev2; tea => //. eapply eval_wellformed in ev1; tea => //. econstructor; eauto. - rewrite -(optimize_csubst _ 1) //. + rewrite -(unbox_csubst _ 1) //. apply IHev3. eapply wellformed_csubst => //. - move/andP => [] clb0 clb1. intuition auto. eapply eval_wellformed in ev1; tea => //. forward IHev2 by eapply wellformed_csubst => //. - econstructor; eauto. rewrite -(optimize_csubst _ 1) //. + econstructor; eauto. rewrite -(unbox_csubst _ 1) //. - - move/andP => [] /andP[] hl wfd wfbrs. rewrite optimize_mkApps in IHev1. + - move/andP => [] /andP[] hl wfd wfbrs. rewrite unbox_mkApps in IHev1. eapply eval_wellformed in ev1 => //. move/wf_mkApps: ev1 => [] wfc' wfargs. eapply nth_error_forallb in wfbrs; tea. rewrite Nat.add_0_r in wfbrs. forward IHev2. eapply wellformed_iota_red; tea => //. - rewrite optimize_iota_red in IHev2 => //. now rewrite e4. + rewrite unbox_iota_red in IHev2 => //. now rewrite e4. econstructor; eauto. - rewrite -is_propositional_cstr_optimize //. tea. + rewrite -is_propositional_cstr_unbox //. tea. rewrite nth_error_map e2 //. len. len. - congruence. @@ -1907,22 +1877,22 @@ Proof. forward IHev2. eapply wellformed_substl; tea => //. rewrite forallb_repeat //. len. rewrite e1 /= Nat.add_0_r in wfbrs. now move/andP: wfbrs. - rewrite optimize_substl in IHev2 => //. + rewrite unbox_substl in IHev2 => //. rewrite forallb_repeat //. len. rewrite e1 /= Nat.add_0_r in wfbrs. now move/andP: wfbrs. eapply eval_iota_sing => //; eauto. - rewrite -is_propositional_optimize //. + rewrite -is_propositional_unbox //. rewrite e1 //. simpl. rewrite map_repeat in IHev2 => //. - - move/andP => [] clf cla. rewrite optimize_mkApps in IHev1. + - move/andP => [] clf cla. rewrite unbox_mkApps in IHev1. simpl in *. eapply eval_wellformed in ev1 => //. move/wf_mkApps: ev1 => [] wff wfargs. eapply eval_fix; eauto. rewrite map_length. - eapply optimize_cunfold_fix; tea. - rewrite optimize_mkApps in IHev3. apply IHev3. + eapply unbox_cunfold_fix; tea. + rewrite unbox_mkApps in IHev3. apply IHev3. rewrite wellformed_mkApps // wfargs. eapply eval_wellformed in ev2; tas => //. rewrite ev2 /= !andb_true_r. eapply wellformed_cunfold_fix; tea. @@ -1931,16 +1901,16 @@ Proof. eapply eval_wellformed in ev1 => //. move/wf_mkApps: ev1 => [] clfix clargs. eapply eval_wellformed in ev2; tas => //. - rewrite optimize_mkApps in IHev1 |- *. + rewrite unbox_mkApps in IHev1 |- *. simpl in *. eapply eval_fix_value. auto. auto. auto. - eapply optimize_cunfold_fix; eauto. + eapply unbox_cunfold_fix; eauto. now rewrite map_length. - move/andP => [] clf cla. eapply eval_wellformed in ev1 => //. eapply eval_wellformed in ev2; tas => //. simpl in *. eapply eval_fix'. auto. auto. - eapply optimize_cunfold_fix; eauto. + eapply unbox_cunfold_fix; eauto. eapply IHev2; tea. eapply IHev3. apply/andP; split => //. eapply wellformed_cunfold_fix; tea. now cbn. @@ -1951,9 +1921,9 @@ Proof. forward IHev2. rewrite hl wellformed_mkApps // /= wfargs clbrs !andb_true_r. eapply wellformed_cunfold_cofix; tea => //. - rewrite !optimize_mkApps /= in IHev1, IHev2. + rewrite !unbox_mkApps /= in IHev1, IHev2. eapply eval_cofix_case. tea. - eapply optimize_cunfold_cofix; tea. + eapply unbox_cunfold_cofix; tea. exact IHev2. - move/andP => [] hl hd. @@ -1964,13 +1934,13 @@ Proof. forward IHev2. { rewrite /= wellformed_mkApps // wfargs andb_true_r. eapply wellformed_cunfold_cofix; tea. } - rewrite optimize_mkApps /= in IHev1. + rewrite unbox_mkApps /= in IHev1. eapply eval_cofix_case. eauto. - eapply optimize_cunfold_cofix; tea. - rewrite optimize_mkApps in IHev2 => //. + eapply unbox_cunfold_cofix; tea. + rewrite unbox_mkApps in IHev2 => //. - rewrite /declared_constant in isdecl. - move: (lookup_env_optimize c wfΣ). + move: (lookup_env_unbox c wfΣ). rewrite isdecl /= //. intros hl. econstructor; tea. cbn. rewrite e //. @@ -1987,15 +1957,15 @@ Proof. rewrite (constructor_isprop_pars_decl_constructor H) in e1. noconf e1. forward IHev1 by auto. forward IHev2. eapply nth_error_forallb in wfargs; tea. - rewrite optimize_mkApps /= in IHev1. + rewrite unbox_mkApps /= in IHev1. eapply eval_iota; tea. - rewrite /constructor_isprop_pars_decl -lookup_constructor_optimize // H //= //. + rewrite /constructor_isprop_pars_decl -lookup_constructor_unbox // H //= //. rewrite H0 H1; reflexivity. cbn. reflexivity. len. len. rewrite skipn_length. lia. unfold iota_red. cbn. - rewrite (substl_rel _ _ (optimize Σ a)) => //. + rewrite (substl_rel _ _ (unbox Σ a)) => //. { eapply nth_error_forallb in wfargs; tea. - eapply wf_optimize in wfargs => //. + eapply wf_unbox in wfargs => //. now eapply wellformed_closed in wfargs. } pose proof (wellformed_projection_args wfΣ hl'). cbn in H1. rewrite nth_error_rev. len. rewrite skipn_length. lia. @@ -2004,7 +1974,7 @@ Proof. rewrite e2 -H1. assert((ind_npars mdecl + cstr_nargs cdecl - ind_npars mdecl) = cstr_nargs cdecl) by lia. rewrite H3. - eapply (f_equal (option_map (optimize Σ))) in e3. + eapply (f_equal (option_map (unbox Σ))) in e3. cbn in e3. rewrite -e3. f_equal. f_equal. lia. - congruence. @@ -2018,7 +1988,7 @@ Proof. rewrite (lookup_constructor_lookup_inductive H) /=. intros [= eq <-]. eapply eval_iota_sing => //; eauto. - rewrite -is_propositional_optimize // /inductive_isprop_and_pars + rewrite -is_propositional_unbox // /inductive_isprop_and_pars (lookup_constructor_lookup_inductive H) //=. congruence. have lenarg := wellformed_projection_args wfΣ hl'. rewrite (substl_rel _ _ tBox) => //. @@ -2026,10 +1996,10 @@ Proof. now constructor. - move/andP=> [] clf cla. - rewrite optimize_mkApps. + rewrite unbox_mkApps. eapply eval_construct; tea. - rewrite -lookup_constructor_optimize //. exact e0. - rewrite optimize_mkApps in IHev1. now eapply IHev1. + rewrite -lookup_constructor_unbox //. exact e0. + rewrite unbox_mkApps in IHev1. now eapply IHev1. now len. now eapply IHev2. @@ -2041,18 +2011,18 @@ Proof. eapply eval_to_value in ev1. destruct ev1; simpl in *; eauto. * depelim a0. - + destruct t => //; rewrite optimize_mkApps /=. + + destruct t => //; rewrite unbox_mkApps /=. + now rewrite /= !orb_false_r orb_true_r in i. * destruct with_guarded_fix. + move: i. rewrite !negb_or. - rewrite optimize_mkApps !isFixApp_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + rewrite unbox_mkApps !isFixApp_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. rewrite !andb_true_r. rtoProp; intuition auto; destruct v => /= //. + move: i. rewrite !negb_or. - rewrite optimize_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + rewrite unbox_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. destruct v => /= //. - intros; rtoProp; intuition eauto. @@ -2063,18 +2033,18 @@ Proof. primProp; depelim H0; intuition eauto. - destruct t => //. all:constructor; eauto. - cbn [atom optimize] in i |- *. - rewrite -lookup_constructor_optimize //. + cbn [atom unbox] in i |- *. + rewrite -lookup_constructor_unbox //. destruct args; cbn in *; eauto. Qed. From MetaCoq.Erasure Require Import EEtaExpanded. -Lemma optimize_expanded {Σ : GlobalContextMap.t} t : expanded Σ t -> expanded Σ (optimize Σ t). +Lemma unbox_expanded {Σ : GlobalContextMap.t} t : expanded Σ t -> expanded Σ (unbox Σ t). Proof. induction 1 using expanded_ind. all:try solve[constructor; eauto; solve_all]. - all:rewrite ?optimize_mkApps. + all:rewrite ?unbox_mkApps. - eapply expanded_mkApps_expanded => //. solve_all. - cbn -[GlobalContextMap.inductive_isprop_and_pars]. rewrite GlobalContextMap.lookup_projection_spec. @@ -2083,39 +2053,39 @@ Proof. destruct proj as [[ind npars] arg]. econstructor; eauto. repeat constructor. - cbn. eapply expanded_tFix. solve_all. - rewrite isLambda_optimize //. + rewrite isLambda_unbox //. - 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 unbox_expanded_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded Σ t -> expanded (unbox_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_unbox // /= H //. 1-2:eauto. auto. solve_all. Qed. -Lemma optimize_expanded_decl {Σ : GlobalContextMap.t} t : expanded_decl Σ t -> expanded_decl Σ (optimize_decl Σ t). +Lemma unbox_expanded_decl {Σ : GlobalContextMap.t} t : expanded_decl Σ t -> expanded_decl Σ (unbox_decl Σ t). Proof. destruct t as [[[]]|] => /= //. unfold expanded_constant_decl => /=. - apply optimize_expanded. + apply unbox_expanded. Qed. -Lemma optimize_expanded_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded_decl Σ t -> expanded_decl (optimize_env Σ) t. +Lemma unbox_expanded_decl_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> expanded_decl Σ t -> expanded_decl (unbox_env Σ) t. Proof. destruct t as [[[]]|] => /= //. unfold expanded_constant_decl => /=. - apply optimize_expanded_irrel. + apply unbox_expanded_irrel. Qed. -Lemma optimize_wellformed_term {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : +Lemma unbox_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). + EWellformed.wellformed (efl := disable_projections_env_flag efl) Σ n (unbox Σ t). Proof. intros hbox hrel wfΣ. induction t in n |- * using EInduction.term_forall_list_ind => //. @@ -2132,50 +2102,50 @@ Proof. rewrite hrel IHt //= andb_true_r. have hargs' := wellformed_projection_args wfΣ hl'. apply Nat.ltb_lt. len. - - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now eapply isLambda_optimize. + - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now eapply isLambda_unbox. - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. - cbn. rtoProp; intuition eauto; solve_all_k 6. Qed. Import EWellformed. -Lemma optimize_wellformed_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : +Lemma unbox_wellformed_irrel {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : wf_glob Σ -> forall n, wellformed (efl := disable_projections_env_flag efl) Σ n t -> - wellformed (efl := disable_projections_env_flag efl) (optimize_env Σ) n t. + wellformed (efl := disable_projections_env_flag efl) (unbox_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_unbox //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. repeat (rtoProp; intuition auto). destruct has_axioms => //. cbn in *. destruct (cst_body c) => //. - - rewrite lookup_env_optimize //. + - rewrite lookup_env_unbox //. 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_optimize //. + - rewrite lookup_env_unbox //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => /= //. intros; rtoProp; intuition auto; solve_all. Qed. -Lemma optimize_wellformed {efl : EEnvFlags} {Σ : GlobalContextMap.t} n t : +Lemma unbox_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). + EWellformed.wellformed (efl := disable_projections_env_flag efl) (unbox_env Σ) n (unbox Σ t). Proof. - intros. apply optimize_wellformed_irrel => //. - now apply optimize_wellformed_term. + intros. apply unbox_wellformed_irrel => //. + now apply unbox_wellformed_term. Qed. From MetaCoq.Erasure Require Import EGenericGlobalMap. #[local] -Instance GT : GenTransform := { gen_transform := optimize; gen_transform_inductive_decl := id }. +Instance GT : GenTransform := { gen_transform := unbox; gen_transform_inductive_decl := id }. #[local] Instance GTID : GenTransformId GT. Proof. @@ -2186,47 +2156,47 @@ Instance GTExt efl : GenTransformExtends efl (disable_projections_env_flag efl) Proof. intros Σ t n wfΣ Σ' ext wf wf'. unfold gen_transform, GT. - eapply wellformed_optimize_extends; tea. + eapply wellformed_unbox_extends; tea. Qed. #[local] Instance GTWf efl : GenTransformWf efl (disable_projections_env_flag efl) GT. Proof. refine {| gen_transform_pre := fun _ _ => has_tBox /\ has_tRel |}; auto. intros Σ n t [] wfΣ wft. - now apply optimize_wellformed. + now apply unbox_wellformed. Defined. -Lemma optimize_extends_env {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : +Lemma unbox_extends_env {efl : EEnvFlags} {Σ Σ' : GlobalContextMap.t} : has_tApp -> extends Σ Σ' -> wf_glob Σ -> wf_glob Σ' -> - extends (optimize_env Σ) (optimize_env Σ'). + extends (unbox_env Σ) (unbox_env Σ'). Proof. 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). +Lemma unbox_env_eq {efl : EEnvFlags} (Σ : GlobalContextMap.t) : wf_glob Σ -> unbox_env Σ = unbox_env' Σ.(GlobalContextMap.global_decls) Σ.(GlobalContextMap.wf). Proof. intros wf. now apply (gen_transform_env_eq (gt := GTExt efl)). Qed. -Lemma optimize_env_expanded {efl : EEnvFlags} {Σ : GlobalContextMap.t} : - wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (optimize_env Σ). +Lemma unbox_env_expanded {efl : EEnvFlags} {Σ : GlobalContextMap.t} : + wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (unbox_env Σ). Proof. unfold expanded_global_env; move=> wfg. - rewrite optimize_env_eq //. + rewrite unbox_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 -(unbox_env_eq Σ'). cbn. now depelim wfg. + eapply (unbox_expanded_decl_irrel (Σ := Σ')). now depelim wfg. + now unshelve eapply (unbox_expanded_decl (Σ:=Σ')). Qed. Lemma Pre_glob efl Σ wf : @@ -2237,36 +2207,37 @@ Proof. split => //. destruct d as [[[|]]|] => //=. Qed. -Lemma optimize_env_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : +Lemma unbox_env_wf {efl : EEnvFlags} {Σ : GlobalContextMap.t} : has_tBox -> has_tRel -> - wf_glob Σ -> wf_glob (efl := disable_projections_env_flag efl) (optimize_env Σ). + wf_glob Σ -> wf_glob (efl := disable_projections_env_flag efl) (unbox_env Σ). Proof. intros hasb hasre wfg. eapply (gen_transform_env_wf (gt := GTExt efl)) => //. now apply Pre_glob. Qed. -Definition optimize_program (p : eprogram_env) := - (optimize_env p.1, optimize p.1 p.2). +Definition unbox_program (p : eprogram_env) := + (unbox_env p.1, unbox p.1 p.2). -Definition optimize_program_wf {efl : EEnvFlags} (p : eprogram_env) {hastbox : has_tBox} {hastrel : has_tRel} : - wf_eprogram_env efl p -> wf_eprogram (disable_projections_env_flag efl) (optimize_program p). +Definition unbox_program_wf {efl : EEnvFlags} (p : eprogram_env) {hastbox : has_tBox} {hastrel : has_tRel} : + wf_eprogram_env efl p -> wf_eprogram (disable_projections_env_flag efl) (unbox_program p). Proof. intros []; split. - now eapply optimize_env_wf. - cbn. now eapply optimize_wellformed. + now eapply unbox_env_wf. + cbn. now eapply unbox_wellformed. Qed. -Definition optimize_program_expanded {efl} (p : eprogram_env) : +Definition unbox_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 (unbox_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, unbox_env_expanded => //. now eapply isEtaExp_env_expanded_global_env. eapply expanded_isEtaExp. - eapply optimize_expanded_irrel => //. - now apply optimize_expanded, isEtaExp_expanded. + eapply unbox_expanded_irrel => //. + now apply unbox_expanded, isEtaExp_expanded. Qed. +*) \ No newline at end of file diff --git a/test-suite/erasure_live_test.v b/test-suite/erasure_live_test.v index c780c2996..d9796c1f1 100644 --- a/test-suite/erasure_live_test.v +++ b/test-suite/erasure_live_test.v @@ -38,10 +38,12 @@ Definition exintrotest := Eval lazy in test exintro. Definition ex_type_introtest := Eval lazy in testty exintro. -Definition idnat := ((fun (X : Set) (x : X) => x) nat). +Definition id := fun (X : Set) (x : X) => x. +Definition idnat := (id nat). MetaCoq Quote Recursively Definition idnatc := idnat. Time Definition test_idnat := Eval lazy in test idnatc. +Time Definition testty_idnat := Eval lazy in testty idnatc. (** Check that optimization of singleton pattern-matchings work *) Definition singlelim := ((fun (X : Set) (x : X) (e : x = x) => From 9ee34f6ba35173178cde5c5638513949d4680712 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 23 Feb 2024 14:20:37 +0100 Subject: [PATCH 15/28] Unsafe inlining and beta-reduction transfomations --- erasure-plugin/_PluginProject.in | 4 + erasure-plugin/src/g_metacoq_erasure.mlg | 5 +- .../src/metacoq_erasure_plugin.mlpack | 2 + erasure-plugin/theories/ETransform.v | 6 - erasure-plugin/theories/Erasure.v | 41 ++-- erasure-plugin/theories/ErasureCorrectness.v | 10 +- erasure/_CoqProject.in | 2 + erasure/theories/EBeta.v | 120 ++++++++++++ erasure/theories/EEnvMap.v | 13 ++ erasure/theories/EInlining.v | 176 ++++++++++++++++++ erasure/theories/EProgram.v | 9 +- erasure/theories/EUnboxing.v | 1 - 12 files changed, 357 insertions(+), 32 deletions(-) create mode 100644 erasure/theories/EBeta.v create mode 100644 erasure/theories/EInlining.v diff --git a/erasure-plugin/_PluginProject.in b/erasure-plugin/_PluginProject.in index dca62f1cd..dae86a499 100644 --- a/erasure-plugin/_PluginProject.in +++ b/erasure-plugin/_PluginProject.in @@ -112,6 +112,10 @@ src/eOptimizePropDiscr.mli src/eOptimizePropDiscr.ml src/optimizePropDiscr.mli src/optimizePropDiscr.ml +src/eInlining.mli +src/eInlining.ml +src/eBeta.mli +src/eBeta.ml src/eProgram.mli src/eProgram.ml src/eInlineProjections.mli diff --git a/erasure-plugin/src/g_metacoq_erasure.mlg b/erasure-plugin/src/g_metacoq_erasure.mlg index ba9dce88e..bb3a14cb9 100644 --- a/erasure-plugin/src/g_metacoq_erasure.mlg +++ b/erasure-plugin/src/g_metacoq_erasure.mlg @@ -59,11 +59,12 @@ let default_config = let make_erasure_config config = let open Erasure0 in - { enable_cofix_to_fix = config.unsafe; + { enable_unsafe = config.unsafe; enable_typed_erasure = config.typed; enable_fast_remove_params = config.fast; dearging_config = default_dearging_config; - inductives_mapping = [] } + inductives_mapping = []; + inlining = Kernames.KernameSet.empty } let time_opt config str fn arg = if config.time then diff --git a/erasure-plugin/src/metacoq_erasure_plugin.mlpack b/erasure-plugin/src/metacoq_erasure_plugin.mlpack index c689fdf8b..bbded2dd3 100644 --- a/erasure-plugin/src/metacoq_erasure_plugin.mlpack +++ b/erasure-plugin/src/metacoq_erasure_plugin.mlpack @@ -58,6 +58,8 @@ ESpineView EPretty Extract EEtaExpanded +EInlining +EBeta ERemoveParams ErasureFunction ErasureFunctionProperties diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index b38027537..3af3e47a8 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -23,7 +23,6 @@ Definition build_wf_env_from_env {cf : checker_flags} (Σ : global_env_map) (wf wf_env_map_repr := Σ.(trans_env_repr); |}. - Notation NormalizationIn_erase_pcuic_program_1 p := (@PCUICTyping.wf_ext config.extraction_checker_flags p -> PCUICSN.NormalizationIn (cf:=config.extraction_checker_flags) (no:=PCUICSN.extraction_normalizing) p) (only parsing). @@ -601,11 +600,6 @@ Section Dearging. End Dearging. -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. diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 1eb37df4d..f24a1d62c 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -6,7 +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.Erasure Require Import EProgram EInlining EBeta. From MetaCoq.ErasurePlugin Require Import ETransform. Import PCUICProgram. @@ -31,11 +31,12 @@ Import EWcbvEval. Local Obligation Tactic := program_simpl. Record erasure_configuration := { - enable_cofix_to_fix : bool; + enable_unsafe : bool; enable_typed_erasure : bool; enable_fast_remove_params : bool; dearging_config : dearging_config; - inductives_mapping : EReorderCstrs.inductives_mapping + inductives_mapping : EReorderCstrs.inductives_mapping; + inlining : KernameSet.t }. Definition default_dearging_config := @@ -45,19 +46,21 @@ Definition default_dearging_config := (* This runs the cofix -> fix translation which is not entirely verified yet *) Definition default_erasure_config := - {| enable_cofix_to_fix := true; + {| enable_unsafe := true; dearging_config := default_dearging_config; enable_typed_erasure := true; enable_fast_remove_params := true; - inductives_mapping := [] |}. + inductives_mapping := []; + inlining := KernameSet.empty |}. (* This runs only the verified phases without the typed erasure and "fast" remove params *) Definition safe_erasure_config := - {| enable_cofix_to_fix := false; + {| enable_unsafe := false; enable_typed_erasure := false; enable_fast_remove_params := false; dearging_config := default_dearging_config; - inductives_mapping := [] |}. + inductives_mapping := []; + inlining := KernameSet.empty |}. Axiom assume_welltyped_template_program_expansion : forall p (wtp : ∥ wt_template_program_env p ∥), @@ -95,7 +98,7 @@ Definition final_wcbv_flags := {| Program Definition optional_unsafe_transforms econf := let efl := EConstructorsAsBlocks.switch_cstr_as_blocks (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags)) in - ETransform.optional_self_transform econf.(enable_cofix_to_fix) + ETransform.optional_self_transform econf.(enable_unsafe) ((* Rebuild the efficient lookup table *) rebuild_wf_env_transform (efl := efl) false false ▷ (* Coinductives & cofixpoints are translated to inductive types and thunked fixpoints *) @@ -103,7 +106,10 @@ Program Definition optional_unsafe_transforms econf := (has_app := eq_refl) (has_box := eq_refl) (has_rel := eq_refl) (has_pars := eq_refl) (has_cstrblocks := eq_refl) ▷ reorder_cstrs_transformation efl final_wcbv_flags econf.(inductives_mapping) ▷ rebuild_wf_env_transform (efl := efl) false false ▷ - unbox_transformation efl final_wcbv_flags). + unbox_transformation efl final_wcbv_flags ▷ + inline_transformation efl final_wcbv_flags econf.(inlining) ▷ + forget_inlining_info_transformation efl final_wcbv_flags ▷ + betared_transformation efl final_wcbv_flags). Program Definition verified_lambdabox_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) @@ -258,8 +264,7 @@ Qed. Next Obligation. unfold optional_unsafe_transforms. cbn. - unfold optional_self_transform. cbn. - destruct enable_cofix_to_fix => //. + destruct enable_unsafe => //. Qed. Local Obligation Tactic := intros; eauto. @@ -358,7 +363,7 @@ Next Obligation. cbn in H. split; cbn; intuition eauto. Qed. Next Obligation. - cbn in H. unfold optional_unsafe_transforms. destruct enable_cofix_to_fix => //. + cbn in H. unfold optional_unsafe_transforms. destruct enable_unsafe => //. Qed. Next Obligation. cbn in H. split; cbn; intuition eauto. @@ -403,7 +408,7 @@ Program Definition run_erase_program {guard : abstract_guard_impl} econf := Next Obligation. Proof. unfold optional_unsafe_transforms. - destruct enable_cofix_to_fix => //. + destruct enable_unsafe => //. Qed. Program Definition erase_and_print_template_program econf (p : Ast.Env.program) : string := @@ -418,21 +423,23 @@ Next Obligation. Qed. Definition erasure_fast_config := - {| enable_cofix_to_fix := false; + {| enable_unsafe := false; dearging_config := default_dearging_config; enable_typed_erasure := false; enable_fast_remove_params := true; - inductives_mapping := [] |}. + inductives_mapping := []; + inlining := KernameSet.empty |}. Program Definition erase_fast_and_print_template_program (p : Ast.Env.program) : string := erase_and_print_template_program erasure_fast_config p. Definition typed_erasure_config := - {| enable_cofix_to_fix := false; + {| enable_unsafe := false; dearging_config := default_dearging_config; enable_typed_erasure := true; enable_fast_remove_params := true; - inductives_mapping := [] |}. + inductives_mapping := []; + inlining := KernameSet.empty |}. (* Parameterized by a configuration for dearging, allowing to, e.g., override masks. *) Program Definition typed_erase_and_print_template_program (p : Ast.Env.program) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index d7cee5569..1c9100f9f 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -5,10 +5,10 @@ From MetaCoq.Utils Require Import bytestring utils. From MetaCoq.PCUIC Require PCUICAst PCUICAstUtils PCUICProgram. From MetaCoq.PCUIC Require Import PCUICNormal. From MetaCoq.SafeChecker Require Import PCUICErrors PCUICWfEnvImpl. -From MetaCoq.Erasure Require EAstUtils ErasureCorrectness EPretty Extract EConstructorsAsBlocks. +From MetaCoq.Erasure Require EAstUtils ErasureCorrectness EPretty Extract EProgram EConstructorsAsBlocks. From MetaCoq.Erasure Require Import EWcbvEvalNamed ErasureFunction ErasureFunctionProperties. From MetaCoq.ErasurePlugin Require Import ETransform Erasure. -Import PCUICProgram. +Import EProgram PCUICProgram. Import PCUICTransform (template_to_pcuic_transform, pcuic_expand_lets_transform). (* This is the total erasure function + @@ -1742,7 +1742,7 @@ Section PCUICErase. match goal with [ |- context [ @erase ?X_type ?X ?nin ?G (tApp _ _) ?wt ] ] => unshelve epose proof (@erase_mkApps X_type X nin G t [u] wt (wt'_erase_pcuic_program (Σ, t) prf3 prf0)) - end. + end. assert (hargs : forall Σ : global_env_ext, Σ ∼_ext env -> ∥ All (welltyped Σ []) [u] ∥). { cbn; intros ? ->. do 2 constructor; auto. destruct prf4 as [[T HT]]. eexists; eapply HT. } specialize (H hargs). @@ -2120,7 +2120,7 @@ Section pipeline_cond. Opaque compose. - Lemma verified_erasure_pipeline_lookup_env_in kn decl (efl := EInlineProjections.switch_no_params all_env_flags) + Lemma verified_erasure_pipeline_lookup_env_in kn decl (efl := EInlineProjections.switch_no_params all_env_flags) {has_rel : has_tRel} {has_box : has_tBox} : EGlobalEnv.lookup_env Σ_t kn = Some decl -> exists decl', @@ -2216,7 +2216,7 @@ Section pipeline_theorem. Let Σ_v := (transform verified_erasure_pipeline (Σ, v) (precond2 _ _ _ _ expΣ expt typing _ _ Heval)).1. Let v_t := compile_value_box (PCUICExpandLets.trans_global_env Σ) v []. - Lemma verified_erasure_pipeline_extends (efl := EInlineProjections.switch_no_params all_env_flags) + Lemma verified_erasure_pipeline_extends (efl := EInlineProjections.switch_no_params all_env_flags) {has_rel : has_tRel} {has_box : has_tBox} : EGlobalEnv.extends Σ_v Σ_t. Proof. diff --git a/erasure/_CoqProject.in b/erasure/_CoqProject.in index 4d0ec4f9b..44a814a90 100644 --- a/erasure/_CoqProject.in +++ b/erasure/_CoqProject.in @@ -30,6 +30,8 @@ theories/ErasureFunction.v theories/ErasureFunctionProperties.v theories/EExtends.v theories/EOptimizePropDiscr.v +theories/EInlining.v +theories/EBeta.v theories/EEtaExpandedFix.v theories/EEtaExpanded.v theories/EProgram.v diff --git a/erasure/theories/EBeta.v b/erasure/theories/EBeta.v new file mode 100644 index 000000000..75adc987d --- /dev/null +++ b/erasure/theories/EBeta.v @@ -0,0 +1,120 @@ +From Coq Require Import List. +From Coq Require Import String. +From MetaCoq.Utils Require Import utils. +From MetaCoq.Common Require Import BasicAst. +From MetaCoq.Erasure Require Import EPrimitive EAst EAstUtils ELiftSubst EProgram. + +Definition map_subterms (f : term -> term) (t : term) : term := + match t with + | tEvar n ts => tEvar n (map f ts) + | tLambda na body => tLambda na (f body) + | tLetIn na val body => tLetIn na (f val) (f body) + | tApp hd arg => tApp (f hd) (f arg) + | tCase p disc brs => + tCase p (f disc) (map (on_snd f) brs) + | tProj p t => tProj p (f t) + | tFix def i => tFix (map (map_def f) def) i + | tCoFix def i => tCoFix (map (map_def f) def) i + | tPrim p => tPrim (map_prim f p) + | tLazy t => tLazy (f t) + | tForce t => tForce (f t) + | tConstruct ind n args => tConstruct ind n (map f args) + | t => t + end. + +Section betared. + Fixpoint decompose_lam (t : term) {struct t} : list name × term := + match t with + | tLambda na B => + let (ns, B0) := decompose_lam B in + (na :: ns, B0) + | _ => ([], t) + end. + + Fixpoint beta_body (body : term) (args : list term) {struct args} : term := + match args with + | [] => body + | a :: args => + match body with + | tLambda na body => beta_body (body{0 := a}) args + | _ => mkApps body (a :: args) + end + end. + + Fixpoint betared_aux (args : list term) (t : term) : term := + match t with + | tApp hd arg => betared_aux (betared_aux [] arg :: args) hd + | tLambda na body => + let b := betared_aux [] body in + beta_body (tLambda na b) args + | t => mkApps (map_subterms (betared_aux []) t) args + end. + + Definition betared : term -> term := betared_aux []. + + Definition betared_in_constant_body cst := + {| cst_body := option_map betared (cst_body cst); |}. + + Definition betared_in_decl d := + match d with + | ConstantDecl cst => ConstantDecl (betared_in_constant_body cst) + | _ => d + end. + +End betared. + +Definition betared_env (Σ : global_declarations) : global_declarations := + map (fun '(kn, decl) => (kn, betared_in_decl decl)) Σ. + +Definition betared_program (p : program) : program := + (betared_env p.1, betared p.2). + +From MetaCoq.Erasure Require Import EProgram EWellformed EWcbvEval. +From MetaCoq.Common Require Import Transform. + +Axiom trust_betared_wf : + forall efl : EEnvFlags, + WcbvFlags -> + forall (input : Transform.program _ term), + wf_eprogram efl input -> wf_eprogram efl (betared_program input). +Axiom trust_betared_pres : + forall (efl : EEnvFlags) (wfl : WcbvFlags) (p : Transform.program _ term) + (v : term), + wf_eprogram efl p -> + eval_eprogram wfl p v -> + exists v' : term, + eval_eprogram wfl (betared_program p) v' /\ v' = betared v. + +Import Transform. + +Program Definition betared_transformation (efl : EEnvFlags) (wfl : WcbvFlags) : + Transform.t _ _ EAst.term EAst.term _ _ + (eval_eprogram wfl) (eval_eprogram wfl) := + {| name := "betared "; + transform p _ := betared_program p ; + pre p := wf_eprogram efl p ; + post p := wf_eprogram efl p ; + obseq p hp p' v v' := v' = betared v |}. + +Next Obligation. + now apply trust_betared_wf. +Qed. +Next Obligation. + now eapply trust_betared_pres. +Qed. + +Import EProgram EGlobalEnv. + +#[global] +Axiom betared_transformation_ext : + forall (efl : EEnvFlags) (wfl : WcbvFlags), + TransformExt.t (betared_transformation efl wfl) + (fun p p' => extends p.1 p'.1) (fun p p' => extends p.1 p'.1). + +#[global] +Axiom betared_transformation_ext' : + forall (efl : EEnvFlags) (wfl : WcbvFlags), + TransformExt.t (betared_transformation efl wfl) + extends_eprogram extends_eprogram. + + diff --git a/erasure/theories/EEnvMap.v b/erasure/theories/EEnvMap.v index 8687285df..c08a5cf14 100644 --- a/erasure/theories/EEnvMap.v +++ b/erasure/theories/EEnvMap.v @@ -69,6 +69,19 @@ Module GlobalContextMap. rewrite (EnvMap.lookup_spec Σ.(global_decls)) //; apply Σ. Qed. + Definition lookup_constant Σ kn : option constant_body := + decl <- lookup_env Σ kn;; + match decl with + | ConstantDecl cb => ret cb + | InductiveDecl _ => None + end. + + Lemma lookup_constant_spec Σ kn : lookup_constant Σ kn = EGlobalEnv.lookup_constant Σ kn. + Proof. + rewrite /lookup_constant /EGlobalEnv.lookup_constant. + rewrite lookup_env_spec //. + Qed. + Definition lookup_minductive Σ kn : option mutual_inductive_body := decl <- lookup_env Σ kn;; match decl with diff --git a/erasure/theories/EInlining.v b/erasure/theories/EInlining.v new file mode 100644 index 000000000..529e5f9ff --- /dev/null +++ b/erasure/theories/EInlining.v @@ -0,0 +1,176 @@ +From Coq Require Import List String Arith Lia ssreflect ssrbool. +Import ListNotations. +From Equations Require Import Equations. +Set Equations Transparent. + +From MetaCoq.PCUIC Require Import PCUICAstUtils. +From MetaCoq.Utils Require Import MCList bytestring utils monad_utils. +From MetaCoq.Erasure Require Import EPrimitive EAst EEnvMap EInduction EGlobalEnv. + +Import Kernames. +Import MCMonadNotation. + +(* Inlining hints *) +Definition inlining := KernameSet.t. +(* Fast maps for actual inlinings *) +Definition constants_inlining := KernameMap.t term. + +Section Inline. + Context (inlining : constants_inlining). + + Equations inline (t : term) : term := + | tVar na => tVar na + | tLambda nm bod => tLambda nm (inline bod) + | tLetIn nm dfn bod => tLetIn nm dfn bod + | tApp fn arg => tApp (inline fn) (inline arg) + | tConst nm with KernameMap.find nm inlining := + { | Some body := (* Already inlined body *) body + | None := tConst nm } + | tConstruct i m args => tConstruct i m (map inline args) + | tCase i mch brs => tCase i mch (map (on_snd inline) brs) + | tFix mfix idx => tFix (map (map_def inline) mfix) idx + | tCoFix mfix idx => tCoFix (map (map_def inline) mfix) idx + | tProj p bod => tProj p (inline bod) + | tPrim p => tPrim (map_prim inline p) + | tLazy t => tLazy (inline t) + | tForce t => tForce (inline t) + | tRel n => tRel n + | tBox => tBox + | tEvar ev args => tEvar ev (map inline args). +End Inline. + +Definition inline_constant_decl inlining cb := + {| cst_body := option_map (inline inlining) cb.(cst_body) |}. + +Definition inline_decl inlining d : (kername * global_decl) := + match d.2 with + | ConstantDecl cb => (d.1, ConstantDecl (inline_constant_decl inlining cb)) + | InductiveDecl idecl => d + end. + +Definition inline_add (inlining : inlining) (acc : global_context × constants_inlining) decl := + let '(Σ, inls) := acc in + if KernameSet.mem decl.1 inlining then + match decl.2 with + | ConstantDecl {| cst_body := Some body |} => KernameMap.add decl.1 body inls + | _ => inls + end + else inls. + +Definition inline_env (inlining : KernameSet.t) Σ : global_context × constants_inlining := + let inline_decl '(Σ, inls) decl := + let inldecl := inline_decl inls decl in + (inldecl :: Σ, inline_add inlining (Σ, inls) inldecl) + in + fold_left (inline_decl) (rev Σ) ([], KernameMap.empty _). + +Definition inlined_program := + (global_context × constants_inlining) × term. + +Definition inlined_program_inlinings (pr : inlined_program) : constants_inlining := + pr.1.2. + +Coercion inlined_program_inlinings : inlined_program >-> constants_inlining. + +Definition inline_program inlining (p : program) : inlined_program := + let '(Σ', inls) := inline_env inlining p.1 in + (Σ', inls, inline inls p.2). + +From MetaCoq.Erasure Require Import EProgram EWellformed EWcbvEval. +From MetaCoq.Common Require Import Transform. + +Definition forget_inlining_info (pr : inlined_program) : eprogram := + let '((Σ', inls), p) := pr in + (Σ', p). + +Coercion forget_inlining_info : inlined_program >-> eprogram. + +Definition eval_inlined_program wfl (pr : inlined_program) := eval_eprogram wfl pr. + +Axiom trust_inlining_wf : + forall efl : EEnvFlags, + WcbvFlags -> + forall inlining : inlining, + forall (input : Transform.program _ term), + wf_eprogram efl input -> wf_eprogram efl (inline_program inlining input). +Axiom trust_inlining_pres : + forall (efl : EEnvFlags) (wfl : WcbvFlags) inlining (p : Transform.program _ term) + (v : term), + wf_eprogram efl p -> + eval_eprogram wfl p v -> + exists v' : term, + let ip := inline_program inlining p in + eval_eprogram wfl ip v' /\ v' = inline ip v. + +Import Transform. + +Program Definition inline_transformation (efl : EEnvFlags) (wfl : WcbvFlags) inlining : + Transform.t _ _ EAst.term EAst.term _ _ + (eval_eprogram wfl) (eval_inlined_program wfl) := + {| name := "inlining "; + transform p _ := inline_program inlining p ; + pre p := wf_eprogram efl p ; + post (p : inlined_program) := wf_eprogram efl p ; + obseq p hp (p' : inlined_program) v v' := v' = inline p' v |}. + +Next Obligation. + now apply trust_inlining_wf. +Qed. +Next Obligation. + now eapply trust_inlining_pres. +Qed. + +#[global] +Axiom trust_inline_transformation_ext : + forall (efl : EEnvFlags) (wfl : WcbvFlags) inlining, + TransformExt.t (inline_transformation efl wfl inlining) + (fun p p' => extends p.1 p'.1) (fun p p' => extends p.1.1 p'.1.1). + +Definition extends_inlined_eprogram (p q : inlined_program) := + extends p.1.1 q.1.1 /\ p.2 = q.2. + +#[global] +Axiom trust_inline_transformation_ext' : + forall (efl : EEnvFlags) (wfl : WcbvFlags) inlining, + TransformExt.t (inline_transformation efl wfl inlining) + extends_eprogram extends_inlined_eprogram. + + +Program Definition forget_inlining_info_transformation (efl : EEnvFlags) (wfl : WcbvFlags) : + Transform.t _ _ EAst.term EAst.term _ _ + (eval_inlined_program wfl) (eval_eprogram wfl) := + {| name := "forgetting about inlining info"; + transform p _ := (p.1.1, p.2) ; + pre (p : inlined_program) := wf_eprogram efl p ; + post (p : eprogram) := wf_eprogram efl p ; + obseq p hp p' v v' := v' = v |}. + + Next Obligation. + destruct input as [[Σ inls] t]. + exact p. + Qed. + Next Obligation. + exists v; split => //. subst p'. + now destruct p as [[Σ inls] t]. + Qed. + +#[global] +Lemma forget_inlining_info_transformation_ext : + forall (efl : EEnvFlags) (wfl : WcbvFlags), + TransformExt.t (forget_inlining_info_transformation efl wfl) + (fun p p' => extends p.1.1 p'.1.1) (fun p p' => extends p.1 p'.1). +Proof. + intros. + red. now intros [[] ?] [[] ?]; cbn. +Qed. + +#[global] +Lemma forget_inlining_info_transformation_ext' : + forall (efl : EEnvFlags) (wfl : WcbvFlags), + TransformExt.t (forget_inlining_info_transformation efl wfl) + extends_inlined_eprogram extends_eprogram. +Proof. + intros ? ? [[] ?] [[] ?]; cbn. + now rewrite /extends_inlined_eprogram /extends_eprogram /=. +Qed. + diff --git a/erasure/theories/EProgram.v b/erasure/theories/EProgram.v index f069d6c25..925bc9f76 100644 --- a/erasure/theories/EProgram.v +++ b/erasure/theories/EProgram.v @@ -5,7 +5,8 @@ From MetaCoq.Utils Require Import utils. From MetaCoq.Common Require Import Transform config BasicAst. From MetaCoq.PCUIC Require PCUICAst PCUICAstUtils PCUICProgram. (* From MetaCoq.SafeChecker Require Import PCUICErrors PCUICWfEnvImpl. *) -From MetaCoq.Erasure Require EAstUtils EWellformed EEnvMap EWcbvEval. +From MetaCoq.Erasure Require EAstUtils EWellformed EEnvMap EGlobalEnv EWcbvEval. +Import EEnvMap. Import bytestring. Local Open Scope bs. @@ -95,3 +96,9 @@ Module TransformExt. End Comp. End TransformExt. + +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. diff --git a/erasure/theories/EUnboxing.v b/erasure/theories/EUnboxing.v index 4bd0d9a73..5cf20e76c 100644 --- a/erasure/theories/EUnboxing.v +++ b/erasure/theories/EUnboxing.v @@ -146,7 +146,6 @@ Section unbox. | view_other u nconstr => mkApps (unbox f) (map unbox v) end. Proof using Type. - Admitted. intros napp. rewrite unbox_equation_1. destruct (TermSpineView.view_mkApps (TermSpineView.view (mkApps f v)) napp nnil) as [hna [hv' ->]]. simp unbox; rewrite -unbox_equation_1. From 2cbbd67089305f50ea15f9f8f4bfc9d8ed66ab0a Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 29 Feb 2024 09:32:40 +0100 Subject: [PATCH 16/28] Implement a general Show typeclass in MetaCoq.Utils (#1063) * Implement a general Show typeclass in MetaCoq.Utils * Fix dependencies of new template-coq plugin --- common/theories/Primitive.v | 15 ----- erasure-plugin/theories/Erasure.v | 2 + erasure/theories/EBeta.v | 12 ++-- erasure/theories/EPretty.v | 4 +- pcuic/theories/utils/PCUICPretty.v | 14 ++-- template-coq/_PluginProject | 10 +-- template-coq/_PluginProject.in | 10 +-- .../gen-src/metacoq_template_plugin.mlpack | 4 +- utils/_CoqProject | 1 + utils/theories/MCUtils.v | 3 +- utils/theories/Show.v | 64 +++++++++++++++++++ 11 files changed, 99 insertions(+), 40 deletions(-) create mode 100644 utils/theories/Show.v diff --git a/common/theories/Primitive.v b/common/theories/Primitive.v index c3de3770a..f8283f6fc 100644 --- a/common/theories/Primitive.v +++ b/common/theories/Primitive.v @@ -9,18 +9,3 @@ Variant prim_tag := | primFloat | primArray. Derive NoConfusion EqDec for prim_tag. - -Definition string_of_prim_int (i:Uint63.int) : string := - (* Better? DecimalString.NilZero.string_of_uint (BinNat.N.to_uint (BinInt.Z.to_N (Int63.to_Z i))). ? *) - string_of_Z (Uint63.to_Z i). - -Definition string_of_float (f : PrimFloat.float) : string := - match Prim2SF f with - | S754_zero sign => if sign then "-0" else "0" - | S754_infinity sign => if sign then "-INFINITY" else "INFINITY" - | S754_nan => "NAN" - | S754_finite sign p z => - let abs := "0x" ++ bytestring.String.of_string (Numbers.HexadecimalString.NilZero.string_of_uint (Pos.to_hex_uint p)) ++ "p" ++ - bytestring.String.of_string (Numbers.DecimalString.NilZero.string_of_int (Z.to_int z)) - in if sign then "-" ++ abs else abs - end. \ No newline at end of file diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index f24a1d62c..261a273b2 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -109,6 +109,8 @@ Program Definition optional_unsafe_transforms econf := unbox_transformation efl final_wcbv_flags ▷ inline_transformation efl final_wcbv_flags econf.(inlining) ▷ forget_inlining_info_transformation efl final_wcbv_flags ▷ + (* Heuristically do it twice for more beta-normal terms *) + betared_transformation efl final_wcbv_flags ▷ betared_transformation efl final_wcbv_flags). Program Definition verified_lambdabox_pipeline {guard : abstract_guard_impl} diff --git a/erasure/theories/EBeta.v b/erasure/theories/EBeta.v index 75adc987d..a4b9cd1cf 100644 --- a/erasure/theories/EBeta.v +++ b/erasure/theories/EBeta.v @@ -19,17 +19,13 @@ Definition map_subterms (f : term -> term) (t : term) : term := | tLazy t => tLazy (f t) | tForce t => tForce (f t) | tConstruct ind n args => tConstruct ind n (map f args) - | t => t + | tRel n => tRel n + | tVar na => tVar na + | tConst kn => tConst kn + | tBox => tBox end. Section betared. - Fixpoint decompose_lam (t : term) {struct t} : list name × term := - match t with - | tLambda na B => - let (ns, B0) := decompose_lam B in - (na :: ns, B0) - | _ => ([], t) - end. Fixpoint beta_body (body : term) (args : list term) {struct args} : term := match args with diff --git a/erasure/theories/EPretty.v b/erasure/theories/EPretty.v index 97e6a2081..9c99f2e64 100644 --- a/erasure/theories/EPretty.v +++ b/erasure/theories/EPretty.v @@ -86,8 +86,8 @@ Module PrintTermTree. Definition print_prim (soft : EAst.term -> Tree.t) (p : @prim_val EAst.term) : Tree.t := match p.π2 return Tree.t with - | primIntModel f => "(int: " ^ Primitive.string_of_prim_int f ^ ")" - | primFloatModel f => "(float: " ^ Primitive.string_of_float f ^ ")" + | primIntModel f => "(int: " ^ show f ^ ")" + | primFloatModel f => "(float: " ^ show f ^ ")" | primArrayModel a => "(array:" ^ soft a.(array_default) ^ " , " ^ string_of_list soft a.(array_value) ^ ")" end. diff --git a/pcuic/theories/utils/PCUICPretty.v b/pcuic/theories/utils/PCUICPretty.v index 8a4ae2a42..94065d591 100644 --- a/pcuic/theories/utils/PCUICPretty.v +++ b/pcuic/theories/utils/PCUICPretty.v @@ -119,8 +119,8 @@ Module PrintTermTree. Definition print_prim (soft : term -> Tree.t) (p : prim_val) : Tree.t := match p.π2 return Tree.t with - | primIntModel f => "(int: " ^ Primitive.string_of_prim_int f ^ ")" - | primFloatModel f => "(float: " ^ Primitive.string_of_float f ^ ")" + | primIntModel f => "(int: " ^ show f ^ ")" + | primFloatModel f => "(float: " ^ show f ^ ")" | primArrayModel a => "(array:" ^ string_of_list soft a.(array_value) ^ ")" end. @@ -128,7 +128,7 @@ Module PrintTermTree. Context (print_term : list ident -> bool -> bool -> term -> t). Definition print_def {A} (f : A -> t) (g : A -> t) (def : def A) := - string_of_name (binder_name (dname def)) ^ " { struct " ^ string_of_nat (rarg def) ^ " }" ^ + string_of_name (binder_name (dname def)) ^ " { struct " ^ show (rarg def) ^ " }" ^ " : " ^ f (dtype def) ^ " := " ^ nl ^ g (dbody def). Definition print_defs Γ (defs : mfixpoint term) := @@ -344,5 +344,11 @@ Definition print_context Σ Γ Δ : string := Definition print_env (short : bool) (prefix : nat) Σ := Tree.to_string (PrintTermTree.print_env short prefix Σ). +#[export] Instance show_env : Show global_env := + fun Σ => print_env false (List.length Σ.(declarations)) Σ. + Definition print_program (short : bool) (prefix : nat) (p : program) : string := - Tree.to_string (PrintTermTree.print_program short prefix p). \ No newline at end of file + Tree.to_string (PrintTermTree.print_program short prefix p). + +#[export] Instance show_program : Show program := + fun p => print_program false (List.length p.1.(declarations)) p. diff --git a/template-coq/_PluginProject b/template-coq/_PluginProject index e8ff059f6..cb59a88eb 100644 --- a/template-coq/_PluginProject +++ b/template-coq/_PluginProject @@ -68,8 +68,6 @@ gen-src/decidableType.ml gen-src/decidableType.mli gen-src/decimal.ml gen-src/decimal.mli -gen-src/decimalString.ml -gen-src/decimalString.mli gen-src/denoter.ml # gen-src/depElim.ml # gen-src/depElim.mli @@ -103,8 +101,8 @@ gen-src/floatOps.ml gen-src/floatOps.mli gen-src/hexadecimal.ml gen-src/hexadecimal.mli -gen-src/hexadecimalString.ml -gen-src/hexadecimalString.mli +# gen-src/hexadecimalString.ml +# gen-src/hexadecimalString.mli gen-src/induction.ml gen-src/induction.mli # gen-src/init.ml @@ -143,6 +141,10 @@ gen-src/mCRelations.ml gen-src/mCRelations.mli gen-src/mCString.ml gen-src/mCString.mli +gen-src/sint63.mli +gen-src/sint63.ml +gen-src/show.mli +gen-src/show.ml # gen-src/mCUint63.ml # gen-src/mCUint63.mli gen-src/mCUtils.ml diff --git a/template-coq/_PluginProject.in b/template-coq/_PluginProject.in index 1532dfeba..110bb8575 100644 --- a/template-coq/_PluginProject.in +++ b/template-coq/_PluginProject.in @@ -66,8 +66,6 @@ gen-src/decidableType.ml gen-src/decidableType.mli gen-src/decimal.ml gen-src/decimal.mli -gen-src/decimalString.ml -gen-src/decimalString.mli gen-src/denoter.ml # gen-src/depElim.ml # gen-src/depElim.mli @@ -101,8 +99,8 @@ gen-src/floatOps.ml gen-src/floatOps.mli gen-src/hexadecimal.ml gen-src/hexadecimal.mli -gen-src/hexadecimalString.ml -gen-src/hexadecimalString.mli +# gen-src/hexadecimalString.ml +# gen-src/hexadecimalString.mli gen-src/induction.ml gen-src/induction.mli # gen-src/init.ml @@ -141,6 +139,10 @@ gen-src/mCRelations.ml gen-src/mCRelations.mli gen-src/mCString.ml gen-src/mCString.mli +gen-src/sint63.mli +gen-src/sint63.ml +gen-src/show.mli +gen-src/show.ml # gen-src/mCUint63.ml # gen-src/mCUint63.mli gen-src/mCUtils.ml diff --git a/template-coq/gen-src/metacoq_template_plugin.mlpack b/template-coq/gen-src/metacoq_template_plugin.mlpack index 9886d3375..aeeae114c 100644 --- a/template-coq/gen-src/metacoq_template_plugin.mlpack +++ b/template-coq/gen-src/metacoq_template_plugin.mlpack @@ -25,7 +25,6 @@ Caml_byte ByteCompare ByteCompareSpec String0 -HexadecimalString Orders OrdersAlt OrdersTac @@ -74,8 +73,9 @@ Zpower SpecFloat PrimFloat FloatOps -DecimalString MCString +Sint63 +Show MCUtils Signature All_Forall diff --git a/utils/_CoqProject b/utils/_CoqProject index 9f0b08c4e..600c73719 100644 --- a/utils/_CoqProject +++ b/utils/_CoqProject @@ -31,6 +31,7 @@ theories/MC_ExtrOCamlInt63.v theories/MC_ExtrOCamlZPosInt.v theories/ReflectEq.v theories/monad_utils.v +theories/Show.v theories/utils.v # extra tactics diff --git a/utils/theories/MCUtils.v b/utils/theories/MCUtils.v index d066b02e4..15dba02bb 100644 --- a/utils/theories/MCUtils.v +++ b/utils/theories/MCUtils.v @@ -1,6 +1,6 @@ From Coq Require Import Nat ZArith Bool. -Require Export MCPrelude +From MetaCoq.Utils Require Export MCPrelude MCReflect All_Forall MCArith @@ -23,6 +23,7 @@ Require Export MCPrelude MCTactics.UniquePose ReflectEq bytestring + Show . Tactic Notation "destruct" "?" := diff --git a/utils/theories/Show.v b/utils/theories/Show.v new file mode 100644 index 000000000..9f3f38704 --- /dev/null +++ b/utils/theories/Show.v @@ -0,0 +1,64 @@ +From Coq Require Import PArith Sint63 String Uint63 PrimFloat SpecFloat FloatOps. +From MetaCoq.Utils Require Import bytestring MCString. + +Local Open Scope bs_scope. + +Class Show (A : Type) := show : A -> string. +Global Hint Mode Show ! : typeclass_instances. + +#[export] Instance show_bytestring : Show string := fun x => x. + +#[export] Instance show_string : Show String.string := bytestring.String.of_string. + +Definition string_show {A} {show : Show A} : A -> String.string := + fun a => bytestring.String.to_string (show a). + +#[export] Instance nat_show : Show nat := string_of_nat. + +Definition string_of_bool b := if (b : bool) then "true" else "false". + +#[export] Instance show_bool : Show bool := string_of_bool. +#[export] Instance show_pair {A B} {showA : Show A} {showB : Show B}: Show (A * B) := + { show '(x, y) := "(" ++ show x ++ ", " ++ show y ++ ")" }. + +#[export] Instance show_sum {A B} {showA : Show A} {showB : Show B}: Show (A + B) := + { show x := match x with + | inl x => "(inl " ++ show x ++ ")" + | inr x => "(inr " ++ show x ++ ")" + end + }. + +Definition string_of_option {A} (fn : A -> string) (x : option A) : string := + match x with + | None => "None" + | Some x => "(Some " ++ fn x ++ ")" + end. + +#[export] Instance show_option {A} {showA : Show A}: Show (option A) := string_of_option show. + +#[export] Instance show_list {A} {SA : Show A} : Show (list A) := string_of_list show. + +Import SpecFloat. +Definition string_of_specfloat (f : SpecFloat.spec_float) := + match f with + | S754_zero sign => if sign then "-0" else "0" + | S754_infinity sign => if sign then "-infinity" else "infinity" + | S754_nan => "nan" + | S754_finite sign p z => + let num := string_of_positive p ++ "p" ++ string_of_Z z in + if sign then "-" ++ num else num + end. + +Definition string_of_prim_int (i:Uint63.int) : string := + (* Better? DecimalString.NilZero.string_of_uint (BinNat.N.to_uint (BinInt.Z.to_N (Int63.to_Z i))). ? *) + string_of_Z (Uint63.to_Z i). + +Definition string_of_float (f : PrimFloat.float) : string := + string_of_specfloat (FloatOps.Prim2SF f). + +#[export] Instance show_uint : Show PrimInt63.int := string_of_prim_int. +#[export] Instance show_sint : Show int_wrapper := { show x := string_of_Z (Sint63.to_Z (x.(int_wrap))) }. +#[export] Instance show_specfloat : Show SpecFloat.spec_float := string_of_specfloat. +#[export] Instance show_float : Show PrimFloat.float := string_of_float. +#[export] Instance show_positive : Show positive := string_of_positive. +#[export] Instance show_Z : Show Z := string_of_Z. From c8946aaf87ac540bb4493396e50ff78d52b7dd66 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Sat, 2 Mar 2024 00:01:49 +0100 Subject: [PATCH 17/28] WIP improving ewcbvevalnamed, needs to consider cofixes --- erasure/theories/EWcbvEvalNamed.v | 94 +++++++++++++++++++++++-------- 1 file changed, 71 insertions(+), 23 deletions(-) diff --git a/erasure/theories/EWcbvEvalNamed.v b/erasure/theories/EWcbvEvalNamed.v index ffef00c29..d6404be50 100644 --- a/erasure/theories/EWcbvEvalNamed.v +++ b/erasure/theories/EWcbvEvalNamed.v @@ -936,6 +936,10 @@ Fixpoint annotate (s : list ident) (u : term) {struct u} : term := let nms := gen_many_fresh s (map dname mfix) 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 + | tCoFix 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 (List.rev (gen_many_fresh s ((map dname mfix))) ++ s)) d) mfix nms in + tFix mfix' idx | tPrim p => tPrim (map_prim (annotate s) p) | tLazy t => tLazy (annotate s t) | tForce t => tForce (annotate s t) @@ -1038,58 +1042,102 @@ Qed. Local Hint Resolve incl_tl incl_appr incl_appl : core. -Lemma wellformed_annotate' Σ Γ Γ' s : +Definition switch_term_flags_to_named (efl : ETermFlags) := + {| + has_tBox := has_tBox; + has_tRel := false; + has_tVar := true; + 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 := has_tCase; + has_tProj := has_tProj; + has_tFix := has_tFix; + has_tCoFix := has_tCoFix; + has_tPrim := has_tPrim; + has_tLazy_Force := has_tLazy_Force + |}. + +Definition switch_env_flags_to_named (efl : EEnvFlags) := + {| + has_axioms := has_axioms; + has_cstr_params := has_cstr_params; + term_switches := switch_term_flags_to_named efl.(term_switches); + cstr_as_blocks := efl.(cstr_as_blocks) |}. + +Lemma wellformed_annotate' efl Σ Γ Γ' s : incl Γ' Γ -> - wellformed (efl := extraction_env_flags) Σ #|Γ| s -> wellformed (efl := named_extraction_env_flags) (annotate_env Γ' Σ) #|Γ| (annotate Γ s). + @wellformed efl Σ #|Γ| s -> + @wellformed (switch_env_flags_to_named efl) (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. + induction s in Γ, Hwf, Γ', Hincl |- * using EInduction.term_forall_list_ind; cbn in *; rtoProp; unshelve intuition eauto. + - eapply Nat.ltb_lt in H0. destruct nth_error eqn:Eq; eauto. eapply nth_error_None in Eq. lia. + - solve_all. - destruct n; cbn. 2: destruct in_dec. all: eapply (IHs (_ :: Γ)); cbn; eauto. - - split; eauto. destruct n; cbn. 2: destruct in_dec; cbn. + - destruct n; cbn. 2: destruct in_dec; cbn. all: eapply (IHs2 (_ :: Γ)); cbn; 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. + now erewrite lookup_annotate_env; eauto. + - destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. + erewrite lookup_annotate_env; eauto. cbn. + destruct nth_error 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. destruct nth_error as [ [] | ]; cbn in *; eauto. repeat split. len. solve_all. + destruct cstr_as_blocks; eauto. + rtoProp; intuition eauto. solve_all. destruct args => //. - destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. erewrite lookup_annotate_env; eauto. cbn. destruct nth_error as [ [] | ]; cbn in *; eauto. + - destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. + destruct nth_error as [ [] | ]; cbn in *; eauto. repeat split. eauto. solve_all. rewrite map_length. rewrite <- app_length. 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. - { 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 ? [[] ]. + - destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. + erewrite lookup_annotate_env; eauto. cbn. + destruct nth_error as [ [] | ]; cbn in *; eauto. + + - clear - H1. solve_all. + generalize (List.rev (gen_many_fresh Γ (map dname m)) ++ Γ). + generalize (gen_many_fresh Γ (map dname m)). + induction H1. + + econstructor. + + destruct x; cbn in *. destruct l0 => //=. cbn. + constructor; cbn. + destruct dbody; cbn in *; eauto. + eapply IHAll. + - 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 H2. 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. + revert H2. 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| + #|Γ|). + revert H2. generalize (#|m| + #|Γ|). intros. - induction m in Γ, n, n0, l, H1 |- *. - - econstructor. - - invs H1. cbn. destruct a; cbn. destruct dname; cbn; econstructor; eauto. - } + induction m in Γ, n, n0, l, H2 |- *. + + econstructor. + + invs H2. cbn. destruct a; cbn. destruct dname; cbn; econstructor; eauto. + - unfold wf_fix in *. rtoProp; intuition auto. solve_all. destruct x; cbn in *. + eapply (a #|fix_context. - repeat solve_all. + Qed. Lemma wellformed_annotate Σ Γ s : From e8e5ede7ac77b7ae5966d65b6d692cc0cfd6da9a Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 6 Mar 2024 10:11:08 +0100 Subject: [PATCH 18/28] Split unsafe passes, update EWcbvEvalNamed to be more parametric on flags --- Makefile | 4 +- erasure-plugin/src/g_metacoq_erasure.mlg | 4 +- erasure-plugin/theories/Erasure.v | 98 ++++++++++++++------ erasure/theories/EWcbvEvalNamed.v | 108 +++++++++++++++-------- 4 files changed, 150 insertions(+), 64 deletions(-) diff --git a/Makefile b/Makefile index a9889ce64..ce0b5b70d 100644 --- a/Makefile +++ b/Makefile @@ -12,7 +12,7 @@ ifeq '$(METACOQ_CONFIG)' 'local' export OCAMLPATH endif -.PHONY: printconf all utils template-coq pcuic erasure install uninstall html clean mrproper .merlin test-suite translations quotation +.PHONY: printconf all utils template-coq pcuic erasure install uninstall html clean mrproper safechecker-plugin .merlin test-suite translations quotation printconf: ifeq '$(METACOQ_CONFIG)' 'local' @@ -26,7 +26,7 @@ else endif endif -install: all +install: all $(MAKE) -C utils install $(MAKE) -C common install $(MAKE) -C template-coq install diff --git a/erasure-plugin/src/g_metacoq_erasure.mlg b/erasure-plugin/src/g_metacoq_erasure.mlg index bb3a14cb9..8b29da0e2 100644 --- a/erasure-plugin/src/g_metacoq_erasure.mlg +++ b/erasure-plugin/src/g_metacoq_erasure.mlg @@ -59,12 +59,12 @@ let default_config = let make_erasure_config config = let open Erasure0 in - { enable_unsafe = config.unsafe; + { enable_unsafe = if config.unsafe then all_unsafe_passes else no_unsafe_passes ; enable_typed_erasure = config.typed; enable_fast_remove_params = config.fast; dearging_config = default_dearging_config; inductives_mapping = []; - inlining = Kernames.KernameSet.empty } + inlined_constants = Kernames.KernameSet.empty } let time_opt config str fn arg = if config.time then diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 261a273b2..4b038a62a 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -30,13 +30,20 @@ Import EWcbvEval. Local Obligation Tactic := program_simpl. +Record unsafe_passes := + { fix_to_lazy : bool; + reorder_constructors : bool; + inlining : bool; + unboxing : bool; + betared : bool }. + Record erasure_configuration := { - enable_unsafe : bool; + enable_unsafe : unsafe_passes; enable_typed_erasure : bool; enable_fast_remove_params : bool; dearging_config : dearging_config; inductives_mapping : EReorderCstrs.inductives_mapping; - inlining : KernameSet.t + inlined_constants : KernameSet.t }. Definition default_dearging_config := @@ -44,23 +51,44 @@ Definition default_dearging_config := do_trim_const_masks := true; do_trim_ctor_masks := false |}. -(* This runs the cofix -> fix translation which is not entirely verified yet *) + +Definition make_unsafe_passes b := + {| fix_to_lazy := b; + reorder_constructors := b; + inlining := b; + unboxing := b; + betared := b |}. + +Definition no_unsafe_passes := make_unsafe_passes false. +Definition all_unsafe_passes := make_unsafe_passes true. + +(* This runs the cofix -> fix/lazy translation as well as inlining and + beta-redex simplification, which are not verified. It does not change + representation by reordering constructors or unboxing. *) + +Definition default_unsafe_passes := + {| fix_to_lazy := true; + reorder_constructors := false; + inlining := true; + unboxing := false; + betared := true |}. + Definition default_erasure_config := - {| enable_unsafe := true; + {| enable_unsafe := default_unsafe_passes; dearging_config := default_dearging_config; enable_typed_erasure := true; enable_fast_remove_params := true; inductives_mapping := []; - inlining := KernameSet.empty |}. + inlined_constants := KernameSet.empty |}. (* This runs only the verified phases without the typed erasure and "fast" remove params *) Definition safe_erasure_config := - {| enable_unsafe := false; + {| enable_unsafe := no_unsafe_passes; enable_typed_erasure := false; enable_fast_remove_params := false; dearging_config := default_dearging_config; inductives_mapping := []; - inlining := KernameSet.empty |}. + inlined_constants := KernameSet.empty |}. Axiom assume_welltyped_template_program_expansion : forall p (wtp : ∥ wt_template_program_env p ∥), @@ -96,22 +124,40 @@ Definition final_wcbv_flags := {| with_constructor_as_block := true |}. Program Definition optional_unsafe_transforms econf := + let passes := econf.(enable_unsafe) in let efl := EConstructorsAsBlocks.switch_cstr_as_blocks (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags)) in - ETransform.optional_self_transform econf.(enable_unsafe) + ETransform.optional_self_transform passes.(fix_to_lazy) ((* Rebuild the efficient lookup table *) rebuild_wf_env_transform (efl := efl) false false ▷ (* Coinductives & cofixpoints are translated to inductive types and thunked fixpoints *) coinductive_to_inductive_transformation efl - (has_app := eq_refl) (has_box := eq_refl) (has_rel := eq_refl) (has_pars := eq_refl) (has_cstrblocks := eq_refl) ▷ - reorder_cstrs_transformation efl final_wcbv_flags econf.(inductives_mapping) ▷ - rebuild_wf_env_transform (efl := efl) false false ▷ - unbox_transformation efl final_wcbv_flags ▷ - inline_transformation efl final_wcbv_flags econf.(inlining) ▷ - forget_inlining_info_transformation efl final_wcbv_flags ▷ - (* Heuristically do it twice for more beta-normal terms *) - betared_transformation efl final_wcbv_flags ▷ - betared_transformation efl final_wcbv_flags). + (has_app := eq_refl) (has_box := eq_refl) (has_rel := eq_refl) (has_pars := eq_refl) (has_cstrblocks := eq_refl)) ▷ + ETransform.optional_self_transform passes.(reorder_constructors) + (reorder_cstrs_transformation efl final_wcbv_flags econf.(inductives_mapping)) ▷ + ETransform.optional_self_transform passes.(unboxing) + (rebuild_wf_env_transform (efl := efl) false false ▷ + unbox_transformation efl final_wcbv_flags) ▷ + ETransform.optional_self_transform passes.(inlining) + (inline_transformation efl final_wcbv_flags econf.(inlined_constants) ▷ + forget_inlining_info_transformation efl final_wcbv_flags) ▷ + (* Heuristically do it twice for more beta-normal terms *) + ETransform.optional_self_transform passes.(betared) + (betared_transformation efl final_wcbv_flags ▷ + betared_transformation efl final_wcbv_flags). + +Next Obligation. + destruct (enable_unsafe econf) as [[] [] [] [] []]; cbn in * => //; intuition auto. +Qed. +Next Obligation. + destruct (enable_unsafe econf) as [[] [] [] [] []]; cbn in * => //; intuition auto. +Qed. +Next Obligation. + destruct (enable_unsafe econf) as [[] [] [] [] []]; cbn in * => //; intuition auto. +Qed. +Next Obligation. + destruct (enable_unsafe econf) as [[] [] [] [] []]; cbn in * => //; intuition auto. +Qed. Program Definition verified_lambdabox_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) @@ -266,7 +312,7 @@ Qed. Next Obligation. unfold optional_unsafe_transforms. cbn. - destruct enable_unsafe => //. + destruct enable_unsafe as [[] ? ? ? ?]=> //. Qed. Local Obligation Tactic := intros; eauto. @@ -365,7 +411,9 @@ Next Obligation. cbn in H. split; cbn; intuition eauto. Qed. Next Obligation. - cbn in H. unfold optional_unsafe_transforms. destruct enable_unsafe => //. + cbn in H. unfold optional_unsafe_transforms. + cbn. + destruct enable_unsafe as [[] ? ? ? ?]=> //. Qed. Next Obligation. cbn in H. split; cbn; intuition eauto. @@ -409,8 +457,8 @@ Program Definition run_erase_program {guard : abstract_guard_impl} econf := else run (erasure_pipeline ▷ (optional_unsafe_transforms econf)). Next Obligation. Proof. - unfold optional_unsafe_transforms. - destruct enable_unsafe => //. + unfold optional_unsafe_transforms; cbn. + destruct enable_unsafe as [[] ? ? ? ?]=> //. Qed. Program Definition erase_and_print_template_program econf (p : Ast.Env.program) : string := @@ -425,23 +473,23 @@ Next Obligation. Qed. Definition erasure_fast_config := - {| enable_unsafe := false; + {| enable_unsafe := no_unsafe_passes; dearging_config := default_dearging_config; enable_typed_erasure := false; enable_fast_remove_params := true; inductives_mapping := []; - inlining := KernameSet.empty |}. + inlined_constants := KernameSet.empty |}. Program Definition erase_fast_and_print_template_program (p : Ast.Env.program) : string := erase_and_print_template_program erasure_fast_config p. Definition typed_erasure_config := - {| enable_unsafe := false; + {| enable_unsafe := no_unsafe_passes; dearging_config := default_dearging_config; enable_typed_erasure := true; enable_fast_remove_params := true; inductives_mapping := []; - inlining := KernameSet.empty |}. + inlined_constants := KernameSet.empty |}. (* Parameterized by a configuration for dearging, allowing to, e.g., override masks. *) Program Definition typed_erase_and_print_template_program (p : Ast.Env.program) diff --git a/erasure/theories/EWcbvEvalNamed.v b/erasure/theories/EWcbvEvalNamed.v index d6404be50..217f85d12 100644 --- a/erasure/theories/EWcbvEvalNamed.v +++ b/erasure/theories/EWcbvEvalNamed.v @@ -939,7 +939,7 @@ Fixpoint annotate (s : list ident) (u : term) {struct u} : term := | tCoFix 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 (List.rev (gen_many_fresh s ((map dname mfix))) ++ s)) d) mfix nms in - tFix mfix' idx + tCoFix mfix' idx | tPrim p => tPrim (map_prim (annotate s) p) | tLazy t => tLazy (annotate s t) | tForce t => tForce (annotate s t) @@ -1134,39 +1134,67 @@ Proof. induction m in Γ, n, n0, l, H2 |- *. + econstructor. + invs H2. cbn. destruct a; cbn. destruct dname; cbn; econstructor; eauto. - - unfold wf_fix in *. rtoProp; intuition auto. solve_all. destruct x; cbn in *. - eapply (a #|fix_context. + - 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 ? [i]. + 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 i0. exact i0. + eapply incl_appr. eauto. + } + 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. - repeat solve_all. - + - solve_all_k 6. Qed. -Lemma wellformed_annotate Σ Γ s : - wellformed (efl := extraction_env_flags) Σ #|Γ| s -> wellformed (efl := named_extraction_env_flags) (annotate_env Γ Σ) #|Γ| (annotate Γ s). +Lemma wellformed_annotate Σ Γ s efl : + wellformed (efl := efl) Σ #|Γ| s -> + wellformed (efl := switch_env_flags_to_named efl) (annotate_env Γ Σ) #|Γ| (annotate Γ s). Proof. eapply wellformed_annotate'. firstorder. Qed. -Lemma nclosed_represents Σ Γ E s : - wellformed Σ #|Γ| s -> Γ ;;; E ⊩ annotate Γ s ~ s. +Lemma nclosed_represents efl Σ Γ E s : + ~~ has_tBox -> ~~ has_tVar -> ~~ has_tEvar -> ~~ has_tCoFix -> ~~ has_tProj -> + wellformed (efl := efl) Σ #|Γ| s -> Γ ;;; E ⊩ annotate Γ s ~ s. Proof. - intros Hwf. - induction s in Γ, Hwf |- * using EInduction.term_forall_list_ind; cbn in *; rtoProp; unshelve eauto. intuition. - - eapply Nat.ltb_lt in Hwf. destruct nth_error eqn:Eq; eauto. + intros nb nv nev ncof nproj Hwf. + induction s in Γ, Hwf |- * using EInduction.term_forall_list_ind; + cbn -[lookup_constructor lookup_inductive] in *; rtoProp; unshelve eauto. + - now rewrite Hwf in nb. + - eapply Nat.ltb_lt in H0. destruct nth_error eqn:Eq; eauto. eapply nth_error_None in Eq. lia. - - econstructor. solve_all. clear - H1. induction H1; econstructor; eauto. eapply p, p. - - econstructor. eauto. - { eapply All2_All2_Set. solve_all. eapply All_All2. eapply H0. intros [] []; cbn in *. + - now rewrite Hwf in nv. + - now rewrite H in nev. + - econstructor. destruct cstr_as_blocks. + destruct lookup_constructor as [[[mib oib] ctor]|] => //. cbn in *. + rtoProp. + { eapply All2_All2_Set. solve_all. } + destruct args => //. + - econstructor; eauto. + eapply All2_All2_Set; solve_all. + { eapply All_All2. eapply H1. cbn. intros [] []; cbn in *. len. now rewrite gen_many_fresh_length. } solve_all. - clear - Γ H0. induction H0; econstructor; eauto. + clear - Γ H1. induction H1; econstructor; eauto. rename x into br. exists (gen_many_fresh Γ br.1). cbn. split. + eapply All2_All2_Set. solve_all. now eapply All2_refl. + split. * eapply p. rewrite app_length gen_many_fresh_length. eapply p. * eapply NoDup_gen_many_fresh. + - now rewrite H in nproj. - eapply represents_tFix with (nms := gen_many_fresh Γ (map dname m)). - 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. + 1:{ solve_all. generalize (List.rev (gen_many_fresh Γ (map dname m)) ++ Γ). clear - H1. + induction H1 in Γ; cbn. econstructor. intros. destruct x, dname; cbn. all: econstructor. - cbn in *. destruct p, dbody; cbn in *; try congruence. - eapply IHAll. - cbn in *. destruct p, dbody; cbn in *; try congruence. @@ -1179,16 +1207,17 @@ Proof. - 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 ? [[] ]. + eapply All_impl in H2. 2:{ intros ? [[] ]. 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. + revert H2. generalize ((List.rev (gen_many_fresh Γ (map dname m)) ++ Γ)). - intros. induction H1 in Γ |- *. + intros. induction H2 in Γ |- *. - econstructor. - cbn. destruct x; cbn. destruct dname; cbn; econstructor; eauto. } + - now rewrite H in ncof. - constructor; solve_all. depelim H1; cbn; solve_all; try econstructor. destruct a; constructor; solve_all. eapply All2_All2_Set. solve_all. Qed. @@ -1798,7 +1827,10 @@ Proof. rewrite in_app_iff in H2; eauto. Qed. -Lemma eval_to_eval_named (Σ Σ' : global_context) E s t u : +Lemma eval_to_eval_named (efl : EEnvFlags) (Σ Σ' : global_context) E s t u : + ~~ has_cstr_params -> + has_tApp -> + ~~ has_tBox -> ~~ has_tVar -> ~~ has_tEvar -> ~~ has_tProj -> ~~ has_tCoFix -> 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)) => @@ -1812,21 +1844,21 @@ Lemma eval_to_eval_named (Σ Σ' : global_context) E s t u : wellformed Σ 0 s -> ∑ v, ⊩ v ~ t × eval Σ' E u v. Proof. - intros HwfΣ HΣ HΣind HE Hsunny Heval Hrep Hwell. + intros haspars hasapp nbox nvar nevar nproj ncof HwfΣ HΣ HΣind HE Hsunny Heval Hrep Hwell. revert u E Hrep Hsunny HE. pattern s, t. revert Heval. eapply eval_preserve_mkApps_ind with (Σ := Σ); cbn -[substl] in *; rtoProp. 1: reflexivity. - 1: eapply Qpreserves_wellformed with (efl := extraction_env_flags). + 1: eapply Qpreserves_wellformed with (efl := efl). 1: eauto. 1:{ intros. cbn in *. eapply eval_wellformed; cbn; eauto. } 1: eauto. - 1: eauto. + 1: eauto. 1:auto. all: try congruence. all: intros; rtoProp. all: repeat match reverse goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. - - cbn in i0. congruence. + - cbn in i0 => //. now rewrite i0 in nbox. - invs Hrep. + invs H3. + cbn in Hsunny. rtoProp. @@ -1858,7 +1890,7 @@ Proof. destruct g; cbn in *; try congruence. red in EEE. unfold wf_minductive in EEE. rtoProp. eapply andb_true_iff in EEE as [Hpars _]. - cbn in Hpars. eapply Nat.eqb_eq in Hpars. + cbn in Hpars. destruct has_cstr_params => //. eapply Nat.eqb_eq in Hpars. destruct nth_error eqn:EEE; cbn in *; try congruence. destruct (nth_error (EAst.ind_ctors o) c) eqn:EEEE; cbn in *; try congruence. now invs EE. @@ -1928,8 +1960,8 @@ Proof. edestruct s1 as (v' & IH1' & IH2'); eauto. destruct d; cbn -[substl] in *; subst. - eapply All2_Set_All2 in H13. - eapply All2_nth_error_Some_r in H13 as Hnth; eauto. + eapply All2_Set_All2 in H14. + eapply All2_nth_error_Some_r in H14 as Hnth; eauto. destruct Hnth as ([na' b] & Hb1 & [Hlam Hb2]); cbn -[substl] in *. destruct b; invs Hlam. @@ -1975,8 +2007,8 @@ Proof. } { 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. + { clear - H14. unfold fix_env, fix_subst. + eapply All2_length in H14 as Hlen. rewrite Hlen. clear Hlen. generalize #|mfix|. induction n; econstructor; eauto. } @@ -2008,6 +2040,7 @@ Proof. edestruct s0 as (v1 & Hv1 & Hv2); eauto. invs Hv1; destruct args using rev_ind; cbn in *; try congruence. all: match goal with [H : _ = mkApps _ _ |- _ ] => now rewrite mkApps_app in H end. + - now rewrite H in nproj. - invs Hrep. + invs H3. + assert (∑ decl' body', declared_constant Σ' c decl' × sunny [] body' × cst_body decl' = Some body' × [];;; [] ⊩ body' ~ body) as (decl' & Hdecl & body' & Hsunny' & e' & H'). @@ -2031,12 +2064,13 @@ Proof. } edestruct s0 as (v & Hv1 & Hv2). 1: eauto. eauto. econstructor. eexists. split. eauto. econstructor; eauto. + - now rewrite H in nproj. - invs Hrep. + invs H2. + cbn in Hsunny. rtoProp. eapply eval_to_value in ev as Hval. invs Hval. * depelim X. - ** destruct t1; cbn in *; try congruence. + ** destruct t1; cbn in *; try congruence. rtoProp. now rewrite H3 in ncof. ** now rewrite /= in H. * invs H4; cbn in *; eauto. * invs H1; cbn in *; eauto; try congruence. @@ -2085,7 +2119,8 @@ Proof. 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 *. + cbn in H1. unfold wf_minductive in H1. rtoProp. cbn in *. + destruct has_cstr_params => //. 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. @@ -2133,7 +2168,7 @@ Proof. eapply All2_map2_left. eapply All2_swap; eauto. eauto. symmetry. eapply All2_length; eauto. intros. - cbn in *. destruct H1 as (([? []] & ?) & ?). + cbn in *. destruct H5 as (([? []] & ?) & ?). rewrite app_nil_r in r. all: eauto. Unshelve. all: repeat econstructor. @@ -2249,7 +2284,10 @@ Proof. - solve_all. Qed. -Lemma eval_to_eval_named_full (Σ Σ' : global_context) s t : +Lemma eval_to_eval_named_full (efl : EEnvFlags) (Σ Σ' : global_context) s t : + ~~ has_cstr_params -> + has_tApp -> + ~~ has_tBox -> ~~ has_tVar -> ~~ has_tEvar -> ~~ has_tProj -> ~~ has_tCoFix -> 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)) => @@ -2262,5 +2300,5 @@ Lemma eval_to_eval_named_full (Σ Σ' : global_context) s t : Proof. intros. eapply eval_to_eval_named; eauto. eapply sunny_annotate. - eapply nclosed_represents. eauto. + eapply nclosed_represents; eauto. Qed. From ec34006a3cc785d16e2178bb91f48f27ccba1a56 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 7 Mar 2024 08:38:28 +0100 Subject: [PATCH 19/28] Fix typo --- erasure-plugin/src/g_metacoq_erasure.mlg | 2 +- erasure-plugin/theories/Erasure.v | 8 ++++---- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/erasure-plugin/src/g_metacoq_erasure.mlg b/erasure-plugin/src/g_metacoq_erasure.mlg index 8b29da0e2..3efd45e9e 100644 --- a/erasure-plugin/src/g_metacoq_erasure.mlg +++ b/erasure-plugin/src/g_metacoq_erasure.mlg @@ -105,7 +105,7 @@ To show this help message type:\n\ MetaCoq Erase -help.\n\n\ Valid options:\n\ -typed : Run the typed erasure pipeline including a dearging phase. By default we run the pipeline without this phase.\n\ --unsafe : Run also partially verified passes of the pipeline. This includes the cofix to fix translation.\n\ +-unsafe : Run also partially verified passes of the pipeline. This includes the cofix to lazy translation.\n\ -time : Time each compilation phase\n\ -bypass-qeds : Bypass the use of Qed and quote opaque proofs. Beware, this can result in large memory\n\ consumption due to reification of large proof terms.\n\ diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index 4b038a62a..b43908c2f 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -31,7 +31,7 @@ Import EWcbvEval. Local Obligation Tactic := program_simpl. Record unsafe_passes := - { fix_to_lazy : bool; + { cofix_to_lazy : bool; reorder_constructors : bool; inlining : bool; unboxing : bool; @@ -53,7 +53,7 @@ Definition default_dearging_config := Definition make_unsafe_passes b := - {| fix_to_lazy := b; + {| cofix_to_lazy := b; reorder_constructors := b; inlining := b; unboxing := b; @@ -67,7 +67,7 @@ Definition all_unsafe_passes := make_unsafe_passes true. representation by reordering constructors or unboxing. *) Definition default_unsafe_passes := - {| fix_to_lazy := true; + {| cofix_to_lazy := true; reorder_constructors := false; inlining := true; unboxing := false; @@ -127,7 +127,7 @@ Program Definition optional_unsafe_transforms econf := let passes := econf.(enable_unsafe) in let efl := EConstructorsAsBlocks.switch_cstr_as_blocks (EInlineProjections.disable_projections_env_flag (ERemoveParams.switch_no_params EWellformed.all_env_flags)) in - ETransform.optional_self_transform passes.(fix_to_lazy) + ETransform.optional_self_transform passes.(cofix_to_lazy) ((* Rebuild the efficient lookup table *) rebuild_wf_env_transform (efl := efl) false false ▷ (* Coinductives & cofixpoints are translated to inductive types and thunked fixpoints *) From 690e443776a0ad1acaf39de5e065056d7711dfc7 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 8 Mar 2024 09:51:00 +0100 Subject: [PATCH 20/28] Fix inlining and reorder constructors which were not translating let bindings... oops! --- erasure/theories/EInlining.v | 2 +- erasure/theories/EReorderCstrs.v | 2 +- make-opam-files.sh | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/erasure/theories/EInlining.v b/erasure/theories/EInlining.v index 529e5f9ff..caff8e359 100644 --- a/erasure/theories/EInlining.v +++ b/erasure/theories/EInlining.v @@ -21,7 +21,7 @@ Section Inline. Equations inline (t : term) : term := | tVar na => tVar na | tLambda nm bod => tLambda nm (inline bod) - | tLetIn nm dfn bod => tLetIn nm dfn bod + | tLetIn nm dfn bod => tLetIn nm (inline dfn) (inline bod) | tApp fn arg => tApp (inline fn) (inline arg) | tConst nm with KernameMap.find nm inlining := { | Some body := (* Already inlined body *) body diff --git a/erasure/theories/EReorderCstrs.v b/erasure/theories/EReorderCstrs.v index 071d6b57a..3e029f23e 100644 --- a/erasure/theories/EReorderCstrs.v +++ b/erasure/theories/EReorderCstrs.v @@ -48,7 +48,7 @@ Section Reorder. Equations reorder (t : term) : term := | tVar na => tVar na | tLambda nm bod => tLambda nm (reorder bod) - | tLetIn nm dfn bod => tLetIn nm dfn bod + | tLetIn nm dfn bod => tLetIn nm (reorder dfn) (reorder bod) | tApp fn arg => tApp (reorder fn) (reorder arg) | tConst nm => tConst nm | tConstruct i m args => tConstruct i (lookup_constructor_ordinal i m) (map reorder args) diff --git a/make-opam-files.sh b/make-opam-files.sh index 1de13c702..2d88765ce 100755 --- a/make-opam-files.sh +++ b/make-opam-files.sh @@ -21,7 +21,7 @@ do opamf=${f/.opam/}; target=$1/$opamf/$opamf.$2/opam; echo $opamf; - mkdir $1/$opamf/$opamf.$2 + mkdir -p $1/$opamf/$opamf.$2 gsed -e "/^version:.*/d" $f > $target echo url { >> $target echo " src:" \"$3\" >> $target From 8a42c5bf6ab163a3032555e3a4a6471ee49e69fc Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 15 Mar 2024 17:24:44 +0100 Subject: [PATCH 21/28] Fix issue #1042 MetaCoq Run does not support evars. --- template-coq/src/g_template_coq.mlg | 19 ++++++++--------- template-coq/src/run_template_monad.ml | 20 +++++++++--------- test-suite/issue1042.v | 28 ++++++++++++++++++++++++++ 3 files changed, 48 insertions(+), 19 deletions(-) create mode 100644 test-suite/issue1042.v diff --git a/template-coq/src/g_template_coq.mlg b/template-coq/src/g_template_coq.mlg index 7639ed088..87d9acf3f 100644 --- a/template-coq/src/g_template_coq.mlg +++ b/template-coq/src/g_template_coq.mlg @@ -51,14 +51,14 @@ let to_ltac_val c = Tacinterp.Value.of_constr c let run_template_program ~pm env evm ~poly pgm = Run_template_monad.run_template_program_rec ~poly (fun ~st _ _ _ -> st) ~st:pm env (evm, pgm) -let fresh_env () = +let fresh_env () = let env = Global.env () in let sigma = Evd.from_env env in env, sigma let to_constr_evars sigma c = EConstr.to_constr ~abort_on_undefined_evars:false sigma c } - + (** ********* Commands ********* *) VERNAC COMMAND EXTEND TemplateCoq_Test_Quote CLASSIFIED AS QUERY STATE program @@ -76,7 +76,7 @@ VERNAC COMMAND EXTEND TemplateCoq_Quote_Definition CLASSIFIED AS SIDEFF STATE pr { let (env, evm) = fresh_env () in let (evm, def) = Constrintern.interp_open_constr env evm def in let (evm, pgm) = EConstr.fresh_global env evm (Lazy.force Template_monad.ptmQuoteDefinition) in - let pgm = Constr.mkApp (EConstr.to_constr evm pgm, [|Constr_quoter.quote_ident name; Constr.mkRel 0; + let pgm = Constr.mkApp (EConstr.to_constr evm pgm, [|Constr_quoter.quote_ident name; Constr.mkRel 0; to_constr_evars evm def|]) in run_template_program env evm ~poly pgm } END @@ -98,7 +98,7 @@ VERNAC COMMAND EXTEND TemplateCoq_Quote_Recursively_Definition CLASSIFIED AS SID { let (env, evm) = fresh_env () in let (evm, def) = Constrintern.interp_open_constr env evm def in let (evm, pgm) = EConstr.fresh_global env evm (Lazy.force Template_monad.ptmQuoteRecDefinition) in - let pgm = Constr.mkApp (EConstr.to_constr evm pgm, [|Constr_quoter.quote_ident name; Constr.mkRel 0; + let pgm = Constr.mkApp (EConstr.to_constr evm pgm, [|Constr_quoter.quote_ident name; Constr.mkRel 0; to_constr_evars evm def|]) in run_template_program env evm ~poly pgm } END @@ -108,7 +108,7 @@ VERNAC COMMAND EXTEND TemplateCoq_Test_Unquote CLASSIFIED AS QUERY STATE program { let (env, evm) = fresh_env () in let (evm, def) = Constrintern.interp_open_constr env evm def in let (evm, pgm) = EConstr.fresh_global env evm (Lazy.force Template_monad.ptmTestUnquote) in - let pgm = Constr.mkApp (EConstr.to_constr evm pgm, + let pgm = Constr.mkApp (EConstr.to_constr evm pgm, [|to_constr_evars evm def|]) in run_template_program env evm ~poly pgm } END @@ -119,7 +119,7 @@ VERNAC COMMAND EXTEND TemplateCoq_Make_Definition CLASSIFIED AS SIDEFF STATE pro let (evm, def) = Constrintern.interp_open_constr env evm def in let (evm, pgm) = EConstr.fresh_global env evm (Lazy.force Template_monad.ptmMkDefinition) in let pgm = Constr.mkApp (EConstr.to_constr evm pgm, - [|Constr_quoter.quote_ident name; + [|Constr_quoter.quote_ident name; to_constr_evars evm def|]) in run_template_program env evm ~poly pgm } END @@ -129,7 +129,7 @@ VERNAC COMMAND EXTEND TemplateCoq_Make_Inductive CLASSIFIED AS SIDEFF STATE prog { let (env, evm) = fresh_env () in let (evm, def) = Constrintern.interp_open_constr env evm def in let (evm, pgm) = EConstr.fresh_global env evm (Lazy.force Template_monad.ptmMkInductive) in - let pgm = Constr.mkApp (EConstr.to_constr evm pgm, + let pgm = Constr.mkApp (EConstr.to_constr evm pgm, [|Constr_quoter.quote_bool false; to_constr_evars evm def|]) in run_template_program env evm ~poly pgm } END @@ -137,8 +137,9 @@ END VERNAC COMMAND EXTEND TemplateCoq_Run_Template_Program CLASSIFIED AS SIDEFF STATE program | #[ poly = polymorphic ] [ "MetaCoq" "Run" constr(def) ] -> { let (env, evm) = fresh_env () in - let (evm, def) = Constrintern.interp_open_constr env evm def in - let pgm = to_constr_evars evm def in + let (pgm, ctx) = Constrintern.interp_constr env evm def in + let evm = Evd.from_ctx ctx in + let pgm = EConstr.to_constr ~abort_on_undefined_evars:true evm pgm in run_template_program env evm ~poly pgm } END diff --git a/template-coq/src/run_template_monad.ml b/template-coq/src/run_template_monad.ml index 189cc834e..05d1629ef 100644 --- a/template-coq/src/run_template_monad.ml +++ b/template-coq/src/run_template_monad.ml @@ -70,12 +70,12 @@ let rec unquote_pos trm : int = let (h,args) = app_full trm [] in match args with [x] -> - if constr_equall h cposI then + if constr_equall h cposI then (2 * unquote_pos x + 1) else if constr_equall h cposO then (2 * unquote_pos x) else not_supported_verb trm "unquote_pos" - | [] -> + | [] -> if constr_equall h cposzero then 1 else not_supported_verb trm "unquote_pos" | _ -> bad_term_verb trm "unquote_pos" @@ -87,7 +87,7 @@ let unquote_Z trm : int = if constr_equall h cZpos then unquote_pos x else if constr_equall h cZneg then - unquote_pos x else not_supported_verb trm "unquote_pos" - | [] -> + | [] -> if constr_equall h cZ0 then 0 else not_supported_verb trm "unquote_pos" | _ -> bad_term_verb trm "unquote_pos" @@ -96,12 +96,12 @@ let unquote_constraint_type trm (* of type constraint_type *) : constraint_type let (h,args) = app_full trm [] in match args with [x] -> - if constr_equall h tunivLe then + if constr_equall h tunivLe then let n = unquote_Z x in if n = 0 then Univ.Le else Univ.Lt else not_supported_verb trm "unquote_constraint_type" - | [] -> + | [] -> if constr_equall h tunivEq then Univ.Eq else not_supported_verb trm "unquote_constraint_type" | _ -> bad_term_verb trm "unquote_constraint_type" @@ -176,7 +176,7 @@ let denote_variance trm (* of type Variance *) : Variance.t = else if constr_equall trm cInvariant then Variance.Invariant else not_supported_verb trm "denote_variance" - + let denote_variance evm trm (* of type Variance.t list *) : _ * Variance.t array = let variances = List.map denote_variance (unquote_list trm) in evm, Array.of_list variances @@ -243,9 +243,9 @@ let unquote_one_inductive_entry env evm trm (* of type one_inductive_entry *) : let map_option f o = match o with | Some x -> Some (f x) - | None -> None + | None -> None -let denote_decl env evm d = +let denote_decl env evm d = let (h, args) = app_full d [] in if constr_equall h tmkdecl then match args with @@ -262,7 +262,7 @@ let denote_decl env evm d = let denote_context env evm ctx = fold_env_evm_right denote_decl env evm (unquote_list ctx) - + let unquote_mutual_inductive_entry env evm trm (* of type mutual_inductive_entry *) : _ * _ * Entries.mutual_inductive_entry = let (h,args) = app_full trm [] in if constr_equall h tBuild_mutual_inductive_entry then @@ -301,7 +301,7 @@ let declare_inductive (env: Environ.env) (evm: Evd.evar_map) (infer_univs : bool let evm' = Evd.from_env env in let evm', ctx, mind = unquote_mutual_inductive_entry env evm' mind in let () = DeclareUctx.declare_universe_context ~poly:false ctx in - let evm, mind = + let evm, mind = if infer_univs then let ctx, mind = Tm_util.RetypeMindEntry.infer_mentry_univs env evm' mind in debug (fun () -> Pp.(str "Declaring universe context " ++ Univ.pr_universe_context_set (Level.pr) ctx)); diff --git a/test-suite/issue1042.v b/test-suite/issue1042.v new file mode 100644 index 000000000..819c11f0b --- /dev/null +++ b/test-suite/issue1042.v @@ -0,0 +1,28 @@ +From MetaCoq.Utils Require Import utils. +From MetaCoq.Template Require Import All. +Import MCMonadNotation. + + +(*Exemple with a typing error*) +Definition ident_term (a : term) : term := a. + +Definition quote_inductive {X}(inductive:X): TemplateMonad _ := + quote <- tmQuote inductive;; + tmReturn quote. + +Inductive tm : Set := . + +Definition d1 : TemplateMonad unit. +(* Set Debug "backtrace". *) +Fail MetaCoq Run( + quote <- quote_inductive tm;; + constructor <- ident_term quote;; + tmPrint constructor) +. +Fail run_template_program (quote <- quote_inductive tm;; +constructor <- ident_term quote;; +tmPrint constructor) ltac:(fun x => idtac). + Fail refine ( + quote <- quote_inductive tm;; + constructor <- ident_term quote;; + tmPrint constructor). \ No newline at end of file From 05cffdfa5ae5c16c914b4720852cd2b81db08b60 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 18 Mar 2024 10:11:56 +0100 Subject: [PATCH 22/28] Fix deprecation and unnecessary module --- template-coq/gen-src/metacoq_template_plugin.mlpack | 1 - template-coq/src/quoter.ml | 2 +- 2 files changed, 1 insertion(+), 2 deletions(-) diff --git a/template-coq/gen-src/metacoq_template_plugin.mlpack b/template-coq/gen-src/metacoq_template_plugin.mlpack index eb0d5c506..214a395fb 100644 --- a/template-coq/gen-src/metacoq_template_plugin.mlpack +++ b/template-coq/gen-src/metacoq_template_plugin.mlpack @@ -118,7 +118,6 @@ TemplateProgram Wf NoConfusion -MCUint63 Init FloatClass DepElim diff --git a/template-coq/src/quoter.ml b/template-coq/src/quoter.ml index aa3db4dc5..752714349 100644 --- a/template-coq/src/quoter.ml +++ b/template-coq/src/quoter.ml @@ -264,7 +264,7 @@ struct let cofixpoint_arities ts = let cofix_arity t = - let ctx, _concl = Term.decompose_prod_assum t in + let ctx, _concl = Term.decompose_prod_decls t in Context.Rel.nhyps ctx in Array.map cofix_arity ts From 96b80c94f753e3bc92a8f932e01208017bb58594 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 18 Mar 2024 10:33:50 +0100 Subject: [PATCH 23/28] Fix script due to misclassification of implicit arg as potential typeclass instance --- erasure-plugin/theories/ErasureCorrectness.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/erasure-plugin/theories/ErasureCorrectness.v b/erasure-plugin/theories/ErasureCorrectness.v index ceb114063..cbb1c577a 100644 --- a/erasure-plugin/theories/ErasureCorrectness.v +++ b/erasure-plugin/theories/ErasureCorrectness.v @@ -782,11 +782,11 @@ Section ErasureFunction. set (eqdecls := (fun Σ => _)) at 9. clearbody eqdecls. set (deps := term_global_deps _). set (nin := (fun (n : nat) => _)). clearbody nin. - epose proof (@erase_global_deps_fast_erase_global_deps deps optimized_abstract_env_impl Σ' (PCUICAst.PCUICEnvironment.declarations Σ) nin) as [nin2 eq]. + epose proof (@erase_global_deps_fast_erase_global_deps deps optimized_abstract_env_impl Σ' (PCUICAst.PCUICEnvironment.declarations Σ) nin _ _) as [nin2 eq]. rewrite /erase_global_fast. erewrite eq. clear eq nin. set (eg := erase_global_deps _ _ _ _ _ _). - unshelve epose proof (erase_correct optimized_abstract_env_impl Σ' Σ.2 _ f v _ _ deps _ _ _ eq_refl _ eq_refl _ Σ eq_refl); eauto. + unshelve epose proof (erase_correct optimized_abstract_env_impl Σ' Σ.2 _ f v _ _ deps _ eqdecls _ eq_refl _ eq_refl _ Σ eq_refl); eauto. { eapply Kernames.KernameSet.subset_spec. rewrite /deps -/env'. cbn [fst snd]. apply Kernames.KernameSetProp.subset_refl. } { cbn => ? -> //. } destruct H as [v'' [ervv'' [ev]]]. From 9af0b0fc7b8c3a2989aae8a11e6a6b24695128ef Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 18 Mar 2024 10:35:34 +0100 Subject: [PATCH 24/28] Fix after merge --- template-coq/_PluginProject | 4 ++-- template-coq/_PluginProject.in | 11 ++--------- 2 files changed, 4 insertions(+), 11 deletions(-) diff --git a/template-coq/_PluginProject b/template-coq/_PluginProject index cc5f3c211..314e020ac 100644 --- a/template-coq/_PluginProject +++ b/template-coq/_PluginProject @@ -105,8 +105,8 @@ gen-src/hexadecimal.ml gen-src/hexadecimal.mli # gen-src/hexadecimalString.ml # gen-src/hexadecimalString.mli -gen-src/induction.ml -gen-src/induction.mli +gen-src/induction0.ml +gen-src/induction0.mli gen-src/init.ml gen-src/init.mli gen-src/int0.ml diff --git a/template-coq/_PluginProject.in b/template-coq/_PluginProject.in index db0886636..dcecb5541 100644 --- a/template-coq/_PluginProject.in +++ b/template-coq/_PluginProject.in @@ -101,17 +101,10 @@ gen-src/floatOps.ml gen-src/floatOps.mli gen-src/hexadecimal.ml gen-src/hexadecimal.mli -<<<<<<< HEAD -gen-src/hexadecimalString.ml -gen-src/hexadecimalString.mli -gen-src/induction0.ml -gen-src/induction0.mli -======= # gen-src/hexadecimalString.ml # gen-src/hexadecimalString.mli -gen-src/induction.ml -gen-src/induction.mli ->>>>>>> coq-8.18 +gen-src/induction0.ml +gen-src/induction0.mli gen-src/init.ml gen-src/init.mli gen-src/int0.ml From f55eb7c21ddfefde4cb08116232428e38ed2ee33 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Tue, 19 Mar 2024 10:24:20 +0100 Subject: [PATCH 25/28] Update make_opam_files to upload the archive to the releases folder on github to avoid issues with github rebuilding the tarballs --- make-opam-files.sh | 21 +++++++++++++++++++-- 1 file changed, 19 insertions(+), 2 deletions(-) diff --git a/make-opam-files.sh b/make-opam-files.sh index 2d88765ce..f448ac508 100755 --- a/make-opam-files.sh +++ b/make-opam-files.sh @@ -6,16 +6,33 @@ then exit 0 fi +archive=`basename $3` +tag=${archive/.tar.gz/} + echo "Target directory: " $1 echo "Target version: " $2 echo "Releases package: " $3 +echo "Archive:" $archive +echo "Tag:" $tag + +if [ -f $archive ] +then + echo "Removing existing archive!" + rm $archive +fi wget $3 -archive=`basename $3` + hash=`shasum -a 512 $archive | cut -f 1 -d " "` echo "Shasum = " $hash +echo "Uploading to release assets" + +gh release upload $tag $archive + +release=https://github.com/MetaCoq/metacoq/releases/download/$tag/$archive + for f in *.opam; do opamf=${f/.opam/}; @@ -24,7 +41,7 @@ do mkdir -p $1/$opamf/$opamf.$2 gsed -e "/^version:.*/d" $f > $target echo url { >> $target - echo " src:" \"$3\" >> $target + echo " src:" \"$release\" >> $target echo " checksum:" \"sha512=$hash\" >> $target echo } >> $target done \ No newline at end of file From a4cd12c5726818ed1f155ab79904ef0de5f9a815 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 6 Jun 2024 18:38:53 +0200 Subject: [PATCH 26/28] Add the force (lazy t) evaluation rule to lambda-box --- erasure-plugin/theories/ETransform.v | 2 +- erasure/theories/EAstUtils.v | 11 ++++ erasure/theories/EConstructorsAsBlocks.v | 36 +++++++++++- erasure/theories/EDeps.v | 1 + erasure/theories/EEtaExpandedFix.v | 5 +- erasure/theories/EImplementBox.v | 13 +++++ erasure/theories/EInlineProjections.v | 7 ++- erasure/theories/EOptimizePropDiscr.v | 6 +- erasure/theories/ERemoveParams.v | 14 ++++- erasure/theories/EWcbvEval.v | 33 ++++++----- .../EWcbvEvalCstrsAsBlocksFixLambdaInd.v | 22 ++++++-- erasure/theories/EWcbvEvalCstrsAsBlocksInd.v | 24 ++++++-- erasure/theories/EWcbvEvalEtaInd.v | 51 +++++++++++------ erasure/theories/EWcbvEvalInd.v | 9 ++- erasure/theories/EWcbvEvalNamed.v | 55 +++++++++++++++---- erasure/theories/ErasureCorrectness.v | 12 ++++ erasure/theories/Typed/OptimizeCorrectness.v | 20 ++++--- erasure/theories/Typed/WcbvEvalAux.v | 1 + 18 files changed, 249 insertions(+), 73 deletions(-) diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index 3af3e47a8..06a08372a 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -976,7 +976,7 @@ Program Definition coinductive_to_inductive_transformation (efl : EEnvFlags) {has_cstrblocks : cstr_as_blocks = true} : Transform.t _ _ EAst.term EAst.term _ _ (eval_eprogram_env block_wcbv_flags) (eval_eprogram block_wcbv_flags) := - {| name := "transforming co-inductive to inductive types"; + {| name := "transforming co-inductive to lazy inductive types"; transform p _ := ECoInductiveToInductive.trans_program p ; pre p := wf_eprogram_env efl p ; post p := wf_eprogram efl p ; diff --git a/erasure/theories/EAstUtils.v b/erasure/theories/EAstUtils.v index de7c2655c..4f3403da3 100644 --- a/erasure/theories/EAstUtils.v +++ b/erasure/theories/EAstUtils.v @@ -318,9 +318,16 @@ Definition is_box c := | _ => false end. +Definition isLazy c := + match c with + | tLazy _ => true + | _ => false + end. + Definition isFixApp t := isFix (head t). Definition isConstructApp t := isConstruct (head t). Definition isPrimApp t := isPrim (head t). +Definition isLazyApp t := isLazy (head t). Lemma isFixApp_mkApps f l : isFixApp (mkApps f l) = isFixApp f. Proof. rewrite /isFixApp head_mkApps //. Qed. @@ -328,6 +335,8 @@ Lemma isConstructApp_mkApps f l : isConstructApp (mkApps f l) = isConstructApp f Proof. rewrite /isConstructApp head_mkApps //. Qed. Lemma isPrimApp_mkApps f l : isPrimApp (mkApps f l) = isPrimApp f. Proof. rewrite /isPrimApp head_mkApps //. Qed. +Lemma isLazyApp_mkApps f l : isLazyApp (mkApps f l) = isLazyApp f. +Proof. rewrite /isLazyApp head_mkApps //. Qed. Lemma is_box_mkApps f a : is_box (mkApps f a) = is_box f. Proof. @@ -347,6 +356,8 @@ Lemma nisBox_mkApps f args : ~~ isBox f -> ~~ isBox (mkApps f args). Proof. destruct args using rev_case => //. rewrite mkApps_app /= //. Qed. Lemma nisPrim_mkApps f args : ~~ isPrim f -> ~~ isPrim (mkApps f args). Proof. destruct args using rev_case => //. rewrite mkApps_app /= //. Qed. +Lemma nisLazy_mkApps f args : ~~ isLazy f -> ~~ isLazy (mkApps f args). +Proof. destruct args using rev_case => //. rewrite mkApps_app /= //. Qed. Definition string_of_def {A : Set} (f : A -> string) (def : def A) := "(" ^ string_of_name (dname def) ^ "," ^ f (dbody def) ^ "," diff --git a/erasure/theories/EConstructorsAsBlocks.v b/erasure/theories/EConstructorsAsBlocks.v index 0032bef87..c4453bbae 100644 --- a/erasure/theories/EConstructorsAsBlocks.v +++ b/erasure/theories/EConstructorsAsBlocks.v @@ -649,6 +649,28 @@ Proof. eapply IHt1. cbn in Hwf'. rtoProp. intuition. Qed. +Lemma transform_blocks_isLazyApp {efl : EEnvFlags} {Σ : GlobalContextMap.t} t : + has_cstr_params = false -> + wf_glob Σ -> wellformed Σ 0 t -> + isLazyApp (transform_blocks Σ t) = isLazyApp t. +Proof. + intros haspars Hwf Hwf'. + induction t; try now cbn; eauto. + eapply transform_blocks_tApp; eauto. + destruct decompose_app. + destruct construct_viewc. + - rewrite GlobalContextMap.lookup_constructor_pars_args_spec. + destruct lookup_constructor_pars_args as [ [[]] | ]; eauto. + cbn. destruct chop. intros (? & ? & ?). subst. + rewrite -[tApp _ _](mkApps_app _ _ [t2]). + rewrite !isLazyApp_mkApps. cbn. reflexivity. + - change (tApp t1 t2) with (mkApps t1 [t2]). + change (tApp (transform_blocks Σ t1) (transform_blocks Σ t2)) with + (mkApps (transform_blocks Σ t1) [transform_blocks Σ t2]). + rewrite !isLazyApp_mkApps. + eapply IHt1. cbn in Hwf'. rtoProp. intuition. +Qed. + Lemma lookup_env_transform_blocks {Σ : GlobalContextMap.t} kn : lookup_env (transform_blocks_env Σ) kn = option_map (transform_blocks_decl Σ) (lookup_env Σ kn). @@ -1125,17 +1147,20 @@ Proof. * rewrite GlobalContextMap.lookup_constructor_pars_args_spec; destruct lookup_constructor_pars_args as [ [[]] | ] eqn:hpa; eauto. cbn [plus]. destruct chop eqn:heqch. - intros [hl [ht ha]]. rewrite ht in H1. rewrite isConstructApp_mkApps isPrimApp_mkApps orb_true_r in H1 => //. + intros [hl [ht ha]]. rewrite ht in H1. rewrite isConstructApp_mkApps + isPrimApp_mkApps isLazyApp_mkApps orb_true_r in H1 => //. * eapply eval_app_cong; eauto. revert H1. destruct f'; try now cbn; tauto. intros H. cbn in H. rewrite transform_blocks_isConstructApp; eauto. rewrite transform_blocks_isPrimApp; eauto. - rewrite negb_or in H. move/andP: H => [] ncstr nprim. + rewrite transform_blocks_isLazyApp; eauto. + rewrite !negb_or in H. move/andP: H => [] /andP [] ncstr nprim nlazy. destruct (isConstructApp (tApp f'1 f'2)) eqn:heq'. -- cbn in ncstr. congruence. - -- eapply transform_blocks_tApp; eauto. clear -nprim. + -- eapply transform_blocks_tApp; eauto. clear -nprim nlazy. + move/negbTE: nlazy ->. move/negbTE: nprim -> => /=. destruct decompose_app. destruct construct_viewc; try now cbn; eauto. rewrite GlobalContextMap.lookup_constructor_pars_args_spec; @@ -1186,6 +1211,11 @@ Proof. eapply All2_over_undep in a. eapply All2_Set_All2 in ev. eapply All2_All2_Set. solve_all. now destruct b. now destruct a0. + - intros evl evt' [evt wft wflt etat etalt]. + intros [evt'' wft' wfv etat' etav]. + simp transform_blocks; rewrite -!transform_blocks_equation_1. + econstructor; eauto. + now simp transform_blocks in evt; rewrite -!transform_blocks_equation_1 in evt. - intros. destruct t; try solve [constructor; cbn in H, H0 |- *; try congruence]. cbn -[lookup_constructor] in H |- *. destruct args => //. destruct lookup_constructor eqn:hl => //. diff --git a/erasure/theories/EDeps.v b/erasure/theories/EDeps.v index b1180b5cc..7bee22452 100644 --- a/erasure/theories/EDeps.v +++ b/erasure/theories/EDeps.v @@ -354,6 +354,7 @@ Proof. - depelim er; depelim X; constructor; eauto. eapply All2_over_undep in a0. solve_all. - easy. + - easy. Qed. #[global] diff --git a/erasure/theories/EEtaExpandedFix.v b/erasure/theories/EEtaExpandedFix.v index a094a3021..f0a877898 100644 --- a/erasure/theories/EEtaExpandedFix.v +++ b/erasure/theories/EEtaExpandedFix.v @@ -1056,7 +1056,7 @@ Lemma eval_etaexp {fl : WcbvFlags} {efl : EEnvFlags} {wcon : with_constructor_as Proof. intros etaΣ wfΣ. induction 1 as [ | ? ? ? ? ? ? ? ? IHs | | | | | ? ? ? ? ? ? ? ? ? ? ? IHs | ? ? ? ? ? ? ? ? ? ? ? IHs - | ? ? ? ? ? ? ? ? ? ? IHs | | | | | | | | | | | ] using eval_mkApps_rect; try now congruence. + | ? ? ? ? ? ? ? ? ? ? IHs | | | | | | | | | | | | ] using eval_mkApps_rect; try now congruence. all:try simp isEtaExp; rewrite -?isEtaExp_equation_1 => //. 6:{ move/isEtaExp_tApp'. @@ -1375,6 +1375,7 @@ Proof. - solve_all. depelim X; solve_all. eapply All2_over_undep in a. subst a0 a'; depelim H; constructor; solve_all. solve_all. + - simp_eta in IHeval1. eauto. Qed. Lemma isEtaExp_fixapp_mon {mfix idx n n'} : n <= n' -> isEtaExp_fixapp mfix idx n -> isEtaExp_fixapp mfix idx n'. @@ -1961,6 +1962,8 @@ Proof. - intros hexp. simp_eta in hexp. depelim X; repeat constructor; eauto. eapply All2_over_undep in a. subst a0 a'. solve_all. depelim hexp; cbn in *. destruct p. eapply All2_All2_Set. solve_all. solve_all. depelim hexp. destruct p. solve_all. + - intros hexp. simp_eta in hexp. econstructor; eauto. apply IHeval2. + specialize (IHeval1 hexp). eapply eval_etaexp in IHeval1. now simp_eta in IHeval1. all:eauto. - intros hexp. now eapply eval_atom. Unshelve. all: eauto. Qed. diff --git a/erasure/theories/EImplementBox.v b/erasure/theories/EImplementBox.v index 99deb0a5f..0817566d3 100644 --- a/erasure/theories/EImplementBox.v +++ b/erasure/theories/EImplementBox.v @@ -388,6 +388,15 @@ Proof. rewrite (isPrimApp_mkApps _ [_]). eauto. Qed. +Lemma implement_box_isLazyApp {efl : EEnvFlags} {Σ : global_declarations} t : + isLazyApp (implement_box t) = isLazyApp t. +Proof. + induction t; try now cbn; eauto. + simp implement_box. + rewrite (isLazyApp_mkApps _ [t2]). + rewrite (isLazyApp_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). @@ -670,6 +679,7 @@ Proof. intros H. cbn in H. rewrite implement_box_isConstructApp; eauto. rewrite implement_box_isPrimApp; eauto. + rewrite implement_box_isLazyApp; eauto. - intros; repeat match goal with [H : MCProd.and3 _ _ _ |- _] => destruct H end. simp implement_box in *. eapply eval_construct_block; tea. eauto. @@ -680,6 +690,9 @@ Proof. - intros wf H; depelim H; simp implement_box; repeat constructor. destruct a0. eapply All2_over_undep in a. eapply All2_All2_Set, All2_map. cbn -[implement_box]. solve_all. now destruct H. now destruct a0. + - intros evt evt' [] []. + simp implement_box. simp implement_box in e. + econstructor; eauto. - intros. destruct t; try solve [constructor; cbn in H, H0 |- *; try congruence]. cbn -[lookup_constructor] in H |- *. destruct args => //. Qed. diff --git a/erasure/theories/EInlineProjections.v b/erasure/theories/EInlineProjections.v index a37b97ae7..302b57bc3 100644 --- a/erasure/theories/EInlineProjections.v +++ b/erasure/theories/EInlineProjections.v @@ -710,13 +710,14 @@ Proof. * destruct with_guarded_fix. + move: i. rewrite !negb_or. - rewrite optimize_mkApps !isFixApp_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + rewrite optimize_mkApps !isFixApp_mkApps !isConstructApp_mkApps !isPrimApp_mkApps + !isLazyApp_mkApps. destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. rewrite !andb_true_r. rtoProp; intuition auto; destruct v => /= //. + move: i. rewrite !negb_or. - rewrite optimize_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + rewrite optimize_mkApps !isConstructApp_mkApps !isPrimApp_mkApps !isLazyApp_mkApps. destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. destruct v => /= //. - intros; rtoProp; intuition eauto. @@ -725,6 +726,8 @@ Proof. eapply All2_Set_All2 in ev. eapply All2_All2_Set. primProp. subst a0 a'; cbn in *. depelim H0; cbn in *. intuition auto; solve_all. primProp; depelim H0; intuition eauto. + - intros wf; econstructor; eauto. eapply IHev2. + eapply eval_wellformed in ev1; tea => //. - destruct t => //. all:constructor; eauto. cbn [atom optimize] in i |- *. diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 211562f18..449e36dc8 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -715,22 +715,24 @@ Proof. * destruct with_guarded_fix. + move: i. rewrite !negb_or. - rewrite remove_match_on_box_mkApps !isFixApp_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + rewrite remove_match_on_box_mkApps !isFixApp_mkApps !isConstructApp_mkApps !isPrimApp_mkApps !isLazyApp_mkApps. destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. rewrite !andb_true_r. rtoProp; intuition auto. destruct v => /= //. destruct v => /= //. destruct v => /= //. + destruct v => /= //. + move: i. rewrite !negb_or. - rewrite remove_match_on_box_mkApps !isConstructApp_mkApps !isPrimApp_mkApps. + rewrite remove_match_on_box_mkApps !isConstructApp_mkApps !isPrimApp_mkApps !isLazyApp_mkApps. destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. destruct v => /= //. - depelim X; repeat constructor. eapply All2_over_undep in a. eapply All2_All2_Set. eapply All2_Set_All2 in ev. subst a0 a'; cbn in *. rewrite /test_array_model /= in H. rtoProp; intuition auto; solve_all. eapply e. cbn in H. rewrite /test_array_model /= in H. now move/andP: H => []. + - intros. econstructor; eauto. eapply IHev2. eapply eval_closed in ev1; tea. - destruct t => //. all:constructor; eauto. cbn [atom remove_match_on_box] in i |- *. rewrite -lookup_constructor_remove_match_on_box //. destruct args => //. diff --git a/erasure/theories/ERemoveParams.v b/erasure/theories/ERemoveParams.v index c58d9e5d9..9780299a8 100644 --- a/erasure/theories/ERemoveParams.v +++ b/erasure/theories/ERemoveParams.v @@ -779,6 +779,14 @@ Proof. all:rewrite !isPrimApp_mkApps //. Qed. +Lemma strip_isLazyApp Σ f : + isLazyApp f = isLazyApp (strip Σ f). +Proof. + funelim (strip Σ f); cbn -[strip] => //. + all:rewrite map_InP_spec. + all:rewrite !isLazyApp_mkApps //. +Qed. + Lemma lookup_inductive_pars_is_prop_and_pars {Σ ind b pars} : inductive_isprop_and_pars Σ ind = Some (b, pars) -> lookup_inductive_pars Σ (inductive_mind ind) = Some pars. @@ -1011,8 +1019,8 @@ Proof. - rewrite !strip_tApp //. eapply eval_app_cong; tea. move: H1. eapply contraNN. - rewrite -strip_isLambda -strip_isConstructApp -strip_isFixApp -strip_isBox -strip_isPrimApp //. - rewrite -strip_isFix //. + rewrite -strip_isLambda -strip_isConstructApp -strip_isFixApp -strip_isBox -strip_isPrimApp + -strip_isLazyApp // -strip_isFix //. - rewrite !strip_mkApps // /=. rewrite (lookup_constructor_lookup_inductive_pars H). @@ -1030,6 +1038,8 @@ Proof. eapply All2_over_undep in a. eapply All2_Set_All2 in ev. eapply All2_All2_Set. solve_all. now destruct b. now destruct a0. + - simp_strip. simp_strip in e0. econstructor; tea. + - destruct t => //. all:constructor; eauto. simp strip. cbn [atom strip] in H |- *. diff --git a/erasure/theories/EWcbvEval.v b/erasure/theories/EWcbvEval.v index 7cd775a2a..6a6d78968 100644 --- a/erasure/theories/EWcbvEval.v +++ b/erasure/theories/EWcbvEval.v @@ -38,6 +38,7 @@ Definition atom `{wfl : WcbvFlags} Σ t := | tBox | tCoFix _ _ | tLambda _ _ + | tLazy _ | tFix _ _ => true | tConstruct ind c [] => negb with_constructor_as_block && isSome (lookup_constructor Σ ind c) | _ => false @@ -258,7 +259,7 @@ Section Wcbv. | eval_app_cong f f' a a' : eval f f' -> ~~ (isLambda f' || (if with_guarded_fix then isFixApp f' else isFix f') || isBox f' || isConstructApp f' - || isPrimApp f') -> + || isPrimApp f' || isLazyApp f') -> eval a a' -> eval (tApp f a) (tApp f' a') @@ -272,13 +273,12 @@ Section Wcbv. eval_primitive eval p p' -> eval (tPrim p) (tPrim p') - (* - | eval_lazy : eval (tLazy t) (tLazy t) - | eval_force t v v' : eval t (tLazy v) -> + | eval_force t v v' : + eval t (tLazy v) -> eval v v' -> - eval (tForce t) v' *) + eval (tForce t) v' - (** Atoms are values (includes abstractions, cofixpoints and constructors) *) + (** Atoms are values (includes abstractions, cofixpoints, constructors and lazy) *) | eval_atom t : atom Σ t -> eval t t. Hint Constructors eval : core. @@ -615,7 +615,7 @@ Section eval_rect. isBox f' || isConstructApp f' - || isPrimApp f')) + || isPrimApp f' || isLazyApp f')) (e0 : eval Σ a a'), P a a' e0 → P (tApp f16 a) @@ -626,14 +626,17 @@ Section eval_rect. (forall p p' (ev : eval_primitive (eval Σ) p p'), eval_primitive_ind _ P _ _ ev -> - P (tPrim p) (tPrim p') (eval_prim _ _ _ ev)) + P (tPrim p) (tPrim p') (eval_prim _ _ _ ev)) -> + (forall t t' v (ev1 : eval Σ t (tLazy t')) (ev2 : eval Σ t' v), + P _ _ ev1 -> P _ _ ev2 -> + P (tForce t) v (eval_force _ t t' v ev1 ev2)) - → (∀ (t : term) (i : atom Σ t), + → (∀ (t : term) (i : atom Σ t), P t t (eval_atom Σ t i)) → ∀ (t t0 : term) (e : eval Σ t t0), P t t0 e. Proof using Type. - intros ?????????????????????? H. + intros ??????????????????????? H. revert t t0 H. fix aux 3. move aux at top. @@ -785,6 +788,7 @@ Section Wcbv. + destruct with_guarded_fix. now cbn in i1. now cbn in i1. + constructor. + + now destruct with_guarded_fix; cbn in i1. + cbn in i. destruct with_guarded_fix; cbn in i; rtoProp; intuition auto. * destruct b0 as (ind & c & mdecl & idecl & cdecl & args & [H1 H2 H3 H4]). rewrite -[tApp _ _](mkApps_app _ (firstn n l) [a']). @@ -1058,7 +1062,7 @@ Section Wcbv. rewrite !mkApps_app /=. eapply eval_app_cong; tea. eapply IHargs => //. - rewrite isFixApp_mkApps // /= isConstructApp_mkApps // !negb_or isPrimApp_mkApps. + rewrite isFixApp_mkApps // /= isConstructApp_mkApps // !negb_or isPrimApp_mkApps isLazyApp_mkApps. rtoProp; intuition auto. apply nisLambda_mkApps => //. destruct with_guarded_fix => //; eapply nisFix_mkApps => //. @@ -1304,15 +1308,15 @@ Section Wcbv. cbn in i. rtoProp; intuition auto. + exfalso. rewrite !negb_or in i. specialize (IHev1 _ ev'1); noconf IHev1. cbn in i. rewrite guarded in i. rtoProp; intuition auto. - rewrite isFixApp_mkApps in H3 => //. + rewrite isFixApp_mkApps in H4 => //. + exfalso. rewrite !negb_or in i. specialize (IHev1 _ ev'1); noconf IHev1. cbn in i. rewrite guarded in i. rtoProp; intuition auto. - rewrite isFixApp_mkApps in H3 => //. + rewrite isFixApp_mkApps in H4 => //. + exfalso. rewrite !negb_or in i. specialize (IHev1 _ ev'1); noconf IHev1. cbn in i. rewrite unguarded in i. now cbn in i. + exfalso. rewrite !negb_or in i. specialize (IHev1 _ ev'1); noconf IHev1. cbn in i. rtoProp; intuition auto. - now rewrite isConstructApp_mkApps in H1. + now rewrite isConstructApp_mkApps in H2. + specialize (IHev1 _ ev'1); noconf IHev1. specialize (IHev2 _ ev'2); noconf IHev2. now assert (i0 = i) as -> by now apply uip. @@ -1329,6 +1333,7 @@ Section Wcbv. induction ev in a, ev0 |- *; depelim ev0; eauto. destruct a. f_equal; eauto. specialize (e0 _ e). now noconf e0. + - depelim ev'; go. - depelim ev'; try go. 2: now assert (i0 = i) as -> by now apply uip. exfalso. invs i. rewrite e in H0. destruct args; cbn in H0; invs H0. diff --git a/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v b/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v index dbaed8575..47d73387e 100644 --- a/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v +++ b/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v @@ -334,7 +334,7 @@ Lemma eval_preserve_mkApps_ind : P f11 f' -> (forall t u (ev' : eval Σ t u), eval_depth ev' <= eval_depth ev -> Q 0 t -> P t u) → ~~ (isLambda f' || (if with_guarded_fix then isFixApp f' else isFix f') || isBox f' - || isConstructApp f' || isPrimApp f') → + || isConstructApp f' || isPrimApp f' || isLazyApp f') → eval Σ a a' → P a a' → P' (tApp f11 a) (tApp f' a')) → (∀ ind i mdecl idecl cdecl args args', @@ -349,6 +349,11 @@ Lemma eval_preserve_mkApps_ind : eval_primitive_ind _ (fun x y _ => P x y) _ _ ev -> P' (tPrim p) (tPrim p')) -> + (forall t t' v, eval Σ t (tLazy t') -> eval Σ t' v -> + P t (tLazy t') -> + P t' v -> + P' (tForce t) v) -> + (∀ t : term, atom Σ t → Q 0 t -> P' t t) -> ∀ (t t0 : term), Q 0 t -> eval Σ t t0 → P' t t0. Proof. @@ -359,7 +364,7 @@ Proof. intros. pose proof (p := @Fix_F { t : _ & { t0 : _ & { qt : Q 0 t & eval Σ t t0 }}}). specialize (p (MR lt (fun x => eval_depth x.π2.π2.π2))). - set(foo := existT _ t (existT _ t0 (existT _ X17 H)) : { t : _ & { t0 : _ & { qt : Q 0 t & eval Σ t t0 }}}). + set(foo := existT _ t (existT _ t0 (existT _ X18 H)) : { t : _ & { t0 : _ & { qt : Q 0 t & eval Σ t t0 }}}). change t with (projT1 foo). change t0 with (projT1 (projT2 foo)). revert foo. @@ -369,8 +374,8 @@ Proof. forward p. 2:{ apply p. apply measure_wf, lt_wf. } clear p. - rename X17 into qt. rename X14 into Xcappexp. - rename X15 into Qprim, X16 into Qatom. + rename X18 into qt. rename X14 into Xcappexp. + rename X15 into Qprim, X16 into Qforce, X17 into Qatom. clear t t0 qt H. intros (t & t0 & qt & ev). intros IH. @@ -569,6 +574,15 @@ Proof. + depelim p0. destruct p. eapply and_assum. unshelve eapply IH; tea. cbn; lia. intros; split => //; eapply P'Q; tea. + - eapply qpres in qt; depelim qt => //. + assert (Q 0 (tLazy v)). + { eapply P'Q; tea. eapply (IH _ _ q ev1); tea. now cbn; lia. } + eapply qpres in X14. depelim X14 => //. + eapply Qforce; tea. + + eapply and_assum. eapply (IH _ _ q ev1); eauto; cbn; lia. + intros hp'. split => //. eapply P'Q; tea. + + eapply and_assum. eapply (IH _ _ q0 ev2); eauto; cbn; lia. + intros hp'. split => //. eapply P'Q; tea. - eapply Qatom; tea. Qed. diff --git a/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v b/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v index 158684f81..955c25c0b 100644 --- a/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v +++ b/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v @@ -298,7 +298,7 @@ Lemma eval_preserve_mkApps_ind : P f11 f' -> (forall t u (ev' : eval Σ t u), eval_depth ev' <= eval_depth ev -> Q 0 t -> P t u) → ~~ (isLambda f' || (if with_guarded_fix then isFixApp f' else isFix f') || isBox f' - || isConstructApp f' || isPrimApp f') → + || isConstructApp f' || isPrimApp f' || isLazyApp f') → eval Σ a a' → P a a' → P' (tApp f11 a) (tApp f' a')) → (∀ ind i mdecl idecl cdecl args args', @@ -313,6 +313,11 @@ Lemma eval_preserve_mkApps_ind : eval_primitive_ind _ (fun x y _ => P x y) _ _ ev -> P' (tPrim p) (tPrim p')) -> + (forall t t' v, eval Σ t (tLazy t') -> eval Σ t' v -> + P t (tLazy t') -> + P t' v -> + P' (tForce t) v) -> + (∀ t : term, atom Σ t → Q 0 t -> P' t t) -> ∀ (t t0 : term), Q 0 t -> eval Σ t t0 → P' t t0. Proof. @@ -322,7 +327,7 @@ Proof. intros. pose proof (p := @Fix_F { t : _ & { t0 : _ & { qt : Q 0 t & eval Σ t t0 }}}). specialize (p (MR lt (fun x => eval_depth x.π2.π2.π2))). - set(foo := existT _ t (existT _ t0 (existT _ X16 H)) : { t : _ & { t0 : _ & { qt : Q 0 t & eval Σ t t0 }}}). + set(foo := existT _ t (existT _ t0 (existT _ X17 H)) : { t : _ & { t0 : _ & { qt : Q 0 t & eval Σ t t0 }}}). change t with (projT1 foo). change t0 with (projT1 (projT2 foo)). revert foo. @@ -332,8 +337,8 @@ Proof. forward p. 2:{ apply p. apply measure_wf, lt_wf. } clear p. - rename X16 into qt. rename X13 into Xcappexp. - rename X14 into Qprim, X15 into Qatom. + rename X17 into qt. rename X13 into Xcappexp. + rename X14 into Qprim, X15 into Qforce, X16 into Qatom. clear t t0 qt H. intros (t & t0 & qt & ev). intros IH. @@ -356,7 +361,7 @@ Proof. Ltac myt hyp anda P'Q := eapply hyp; tea; (apply and_assum; [ih|hp' P'Q]). destruct ev. - 1-18:eapply qpres in qt as qt'; depelim qt' => //. + 1-19:eapply qpres in qt as qt'; depelim qt' => //. all:try congruence. - eapply X; tea; (apply and_assum; [ih|hp' P'Q]). - assert (ql : Q 0 (tLambda na b)). @@ -468,6 +473,15 @@ Proof. + depelim p0. destruct p. eapply and_assum. unshelve eapply IH; tea. cbn; lia. intros; split => //; eapply P'Q; tea. + - eapply qpres in qt; depelim qt => //. + assert (Q 0 (tLazy v)). + { eapply P'Q; tea. eapply (IH _ _ q ev1); tea. now cbn; lia. } + eapply qpres in X13. depelim X13 => //. + eapply Qforce; tea. + + eapply and_assum. eapply (IH _ _ q ev1); eauto; cbn; lia. + intros hp'. split => //. eapply P'Q; tea. + + eapply and_assum. eapply (IH _ _ q0 ev2); eauto; cbn; lia. + intros hp'. split => //. eapply P'Q; tea. - eapply Qatom; tea. Qed. diff --git a/erasure/theories/EWcbvEvalEtaInd.v b/erasure/theories/EWcbvEvalEtaInd.v index f2058f16f..442a7ee29 100644 --- a/erasure/theories/EWcbvEvalEtaInd.v +++ b/erasure/theories/EWcbvEvalEtaInd.v @@ -308,7 +308,7 @@ Lemma eval_preserve_mkApps_ind : P f11 f' -> (forall t u (ev' : eval Σ t u), eval_depth ev' <= eval_depth ev -> Q 0 t -> isEtaExp Σ t -> P t u) → ~~ (isLambda f' || (if with_guarded_fix then isFixApp f' else isFix f') || isBox f' || isConstructApp f' || - isPrimApp f') → + isPrimApp f' || isLazyApp f') → eval Σ a a' → P a a' → isEtaExp Σ (tApp f' a') -> P' (tApp f11 a) (tApp f' a')) → @@ -326,11 +326,10 @@ Lemma eval_preserve_mkApps_ind : eval_primitive_ind _ (fun x y _ => P x y) _ _ ev -> P' (tPrim p) (tPrim p')) -> - (* (forall t t', eval Σ t t' -> Q 0 t -> Q 0 t' -> P t t' -> - P' (tLazy t) (tLazy t')) -> - - (forall t t', eval Σ t t' -> Q 0 t -> Q 0 t' -> P t t' -> - P' (tForce t) (tForce t')) -> *) + (forall t t' v, eval Σ t (tLazy t') -> eval Σ t' v -> + P t (tLazy t') -> + P t' v -> + P' (tForce t) v) -> (∀ t : term, atom Σ t → Q 0 t -> isEtaExp Σ t -> P' t t) -> ∀ (t t0 : term), Q 0 t -> isEtaExp Σ t -> eval Σ t t0 → P' t t0. @@ -339,10 +338,10 @@ Proof. assert (qfixs: Qfixs Q) by tc. assert (qcofixs: Qcofixs Q) by tc. intros. - enough (P' t t0 × isEtaExp Σ t0). apply X17. + enough (P' t t0 × isEtaExp Σ t0). apply X18. pose proof (p := @Fix_F { t : _ & { t0 : _ & { qt : Q 0 t & eval Σ t t0 }}}). specialize (p (MR lt (fun x => eval_depth x.π2.π2.π2))). - set(foo := existT _ t (existT _ t0 (existT _ X16 H0)) : { t : _ & { t0 : _ & { qt : Q 0 t & eval Σ t t0 }}}). + set(foo := existT _ t (existT _ t0 (existT _ X17 H0)) : { t : _ & { t0 : _ & { qt : Q 0 t & eval Σ t t0 }}}). move: H. change t with (projT1 foo). change t0 with (projT1 (projT2 foo)). @@ -354,8 +353,9 @@ Proof. forward p. 2:{ apply p. apply measure_wf, lt_wf. } clear p. - rename X16 into qt. rename X13 into Xcappexp. - rename X14 into Qprim. rename X15 into Qatom. + rename X17 into qt. rename X13 into Xcappexp. + rename X14 into Qprim. rename X15 into Qforce. + rename X16 into Qatom. clear t t0 qt H0. intros (t & t0 & qt & ev). intros IH. @@ -381,7 +381,7 @@ Proof. eapply H; tea; (apply and_assum; [ih|hp' P'Q]) end. destruct ev. - 1-18:eapply qpres in qt as qt'; depelim qt' => //. + 1-19:eapply qpres in qt as qt'; depelim qt' => //. - move/isEtaExp_tApp. destruct decompose_app as [hd args] eqn:da. destruct (construct_viewc hd) eqn:cv. @@ -698,29 +698,44 @@ Proof. split; simp_eta. unshelve eapply Qprim. constructor; eauto. constructor. + apply All2_over_undep. cbn in IH. ELiftSubst.solve_all. - depelim H. destruct p as [[] ?]. + depelim H0. destruct p as [[] a0]. clear -ev IH a0 P'Q and_assum. cbn in a0. subst a; cbn in *. induction ev; constructor; eauto. - ** depelim a0. destruct p as []. + ** depelim a0. destruct p as [? []]. eapply and_assum. unshelve eapply IH; tea. cbn. lia. intros []. split => //. split => //. eapply P'Q; tea. ** depelim a0. intuition eauto. eapply IHev; intros. 2:eauto. unshelve eapply IH; tea. cbn; lia. - + ELiftSubst.solve_all. depelim H; destruct p as [[] ?]. + + ELiftSubst.solve_all. depelim H0. destruct p as [[? []] ?]. eapply and_assum. unshelve eapply IH; tea. cbn; lia. intros []; split => //; split => //. eapply P'Q; tea. + ELiftSubst.solve_all. subst a a'; cbn in *. - depelim H; constructor; cbn in *; intuition eauto. + depelim H0; constructor; cbn in *; intuition eauto. unshelve eapply IH. 2:tea. all:eauto. cbn; lia. clear -ev IH b P'Q and_assum. cbn in b. induction ev; constructor; eauto. - ** depelim b. destruct p. unshelve eapply IH. 2:tea. all:eauto. cbn. lia. + ** depelim b. destruct p as [? []]. unshelve eapply IH. 2:tea. all:eauto. cbn. lia. ** depelim b. intuition eauto. eapply IHev; intros. 2:eauto. unshelve eapply IH; tea. cbn; lia. - - intros ise. split => //. eapply Qatom; tea. + - simp_eta => ise. + eapply Qpres in qt. depelim qt => //. + assert (Q 0 (tLazy v)). + { eapply P'Q; tea. eapply (IH _ _ q ev1); tea. now cbn; lia. } + eapply qpres in X13. depelim X13 => //. + assert (isEtaExp Σ (tLazy v)) as etav. + { eapply (IH _ _ q ev1); tea. now cbn; lia. } + simp_eta in etav. + split. + + unshelve eapply Qforce. 2:tea. auto. + eapply and_assum. eapply (IH _ _ q ev1); tea. now cbn; lia. + intros []. split => //. split => //. eapply P'Q; tea. + pose proof (IH _ _ q ev1). forward X13. cbn; lia. specialize (X13 ise) as []. + eapply and_assum. eapply (IH _ _ q0 ev2); tea. now cbn; lia. + now simp_eta in i1. + + eapply (IH _ _ q0 ev2) => //. now cbn; lia. + - intros ise. split => //. eapply Qatom; tea. Qed. - From MetaCoq.Erasure Require Import ELiftSubst. Lemma Qpreserves_wellformed (efl : EEnvFlags) Σ : cstr_as_blocks = false -> diff --git a/erasure/theories/EWcbvEvalInd.v b/erasure/theories/EWcbvEvalInd.v index 05328f1ef..8e0f365b2 100644 --- a/erasure/theories/EWcbvEvalInd.v +++ b/erasure/theories/EWcbvEvalInd.v @@ -238,7 +238,7 @@ Section eval_mkApps_rect. then isFixApp f' else isFix f') || isBox f' || - isConstructApp f' || isPrimApp f')) + isConstructApp f' || isPrimApp f' || isLazyApp f')) (e0 : eval Σ a a'), P a a' → P (tApp f15 a) @@ -246,12 +246,15 @@ Section eval_mkApps_rect. (forall p p' (ev : eval_primitive (eval Σ) p p'), eval_primitive_ind _ (fun x y _ => P x y) _ _ ev -> - P (tPrim p) (tPrim p')) + P (tPrim p) (tPrim p')) -> + (forall t t' v (ev1 : eval Σ t (tLazy t')) (ev2 : eval Σ t' v), + P t (tLazy t') -> P t' v -> + P (tForce t) v) → (∀ t : term, atom Σ t → P t t) → ∀ t t0 : term, eval Σ t t0 → P t t0. Proof using Type. - intros ?????????????????????? H. + intros ??????????????????????? H. pose proof (p := @Fix_F { t : _ & { t0 : _ & eval Σ t t0 }}). specialize (p (MR lt (fun x => eval_depth x.π2.π2))). set(foo := existT _ t (existT _ t0 H) : { t : _ & { t0 : _ & eval Σ t t0 }}). diff --git a/erasure/theories/EWcbvEvalNamed.v b/erasure/theories/EWcbvEvalNamed.v index 02816f1d3..a3ac397de 100644 --- a/erasure/theories/EWcbvEvalNamed.v +++ b/erasure/theories/EWcbvEvalNamed.v @@ -36,7 +36,8 @@ Inductive value : Set := | 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)) -| vPrim (p : EPrimitive.prim_val value). +| vPrim (p : EPrimitive.prim_val value) +| vLazy (p : term) (env : list (ident * value)). Definition environment := list (ident * value). Definition add : ident -> value -> environment -> environment := fun (x : ident) (v : value) env => @@ -148,7 +149,14 @@ Section Wcbv. | eval_prim p p' : eval_primitive (eval Γ) p p' -> - eval Γ (tPrim p) (vPrim p'). + eval Γ (tPrim p) (vPrim p') + + | eval_lazy t : eval Γ (tLazy t) (vLazy t Γ) + | eval_force Γ' t t' v : + eval Γ t (vLazy t' Γ') -> + eval Γ' t' v -> + eval Γ (tForce t) v + . 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), @@ -247,7 +255,11 @@ Section Wcbv. P Γ (tConstruct ind c []) (vConstruct ind c []) (eval_construct_block_empty Γ ind c mdecl idecl cdecl e)) (f10 : ∀ (Γ : environment) (p : prim_val term) (p' : prim_val value) (ev : eval_primitive (eval Γ) p p') (evih : eval_primitive_ind (eval Γ) (P Γ) _ _ ev), - P Γ (tPrim p) (vPrim p') (eval_prim _ _ _ ev)), + P Γ (tPrim p) (vPrim p') (eval_prim _ _ _ ev)) + (f11 :∀ (Γ : environment) t, P Γ (tLazy t) (vLazy t Γ) (eval_lazy _ _)) + (f12 : ∀ Γ Γ' t t' v (ev0 : eval Γ t (vLazy t' Γ')) (ev1 : eval Γ' t' v), + P _ _ _ ev0 -> P _ _ _ ev1 -> + P _ _ _ (eval_force _ _ _ _ _ ev0 ev1)), 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 @@ -263,6 +275,8 @@ Section Wcbv. | @eval_construct_block _ ind c mdecl idecl cdecl args args' e0 l a => f8 Γ ind c mdecl idecl cdecl args args' e0 l a (map_All2_dep _ (F Γ) a) | @eval_construct_block_empty _ ind c mdecl idecl cdecl e0 => f9 Γ ind c mdecl idecl cdecl e0 | @eval_prim _ p p' ev => f10 Γ p p' ev (map_eval_primitive (P := P Γ) (F Γ) ev) + | @eval_lazy _ t => f11 Γ t + | @eval_force _ Γ' t t' v ev ev' => f12 _ _ _ _ _ ev ev' (F Γ _ _ ev) (F Γ' _ _ ev') end. End Wcbv. @@ -307,6 +321,7 @@ with represents_value : value -> term -> Set := | represents_value_tFix vfix i E mfix : 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) | represents_value_tPrim p p' : onPrims represents_value p p' -> represents_value (vPrim p) (tPrim p') +| represents_value_tLazy E t t' : [] ;;; E ⊩ t ~ t' -> represents_value (vLazy t E) (tLazy t') where "Γ ;;; E ⊩ s ~ t" := (represents Γ E s t). Derive Signature for represents represents_value. @@ -413,7 +428,9 @@ Program Definition represents_ind := P0 (vRecClos vfix i E) (tFix mfix i) (represents_value_tFix vfix i E mfix a)) (f13 : forall p p' (o : onPrims represents_value p p'), onPrims_dep _ P0 _ _ o -> - P0 (vPrim p) (tPrim p') (represents_value_tPrim p p' o)), + P0 (vPrim p) (tPrim p') (represents_value_tPrim p p' o)) + (f14 : forall E t t' (r : represents [] E t t'), P [] E t t' r -> + P0 (vLazy t E) (tLazy t') (represents_value_tLazy E t t' r)), fix F (l : list ident) (e : environment) (t t0 : term) (r : l;;; e ⊩ t ~ t0) {struct r} : P l e t t0 r := @@ -450,6 +467,7 @@ Program Definition represents_ind := | represents_value_tConstruct vs ts ind c a => f11 vs ts ind c a _ | represents_value_tFix vfix i E mfix a => f12 vfix i E mfix a _ | represents_value_tPrim p p' r => f13 p p' r (map_onPrims F0 r) + | represents_value_tLazy E t t' r => f14 E t t' r (F [] E t t' r) end for F). @@ -578,7 +596,9 @@ Program Definition represents_value_ind := P0 (vRecClos vfix i E) (tFix mfix i) (represents_value_tFix vfix i E mfix a)) (f13 : forall p p' (o : onPrims represents_value p p'), onPrims_dep _ P0 _ _ o -> - P0 (vPrim p) (tPrim p') (represents_value_tPrim p p' o)), + P0 (vPrim p) (tPrim p') (represents_value_tPrim p p' o)) + (f14 : forall E t t' (r : represents [] E t t'), P [] E t t' r -> + P0 (vLazy t E) (tLazy t') (represents_value_tLazy E t t' r)), fix F (l : list ident) (e : environment) (t t0 : term) @@ -617,6 +637,7 @@ Program Definition represents_value_ind := | represents_value_tConstruct vs ts ind c a => f11 vs ts ind c a _ | represents_value_tFix vfix i E mfix a => f12 vfix i E mfix a _ | represents_value_tPrim p p' r => f13 p p' r (map_onPrims F0 r) + | represents_value_tLazy E t t' r => f14 E t t' r (F [] E t t' r) end for F0). @@ -746,9 +767,10 @@ Definition rep_ind := (represents_value_tFix vfix i E mfix a)) (f13 : forall p p' (o : onPrims represents_value p p'), onPrims_dep _ P0 _ _ o -> P0 (vPrim p) (tPrim p') (represents_value_tPrim p p' o)) + f14 , - (represents_ind P P0 f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 flazy fforce f10 f11 f12 f13, - represents_value_ind P P0 f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 flazy fforce f10 f11 f12 f13)). + (represents_ind P P0 f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 flazy fforce f10 f11 f12 f13 f14, + represents_value_ind P P0 f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 flazy fforce f10 f11 f12 f13 f14)). Local Notation "'⊩' v ~ s" := (represents_value v s) (at level 50). Local Hint Constructors represents : core. @@ -1226,7 +1248,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. @@ -1515,7 +1537,8 @@ Inductive wf : value -> Type := All (fun t => sunny (map fst vfix ++ map fst E) (snd t)) vfix -> All (fun v => wf (snd v)) E -> wf (vRecClos vfix idx E) -| wf_vPrim p : primProp wf p -> wf (vPrim p). +| wf_vPrim p : primProp wf p -> wf (vPrim p) +| wf_vLazy E t : sunny (map fst E) t -> All (fun v => wf (snd v)) E -> wf (vLazy t E). Lemma declared_constant_Forall P Σ c decl : Forall (P ∘ snd) Σ -> @@ -1609,6 +1632,8 @@ Proof. - rtoProp; solve_all. Qed. +Derive NoConfusion for value. + 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 -> @@ -1723,6 +1748,8 @@ Proof. - cbn. econstructor. econstructor. - econstructor. solve_all. depelim evih; depelim Hsunny; try subst a0 a'; cbn; constructor; cbn in *; intuition solve_all. eapply All2_over_undep in a. solve_all. + - econstructor; eauto. + - eapply X in HE; eauto. depelim HE. eapply X0; eauto. Qed. Lemma eval_construct_All2 Σ E ind c args vs mdecl idecl cdecl : @@ -2154,6 +2181,12 @@ Proof. { clear -Hvs; induction Hvs; constructor; intuition eauto. } + constructor; eauto. eapply All2_All2_Set. { clear -Hvs; induction Hvs; constructor; intuition eauto. } + - invs Hrep; cbn in *; rtoProp. depelim H2. + apply s0 in H5 as [v' [reprev t1v]]; eauto. + depelim reprev. + eapply eval_wf in HΣ; tea. depelim HΣ. + eapply s1 in r as [v2 [reprv2 evv2]]; eauto. + exists v2. split => //. econstructor; eauto. - invs Hrep; cbn in *; try congruence; rtoProp. + econstructor. split; eauto. econstructor. eauto. + econstructor. split; eauto. econstructor. @@ -2170,8 +2203,8 @@ Proof. intros. cbn in *. destruct H5 as (([? []] & ?) & ?). rewrite app_nil_r in r. all: eauto. - - Unshelve. all: repeat econstructor. + + exists (vLazy t1 E). now split; econstructor. + Unshelve. all: repeat econstructor. Qed. Lemma concat_sing {X} l : diff --git a/erasure/theories/ErasureCorrectness.v b/erasure/theories/ErasureCorrectness.v index 4a12486ef..4dcabe930 100644 --- a/erasure/theories/ErasureCorrectness.v +++ b/erasure/theories/ErasureCorrectness.v @@ -40,6 +40,14 @@ Notation "Σ ⊢ s ⇓ t" := (EWcbvEval.eval Σ s t) (at level 50, s, t at next Import ssrbool. +Lemma erases_lazy Σ Γ t t' : Σ ;;; Γ |- t ⇝ℇ t' -> EAstUtils.isLazyApp t' -> False. +Proof. + rewrite /EAstUtils.isLazyApp /EAstUtils.head /EAstUtils.decompose_app. + generalize (@nil EAst.term) as l; intros l. + induction 1 in l |- *; cbn => //=. + eapply IHerases1. +Qed. + Lemma erases_correct (wfl := default_wcbv_flags) Σ t t' v Σ' : wf_ext Σ -> welltyped Σ [] t -> @@ -1150,6 +1158,10 @@ Proof. -- rewrite !negb_or in i. rtoProp; intuition auto. eapply is_PrimApp_erases in H8; tea. now move/negbTE: H8. + -- rewrite !negb_or in i. + rtoProp; intuition auto. + apply/negbTE. apply PCUICNormal.negb_is_true => isl. + eapply erases_lazy in H1; tea. + exists EAst.tBox. split. 2: now constructor; econstructor. econstructor. eapply inversion_App in Hty as [na [A [B [Hf [Ha _]]]]]; auto. diff --git a/erasure/theories/Typed/OptimizeCorrectness.v b/erasure/theories/Typed/OptimizeCorrectness.v index 92bc3b69e..a8f685dfc 100644 --- a/erasure/theories/Typed/OptimizeCorrectness.v +++ b/erasure/theories/Typed/OptimizeCorrectness.v @@ -2017,6 +2017,8 @@ Proof. - depelim X; auto. eapply All2_over_undep in a. eapply All2_Set_All2 in ev. solve_all. subst a0 a'; cbn in *. depelim exp_t; constructor; cbn in *; intuition eauto. solve_all. + - eapply IHev1 in exp_t. eapply IHev2 in exp_t. + eapply is_expanded_aux_upwards; tea. lia. Qed. Lemma valid_case_masks_lift ind c brs n k : @@ -2823,6 +2825,7 @@ Proof with auto with dearg. - depelim X; auto. eapply All2_over_undep in a. eapply All2_Set_All2 in ev. subst a0 a'; cbn -[test_prim] in *. solve_all. depelim H0; constructor; cbn; intuition eauto. solve_all. + - eapply IHev2. eapply eval_closed in ev1; tea. eapply IHev1; eauto. Qed. Lemma declared_constant_dearg Σ k cst : @@ -3452,22 +3455,22 @@ Section dearg. cbn in *; propify. rewrite dearg_single_masked by (now rewrite map_length). rewrite isLambda_mkApps, isFixApp_mkApps, isBox_mkApps, isConstructApp_mkApps;cbn in *. - rewrite isPrimApp_mkApps. - destruct with_guarded_fix;cbn;auto. + rewrite isPrimApp_mkApps, isLazyApp_mkApps in *. + destruct with_guarded_fix;cbn;intuition auto. now rewrite EOptimizePropDiscr.isFix_mkApps;cbn. * rewrite isLambda_mkApps, isFixApp_mkApps, isBox_mkApps, isConstructApp_mkApps in *;cbn in *. propify. destruct with_guarded_fix;cbn in *; intuition. * unfold dearg_case. destruct with_guarded_fix;cbn. - now rewrite isLambda_mkApps, isFixApp_mkApps, isBox_mkApps, isConstructApp_mkApps, isPrimApp_mkApps;cbn. - now rewrite isLambda_mkApps, isBox_mkApps, isConstructApp_mkApps, EOptimizePropDiscr.isFix_mkApps, isPrimApp_mkApps;cbn. + now rewrite isLambda_mkApps, isFixApp_mkApps, isBox_mkApps, isConstructApp_mkApps, isPrimApp_mkApps, isLazyApp_mkApps;cbn. + now rewrite isLambda_mkApps, isBox_mkApps, isConstructApp_mkApps, EOptimizePropDiscr.isFix_mkApps, isPrimApp_mkApps, isLazyApp_mkApps;cbn. * unfold dearg_proj. unfold dearg_case. destruct with_guarded_fix;cbn. - ** now rewrite isLambda_mkApps, isFixApp_mkApps, isBox_mkApps, isConstructApp_mkApps, isPrimApp_mkApps;cbn. - ** now rewrite isLambda_mkApps, isBox_mkApps, isConstructApp_mkApps, EOptimizePropDiscr.isFix_mkApps, isPrimApp_mkApps;cbn. - * rewrite !isLambda_mkApps, !isFixApp_mkApps, !EOptimizePropDiscr.isFix_mkApps, !isBox_mkApps, isConstructApp_mkApps, isPrimApp_mkApps in * + ** now rewrite isLambda_mkApps, isFixApp_mkApps, isBox_mkApps, isConstructApp_mkApps, isPrimApp_mkApps, isLazyApp_mkApps;cbn. + ** now rewrite isLambda_mkApps, isBox_mkApps, isConstructApp_mkApps, EOptimizePropDiscr.isFix_mkApps, isPrimApp_mkApps, isLazyApp_mkApps;cbn. + * rewrite !isLambda_mkApps, !isFixApp_mkApps, !EOptimizePropDiscr.isFix_mkApps, !isBox_mkApps, isConstructApp_mkApps, isPrimApp_mkApps, isLazyApp_mkApps in * by now destruct hd. rewrite map_length. destruct with_guarded_fix;cbn;auto; @@ -4971,6 +4974,9 @@ Proof. unshelve eapply IHb0; tea. cbn in deriv_len. lia. cbn in *; unfold test_array_model in *; subst a a'; cbn in *. unshelve eapply IH; tea; rtoProp; intuition eauto. lia. + + facts. econstructor. specialize (IH _ _ clos_t valid_t exp_t ev1). + cbn in IH. apply IH. lia. + now forward (IH v _ H2 H4 H6 ev2). + destruct t; cbn in *; try destruct y; try congruence; now constructor. Qed. End dearg_correct. diff --git a/erasure/theories/Typed/WcbvEvalAux.v b/erasure/theories/Typed/WcbvEvalAux.v index e24e82804..11b2d6413 100644 --- a/erasure/theories/Typed/WcbvEvalAux.v +++ b/erasure/theories/Typed/WcbvEvalAux.v @@ -307,6 +307,7 @@ Fixpoint deriv_length {Σ t v} (ev : Σ e⊢ t ⇓ v) : nat := S (deriv_length ev1 + deriv_length ev2 + deriv_length ev3) | eval_construct_block _ _ _ _ _ args _ _ _ _ _ => S #|args| | eval_prim _ _ ev => S (eval_prim_length (@deriv_length _) ev) + | eval_force _ _ _ ev ev' => S (deriv_length ev + deriv_length ev') end. Lemma deriv_length_min {Σ t v} (ev : Σ e⊢ t ⇓ v) : From 462f3abaab68e5ca3fd861fcee29c660a2570b5b Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 22 Jul 2024 14:46:51 +0200 Subject: [PATCH 27/28] Proof that reordering of constructors is correct (#1095) * WIP on reordering constructors * New/old tag reasoning * WIP correctness of constructor reordering. main lemma proven * Reordering preserves well-formedness * Reorderin * Admit free proofs of preservation for wf and substitution * WIP adapting to an additional mapping argument for transforms and stronger wellformedness property on extracted terms * Full composed pipeline with mapping of constructors * Fixes due to removal of useless -fast flag and change for reordering of constructors, now verified * Remove option to reorder constructors, now safely done always * Fix metacoq_tour * Fix a remaining todo * Remove generated files --- .vscode/metacoq.code-workspace | 1 + erasure-plugin/clean_extraction.sh | 2 +- erasure-plugin/src/g_metacoq_erasure.mlg | 15 +- erasure-plugin/theories/ETransform.v | 69 +- erasure-plugin/theories/Erasure.v | 890 ++++++++-- erasure/theories/ECoInductiveToInductive.v | 3 +- erasure/theories/EConstructorsAsBlocks.v | 2 +- erasure/theories/EDeps.v | 46 +- erasure/theories/EGenericGlobalMap.v | 2 +- erasure/theories/EGenericMapEnv.v | 2 +- erasure/theories/EImplementBox.v | 3 +- erasure/theories/EInlineProjections.v | 14 +- erasure/theories/EOptimizePropDiscr.v | 6 +- erasure/theories/EProgram.v | 3 + erasure/theories/ERemoveParams.v | 5 +- erasure/theories/EReorderCstrs.v | 1541 ++++++++++++++++- .../EWcbvEvalCstrsAsBlocksFixLambdaInd.v | 1 - erasure/theories/EWcbvEvalCstrsAsBlocksInd.v | 1 + erasure/theories/EWcbvEvalNamed.v | 9 +- erasure/theories/EWellformed.v | 26 +- erasure/theories/ErasureCorrectness.v | 9 +- erasure/theories/ErasureFunction.v | 1 + erasure/theories/ErasureFunctionProperties.v | 12 +- erasure/theories/ErasureProperties.v | 35 +- erasure/theories/Extract.v | 1 + erasure/theories/Typed/OptimizeCorrectness.v | 10 +- examples/metacoq_tour.v | 2 +- safechecker-plugin/clean_extraction.sh | 2 +- template-coq/_PluginProject | 248 --- template-coq/_TemplateCoqProject | 32 - test-suite/erasure_live_test.v | 15 +- 31 files changed, 2456 insertions(+), 552 deletions(-) delete mode 100644 template-coq/_PluginProject delete mode 100644 template-coq/_TemplateCoqProject diff --git a/.vscode/metacoq.code-workspace b/.vscode/metacoq.code-workspace index f69bd9643..acfd61685 100644 --- a/.vscode/metacoq.code-workspace +++ b/.vscode/metacoq.code-workspace @@ -106,5 +106,6 @@ "**/.DS_Store": true, "**/Thumbs.db": true }, + "coq-lsp.check_only_on_request": true, } } diff --git a/erasure-plugin/clean_extraction.sh b/erasure-plugin/clean_extraction.sh index 4187f6ddc..562a941fa 100755 --- a/erasure-plugin/clean_extraction.sh +++ b/erasure-plugin/clean_extraction.sh @@ -9,7 +9,7 @@ fi shopt -s nullglob # make the for loop do nothnig when there is no *.ml* files -files=`cat ../template-coq/_PluginProject | grep "^[^#].*mli\?$" | sed -e s/gen-src/src/` +files=`cat ../template-coq/_PluginProject.in | grep "^[^#].*mli\?$" | sed -e s/gen-src/src/` if [[ ! -f "src/metacoq_erasure_plugin.cmxs" || "src/metacoq_erasure_plugin.cmxs" -ot "theories/Extraction.vo" ]] diff --git a/erasure-plugin/src/g_metacoq_erasure.mlg b/erasure-plugin/src/g_metacoq_erasure.mlg index 3efd45e9e..84d490c08 100644 --- a/erasure-plugin/src/g_metacoq_erasure.mlg +++ b/erasure-plugin/src/g_metacoq_erasure.mlg @@ -25,7 +25,6 @@ type erasure_command_args = | Time | Typed | BypassQeds - | Fast let pr_char c = str (Char.escaped c) @@ -46,24 +45,19 @@ type erasure_config = { unsafe : bool; time : bool; typed : bool; - bypass_qeds : bool; - fast : bool; - } + bypass_qeds : bool } let default_config = { unsafe = false; time = false; typed = false; - bypass_qeds = false; - fast = false } + bypass_qeds = false } let make_erasure_config config = let open Erasure0 in { enable_unsafe = if config.unsafe then all_unsafe_passes else no_unsafe_passes ; enable_typed_erasure = config.typed; - enable_fast_remove_params = config.fast; dearging_config = default_dearging_config; - inductives_mapping = []; inlined_constants = Kernames.KernameSet.empty } let time_opt config str fn arg = @@ -76,8 +70,9 @@ let check config env evm c = let time str f arg = time_opt config str f arg in let term = time (str"Quoting") (Ast_quoter.quote_term_rec ~bypass:config.bypass_qeds env evm) (EConstr.to_constr evm c) in let erasure_config = make_erasure_config config in + let mapping = [] in let erase = - time (str"Erasure") (Erasure0.erase_and_print_template_program erasure_config) term + time (str"Erasure") (Erasure0.erase_and_print_template_program erasure_config mapping) term in Feedback.msg_notice (pr_char_list erase) @@ -92,7 +87,6 @@ let interp_erase args env evm c = | Time -> aux {config with time = true} args | Typed -> aux {config with typed = true} args | BypassQeds -> aux {config with bypass_qeds = true} args - | Fast -> aux {config with fast = true} args in aux default_config args in check config env evm c @@ -123,7 +117,6 @@ ARGUMENT EXTEND erase_args | [ "-time" ] -> { Time } | [ "-typed" ] -> { Typed } | [ "-bypass-qeds" ] -> { BypassQeds } -| [ "-fast" ] -> { Fast } END VERNAC COMMAND EXTEND MetaCoqErase CLASSIFIED AS QUERY diff --git a/erasure-plugin/theories/ETransform.v b/erasure-plugin/theories/ETransform.v index 06a08372a..1e58c8f01 100644 --- a/erasure-plugin/theories/ETransform.v +++ b/erasure-plugin/theories/ETransform.v @@ -1020,42 +1020,59 @@ Qed. From MetaCoq.Erasure Require Import EReorderCstrs. -Axiom trust_reorder_cstrs_wf : - forall efl : EEnvFlags, - WcbvFlags -> - forall (m : inductives_mapping) (input : Transform.program E.global_context term), - wf_eprogram efl input -> wf_eprogram efl (reorder_program m input). -Axiom trust_reorder_cstrs_pres : - forall (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping) (p : Transform.program E.global_context term) - (v : term), - wf_eprogram efl p -> - eval_eprogram wfl p v -> exists v' : term, eval_eprogram wfl (reorder_program m p) v' /\ v' = reorder m v. +Definition eval_eprogram_mapping (wfl : WcbvFlags) (p : inductives_mapping * eprogram) t := + eval_eprogram wfl p.2 t. -Program Definition reorder_cstrs_transformation (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping) : - Transform.t _ _ EAst.term EAst.term _ _ - (eval_eprogram wfl) (eval_eprogram wfl) := - {| name := "reoder inductive constructors "; - transform p _ := EReorderCstrs.reorder_program m p ; - pre p := wf_eprogram efl p ; - post p := wf_eprogram efl p ; - obseq p hp p' v v' := v' = EReorderCstrs.reorder m v |}. +Definition eval_eprogram_env_mapping (wfl : WcbvFlags) (p : inductives_mapping * eprogram_env) t := + eval_eprogram_env wfl p.2 t. + +Definition to_program (e : eprogram_env) : eprogram := + (e.1, e.2). + +Program Definition reorder_cstrs_transformation {efl : EEnvFlags} {wca : cstr_as_blocks = false} {has_app : has_tApp} + (wfl : WcbvFlags) {wcon : with_constructor_as_block = false} : + Transform.t _ _ _ EAst.term _ _ + (eval_eprogram_env_mapping wfl) (eval_eprogram wfl) := + {| name := "reorder inductive constructors "; + transform p _ := EReorderCstrs.reorder_program p.1 (to_program p.2) ; + pre p := [/\ wf_eprogram_env efl p.2, EEtaExpandedFix.expanded_eprogram_env p.2 & wf_inductives_mapping p.2.1 p.1] ; + post p := wf_eprogram efl p /\ EEtaExpandedFix.expanded_eprogram p; + obseq p hp p' v v' := v' = EReorderCstrs.reorder p.1 v |}. Next Obligation. - move=> efl wfl m. cbn. now apply trust_reorder_cstrs_wf. + move=> efl wca hasapp wfl wcb [m p] [wfp exp wfm]. split => //. + now unshelve eapply reorder_program_wf. cbn. + now eapply reorder_program_expanded_fix. Qed. Next Obligation. - red. eapply trust_reorder_cstrs_pres. + red. intros efl wca hasapp wfl wcb [m p] v [wfp wfm] evp. + destruct evp as [ev]. + unshelve eapply EReorderCstrs.optimize_correct in ev; trea. + 2,3:apply wfp. + exists (reorder m v). split => //. Qed. #[global] -Axiom trust_reorder_cstrs_transformation_ext : forall (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping), - TransformExt.t (reorder_cstrs_transformation efl wfl m) - (fun p p' => extends p.1 p'.1) (fun p p' => extends p.1 p'.1). +Instance reorder_cstrs_transformation_ext {efl : EEnvFlags} (wca : cstr_as_blocks = false) (has_app : has_tApp) (wfl : WcbvFlags) + {wcon : with_constructor_as_block = false} : + TransformExt.t (reorder_cstrs_transformation (wca := wca) (has_app := has_app) wfl (wcon:=wcon)) + (fun p p' => p.1 = p'.1 /\ extends p.2.1 p'.2.1) (fun p p' => extends p.1 p'.1). +Proof. + red. intros p p' pr pr' [eq ext]. + red; cbn. rewrite -eq. eapply EReorderCstrs.optimize_extends_env; eauto. + move: pr'; cbn. now intros []. apply pr. apply pr'. +Qed. #[global] -Axiom trust_reorder_cstrs_transformation_ext' : forall (efl : EEnvFlags) (wfl : WcbvFlags) (m : inductives_mapping), - TransformExt.t (reorder_cstrs_transformation efl wfl m) - extends_eprogram extends_eprogram. +Instance reorder_cstrs_transformation_ext' {efl : EEnvFlags} (wca : cstr_as_blocks = false) (has_app : has_tApp) (wfl : WcbvFlags) + {wcon : with_constructor_as_block = false} : + TransformExt.t (reorder_cstrs_transformation (wca := wca) (has_app := has_app) wfl (wcon:=wcon)) + (fun p p' => p.1 = p'.1 /\ extends_eprogram_env p.2 p'.2) extends_eprogram. +Proof. + red. intros p p' pr pr' [eq [ext eq']]. cbn. + red. cbn. rewrite -eq -eq'. split => //. eapply EReorderCstrs.optimize_extends_env; eauto. + move: pr'; cbn. now intros []. apply pr. apply pr'. +Qed. From MetaCoq.Erasure Require Import EUnboxing. diff --git a/erasure-plugin/theories/Erasure.v b/erasure-plugin/theories/Erasure.v index a36fbe6da..83d2ed7d6 100644 --- a/erasure-plugin/theories/Erasure.v +++ b/erasure-plugin/theories/Erasure.v @@ -34,7 +34,6 @@ Local Obligation Tactic := program_simpl. Record unsafe_passes := { cofix_to_lazy : bool; - reorder_constructors : bool; inlining : bool; unboxing : bool; betared : bool }. @@ -42,9 +41,7 @@ Record unsafe_passes := Record erasure_configuration := { enable_unsafe : unsafe_passes; enable_typed_erasure : bool; - enable_fast_remove_params : bool; dearging_config : dearging_config; - inductives_mapping : EReorderCstrs.inductives_mapping; inlined_constants : KernameSet.t }. @@ -56,7 +53,6 @@ Definition default_dearging_config := Definition make_unsafe_passes b := {| cofix_to_lazy := b; - reorder_constructors := b; inlining := b; unboxing := b; betared := b |}. @@ -65,12 +61,10 @@ Definition no_unsafe_passes := make_unsafe_passes false. Definition all_unsafe_passes := make_unsafe_passes true. (* This runs the cofix -> fix/lazy translation as well as inlining and - beta-redex simplification, which are not verified. It does not change - representation by reordering constructors or unboxing. *) + beta-redex simplification, which are not verified. It does do unboxing. *) Definition default_unsafe_passes := {| cofix_to_lazy := true; - reorder_constructors := false; inlining := true; unboxing := false; betared := true |}. @@ -79,17 +73,13 @@ Definition default_erasure_config := {| enable_unsafe := default_unsafe_passes; dearging_config := default_dearging_config; enable_typed_erasure := true; - enable_fast_remove_params := true; - inductives_mapping := []; inlined_constants := KernameSet.empty |}. (* This runs only the verified phases without the typed erasure and "fast" remove params *) Definition safe_erasure_config := {| enable_unsafe := no_unsafe_passes; enable_typed_erasure := false; - enable_fast_remove_params := false; dearging_config := default_dearging_config; - inductives_mapping := []; inlined_constants := KernameSet.empty |}. Axiom assume_welltyped_template_program_expansion : @@ -135,8 +125,6 @@ Program Definition optional_unsafe_transforms econf := (* Coinductives & cofixpoints are translated to inductive types and thunked fixpoints *) coinductive_to_inductive_transformation efl (has_app := eq_refl) (has_box := eq_refl) (has_rel := eq_refl) (has_pars := eq_refl) (has_cstrblocks := eq_refl)) ▷ - ETransform.optional_self_transform passes.(reorder_constructors) - (reorder_cstrs_transformation efl final_wcbv_flags econf.(inductives_mapping)) ▷ ETransform.optional_self_transform passes.(unboxing) (rebuild_wf_env_transform (efl := efl) false false ▷ unbox_transformation efl final_wcbv_flags) ▷ @@ -149,26 +137,22 @@ Program Definition optional_unsafe_transforms econf := betared_transformation efl final_wcbv_flags). Next Obligation. - destruct (enable_unsafe econf) as [[] [] [] [] []]; cbn in * => //; intuition auto. + destruct (enable_unsafe econf) as [[] [] [] []]; cbn in * => //; intuition auto. Qed. Next Obligation. - destruct (enable_unsafe econf) as [[] [] [] [] []]; cbn in * => //; intuition auto. + destruct (enable_unsafe econf) as [[] [] [] []]; cbn in * => //; intuition auto. Qed. Next Obligation. - destruct (enable_unsafe econf) as [[] [] [] [] []]; cbn in * => //; intuition auto. -Qed. -Next Obligation. - destruct (enable_unsafe econf) as [[] [] [] [] []]; cbn in * => //; intuition auto. + destruct (enable_unsafe econf) as [[] [] [] []]; cbn in * => //; intuition auto. 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 _ _ (* Standard evaluation, with cases on prop, guarded fixpoints, applied constructors *) - (EProgram.eval_eprogram_env default_wcbv_flags) + (eval_eprogram_env default_wcbv_flags) (* Target evaluation, with no more cases on prop, unguarded fixpoints, constructors as block *) (EProgram.eval_eprogram final_wcbv_flags) := - (* 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. *) @@ -203,9 +187,276 @@ Next Obligation. now eapply ETransform.expanded_eprogram_env_expanded_eprogram_cstrs. Qed. +Program Definition verified_lambdabox_pipeline_mapping {guard : abstract_guard_impl} + (efl := EWellformed.all_env_flags) + : Transform.t _ _ _ EAst.term _ _ + (* Standard evaluation, with cases on prop, guarded fixpoints, applied constructors *) + (eval_eprogram_env_mapping default_wcbv_flags) + (* Target evaluation, with no more cases on prop, unguarded fixpoints, constructors as block *) + (EProgram.eval_eprogram final_wcbv_flags) := + (* Reorder constructors according to the mapping *) + reorder_cstrs_transformation (wca := eq_refl) (has_app := eq_refl) default_wcbv_flags (wcon := eq_refl) ▷ + (* Rebuild the efficient lookup table *) + rebuild_wf_env_transform (efl := efl) true true ▷ + (* Extract as usual *) + verified_lambdabox_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 eval_pcuic_program_mapping (p : inductives_mapping * pcuic_program) t := + PCUICTransform.eval_pcuic_program p.2 t. + +Import EReorderCstrs (wf_tags_list). + +Definition wf_pcuic_inductive_mapping (Σ : PCUICAst.PCUICEnvironment.global_env) (m : inductive_mapping) : bool := + let '(ind, (_, cstrs)) := m in + match PCUICAst.PCUICLookup.lookup_inductive Σ ind with + | Some (mib, oib) => wf_tags_list cstrs #|oib.(PCUICAst.PCUICEnvironment.ind_ctors)| + | None => true + end. + +Definition wf_pcuic_inductives_mapping Σ m := forallb (wf_pcuic_inductive_mapping Σ) m. + +Import ErasureFunction ErasureFunctionProperties. + +Import PCUICAst.PCUICEnvironment Extract EReorderCstrs EDeps EGlobalEnv. + +Lemma lookup_env_filter Σ f kn : + fresh_global kn Σ -> + lookup_env (filter f Σ) kn = None. +Proof. + induction 1; cbn => //. + destruct f. cbn. case: eqb_spec. intros ->. contradiction. + eauto. eauto. +Qed. + +Lemma lookup_env_filter_deps {efl : EEnvFlags} {Σ f i p} : + wf_glob Σ -> + lookup_env Σ i = Some p -> + lookup_env (filter f Σ) i = Some p \/ + lookup_env (filter f Σ) i = None. +Proof. + move=> wfΣ. + induction Σ in wfΣ |- *; cbn => //. + case: eqb_spec. + - intros -> [= <-]. destruct f. left. cbn. now rewrite eqb_refl. + right. depelim wfΣ. cbn. now apply lookup_env_filter. + - intros diff hl. + destruct f. cbn. case: eqb_spec. + intros ->. contradiction. intros _. depelim wfΣ. + now specialize (IHΣ wfΣ hl). + depelim wfΣ. + now specialize (IHΣ wfΣ hl). +Qed. + +Lemma lookup_env_filter_deps_None {efl : EEnvFlags} {Σ f i} : + wf_glob Σ -> + lookup_env Σ i = None -> + lookup_env (filter f Σ) i = None. +Proof. + move=> wfΣ. + induction Σ in wfΣ |- *; cbn => //. + case: eqb_spec. + - intros -> [= <-]. + - intros diff hl. + destruct f. cbn. case: eqb_spec. + intros ->. contradiction. intros _. depelim wfΣ. + now specialize (IHΣ wfΣ hl). + depelim wfΣ. + now specialize (IHΣ wfΣ hl). +Qed. + +Lemma lookup_inductive_filter_deps_Some {efl : EEnvFlags} {Σ f i p} : + wf_glob Σ -> + lookup_inductive Σ i = Some p -> + lookup_inductive (filter f Σ) i = Some p \/ + lookup_inductive (filter f Σ) i = None. +Proof. + move => wfΣ. + rewrite /lookup_inductive /lookup_minductive. + destruct lookup_env eqn:hle => //=. + eapply lookup_env_filter_deps in hle as [-> | ->] => //; destruct g => //=. now left. + now right. +Qed. + +Lemma lookup_inductive_filter_deps_None {efl : EEnvFlags} {Σ f i} : + wf_glob Σ -> + lookup_inductive Σ i = None -> + lookup_inductive (filter f Σ) i = None. +Proof. + move => wfΣ. + rewrite /lookup_inductive /lookup_minductive. + destruct lookup_env eqn:hle => //=; try destruct g => //=. + eapply (lookup_env_filter_deps) in hle as [-> | ->] => //; destruct g => //=. + eapply (lookup_env_filter_deps) in hle as [-> | ->] => //; destruct g => //=. + eapply (lookup_env_filter_deps_None) in hle as -> => //; destruct g => //=. +Qed. + +Lemma erases_global_lookup_env_constant Σ Σ' kn : + erases_global Σ Σ' -> + forall d, PCUICAst.PCUICEnvironment.lookup_env Σ kn = Some (ConstantDecl d) -> + exists d', lookup_env Σ' kn = Some (EAst.ConstantDecl d'). +Proof. + destruct Σ as [? ? ?]. rewrite /erases_global //=. + induction 1; cbn => //. + - intros d. case: eqb_spec. + + intros ->. intros [=]; subst. now eexists. + + intros diff. cbn in *. eauto. + - intros d. case: eqb_spec. + + intros ->. intros [=]; subst. + + intros diff. cbn in *. eauto. +Qed. + +Lemma erases_global_lookup_env_inductive Σ Σ' kn : + erases_global Σ Σ' -> + forall d, PCUICAst.PCUICEnvironment.lookup_env Σ kn = Some (InductiveDecl d) -> + exists d', lookup_env Σ' kn = Some (EAst.InductiveDecl d') /\ erases_mutual_inductive_body d d'. +Proof. + destruct Σ as [? ? ?]. rewrite /erases_global //=. + induction 1; cbn => //. + - intros d. case: eqb_spec. + + intros ->. intros [=]; subst. + + intros diff. cbn in *. eauto. + - intros d. case: eqb_spec. + + intros ->. intros [=]; subst. now eexists. + + intros diff. cbn in *. eauto. +Qed. + +Lemma erases_global_lookup_env_none Σ Σ' kn : + erases_global Σ Σ' -> + PCUICAst.PCUICEnvironment.lookup_env Σ kn = None -> + lookup_env Σ' kn = None. +Proof. + destruct Σ as [? ? ?]. rewrite /erases_global //=. + induction 1; cbn => //. + - case: eqb_spec. + + intros ->. intros [=]; subst. + + intros diff. cbn in *. eauto. + - case: eqb_spec. + + intros ->. intros [=]; subst. + + intros diff. cbn in *. eauto. +Qed. + +Lemma erases_wf_inductives_mapping {Σ : global_env_ext} {Σ' deps m} : + PCUICTyping.wf_ext Σ -> + erases_global Σ Σ' -> + wf_pcuic_inductives_mapping Σ m -> + EReorderCstrs.wf_inductives_mapping (filter_deps deps Σ') m. +Proof. + move=> wfΣ er. + have wfΣ' := ErasureCorrectness.erases_global_wf_glob _ (fst wfΣ) er. + rewrite /wf_pcuic_inductives_mapping /wf_inductives_mapping. + solve_all. move: H. + rewrite /wf_pcuic_inductive_mapping /wf_inductive_mapping. + destruct x as [ind [na tags]]. + have eral := EDeps.erases_global_all_deps _ _ (PCUICElimination.wf_ext_wf _ wfΣ) er. + have trind := ErasureProperties.trans_lookup_inductive eral ind. + destruct PCUICAst.PCUICLookup.lookup_inductive as [[mib oib]|] eqn:li. + specialize (trind _ _ eq_refl). + - intros wftags. + destruct trind as [mib' [oib' [hli' hctors]]]. + rewrite (filter_deps_filter (efl := all_env_flags)) //. + set (f := fun x => _). + eapply (lookup_inductive_filter_deps_Some (efl := all_env_flags) (f:=f)) in hli' as [H|H] => //; rewrite H //. + congruence. + - intros _. clear trind. + destruct (lookup_inductive (filter_deps deps Σ') ind) eqn:li' => //. + apply False_rect. move: li'. + rewrite (filter_deps_filter (efl:=all_env_flags)) //. + generalize (fun x : KernameSet.elt × EAst.global_decl => flip KernameSet.mem (global_deps Σ' deps) x.1). + intros f li'. + have nli : lookup_inductive Σ' ind = None. + { clear -wfΣ wfΣ' er li. + move: li. rewrite /PCUICAst.PCUICLookup.lookup_inductive /lookup_inductive. + rewrite /PCUICAst.PCUICLookup.lookup_inductive_gen /lookup_minductive. + rewrite /PCUICAst.PCUICLookup.lookup_minductive_gen /lookup_minductive. + destruct PCUICAst.PCUICEnvironment.lookup_env eqn:hle => //=. destruct g => //. + eapply erases_global_lookup_env_constant in hle as [? ->]; tea => //. + eapply erases_global_lookup_env_inductive in hle as [? [-> ?]]; tea => //=. + depelim H. + destruct nth_error eqn:hnth => // _. + destruct (nth_error (EAst.ind_bodies _) _) eqn:hnth' => //=. + eapply nth_error_Some_length in hnth'. eapply nth_error_None in hnth. eapply Forall2_length in H. lia. + intros _. + eapply erases_global_lookup_env_none in hle as ->; tea => //. } + now rewrite (lookup_inductive_filter_deps_None wfΣ' nli) in li'. +Qed. + + +Program Definition erase_transform_mapping {guard : abstract_guard_impl} : Transform.t _ _ _ _ PCUICAst.term EAst.term + eval_pcuic_program_mapping (eval_eprogram_env_mapping EWcbvEval.default_wcbv_flags) := + {| name := "erasure"; + pre p := wf_pcuic_inductives_mapping p.2.1 p.1 /\ pre erase_transform p.2 ; + transform p hp := let nhs := proj2 hp in + (p.1, transform erase_transform p.2 nhs) ; + post p := EReorderCstrs.wf_inductives_mapping p.2.1.(EEnvMap.GlobalContextMap.global_decls) p.1 /\ post erase_transform p.2; + obseq p hp p' := obseq erase_transform p.2 (proj2 hp) p'.2 |}. + +Local Obligation Tactic := idtac. +Next Obligation. + cbn -[erase_transform]. intros. + split. 2:eapply (correctness erase_transform). + destruct input as [m prog]; cbn [fst snd] in *. + destruct p. + unfold erase_transform; cbn [transform]. + unfold erase_program, erase_pcuic_program; + set (egf := erase_global_fast _ _ _ _ _); + set (ef := ErasureFunction.erase _ _ _ _ _); + cbn -[egf ef]; unfold erase_global_fast in egf; rewrite /egf; clear egf ef; + set (deps := EAstUtils.term_global_deps _); clearbody deps; + match goal with + |- context [erase_global_deps_fast ?deps ?X_type ?X ?decls (normalization_in:=?normalization_in) ?hprefix ] => + have heq := @erase_global_deps_fast_erase_global_deps deps X_type X decls normalization_in hprefix + end; cbn in heq; specialize (heq (fun Σ eq => f_equal declarations eq)); + destruct heq as [nin' eq]; rewrite eq; clear eq. + rewrite erase_global_deps_erase_global. + cbn -[ErasureFunction.erase_global]. + match goal with + |- context [@erase_global ?X_type ?X ?decls ?SN ?eq] => + have erg := @erases_global_erase_global X_type X _ _ SN eq + end. + cbn in erg. specialize (erg _ eq_refl). + cbn in p. destruct p as [[wt] ?]. + set (eg := erase_global _ _ _) in erg |- *. clearbody eg. + have := erases_wf_inductives_mapping (Σ' := eg) (fst wt) erg i. + intros h. eauto. +Qed. + +Next Obligation. + intros g p v pr ev. cbn. + now apply (preservation erase_transform). +Qed. + +Program Definition pcuic_expand_lets_transform_mapping {cf : checker_flags} K : Transform.t _ _ _ _ PCUICAst.term PCUICAst.term + eval_pcuic_program_mapping eval_pcuic_program_mapping := + {| name := "let expansion in constructors"; + pre p := wf_pcuic_inductives_mapping p.2.1 p.1 /\ pre (pcuic_expand_lets_transform K) p.2 ; + transform p hp := let nhs := proj2 hp in + (p.1, transform (pcuic_expand_lets_transform K) p.2 nhs) ; + post p := wf_pcuic_inductives_mapping p.2.1 p.1 /\ post (pcuic_expand_lets_transform K) p.2; + obseq p hp p' := obseq (pcuic_expand_lets_transform K) p.2 (proj2 hp) p'.2 |}. +Next Obligation. + cbn. intros cf K input []. split => //. + - move: H; rewrite /wf_pcuic_inductives_mapping. + solve_all. move: H; rewrite /wf_pcuic_inductive_mapping. + destruct x as [ind [na tags]]. + rewrite /PCUICAst.PCUICLookup.lookup_inductive /PCUICAst.PCUICLookup.lookup_inductive_gen /PCUICAst.PCUICLookup.lookup_minductive_gen. + rewrite PCUICExpandLetsCorrectness.trans_lookup. destruct PCUICAst.PCUICEnvironment.lookup_env => //. destruct g => //=. + rewrite nth_error_mapi. destruct nth_error => //=. now len. + - unshelve eapply (correctness (pcuic_expand_lets_transform K)). apply H0. +Qed. +Next Obligation. + intros cf K p v pr ev. now eapply (preservation (pcuic_expand_lets_transform K)). +Qed. + Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - Transform.t _ _ - PCUICAst.term EAst.term _ _ + Transform.t _ _ _ EAst.term _ _ PCUICTransform.eval_pcuic_program (EProgram.eval_eprogram final_wcbv_flags) := (* a bunch of nonsense for normalization preconditions *) @@ -219,6 +470,37 @@ Program Definition verified_erasure_pipeline {guard : abstract_guard_impl} (efl (* Erasure of proofs terms in Prop and types *) erase_transform ▷ verified_lambdabox_pipeline. +Next Obligation. + intros guard efl K T1 p. + cbn. intuition eauto. now eapply H2. now eapply H2. +Qed. +Next Obligation. + intros guard efl K T1 p. cbn. intuition auto. +Qed. + +Program Definition verified_erasure_pipeline_mapping {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : + Transform.t _ _ _ EAst.term _ _ + eval_pcuic_program_mapping + (EProgram.eval_eprogram final_wcbv_flags) := + (* 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_mapping (K _ T1) ▷ + (* Erasure of proofs terms in Prop and types *) + erase_transform_mapping ▷ + verified_lambdabox_pipeline_mapping. +Next Obligation. + intros guard efl K T1 p. + cbn. intuition eauto. now eapply H3. now eapply H3. +Qed. +Next Obligation. + intros guard efl K T1 p. cbn. intuition auto. + split; eauto. +Qed. Import EGlobalEnv EWellformed. @@ -236,24 +518,143 @@ Proof. eexists;reflexivity. Qed. -Lemma verified_lambdabox_pipeline_extends {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : - TransformExt.t verified_lambdabox_pipeline (fun p p' => extends (EEnvMap.GlobalContextMap.global_decls p.1) - (EEnvMap.GlobalContextMap.global_decls p'.1)) (fun p p' => extends p.1 p'.1). +Instance verified_lambdabox_pipeline_extends {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : + TransformExt.t verified_lambdabox_pipeline + (fun p p' => + extends (EEnvMap.GlobalContextMap.global_decls p.1) (EEnvMap.GlobalContextMap.global_decls p'.1)) (fun p p' => extends p.1 p'.1). Proof. unfold verified_lambdabox_pipeline. tc. Qed. +Instance verified_lambdabox_pipeline_mapping_extends {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : + TransformExt.t verified_lambdabox_pipeline_mapping + (fun p p' => + p.1 = p'.1 /\ + extends (EEnvMap.GlobalContextMap.global_decls p.2.1) (EEnvMap.GlobalContextMap.global_decls p'.2.1)) (fun p p' => extends p.1 p'.1). +Proof. + unfold verified_lambdabox_pipeline_mapping. tc. +Qed. + 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. unfold verified_lambdabox_pipeline. tc. Qed. -Program Definition pre_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : +Lemma verified_lambdabox_pipeline_mapping_extends' {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : + TransformExt.t verified_lambdabox_pipeline_mapping (fun p p' => p.1 = p'.1 /\ extends_eprogram_env p.2 p'.2) extends_eprogram. +Proof. + unfold verified_lambdabox_pipeline_mapping. tc. +Qed. + +Definition eval_template_program_mapping (p : inductives_mapping * template_program) t := + TemplateProgram.eval_template_program p.2 t. + +Definition eval_template_program_env_mapping (p : inductives_mapping * template_program_env) t := + TemplateProgram.eval_template_program_env p.2 t. + +Definition wf_template_inductive_mapping (Σ : Env.global_env) (m : inductive_mapping) : bool := + let '(ind, (_, cstrs)) := m in + match Template.Ast.lookup_inductive Σ ind with + | Some (mib, oib) => wf_tags_list cstrs #|oib.(Env.ind_ctors)| + | None => true + end. + +Definition wf_template_inductives_mapping Σ m := forallb (wf_template_inductive_mapping Σ) m. + +Program Definition build_template_program_env_mapping {cf : checker_flags} K : + Transform.t _ _ _ _ _ _ eval_template_program_mapping eval_template_program_env_mapping := + {| name := name (build_template_program_env K); + pre p := wf_template_inductives_mapping p.2.1 p.1 /\ pre (build_template_program_env K) p.2 ; + transform p hp := let nhs := proj2 hp in + (p.1, transform (build_template_program_env K) p.2 nhs) ; + post p := wf_template_inductives_mapping p.2.1 p.1 /\ post (build_template_program_env K) p.2; + obseq p hp p' := obseq (build_template_program_env K) p.2 (proj2 hp) p'.2 |}. +Next Obligation. + cbn; intros cf K input [? []]. intuition auto. +Qed. +Next Obligation. + cbn; intros cf K ? ? [? []] ev. cbn. + now unshelve eapply (preservation (build_template_program_env K)). +Qed. + +Lemma lookup_lookup_global_env_None (Σ : GlobalEnvMap.t) kn : + Env.lookup_env Σ kn = None -> + Env.lookup_global (eta_global_declarations Σ (Env.declarations Σ)) kn = None. +Proof. + intros hl. + induction Σ => // /=. induction env => //=. cbn in hl. + rewrite /eta_global_declarations. cbn in *. rewrite lookup_global_map_on_snd hl //. +Qed. + +Lemma eta_lookup_inductive_ctors (Σ : GlobalEnvMap.t) ind : + option_map (fun '(mib, oib) => #|Env.ind_ctors oib|) (Ast.lookup_inductive Σ ind) = option_map (fun '(mib, oib) => #|Env.ind_ctors oib|) (Ast.lookup_inductive (eta_expand_global_env Σ) ind). +Proof. + rewrite /Ast.lookup_inductive /eta_expand_global_env /lookup_inductive_gen /lookup_minductive_gen. + destruct Env.lookup_env eqn:le => //. + cbn. eapply eta_lookup_global in le. rewrite le /=. + destruct g => //=. rewrite nth_error_map. destruct nth_error => //=. len. + cbn. now eapply lookup_lookup_global_env_None in le as ->. +Qed. + +Program Definition eta_expand_mapping K : + Transform.t _ _ _ _ _ _ eval_template_program_env_mapping eval_template_program_mapping := + {| name := name (eta_expand K); + pre p := wf_template_inductives_mapping p.2.1 p.1 /\ pre (eta_expand K) p.2 ; + transform p hp := let nhs := proj2 hp in + (p.1, transform (eta_expand K) p.2 nhs) ; + post p := wf_template_inductives_mapping p.2.1 p.1 /\ post (eta_expand K) p.2; + obseq p hp p' := obseq (eta_expand K) p.2 (proj2 hp) p'.2 |}. +Next Obligation. + cbn; intros K input [? ?]. split. 2:now unshelve eapply (correctness (eta_expand K)). + move: H. rewrite /wf_template_inductives_mapping. + solve_all. move: H. + rewrite /wf_template_inductive_mapping. + destruct x as [ind [na tags]]. + have eqe := eta_lookup_inductive_ctors input.2.1 ind. + do 2 destruct Ast.lookup_inductive as [[? ?]|]; cbn in *; congruence. +Qed. +Next Obligation. + cbn; intros K ? ? [? []] ev. cbn. + now unshelve eapply (preservation (eta_expand K)). +Qed. + +Program Definition template_to_pcuic_mapping K : + Transform.t _ _ _ _ _ _ eval_template_program_mapping eval_pcuic_program_mapping := + {| name := name (template_to_pcuic_transform K); + pre p := wf_template_inductives_mapping p.2.1 p.1 /\ pre (template_to_pcuic_transform K) p.2 ; + transform p hp := let nhs := proj2 hp in + (p.1, transform (template_to_pcuic_transform K) p.2 nhs) ; + post p := wf_pcuic_inductives_mapping p.2.1 p.1 /\ post (template_to_pcuic_transform K) p.2; + obseq p hp p' := obseq (template_to_pcuic_transform K) p.2 (proj2 hp) p'.2 |}. +Next Obligation. + cbn; intros K input [wfm pre]. + pose proof (correctness (template_to_pcuic_transform K) _ pre) as post. + split. 2:now unshelve eapply (correctness (template_to_pcuic_transform K)). + move: wfm. rewrite /wf_template_inductives_mapping. + solve_all. move: H. + rewrite /wf_template_inductive_mapping /wf_pcuic_inductive_mapping. + unfold Ast.lookup_inductive, PCUICAst.PCUICLookup.lookup_inductive, Ast.lookup_inductive_gen, + PCUICAst.PCUICLookup.lookup_inductive_gen, + PCUICAst.PCUICLookup.lookup_minductive_gen, + Ast.lookup_minductive_gen, lookup_minductive_gen. + cbn. destruct post as [[[wtr _]] _]. cbn in wtr. + destruct x as [ind [na tags]]. + destruct pre as [[wta] _]. + have trl := TemplateToPCUICCorrectness.trans_lookup input.2.1 (inductive_mind ind) wta.1 wtr.1. + destruct Env.lookup_env; cbn in *; rewrite trl //. + destruct g => //=. rewrite nth_error_map. destruct nth_error => //=. len. +Qed. +Next Obligation. + cbn; intros K ? ? [? []] ev. cbn. + now unshelve eapply (preservation (template_to_pcuic_transform K)). +Qed. + +Program Definition pre_erasure_pipeline (efl := EWellformed.all_env_flags) : Transform.t _ _ - Ast.term PCUICAst.term _ _ - TemplateProgram.eval_template_program - PCUICTransform.eval_pcuic_program := + _ _ _ _ + 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 @@ -261,7 +662,7 @@ Program Definition pre_erasure_pipeline {guard : abstract_guard_impl} (efl := EW (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 T3 (p:Env.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) ▷ @@ -269,22 +670,75 @@ Program Definition pre_erasure_pipeline {guard : abstract_guard_impl} (efl := EW eta_expand (K _ T3) ▷ (* Casts are removed, application is binary, case annotations are inferred from the global environment *) template_to_pcuic_transform (K _ T2). +Next Obligation. + intros efl K T1 T2 T3 T4 p. cbn. intuition eauto. +Qed. +Next Obligation. + intros efl K T1 T2 T3 T4 p. cbn. intuition eauto. +Qed. -Program Definition erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : Transform.t _ _ - Ast.term EAst.term _ _ - TemplateProgram.eval_template_program +Program Definition pre_erasure_pipeline_mapping (efl := EWellformed.all_env_flags) : + Transform.t _ _ + _ _ _ _ + eval_template_program_mapping + eval_pcuic_program_mapping := + (* 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:Env.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_mapping (K _ T4) ▷ + (* Eta-expand constructors and fixpoint *) + eta_expand_mapping (K _ T3) ▷ + (* Casts are removed, application is binary, case annotations are inferred from the global environment *) + template_to_pcuic_mapping (K _ T2). +Next Obligation. + intros efl K T1 T2 T3 T4 p. cbn. intuition eauto. +Qed. +Next Obligation. + intros efl K T1 T2 T3 T4 p. cbn. intuition eauto. +Qed. + +Program Definition erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : Transform.t _ _ _ _ _ _ + eval_template_program (EProgram.eval_eprogram final_wcbv_flags) := pre_erasure_pipeline ▷ verified_erasure_pipeline. -Program Definition verified_lambdabox_typed_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) econf : - Transform.t _ _ EAst.term EAst.term _ _ - (EProgram.eval_eprogram_env {| with_prop_case := false; with_guarded_fix := true; with_constructor_as_block := false |}) +Next Obligation. + intros guard efl p; cbn. intuition auto. +Qed. + +Program Definition erasure_pipeline_mapping {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) : Transform.t _ _ _ _ _ _ + eval_template_program_mapping + (EProgram.eval_eprogram final_wcbv_flags) := + pre_erasure_pipeline_mapping ▷ + verified_erasure_pipeline_mapping. + +Next Obligation. + intros guard efl p; cbn. intuition auto. +Qed. + +Local Obligation Tactic := program_simpl. + +Program Definition verified_lambdabox_typed_pipeline (efl := EWellformed.all_env_flags) + (wfl := {| with_prop_case := false; with_guarded_fix := true; with_constructor_as_block := false |}) econf : + Transform.t _ _ _ EAst.term _ _ + (eval_eprogram_env_mapping wfl) (EProgram.eval_eprogram final_wcbv_flags) := + (* Reorder constructors according to the mapping *) + reorder_cstrs_transformation (wca := eq_refl) (has_app := eq_refl) wfl (wcon := eq_refl) ▷ + (* Rebuild the efficient lookup table *) + rebuild_wf_env_transform (efl := efl) true 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 (fl := {| with_prop_case := false; with_guarded_fix := true; with_constructor_as_block := false |}) (wcon := eq_refl) eq_refl ▷ + guarded_to_unguarded_fix (fl := {| with_prop_case := false; with_guarded_fix := true; with_constructor_as_block := false |}) (wcon := eq_refl) eq_refl ▷ (* Remove all constructor parameters *) remove_params_optimization (wcon := eq_refl) ▷ (* Rebuild the efficient lookup table *) @@ -319,12 +773,254 @@ Qed. Local Obligation Tactic := intros; eauto. +Definition eval_typed_eprogram_mapping (wfl : WcbvFlags) (p : inductives_mapping × (ExAst.global_env × EAst.term)) t := + eval_typed_eprogram wfl p.2 t. + +Lemma wf_inductive_mapping_trans {Σ Σ' m} : + (forall kn decl, ExAst.lookup_inductive Σ' kn = Some decl -> + exists decl', PCUICAst.lookup_inductive Σ kn = Some decl' /\ #|ind_ctors decl'.2| = #|ExAst.ind_ctors decl|) -> + wf_pcuic_inductive_mapping Σ m -> + wf_inductive_mapping (ExAst.trans_env Σ') m. +Proof. + intros hkn. + rewrite /wf_pcuic_inductive_mapping /wf_inductive_mapping. + destruct m as [ind [na tags]]. + destruct lookup_inductive as [decl|]eqn:hl => //. + have [pdecl [hl' htr]] : exists p, ExAst.lookup_inductive Σ' ind = Some p /\ ExAst.trans_oib p = decl.2. + { move: hl; rewrite /lookup_inductive /ExAst.lookup_inductive. + rewrite /lookup_minductive /= /ExAst.lookup_minductive. + rewrite OptimizeCorrectness.lookup_env_trans_env. + destruct ExAst.lookup_env => //=. destruct g => //=. + rewrite nth_error_map //=. destruct nth_error => //=. intros [= <-]. + now eexists. } + eapply hkn in hl' as [decl' [hl''hct]]. + rewrite hl''hct. destruct decl'; cbn in *. destruct decl. cbn in *. subst o0. + rewrite H. rewrite /ExAst.trans_oib /=. + rewrite /ExAst.trans_ctors. now rewrite map_length. +Qed. + +From Equations Require Import Equations. + +Lemma nth_error_map_In {A B} (l : list A) (f : forall x, In x l -> B) n o : + nth_error (map_In l f) n = Some o -> + exists a (ha : In a l), nth_error l n = Some a /\ o = f a ha. +Proof. + funelim (map_In l f); cbn. + - rewrite nth_error_nil => //. + - destruct n; cbn. + * intros [= <-]; now eexists. + * intros h. have [a [hin []]]:= H _ _ h. + intros -> ->. exists a. exists (or_intror hin). split => //. +Qed. + +Lemma map_In_length {A B} (l : list A) (f : forall x, In x l -> B) : + #|map_In l f| = #|l|. +Proof. funelim (map_In l f); cbn; auto; lia. Qed. + +Lemma lookup_inductive_erase_global_decls_deps {X_type X} {decls : global_env} {prop incl ignore} {kn} {decl} : + ExAst.lookup_inductive (@Erasure.erase_global_decls_deps_recursive X_type X (declarations decls) (universes decls) (retroknowledge decls) prop incl ignore) kn = Some decl -> + exists decl', PCUICAst.lookup_inductive decls kn = Some decl' /\ #|ind_ctors decl'.2| = #|ExAst.ind_ctors decl|. +Proof. + rewrite /ExAst.lookup_inductive /ExAst.lookup_minductive. + destruct decls as [univs' decls retro]. + cbn. unfold ExAst.lookup_env. + induction decls as [|[kn' []]] in X, prop, incl |- *. cbn => //. + - cbn [Erasure.erase_global_decls_deps_recursive] => //. + destruct KernameSet.mem; cbn [ExAst.lookup_env Erasure.erase_global_decl]; + try destruct Erasure.erase_constant_decl => //=; + unfold PCUICAst.lookup_inductive; + unfold PCUICAst.PCUICEnvironment.lookup_env; cbn; + unfold PCUICAst.lookup_inductive_gen, PCUICAst.lookup_minductive_gen. + * unfold ExAst.lookup_env. cbn. + case: eqb_spec. + + intros <-. now cbn. + + intros diff. cbn in IHdecls. eapply IHdecls. + * case: eqb_spec. + + intros <-. now cbn. + + intros diff. eapply IHdecls. + * case: eqb_spec. + + intros heq. + pose proof (PCUICWfEnv.abstract_env_exists X) as [[Σ abs]]. + pose proof (PCUICWfEnv.abstract_env_wf _ abs) as []. + pose proof (prop _ abs). subst Σ. cbn in X0. + eapply wf_fresh_globals in X0. cbn in X0. + depelim X0. clear abs. + match goal with + |- context [@Erasure.erase_global_decls_deps_recursive ?X_type ?X ?decls ?univs ?retro ?prop ?incl ?ignore ] => + set (iprop := prop); set (iincl := incl) + (* have hcor := @ErasureCorrectness.erase_global_decls_deps_recursive_correct X_type X decls univs retro prop incl ignore (fun _ => eq_refl) *) + end. + intros h; specialize (IHdecls _ iprop iincl h). + destruct IHdecls as [decl'' []]. + subst kn'. move: H0. + rewrite /PCUICAst.lookup_inductive /PCUICAst.lookup_inductive_gen. + rewrite /PCUICAst.lookup_minductive /PCUICAst.lookup_minductive_gen. cbn. + destruct lookup_global eqn:hl => //=. + eapply PCUICAst.fresh_global_iff_lookup_global_None in H. rewrite H in hl. congruence. + + intros diff. eapply IHdecls. + - cbn [Erasure.erase_global_decls_deps_recursive] => //. + destruct KernameSet.mem; cbn [ExAst.lookup_env Erasure.erase_global_decl]; + unfold PCUICAst.lookup_inductive; + unfold PCUICAst.PCUICEnvironment.lookup_env; cbn; + unfold PCUICAst.lookup_inductive_gen, PCUICAst.lookup_minductive_gen. + * unfold ExAst.lookup_env. cbn. + case: eqb_spec. + + intros <-. cbn. + destruct nth_error eqn:hmi => //. intros [= <-]. + eapply nth_error_map_In in hmi as [oib [hin [hnth ->]]]. + exists (m, oib); rewrite hnth. split => //=. + now rewrite /Erasure.erase_ind_body map_In_length. + + intros diff. eapply IHdecls. + * case: eqb_spec. + + intros heq. + pose proof (PCUICWfEnv.abstract_env_exists X) as [[Σ abs]]. + pose proof (PCUICWfEnv.abstract_env_wf _ abs) as []. + pose proof (prop _ abs). subst Σ. cbn in X0. + eapply wf_fresh_globals in X0. cbn in X0. + depelim X0. clear abs. + match goal with + |- context [@Erasure.erase_global_decls_deps_recursive ?X_type ?X ?decls ?univs ?retro ?prop ?incl ?ignore ] => + set (iprop := prop); set (iincl := incl) + (* have hcor := @ErasureCorrectness.erase_global_decls_deps_recursive_correct X_type X decls univs retro prop incl ignore (fun _ => eq_refl) *) + end. + intros h; specialize (IHdecls _ iprop iincl h). + destruct IHdecls as [decl'' []]. + subst kn'. move: H0. + rewrite /PCUICAst.lookup_inductive /PCUICAst.lookup_inductive_gen. + rewrite /PCUICAst.lookup_minductive /PCUICAst.lookup_minductive_gen. cbn. + destruct lookup_global eqn:hl => //=. + eapply PCUICAst.fresh_global_iff_lookup_global_None in H. rewrite H in hl. congruence. + + intros diff. eapply IHdecls. +Qed. + +Local Obligation Tactic := program_simpl. + +Program Definition typed_erase_transform_mapping : + Transform.t _ _ _ _ _ _ eval_pcuic_program_mapping (eval_typed_eprogram_mapping default_wcbv_flags) := + {| name := name typed_erase_transform; + pre p := wf_pcuic_inductives_mapping p.2.1 p.1 /\ pre typed_erase_transform p.2 ; + transform p hp := let nhs := proj2 hp in + (p.1, transform typed_erase_transform p.2 nhs) ; + post p := wf_inductives_mapping (ExAst.trans_env p.2.1) p.1 /\ post typed_erase_transform p.2; + obseq p hp p' := obseq typed_erase_transform p.2 (proj2 hp) p'.2 |}. + +Next Obligation. +Proof. + destruct s. cbn -[ExAst.trans_env typed_erase_transform]. split => //. + 2:{ eapply (correctness typed_erase_transform). } + move: i; rewrite /wf_pcuic_inductives_mapping /wf_inductives_mapping. + solve_all. + destruct w as [wf wt]. + eapply wf_inductive_mapping_trans; tea. + intros kn d. + eapply lookup_inductive_erase_global_decls_deps. +Qed. +Next Obligation. + cbn. intros p v pr. eapply (preservation typed_erase_transform). +Qed. + +Lemma lookup_inductive_map_on_snd_remove_decl Σ Σ' kn : + lookup_inductive (map (on_snd (EOptimizePropDiscr.remove_match_on_box_decl Σ')) Σ) kn = lookup_inductive Σ kn. +Proof. + rewrite /lookup_inductive /lookup_minductive. + rewrite EGenericGlobalMap.lookup_env_map_snd. destruct lookup_env => //=. + destruct g => //. +Qed. + +Program Definition remove_match_on_box_typed_transform_mapping {fl} {wcon : with_constructor_as_block = false} {efl : EEnvFlags} {hastrel : has_tRel} {hastbox : has_tBox} : + Transform.t _ _ _ _ _ _ (eval_typed_eprogram_mapping fl) (eval_typed_eprogram_mapping (disable_prop_cases fl)) := + let tr := (remove_match_on_box_typed_transform (fl:=fl) (wcon := wcon) (hastrel := hastrel) (hastbox := hastbox)) in + {| name := name tr; + pre p := wf_inductives_mapping (ExAst.trans_env p.2.1) p.1 /\ pre tr p.2 ; + transform p hp := let nhs := proj2 hp in (p.1, transform tr p.2 nhs) ; + post p := wf_inductives_mapping (ExAst.trans_env p.2.1) p.1 /\ post tr p.2; + obseq p hp p' := obseq tr p.2 (proj2 hp) p'.2 |}. + +Next Obligation. +Proof. + split. 2:{ cbn. now unshelve eapply (correctness remove_match_on_box_typed_transform). } + cbn. + move: i; rewrite /wf_inductives_mapping; solve_all. + move: H; rewrite /wf_inductive_mapping. + rewrite /remove_match_on_box_typed_env. + rewrite -map_trans_env. + destruct x as [ind [na tags]]. now rewrite lookup_inductive_map_on_snd_remove_decl. +Qed. +Next Obligation. + intros p v pr ev. + now unshelve eapply (preservation remove_match_on_box_typed_transform). +Qed. + +Definition eval_typed_eprogram_masks_mapping (fl : WcbvFlags) (p : inductives_mapping × ((option Optimize.dearg_set × ExAst.global_env) × EAst.term)) t := + eval_typed_eprogram_masks fl p.2 t. + +Program Definition dearging_checks_transform_mapping {efl : EEnvFlags} (cf : ETransform.dearging_config) {hastrel : has_tRel} {hastbox : has_tBox} : + Transform.t _ _ _ _ _ _ (eval_typed_eprogram_mapping opt_wcbv_flags) (eval_typed_eprogram_masks_mapping opt_wcbv_flags) := + let tr := (dearging_checks_transform (efl:=efl) cf (hastrel := hastrel) (hastbox := hastbox)) in + {| name := name tr; + pre p := wf_inductives_mapping (ExAst.trans_env p.2.1) p.1 /\ pre tr p.2 ; + transform p hp := let nhs := proj2 hp in (p.1, transform tr p.2 nhs) ; + post p := wf_inductives_mapping (ExAst.trans_env p.2.1.2) p.1 /\ post tr p.2; + obseq p hp p' := obseq tr p.2 (proj2 hp) p'.2 |}. +Next Obligation. +Proof. + split => //. +Qed. +Next Obligation. +Proof. + intros p v pr ev. now unshelve eapply (preservation (dearging_checks_transform cf)). +Qed. + +Program Definition dearging_transform_mapping (cf : ETransform.dearging_config) : + Transform.t _ _ _ _ _ _ (eval_typed_eprogram_masks_mapping opt_wcbv_flags) (eval_eprogram_mapping opt_wcbv_flags) := + let tr := (dearging_transform cf) in + {| name := name tr; + pre p := wf_inductives_mapping (ExAst.trans_env p.2.1.2) p.1 /\ pre tr p.2 ; + transform p hp := let nhs := proj2 hp in (p.1, transform tr p.2 nhs) ; + post p := wf_inductives_mapping p.2.1 p.1 /\ post tr p.2; + obseq p hp p' := obseq tr p.2 (proj2 hp) p'.2 |}. +Next Obligation. +Proof. + split => //. + 2:{ now unshelve eapply (correctness (dearging_transform cf)). } + move: H; rewrite /wf_inductives_mapping; solve_all. + move: H; rewrite /wf_inductive_mapping. destruct x as [ind [na tags]]. + rewrite /dearg. destruct input.2.1.1 => //. + rewrite /OptimizeCorrectness.dearg_env. + rewrite /lookup_inductive /lookup_minductive /=. + rewrite !OptimizeCorrectness.lookup_env_trans_env. + rewrite OptimizeCorrectness.lookup_env_debox_env_types option_map_two /=. + rewrite OptimizeCorrectness.lookup_env_dearg_env option_map_two /=. + destruct ExAst.lookup_env => //=; destruct g => //=. + rewrite !nth_error_map option_map_two /=. + rewrite /Optimize.dearg_mib /=. destruct Optimize.get_mib_masks => //=. + rewrite !nth_error_mapi. + 1-2:destruct nth_error eqn:hnth => //=; len. + intros _. destruct o => //=. destruct p => //. +Qed. +Next Obligation. +Proof. + intros p v pr ev. now unshelve eapply (preservation (dearging_transform cf)). +Qed. + +Program Definition rebuild_wf_env_transform_mapping {fl : WcbvFlags} {efl : EEnvFlags} (with_exp with_fix : bool) : + Transform.t _ _ _ _ _ _ (eval_eprogram_mapping fl) (eval_eprogram_env_mapping fl) := + let tr := @rebuild_wf_env_transform fl efl with_exp with_fix in + {| name := name tr; + pre p := wf_inductives_mapping p.2.1 p.1 /\ pre tr p.2 ; + transform p hp := let nhs := proj2 hp in (p.1, transform tr p.2 nhs) ; + post p := wf_inductives_mapping p.2.1.(EEnvMap.GlobalContextMap.global_decls) p.1 /\ post tr p.2; + obseq p hp p' := obseq tr p.2 (proj2 hp) p'.2 |}. +Next Obligation. +Proof. + intros p v pr ev. now unshelve eapply (preservation (rebuild_wf_env_transform _ _)). +Qed. + Program Definition verified_typed_erasure_pipeline {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) econf : - Transform.t _ _ - PCUICAst.term EAst.term _ _ - PCUICTransform.eval_pcuic_program + Transform.t _ _ _ _ _ _ + eval_pcuic_program_mapping (EProgram.eval_eprogram final_wcbv_flags) := (* a bunch of nonsense for normalization preconditions *) let K ty (T : ty -> _) p @@ -333,32 +1029,30 @@ Program Definition verified_typed_erasure_pipeline {guard : abstract_guard_impl} (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) ▷ + pcuic_expand_lets_transform_mapping (K _ T1) ▷ (* Erasure of proofs terms in Prop and types *) - typed_erase_transform ▷ + typed_erase_transform_mapping ▷ (* Remove match on box early for dearging *) - remove_match_on_box_typed_transform (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ + remove_match_on_box_typed_transform_mapping (wcon := eq_refl) (hastrel := eq_refl) (hastbox := eq_refl) ▷ (* Check if the preconditions for dearging are valid, otherwise dearging will be the identity *) - dearging_checks_transform econf.(dearging_config) (hastrel := eq_refl) (hastbox := eq_refl) ▷ - dearging_transform econf.(dearging_config) ▷ - rebuild_wf_env_transform true true ▷ + dearging_checks_transform_mapping econf.(dearging_config) (hastrel := eq_refl) (hastbox := eq_refl) ▷ + dearging_transform_mapping econf.(dearging_config) ▷ + rebuild_wf_env_transform_mapping true true ▷ verified_lambdabox_typed_pipeline econf. Next Obligation. - cbn in H. split; cbn; intuition eauto. + unfold pre_remove_match_on_box_typed. split; intuition eauto. Qed. Next Obligation. - cbn in H |- *; intuition eauto. + split; intuition auto. Qed. Program Definition typed_erasure_pipeline {guard : abstract_guard_impl} - (efl := EWellformed.all_env_flags) - econf : - Transform.t _ _ - Ast.term EAst.term _ _ - TemplateProgram.eval_template_program + (efl := EWellformed.all_env_flags) econf : + Transform.t _ _ _ _ _ _ + eval_template_program_mapping (EProgram.eval_eprogram final_wcbv_flags) := - pre_erasure_pipeline ▷ + pre_erasure_pipeline_mapping ▷ verified_typed_erasure_pipeline econf. (* At the end of erasure we get a well-formed program (well-scoped globally and localy), without @@ -369,61 +1063,6 @@ Program Definition typed_erasure_pipeline {guard : abstract_guard_impl} Import EGlobalEnv EWellformed. -Program Definition erasure_pipeline_fast {guard : abstract_guard_impl} (efl := EWellformed.all_env_flags) econf := - (* 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_template_program_env (K _ T4) ▷ - eta_expand (K _ T3) ▷ - template_to_pcuic_transform (K _ T2) ▷ - pcuic_expand_lets_transform (K _ T1) ▷ - erase_transform ▷ - 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 false ▷ - 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 false ▷ - 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 false ▷ - constructors_as_blocks_transformation (efl := efl) (has_app := eq_refl) (has_pars := eq_refl) (has_rel := eq_refl) (has_box := eq_refl) (has_cstrblocks := eq_refl) ▷ - optional_unsafe_transforms econf. -Next Obligation. - destruct H; split => //. now eapply ETransform.expanded_eprogram_env_expanded_eprogram_cstrs. -Qed. -Next Obligation. - cbn in H. split; cbn; intuition eauto. -Qed. -Next Obligation. - cbn in H. split; cbn; intuition eauto. -Qed. -Next Obligation. - cbn in H. split; cbn; intuition eauto. -Qed. -Next Obligation. - cbn in H. split; cbn; intuition eauto. -Qed. -Next Obligation. - cbn in H. split; cbn; intuition eauto. -Qed. -Next Obligation. - cbn in H. unfold optional_unsafe_transforms. - cbn. - destruct enable_unsafe as [[] ? ? ? ?]=> //. -Qed. -Next Obligation. - cbn in H. split; cbn; intuition eauto. -Qed. - -Definition run_erase_program_fast {guard : abstract_guard_impl} (econf : erasure_configuration) := - run (erasure_pipeline_fast econf). - Local Open Scope string_scope. Axiom fake_guard_impl_properties : @@ -447,6 +1086,8 @@ Global Existing Instance fake_normalization. the environment is wellformed and the term well-typed (e.g. when it comes directly from a Coq definition). *) +Axiom assume_that_we_only_erase_on_wellformed_inductive_mappings : forall {cf : checker_flags}, + forall (p : Ast.Env.program) (m : inductives_mapping), wf_template_inductives_mapping p.1 m. Axiom assume_that_we_only_erase_on_welltyped_programs : forall {cf : checker_flags}, forall (p : Ast.Env.program), squash (TemplateProgram.wt_template_program p). @@ -454,19 +1095,19 @@ Axiom assume_that_we_only_erase_on_welltyped_programs : forall {cf : checker_fla (* This also optionally runs the cofix to fix translation *) Program Definition run_erase_program {guard : abstract_guard_impl} econf := if econf.(enable_typed_erasure) then run (typed_erasure_pipeline econf) - else if econf.(enable_fast_remove_params) then - run (erasure_pipeline_fast econf) - else run (erasure_pipeline ▷ (optional_unsafe_transforms econf)). + else run (erasure_pipeline_mapping ▷ (optional_unsafe_transforms econf)). Next Obligation. Proof. unfold optional_unsafe_transforms; cbn. destruct enable_unsafe as [[] ? ? ? ?]=> //. Qed. -Program Definition erase_and_print_template_program econf (p : Ast.Env.program) : string := - let p' := run_erase_program econf p _ in +Program Definition erase_and_print_template_program econf (m : inductives_mapping) (p : Ast.Env.program) : string := + let p' := run_erase_program econf (m, p) _ in time "Pretty printing" EPretty.print_program p'. Next Obligation. + split. + now eapply assume_that_we_only_erase_on_wellformed_inductive_mappings. split. now eapply assume_that_we_only_erase_on_welltyped_programs. cbv [PCUICWeakeningEnvSN.normalizationInAdjustUniversesIn]. @@ -474,26 +1115,13 @@ Next Obligation. split; typeclasses eauto. Qed. -Definition erasure_fast_config := - {| enable_unsafe := no_unsafe_passes; - dearging_config := default_dearging_config; - enable_typed_erasure := false; - enable_fast_remove_params := true; - inductives_mapping := []; - inlined_constants := KernameSet.empty |}. - -Program Definition erase_fast_and_print_template_program (p : Ast.Env.program) : string := - erase_and_print_template_program erasure_fast_config p. - Definition typed_erasure_config := {| enable_unsafe := no_unsafe_passes; dearging_config := default_dearging_config; enable_typed_erasure := true; - enable_fast_remove_params := true; - inductives_mapping := []; inlined_constants := KernameSet.empty |}. -(* Parameterized by a configuration for dearging, allowing to, e.g., override masks. *) +(* TODO: Parameterize by a configuration for dearging, allowing to, e.g., override masks. *) Program Definition typed_erase_and_print_template_program (p : Ast.Env.program) : string := - erase_and_print_template_program typed_erasure_config p. + erase_and_print_template_program typed_erasure_config [] p. diff --git a/erasure/theories/ECoInductiveToInductive.v b/erasure/theories/ECoInductiveToInductive.v index 485284539..602321dcb 100644 --- a/erasure/theories/ECoInductiveToInductive.v +++ b/erasure/theories/ECoInductiveToInductive.v @@ -383,6 +383,7 @@ Proof. unfold lookup_inductive in hl. destruct lookup_minductive => //. } - rewrite !GlobalContextMap.lookup_inductive_kind_spec. + move: H2; rewrite /wf_brs. destruct lookup_inductive as [[mdecl idecl]|] eqn:hl => //. assert (map (on_snd (trans Σ)) l = map (on_snd (trans Σ')) l) as -> by solve_all. rewrite (extends_lookup_inductive_kind H0 H1) //. @@ -1013,7 +1014,7 @@ Proof. 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_trans //. + - rewrite /wf_brs; cbn; rewrite lookup_env_trans //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => /= //. diff --git a/erasure/theories/EConstructorsAsBlocks.v b/erasure/theories/EConstructorsAsBlocks.v index c4453bbae..63f188d20 100644 --- a/erasure/theories/EConstructorsAsBlocks.v +++ b/erasure/theories/EConstructorsAsBlocks.v @@ -738,7 +738,7 @@ Proof. 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. + - len. move: H2. rewrite /wf_brs. 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. diff --git a/erasure/theories/EDeps.v b/erasure/theories/EDeps.v index 7bee22452..cfd1c70e3 100644 --- a/erasure/theories/EDeps.v +++ b/erasure/theories/EDeps.v @@ -61,9 +61,10 @@ Proof. - depelim er. econstructor; eauto. - depelim er. - econstructor; eauto. + econstructor; eauto. now len. + clear H3. induction X; [easy|]. - depelim H3. + depelim H4. constructor; [|easy]. now cbn. - depelim er. @@ -115,9 +116,9 @@ Proof. - depelim er. econstructor; eauto. - depelim er. - econstructor; eauto. + econstructor; eauto. now len. clear H3. induction X; [easy|]. - depelim H3. + depelim H4. constructor; [|easy]. now cbn. - depelim er. @@ -175,9 +176,10 @@ Proof. - depelim er. cbn. econstructor; eauto. - depelim er. - econstructor; [easy|easy|easy|easy|easy|]. + econstructor; [easy|easy|easy|easy|easy| | ]. + now len. clear H3. induction X; [easy|]. - depelim H3. + depelim H4. constructor; [|easy]. now cbn. - depelim er. @@ -290,14 +292,14 @@ Proof. unfold EGlobalEnv.iota_red. apply erases_deps_substl. + intuition auto. - apply erases_deps_mkApps_inv in H4. + apply erases_deps_mkApps_inv in H5. now apply Forall_rev, Forall_skipn. + eapply nth_error_forall in e2; [|now eauto]. assumption. - congruence. - depelim er. subst brs; cbn in *. - depelim H3. + depelim H4. cbn in *. apply IHev2. apply erases_deps_substl; [|easy]. @@ -325,9 +327,9 @@ Proof. - depelim er. specialize (IHev1 er). apply erases_deps_mkApps_inv in IHev1 as (? & ?). - depelim H4. + depelim H5. apply IHev2. - econstructor; [easy|easy|easy|easy| |easy]. + econstructor; [easy|easy|easy|easy| |easy|easy ]. apply erases_deps_mkApps; [|easy]. now eapply erases_deps_cunfold_cofix; eauto. - depelim er. @@ -398,6 +400,7 @@ Lemma erases_deps_forall_ind Σ Σ' erases_one_inductive_body idecl idecl' -> erases_deps Σ Σ' discr -> P discr -> + #|idecl'.(ind_ctors)| = #|brs| -> Forall (fun br : _ × Extract.E.term => erases_deps Σ Σ' br.2) brs -> Forall (fun br => P br.2) brs -> P (Extract.E.tCase p discr brs)) @@ -438,7 +441,7 @@ Proof. - eauto. - eapply Hcase; try eassumption. + now apply f. - + revert brs H3. + + clear H3; revert brs H4. fix f' 2. intros brs []; [now constructor|]. constructor; [now apply f|now apply f']. @@ -622,6 +625,14 @@ Proof. repeat eapply conj; try eassumption. cbn in *. now rewrite H8, H9. Qed. +Lemma Forall2_nth_error_left {A B} {P} {l : list A} {l' : list B} : Forall2 P l l' -> + forall n x, nth_error l n = Some x -> + exists x', nth_error l' n = Some x' /\ P x x'. +Proof. + induction 1; destruct n; simpl; auto; try discriminate. + intros x' [= ->]. eexists; eauto. +Qed. + Lemma erases_deps_single Σ Σ' Γ t T et : wf_ext Σ -> Σ;;; Γ |- t : T -> @@ -653,6 +664,11 @@ Proof. econstructor; eauto. destruct H2. destruct x1; eauto. destruct H1. eapply Forall2_All2 in H2. eapply All2_nth_error in H2; eauto. + { eapply Forall2_length in H0. eapply All2i_length in brs_ty. + eapply All2_length in X. rewrite <- H0. + depelim H2. destruct H1 as []. + destruct x1 as []. eapply Forall2_nth_error_left in H2 as [x' []]; tea. + rewrite H4 in H2; noconf H2. depelim H7. eapply Forall2_length in H2. lia. } clear -wf brs_ty X H0 Σer. subst predctx ptm. clear X. @@ -711,14 +727,6 @@ Proof. depelim H0; depelim p1; depelim X; cbn in *; try constructor; cbn; intuition eauto. solve_all. Qed. -Lemma Forall2_nth_error_left {A B} {P} {l : list A} {l' : list B} : Forall2 P l l' -> - forall n x, nth_error l n = Some x -> - exists x', nth_error l' n = Some x' /\ P x x'. -Proof. - induction 1; destruct n; simpl; auto; try discriminate. - intros x' [= ->]. eexists; eauto. -Qed. - Lemma erases_global_all_deps Σ Σ' : wf Σ -> erases_global Σ Σ' -> diff --git a/erasure/theories/EGenericGlobalMap.v b/erasure/theories/EGenericGlobalMap.v index 797eedcf3..d209122f9 100644 --- a/erasure/theories/EGenericGlobalMap.v +++ b/erasure/theories/EGenericGlobalMap.v @@ -324,7 +324,7 @@ Proof. 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 //. + - rewrite /wf_brs. cbn. rewrite lookup_env_gen_transform //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. rewrite gen_transform_inductive_decl_id. diff --git a/erasure/theories/EGenericMapEnv.v b/erasure/theories/EGenericMapEnv.v index c5a0b4853..1fa0d2bed 100644 --- a/erasure/theories/EGenericMapEnv.v +++ b/erasure/theories/EGenericMapEnv.v @@ -279,7 +279,7 @@ Proof. 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 //. + - rewrite /wf_brs; cbn. rewrite lookup_env_gen_transform //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => /= //. diff --git a/erasure/theories/EImplementBox.v b/erasure/theories/EImplementBox.v index 0817566d3..f87ee87fb 100644 --- a/erasure/theories/EImplementBox.v +++ b/erasure/theories/EImplementBox.v @@ -481,7 +481,8 @@ Proof. all: destruct cstr_as_blocks; rtoProp; try split; eauto. + solve_all. + destruct block_args; cbn in *; eauto. - - rewrite lookup_inductive_implement_box. intuition auto. solve_all. + - rewrite /wf_brs; cbn -[lookup_inductive implement_box]. + rewrite lookup_inductive_implement_box. intuition auto. solve_all. solve_all. replace (#|x.1| + S m) with ((#|x.1| + m) + 1) by lia. eapply wellformed_lift. eauto. - rewrite lookup_constructor_implement_box. intuition auto. diff --git a/erasure/theories/EInlineProjections.v b/erasure/theories/EInlineProjections.v index 302b57bc3..c1f930195 100644 --- a/erasure/theories/EInlineProjections.v +++ b/erasure/theories/EInlineProjections.v @@ -74,7 +74,7 @@ Qed. Lemma wellformed_projection_args {efl : EEnvFlags} {Σ p mdecl idecl cdecl pdecl} : wf_glob Σ -> lookup_projection Σ p = Some (mdecl, idecl, cdecl, pdecl) -> - p.(proj_arg) < cdecl.(cstr_nargs). + #|idecl.(ind_ctors)| = 1 /\ p.(proj_arg) < cdecl.(cstr_nargs). Proof. intros wfΣ. rewrite /lookup_projection /lookup_constructor /lookup_inductive. @@ -161,10 +161,10 @@ Section optimize. - rewrite GlobalContextMap.lookup_projection_spec. destruct lookup_projection as [[[[mdecl idecl] cdecl] pdecl]|] eqn:hl; auto => //. simpl. - have arglen := wellformed_projection_args wfΣ hl. + have [ncstrs arglen] := wellformed_projection_args wfΣ hl. apply lookup_projection_lookup_constructor, lookup_constructor_lookup_inductive in hl. - rewrite hl /= andb_true_r. - rewrite IHt //=; len. apply Nat.ltb_lt. + rewrite /wf_brs hl /= andb_true_r. + rewrite IHt //=; len. rewrite ncstrs. apply Nat.ltb_lt. lia. - len. rtoProp; solve_all. solve_all. now eapply isLambda_optimize. solve_all. @@ -795,9 +795,11 @@ Proof. rewrite GlobalContextMap.lookup_projection_spec. destruct lookup_projection as [[[[mdecl idecl] cdecl] pdecl]|] eqn:hl'; auto => //. simpl. + rewrite /wf_brs. rewrite (lookup_constructor_lookup_inductive (lookup_projection_lookup_constructor hl')) /=. rewrite hrel IHt //= andb_true_r. - have hargs' := wellformed_projection_args wfΣ hl'. + have [-> hargs'] := wellformed_projection_args wfΣ hl'. + rtoProp; intuition auto. apply Nat.ltb_lt. len. - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. now eapply isLambda_optimize. - cbn. unfold wf_fix; rtoProp; intuition auto; solve_all. @@ -823,7 +825,7 @@ Proof. 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_optimize //. + - rewrite /wf_brs; cbn; rewrite lookup_env_optimize //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => /= //. diff --git a/erasure/theories/EOptimizePropDiscr.v b/erasure/theories/EOptimizePropDiscr.v index 449e36dc8..8c0f5bd2f 100644 --- a/erasure/theories/EOptimizePropDiscr.v +++ b/erasure/theories/EOptimizePropDiscr.v @@ -378,7 +378,9 @@ Proof. - 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 (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). + assert (iss : isSome (lookup_inductive Σ p.1)). + { move: H2; rewrite /wf_brs; destruct lookup_inductive => //. } + rewrite (extends_inductive_isprop_and_pars H0 H1 iss). destruct inductive_isprop_and_pars as [[[]]|]. destruct map => //. f_equal; eauto. destruct l0 => //. destruct p0 => //. f_equal; eauto. @@ -849,7 +851,7 @@ Proof. 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_remove_match_on_box //. + - rewrite /wf_brs; cbn. rewrite lookup_env_remove_match_on_box //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => /= //. diff --git a/erasure/theories/EProgram.v b/erasure/theories/EProgram.v index 925bc9f76..e05c56212 100644 --- a/erasure/theories/EProgram.v +++ b/erasure/theories/EProgram.v @@ -18,6 +18,9 @@ Local Obligation Tactic := program_simpl. Import EGlobalEnv EWellformed. +Definition inductive_mapping : Set := Kernames.inductive * (bytestring.string * list nat). +Definition inductives_mapping := list inductive_mapping. + Definition eprogram := (EAst.global_context * EAst.term). Definition eprogram_env := (EEnvMap.GlobalContextMap.t * EAst.term). diff --git a/erasure/theories/ERemoveParams.v b/erasure/theories/ERemoveParams.v index 9780299a8..f1a992987 100644 --- a/erasure/theories/ERemoveParams.v +++ b/erasure/theories/ERemoveParams.v @@ -1094,9 +1094,10 @@ Proof. destruct g eqn:hg => /= //. subst g. destruct nth_error => //. destruct nth_error => //. - cbn -[strip]. + rtoProp. move: H2. rewrite /wf_brs; cbn -[strip]. rewrite lookup_env_strip. cbn in H1. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. - destruct nth_error => //. rtoProp; intuition auto. + destruct nth_error => //. rtoProp; intuition auto. len. simp_strip. toAll; solve_all. toAll. solve_all. - cbn -[strip] in H0 |- *. @@ -1142,7 +1143,7 @@ Proof. destruct cstr_as_blocks => //; repeat split; eauto. destruct nth_error => /= //. destruct nth_error => /= //. - - rewrite lookup_env_strip //. + - rewrite /wf_brs; cbn; rewrite lookup_env_strip //. destruct lookup_env eqn:hl => // /=. destruct g eqn:hg => /= //. subst g. destruct nth_error => /= //. diff --git a/erasure/theories/EReorderCstrs.v b/erasure/theories/EReorderCstrs.v index 3e029f23e..71f288ea0 100644 --- a/erasure/theories/EReorderCstrs.v +++ b/erasure/theories/EReorderCstrs.v @@ -1,37 +1,81 @@ -From Coq Require Import List String Arith Lia. +From Coq Require Import List String Arith Lia ssreflect ssrbool Morphisms. Import ListNotations. From Equations Require Import Equations. Set Equations Transparent. From MetaCoq.PCUIC Require Import PCUICAstUtils. From MetaCoq.Utils Require Import MCList bytestring utils monad_utils. -From MetaCoq.Erasure Require Import EPrimitive EAst ESpineView EEtaExpanded EInduction ERemoveParams Erasure EGlobalEnv. +From MetaCoq.Erasure Require Import EProgram EPrimitive EAst ESpineView EEtaExpanded EInduction ERemoveParams Erasure EGlobalEnv + EAstUtils ELiftSubst EWellformed ECSubst EWcbvEval. Import Kernames. Import MCMonadNotation. -Definition inductive_mapping : Set := Kernames.inductive * (bytestring.string * list nat). -Definition inductives_mapping := list inductive_mapping. +Lemma lookup_declared_constructor {Σ id mdecl idecl cdecl} : + lookup_constructor Σ id.1 id.2 = Some (mdecl, idecl, cdecl) -> + declared_constructor Σ id mdecl idecl cdecl. +Proof. + rewrite /lookup_constructor /declared_constructor. + rewrite /declared_inductive /lookup_inductive. + rewrite /declared_minductive /lookup_minductive. + destruct lookup_env => //=. destruct g => //=. + destruct nth_error eqn:hn => //. destruct (nth_error _ id.2) eqn:hn' => //. + intros [= <- <- <-]. intuition auto. +Qed. -Fixpoint lookup_inductive {A} (Σ : list (inductive × A)) (kn : inductive) {struct Σ} : option A := +Fixpoint lookup_inductive_assoc {A} (Σ : list (inductive × A)) (kn : inductive) {struct Σ} : option A := match Σ with | [] => None - | d :: tl => if kn == d.1 then Some d.2 else lookup_inductive tl kn + | d :: tl => if kn == d.1 then Some d.2 else lookup_inductive_assoc tl kn end. -Section Reorder. - Context (Σ : global_declarations). - Context (mapping : inductives_mapping). +Section Tags. - Definition lookup_constructor_mapping i m := - '(tyna, tags) <- lookup_inductive mapping i ;; - List.nth_error tags m. - - Definition lookup_constructor_ordinal i m := - match lookup_constructor_mapping i m with - | None => m - | Some m' => m' + Fixpoint find_tag (l : list nat) (idx : nat) (tag : nat) : option nat := + match l with + | [] => None + | tag' :: tags => if tag == tag' then Some idx + else find_tag tags (S idx) tag end. + (* e.g. for bool: [ 1 0 ], i.e true (0 in Coq) is mapped to 1 and false (1 in Coq) is mapped to 0 *) + Definition new_tag tags tag := find_tag tags 0 tag. + Definition old_tag (tags : list nat) tag := nth_error tags tag. + + (*Lemma old_of_new tags oldidx : + old_tag tags oldidx >>= new_tag tags = Some oldidx. + Proof. + rewrite /old_tag /new_tag. + destruct nth_error eqn:hnth => //=. 2:admit. + revert hnth. + rewrite -{2}[oldidx]Nat.add_0_r. generalize 0. + induction tags in oldidx, n |- *. + - intros n0. now rewrite nth_error_nil. + - cbn. intros n0 hnth. case: eqb_spec. + intros ->. destruct oldidx => //. (* tags are unique *) admit. + intros neq. + destruct oldidx. + * cbn in hnth. now noconf hnth. + * cbn in hnth. rewrite (IHtags oldidx) //. f_equal. lia. + Qed.*) + + Lemma new_tag_spec tags newidx oldidx : + new_tag tags newidx = Some oldidx -> + old_tag tags oldidx = Some newidx. + Proof. + rewrite /old_tag /new_tag. + rewrite -{2}[oldidx]Nat.sub_0_r. generalize 0. + induction tags in oldidx |- *. + - intros n0. now rewrite nth_error_nil. + - cbn. intros n0. case: eqb_spec. + * move=> -> [= ->]. now rewrite Nat.sub_diag. + * intros neq h. specialize (IHtags _ _ h). + case H: (oldidx - n0) => [|n]. + + cbn. assert (oldidx - S n0 = 0). lia. rewrite H0 in IHtags. + destruct tags; cbn in * => //. + noconf IHtags. rewrite eqb_refl in h. noconf h. lia. + + cbn. destruct oldidx; try lia. rewrite Nat.sub_succ in IHtags. + assert (oldidx - n0 = n) by lia. now rewrite H0 in IHtags. + Qed. Definition reorder_list_opt {A} tags (brs : list A) : option (list A) := mapopt (nth_error brs) tags. @@ -39,8 +83,277 @@ Section Reorder. Definition reorder_list {A} tags (l : list A) := option_get l (reorder_list_opt tags l). + Fixpoint one_to_one_map n l := + match n with + | 0 => l == nil + | S n => + existsb (fun x => x == n) l && one_to_one_map n (remove Nat.eq_dec n l) + end. + + Definition wf_tags (l : list nat) := + forallb (fun n => n + Forall (fun tag => tag <> idx) tags. + Proof. + induction tags in n |- *; cbn; auto. + case: eqb_spec => [->|neq] // h. + specialize (IHtags _ h). constructor; auto. + Qed. + + Lemma one_to_one_map_spec l : one_to_one_map #|l| l -> + forall i, i < #|l| <-> In i l. + Proof. + move: (Nat.le_refl #|l|). + generalize #|l| at 2 3 4 as n. induction n in l |- *; cbn. + - move=> hl _ i; split => //. lia. move: l hl; case => //=; try lia. + - move=> hl /andP[] /existsb_exists => [[x [inx xeqn]]] hone i. + specialize (IHn (remove Nat.eq_dec n l)). + apply eqb_eq in xeqn. subst x. + forward IHn. + have hlt := remove_length_lt Nat.eq_dec l n inx. lia. + specialize (IHn hone). + case: (eqb_spec i n) => [->|neq]. intuition auto. + split => hi. assert (i < n) by lia. + now have := proj1 (IHn _) H => /(in_remove Nat.eq_dec). + suff: i < n by lia. apply <- IHn. + apply (in_in_remove Nat.eq_dec) => //. + Qed. + + Lemma find_tag_wf tags idx : + wf_tags tags -> + new_tag tags idx = None -> + #|tags| <= idx. + Proof. + rewrite /wf_tags => /andP[] /forallb_All hwf. + move => /one_to_one_map_spec hone /find_tag_None /Forall_forall hdiff. + case: (Nat.leb_spec #|tags| idx) => // hlt. + elim (hdiff idx). apply hone, hlt. reflexivity. + Qed. + + Lemma mapopt_spec {A B} (f : A -> option B) (l : list A) : + (forall i x, nth_error l i = Some x -> exists x', f x = Some x') -> + exists l', mapopt f l = Some l' /\ + (forall i x, nth_error l i = Some x -> exists x', f x = Some x' /\ nth_error l' i = f x). + Proof. + induction l; cbn. + - intros hf. exists []. split => // i x. rewrite nth_error_nil //. + - intros hf. forward IHl. + { intros i x hx. apply (hf (S i) x hx). } + destruct IHl as [l' [exl' hl']]. + specialize (hf 0 a eq_refl) as [x' eqx']. + destruct (f a) eqn:fa. + * noconf eqx'. rewrite exl'. + eexists; split => //. + intros i x hnth. destruct i; cbn in *. now noconf hnth. + now apply hl'. + * discriminate. + Qed. + + Lemma mapopt_inv_spec {B} (f : nat -> option B) (l : list nat) : + (forall i x, nth_error l i = Some x -> exists x', f x = Some x') -> + exists l', mapopt f l = Some l' /\ #|l| = #|l'| /\ + (forall i x, nth_error l' i = Some x -> exists x', nth_error l i = Some x' /\ f x' = Some x). + Proof. + induction l; cbn. + - intros hf. exists []. split => //; split => // i x. rewrite nth_error_nil //. + - intros hf. forward IHl. + { intros i x hx. apply (hf (S i) x hx). } + destruct IHl as [l' [exl' hl']]. + specialize (hf 0 a eq_refl) as [x' eqx']. + rewrite eqx' exl'. eexists; split => //= //. split => //. lia. + intros. + destruct hl' as [hl hl']. destruct i. cbn in *. noconf H. exists a. now split => //. + now apply hl'. + Qed. + + (*Lemma mapopt_spec' {B} (f : nat -> option B) (l : list nat) : + (forall i x, nth_error l i = Some x -> exists x', f x = Some x') -> + exists l', mapopt f l = Some l' /\ + (forall i x, nth_error l i = Some x -> exists b, f x = Some b /\ nth_error l' x = Some b). + Proof. + induction l; cbn. + - intros hf. exists []. split => // i x. rewrite nth_error_nil //. + - intros hf. forward IHl. + { intros i x hx. apply (hf (S i) x hx). } + destruct IHl as [l' [exl' hl']]. + pose proof (hf 0 a eq_refl) as [x' eqx']. + destruct (f a) eqn:fa => //. + noconf eqx'. rewrite exl'. + eexists; split => //. + intros i x hnth. destruct i; cbn in *. noconf hnth. + * exists x'. split => //. destruct a; cbn. congruence. + apply hl'. + * discriminate. + Qed. *) + + Lemma reorder_list_opt_spec {A} (l : list A) (tags : list nat) (htags : wf_tags tags) : + #|l| = #|tags| -> + exists l', reorder_list_opt tags l = Some l' /\ + (forall i k, old_tag tags i = Some k -> exists x, nth_error l k = Some x /\ nth_error l' i = nth_error l k). + Proof. + rewrite /reorder_list_opt. + rewrite /wf_tags in htags. + intros hlen. + move/andP: htags => [] htags _. + have optH := mapopt_spec (nth_error l) tags. + forward optH. + { intros i x hnth. solve_all. eapply All_nth_error in htags; tea. apply Nat.ltb_lt in htags. + rewrite -hlen in htags. + apply nth_error_Some in htags. destruct (nth_error l x) eqn:hnth'; try congruence. now eexists. } + destruct optH as [l' [heq hl']]. + setoid_rewrite heq. eexists; split => //. + Qed. + + Lemma reorder_list_opt_In {A} (l : list A) (tags : list nat) l' : + reorder_list_opt tags l = Some l' -> + (forall x, In x l' -> In x l). + Proof. + rewrite /reorder_list_opt. + induction tags in l, l' |- *; cbn. + - intros [= <-] x []. + - destruct nth_error eqn:hnth => //. + destruct mapopt eqn:hmap => //. + intros [= <-] x [->|]. + now eapply nth_error_In in hnth. + eapply IHtags; tea. + Qed. + + (* Lemma reorder_list_spec {A} (tags : list nat) (l : list A) n i : + wf_tags tags -> #|l| = #|tags| -> + nth_error tags i = Some n -> + nth_error (reorder_list tags l) i = nth_error l n. + Proof. + intros wft hlt hnth. + rewrite /reorder_list. + now have [l' [-> hnth']] := reorder_list_opt_spec l tags wft hlt. + Qed. *) + + (* Definition reorder_list_spec_maps {A} (l : list A) (tags : list nat) : + forall l', reorder_list_opt tags l = Some l' -> + (forall i k, maps_to tags i k -> nth_error l' k = nth_error l i). + Proof. + intros l'. + induction l; cbn. + - destruct l'; cbn. + intros [=] i k. now rewrite !nth_error_nil. + rewrite /reorder_list_opt /=. + destruct tags => //=. now rewrite nth_error_nil. + - rewrite /reorder_list_opt. *) + + Definition inj_tags (tags : list nat) := + forall i i' k, nth_error tags i = Some k -> nth_error tags i' = Some k -> i = i'. + + Lemma reorder_list_opt_spec' {A} (l : list A) (tags : list nat) (htags : wf_tags tags) : + #|l| = #|tags| -> + (* inj_tags tags -> *) + exists l', reorder_list_opt tags l = Some l' /\ + (forall oldidx a, nth_error l' oldidx = Some a -> + exists newidx, old_tag tags oldidx = Some newidx /\ nth_error l newidx = Some a). + Proof. + rewrite /reorder_list_opt. + rewrite /wf_tags in htags. + move: htags => /andP[] htags hone. + intros hlen. + have optH := mapopt_inv_spec (nth_error l) tags. + forward optH. + { intros i x hnth. solve_all. eapply All_nth_error in htags; tea. apply Nat.ltb_lt in htags. + rewrite -hlen in htags. + apply nth_error_Some in htags. destruct (nth_error l x) eqn:hnth'; try congruence. now eexists. } + destruct optH as [l' [heq hl']]. + setoid_rewrite heq. eexists; split => //. + destruct hl' as [hlen' hl']. + intros newidx a hnth. + specialize (hl' _ _ hnth). + destruct hl' as [x' [eqx' hl']]. exists x'. split => //. + Qed. + + Lemma reorder_list_spec' {A} (tags : list nat) (l : list A) n i x : + wf_tags tags -> #|l| = #|tags| -> + nth_error tags i = Some n -> (* tags[0] = 1 *) + nth_error l n = Some x -> (* l[1] = info *) + nth_error (reorder_list tags l) i = Some x. (* l'[0] = info*) + Proof. + intros wft hlt hnth hnth'. + rewrite /reorder_list. + have [l' [-> hnth'']] := reorder_list_opt_spec l tags wft hlt. + cbn. specialize (hnth'' _ _ hnth). destruct hnth'' as [? []]. congruence. + Qed. + + Lemma reorder_list_spec_inv {A} (tags : list nat) (l : list A) n x : + wf_tags tags -> #|l| = #|tags| -> + nth_error (reorder_list tags l) n = Some x -> (* n is a new index, i an old one *) + exists i, nth_error tags n = Some i /\ nth_error l i = Some x. + Proof. + intros wft hlt. + rewrite /reorder_list. + have [l' [eq hnth'']] := reorder_list_opt_spec' l tags wft hlt; rewrite eq /= => hnth. + specialize (hnth'' _ _ hnth) as [oldidx []]. exists oldidx; now split. + Qed. + + + Lemma reorder_list_old_tag {A} (tags : list nat) (l : list A) oldidx newidx : + wf_tags tags -> #|l| = #|tags| -> + old_tag tags newidx = Some oldidx -> + nth_error (reorder_list tags l) newidx = nth_error l oldidx. + Proof. + rewrite /old_tag. + intros wft hlen ht. + destruct (nth_error l) eqn:hl => //=. + eapply (reorder_list_spec' tags l) in ht; tea. + move: wft => /andP[] wft hone. + solve_all. eapply All_nth_error in wft; tea. + apply Nat.ltb_lt in wft. rewrite -hlen in wft. apply nth_error_None in hl. lia. + Qed. + +(* Lemma reorder_list_old_tag' {A} (tags : list nat) (l : list A) oldidx newidx : + wf_tags tags -> #|l| = #|tags| -> + old_tag tags oldidx = Some newidx -> + nth_error (reorder_list tags l) oldidx = nth_error l newidx. + Proof. + rewrite /old_tag. + intros wft hlen ht. + destruct (nth_error l) eqn:hl => //=. + eapply (reorder_list_spec' tags l) in ht; tea. + move/andP: wft => [] wft hone. + unfold wf_tags in wft. solve_all. eapply All_nth_error in wft; tea. + apply Nat.ltb_lt in wft. rewrite -hlen in wft. apply nth_error_None in hl. lia. + Qed. +*) + + (* Lemma reorder_list_old_tag_inv {A} (tags : list nat) (l : list A) oldidx newidx : + wf_tags tags -> #|l| = #|tags| -> + old_tag tags oldidx = Some newidx -> + nth_error (reorder_list tags l) newidx = nth_error l oldidx. + Proof. + rewrite /old_tag. + intros wft hlen ht. + destruct (nth_error l) eqn:hl => //=. + eapply (reorder_list_spec tags l) in ht; tea. + unfold wf_tags in wft. solve_all. eapply All_nth_error in wft; tea. + apply Nat.ltb_lt in wft. rewrite -hlen in wft. apply nth_error_None in hl. lia. + Qed. *) + + +End Tags. + +Section Reorder. + Context (Σ : global_declarations). + Context (mapping : inductives_mapping). + + Definition lookup_constructor_mapping i c := + '(tyna, tags) <- lookup_inductive_assoc mapping i ;; + new_tag tags c. + + Definition lookup_constructor_ordinal i c := + match lookup_constructor_mapping i c with + | None => c + | Some c' => c' + end. + Definition reorder_branches (i : inductive) (brs : list (list BasicAst.name × term)) : list (list BasicAst.name × term) := - match lookup_inductive mapping i with + match lookup_inductive_assoc mapping i with | None => brs | Some (tyna, tags) => reorder_list tags brs end. @@ -52,11 +365,11 @@ Section Reorder. | tApp fn arg => tApp (reorder fn) (reorder arg) | tConst nm => tConst nm | tConstruct i m args => tConstruct i (lookup_constructor_ordinal i m) (map reorder args) - | tCase i mch brs => tCase i mch (reorder_branches i.1 (map (on_snd reorder) brs)) + | tCase i mch brs => tCase i (reorder mch) (reorder_branches i.1 (map (on_snd reorder) brs)) | tFix mfix idx => tFix (map (map_def reorder) mfix) idx | tCoFix mfix idx => tCoFix (map (map_def reorder) mfix) idx - | tProj (Kernames.mkProjection ind i arg) bod => - tProj (Kernames.mkProjection ind i (lookup_constructor_ordinal ind arg)) (reorder bod) + | tProj p bod => + tProj p (reorder bod) | tPrim p => tPrim (map_prim reorder p) | tLazy t => tLazy (reorder t) | tForce t => tForce (reorder t) @@ -68,25 +381,26 @@ Section Reorder. {| cst_body := option_map reorder cb.(cst_body) |}. Definition reorder_one_ind kn i (oib : one_inductive_body) : one_inductive_body := - match lookup_inductive mapping {| inductive_mind := kn; inductive_ind := i |} with + match lookup_inductive_assoc mapping {| inductive_mind := kn; inductive_ind := i |} with | None => oib | Some (tyna, tags) => {| ind_name := oib.(ind_name); ind_propositional := oib.(ind_propositional); ind_kelim := oib.(ind_kelim); ind_ctors := reorder_list tags oib.(ind_ctors); - ind_projs := reorder_list tags oib.(ind_projs) |} + ind_projs := oib.(ind_projs) |} end. Definition reorder_inductive_decl kn idecl := - {| ind_finite := idecl.(ind_finite); ind_npars := 0; + {| ind_finite := idecl.(ind_finite); ind_npars := idecl.(ind_npars); ind_bodies := mapi (reorder_one_ind kn) idecl.(ind_bodies) |}. Definition reorder_decl d := - match d.2 with - | ConstantDecl cb => (d.1, ConstantDecl (reorder_constant_decl cb)) - | InductiveDecl idecl => (d.1, InductiveDecl (reorder_inductive_decl d.1 idecl)) - end. + let d' := match d.2 with + | ConstantDecl cb => ConstantDecl (reorder_constant_decl cb) + | InductiveDecl idecl => InductiveDecl (reorder_inductive_decl d.1 idecl) + end in + (d.1, d'). Definition reorder_env Σ := map (reorder_decl) Σ. @@ -95,3 +409,1174 @@ End Reorder. Definition reorder_program mapping (p : program) : program := (reorder_env mapping p.1, reorder mapping p.2). + +Definition wf_tags_list (tags : list nat) (n : nat) := + (#|tags| == n) && wf_tags tags. + +Lemma nth_error_reorder {A} {l : list A} {tags n newidx} : + wf_tags_list tags #|l| -> + old_tag tags newidx = Some n -> + nth_error (reorder_list tags l) newidx = nth_error l n. +Proof. + move=> /andP [] h. apply eqb_eq in h. move=> wft hnw. + pose proof (reorder_list_old_tag tags l n newidx wft). + now apply H. +Qed. + +Definition wf_reordering ncstrs cstrs := + (List.length cstrs == ncstrs) && + one_to_one_map ncstrs cstrs. + +Definition wf_inductive_mapping (Σ : global_declarations) (m : inductive_mapping) : bool := + let '(ind, (_, cstrs)) := m in + match EGlobalEnv.lookup_inductive Σ ind with + | Some (mib, oib) => wf_tags_list cstrs #|oib.(ind_ctors)| + | None => true + end. + +Definition wf_inductives_mapping Σ (m : inductives_mapping) : bool := + forallb (wf_inductive_mapping Σ) m. + +Section reorder_proofs. + Context {efl : EEnvFlags}. + Context {wca : cstr_as_blocks = false}. + Context (Σ : global_declarations) (m : inductives_mapping). + Context (wfm : wf_inductives_mapping Σ m). + + Notation optimize := (reorder m). + + Lemma optimize_mkApps f l : optimize (mkApps f l) = mkApps (optimize f) (map optimize l). + Proof using Type. + induction l using rev_ind; simpl; auto. + now rewrite mkApps_app /= IHl map_app /= mkApps_app /=. + Qed. + + Lemma map_optimize_repeat_box n : map optimize (repeat tBox n) = repeat tBox n. + Proof using Type. by rewrite map_repeat. Qed. + + (* move to globalenv *) + + Lemma isLambda_optimize t : isLambda t -> isLambda (optimize t). + Proof. destruct t => //. Qed. + Lemma isBox_optimize t : isBox t -> isBox (optimize t). + Proof. destruct t => //. Qed. + + Lemma lookup_env_reorder kn : lookup_env (reorder_env m Σ) kn = option_map (fun decl => snd (reorder_decl m (kn, decl))) (lookup_env Σ kn). + Proof. + clear wfm. induction Σ; cbn; auto. + case: eqb_spec => [->|hneq] //=. + Qed. + + Lemma lookup_constant_reorder c : option_map (reorder_constant_decl m) (lookup_constant Σ c) = lookup_constant (reorder_env m Σ) c. + Proof. + rewrite /lookup_constant lookup_env_reorder. + destruct lookup_env => //=. + destruct g => //. + Qed. + + Lemma lookup_inductive_reorder i : + lookup_inductive (reorder_env m Σ) i = option_map (fun '(mib, oib) => + (reorder_inductive_decl m i.(inductive_mind) mib, reorder_one_ind m i.(inductive_mind) i.(inductive_ind) oib)) + (lookup_inductive Σ i). + Proof. + unfold lookup_inductive, lookup_minductive. + unfold reorder_env. rewrite lookup_env_reorder. + destruct lookup_env => //=. + destruct g => //. + cbn. rewrite nth_error_mapi. + destruct nth_error => //=. + Qed. + + Lemma lookup_inductive_assoc_spec {ind p} : + wf_inductives_mapping Σ m -> + lookup_inductive_assoc m ind = Some p -> + wf_inductive_mapping Σ (ind, p). + Proof. + clear. rewrite /wf_inductives_mapping. + induction m; cbn -[lookup_inductive wf_inductive_mapping] => //. + destruct eq_inductive eqn:heq => //. + - move=> /andP[] wfa wfi. intros [= <-]. + apply eqb_eq in heq. subst ind. destruct a; apply wfa. + - move=> /andP[] wfa wfi. intros hl. now apply IHi. + Qed. + + Lemma lookup_inductive_assoc_wf_tags {ind mib oib p} : + lookup_inductive Σ ind = Some (mib, oib) -> + lookup_inductive_assoc m ind = Some p -> + wf_tags_list (snd p) #|oib.(ind_ctors)|. + Proof. + move=> hl. + move/(lookup_inductive_assoc_spec wfm). + rewrite /wf_inductive_mapping hl. + now destruct p. + Qed. + + Lemma ind_ctors_reorder {ind mib oib p} : + lookup_inductive Σ ind = Some (mib, oib) -> + lookup_inductive_assoc m ind = Some p -> + ind_ctors (reorder_one_ind m ind.(inductive_mind) ind.(inductive_ind) oib) = + reorder_list p.2 (ind_ctors oib). + Proof. + intros hli hia. + unfold reorder_one_ind. + destruct ind; rewrite hia. destruct p. now cbn. + Qed. + + Lemma ind_ctors_no_reorder {ind mib oib} : + lookup_inductive Σ ind = Some (mib, oib) -> + lookup_inductive_assoc m ind = None -> + ind_ctors (reorder_one_ind m ind.(inductive_mind) ind.(inductive_ind) oib) = + ind_ctors oib. + Proof. + intros hli hia. + unfold reorder_one_ind. + now destruct ind; rewrite hia. + Qed. + + Lemma nth_reorder_list_ctors {i mib oib na tags} : + lookup_inductive Σ i = Some (mib, oib) -> + lookup_inductive_assoc m i = Some (na, tags) -> + forall newidx n, old_tag tags newidx = Some n -> + nth_error (reorder_list tags oib.(ind_ctors)) newidx = nth_error oib.(ind_ctors) n. + Proof. + move=> hli hia idx hnew. + apply nth_error_reorder => //. + eapply (lookup_inductive_assoc_wf_tags (p:=(na, tags))); tea. + Qed. + + Lemma map_opt_length {A B} (f : A -> option B) l : + match mapopt f l with + | None => True + | Some l' => #|l'| = #|l| + end. + Proof. + induction l; cbn; auto. + destruct (f a) => //. + destruct (mapopt f l) => //=. now f_equal. + Qed. + + Lemma reorder_list_length {A} tags (l : list A): #|tags| = #|l| -> #|reorder_list tags l| = #|l|. + Proof. + move=> hl. rewrite /reorder_list. + case hr: reorder_list_opt => [l'|] //=. + move: hr; rewrite /reorder_list_opt. + have := map_opt_length (nth_error l) tags. + destruct mapopt => //. congruence. + Qed. + + Lemma lookup_inductive_assoc_None_reorder_one_ind i oib : + lookup_inductive_assoc m i = None -> + reorder_one_ind m (inductive_mind i) (inductive_ind i) oib = oib. + Proof. + rewrite /reorder_one_ind. destruct i; move => -> //. + Qed. +(* + Lemma lookup_inductive_assoc_None_reorder_inductive_decl i mib : + lookup_inductive_assoc m i = None -> + reorder_inductive_decl m (inductive_mind i) mib = mib. + Proof. + rewrite /reorder_inductive_decl. + intros hl. destruct mib; f_equal. + cbn; induction ind_bodies => //=. f_equal; eauto. + rewrite (lookup_inductive_assoc_None_reorder_one_ind {| inductive_mind := inductive_mind i; inductive_ind := 0 |}). //. + + Qed. *) + Arguments eqb_eq {A H x y}. + Lemma lookup_constructor_reorder i n : + option_map (fun '(mib, oib, c) => (reorder_inductive_decl m (inductive_mind i) mib, reorder_one_ind m (inductive_mind i) (inductive_ind i) oib, c)) (lookup_constructor Σ i n) = + lookup_constructor (reorder_env m Σ) i (lookup_constructor_ordinal m i n). + Proof. + rewrite /lookup_constructor lookup_inductive_reorder. + destruct lookup_inductive as [[mib oib]|] eqn:hind => //=. + destruct nth_error eqn:hnth => //=. + destruct (lookup_inductive_assoc m i) as [(na, tags)|] eqn:hl. + have /andP[] := lookup_inductive_assoc_wf_tags hind hl => //= /eqb_eq => hlen wft. + rewrite (ind_ctors_reorder hind hl). cbn. + destruct (nth_error _ (lookup_constructor_ordinal _ _ _)) eqn:hnth'. + rewrite /lookup_constructor_ordinal /lookup_constructor_mapping in hnth'. + rewrite hl /= in hnth'. + destruct (new_tag tags n) as [newidx|] eqn:ht. + - eapply new_tag_spec in ht. + now rewrite (nth_reorder_list_ctors hind hl _ n) in hnth'. + - move/nth_error_Some_length: hnth. + rewrite /new_tag in ht => hn. + eapply find_tag_wf in ht => //. lia. + - eapply nth_error_None in hnth'. + apply nth_error_Some_length in hnth. + rewrite reorder_list_length in hnth' => //. + move: hnth'. + rewrite /lookup_constructor_ordinal /lookup_constructor_mapping hl /=. + destruct (new_tag tags n) as [newidx|] eqn:ht. + eapply new_tag_spec in ht. + apply nth_error_Some_length in ht. lia. + rewrite /new_tag in ht => hn. + eapply find_tag_wf in ht => //. lia. + - rewrite /lookup_constructor_ordinal /lookup_constructor_mapping hl //=. + destruct i. + rewrite /reorder_one_ind hl //=. now rewrite hnth. + - apply nth_error_None in hnth. + rewrite /lookup_constructor_ordinal /lookup_constructor_mapping /=. + destruct lookup_inductive_assoc eqn:hl => //=. + have /andP[] := lookup_inductive_assoc_wf_tags hind hl => //= /(eqb_eq (A:=nat)) => hlen wft. + destruct p as [na tags]. + destruct (new_tag tags n) as [newidx|] eqn:ht. + + eapply new_tag_spec in ht. + move/andP: wft => [] ht' _. + eapply forallb_All in ht'. eapply All_nth_error in ht'; tea. apply Nat.ltb_lt in ht'. lia. + + rewrite /reorder_one_ind. destruct i; rewrite hl //=. + destruct nth_error eqn:hnth' => //. + apply nth_error_Some_length in hnth'. rewrite reorder_list_length in hnth'. cbn in hlen; lia. lia. + + rewrite lookup_inductive_assoc_None_reorder_one_ind //=. + destruct nth_error eqn:hnth' => //. + apply nth_error_Some_length in hnth'. lia. + Qed. + + Lemma ind_projs_reorder mind ind oib : + ind_projs (reorder_one_ind m mind ind oib) = ind_projs oib. + Proof. + rewrite /reorder_one_ind. destruct lookup_inductive_assoc as [[na tags]|]=> //. + Qed. + + Lemma wf_glob_ind_projs {p pinfo} : + wf_glob Σ -> + lookup_projection Σ p = Some pinfo -> + #|(ind_ctors pinfo.1.1.2)| = 1. + Proof. + intros wf. + rewrite /lookup_projection /lookup_constructor /lookup_inductive /lookup_minductive. + destruct lookup_env eqn:hl => //. + have := lookup_env_wellformed wf hl. + destruct g as [|mib] => //=. + destruct nth_error eqn:hind => //=. + destruct ind_ctors eqn:hctors => //=. + destruct (nth_error (ind_projs o) _) eqn:hprojs => //=. + intros wfmind [= <-]. cbn. + move: wfmind; rewrite /wf_minductive. + move/andP=> [] _ /forallb_All ha. + eapply All_nth_error in ha; tea. + move: ha; rewrite /wf_inductive /wf_projections. + destruct ind_projs => //. now rewrite nth_error_nil in hprojs. + rewrite hctors. destruct l => //. + Qed. + + + Lemma lookup_projection_ordinal p : + wf_glob Σ -> + isSome (lookup_projection Σ p) -> + lookup_constructor_ordinal m (proj_ind p) 0 = 0. + Proof. + move=> wf. + case hl: lookup_projection => [pro|] //= _. + have wfpro := wf_glob_ind_projs wf hl. move: hl. + rewrite /lookup_projection /lookup_constructor_ordinal. + destruct lookup_constructor as [[[mib oib] cb]|] eqn:hlc => //=. + destruct nth_error eqn:nthp => //=. intros [= <-]. cbn in wfpro. + rewrite /lookup_constructor_mapping. + destruct lookup_inductive_assoc as [[na tags]|] eqn:hla => //=. + destruct new_tag eqn:hnt => //=. + eapply new_tag_spec in hnt. + eapply lookup_inductive_assoc_spec in hla; tea. + rewrite /wf_inductive_mapping in hla. + eapply lookup_constructor_lookup_inductive in hlc. rewrite hlc in hla. + move/andP: hla. rewrite wfpro. rewrite /wf_tags => [] [] taglen /andP[] /forallb_All ht. + destruct tags as [|] => //. destruct tags as [|] => //. + destruct n => //. cbn in hnt. now rewrite nth_error_nil in hnt. + Qed. + + Lemma lookup_projection_reorder p : + wf_glob Σ -> + isSome (lookup_projection Σ p) -> + lookup_projection (reorder_env m Σ) p = option_map (fun '(((mib, oib), c), pb) => + (reorder_inductive_decl m p.(proj_ind).(inductive_mind) mib, + reorder_one_ind m p.(proj_ind).(inductive_mind) p.(proj_ind).(inductive_ind) oib, + c, pb)) + (lookup_projection Σ p). + Proof. + intros wf iss. unfold lookup_projection. + rewrite -{1}(lookup_projection_ordinal _ wf iss) -lookup_constructor_reorder. + destruct lookup_constructor eqn:hlc => //=. + destruct p0 as [[mib oib] cb] => //=. + rewrite ind_projs_reorder //=. + destruct nth_error eqn:nthp => //=. + Qed. + + Lemma All_reorder_list {A tags} {P : A -> Prop} {l} : + All P l -> + All P (reorder_list tags l). + Proof. + rewrite /reorder_list. + destruct reorder_list_opt as [l'|] eqn:hre => //=. + have hin := reorder_list_opt_In _ _ _ hre. + move/(All_Forall). + move/(Forall_forall _ _) => hl. + apply Forall_All, Forall_forall. intuition eauto. + Qed. + + Lemma wf_ind_mapping_wf_brs {ind n nas tags} : wf_brs Σ ind n -> + lookup_inductive_assoc m ind = Some (nas, tags) -> + #|tags| = n. + Proof. + rewrite /wf_brs. + destruct lookup_inductive as [[mib oib]|] eqn:hli => //. + move=> /eqb_eq hlen hla. + have := lookup_inductive_assoc_wf_tags hli hla. + rewrite /wf_tags_list /= => /andP[] /eqb_eq hlt _. lia. + Qed. + + Ltac rtoProp ::= + repeat match goal with + | H : is_true (_ && _) |- _ => apply andb_and in H; destruct H + | |- context [is_true (_ && _)] => rewrite andb_and + | H : is_true (_ || _) |- _ => move/orP: H => H; destruct H + | |- context [is_true (_ || _)] => apply/orP + end. + + + + Lemma wf_optimize t k : + wf_glob Σ -> + wellformed Σ k t -> wellformed (reorder_env m Σ) k (optimize t). + Proof using Type wfm wca. + intros wfΣ. + induction t in k |- * using EInduction.term_forall_list_ind; simpl; auto; + intros; try easy; + rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + unfold wf_fix_gen, test_def in *; + simpl closed in *; try solve [simpl subst; simpl closed; f_equal; auto; bool; solve_all]; try easy. + - bool. rewrite -lookup_constant_reorder. destruct lookup_constant => //=; bool. + now destruct (cst_body c) => //=. + - rewrite wca in H *. move/andP: H => [] /andP[] -> iss isnil /=. + rewrite -lookup_constructor_reorder. + destruct lookup_constructor eqn:hl => //=. + destruct args => //. + - rtoProp; intuition auto; solve_all. + * rewrite /reorder_branches. + destruct lookup_inductive_assoc as [[na tags]|] eqn:hl => //=. + have lenreo := wf_ind_mapping_wf_brs H0 hl. + rewrite reorder_list_length. len. len. + move: H0. rewrite /wf_brs. destruct p as [[mind ind] i]. + rewrite lookup_inductive_reorder. destruct lookup_inductive as [[mib oib]|]=> //=. + rewrite /reorder_one_ind hl /=. move/eqb_eq => hl'. rewrite reorder_list_length //. lia. + now apply Nat.eqb_eq. len. + move: H0. rewrite /wf_brs. destruct p as [[mind ind] i]. + rewrite lookup_inductive_reorder. destruct lookup_inductive as [[mib oib]|]=> //=. + rewrite /reorder_one_ind hl /=. move/eqb_eq => hl'. now apply Nat.eqb_eq. + * rewrite /reorder_branches. + destruct lookup_inductive_assoc as [[nas tags]|]. + eapply All_reorder_list. + all:solve_all. + - rtoProp; intuition auto. + rewrite lookup_projection_reorder //. + destruct lookup_projection => //. + - rtoProp; intuition auto; solve_all. + destruct (dbody x) => //. + - rtoProp; intuition auto; solve_all. + depelim H1; constructor; eauto. intuition auto. + cbn; solve_all. + Qed. + + Instance mapopt_ext {A B} : + Proper (pointwise_relation A eq ==> eq ==> eq) (@mapopt A B). + Proof. + intros f g eqfg l ? <-. + induction l; cbn; auto. + rewrite (eqfg a). destruct (g a) => //. + now rewrite IHl. + Qed. + + Lemma mapopt_option_map {A B C} (f : A -> option B) (g : B -> C) l : + mapopt (fun x => option_map g (f x)) l = option_map (map g) (mapopt f l). + Proof. + induction l; cbn; auto. + destruct (f a) => //=. + rewrite IHl. destruct (mapopt f l) => //. + Qed. + + Lemma reorder_list_opt_map {A B} tags (f : A -> B) (l : list A) : + reorder_list_opt tags (map f l) = option_map (map f) (reorder_list_opt tags l). + Proof. + rewrite /reorder_list_opt. + have req : pointwise_relation nat eq (nth_error (map f l)) (fun x => option_map f (nth_error l x)). + { intros x. now rewrite nth_error_map. } + setoid_rewrite req. now rewrite mapopt_option_map. + Qed. + + Lemma reorder_branches_map i f brs : + reorder_branches m i (map f brs) = + map f (reorder_branches m i brs). + Proof. + rewrite /reorder_branches. + destruct lookup_inductive_assoc as [[na tags]|] eqn:hl => //. + rewrite /reorder_list. + rewrite reorder_list_opt_map. + destruct (reorder_list_opt tags brs) => //=. + Qed. + + Lemma optimize_csubst {a k b} n : + wf_glob Σ -> + wellformed Σ (k + n) b -> + optimize (ECSubst.csubst a k b) = ECSubst.csubst (optimize a) k (optimize b). + Proof using Type wfm wca. + intros wfΣ. + induction b in k |- * using EInduction.term_forall_list_ind; simpl; auto; + intros wft; try easy; + rewrite -> ?map_map_compose, ?compose_on_snd, ?compose_map_def, ?map_length; + unfold wf_fix, test_def in *; + simpl closed in *; try solve [simpl subst; simpl closed; f_equal; auto; rtoProp; solve_all]; try easy. + - destruct (k ?= n0)%nat; auto. + - f_equal. rtoProp. rewrite wca in H0. now destruct args; inv H0. + - move/andP: wft => [] hasc /andP[] /andP[] hi hb hl. rewrite IHb. f_equal. unfold on_snd; solve_all. + repeat toAll. f_equal. solve_all. + rewrite -!reorder_branches_map map_map_compose. cbn. f_equal. + unfold on_snd; cbn. + solve_all. f_equal. unfold optimize in *. + rewrite a0 //. red; rewrite -b0. lia_f_equal. + - move/andP: wft => [] /andP[] hasf hp /andP[] hb hwfm. + f_equal. solve_all. + rewrite /map_def; destruct x => //=. f_equal. + apply a0; cbn in *. red; rewrite -b0. lia_f_equal. + - move/andP: wft => [] hasco /andP[] hp hb. + f_equal. solve_all. + rewrite /map_def; destruct x => //=. f_equal. + apply a0; cbn in *. red; rewrite -b. lia_f_equal. + Qed. + + Lemma optimize_substl s t : + wf_glob Σ -> + forallb (wellformed Σ 0) s -> + wellformed Σ #|s| t -> + optimize (substl s t) = substl (map optimize s) (optimize t). + Proof using Type wfm wca. + intros wfΣ. induction s in t |- *; simpl; auto. + move/andP => [] cla cls wft. + rewrite IHs //. eapply wellformed_csubst => //. + f_equal. rewrite (optimize_csubst (S #|s|)) //. + Qed. + + Lemma optimize_iota_red pars args br : + wf_glob Σ -> + forallb (wellformed Σ 0) args -> + wellformed Σ #|skipn pars args| br.2 -> + optimize (EGlobalEnv.iota_red pars args br) = EGlobalEnv.iota_red pars (map optimize args) (on_snd optimize br). + Proof using Type wfm wca. + intros wfΣ wfa wfbr. + unfold EGlobalEnv.iota_red. + rewrite optimize_substl //. + rewrite forallb_rev forallb_skipn //. + now rewrite List.rev_length. + 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). + Proof using Type. + unfold EGlobalEnv.fix_subst. + rewrite map_length. + generalize #|mfix|. + induction n; simpl; auto. + f_equal; auto. + Qed. + + Lemma optimize_cofix_subst mfix : EGlobalEnv.cofix_subst (map (map_def optimize) mfix) = map optimize (EGlobalEnv.cofix_subst mfix). + Proof using Type. + unfold EGlobalEnv.cofix_subst. + rewrite map_length. + generalize #|mfix|. + induction n; simpl; auto. + f_equal; auto. + Qed. + + Lemma optimize_cunfold_fix mfix idx n f : + wf_glob Σ -> + wellformed Σ 0 (tFix mfix idx) -> + cunfold_fix mfix idx = Some (n, f) -> + cunfold_fix (map (map_def optimize) mfix) idx = Some (n, optimize f). + Proof using Type wfm wca. + intros wfΣ hfix. + unfold cunfold_fix. + rewrite nth_error_map. + cbn in hfix. move/andP: hfix => [] /andP[] hasfix hlam /andP[] hidx hfix. + destruct nth_error eqn:hnth => //. + intros [= <- <-] => /=. f_equal. + rewrite optimize_substl //. eapply wellformed_fix_subst => //. + rewrite fix_subst_length. + eapply nth_error_forallb in hfix; tea. now rewrite Nat.add_0_r in hfix. + now rewrite optimize_fix_subst. + Qed. + + Lemma optimize_cunfold_cofix mfix idx n f : + wf_glob Σ -> + wellformed Σ 0 (tCoFix mfix idx) -> + cunfold_cofix mfix idx = Some (n, f) -> + cunfold_cofix (map (map_def optimize) mfix) idx = Some (n, optimize f). + Proof using Type wfm wca. + intros wfΣ hfix. + unfold cunfold_cofix. + rewrite nth_error_map. + cbn in hfix. move/andP: hfix => [] hasfix /andP[] hidx hfix. + destruct nth_error eqn:hnth => //. + intros [= <- <-] => /=. f_equal. + rewrite optimize_substl //. eapply wellformed_cofix_subst => //. + rewrite cofix_subst_length. + eapply nth_error_forallb in hfix; tea. now rewrite Nat.add_0_r in hfix. + now rewrite optimize_cofix_subst. + Qed. + +End reorder_proofs. + +Import EGlobalEnv EExtends. + +Lemma extends_lookup_projection {efl : EEnvFlags} {Σ Σ' p} : extends Σ Σ' -> wf_glob Σ' -> + isSome (lookup_projection Σ p) -> + lookup_projection Σ p = lookup_projection Σ' p. +Proof. + intros ext wf. + unfold lookup_projection. + destruct lookup_constructor as [[[mdecl idecl] cdecl]|] eqn:hl => //. + simpl. + rewrite (extends_lookup_constructor wf ext _ _ _ hl) //. +Qed. + +(* +Lemma wellformed_optimize_extends {wfl: EEnvFlags} {Σ} t : + forall n, EWellformed.wellformed Σ n t -> + forall {Σ'}, extends Σ Σ' -> wf_glob Σ' -> + optimize Σ t = optimize Σ' t. +Proof. + induction t using EInduction.term_forall_list_ind; cbn -[lookup_constant lookup_inductive + GlobalContextMap.lookup_projection]; intros => //. + all:unfold wf_fix_gen in *; rtoProp; intuition auto. + 5:{ destruct cstr_as_blocks; rtoProp. f_equal; eauto; solve_all. destruct args; cbn in *; eauto. } + all:f_equal; eauto; solve_all. + - rewrite !GlobalContextMap.lookup_projection_spec. + rewrite -(extends_lookup_projection H0 H1 H3). + destruct lookup_projection as [[[[]]]|]. f_equal; eauto. + now cbn in H3. +Qed. + +Lemma wellformed_reorder_decl_extends {wfl: EEnvFlags} {Σ} t : + wf_global_decl Σ t -> + forall {Σ' : GlobalContextMap.t}, extends Σ Σ' -> wf_glob Σ' -> + reorder_decl Σ t = reorder_decl Σ' t. +Proof. + destruct t => /= //. + intros wf Σ' ext wf'. f_equal. unfold optimize_constant_decl. f_equal. + destruct (cst_body c) => /= //. f_equal. + now eapply wellformed_optimize_extends. +Qed. + +Lemma lookup_env_reorder_env_Some {efl : EEnvFlags} {Σ} kn d : + wf_glob Σ -> + GlobalContextMap.lookup_env Σ kn = Some d -> + ∑ Σ' : GlobalContextMap.t, + [× extends_prefix Σ' Σ, wf_global_decl Σ' d & + lookup_env (reorder_env Σ) kn = Some (reorder_decl Σ' d)]. +Proof. + rewrite GlobalContextMap.lookup_env_spec. + destruct Σ as [Σ map repr wf]. + induction Σ in map, repr, wf |- *; simpl; auto => //. + intros wfg. + case: eqb_specT => //. + - intros ->. cbn. intros [= <-]. + exists (GlobalContextMap.make Σ (fresh_globals_cons_inv wf)). split. + now eexists [_]. + cbn. now depelim wfg. + f_equal. symmetry. eapply wellformed_reorder_decl_extends. cbn. now depelim wfg. + 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 Σ')). + cbn in IHΣ. 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. + eapply map_ext_in => kn hin. unfold on_snd. f_equal. + symmetry. eapply wellformed_reorder_decl_extends => //. cbn. + eapply lookup_env_In in hin. 2:now depelim wfg. + depelim wfg. eapply lookup_env_wellformed; tea. + 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). +Proof. + induction Σ; cbn; auto. + case: eqb_spec => //. +Qed. + +Lemma lookup_env_reorder_env_None {efl : EEnvFlags} {Σ} kn : + GlobalContextMap.lookup_env Σ kn = None -> + lookup_env (reorder_env Σ) kn = None. +Proof. + rewrite GlobalContextMap.lookup_env_spec. + destruct Σ as [Σ map repr wf]. + cbn. intros hl. rewrite lookup_env_map_snd hl //. +Qed. +*) + +Section reorder_mapping. + Context {efl : EEnvFlags}. + Context {wca : cstr_as_blocks = false}. + + Context (mapping : inductives_mapping). + Context (Σ : global_context). + Context (wfm : wf_inductives_mapping Σ mapping). + Notation reorder := (reorder mapping). + Notation reorder_decl := (reorder_decl mapping). + Notation reorder_env := (reorder_env mapping). + +Lemma is_propositional_optimize ind : + wf_glob Σ -> + inductive_isprop_and_pars Σ ind = inductive_isprop_and_pars (reorder_env Σ) ind. +Proof. + rewrite /inductive_isprop_and_pars => wf. + rewrite /lookup_inductive /lookup_minductive. + rewrite lookup_env_reorder. + destruct lookup_env as [[decl|]|] => //=. + rewrite nth_error_mapi. destruct nth_error => //=. + destruct o => //=. rewrite /reorder_one_ind /=. + destruct lookup_inductive_assoc as [[na tags]|]=> //=. +Qed. + +Lemma lookup_inductive_pars_optimize ind : + wf_glob Σ -> + EGlobalEnv.lookup_inductive_pars Σ ind = EGlobalEnv.lookup_inductive_pars (reorder_env Σ) ind. +Proof. + rewrite /lookup_inductive_pars => wf. + rewrite /lookup_inductive /lookup_minductive. + rewrite lookup_env_reorder. + destruct lookup_env as [[decl|]|] => //. +Qed. + +Lemma is_propositional_cstr_optimize ind c : + wf_glob Σ -> + constructor_isprop_pars_decl Σ ind c = constructor_isprop_pars_decl (reorder_env Σ) ind (lookup_constructor_ordinal mapping ind c). +Proof. + rewrite /constructor_isprop_pars_decl => wf. + rewrite -lookup_constructor_reorder //. + destruct lookup_constructor as [[[mib oib] cb]|] => //=. + rewrite /reorder_one_ind. + destruct lookup_inductive_assoc as [[na tags]|]=> //. +Qed. + +Lemma closed_iota_red pars c args brs br : + forallb (closedn 0) args -> + nth_error brs c = Some br -> + #|skipn pars args| = #|br.1| -> + closedn #|br.1| br.2 -> + closed (iota_red pars args br). +Proof. + intros clargs hnth hskip clbr. + rewrite /iota_red. + eapply ECSubst.closed_substl => //. + now rewrite forallb_rev forallb_skipn. + now rewrite List.rev_length hskip Nat.add_0_r. +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 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 {hasapp : 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 substl_closed s t : closed t -> substl s t = t. +Proof. + induction s in t |- *; cbn => //. + intros clt. rewrite csubst_closed //. now apply IHs. +Qed. + +Lemma substl_rel s k a : + closed a -> + nth_error s k = Some a -> + substl s (tRel k) = a. +Proof. + intros cla. + induction s in k |- *. + - rewrite nth_error_nil //. + - destruct k => //=. + * intros [= ->]. rewrite substl_closed //. + * intros hnth. now apply IHs. +Qed. + +(* EEnvFlags is reorder_env_flags *) +Lemma optimize_correct {fl} t v {withc : with_constructor_as_block = false} : + has_tApp -> + wf_glob Σ -> + @eval fl Σ t v -> + wellformed Σ 0 t -> + @eval fl (reorder_env Σ) (reorder t) (reorder v). +Proof. + intros hastapp wfΣ ev. + induction ev; simpl in *. + + - move/andP => [] /andP[] hasapp cla clt. econstructor; eauto. + - move/andP => [] /andP[] hasapp clf cla. + eapply eval_wellformed in ev2; tea => //. + eapply eval_wellformed in ev1; tea => //. + econstructor; eauto. + move: ev1 => /= /andP[] hasl ev1. + rewrite -(optimize_csubst _ _ wfm 1) //. + apply IHev3. eapply wellformed_csubst => //. + + - move/andP => [] /andP[] haslet clb0 clb1. + intuition auto. + eapply eval_wellformed in ev1; tea => //. + forward IHev2 by eapply wellformed_csubst => //. + econstructor; eauto. rewrite -(optimize_csubst _ _ wfm 1) //. + + - move/andP => [] hascase /andP[] /andP[] hl wfd wfbrs. + rewrite optimize_mkApps in IHev1. + eapply eval_wellformed in ev1 => //. + move/(@wf_mkApps hastapp): ev1 => [] wfc' wfargs. + eapply nth_error_forallb in wfbrs; tea. + rewrite Nat.add_0_r in wfbrs. + forward IHev2. eapply wellformed_iota_red; tea => //. + rewrite (optimize_iota_red _ _ wfm) in IHev2 => //. now rewrite e4. + econstructor; eauto. + rewrite -is_propositional_cstr_optimize //. tea. + rewrite reorder_branches_map nth_error_map. + rewrite /reorder_branches. + rewrite /lookup_constructor_ordinal /lookup_constructor_mapping. + destruct lookup_inductive_assoc as [[na tags]|] eqn:hla => //=. + rewrite /wf_brs in hl. + destruct lookup_inductive as [[mib oib]|] eqn:li => //. apply Nat.eqb_eq in hl. + have wftags := lookup_inductive_assoc_wf_tags _ _ wfm li hla. + have wfmap := lookup_inductive_assoc_spec _ _ wfm hla. + have wfbrl : #|brs| = #|ind_ctors oib|. lia. + have wftbrs : wf_tags_list tags #|brs|. + { now move: wftags; rewrite /wf_tags_list -wfbrl. } + destruct new_tag eqn:hnewt => //=. + * eapply new_tag_spec in hnewt. + rewrite (nth_error_reorder wftbrs hnewt) e2 //. + * move: wftbrs => /andP[] htbrs wftag. + eapply find_tag_wf in wftag; tea. apply nth_error_Some_length in e2. + apply eqb_eq in htbrs. lia. + * rewrite e2 //. + * len. + * len. + + - congruence. + + - move/andP => [] hascase /andP[] /andP[] hl wfd wfbrs. + eapply eval_wellformed in ev1 => //. + move: hl; rewrite /wf_brs. + destruct lookup_inductive as [[mib oib]|] eqn:li => //. move/Nat.eqb_eq. + len. cbn in wfbrs. subst brs. cbn in wfbrs. rewrite Nat.add_0_r in wfbrs. + move/andP: wfbrs => [] wfbrs _. cbn; intros hlen. + forward IHev2. eapply wellformed_substl; tea => //. + rewrite forallb_repeat //. len. + eapply eval_iota_sing; eauto. + rewrite -is_propositional_optimize //. + rewrite /reorder_branches. + destruct lookup_inductive_assoc as [[na tags]|] eqn:hla => //=. + have wftags := lookup_inductive_assoc_wf_tags _ _ wfm li hla. + have wfmap := lookup_inductive_assoc_spec _ _ wfm hla. + have wftbrs : wf_tags_list tags #|[(n, f4)]|. + { move: wftags; rewrite /wf_tags_list //=. now rewrite hlen. } + rewrite //=. + move: wftbrs => /andP[] /Nat.eqb_eq //=. + destruct tags => //=. destruct tags => //= _. + rewrite /wf_tags => /andP[] ht _. move: ht => //= /andP[] /Nat.ltb_lt. + destruct n0 => //. destruct n0 => //. cbn -[substl]. + rewrite (optimize_substl Σ) in IHev2 => //. + * now rewrite forallb_repeat. + * now len. + * now rewrite map_repeat in IHev2. + + - move/andP => [] /andP[] hasapp clf cla. rewrite optimize_mkApps in IHev1. + simpl in *. + eapply eval_wellformed in ev1 => //. + move/(@wf_mkApps hastapp): ev1 => [] wff wfargs. + eapply eval_fix; eauto. + rewrite map_length. + unshelve (eapply optimize_cunfold_fix; tea); tea. + rewrite optimize_mkApps in IHev3. apply IHev3. + rewrite wellformed_mkApps // wfargs. + eapply eval_wellformed in ev2; tas => //. rewrite ev2 /= !andb_true_r. + rewrite hastapp. + eapply wellformed_cunfold_fix; tea. + + - move/andP => [] /andP[] hasapp clf cla. + eapply eval_wellformed in ev1 => //. + move/(@wf_mkApps hastapp): ev1 => [] clfix clargs. + eapply eval_wellformed in ev2; tas => //. + rewrite optimize_mkApps in IHev1 |- *. + simpl in *. eapply eval_fix_value. auto. auto. auto. + unshelve (eapply optimize_cunfold_fix; eauto); tea. + now rewrite map_length. + + - move/andP => [] /andP[] hasapp clf cla. + eapply eval_wellformed in ev1 => //. + eapply eval_wellformed in ev2; tas => //. + simpl in *. eapply eval_fix'. auto. auto. + unshelve (eapply optimize_cunfold_fix; eauto); tea. + eapply IHev2; tea. eapply IHev3. + apply/andP; split => //. + rewrite hastapp. + eapply wellformed_cunfold_fix; tea. + + - move/andP => [] hascase /andP[] /andP[] hl cd clbrs. specialize (IHev1 cd). + eapply eval_wellformed in ev1; tea => //. + move/(@wf_mkApps hastapp): ev1 => [] wfcof wfargs. + forward IHev2. + rewrite hl wellformed_mkApps // /= wfargs clbrs !andb_true_r. + rewrite hascase. + eapply wellformed_cunfold_cofix; tea => //. + rewrite !optimize_mkApps /= in IHev1, IHev2. + eapply eval_cofix_case. tea. + unshelve (eapply optimize_cunfold_cofix; tea); tea. + exact IHev2. + + - move/andP => [] /andP[] hasproj hl hd. + destruct lookup_projection as [[[[mdecl idecl] cdecl] pdecl]|] eqn:hl' => //. + eapply eval_wellformed in ev1 => //. + move/(@wf_mkApps hastapp): ev1 => [] wfcof wfargs. + forward IHev2. + { rewrite /= wellformed_mkApps // wfargs andb_true_r hasproj andb_true_r. + eapply wellformed_cunfold_cofix; tea. } + rewrite optimize_mkApps /= in IHev1. + eapply eval_cofix_proj. eauto. + unshelve (eapply optimize_cunfold_cofix; tea); tea. + rewrite optimize_mkApps in IHev2 => //. + + - rewrite /declared_constant in isdecl. + rewrite /lookup_constant. rewrite isdecl /= => _. + destruct decl; cbn in e; subst cst_body. + econstructor. + rewrite /declared_constant. + rewrite lookup_env_reorder isdecl //=. cbn. reflexivity. + eapply IHev. + eapply lookup_env_wellformed in wfΣ; tea. + move: wfΣ. now rewrite /wf_global_decl /=. + + - move=> /andP[] /andP[] hasproj iss cld. + eapply eval_wellformed in ev1; tea => //. + move/(@wf_mkApps hastapp): ev1 => [] wfc wfargs. + destruct lookup_projection as [[[[mdecl idecl] cdecl'] pdecl]|] eqn:hl' => //. + pose proof (lookup_projection_lookup_constructor hl'). + rewrite (constructor_isprop_pars_decl_constructor H) in e1. noconf e1. + forward IHev1 by auto. + forward IHev2. eapply nth_error_forallb in wfargs; tea. + rewrite optimize_mkApps /= in IHev1. + have lp := lookup_projection_reorder _ _ wfm p wfΣ. + forward lp. now rewrite hl'. + rewrite hl' /= in lp. + have lpo := (lookup_projection_ordinal _ _ wfm p wfΣ). + forward lpo by now rewrite hl'. + rewrite lpo in IHev1. + eapply eval_proj; tea. + rewrite -lpo -is_propositional_cstr_optimize //. + rewrite /constructor_isprop_pars_decl // H //= // H0 H1 //. + len. rewrite nth_error_map e3 //. + + - congruence. + + - move/andP => [] /andP[] hasproj hl hd. + eapply eval_proj_prop; eauto. + rewrite -is_propositional_optimize //. + + - move/andP=> [] /andP[] hasapp clf cla. + rewrite optimize_mkApps. + eapply eval_construct; tea. + rewrite -lookup_constructor_reorder //. now rewrite e0 //=. + rewrite optimize_mkApps in IHev1. now eapply IHev1. + now len. + now eapply IHev2. + + - congruence. + + - move/andP => [] /andP[] hasapp clf cla. + specialize (IHev1 clf). specialize (IHev2 cla). + eapply eval_app_cong; eauto. + eapply eval_to_value in ev1. + destruct ev1; simpl in *; eauto. + * depelim a0. + + destruct t => //; rewrite optimize_mkApps /=. + + now rewrite /= !orb_false_r orb_true_r in i. + * destruct with_guarded_fix. + + move: i. + rewrite !negb_or. + rewrite optimize_mkApps !isFixApp_mkApps !isConstructApp_mkApps !isPrimApp_mkApps + !isLazyApp_mkApps. + destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. + rewrite !andb_true_r. + rtoProp; intuition auto; destruct v => /= //. + + move: i. + rewrite !negb_or. + rewrite optimize_mkApps !isConstructApp_mkApps !isPrimApp_mkApps !isLazyApp_mkApps. + destruct args using rev_case => // /=. rewrite map_app !mkApps_app /= //. + destruct v => /= //. + - intros; rtoProp; intuition eauto. + depelim X; repeat constructor. + eapply All2_over_undep in a. + eapply All2_Set_All2 in ev. eapply All2_All2_Set. primProp. + subst a0 a'; cbn in *. depelim H0; cbn in *. intuition auto; solve_all. + primProp; depelim H0; intuition eauto. + - move=> /andP[] haslazy wf. econstructor; eauto. eapply IHev2. + eapply eval_wellformed in ev1; tea. move/andP: ev1 => []; tea => //. + - destruct t => //. + all:constructor; eauto. + cbn [atom reorder] in i |- *. + rewrite -lookup_constructor_reorder //. + destruct args. 2:cbn in *; eauto. + cbn -[lookup_constructor]. rtoProp; intuition auto. + destruct lookup_constructor => //. +Qed. +End reorder_mapping. + +Lemma wf_inductive_mapping_inv {efl : EEnvFlags} d Σ m : + wf_glob (d :: Σ) -> + wf_inductives_mapping (d :: Σ) m -> + (match d.2 with + | InductiveDecl mib => + (forall i oib, nth_error mib.(ind_bodies) i = Some oib -> + match lookup_inductive_assoc m {| inductive_mind := d.1; inductive_ind := i |} with + | Some (na, tags) => wf_tags_list tags #|oib.(ind_ctors)| + | None => true + end) + | ConstantDecl _ => True + end) /\ wf_inductives_mapping Σ m. +Proof. + rewrite /wf_inductives_mapping. + intros wfΣ wfm. split; revgoals. solve_all. + move: H; rewrite /wf_inductive_mapping. + destruct x as [ind [na tags]]. + destruct lookup_inductive as [[mib oib]|] eqn:li => //=. + depelim wfΣ; cbn. + move: li. cbn. case: eqb_spec. + * intros <-. + destruct d0 => //. destruct nth_error eqn:hnth => //. + intros [= -> ->]. + destruct lookup_env eqn:hle. + have:= lookup_env_Some_fresh hle. contradiction. + auto. + * intros neq. destruct lookup_env => //=. destruct g => //. + destruct nth_error => //. now intros [= -> ->]. + * depelim wfΣ. move: li. cbn; case: eqb_spec. + + intros <-. destruct d0 => //. + destruct lookup_env eqn:hle => //. + have:= lookup_env_Some_fresh hle. contradiction. + destruct lookup_env eqn:hle => //. + have:= lookup_env_Some_fresh hle. contradiction. + + intros neq. destruct lookup_env => //=. destruct g => //. + destruct nth_error => //. + * depelim wfΣ; cbn. + destruct d0 => //. + intros i oib hnth. + destruct lookup_inductive_assoc as [[na tags]|] eqn:hla => //. + have hla' := lookup_inductive_assoc_wf_tags _ _ wfm _ hla. + eapply hla'. + rewrite /lookup_inductive /lookup_minductive /=. + rewrite eqb_refl. rewrite hnth. reflexivity. +Qed. + +From MetaCoq.Erasure Require Import EEtaExpanded. + +Lemma optimize_expanded {Σ m} t : + wf_inductives_mapping Σ m -> expanded Σ t -> expanded (reorder_env m Σ) (reorder m t). +Proof. + intros wfm. + induction 1 using expanded_ind. + all:try solve[constructor; eauto; solve_all]. + all:rewrite ?optimize_mkApps. + - eapply expanded_mkApps_expanded => //. solve_all. + - cbn. econstructor; eauto. + rewrite /reorder_branches. + destruct lookup_inductive_assoc => //; solve_all. + eapply All_reorder_list. solve_all. + - cbn. eapply expanded_tFix. solve_all. + rewrite isLambda_optimize //. + - cbn. + eapply declared_constructor_lookup in H. + have hl := lookup_constructor_reorder _ _ wfm ind idx. rewrite H /= in hl. + eapply expanded_tConstruct_app; tea. + eapply lookup_declared_constructor. now symmetry. + now len. now solve_all. +Qed. + +Lemma optimize_expanded_decl {Σ m kn t} : + wf_inductives_mapping Σ m -> + expanded_decl Σ t -> expanded_decl (reorder_env m Σ) (reorder_decl m (kn, t)).2. +Proof. + destruct t as [[[]]|] => /= //. + unfold expanded_constant_decl => /=. + intros. now apply optimize_expanded. +Qed. + +Lemma reorder_env_expanded {efl : EEnvFlags} {Σ m} : + wf_inductives_mapping Σ m -> + wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (reorder_env m Σ). +Proof. + intros wfm. + unfold expanded_global_env; move=> wfg exp. + induction exp; cbn; constructor; auto. + cbn in IHexp. + unshelve eapply IHexp; tea. eapply wf_inductive_mapping_inv; tea. now depelim wfg. cbn. + unshelve eapply (optimize_expanded_decl). now eapply wf_inductive_mapping_inv. + now destruct decl. +Qed. + +From MetaCoq.Erasure Require Import EEtaExpandedFix. + +Lemma optimize_expanded_fix {Σ Γ m} t : + wf_inductives_mapping Σ m -> expanded Σ Γ t -> expanded (reorder_env m Σ) Γ (reorder m t). +Proof. + intros wfm. + induction 1 using expanded_ind. + all:try solve[constructor; eauto; solve_all]. + all:rewrite ?optimize_mkApps. + - cbn. eapply expanded_tRel_app; tea. len. solve_all. + - cbn. econstructor; eauto. + 2:solve_all. + destruct f3 => //. + - cbn. econstructor; eauto. + rewrite /reorder_branches. + destruct lookup_inductive_assoc => //; solve_all. + eapply All_reorder_list. solve_all. + - cbn. eapply expanded_tFix. solve_all. + rewrite isLambda_optimize //. + now rewrite rev_map_spec map_map_compose /= -rev_map_spec. + solve_all. destruct args => //. rewrite nth_error_map /= H4 //. + len. + - cbn. + eapply declared_constructor_lookup in H. + have hl := lookup_constructor_reorder _ _ wfm ind idx. rewrite H /= in hl. + eapply expanded_tConstruct_app; tea. + eapply lookup_declared_constructor. now symmetry. + now len. now solve_all. +Qed. + +Lemma optimize_expanded_decl_fix {Σ m kn t} : + wf_inductives_mapping Σ m -> + expanded_decl Σ t -> expanded_decl (reorder_env m Σ) (reorder_decl m (kn, t)).2. +Proof. + destruct t as [[[]]|] => /= //. + unfold expanded_constant_decl => /=. + intros. now apply optimize_expanded_fix. +Qed. + +Lemma reorder_env_expanded_fix {efl : EEnvFlags} {Σ m} : + wf_inductives_mapping Σ m -> + wf_glob Σ -> expanded_global_env Σ -> expanded_global_env (reorder_env m Σ). +Proof. + intros wfm. + unfold expanded_global_env; move=> wfg exp. + induction exp; cbn; constructor; auto. + cbn in IHexp. + unshelve eapply IHexp; tea. eapply wf_inductive_mapping_inv; tea. now depelim wfg. cbn. + unshelve eapply (optimize_expanded_decl_fix). now eapply wf_inductive_mapping_inv. + now destruct decl. +Qed. + +Lemma optimize_extends_env {efl : EEnvFlags} {Σ Σ'} m : + wf_inductives_mapping Σ' m -> + extends Σ Σ' -> + wf_glob Σ -> + wf_glob Σ' -> + extends (reorder_env m Σ) (reorder_env m Σ'). +Proof. + intros hast ext wf wf'. + rewrite /extends => kn decl. + rewrite !lookup_env_reorder. + specialize (ext kn). + destruct lookup_env eqn:hle => //. specialize (ext _ eq_refl). + rewrite ext. auto. +Qed. + +Lemma reorder_env_wf {efl : EEnvFlags} {Σ m} + {wcb : cstr_as_blocks = false} : + wf_inductives_mapping Σ m -> + wf_glob Σ -> wf_glob (reorder_env m Σ). +Proof. + intros wfg wfΣ. + induction wfΣ; cbn. constructor; eauto. + have wfdΣ : (wf_glob ((kn, d) :: Σ)) by constructor; auto. + have [wfd wfinv] := (wf_inductive_mapping_inv _ _ _ wfdΣ wfg). cbn in wfd. + constructor; eauto. destruct d; cbn. + - unfold wf_global_decl in H. cbn in H. + destruct (cst_body c) => //=. cbn in H. + unshelve (eapply wf_optimize); eauto. + - cbn in H. rewrite /wf_minductive in H. cbn in H. + rewrite /wf_minductive /=. + move/andP: H => [] -> H /=. solve_all. + have hfa := (forall_nth_error_Alli (fun i oib => + is_true (match lookup_inductive_assoc m {| inductive_mind := kn; inductive_ind := i |} with + | Some p => let (_, tags) := p in wf_tags_list tags #|ind_ctors oib| + | None => true + end)) 0 (ind_bodies m0) wfd). clear wfd. + eapply Alli_All_mix in hfa; tea. clear H. + eapply (Alli_All (P:=fun _ x => wf_inductive x) (n:=0)); eauto. + eapply (fst (Alli_mapi _ _ _)). + eapply Alli_impl; tea. cbn. + intros n x [hla wf]. + rewrite /reorder_one_ind. + destruct lookup_inductive_assoc as [[na tags]|] eqn:hla' => //. + move: wf. rewrite /wf_inductive /wf_projections /=. + destruct ind_projs => //. destruct ind_ctors => //. destruct l0 => //=. + move: hla; rewrite /wf_tags_list. destruct tags => //=. + destruct tags => //=. rewrite /wf_tags. move/andP=> [] hf _. + cbn in hf. rewrite andb_true_r in hf. apply Nat.leb_le in hf. + have -> : n0 = 0 by lia. now cbn. + - cbn. apply ErasureFunction.fresh_global_In. + rewrite map_map_compose /=. intros hin'. apply ErasureFunction.fresh_global_In in H0. + now apply H0. +Qed. + +From MetaCoq.Erasure Require Import EProgram. + +Definition reorder_program_wf {efl : EEnvFlags} {wca : cstr_as_blocks = false} (p : eprogram) m (wfm : wf_inductives_mapping p.1 m) : + wf_eprogram efl p -> wf_eprogram efl (reorder_program m p). +Proof. + intros []; split. + now unshelve eapply reorder_env_wf. + cbn. now eapply (@wf_optimize _ wca). +Qed. + +Definition reorder_program_expanded {efl : EEnvFlags} (p : eprogram) m (wfm : wf_inductives_mapping p.1 m) : + wf_eprogram efl p -> + expanded_eprogram_cstrs p -> expanded_eprogram_cstrs (reorder_program m p). +Proof. + unfold expanded_eprogram_env_cstrs. + move=> [wfe wft] /andP[] etae etat. + apply/andP; split. + cbn. eapply EEtaExpanded.expanded_global_env_isEtaExp_env, reorder_env_expanded => //. + now eapply isEtaExp_env_expanded_global_env. + eapply EEtaExpanded.expanded_isEtaExp. + now apply optimize_expanded, EEtaExpanded.isEtaExp_expanded. +Qed. + +Definition reorder_program_expanded_fix {efl : EEnvFlags} (p : eprogram) m (wfm : wf_inductives_mapping p.1 m) : + wf_eprogram efl p -> + expanded_eprogram p -> expanded_eprogram (reorder_program m p). +Proof. + unfold expanded_eprogram. + move=> [wfe wft] [] etae etat. + split. now eapply reorder_env_expanded_fix. + now eapply optimize_expanded_fix. +Qed. diff --git a/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v b/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v index 47d73387e..5202cef2f 100644 --- a/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v +++ b/erasure/theories/EWcbvEvalCstrsAsBlocksFixLambdaInd.v @@ -618,7 +618,6 @@ Proof. split; intros; rtoProp; intuition auto; solve_all. - red. move=> hascase n ci discr brs. simpl. - destruct lookup_inductive eqn:hl => /= //. intros; rtoProp; intuition auto; solve_all. - red. simpl. move=> hasproj n p discr wf discr' wf'. simpl. rtoProp; intuition auto. diff --git a/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v b/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v index 955c25c0b..bfedd9f57 100644 --- a/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v +++ b/erasure/theories/EWcbvEvalCstrsAsBlocksInd.v @@ -517,6 +517,7 @@ Proof. split; intros; rtoProp; intuition auto; solve_all. - red. move=> hascase n ci discr brs. simpl. + rewrite /wf_brs. destruct lookup_inductive eqn:hl => /= //. intros; rtoProp; intuition auto; solve_all. - red. simpl. move=> hasproj n p discr wf discr' wf'. diff --git a/erasure/theories/EWcbvEvalNamed.v b/erasure/theories/EWcbvEvalNamed.v index a3ac397de..8fcd94a95 100644 --- a/erasure/theories/EWcbvEvalNamed.v +++ b/erasure/theories/EWcbvEvalNamed.v @@ -1118,12 +1118,11 @@ Proof. repeat split. len. solve_all. destruct cstr_as_blocks; eauto. rtoProp; intuition eauto. solve_all. destruct args => //. - - destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. + - move: H0; rewrite /wf_brs; cbn; destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. erewrite lookup_annotate_env; eauto. cbn. - destruct nth_error as [ [] | ]; cbn in *; eauto. - - destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. - destruct nth_error as [ [] | ]; cbn in *; eauto. - repeat split. eauto. + destruct nth_error as [ [] | ]; cbn in *; eauto. len. + - move: H0; rewrite /wf_brs; cbn; destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. + destruct nth_error as [ [] | ]; cbn in *; eauto. intros eq. solve_all. rewrite map_length. rewrite <- app_length. eapply a; eauto. len. rewrite gen_many_fresh_length. eauto. - destruct lookup_env as [ [] | ] eqn:E; cbn in *; eauto. diff --git a/erasure/theories/EWellformed.v b/erasure/theories/EWellformed.v index 25037ce0e..84f333032 100644 --- a/erasure/theories/EWellformed.v +++ b/erasure/theories/EWellformed.v @@ -107,7 +107,8 @@ Section wf. - all occuring constructors are defined, - all occuring constants are defined, and - all occuring fixpoints have lambdas as bodies - - if has_axioms is false, all occuring constants have bodies *) + - if has_axioms is false, all occuring constants have bodies + - all cases are exhaustive w.r.t the global declaration *) Definition wf_fix_gen (wf : nat -> term -> bool) k mfix idx := let k' := List.length mfix + k in @@ -115,6 +116,12 @@ Section wf. Definition is_nil {A} (l : list A) := match l with [] => true | _ => false end. + Definition wf_brs ind brsl := + match lookup_inductive Σ ind with + | Some (mib, oib) => #|oib.(ind_ctors)| == brsl + | None => false + end. + Fixpoint wellformed k (t : term) : bool := match t with | tRel i => has_tRel && Nat.ltb i k @@ -124,7 +131,7 @@ Section wf. | tLetIn na b b' => has_tLetIn && wellformed k b && wellformed (S k) b' | tCase ind c brs => has_tCase && let brs' := List.forallb (fun br => wellformed (#|br.1| + k) br.2) brs in - isSome (lookup_inductive Σ ind.1) && wellformed k c && brs' + wf_brs ind.1 #|brs| && wellformed k c && brs' | tProj p c => has_tProj && isSome (lookup_projection Σ p) && wellformed k c | tFix mfix idx => has_tFix && List.forallb (isLambda ∘ dbody) mfix && wf_fix_gen wellformed k mfix idx | tCoFix mfix idx => has_tCoFix && wf_fix_gen wellformed k mfix idx @@ -440,6 +447,19 @@ Proof. intros wfΣ' Hl; apply Hl. Qed. + +Lemma extends_lookup_inductive {efl} {Σ Σ'} : + wf_glob Σ' -> extends Σ Σ' -> + forall ind b, lookup_inductive Σ ind = Some b -> + lookup_inductive Σ' ind = Some b. +Proof. + intros wf ex ind b. + rewrite /lookup_inductive /lookup_minductive. + destruct lookup_env eqn:lookup => //=. + now rewrite (extends_lookup wf ex lookup). +Qed. + + Lemma extends_lookup_constructor {efl} {Σ Σ'} : wf_glob Σ' -> extends Σ Σ' -> forall ind c b, lookup_constructor Σ ind c = Some b -> @@ -482,6 +502,8 @@ Proof. all:try destruct g => //. - destruct cstr_as_blocks; eauto; solve_all. destruct lookup_constructor_pars_args as [ [] | ]; rtoProp; repeat solve_all. + - move: H0. rewrite /wf_brs. destruct lookup_inductive eqn:hl => //. + now rewrite (extends_lookup_inductive wf ex _ _ hl). - move/andP: H0 => [] hn hf. unfold wf_fix. rewrite hn /=. solve_all. - move/andP: H0 => [] hn hf. unfold wf_fix. rewrite hn /=. solve_all. Qed. diff --git a/erasure/theories/ErasureCorrectness.v b/erasure/theories/ErasureCorrectness.v index 4dcabe930..1dfd6fec1 100644 --- a/erasure/theories/ErasureCorrectness.v +++ b/erasure/theories/ErasureCorrectness.v @@ -195,9 +195,10 @@ Proof. eapply Construct_Ind_ind_eq in X0; tea. 2:{ eapply declared_constructor_from_gen; tea. } destruct X0 as (((([_ Ru] & cu) & cpu) & ?) & (parsubst & argsubst & parsubst' & argsubst' & [])). invs He. - + depelim Hed. + + depelim Hed. rename H5 into Hlen. rename H1 into decli. rename H2 into decli'. rename H3 into em. rename Hed into er. + rename H6 into H5. rename H7 into H6. rename H8 into H7. edestruct (IHeval1 _ scrut_ty) as (v' & Hv' & [He_v']); eauto. pose proof e0 as Hnth. assert (lenppars : ind_npars mdecl = #|pparams p|). @@ -838,7 +839,9 @@ Proof. 2:{ now eapply nth_error_all in a; tea. } invs He. * edestruct IHeval1 as (? & ? & ?); eauto. now depelim Hed. - depelim Hed. + depelim Hed. rename H5 into Hlen. + rename H6 into H5. rename H7 into H6. rename H8 into H7. + rename H9 into H8. rename H10 into H9. unshelve eapply declared_inductive_to_gen in H1, decli; eauto. destruct (declared_inductive_inj H1 decli). subst mdecl0 idecl0. rename H0 into decli'. rename H1 into decli''. rename H2 into er. rename H3 into H0. @@ -1368,7 +1371,7 @@ Proof. depelim etaΣ. eauto. depelim etaΣ. - solve_all. rewrite -b2. len. eapply H7 => //. + solve_all. rewrite -b2. len. eapply H8 => //. exact a0. - intros Γ0 v etaΣ. move=> /erases_mkApps_inv; intros [(?&?&?&?&_&?&?&?)|(?&?&?&?&?)]; subst. diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index cbebb27ab..34c0e307a 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -169,6 +169,7 @@ Proof. - destruct p. rewrite KernameSet.union_spec knset_in_fold_left in H0. destruct H0. apply KernameSet.singleton_spec in H0. subst kn. cbn in *. + move: H1; rewrite /EWellformed.wf_brs; cbn. destruct (EGlobalEnv.lookup_env Σ _) eqn:E. destruct g => //. now eapply lookup_env_In in E. easy. clear H1. diff --git a/erasure/theories/ErasureFunctionProperties.v b/erasure/theories/ErasureFunctionProperties.v index c6b674c45..be6884594 100644 --- a/erasure/theories/ErasureFunctionProperties.v +++ b/erasure/theories/ErasureFunctionProperties.v @@ -245,6 +245,9 @@ Proof. eapply erases_deps_tCase; eauto. apply declared_inductive_from_gen; auto. split; eauto. split; eauto. + apply All2_length in X. rewrite -X. + destruct H5. apply Forall2_length in H5. rewrite -H5. + now apply Forall2_length in wf_brs. destruct H1. eapply In_Forall in H3. eapply All_Forall. eapply Forall_All in H3. @@ -350,6 +353,7 @@ Proof. eapply extends_decls_extends, strictly_extends_decls_extends_decls. econstructor; try reflexivity. eexists [(_, _)]; reflexivity. - econstructor; eauto. eapply declared_inductive_from_gen. + rename H3 into Hlen, H4 into H3, H5 into H4. inv wfΣ. inv X. assert (wf Σ) by (inversion H5;econstructor; eauto). unshelve eapply declared_inductive_to_gen in H; eauto. @@ -1451,11 +1455,11 @@ Proof. - cbn. red in H0. rewrite H0 //. - cbn -[lookup_constructor]. cbn. now destruct H0 as [[-> ->] ->]. - - cbn in *. move/andP: H5 => [] cld clbrs. + - cbn in *. move/andP: H6 => [] cld clbrs. cbn. apply/andP; split. apply/andP; split. - * now destruct H0 as [-> ->]. - * now move/andP: H6. - * move/andP: H6; PCUICAstUtils.solve_all. + * rewrite /wf_brs. cbn. destruct H0 as [-> ->]. now apply Nat.eqb_eq. + * now move/andP: H7. + * move/andP: H7; PCUICAstUtils.solve_all. - cbn -[lookup_projection] in *. apply/andP; split; eauto. now rewrite (declared_projection_lookup H0). - cbn in H, H0 |- *. rtoProp; intuition eauto. solve_all_k 7. diff --git a/erasure/theories/ErasureProperties.v b/erasure/theories/ErasureProperties.v index 125ca26b8..06b919b1a 100644 --- a/erasure/theories/ErasureProperties.v +++ b/erasure/theories/ErasureProperties.v @@ -475,6 +475,12 @@ Section wellscoped. Context (Σ : global_env). Import ssrbool. + Definition wf_brs ind brsl := + match lookup_inductive Σ ind with + | Some (mib, oib) => #|oib.(ind_ctors)| == brsl + | None => false + end. + Fixpoint wellformed (t : term) : bool := match t with | tRel i => true @@ -485,7 +491,7 @@ Section wellscoped. | tLetIn na b ty b' => wellformed b && wellformed ty && wellformed b' | tCase ind p c brs => let brs' := forallb (wellformed ∘ bbody) brs in - isSome (lookup_inductive Σ ind.(ci_ind)) && wellformed c && brs' + wf_brs ind.(ci_ind) #|brs| && wellformed c && brs' | tProj p c => isSome (lookup_projection Σ p) && wellformed c | tFix mfix idx => (idx isSome (EGlobalEnv.lookup_inductive Σ' kn). - Proof using g. - destruct g. - destruct (lookup_inductive Σ kn) as [[]|] eqn:hl => /= // _. + Lemma trans_lookup_inductive kn mib oib : + lookup_inductive Σ kn = Some (mib, oib) -> + exists mib' oib', EGlobalEnv.lookup_inductive Σ' kn = Some (mib', oib') /\ #|oib.(ind_ctors)| = #|oib'.(EAst.ind_ctors)|. + Proof. destruct g. + destruct (lookup_inductive Σ kn) as [[]|] eqn:hl => /= // [=]. intros -> ->. eapply lookup_inductive_declared in hl. eapply declared_inductive_from_gen in hl. - specialize (H0 kn m o hl) as [? [? [d _]]]. - now rewrite (EGlobalEnv.declared_inductive_lookup d). + specialize (H0 kn mib oib hl) as [? [? [d er]]]. + exists x, x0. + split. now rewrite (EGlobalEnv.declared_inductive_lookup d). + depelim er. depelim hl. depelim d. + eapply Forall2_nth_error_left in H4 as [x' []]; tea. + rewrite H3 in H4. noconf H4. depelim H6. now apply Forall2_length in H4. Qed. Lemma trans_lookup_constructor kn c : isSome (lookup_constructor Σ kn c) -> isSome (EGlobalEnv.lookup_constructor Σ' kn c). @@ -636,7 +649,11 @@ Proof. - eapply trans_lookup_constructor in wfa; tea. now rewrite wfa. - move/andP: wfa => [] /andP[] lookup wfc wfbrs. apply/andP. split. apply/andP. split; eauto. - eapply trans_lookup_inductive; tea. + { rewrite /wf_brs. + move: lookup. rewrite /ErasureProperties.wf_brs. + destruct lookup_inductive as [[mib oib]|] eqn:hli => //. + eapply trans_lookup_inductive in hli as [mib' [oib' [-> <-]]]; tea. + move/Nat.eqb_eq ->. apply All2_length in X1. now apply Nat.eqb_eq. } eapply All_forallb. unfold tCaseBrsProp_k in X0. eapply All2_All_mix_left in X1; eauto. eapply forallb_All in wfbrs. diff --git a/erasure/theories/Extract.v b/erasure/theories/Extract.v index 3a3f72ee9..ea9d49a2c 100644 --- a/erasure/theories/Extract.v +++ b/erasure/theories/Extract.v @@ -338,6 +338,7 @@ Inductive erases_deps (Σ : global_env) (Σ' : E.global_declarations) : E.term - erases_mutual_inductive_body mdecl mdecl' -> erases_one_inductive_body idecl idecl' -> erases_deps Σ Σ' discr -> + #|EAst.ind_ctors idecl'| = #|brs| -> Forall (fun br => erases_deps Σ Σ' br.2) brs -> erases_deps Σ Σ' (E.tCase p discr brs) | erases_deps_tProj p mdecl idecl cdecl pdecl mdecl' idecl' cdecl' pdecl' t : diff --git a/erasure/theories/Typed/OptimizeCorrectness.v b/erasure/theories/Typed/OptimizeCorrectness.v index a8f685dfc..7850b06fc 100644 --- a/erasure/theories/Typed/OptimizeCorrectness.v +++ b/erasure/theories/Typed/OptimizeCorrectness.v @@ -3611,6 +3611,7 @@ Section dearg. erewrite lookup_ctor_trans_env_inv; tea. - destruct p. rewrite wellformed_mkApps; try easy. unfold dearg_case. + rewrite /wf_brs in H. destruct (EGlobalEnv.lookup_inductive _ _) as [[mib oib]|] eqn:hl => //. assert (decl_ind :declared_inductive (trans_env Σ) i mib oib). { move: hl. unfold EGlobalEnv.lookup_inductive. cbn. @@ -3619,15 +3620,18 @@ Section dearg. specialize (valid_ind_mask_inductive _ _ _ _ valid_Σ decl_ind) as [mask [Hmask Hparams]]. rewrite Hmask. rtoProp; intuition eauto; solve_all. - cbn [wellformed]. rtoProp; intuition eauto. + cbn [wellformed]. rtoProp; intuition eauto. len. + rewrite mapi_length map_length. rewrite /wf_brs. { unfold EGlobalEnv.lookup_inductive. cbn. move: hl. cbn. rewrite !lookup_env_trans_env lookup_env_dearg_env. destruct lookup_env => //=. destruct g => //=. rewrite !nth_error_map. unfold dearg_mib. rewrite Hmask. cbn. - rewrite nth_error_mapi. destruct nth_error => //. } + rewrite nth_error_mapi. destruct nth_error => //=. + intros [= <- <-]. + move: H; cbn. now rewrite /trans_ctors !map_length !mapi_length. } cbn. - unfold mapi. clear clos_args IHt. + unfold mapi. clear clos_args IHt H. unfold valid_case_masks in H3. rewrite Hmask in H3. move/andP: H3 => [] _ hbrs. eapply alli_Alli in hbrs. diff --git a/examples/metacoq_tour.v b/examples/metacoq_tour.v index 3830aa3de..5b5ba9c1c 100644 --- a/examples/metacoq_tour.v +++ b/examples/metacoq_tour.v @@ -99,7 +99,7 @@ From MetaCoq.ErasurePlugin Require Import Erasure Loader. (** Running erasure live in Coq *) Definition test (p : Ast.Env.program) : string := - erase_and_print_template_program default_erasure_config p. + erase_and_print_template_program default_erasure_config [] p. MetaCoq Quote Recursively Definition zero := 0. diff --git a/safechecker-plugin/clean_extraction.sh b/safechecker-plugin/clean_extraction.sh index 44164597f..85e2f761a 100755 --- a/safechecker-plugin/clean_extraction.sh +++ b/safechecker-plugin/clean_extraction.sh @@ -9,7 +9,7 @@ fi shopt -s nullglob # make the for loop do nothnig when there is no *.ml* files -files=`cat ../template-coq/_PluginProject | grep "^[^#].*mli\?$" | sed -e s/gen-src/src/` +files=`cat ../template-coq/_PluginProject.in | grep "^[^#].*mli\?$" | sed -e s/gen-src/src/` if [[ ! -f "src/metacoq_safechecker_plugin.cmxs" || "src/metacoq_safechecker_plugin.cmxs" -ot "theories/Extraction.vo" ]] diff --git a/template-coq/_PluginProject b/template-coq/_PluginProject deleted file mode 100644 index 314e020ac..000000000 --- a/template-coq/_PluginProject +++ /dev/null @@ -1,248 +0,0 @@ -# DO NOT EDIT THIS FILE: autogenerated from ./configure.sh --R ../utils/theories MetaCoq.Utils -R ../common/theories MetaCoq.Common --I gen-src --Q gen-src MetaCoq.Template --R theories MetaCoq.Template -META.coq-metacoq-template-ocaml --I . - -# Generated Code by `ls -1 gen-src/*.ml gen-src/*.mli` in `template-coq/` after having compiled `Extraction.v` -gen-src/all_Forall.ml -gen-src/all_Forall.mli -gen-src/ascii.ml -gen-src/ascii.mli -gen-src/ast0.ml -gen-src/ast0.mli -gen-src/astUtils.ml -gen-src/astUtils.mli -gen-src/ast_denoter.ml -gen-src/ast_quoter.ml -gen-src/basicAst.ml -gen-src/basicAst.mli -gen-src/basics.ml -gen-src/basics.mli -gen-src/binInt.ml -gen-src/binInt.mli -gen-src/binNat.ml -gen-src/binNat.mli -gen-src/binNums.ml -gen-src/binNums.mli -gen-src/binPos.ml -gen-src/binPos.mli -gen-src/binPosDef.ml -gen-src/binPosDef.mli -gen-src/decidableClass.mli -gen-src/decidableClass.ml -gen-src/bool.ml -gen-src/bool.mli -gen-src/byte.ml -gen-src/byte.mli -gen-src/byte0.ml -gen-src/byte0.mli -gen-src/byteCompare.ml -gen-src/byteCompare.mli -gen-src/byteCompareSpec.ml -gen-src/byteCompareSpec.mli -gen-src/bytestring.ml -gen-src/bytestring.mli -gen-src/cRelationClasses.ml -gen-src/cRelationClasses.mli -gen-src/caml_byte.ml -gen-src/caml_byte.mli -gen-src/caml_bytestring.ml -gen-src/caml_nat.ml -gen-src/caml_nat.mli -gen-src/carryType.ml -gen-src/carryType.mli -gen-src/classes0.ml -gen-src/classes0.mli -gen-src/common0.ml -gen-src/common0.mli -gen-src/compare_dec.ml -gen-src/compare_dec.mli -gen-src/config0.ml -gen-src/config0.mli -gen-src/coreTactics.ml -gen-src/coreTactics.mli -gen-src/datatypes.ml -gen-src/datatypes.mli -gen-src/decidableType.ml -gen-src/decidableType.mli -gen-src/decimal.ml -gen-src/decimal.mli -gen-src/denoter.ml -gen-src/depElim.ml -gen-src/depElim.mli -gen-src/envMap.ml -gen-src/envMap.mli -gen-src/environment.ml -gen-src/environment.mli -gen-src/environmentTyping.ml -gen-src/environmentTyping.mli -gen-src/eqDec.ml -gen-src/eqDec.mli -gen-src/eqDecInstances.ml -gen-src/eqDecInstances.mli -gen-src/eqdepFacts.ml -gen-src/eqdepFacts.mli -gen-src/equalities.ml -gen-src/equalities.mli -gen-src/extractable.ml -gen-src/extractable.mli -gen-src/fMapAVL.ml -gen-src/fMapAVL.mli -gen-src/fMapFacts.ml -gen-src/fMapFacts.mli -gen-src/fMapInterface.ml -gen-src/fMapInterface.mli -gen-src/fMapList.ml -gen-src/fMapList.mli -gen-src/floatClass.ml -gen-src/floatClass.mli -gen-src/floatOps.ml -gen-src/floatOps.mli -gen-src/hexadecimal.ml -gen-src/hexadecimal.mli -# gen-src/hexadecimalString.ml -# gen-src/hexadecimalString.mli -gen-src/induction0.ml -gen-src/induction0.mli -gen-src/init.ml -gen-src/init.mli -gen-src/int0.ml -gen-src/int0.mli -gen-src/kernames.ml -gen-src/kernames.mli -gen-src/liftSubst.ml -gen-src/liftSubst.mli -gen-src/list0.ml -gen-src/list0.mli -gen-src/logic0.ml -gen-src/logic0.mli -gen-src/logic1.ml -gen-src/logic1.mli -gen-src/logic2.ml -gen-src/logic2.mli -gen-src/mCCompare.ml -gen-src/mCCompare.mli -gen-src/mCFSets.ml -gen-src/mCFSets.mli -gen-src/mCList.ml -gen-src/mCList.mli -gen-src/mCMSets.ml -gen-src/mCMSets.mli -gen-src/mCOption.ml -gen-src/mCOption.mli -gen-src/mCPrelude.ml -gen-src/mCPrelude.mli -gen-src/mCProd.ml -gen-src/mCProd.mli -gen-src/mCReflect.ml -gen-src/mCReflect.mli -gen-src/mCRelations.ml -gen-src/mCRelations.mli -gen-src/mCString.ml -gen-src/mCString.mli -gen-src/sint63.mli -gen-src/sint63.ml -gen-src/show.mli -gen-src/show.ml -# gen-src/mCUint63.ml -# gen-src/mCUint63.mli -gen-src/mCUtils.ml -gen-src/mCUtils.mli -gen-src/mSetAVL.ml -gen-src/mSetAVL.mli -gen-src/mSetDecide.ml -gen-src/mSetDecide.mli -gen-src/mSetFacts.ml -gen-src/mSetFacts.mli -gen-src/mSetInterface.ml -gen-src/mSetInterface.mli -gen-src/mSetList.ml -gen-src/mSetList.mli -gen-src/mSetProperties.ml -gen-src/mSetProperties.mli -gen-src/monad_utils.ml -gen-src/monad_utils.mli -gen-src/nat0.ml -gen-src/nat0.mli -gen-src/noConfusion.ml -gen-src/noConfusion.mli -gen-src/number.ml -gen-src/number.mli -gen-src/orderedType0.ml -gen-src/orderedType0.mli -gen-src/orderedTypeAlt.ml -gen-src/orderedTypeAlt.mli -gen-src/orders.ml -gen-src/orders.mli -gen-src/ordersAlt.ml -gen-src/ordersAlt.mli -gen-src/ordersFacts.ml -gen-src/ordersFacts.mli -gen-src/ordersLists.ml -gen-src/ordersLists.mli -gen-src/ordersTac.ml -gen-src/ordersTac.mli -gen-src/peanoNat.ml -gen-src/peanoNat.mli -gen-src/plugin_core.ml -gen-src/plugin_core.mli -gen-src/pretty.ml -gen-src/pretty.mli -gen-src/primFloat.ml -gen-src/primFloat.mli -gen-src/primInt63.ml -gen-src/primInt63.mli -gen-src/primitive.ml -gen-src/primitive.mli -gen-src/quoter.ml -gen-src/reflect.ml -gen-src/reflect.mli -gen-src/reflectEq.ml -gen-src/reflectEq.mli -gen-src/reification.ml -gen-src/relation.ml -gen-src/relation.mli -gen-src/run_extractable.ml -gen-src/run_extractable.mli -gen-src/signature.ml -gen-src/signature.mli -gen-src/specFloat.ml -gen-src/specFloat.mli -gen-src/specif.ml -gen-src/specif.mli -gen-src/string0.ml -gen-src/string0.mli -gen-src/sumbool.ml -gen-src/sumbool.mli -gen-src/templateEnvMap.ml -gen-src/templateEnvMap.mli -gen-src/templateProgram.ml -gen-src/templateProgram.mli -gen-src/termEquality.ml -gen-src/termEquality.mli -gen-src/tm_util.ml -gen-src/transform.ml -gen-src/transform.mli -gen-src/typing0.ml -gen-src/typing0.mli -gen-src/uint0.ml -gen-src/uint0.mli -gen-src/universes0.ml -gen-src/universes0.mli -gen-src/wf.ml -gen-src/wf.mli -gen-src/zArith_dec.ml -gen-src/zArith_dec.mli -gen-src/zbool.ml -gen-src/zbool.mli -gen-src/zeven.ml -gen-src/zeven.mli -gen-src/zpower.ml -gen-src/zpower.mli - -gen-src/metacoq_template_plugin.mlpack - -theories/ExtractableLoader.v diff --git a/template-coq/_TemplateCoqProject b/template-coq/_TemplateCoqProject deleted file mode 100644 index 05f7e2517..000000000 --- a/template-coq/_TemplateCoqProject +++ /dev/null @@ -1,32 +0,0 @@ -# DO NOT EDIT THIS FILE: autogenerated from ./configure.sh --R ../utils/theories MetaCoq.Utils -R ../common/theories MetaCoq.Common --I src --Q src MetaCoq.Template --R theories MetaCoq.Template -META.coq-metacoq-template-coq --I . - -# the MetaCoq plugin -src/tm_util.ml - -src/reification.ml -src/quoter.ml -src/denoter.ml -src/constr_reification.ml -src/constr_quoter.ml -src/constr_denoter.ml - -src/g_template_coq.mlg -src/template_coq.mlpack -src/template_monad.mli -src/template_monad.ml -src/run_template_monad.mli -src/run_template_monad.ml -src/plugin_core.mli -src/plugin_core.ml - -theories/Loader.v -theories/All.v - -# A checker of well-formedness in MetaCoq and Coq -theories/TemplateCheckWf.v diff --git a/test-suite/erasure_live_test.v b/test-suite/erasure_live_test.v index d9796c1f1..79d10fdfc 100644 --- a/test-suite/erasure_live_test.v +++ b/test-suite/erasure_live_test.v @@ -18,14 +18,11 @@ Unset MetaCoq Debug. #[local] Existing Instance config.extraction_checker_flags. Definition test (p : Ast.Env.program) : string := - erase_and_print_template_program default_erasure_config p. + erase_and_print_template_program default_erasure_config [] p. Definition testty (p : Ast.Env.program) : string := typed_erase_and_print_template_program p. -Definition test_fast (p : Ast.Env.program) : string := - erase_fast_and_print_template_program p. - MetaCoq Quote Recursively Definition zero := 0. Definition zerocst := Eval lazy in test zero. @@ -53,12 +50,7 @@ Definition singlelim := ((fun (X : Set) (x : X) (e : x = x) => Definition erase {A} (a : A) : TemplateMonad unit := aq <- tmQuoteRec a ;; - s <- tmEval lazy (erase_and_print_template_program default_erasure_config aq) ;; - tmMsg s. - -Definition erase_fast {A} (a : A) : TemplateMonad unit := - aq <- tmQuoteRec a ;; - s <- tmEval lazy (erase_fast_and_print_template_program aq) ;; + s <- tmEval lazy (erase_and_print_template_program default_erasure_config [] aq) ;; tmMsg s. MetaCoq Run (erase 0). @@ -80,7 +72,6 @@ Definition vplus0123 := (vplus v01 v23). Set MetaCoq Timing. Time MetaCoq Run (tmEval hnf vplus0123 >>= erase). -Time MetaCoq Run (tmEval hnf vplus0123 >>= erase_fast). (** Cofix *) From Coq Require Import StreamMemo. @@ -352,7 +343,7 @@ MetaCoq Quote Recursively Definition cbv_provedCopyx := Definition ans_provedCopyx := Eval lazy in (test cbv_provedCopyx). MetaCoq Quote Recursively Definition p_provedCopyx := provedCopyx. (* program *) -Time Definition P_provedCopyx := Eval lazy in (test_fast cbv_provedCopyx). +Time Definition P_provedCopyx := Eval lazy in (test cbv_provedCopyx). (* We don't run this one every time as it is really expensive *) (*Time Definition P_provedCopyxvm := Eval vm_compute in (test p_provedCopyx).*) From b1b6be23e00fefba6a03c153a458e72acb463daa Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Mon, 22 Jul 2024 16:24:42 +0200 Subject: [PATCH 28/28] make-opam-files now computes the appropriate ci-skip line --- make-opam-files.sh | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/make-opam-files.sh b/make-opam-files.sh index f448ac508..38bbba529 100755 --- a/make-opam-files.sh +++ b/make-opam-files.sh @@ -33,15 +33,20 @@ gh release upload $tag $archive release=https://github.com/MetaCoq/metacoq/releases/download/$tag/$archive +skipline="" + for f in *.opam; do opamf=${f/.opam/}; target=$1/$opamf/$opamf.$2/opam; echo $opamf; mkdir -p $1/$opamf/$opamf.$2 + skipline="$skipline $opamf.$2" gsed -e "/^version:.*/d" $f > $target echo url { >> $target echo " src:" \"$release\" >> $target echo " checksum:" \"sha512=$hash\" >> $target echo } >> $target -done \ No newline at end of file +done + +echo "ci-skip:" $skipline \ No newline at end of file