From ed9a17c3275dd521ebdd17024aa2d06df08c3988 Mon Sep 17 00:00:00 2001 From: Kenji Maillard Date: Wed, 30 Sep 2020 11:32:17 +0200 Subject: [PATCH 01/26] inversion for tProd --- .gitignore | 3 +++ pcuic/theories/PCUICConversion.v | 31 +++++++++++++++++++++++++++++-- 2 files changed, 32 insertions(+), 2 deletions(-) diff --git a/.gitignore b/.gitignore index aafae18fd..4049a0369 100644 --- a/.gitignore +++ b/.gitignore @@ -397,6 +397,9 @@ vs.mli *.v.timing time-of-build.log time-of-build-pretty.log +_build/** +extraction/src/*.ml +extraction/src/*.mli erasure/src/orderedTypeAlt.ml erasure/src/orderedTypeAlt.mli erasure/src/orderedTypeEx.ml diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index f2fc1b933..a71569c42 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -337,6 +337,8 @@ Proof. repeat constructor. Qed. +Derive Signature for conv. + Lemma conv_Prod_l_inv {cf:checker_flags} (Σ : global_env_ext) Γ na dom codom T : wf Σ -> Σ ;;; Γ |- tProd na dom codom = T -> @@ -1273,6 +1275,32 @@ Section Inversions. + eapply conv_cumul. eapply conv_Prod_l. assumption. Qed. + Lemma Prod_conv_cum_inv : + forall {Γ leq na1 na2 A1 A2 B1 B2}, + conv_cum leq Σ Γ (tProd na1 A1 B1) (tProd na2 A2 B2) -> + ∥ Σ ;;; Γ |- A1 = A2 ∥ /\ + conv_cum leq Σ (Γ,, vass na1 A1) B1 B2. + Proof. + intros *; destruct leq; simpl. + - intros [[na1' [A1' [B1' [[Hred eqA] eqB]]]]%conv_Prod_r_inv]. 2: assumption. + apply invert_red_prod in Hred. 2: assumption. + destruct Hred as [? [? [[[= ? ?%eq_sym ?%eq_sym] redA] redB]]]. + subst. + assert (Σ;;;Γ |- A1 = A2) as HA by + (etransitivity; [apply red_conv, redA | assumption]). + split; constructor. 1: assumption. + etransitivity. 1: apply red_conv, redB. + eapply conv_conv_ctx. 1,2 : eassumption. + apply ctx_rel_vass. 1: reflexivity. + now constructor. + - intros [[eqA cumB]%cumul_Prod_Prod_inv]. 2: assumption. + split; constructor. 1: assumption. + eapply cumul_conv_ctx. 1,2: eassumption. + apply ctx_rel_vass. 1: reflexivity. + now constructor. + Qed. + + Lemma cumul_Case_c : forall Γ indn p brs u v, Σ ;;; Γ |- u = v -> @@ -1919,11 +1947,10 @@ Section Inversions. Lemma Lambda_conv_cum_inv : forall leq Γ na1 na2 A1 A2 b1 b2, - wf_local Σ Γ -> conv_cum leq Σ Γ (tLambda na1 A1 b1) (tLambda na2 A2 b2) -> ∥ Σ ;;; Γ |- A1 = A2 ∥ /\ conv_cum leq Σ (Γ ,, vass na1 A1) b1 b2. Proof. - intros * wfΓ. + intros *. destruct leq; simpl in *. - destruct 1. eapply conv_alt_red in X as [l [r [[redl redr] eq]]]. From 34f46bcc07278cf9fd896e987d71a8d538f119a5 Mon Sep 17 00:00:00 2001 From: Kenji Maillard Date: Wed, 16 Sep 2020 15:23:16 +0200 Subject: [PATCH 02/26] adding completeness to ConversionResult + todo --- pcuic/theories/PCUICConversion.v | 9 + safechecker/theories/PCUICSafeChecker.v | 2 +- safechecker/theories/PCUICSafeConversion.v | 501 +++++++++++++++++++-- 3 files changed, 463 insertions(+), 49 deletions(-) diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index a71569c42..a5efa9ad5 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -1464,6 +1464,15 @@ Section Inversions. + apply conv_Case_brs. all: assumption. Qed. + Lemma Case_conv_inv : + forall Γ indn p p' c c' brs brs', + Σ ;;; Γ |- tCase indn p c brs = tCase indn p' c' brs' -> + Σ ;;; Γ |- p = p' /\ Σ ;;; Γ |- c = c' /\ + ∥ All2 (fun u v => u.1 = v.1 × Σ ;;; Γ |- u.2 = v.2) brs brs' ∥. + Proof. + todo "Completeness"%string. + Qed. + Lemma conv_Proj_c : forall Γ p u v, Σ ;;; Γ |- u = v -> diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index 235587962..fcad4449f 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -484,7 +484,7 @@ Section Typecheck. match leqb_term Σ G t u with true => ret _ | false => match iscumul Γ t ht u hu with | Success _ => ret _ - | Error e => (* fallback *) (* nico *) + | Error e h => (* fallback *) (* nico *) let t' := hnf Γ t ht in let u' := hnf Γ u hu in (* match leq_term (snd Σ) t' u' with true => ret _ | false => *) diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 7e88e0eef..541702927 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -17,6 +17,17 @@ Set Default Goal Selector "!". Module PSR := PCUICSafeReduce. +(* Removing the definitions generated by equations *) +Add Search Blacklist "obligation". +Add Search Blacklist "equation". +Add Search Blacklist "clause". +Add Search Blacklist "Functional". +Add Search Blacklist "graph". +Add Search Blacklist "elim". +(* Removing the eliminators *) +Add Search Blacklist "ind". +Add Search Blacklist "rec". + (** * Conversion for PCUIC without fuel Following PCUICSafereduce, we derive a fuel-free implementation of @@ -732,7 +743,7 @@ Section Conversion. Inductive ConversionResult (P : Prop) := | Success (h : P) - | Error (e : ConversionError). + | Error (e : ConversionError) (h : ~P). Arguments Success {_} _. Arguments Error {_} _. @@ -772,6 +783,7 @@ Section Conversion. Ret s' Γ t1' π1' t2' π2'. Notation yes := (Success _) (only parsing). + Notation no := (fun e => Error e _) (only parsing). (* TODO NOTE repack could also take another argument of type @@ -779,7 +791,7 @@ Section Conversion. to keep a full error trace (we currently only do it partially). *) Notation repack e := - (match e with Success h => Success _ | Error er => Error er end) + (match e with Success h => Success _ | Error er h => Error er _ end) (only parsing). Notation isconv_red_raw leq t1 π1 t2 π2 aux := @@ -1103,6 +1115,80 @@ Section Conversion. * eapply red_conv_cum_r. all: eauto. * eapply conv_context_sym. all: auto. Qed. + Next Obligation. + rename H into e; apply h; clear h. + + (* FIXME: the script below is essentially a copy-paste of the previous obligation + with only the hypothesis and the goal swapped... *) + destruct hΣ as [wΣ]. + unfold zipp in e. rewrite <- e1, <- e2 in e. + + match type of eq1 with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r1] ; + pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d1 ; + pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c1 + end. + rewrite <- eq1 in r1. + rewrite <- eq1 in d1. cbn in d1. + rewrite <- eq1 in c1. cbn in c1. + rewrite stack_context_appstack in c1. cbn in c1. + + match type of eq2 with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; + pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; + pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 + end. + rewrite <- eq2 in r2. + rewrite <- eq2 in d2. cbn in d2. + rewrite <- eq2 in c2. cbn in c2. + rewrite stack_context_appstack in c2. cbn in c2. + + clear eq1 eq2. + + pose proof (decompose_stack_eq _ _ _ (eq_sym e1)). subst. + case_eq (decompose_stack π1'). intros args1' ρ1' e1'. + rewrite e1' in d1. cbn in d1. + rewrite decompose_stack_appstack in d1. cbn in d1. subst. + pose proof (decompose_stack_eq _ _ _ e1'). subst. + + pose proof (decompose_stack_eq _ _ _ (eq_sym e2)). subst. + case_eq (decompose_stack π2'). intros args2' ρ2' e2'. + rewrite e2' in d2. cbn in d2. + rewrite decompose_stack_appstack in d2. cbn in d2. subst. + pose proof (decompose_stack_eq _ _ _ e2'). subst. + + rewrite stack_context_appstack in r1. cbn in r1. + rewrite 2!zipc_appstack in r1. cbn in r1. + + rewrite stack_context_appstack in r2. cbn in r2. + rewrite 2!zipc_appstack in r2. cbn in r2. + + rewrite 2!stack_cat_appstack. + unfold zipp. + rewrite 2!decompose_stack_appstack. + rewrite decompose_stack_twice with (1 := eq_sym e1). + rewrite decompose_stack_twice with (1 := eq_sym e2). + simpl. + rewrite 2!app_nil_r. + rewrite 2!stack_context_appstack in hx. + rewrite stack_context_appstack. + + rewrite stack_context_appstack in e. + + destruct hx as [hx]. + etransitivity. + - eapply red_conv_cum_r; eassumption. + - etransitivity. + + eassumption. + + eapply conv_cum_context_convp. + * assumption. + * eapply red_conv_cum_l. all: eauto. + * eapply conv_context_sym. all: auto. + Qed. + + Opaque reduce_stack. Equations unfold_one_fix (Γ : context) (mfix : mfixpoint term) @@ -1427,6 +1513,8 @@ Section Conversion. now eapply All_local_env_app in wf. Qed. + (* Q: Shouldnt't we have as hypothesis that [c <> c'] ? *) + (* Also, is there an invariant of injectivity on the assignment between kernames and inductives *) Equations(noeqns) unfold_constants (Γ : context) (leq : conv_pb) (c : kername) (u : Instance.t) (π1 : stack) (h1 : wtp Γ (tConst c u) π1) @@ -1446,7 +1534,7 @@ Section Conversion. isconv_red leq (subst_instance_constr u b) π1 (tConst c' u') π2 aux ; (* Both Inductive or not found *) - | _ := Error (NotFoundConstants c c') + | _ := Error (NotFoundConstants c c') _ } }. Next Obligation. @@ -1467,6 +1555,14 @@ Section Conversion. eapply red_conv_cum_r ; try assumption. eapply red_zipp. eapply red_const. eassumption. Qed. + Next Obligation. + (* Contraposition of previous goal *) + apply h; clear h. + destruct hΣ. + etransitivity ; try eassumption. + eapply red_conv_cum_l ; try assumption. + eapply red_zipp. eapply red_const. eassumption. + Qed. Next Obligation. eapply red_wellformed ; auto ; [ exact h1 | ]. constructor. eapply red_zipc. @@ -1483,6 +1579,29 @@ Section Conversion. eapply red_zipp. eapply red_const. eassumption. Qed. Next Obligation. + (* Contraposition of previous goal *) + apply h; clear h. + destruct hΣ as [wΣ]. + etransitivity ; try eassumption. + eapply red_conv_cum_r ; try assumption. + eapply red_zipp. eapply red_const. eassumption. + Qed. + Next Obligation. + (* No definition to unfold... + Should we have an hypothesis ditching that case ? *) + todo "Completeness"%string. + Qed. + Next Obligation. + (* inductive vs non-defined constant *) + (* invert H and discriminate between inductive and constant *) + todo "Completeness"%string. + Qed. + Next Obligation. + (* Not found vs non-defined constant *) + (* invert H and discriminate between inductive and constant *) + todo "Completeness"%string. + Qed. + Next Obligation. eapply red_wellformed ; auto ; [ exact h1 | ]. constructor. eapply red_zipc. eapply red_const. eassumption. Qed. @@ -1496,6 +1615,26 @@ Section Conversion. eapply red_conv_cum_l ; try assumption. eapply red_zipp. eapply red_const. eassumption. Qed. + Next Obligation. + (* Contrapositive of previous obligation *) + apply h; clear h. + destruct hΣ as [wΣ]. + etransitivity ; try eassumption. + eapply red_conv_cum_r ; try assumption. + eapply red_zipp. eapply red_const. eassumption. + Qed. + Next Obligation. + (* Non-defined constant vs inductive *) + todo "Completeness"%string. + Qed. + Next Obligation. + (* Two inductives: missing hypothesis to ditch that case ?*) + todo "Completeness"%string. + Qed. + Next Obligation. + (* Not found vs inductive *) + todo "Completeness"%string. + Qed. Next Obligation. eapply red_wellformed ; auto ; [ exact h1 | ]. constructor. eapply red_zipc. eapply red_const. eassumption. @@ -1510,7 +1649,28 @@ Section Conversion. eapply red_conv_cum_l ; try assumption. eapply red_zipp. eapply red_const. eassumption. Qed. - + Next Obligation. + (* Contrapositive of previous obligation *) + apply h; clear h. + destruct hΣ as [wΣ]. + etransitivity ; try eassumption. + eapply red_conv_cum_r ; try assumption. + eapply red_zipp. eapply red_const. eassumption. + Qed. + Next Obligation. + (* Constant w/o body vs not found *) + todo "Completeness"%string. + Qed. + Next Obligation. + (* Inductive vs not found *) + todo "Completeness"%string. + Qed. + Next Obligation. + (* Not found vs Not found *) + todo "Completeness"%string. + Qed. + + (* TODO MOVE *) Definition eqb_universe_instance u v := forallb2 (check_eqb_universe G) (map Universe.make u) (map Universe.make v). @@ -1586,15 +1746,15 @@ Section Conversion. p c (brs1 ++ [(m,br)]) brs2 π _ p' c' (brs1' ++ [(m', br')]) brs2' π' _ hx _ _ := { | Success h3 := yes ; - | Error e := Error e + | Error e h := no e } ; - | Error e := Error e + | Error e h := no e } ; - | @exist false _ := Error ( + | @exist false eq1 := Error ( CaseBranchNumMismatch ind par (Γ ,,, stack_context π) p c brs1 m br brs2 (Γ ,,, stack_context π') p' c' brs1' m' br' brs2' - ) + ) _ } ; isconv_branches Γ ind par @@ -1672,6 +1832,26 @@ Section Conversion. destruct (eqb_spec m m'). 2: discriminate. constructor. constructor. all: intuition eauto. Qed. + Next Obligation. + (* COntrapositive of previous obligation *) + apply h; clear h. + destruct H as [H]; inversion H; now constructor. + Qed. + Next Obligation. + (* Contrapositive of 3rd obligation above *) + apply h; clear h. + destruct h1 as [h1], H as [h2]. + constructor. inversion h2; clear h2. + destruct X as [_ h2]. apply h2. + Qed. + Next Obligation. + destruct H as [H]; inversion H. + destruct X as [eq_mm' _]. + change (m =? m') with (eqb m m') in eq1. + destruct (eqb_spec m m') as [|F]. 1: discriminate. + apply F, eq_mm'. + Qed. + Equations isconv_branches' (Γ : context) (ind : inductive) (par : nat) @@ -1730,15 +1910,15 @@ Section Conversion. hx _ _ := { | Success h3 := yes ; - | Error e := Error e + | Error e h := no e } ; - | Error e := Error e + | Error e h := no e } ; - | @exist false _ := Error ( + | @exist false eq1 := Error ( FixRargMismatch idx (Γ ,,, stack_context π) u mfix1 mfix2 (Γ ,,, stack_context π') v mfix1' mfix2' - ) + ) _ } ; isconv_fix_types Γ idx mfix1 [] π h mfix1' [] π' h' hx h1 aux := yes ; @@ -1751,11 +1931,19 @@ Section Conversion. FixMfixMismatch idx (Γ ,,, stack_context π) (mfix1 ++ mfix2) (Γ ,,, stack_context π') (mfix1' ++ mfix2') - ). + ) _. Next Obligation. constructor. constructor. Qed. + Next Obligation. + (* Left list is empty *) + destruct H as [H]; inversion H. + Qed. + Next Obligation. + (* Right list is empty *) + destruct H as [H]; inversion H. + Qed. Next Obligation. destruct u. assumption. Qed. @@ -1817,6 +2005,22 @@ Section Conversion. clear eq1. intuition eauto. Qed. + Next Obligation. + (* Contrapositive of previous obligation *) + apply h; clear h. + destruct H as [H]; inversion H. + constructor; assumption. + Qed. + Next Obligation. + destruct H as [H]; inversion H; destruct X as [eq_uv _]. + apply h; clear h; constructor; apply eq_uv. + Qed. + Next Obligation. + destruct H as [H]; inversion H; destruct X as [_ eq_uv]. + change (?ru =? ?rv) with (eqb ru rv) in eq1. + destruct (eqb_spec (rarg u) (rarg v)) as [|neq_uv]; [discriminate|]. + exact (neq_uv eq_uv). + Qed. (* TODO MOVE *) Lemma conv_context_decl : @@ -1861,9 +2065,9 @@ Section Conversion. hx _ _ _ := { | Success h3 := yes ; - | Error e := Error e + | Error e h := no e } ; - | Error e := Error e + | Error e h := no e } ; isconv_fix_bodies Γ idx mfix1 [] π h mfix1' [] π' h' hx h1 ha aux := yes ; @@ -2043,6 +2247,16 @@ Section Conversion. rewrite <- !app_assoc in hh. simpl in hh. assumption. Qed. + Next Obligation. + apply h; clear h. + destruct H as [H]; inversion H; constructor. + rewrite map_app, <- app_assoc; simpl; assumption. + Qed. + Next Obligation. + apply h; clear h. + destruct H as [H]; inversion H; constructor. + rewrite app_context_assoc; apply X. + Qed. Equations isconv_fix (Γ : context) (mfix : mfixpoint term) (idx : nat) (π : stack) @@ -2073,9 +2287,9 @@ Section Conversion. hx _ _ _ := { | Success h2 := yes ; - | Error e := Error e + | Error e h := no e } ; - | Error e := Error e + | Error e h := no e }. Next Obligation. @@ -2099,6 +2313,23 @@ Section Conversion. intros [? ? ? ?] [? ? ? ?] ?. simpl in *. intuition eauto. Qed. + Next Obligation. + (* Contrapositive of previous obligation *) + apply h; clear h. + destruct H as [H]; constructor. + apply (All2_impl H). + rewrite <- fix_context_fix_context_alt. intuition. + Qed. + Next Obligation. + (* Contrapositive of pre-previous obligation *) + apply h; clear h. + destruct H as [H]; constructor; apply (All2_impl H). + intuition. + Qed. + + Lemma wellformed_zipc_tProd_appstack_nil {Γ na A B l ρ} : + wellformed Σ Γ (zipc (tProd na A B) (appstack l ρ)) -> l = []. + Proof. todo "Completeness"%string. Qed. Opaque reduce_stack. Equations(noeqns) _isconv_prog (Γ : context) (leq : conv_pb) @@ -2118,12 +2349,12 @@ Section Conversion. | @exist true eq2 with isconv_args_raw leq (tConst c u) π1 (tConst c' u') π2 aux := { | Success h := yes ; (* Unfold both constants at once *) - | Error e with inspect (lookup_env Σ c) := { + | Error e h with inspect (lookup_env Σ c) := { | @exist (Some (ConstantDecl {| cst_body := Some body |})) eq3 := isconv_red leq (subst_instance_constr u body) π1 (subst_instance_constr u' body) π2 aux ; (* Inductive or not found *) - | @exist _ _ := Error (NotFoundConstant c) + | @exist _ _ := Error (NotFoundConstant c) _ } } ; (* If universes are different, we unfold one of the constants *) @@ -2140,12 +2371,12 @@ Section Conversion. isconv_red leq t1 (Lambda_tm na A1 π1) t2 (Lambda_tm na' A2 π2) aux ; - | Error e := + | Error e h := Error ( LambdaNotConvertibleTypes (Γ ,,, stack_context π1) na A1 t1 (Γ ,,, stack_context π2) na' A2 t2 e - ) + ) _ } ; | prog_view_Prod na A1 B1 na' A2 B2 @@ -2154,12 +2385,12 @@ Section Conversion. isconv_red leq B1 (Prod_r na A1 π1) B2 (Prod_r na' A2 π2) aux ; - | Error e := + | Error e h := Error ( ProdNotConvertibleDomains (Γ ,,, stack_context π1) na A1 B1 (Γ ,,, stack_context π2) na' A2 B2 e - ) + ) _ } ; | prog_view_Case ind par p c brs ind' par' p' c' brs' @@ -2183,20 +2414,20 @@ Section Conversion. | Success h3 with isconv_args_raw leq (tCase (ind, par) p c brs) π1 (tCase (ind', par') p' c' brs') π2 aux := { | Success h4 := yes ; - | Error e := Error e + | Error e h := Error e h } ; - | Error e := Error e + | Error e h := Error e _ } ; - | Error e := Error e + | Error e h := Error e _ } ; - | Error e := Error e + | Error e h := Error e _ } ; | @exist false _ := Error ( CaseOnDifferentInd (Γ ,,, stack_context π1) ind par p c brs (Γ ,,, stack_context π2) ind' par' p' c' brs' - ) + ) _ } ; | @exist false eq3 := isconv_red leq (tCase (ind, par) p cred brs) π1 @@ -2210,14 +2441,14 @@ Section Conversion. | @exist true eq1 with isconv_red_raw Conv c (Proj p π1) c' (Proj p' π2) aux := { | Success h1 := isconv_args leq (tProj p c) π1 (tProj p' c') π2 aux ; - | Error e := Error e + | Error e h := Error e _ } ; | @exist false _ := Error ( DistinctStuckProj (Γ ,,, stack_context π1) p c (Γ ,,, stack_context π2) p' c' - ) + ) _ } ; | prog_view_Fix mfix idx mfix' idx' @@ -2245,16 +2476,16 @@ Section Conversion. | @exist true eq4 with isconv_fix Γ mfix idx π1 _ mfix' idx' π2 _ _ _ aux := { | Success h1 with isconv_args_raw leq (tFix mfix idx) π1 (tFix mfix' idx') π2 aux := { | Success h2 := yes ; - | Error e := Error e + | Error e h := Error e _ } ; - | Error e := Error e + | Error e h := Error e _ } ; | @exist false _ := Error ( CannotUnfoldFix (Γ ,,, stack_context π1) mfix idx (Γ ,,, stack_context π2) mfix' idx' - ) + )_ } } } @@ -2268,7 +2499,7 @@ Section Conversion. DistinctCoFix (Γ ,,, stack_context π1) mfix idx (Γ ,,, stack_context π2) mfix' idx' - ) (* TODO Incomplete *) + ) _ (* TODO Incomplete *) } ; | prog_view_other t1 t2 h := @@ -2364,6 +2595,69 @@ Section Conversion. destruct hΣ. now eapply conv_cum_Lambda. Qed. + Next Obligation. + (* Contrapositive of previous obligation *) + apply h; clear h. + destruct h0 as [h0]. + destruct hx as [hx]. + + unfold zipp in h0. simpl in h0. + unfold zipp in H. simpl in H. + unfold zipp. + case_eq (decompose_stack π1). intros l1 ρ1 e1. + case_eq (decompose_stack π2). intros l2 ρ2 e2. + pose proof (decompose_stack_eq _ _ _ e1). subst. + pose proof (decompose_stack_eq _ _ _ e2). subst. + rewrite stack_context_appstack in h0. + rewrite stack_context_appstack in H. + + destruct ir1 as [_ hl1]. cbn in hl1. + specialize (hl1 eq_refl). + destruct l1 ; try discriminate hl1. clear hl1. + + destruct ir2 as [_ hl2]. cbn in hl2. + specialize (hl2 eq_refl). + destruct l2 ; try discriminate hl2. clear hl2. + rewrite e1, e2 in H. + simpl in *. + + pose proof (wellformed_zipc_stack_context _ hΣ _ _ _ _ _ e1 h1) as + [wf]%wellformed_wf_local. + destruct hΣ. + + apply Lambda_conv_cum_inv in H. 2: assumption. + destruct H as [_ ?]; assumption. + Qed. + Next Obligation. + (* Contrapositive of previous obligation *) + apply h; clear h. + destruct hx as [hx]. + + unfold zipp in H. simpl in H. + unfold zipp. + case_eq (decompose_stack π1). intros l1 ρ1 e1. + case_eq (decompose_stack π2). intros l2 ρ2 e2. + pose proof (decompose_stack_eq _ _ _ e1). subst. + pose proof (decompose_stack_eq _ _ _ e2). subst. + rewrite stack_context_appstack in H. + + destruct ir1 as [_ hl1]. cbn in hl1. + specialize (hl1 eq_refl). + destruct l1 ; try discriminate hl1. clear hl1. + + destruct ir2 as [_ hl2]. cbn in hl2. + specialize (hl2 eq_refl). + destruct l2 ; try discriminate hl2. clear hl2. + rewrite e1, e2 in H. + simpl in *. + + pose proof (wellformed_zipc_stack_context _ hΣ _ _ _ _ _ e1 h1) as + [wf]%wellformed_wf_local. + destruct hΣ. + + apply Lambda_conv_cum_inv in H. 2: assumption. + destruct H as [[?] _]; now constructor. + Qed. (* tProd *) Next Obligation. @@ -2413,6 +2707,45 @@ Section Conversion. simpl. eapply conv_cum_Prod. all: auto. Qed. + Next Obligation. + (* Codomains are not convertible *) + apply h; clear h. + destruct hΣ as [wΣ], h0 as [h0], hx as [hx]. + unfold zipp in h0, H |- *. simpl in *. + + case_eq (decompose_stack π1). intros l1 ρ1 e1. + case_eq (decompose_stack π2). intros l2 ρ2 e2. + pose proof (decompose_stack_eq _ _ _ e1). subst. + pose proof (decompose_stack_eq _ _ _ e2). subst. + rewrite stack_context_appstack in h0, H |- *. + rewrite e1, e2 in H. + pose proof (wellformed_zipc_tProd_appstack_nil h1). + pose proof (wellformed_zipc_tProd_appstack_nil h2). + subst; simpl in *. + apply Prod_conv_cum_inv in H. 2: assumption. + now destruct H as [_ H]. + Qed. + Next Obligation. + (* Domains are not convertible *) + apply h; clear h. + + destruct hΣ as [wΣ], hx as [hx]. + unfold zipp in H |- *. simpl in *. + + case_eq (decompose_stack π1). intros l1 ρ1 e1. + case_eq (decompose_stack π2). intros l2 ρ2 e2. + pose proof (decompose_stack_eq _ _ _ e1). subst. + pose proof (decompose_stack_eq _ _ _ e2). subst. + rewrite stack_context_appstack in H |- *. + rewrite e1, e2 in H. + pose proof (wellformed_zipc_tProd_appstack_nil h1). + pose proof (wellformed_zipc_tProd_appstack_nil h2). + subst; simpl in *. + + apply Prod_conv_cum_inv in H. 2: assumption. + now destruct H. + Qed. + (* tCase *) Next Obligation. @@ -2486,6 +2819,20 @@ Section Conversion. subst. eapply conv_Case. all: assumption. Qed. + Next Obligation. + apply h; clear h. + Search "inv" "Case". + todo "Completeness"%string. + Qed. + Next Obligation. + todo "Completeness"%string. + Qed. + Next Obligation. + todo "Completeness"%string. + Qed. + Next Obligation. + todo "Completeness"%string. + Qed. Next Obligation. eapply red_wellformed ; auto. - exact h1. @@ -2613,6 +2960,9 @@ Section Conversion. induction brs' ; eauto. + eapply conv_context_sym. all: auto. Qed. + Next Obligation. + todo "Completeness"%string. + Qed. (* tProj *) Next Obligation. @@ -2634,6 +2984,12 @@ Section Conversion. constructor. eapply conv_Proj_c. assumption. Qed. + Next Obligation. + todo "Completeness"%string. + Qed. + Next Obligation. + todo "Completeness"%string. + Qed. (* tFix *) Next Obligation. @@ -2791,6 +3147,9 @@ Section Conversion. - eapply red_conv_cum_l. all: eassumption. - rewrite e. assumption. Qed. + Next Obligation. + todo "Completeness"%string. + Qed. Next Obligation. cbn. rewrite zipc_appstack. cbn. apply unfold_one_fix_red_zipp in eq1 as r. @@ -2945,6 +3304,9 @@ Section Conversion. assumption. - assumption. Qed. + Next Obligation. + todo "Completeness"%string. + Qed. Next Obligation. change (true = eqb idx idx') in eq4. destruct (eqb_spec idx idx'). 2: discriminate. @@ -2965,6 +3327,12 @@ Section Conversion. subst. eapply conv_Fix. all: assumption. Qed. + Next Obligation. + todo "Completeness"%string. + Qed. + Next Obligation. + todo "Completeness"%string. + Qed. (* tCoFix *) Next Obligation. @@ -2978,6 +3346,10 @@ Section Conversion. constructor. constructor. eapply eqb_term_spec. auto. Qed. + (* Beware !! There is a TODO in the code *) + Next Obligation. + todo "Completeness"%string. + Qed. (* Fallback *) Next Obligation. @@ -2985,7 +3357,7 @@ Section Conversion. all: try reflexivity. simpl. constructor. Qed. - + (* TODO MOVE *) Lemma App_conv' : forall leq Γ t1 t2 u1 u2, @@ -3035,21 +3407,21 @@ Section Conversion. with aux u1 u2 args1 l1 (coApp t2 (appstack l2 π2)) _ _ _ _ Conv _ I I I := { | Success H1 with _isconv_args' leq Γ t1 (args1 ++ [u1]) l1 π1 _ _ (tApp t2 u2) l2 π2 _ _ _ _ _ := { | Success H2 := yes ; - | Error e := + | Error e h := Error ( StackTailError leq (Γ ,,, stack_context π1) t1 args1 u1 l1 (Γ ,,, stack_context π2) t2 u2 l2 e - ) + ) _ } ; - | Error e := + | Error e h := Error ( StackHeadError leq (Γ ,,, stack_context π1) t1 args1 u1 l1 (Γ ,,, stack_context π2) t2 u2 l2 e - ) + ) _ } ; _isconv_args' leq Γ t1 args1 [] π1 h1 hπ1 t2 [] π2 h2 hπ2 hx h aux := yes ; @@ -3059,7 +3431,7 @@ Section Conversion. StackMismatch (Γ ,,, stack_context π1) t1 args1 l1 (Γ ,,, stack_context π2) t2 l2 - ). + ) _. Next Obligation. unfold zipp. case_eq (decompose_stack π1). intros l1 ρ1 e1. @@ -3072,6 +3444,12 @@ Section Conversion. rewrite stack_context_appstack in h. assumption. Defined. + Next Obligation. + todo "Completeness"%string. + Qed. + Next Obligation. + todo "Completeness"%string. + Qed. Next Obligation. split. 1: reflexivity. eapply positionR_poscat. constructor. @@ -3137,6 +3515,12 @@ Section Conversion. rewrite stack_context_appstack. assumption. Defined. + Next Obligation. + todo "Completeness"%string. + Qed. + Next Obligation. + todo "Completeness"%string. + Qed. Equations(noeqns) _isconv_args (leq : conv_pb) (Γ : context) (t1 : term) (π1 : stack) (h1 : wtp Γ t1 π1) @@ -3149,7 +3533,7 @@ Section Conversion. | @exist (l1, θ1) eq1 with inspect (decompose_stack π2) := { | @exist (l2, θ2) eq2 with _isconv_args' leq Γ t1 [] l1 θ1 _ _ t2 l2 θ2 _ _ _ _ _ := { | Success h := yes ; - | Error e := Error e (* TODO Is it sufficient? *) + | Error e h := Error e _ (* TODO Is it sufficient? *) } } }. @@ -3192,7 +3576,10 @@ Section Conversion. pose proof (decompose_stack_eq _ _ _ (eq_sym eq2)). subst. assumption. Qed. - + Next Obligation. + todo "Completeness"%string. + Qed. + Equations unfold_one_case (Γ : context) (ind : inductive) (par : nat) (p c : term) (brs : list (nat × term)) (h : wellformed Σ Γ (tCase (ind, par) p c brs)) : option term := @@ -3211,6 +3598,7 @@ Section Conversion. | ccview_other t _ := None } }. + Next Obligation. destruct hΣ as [wΣ]. cbn. destruct h as [[T h] | [[ctx [s [h1 _]]]]]; [| discriminate ]. @@ -3509,7 +3897,7 @@ Section Conversion. Conv (Γ ,,, stack_context π1) t1 (Γ ,,, stack_context π2) t2 - ) + ) _ } ; | @exist Cumul eq1 with inspect (leqb_term t1 t2) := { | @exist true eq2 := isconv_args Cumul t1 π1 t2 π2 aux ; @@ -3519,7 +3907,7 @@ Section Conversion. Cumul (Γ ,,, stack_context π1) t1 (Γ ,,, stack_context π2) t2 - ) + ) _ } } } @@ -3670,6 +4058,9 @@ Section Conversion. rewrite zipc_appstack in r2. cbn in r2. assumption. Qed. + Next Obligation. + todo "Completeness"%string. + Qed. Next Obligation. cbn. rewrite zipc_appstack. cbn. apply reducible_head_red_zipp in eq1 as r. @@ -3821,6 +4212,9 @@ Section Conversion. assumption. - eapply conv_context_sym. all: auto. Qed. + Next Obligation. + todo "Completeness"%string. + Qed. Next Obligation. eapply R_stateR. all: simpl. all: try reflexivity. constructor. @@ -3829,6 +4223,9 @@ Section Conversion. constructor. constructor. eapply eqb_term_spec. auto. Qed. + Next Obligation. + todo "Completeness"%string. + Qed. Next Obligation. eapply R_stateR. all: simpl. all: try reflexivity. constructor. @@ -3837,7 +4234,10 @@ Section Conversion. constructor. constructor. eapply leqb_term_spec. auto. Qed. - + Next Obligation. + todo "Completeness"%string. + Qed. + Equations _isconv (s : state) (Γ : context) (t1 : term) (π1 : stack) (h1 : wtp Γ t1 π1) (t2 : term) (π2 : stack) (h2 : wtp Γ t2 π2) @@ -3872,6 +4272,7 @@ Section Conversion. (fun pp f => _) (x := mkpack Γ s t1 π1 t2 π2 _) _ _ _ _. + Obligations. Next Obligation. unshelve eapply _isconv. 1: exact (st pp). all: try assumption. intros s' t1' π1' t2' π2' h1' h2' hx' hR. @@ -3884,12 +4285,16 @@ Section Conversion. Next Obligation. apply R_Acc. assumption. Qed. - - Definition isconv Γ leq t1 π1 h1 t2 π2 h2 hx := + + Program Definition isconv Γ leq t1 π1 h1 t2 π2 h2 hx := match isconv_full Reduction Γ t1 π1 h1 t2 π2 h2 leq hx I I I with | Success _ => Success I - | Error e => Error e + | Error e h => Error e _ end. + Next Obligation. + (* This goal is probably not solvable, should probably change the definition of ConversionResult *) + todo "Completeness"%string. + Qed. Theorem isconv_sound : forall Γ leq t1 π1 h1 t2 π2 h2 hx, From b37c97f15f7f045d3cbf6d106bd9825f52477eed Mon Sep 17 00:00:00 2001 From: Kenji Maillard Date: Tue, 22 Sep 2020 14:13:40 +0200 Subject: [PATCH 03/26] getting to reduction troubles --- pcuic/theories/PCUICConversion.v | 29 ++++++++--- safechecker/theories/PCUICSafeConversion.v | 60 +++++++++++++++++++++- 2 files changed, 81 insertions(+), 8 deletions(-) diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index a5efa9ad5..4b7a0b15d 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -337,7 +337,6 @@ Proof. repeat constructor. Qed. -Derive Signature for conv. Lemma conv_Prod_l_inv {cf:checker_flags} (Σ : global_env_ext) Γ na dom codom T : wf Σ -> @@ -1435,11 +1434,10 @@ Section Inversions. Lemma conv_Case_brs : forall Γ indn p c brs brs', - wf Σ -> All2 (fun u v => u.1 = v.1 × Σ ;;; Γ |- u.2 = v.2) brs brs' -> Σ ;;; Γ |- tCase indn p c brs = tCase indn p c brs'. Proof. - intros Γ [ind n] p c brs brs' wΣ h. + intros Γ [ind n] p c brs brs' h. apply All2_many_OnOne2 in h. induction h. - reflexivity. @@ -1450,13 +1448,12 @@ Section Inversions. Lemma conv_Case : forall Γ indn p p' c c' brs brs', - wf Σ -> Σ ;;; Γ |- p = p' -> Σ ;;; Γ |- c = c' -> All2 (fun u v => u.1 = v.1 × Σ ;;; Γ |- u.2 = v.2) brs brs' -> Σ ;;; Γ |- tCase indn p c brs = tCase indn p' c' brs'. Proof. - intros Γ [ind n] p p' c c' brs brs' wΣ hp hc hbrs. + intros Γ [ind n] p p' c c' brs brs' hp hc hbrs. etransitivity. - eapply conv_Case_p. eassumption. - etransitivity. @@ -1467,12 +1464,30 @@ Section Inversions. Lemma Case_conv_inv : forall Γ indn p p' c c' brs brs', Σ ;;; Γ |- tCase indn p c brs = tCase indn p' c' brs' -> + Σ ;;; Γ |- p = p' × Σ ;;; Γ |- c = c' × + All2 (fun u v => u.1 = v.1 × Σ ;;; Γ |- u.2 = v.2) brs brs'. + Proof. + (* intros Γ indn p p' c c' brs brs' H. + depind H. + + depelim e; split; [| split]. + 1,2: by apply conv_refl. + apply (All2_impl a). + intros x y [? ?] ; split. 1: by []. + by apply conv_refl. *) + (* TODO needs to add a whne hypothesis *) + todo "Completeness"%string. + Qed. + +(* + Lemma Case_conv_cum_inv : + forall leq Γ indn p p' c c' brs brs', + conv_cum leq (tCase indn p c brs) (tCase indn p' c' brs') -> Σ ;;; Γ |- p = p' /\ Σ ;;; Γ |- c = c' /\ ∥ All2 (fun u v => u.1 = v.1 × Σ ;;; Γ |- u.2 = v.2) brs brs' ∥. Proof. todo "Completeness"%string. - Qed. - + Qed. *) + Lemma conv_Proj_c : forall Γ p u v, Σ ;;; Γ |- u = v -> diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 541702927..fe23ae172 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -2329,7 +2329,59 @@ Section Conversion. Lemma wellformed_zipc_tProd_appstack_nil {Γ na A B l ρ} : wellformed Σ Γ (zipc (tProd na A B) (appstack l ρ)) -> l = []. - Proof. todo "Completeness"%string. Qed. + Proof. + (* Proof idea: a sort cannot be applied to anything *) + todo "Completeness"%string. + Qed. + + Notation whne := (whne flags Σ). + + Lemma app_conv_inv : forall Γ t t' u u', + (forall napp, eq_term_upto_univ_napp Σ.1 (eq_universe Σ) (eq_universe Σ) napp t t' -> + eq_term Σ t t') -> + whne Γ t -> + whne Γ t' -> + Σ ;;; Γ |- tApp t u = tApp t' u' -> + Σ ;;; Γ |- t = t'. + Proof. + intros Γ t t' u u' cum_irr wh wh' hconv. + depind hconv. + + intros **. depelim e. eapply conv_refl, cum_irr; eassumption. + + (* Proof idea: invert the reduction r using thw whne hyp to apply the IH *) + todo "Completeness"%string. + + (* Proof idea: invert the reduction r using thw whne hyp to apply the IH *) + todo "Completeness"%string. + Qed. + + + + Lemma zipp_Case_conv_inv : + forall Γ indn p p' c c' brs brs' π π', + let t := tCase indn p c brs in + let t' := tCase indn p' c' brs' in + whne Γ t -> whne Γ t' -> + Σ ;;; Γ |- zipp t π = zipp t' π' -> + Σ ;;; Γ |- t = t'. + Proof. + (* Proof idea: decompose the stacks, prove by induction that the + each side is applied to the same number of arguments. *) + intros Γ indn p p' c c' brs brs' π π' t t' wht wht'. + + destruct (decompose_stack π) as [l s] eqn:Hpi. + destruct (decompose_stack π') as [l' s'] eqn: Hpi'. + unfold zipp; rewrite Hpi, Hpi'; clear s s' Hpi Hpi'. + induction l as [|x l IH] in l' |- *; destruct l' as [|x' l']. + + trivial. + + simpl. intro H. depelim H. + + + assert (cum_irr : forall napp, eq_term_upto_univ_napp Σ.1 (eq_universe Σ) (eq_universe Σ) napp t t' -> eq_term Σ t t'). + - intros n heq; depelim heq; now constructor. + - + + + todo "Completeness"%string. + Qed. Opaque reduce_stack. Equations(noeqns) _isconv_prog (Γ : context) (leq : conv_pb) @@ -2821,6 +2873,12 @@ Section Conversion. Qed. Next Obligation. apply h; clear h. + destruct leq. + + destruct H as [H]; cbn in H. + unfold zipp in H. + destruct (decompose_stack π1) eqn:H1; subst. + destruct (decompose_stack π2) eqn:H2; subst. + cbn in *. Search "inv" "Case". todo "Completeness"%string. Qed. From 9d687112b7d5ca6437d1aaf53a8263b08b59a4c7 Mon Sep 17 00:00:00 2001 From: Kenji Maillard Date: Tue, 22 Sep 2020 21:30:22 +0200 Subject: [PATCH 04/26] Proof ideas on the completeness holes, safechecker compiling --- pcuic/theories/PCUICCumulativity.v | 1 + safechecker/theories/PCUICSafeChecker.v | 10 +- safechecker/theories/PCUICSafeConversion.v | 126 ++++++++++++++++----- 3 files changed, 103 insertions(+), 34 deletions(-) diff --git a/pcuic/theories/PCUICCumulativity.v b/pcuic/theories/PCUICCumulativity.v index 5848e37ed..743fa4b23 100644 --- a/pcuic/theories/PCUICCumulativity.v +++ b/pcuic/theories/PCUICCumulativity.v @@ -249,6 +249,7 @@ Proof. econstructor 2; eauto. Qed. +(* This should be called sym instead of inv, no ? *) Lemma red_conv_conv_inv `{cf : checker_flags} Σ Γ t u v : red (fst Σ) Γ t u -> Σ ;;; Γ |- v = u -> Σ ;;; Γ |- v = t. Proof. diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index fcad4449f..d1bc8dba8 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -483,8 +483,8 @@ Section Typecheck. : typing_result (∥ Σ ;;; Γ |- t <= u ∥) := match leqb_term Σ G t u with true => ret _ | false => match iscumul Γ t ht u hu with - | Success _ => ret _ - | Error e h => (* fallback *) (* nico *) + | ConvSuccess => ret _ + | ConvError e => (* fallback *) (* nico *) let t' := hnf Γ t ht in let u' := hnf Γ u hu in (* match leq_term (snd Σ) t' u' with true => ret _ | false => *) @@ -753,8 +753,8 @@ Section Typecheck. (* We could avoid one useless sort comparison by only comparing *) (* the contexts [pctx] and [indctx] (what is done in Coq). *) match iscumul Γ pty.π1 _ pty' _ with - | Error e => raise (NotCumulSmaller G Γ pty.π1 pty' pty.π1 pty' e) - | Success _ => + | ConvError e => raise (NotCumulSmaller G Γ pty.π1 pty' pty.π1 pty' e) + | ConvSuccess => match map_option_out (build_branches_type ind decl body params u p) with | None => raise (Msg "failure in build_branches_type") | Some btys => @@ -994,7 +994,7 @@ Section Typecheck. Defined. Next Obligation. - rename Heq_anonymous2 into XX2. destruct wildcard'. + rename Heq_anonymous2 into XX2. symmetry in XX2. simpl in *. eapply isconv_sound in XX2. change (eqb ind ind' = true) in H0. destruct (eqb_spec ind ind') as [e|e]; [destruct e|discriminate H0]. diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index fe23ae172..7a8953df9 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -121,7 +121,8 @@ Qed. Section Conversion. Context {cf : checker_flags}. - Context (flags : RedFlags.t). + (* Unused *) + (* Context (flags : RedFlags.t). *) Context (Σ : global_env_ext). Context (hΣ : ∥ wf Σ ∥) (Hφ : ∥ on_udecl Σ.1 Σ.2 ∥). Context (G : universes_graph) (HG : is_graph_of_uctx G (global_ext_uctx Σ)). @@ -2334,6 +2335,7 @@ Section Conversion. todo "Completeness"%string. Qed. +(* Notation whne := (whne flags Σ). Lemma app_conv_inv : forall Γ t t' u u', @@ -2366,7 +2368,7 @@ Section Conversion. (* Proof idea: decompose the stacks, prove by induction that the each side is applied to the same number of arguments. *) intros Γ indn p p' c c' brs brs' π π' t t' wht wht'. - +(* destruct (decompose_stack π) as [l s] eqn:Hpi. destruct (decompose_stack π') as [l' s'] eqn: Hpi'. unfold zipp; rewrite Hpi, Hpi'; clear s s' Hpi Hpi'. @@ -2378,10 +2380,17 @@ Section Conversion. assert (cum_irr : forall napp, eq_term_upto_univ_napp Σ.1 (eq_universe Σ) (eq_universe Σ) napp t t' -> eq_term Σ t t'). - intros n heq; depelim heq; now constructor. - - + *) todo "Completeness"%string. - Qed. + Qed. *) + + Lemma conv_red_full_l : + forall Σ Γ (t u v : term), + red Σ.1 Γ v t -> Σ;;; Γ |- v = u -> Σ;;; Γ |- t = u. + Proof. + admit. + Admitted. Opaque reduce_stack. Equations(noeqns) _isconv_prog (Γ : context) (leq : conv_pb) @@ -2474,7 +2483,7 @@ Section Conversion. } ; | Error e h := Error e _ } ; - | @exist false _ := + | @exist false eqDiff := Error ( CaseOnDifferentInd (Γ ,,, stack_context π1) ind par p c brs @@ -2495,7 +2504,7 @@ Section Conversion. | Success h1 := isconv_args leq (tProj p c) π1 (tProj p' c') π2 aux ; | Error e h := Error e _ } ; - | @exist false _ := + | @exist false eqDiff := Error ( DistinctStuckProj (Γ ,,, stack_context π1) p c @@ -2506,7 +2515,7 @@ Section Conversion. | prog_view_Fix mfix idx mfix' idx' with inspect (eqb_term (tFix mfix idx) (tFix mfix' idx')) := { | @exist true eq1 := isconv_args leq (tFix mfix idx) π1 (tFix mfix' idx') π2 aux ; - | @exist false _ with inspect (unfold_one_fix Γ mfix idx π1 _) := { + | @exist false eqDiff with inspect (unfold_one_fix Γ mfix idx π1 _) := { | @exist (Some (fn, θ)) eq1 with inspect (decompose_stack θ) := { | @exist (l', θ') eq2 @@ -2873,22 +2882,34 @@ Section Conversion. Qed. Next Obligation. apply h; clear h. - destruct leq. - + destruct H as [H]; cbn in H. - unfold zipp in H. - destruct (decompose_stack π1) eqn:H1; subst. - destruct (decompose_stack π2) eqn:H2; subst. - cbn in *. - Search "inv" "Case". + (* Proof idea: + 1) prove that the 2 tCase are in conv_cum relation + because they are whne (needs to add thys hypothesis) + 2) wlog leq = Conv because there is no cumulativity between tCase + 3) invert the conv judgement using that the tCase are whne + (so preserved by wh-reductions) + *) todo "Completeness"%string. Qed. Next Obligation. + apply h ; clear h. + cbn. (* Same idea as previous case *) todo "Completeness"%string. Qed. Next Obligation. + apply h; clear h. + cbn. (* Same idea as previous case *) todo "Completeness"%string. Qed. Next Obligation. + (* Proof idea: + 1) from H, derive that some + tCase (ind, par) _ _ _ and tCase (ind',par') _ _ _ + are in eqb_term + maybe using conv_alt_red + 2) deduce that ind = ind' and par = par' + 3) conclude that it contradicts eqDiff + *) todo "Completeness"%string. Qed. Next Obligation. @@ -3019,7 +3040,30 @@ Section Conversion. + eapply conv_context_sym. all: auto. Qed. Next Obligation. - todo "Completeness"%string. + apply h; clear h. + set (redc := reduce_term _ _ _ _ _ _); + assert (∥ red Σ.1 (Γ ,,, stack_context π1) c redc ∥) + as [hc] + by apply reduce_term_sound. + set (redc' := reduce_term _ _ _ _ _ _); + assert (∥ red Σ.1 (Γ ,,, stack_context π2) c' redc' ∥) + as [hc'] + by apply reduce_term_sound. + destruct hx as [hx]. + destruct hΣ as [ wfΣ]. + (* Proof idea: wlog leq = Conv, *) + destruct leq. + - set (t0 := zipp _ _). + destruct H as [Hconv]. + constructor. + eapply conv_red_full_l. 1: apply red_zipp, red_case_c, hc. + symmetry. + eapply conv_conv_ctx. 1: assumption. + 2: eapply conv_context_sym; eassumption. + eapply conv_red_full_l. 1: apply red_zipp, red_case_c, hc'. + eapply conv_conv_ctx; try eassumption. + symmetry; assumption. + - todo "Completeness"%string. (* see previous comment *) Qed. (* tProj *) @@ -3043,9 +3087,12 @@ Section Conversion. eapply conv_Proj_c. assumption. Qed. Next Obligation. + apply h; clear h. + cbn. todo "Completeness"%string. Qed. Next Obligation. + (* Proof idea: c and c' are whne so from H, p = p' contradicting eqDiff *) todo "Completeness"%string. Qed. @@ -3206,6 +3253,8 @@ Section Conversion. - rewrite e. assumption. Qed. Next Obligation. + apply h; clear h. + (* Idea: a fixpoint should be convertible to its unfolding + *) todo "Completeness"%string. Qed. Next Obligation. @@ -3363,6 +3412,8 @@ Section Conversion. - assumption. Qed. Next Obligation. + apply h; clear h. + (* Idea: a fixpoint should be convertible to its unfolding + *) todo "Completeness"%string. Qed. Next Obligation. @@ -3386,9 +3437,15 @@ Section Conversion. eapply conv_Fix. all: assumption. Qed. Next Obligation. + apply h; clear h. + (* Idea : The fixpoints are neutral (do not unfold) *) + (* so each definitions are convertible *) todo "Completeness"%string. Qed. Next Obligation. + (* Proof idea: reduce to the case where we have eq_term + between tFix using that the fixpoints do not unfold + and then conlude with eqDiff *) todo "Completeness"%string. Qed. @@ -3465,7 +3522,7 @@ Section Conversion. with aux u1 u2 args1 l1 (coApp t2 (appstack l2 π2)) _ _ _ _ Conv _ I I I := { | Success H1 with _isconv_args' leq Γ t1 (args1 ++ [u1]) l1 π1 _ _ (tApp t2 u2) l2 π2 _ _ _ _ _ := { | Success H2 := yes ; - | Error e h := + | Error e herr := Error ( StackTailError leq @@ -3473,7 +3530,7 @@ Section Conversion. (Γ ,,, stack_context π2) t2 u2 l2 e ) _ } ; - | Error e h := + | Error e herr := Error ( StackHeadError leq @@ -3574,9 +3631,11 @@ Section Conversion. assumption. Defined. Next Obligation. + apply herr; clear herr. todo "Completeness"%string. Qed. Next Obligation. + apply herr; clear herr. todo "Completeness"%string. Qed. @@ -3591,7 +3650,7 @@ Section Conversion. | @exist (l1, θ1) eq1 with inspect (decompose_stack π2) := { | @exist (l2, θ2) eq2 with _isconv_args' leq Γ t1 [] l1 θ1 _ _ t2 l2 θ2 _ _ _ _ _ := { | Success h := yes ; - | Error e h := Error e _ (* TODO Is it sufficient? *) + | Error e h := Error e _ } } }. @@ -3635,7 +3694,10 @@ Section Conversion. assumption. Qed. Next Obligation. - todo "Completeness"%string. + apply h; clear h. + pose proof (decompose_stack_eq _ _ _ (eq_sym eq1)). subst. + pose proof (decompose_stack_eq _ _ _ (eq_sym eq2)). subst. + assumption. Qed. Equations unfold_one_case (Γ : context) (ind : inductive) (par : nat) @@ -4117,6 +4179,8 @@ Section Conversion. assumption. Qed. Next Obligation. + apply h; clear h. + (* Contrapositive of previous case *) todo "Completeness"%string. Qed. Next Obligation. @@ -4271,6 +4335,8 @@ Section Conversion. - eapply conv_context_sym. all: auto. Qed. Next Obligation. + apply h; clear h. + (* Contrapositive of previous case *) todo "Completeness"%string. Qed. Next Obligation. @@ -4282,6 +4348,7 @@ Section Conversion. eapply eqb_term_spec. auto. Qed. Next Obligation. + (* Proof idea: t1, t2 are un whnf, convertible but distinct ? *) todo "Completeness"%string. Qed. Next Obligation. @@ -4293,6 +4360,7 @@ Section Conversion. eapply leqb_term_spec. auto. Qed. Next Obligation. + (* Proof idea: t1, t2 are un whnf, in cumulative relation but not in leqb_term relation ? *) todo "Completeness"%string. Qed. @@ -4344,19 +4412,19 @@ Section Conversion. apply R_Acc. assumption. Qed. - Program Definition isconv Γ leq t1 π1 h1 t2 π2 h2 hx := + Inductive ConversionResultSummary := + | ConvSuccess : ConversionResultSummary + | ConvError : ConversionError -> ConversionResultSummary. + + Definition isconv Γ leq t1 π1 h1 t2 π2 h2 hx := match isconv_full Reduction Γ t1 π1 h1 t2 π2 h2 leq hx I I I with - | Success _ => Success I - | Error e h => Error e _ + | Success _ => ConvSuccess + | Error e h => ConvError e end. - Next Obligation. - (* This goal is probably not solvable, should probably change the definition of ConversionResult *) - todo "Completeness"%string. - Qed. - + Theorem isconv_sound : forall Γ leq t1 π1 h1 t2 π2 h2 hx, - isconv Γ leq t1 π1 h1 t2 π2 h2 hx = Success I -> + isconv Γ leq t1 π1 h1 t2 π2 h2 hx = ConvSuccess -> conv_cum leq Σ (Γ ,,, stack_context π1) (zipp t1 π1) (zipp t2 π2). Proof. unfold isconv. @@ -4371,7 +4439,7 @@ Section Conversion. Theorem isconv_term_sound : forall Γ leq t1 h1 t2 h2, - isconv_term Γ leq t1 h1 t2 h2 = Success I -> + isconv_term Γ leq t1 h1 t2 h2 = ConvSuccess -> conv_cum leq Σ Γ t1 t2. Proof. intros Γ leq t1 h1 t2 h2. From 0c2d3a73fca103af6fbfb285ff735d77d965aa64 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Mon, 19 Oct 2020 15:35:15 +0200 Subject: [PATCH 05/26] Some universe stuff for completeness --- safechecker/theories/PCUICSafeConversion.v | 79 ++++++++++++++-------- 1 file changed, 50 insertions(+), 29 deletions(-) diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 7a8953df9..5cb2582b7 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -1514,18 +1514,59 @@ Section Conversion. now eapply All_local_env_app in wf. Qed. - (* Q: Shouldnt't we have as hypothesis that [c <> c'] ? *) + (* TODO MOVE *) + Definition eqb_universe_instance u v := + forallb2 (check_eqb_universe G) (map Universe.make u) (map Universe.make v). + + (* TODO MOVE *) + Lemma eqb_universe_instance_spec : + forall u v, + eqb_universe_instance u v -> + R_universe_instance (eq_universe (global_ext_constraints Σ)) u v. + Proof. + intros u v e. + unfold eqb_universe_instance in e. + eapply forallb2_Forall2 in e. + eapply Forall2_impl. 1: eassumption. + intros. eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). + all: auto. + - eapply wf_ext_global_uctx_invariants. + eapply hΣ'. + - eapply global_ext_uctx_consistent. + eapply hΣ'. + Qed. + + Definition eq_termp leq v v' := + match leq with + | Conv => eq_term Σ v v' + | Cumul => leq_term Σ Σ v v' + end. + + Lemma conv_cum_alt leq Γ t t' : + conv_cum leq Σ Γ t t' -> + ∥∑ v v', (red Σ Γ t v × red Σ Γ t' v') × eq_termp leq v v'∥. + Proof. + destruct leq; cbn. + - intros [c]. + apply conv_alt_red in c. + now constructor. + - intros [c]. + apply cumul_alt in c. + now constructor. + Qed. + (* Also, is there an invariant of injectivity on the assignment between kernames and inductives *) Equations(noeqns) unfold_constants (Γ : context) (leq : conv_pb) (c : kername) (u : Instance.t) (π1 : stack) (h1 : wtp Γ (tConst c u) π1) (c' : kername) (u' : Instance.t) (π2 : stack) (h2 : wtp Γ (tConst c' u') π2) + (ne : c <> c' \/ (c = c' /\ ~eqb_universe_instance u u')) (hx : conv_stack_ctx Γ π1 π2) (aux : Aux Term Γ (tConst c u) π1 (tConst c' u') π2 h2) : ConversionResult (conv_term leq Γ (tConst c u) π1 (tConst c' u') π2) := - unfold_constants Γ leq c u π1 h1 c' u' π2 h2 hx aux + unfold_constants Γ leq c u π1 h1 c' u' π2 h2 ne hx aux with inspect (lookup_env Σ c') := { | @exist (Some (ConstantDecl {| cst_body := Some b |})) eq1 := isconv_red leq (tConst c u) π1 (subst_instance_constr u' b) π2 aux ; @@ -1588,9 +1629,12 @@ Section Conversion. eapply red_zipp. eapply red_const. eassumption. Qed. Next Obligation. - (* No definition to unfold... - Should we have an hypothesis ditching that case ? *) - todo "Completeness"%string. + (* Both c and c' are axioms. Either they are different constants or they are not + convertible because the universes are different. *) + destruct ne. + - apply conv_cum_alt in H as [(?&?&(r1&r2)&e)]. + unfold zipp in r1, r2. + eapply red_zipp_inv in r1. Qed. Next Obligation. (* inductive vs non-defined constant *) @@ -1672,28 +1716,6 @@ Section Conversion. Qed. - (* TODO MOVE *) - Definition eqb_universe_instance u v := - forallb2 (check_eqb_universe G) (map Universe.make u) (map Universe.make v). - - (* TODO MOVE *) - Lemma eqb_universe_instance_spec : - forall u v, - eqb_universe_instance u v -> - R_universe_instance (eq_universe (global_ext_constraints Σ)) u v. - Proof. - intros u v e. - unfold eqb_universe_instance in e. - eapply forallb2_Forall2 in e. - eapply Forall2_impl. 1: eassumption. - intros. eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). - all: auto. - - eapply wf_ext_global_uctx_invariants. - eapply hΣ'. - - eapply global_ext_uctx_consistent. - eapply hΣ'. - Qed. - (* TODO (RE)MOVE *) Lemma destArity_eq_term_upto_univ : forall Re Rle Γ1 Γ2 t1 t2 Δ1 s1, @@ -2402,8 +2424,7 @@ Section Conversion. : ConversionResult (conv_term leq Γ t1 π1 t2 π2) := _isconv_prog Γ leq t1 π1 h1 t2 π2 h2 hx ir1 ir2 aux with prog_viewc t1 t2 := { - | prog_view_App _ _ _ _ := - False_rect _ _ ; + | prog_view_App _ _ _ _ := !; | prog_view_Const c u c' u' with eq_dec c c' := { | left eq1 with inspect (eqb_universe_instance u u') := { From 23a75418b7b345f99071bb0b46b3a032c5ff298a Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Tue, 20 Oct 2020 22:14:56 +0200 Subject: [PATCH 06/26] Finish tConst case Also add an induction principle for red1 t t' /\ whne t -> P since we will be using that a lot. --- pcuic/theories/PCUICEquality.v | 6 +- pcuic/theories/PCUICNormal.v | 500 +++++++++++++++----- pcuic/theories/PCUICReduction.v | 14 + pcuic/theories/PCUICSafeLemmata.v | 19 +- safechecker/theories/PCUICSafeConversion.v | 506 ++++++++++++++------- safechecker/theories/PCUICSafeReduce.v | 3 + template-coq/theories/common/uGraph.v | 18 +- 7 files changed, 761 insertions(+), 305 deletions(-) diff --git a/pcuic/theories/PCUICEquality.v b/pcuic/theories/PCUICEquality.v index 664727314..ed3d446cc 100644 --- a/pcuic/theories/PCUICEquality.v +++ b/pcuic/theories/PCUICEquality.v @@ -1680,11 +1680,11 @@ Proof. Qed. -Lemma eq_term_upto_univ_mkApps_inv Σ Re u l u' l' : +Lemma eq_term_upto_univ_mkApps_inv Σ Re Rle u l u' l' : isApp u = false -> isApp u' = false -> - eq_term_upto_univ Σ Re Re (mkApps u l) (mkApps u' l') -> - eq_term_upto_univ_napp Σ Re Re #|l| u u' * All2 (eq_term_upto_univ Σ Re Re) l l'. + eq_term_upto_univ Σ Re Rle (mkApps u l) (mkApps u' l') -> + eq_term_upto_univ_napp Σ Re Rle #|l| u u' * All2 (eq_term_upto_univ Σ Re Re) l l'. Proof. intros hu hu' h. apply eq_term_upto_univ_mkApps_l_inv in h as hh. diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index 58c7ffca4..853b250e1 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -3,7 +3,7 @@ From Coq Require Import Bool String List Program BinPos Compare_dec Arith Lia. From MetaCoq.Template Require Import config Universes monad_utils utils BasicAst AstUtils UnivSubst. -From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICTyping PCUICInduction. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICLiftSubst PCUICTyping PCUICInduction. Require Import ssreflect. Set Asymmetric Patterns. @@ -476,19 +476,22 @@ Qed. Lemma red1_mkApps_tCoFix_inv Σ Γ mfix id v t' : red1 Σ Γ (mkApps (tCoFix mfix id) v) t' -> - ∑ mfix' v', t' = mkApps (tCoFix mfix' id) v'. + (∑ v', (t' = mkApps (tCoFix mfix id) v') * (OnOne2 (red1 Σ Γ) v v')) + + (∑ mfix', (t' = mkApps (tCoFix mfix' id) v) * (OnOne2 (on_Trel_eq (red1 Σ Γ) dtype (fun x0 : def term => (dname x0, dbody x0, rarg x0))) mfix mfix')) + + (∑ mfix', (t' = mkApps (tCoFix mfix' id) v) * (OnOne2 (on_Trel_eq (red1 Σ (Γ ,,, PCUICLiftSubst.fix_context mfix)) dbody (fun x0 : def term => (dname x0, dtype x0, rarg x0))) mfix mfix')). Proof. revert t'. induction v using rev_ind; intros. - cbn in *. depelim X; help; eauto. - + now eexists _, []. - + now eexists _, []. - rewrite mkApps_snoc in X. depelim X; help; eauto. - + eapply IHv in X as (?&?&->). - exists x0, (x1 ++ [x]). - now rewrite <- mkApps_nested. - + exists mfix, (v ++ [N2]). - now rewrite <- mkApps_nested. + + eapply IHv in X as [ [(? & ? & ?)|(?&?&?)] | (?&?&?)]. + * subst. left. left. exists (x0 ++ [x]). split. now rewrite mkApps_snoc. + now apply OnOne2_app_r. + * subst. left. right. exists x0. split. now rewrite mkApps_snoc. eauto. + * subst. right. exists x0. split. now rewrite mkApps_snoc. eauto. + + left. left. exists (v ++ [N2]). split. + now rewrite mkApps_snoc. + eapply OnOne2_app. econstructor. eauto. Qed. Lemma whne_isConstruct_app flags Σ Γ t : @@ -509,34 +512,166 @@ Proof. solve_discr. Qed. -Lemma whne_pres1 Σ Γ t t' : - red1 Σ Γ t t' -> - whne RedFlags.default Σ Γ t -> - whne RedFlags.default Σ Γ t'. +Section whne_red1_ind. +Context (flags : RedFlags.t). +Context (Σ : global_env). +Context (Γ : context). +Context (P : term -> term -> Prop). + +Lemma whne_red1_ind + (Hrel : forall i body, + RedFlags.zeta flags = false -> + option_map decl_body (nth_error Γ i) = Some (Some body) -> + P (tRel i) (lift0 (S i) body)) + (Hlam_nobeta_1 : forall na A b M', + RedFlags.beta flags = false -> + red1 Σ Γ A M' -> + P (tLambda na A b) (tLambda na M' b)) + (Hlam_nobeta_2 : forall na A b M', + RedFlags.beta flags = false -> + red1 Σ (Γ,, vass na A) b M' -> + P (tLambda na A b) (tLambda na A M')) + (Hevar : forall n l l', + OnOne2 (red1 Σ Γ) l l' -> + P (tEvar n l) (tEvar n l')) + (Hletin_nozeta_1 : + forall t' na B b t, + RedFlags.zeta flags = false -> + red1 Σ Γ (tLetIn na B b t) t' -> + P (tLetIn na B b t) t') + (Hconst_nodelta : forall t' c u, + RedFlags.delta flags = false -> + red1 Σ Γ (tConst c u) t' -> + P (tConst c u) t') + (Hbeta_nobeta : + forall na t b v, + RedFlags.beta flags = false -> + P (tApp (tLambda na t b) v) (b{0 := v})) + (Hfix_nofix : + forall v mfix idx args narg fn, + RedFlags.fix_ flags = false -> + unfold_fix mfix idx = Some (narg, fn) -> + is_constructor narg (args ++ [v]) = true -> + P (mkApps (tFix mfix idx) (args ++ [v])) (mkApps fn (args ++ [v]))) + (Happ_1 : forall f v N1, + whne flags Σ Γ f -> + red1 Σ Γ f N1 -> + P f N1 -> + P (tApp f v) (tApp N1 v)) + (Happ_2 : forall f v N2, + whne flags Σ Γ f -> + red1 Σ Γ v N2 -> + P (tApp f v) (tApp f N2)) + (Hfix_red_arg : forall mfix idx args rarg arg body x, + unfold_fix mfix idx = Some (rarg, body) -> + nth_error args rarg = Some arg -> + whne flags Σ Γ arg -> + OnOne2 (fun a a' => red1 Σ Γ a a') args x -> + (forall t', red1 Σ Γ arg t' -> P arg t') -> + P (mkApps (tFix mfix idx) args) (mkApps (tFix mfix idx) x)) + (Hfix_red_def_type : forall mfix idx args rarg arg body x, + unfold_fix mfix idx = Some (rarg, body) -> + nth_error args rarg = Some arg -> + whne flags Σ Γ arg -> + OnOne2 + (fun x y : def term => + red1 Σ Γ (dtype x) (dtype y) * + ((dname x, dbody x, BasicAst.rarg x) = + (dname y, dbody y, BasicAst.rarg y))) mfix x -> + P (mkApps (tFix mfix idx) args) (mkApps (tFix x idx) args)) + (Hfix_red_def_body : forall mfix idx args rarg arg body x, + unfold_fix mfix idx = Some (rarg, body) -> + nth_error args rarg = Some arg -> + whne flags Σ Γ arg -> + OnOne2 + (fun x y : def term => + red1 Σ (Γ ,,, fix_context mfix) (dbody x) (dbody y) * + ((dname x, dtype x, BasicAst.rarg x) = + (dname y, dtype y, BasicAst.rarg y))) mfix x -> + P (mkApps (tFix mfix idx) args) (mkApps (tFix x idx) args)) + (Hfix_red_type_nofix : forall defs i mfix1, + RedFlags.fix_ flags = false -> + OnOne2 + (fun x y : def term => + red1 Σ Γ (dtype x) (dtype y) * + ((dname x, dbody x, rarg x) = (dname y, dbody y, rarg y))) + defs mfix1 -> + P (tFix defs i) (tFix mfix1 i)) + (Hfix_red_body_nofix : forall defs i mfix1, + RedFlags.fix_ flags = false -> + OnOne2 + (fun x y : def term => + red1 Σ (Γ ,,, fix_context defs) (dbody x) (dbody y) * + ((dname x, dtype x, rarg x) = (dname y, dtype y, rarg y))) + defs mfix1 -> + P (tFix defs i) (tFix mfix1 i)) + (Hcase_discr : forall i p c brs p', + whne flags Σ Γ c -> + red1 Σ Γ p p' -> + P (tCase i p c brs) (tCase i p' c brs)) + (Hcase_motive : forall i p c brs c', + whne flags Σ Γ c -> + red1 Σ Γ c c' -> + P c c' -> + P (tCase i p c brs) (tCase i p c' brs)) + (Hcase_branch : forall i p c brs brs', + whne flags Σ Γ c -> + OnOne2 (on_Trel_eq (red1 Σ Γ) snd fst) brs brs' -> + P (tCase i p c brs) (tCase i p c brs')) + (Hcase_noiota : forall t' i p c brs, + RedFlags.iota flags = false -> + red1 Σ Γ (tCase i p c brs) t' -> + P (tCase i p c brs) t') + (Hproj : forall p c c', + whne flags Σ Γ c -> + red1 Σ Γ c c' -> + P c c' -> + P (tProj p c) (tProj p c')) + (Hproj_cofix_noiota : forall p mfix idx args narg fn, + RedFlags.iota flags = false -> + unfold_cofix mfix idx = Some (narg, fn) -> + P (tProj p (mkApps (tCoFix mfix idx) args)) (tProj p (mkApps fn args))) + (Hproj_noiota : forall i pars narg args u arg, + RedFlags.iota flags = false -> + nth_error args (pars + narg) = Some arg -> + P (tProj (i, pars, narg) (mkApps (tConstruct i 0 u) args)) arg) + (Hproj_discr_noiota : forall p c c', + RedFlags.iota flags = false -> + red1 Σ Γ c c' -> + P (tProj p c) (tProj p c')) : + forall t t', whne flags Σ Γ t -> red1 Σ Γ t t' -> P t t'. Proof. - intros r wh. - induction wh in wh, t, r, t' |- *; cbn in *; try easy. + intros t t' wh r. + induction wh in t, t', wh, r |- *; cbn in *. - depelim r; [congruence|now solve_discr]. - - depelim r; cbn in *. - solve_discr. - - depelim r; try easy. - solve_discr. - - depelim r; try easy. - solve_discr. - - depelim r; try easy. - + depelim wh. - solve_discr. - + destruct args using List.rev_ind; [now solve_discr|]. + - depelim r; [|solve_discr]; eauto. + - depelim r; solve_discr; eauto. + - depelim r; solve_discr. + - depelim r; solve_discr; eauto. + - eauto. + - depelim r; solve_discr. + unfold declared_constant in isdecl. + rewrite H in isdecl. + inv isdecl. + congruence. + - eauto. + - depelim r; eauto. + + depelim wh; solve_discr. + now apply Hbeta_nobeta. + + destruct args as [|? ? _] using List.rev_ind; [now solve_discr|]. rewrite <- mkApps_nested in x. cbn in *. inv x. apply whne_mkApps_inv in wh; [|easy]. destruct wh as [|]. * depelim H. - solve_discr. - rewrite H in e. - inv e. - rewrite nth_error_nil in H0; congruence. + -- solve_discr. + rewrite H in e. + inv e. + rewrite nth_error_nil in H0; congruence. + -- change (tApp ?hd ?a) with (mkApps hd [a]). + rewrite mkApps_nested. + now eapply Hfix_nofix. * destruct H as (?&?&?&?&?&?&?&?&?). inv H. rewrite H0 in e. @@ -550,43 +685,13 @@ Proof. rewrite H0. destruct isConstruct_app eqn:ctor; [|easy]. now eapply whne_isConstruct_app in ctor. } - destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]. - + eapply OnOne2_nth_error in H0; eauto. - destruct H0 as (?&?&[->|]). - * eapply whne_fixapp; eauto. - * apply IHwh in r. - eapply whne_fixapp; eauto. - + unfold unfold_fix in *. - destruct (nth_error mfix idx) eqn:nth; [|easy]. - eapply OnOne2_nth_error in nth; eauto. - inv H. - destruct nth as (?&?&[->|(?&?)]). - * eapply whne_fixapp; eauto. - unfold unfold_fix. - rewrite e. - reflexivity. - * eapply whne_fixapp; eauto. - unfold unfold_fix. - rewrite e. - inv e0. - rewrite H3. - reflexivity. - + unfold unfold_fix in *. - destruct (nth_error mfix idx) eqn:nth; [|easy]. - eapply OnOne2_nth_error in nth; eauto. - inv H. - destruct nth as (?&?&[->|(?&?)]). - * eapply whne_fixapp; eauto. - unfold unfold_fix. - rewrite e. - reflexivity. - * eapply whne_fixapp; eauto. - unfold unfold_fix. - rewrite e. - inv e0. - rewrite H3. - reflexivity. - - depelim r; try easy. + destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]; eauto. + - depelim r; eauto. + destruct args using List.rev_ind; [|rewrite <- mkApps_nested in x; cbn in x; discriminate]. + cbn in *. + unfold is_constructor in e0. + rewrite nth_error_nil in e0; discriminate. + - depelim r; eauto. + apply whne_mkApps_inv in wh; [|easy]. destruct wh as [|(?&?&?&?&?&?&?)]; [|discriminate]. depelim H. @@ -596,7 +701,8 @@ Proof. destruct wh as [|(?&?&?&?&?&?&?)]; [|discriminate]. depelim H. solve_discr. - - depelim r; try easy. + - eauto. + - depelim r; eauto. + solve_discr. + apply whne_mkApps_inv in wh; [|easy]. destruct wh as [|(?&?&?&?&?&?&?)]; [|discriminate]. @@ -606,6 +712,54 @@ Proof. destruct wh as [|(?&?&?&?&?&?&?)]; [|discriminate]. depelim H. solve_discr. + - depelim r; eauto. + solve_discr. +Qed. +End whne_red1_ind. + +Lemma whne_pres1 Σ Γ t t' : + red1 Σ Γ t t' -> + whne RedFlags.default Σ Γ t -> + whne RedFlags.default Σ Γ t'. +Proof. + intros r wh. + apply (whne_red1_ind RedFlags.default Σ Γ (fun _ => whne RedFlags.default Σ Γ)) + with (t := t) (t' := t'); intros; try easy. + - eapply OnOne2_nth_error in H0; eauto. + destruct H0 as (?&?&[->|]). + + eapply whne_fixapp; eauto. + + apply H2 in r0. + eapply whne_fixapp; eauto. + - unfold unfold_fix in *. + destruct (nth_error mfix idx) eqn:nth; [|easy]. + eapply OnOne2_nth_error in nth; eauto. + inv H. + destruct nth as (?&?&[->|(?&?)]). + + eapply whne_fixapp; eauto. + unfold unfold_fix. + rewrite e. + reflexivity. + + eapply whne_fixapp; eauto. + unfold unfold_fix. + rewrite e. + inv e0. + rewrite H3. + reflexivity. + - unfold unfold_fix in *. + destruct (nth_error mfix idx) eqn:nth; [|easy]. + eapply OnOne2_nth_error in nth; eauto. + inv H. + destruct nth as (?&?&[->|(?&?)]). + + eapply whne_fixapp; eauto. + unfold unfold_fix. + rewrite e. + reflexivity. + + eapply whne_fixapp; eauto. + unfold unfold_fix. + rewrite e. + inv e0. + rewrite H4. + reflexivity. Qed. Lemma whne_pres Σ Γ t t' : @@ -664,9 +818,7 @@ Proof. rewrite nth in H. destruct s as [->|(?&?)]; [easy|]. now inv e. - - eapply red1_mkApps_tCoFix_inv in r; eauto. - destruct r as (?&?&->). - apply whnf_cofixapp. + - eapply red1_mkApps_tCoFix_inv in r as [[(?&->&?)|(?&->&?)]|(?&->&?)]; eauto. Qed. Lemma whnf_pres Σ Γ t t' : @@ -771,59 +923,21 @@ Lemma whne_red1_mkApps_tInd Σ Γ ind u args t : False. Proof. intros wh r. - remember (mkApps (tInd ind u) args) as t' eqn:t'eq. - revert t' r args t'eq. - induction wh; intros t' r args' ->; cbn in *; try easy. - - depelim r; [congruence|now solve_discr]. - - depelim r; cbn in *. - solve_discr. - - depelim r; solve_discr. - - depelim r; solve_discr. - unfold PCUICAst.declared_constant in isdecl. - rewrite H in isdecl. - inv isdecl. - congruence. - - depelim r. - + depelim wh; solve_discr. - + destruct args using rev_ind; [cbn in *; congruence|]. - rewrite <- mkApps_nested in x0. - cbn in *. - inv x0. - apply whne_mkApps_inv in wh; [|easy]. - destruct wh as [|(?&?&?&?&?&?&?&?&?)]. - * depelim H; solve_discr. - now rewrite nth_error_nil in H0. - * inv H. - rewrite e in H0. - inv H0. - unfold is_constructor in e0. - erewrite nth_error_app_left in e0; [|eassumption]. - now eapply whne_isConstruct_app. - + destruct args' as [|? ? _] using rev_ind; [cbn in *; congruence|]. - rewrite <- mkApps_nested in x. - cbn in *. - inv x. - eauto. - + destruct args' as [|? ? _] using rev_ind; [cbn in *; congruence|]. - rewrite <- mkApps_nested in x. - cbn in *. - inv x. - eapply whne_mkApps_tInd; eauto. - - apply red1_mkApps_tFix_inv in r. - 2: { rewrite H. - unfold is_constructor. - rewrite H0. - destruct isConstruct_app eqn:ctor; [|easy]. - now eapply whne_isConstruct_app in ctor. } - destruct r as [[(?&?&?)|(?&?&?)]|(?&?&?)]; solve_discr. - - depelim r; solve_discr. - apply whne_mkApps_inv in wh; [|easy]. - destruct wh as [|(?&?&?&?&?&?&?&?&?)]; [|congruence]. - depelim H; solve_discr. - - depelim r; solve_discr. - apply whne_mkApps_inv in wh; [|easy]. - destruct wh as [|(?&?&?&?&?&?&?&?&?)]; [|congruence]. - depelim H; solve_discr. + remember (mkApps (tInd ind u) args) eqn:eq. + revert args eq. + apply (whne_red1_ind + RedFlags.default + Σ Γ + (fun t t' => forall args, t' = mkApps (tInd ind u) args -> False)) + with (t := t) (t' := t0); intros; try easy; solve_discr. + - destruct args using List.rev_ind; [discriminate H1|]. + rewrite <- mkApps_nested in H1. + inv H1. + eauto. + - destruct args using List.rev_ind; [discriminate H0|]. + rewrite <- mkApps_nested in H0. + inv H0. + now apply whne_mkApps_tInd in H. Qed. Lemma whnf_red1_mkApps_tInd Σ Γ ind u args t : @@ -849,7 +963,7 @@ Proof. unfold is_constructor. now rewrite H. } destruct r as [[(?&?&?)|(?&?&?)]|(?&?&?)]; solve_discr. - - apply red1_mkApps_tCoFix_inv in r as (?&?&?); solve_discr. + - eapply red1_mkApps_tCoFix_inv in r as [[(?&?&?)|(?&?&?)]|(?&?&?)]; solve_discr. Qed. Lemma whnf_red_mkApps_tInd Σ Γ ind u args t : @@ -865,6 +979,144 @@ Proof. - apply whnf_red1_mkApps_tInd in X as (?&->); eauto. eapply whnf_pres; eauto. Qed. + +Lemma whne_red1_from_mkApps f Σ Γ hd args t : + RedFlags.beta f = true -> + RedFlags.fix_ f = true -> + isApp hd = false -> + whne f Σ Γ (mkApps hd args) -> + red1 Σ Γ (mkApps hd args) t -> + ∥∑ hd' args', + t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. +Proof. + intros redbeta redfix noapp wh r. + remember (mkApps hd args) eqn:eq. + revert args eq. + apply (whne_red1_ind + f Σ Γ + (fun t t' => (forall args, t = mkApps hd args -> _ : Prop))) + with (t1 := t0) (t' := t); intros; cbn in *; + try solve [congruence]; + try solve [ + match goal with + | [H: ?a = mkApps ?h ?args |- _] => + apply (mkApps_notApp_inj a h [] args) in H as (<-&<-); auto; + constructor; eexists _, []; split; [reflexivity|]; + eauto with pcuic + end]. + - destruct args as [|? ? _] using List.rev_ind; [destruct hd; cbn in *; congruence|]. + rewrite <- mkApps_nested in H1. + inv H1. + destruct (H0 _ eq_refl) as [(?&?&->&?&?)]. + constructor; exists x0, (x1 ++ [x]). + rewrite <- mkApps_nested. + split; [easy|]. + split; [now eauto|]. + apply All2_app; eauto. + - destruct args as [|? ? _] using List.rev_ind; [destruct hd; cbn in *; congruence|]. + rewrite <- mkApps_nested in H0. + inv H0. + constructor; exists hd, (args ++ [N2]). + rewrite <- mkApps_nested. + split; [easy|]. + split; [easy|]. + apply All2_app; eauto with pcuic. + apply All2_same; eauto. + - apply mkApps_notApp_inj in H3 as (<-&<-); auto. + constructor; eexists _, _; split; [reflexivity|]. + split; [eauto|]. + eapply OnOne2_All2; eauto. + - apply mkApps_notApp_inj in H2 as (<-&<-); auto. + constructor; eexists _, _; split; [reflexivity|]. + split; [|now apply All2_same]. + eauto with pcuic. + - apply mkApps_notApp_inj in H2 as (<-&<-); auto. + constructor; eexists _, _; split; [reflexivity|]. + split; [|now apply All2_same]. + eauto with pcuic. +Qed. + +Lemma whnf_red1_from_mkApps f Σ Γ hd args t : + RedFlags.beta f = true -> + RedFlags.fix_ f = true -> + isApp hd = false -> + whnf f Σ Γ (mkApps hd args) -> + red1 Σ Γ (mkApps hd args) t -> + ∥∑ hd' args', + t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. +Proof. + intros nobeta nofix noapp wh r. + depelim wh; + try solve [ + match goal with + | [H: ?a = mkApps ?h ?args |- _] => + apply (mkApps_notApp_inj a h [] args) in H as (<-&<-); auto; + constructor; eexists _, []; split; [reflexivity|]; + eauto with pcuic + end]. + - now eapply whne_red1_from_mkApps. + - apply mkApps_notApp_inj in x as (<-&<-); auto. + apply red1_mkApps_tConstruct_inv in r as (?&->&?). + constructor; eexists _, _; split; [reflexivity|]. + split; [easy|]. + eapply OnOne2_All2; eauto. + - apply mkApps_notApp_inj in x as (<-&<-); auto. + apply red1_mkApps_tInd_inv in r as (?&->&?). + constructor; eexists _, _; split; [reflexivity|]. + split; [easy|]. + eapply OnOne2_All2; eauto. + - apply mkApps_notApp_inj in x as (<-&<-); auto. + apply red1_mkApps_tFix_inv in r as [[(?&->&?)|(?&->&?)]|(?&->&?)]; auto. + 1-3: constructor; eexists _, _; split; [reflexivity|]; + eauto using OnOne2_All2, All2_same with pcuic. + destruct unfold_fix as [(?&?)|]; [|easy]. + unfold is_constructor. + now rewrite H. + - apply mkApps_notApp_inj in x as (<-&<-); auto. + apply red1_mkApps_tCoFix_inv in r as [[(?&->&?)|(?&->&?)]|(?&->&?)]; auto. + 1-3: constructor; eexists _, _; split; [reflexivity|]; + eauto using OnOne2_All2, All2_same with pcuic. +Qed. + +Lemma whnf_red_from_mkApps f Σ Γ hd args t : + RedFlags.beta f = true -> + RedFlags.fix_ f = true -> + isApp hd = false -> + whnf f Σ Γ (mkApps hd args) -> + red Σ Γ (mkApps hd args) t -> + ∥∑ hd' args', + t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. +Proof. + intros nobeta nofix notapp wh r. + remember (mkApps hd args) as from eqn:fromeq. + revert args fromeq. + induction r using red_rect_1n; intros args ->. + - eapply whnf_red1_from_mkApps in r as [(?&?&->&?&?)]; eauto. + now constructor; eexists _, _; split; [reflexivity|]. + - constructor; eexists _, _; split; [reflexivity|]. + split; [easy|]. + now apply All2_same. + - fold (red Σ Γ (mkApps hd args) y) in r1. + apply whnf_red1_from_mkApps in r1. + +Lemma whnf_red_from_mkApps Σ Γ t + whnf RedFlags.default Σ Γ t -> + red Σ Γ t (mkApps (tInd ind u) args) -> + exists args', t = mkApps (tInd ind u) args'. +Proof. + remember (mkApps (tInd ind u) args) as hd eqn:hdeq. + intros wh r. + revert args hdeq. + induction r using red_rect'; intros args ->. + - easy. + - apply whnf_red1_mkApps_tInd in X as (?&->); eauto. + eapply whnf_pres; eauto. +Qed. + +Lemma whne_red1_from_mkApps_tCase f Σ Γ p motive discr brs : + all_except_delta f -> + whnf f Σ Γ (tCase p motive discr brs) -> + red1 Σ Γ (mkApps (tCase (* Definition head_arg_is_constructor mfix idx args := diff --git a/pcuic/theories/PCUICReduction.v b/pcuic/theories/PCUICReduction.v index d51cf4b07..615909708 100644 --- a/pcuic/theories/PCUICReduction.v +++ b/pcuic/theories/PCUICReduction.v @@ -341,6 +341,20 @@ Proof. eapply X0; tea. now apply clos_rt_rtn1_iff. Defined. +Definition red_rect_n1 := red_rect'. +Definition red_rect_1n Σ Γ (P : term -> term -> Type) : + (forall x, P x x) -> + (forall x y z, red1 Σ Γ x y -> red Σ Γ y z -> P y z -> P x z) -> + forall x y, red Σ Γ x y -> P x y. +Proof. + intros Hrefl Hstep x y r. + apply clos_rt_rt1n_iff in r. + induction r; eauto. + eapply Hstep; eauto. + now apply clos_rt_rt1n_iff. +Defined. + + (** Simple lemmas about reduction *) diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index ee6321878..86a4671b1 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -1386,7 +1386,24 @@ Section Lemmata. - simpl in *. destruct h. constructor. eapply cumul_conv_ctx. all: eauto. Qed. - + + Lemma conv_cum_zipp : + forall Γ leq u v π1 π2, + conv_cum leq Σ Γ u v -> + conv_context Σ (Γ,,, stack_context π1) (Γ,,, stack_context π2) -> + conv_cum leq Σ Γ (zipp u π1) (zipp v π2). + Proof. + intros Γ leq u v π1 π2 c_uv c_stacks. + unfold zipp. + destruct (decompose_stack π1) eqn:decomp1, (decompose_stack π2) eqn:decomp2. + apply decompose_stack_eq in decomp1 as ->. + apply decompose_stack_eq in decomp2 as ->. + rewrite !stack_context_appstack in c_stacks. + destruct hΣ. + destruct leq. + - destruct c_uv as [c_uv]. + constructor. + apply mkApps_conv_args; auto. End Lemmata. Lemma Case_Construct_ind_eq {cf:checker_flags} Σ (hΣ : ∥ wf Σ.1 ∥) : diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 5cb2582b7..fd98e55a9 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -764,13 +764,44 @@ Section Conversion. conv leq Σ (Γ ,,, stack_context π) t t' -> { b : bool | if b then conv_term leq Γ t π t' π' else True } end. *) + + Definition isred_full Γ t π := + isApp t = false /\ + whnf nodelta_flags Σ (Γ,,, stack_context π) (zipp t π). + + Lemma isred_full_nobeta Γ t π : + isred_full Γ t π -> + isLambda t -> + isStackApp π = false. + Proof. + intros (?&isr) islam. + destruct t; cbn in *; try easy. + unfold zipp in isr. + destruct π; cbn in *; try easy. + destruct (decompose_stack π) in isr. + depelim isr; rewrite mkApps_tApp in *; try solve [solve_discr]. + apply whne_mkApps_inv in H0; [|easy]. + destruct H0 as [|(?&?&?&?&?&?&?&?)]; [|discriminate]. + depelim H0; solve_discr; discriminate. + Qed. Definition Ret s Γ t π t' π' := forall (leq : conv_pb), conv_stack_ctx Γ π π' -> - (match s with Fallback | Term => isred (t, π) | _ => True end) -> - (match s with Fallback | Term => isred (t', π') | _ => True end) -> - (match s with Args => conv_cum leq Σ (Γ ,,, stack_context π) t t' | _ => True end) -> + match s with + | Fallback + | Term => isred_full Γ t π + | _ => True + end -> + match s with + | Fallback + | Term => isred_full Γ t' π' + | _ => True + end -> + match s with + | Args => conv_cum leq Σ (Γ ,,, stack_context π) t t' + | _ => True + end -> ConversionResult (conv_term leq Γ t π t' π'). Definition Aux s Γ t1 π1 t2 π2 h2 := @@ -1015,38 +1046,62 @@ Section Conversion. Next Obligation. match type of eq1 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as r1 + pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as r1; + pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w1 end. - rewrite <- eq1 in r1. destruct r1 as [ha hl]. split. - - assumption. - - cbn in hl. cbn. intro h. - specialize (hl h). - destruct π1'. - all: try reflexivity. - + cbn. destruct ρ1. - all: try reflexivity. - exfalso. - apply (decompose_stack_not_app _ _ _ _ (eq_sym e1)). - + discriminate hl. + - rewrite <- eq1 in r1. destruct r1 as [ha hl]. + assumption. + - rewrite <- eq1 in w1. + apply (f_equal (snd ∘ decompose_stack ∘ snd)) in eq1. + rewrite reduce_stack_decompose, decompose_stack_appstack in eq1. + rewrite stack_context_stack_cat. + cbn in *. + rewrite <- stack_context_decompose in w1. + rewrite <- (stack_context_decompose π1') in w1 |- *. + rewrite eq1 in w1 |- *. + rewrite <- e1 in w1. + cbn in *. + unfold zipp. + assert (decompose_stack π1' = ((decompose_stack π1').1, (decompose_stack π1').2)) + by (now destruct (decompose_stack π1')). + rewrite eq1 in H. + apply decompose_stack_eq in H as ->. + rewrite stack_cat_appstack, decompose_stack_appstack. + erewrite (decompose_stack_twice _ _ ρ1); eauto. + cbn; rewrite app_nil_r. + unfold zipp in w1. + now destruct (decompose_stack π1'). Qed. Next Obligation. match type of eq2 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as r2 + pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as r2; + pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w2 end. - rewrite <- eq2 in r2. destruct r2 as [ha hl]. split. - - assumption. - - cbn in hl. cbn. intro h. - specialize (hl h). - destruct π2'. - all: try reflexivity. - + cbn. destruct ρ2. - all: try reflexivity. - exfalso. - apply (decompose_stack_not_app _ _ _ _ (eq_sym e2)). - + discriminate hl. + - rewrite <- eq2 in r2. destruct r2 as [ha hl]. + assumption. + - rewrite <- eq2 in w2. + apply (f_equal (snd ∘ decompose_stack ∘ snd)) in eq2. + rewrite reduce_stack_decompose, decompose_stack_appstack in eq2. + rewrite stack_context_stack_cat. + cbn in *. + rewrite <- stack_context_decompose in w2. + rewrite <- (stack_context_decompose π2') in w2 |- *. + rewrite eq2 in w2 |- *. + rewrite <- e2 in w2. + cbn in *. + unfold zipp. + assert (decompose_stack π2' = ((decompose_stack π2').1, (decompose_stack π2').2)) + by (now destruct (decompose_stack π2')). + rewrite eq2 in H. + apply decompose_stack_eq in H as ->. + rewrite stack_cat_appstack, decompose_stack_appstack. + erewrite (decompose_stack_twice _ _ ρ2); eauto. + cbn; rewrite app_nil_r. + unfold zipp in w2. + now destruct (decompose_stack π2'). Qed. Next Obligation. destruct hΣ as [wΣ]. @@ -1530,17 +1585,77 @@ Section Conversion. eapply Forall2_impl. 1: eassumption. intros. eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). all: auto. + + eapply wf_ext_global_uctx_invariants. + eapply hΣ'. + + eapply global_ext_uctx_consistent. + eapply hΣ'. + Qed. + + Lemma eq_universe_instance_spec udecl u v : + consistent_instance_ext Σ udecl u -> + consistent_instance_ext Σ udecl v -> + R_universe_instance (eq_universe (global_ext_constraints Σ)) u v -> + eqb_universe_instance u v. + Proof. + intros consu consv r. + apply Forall2_forallb2. + unfold consistent_instance_ext, consistent_instance, R_universe_instance in *. + destruct udecl; [destruct u,v; cbn in *; try congruence; constructor|]. + destruct consu as (_&all_mem_u&_&_). + destruct consv as (_&all_mem_v&_&_). + apply forallb_Forall in all_mem_u. + apply forallb_Forall in all_mem_v. + apply Forall2_map. + apply Forall2_map_inv in r. + eapply Forall_Forall2_and in r; [|exact all_mem_u]. + eapply Forall_Forall2_and' in r; [|exact all_mem_v]. + clear all_mem_u all_mem_v. + cbn beta in r. + eapply Forall2_impl; [eassumption|]; cbn beta in *. + intros x y ((x_mem&x_eq_y)&y_mem). + eapply eq_universe_spec'; eauto. - eapply wf_ext_global_uctx_invariants. eapply hΣ'. - eapply global_ext_uctx_consistent. eapply hΣ'. + - intros ? ->%UnivExprSet.singleton_spec. + destruct (UnivExpr.get_noprop (UnivExpr.make x)) eqn:np; [|easy]. + cbn. + replace x with (NoPropLevel.to_level t) in *. + 2: { destruct x; cbn in *; try easy. + all: now inv np. } + unfold global_ext_levels in *. + now apply LevelSet.mem_spec in x_mem. + - intros ? ->%UnivExprSet.singleton_spec. + destruct (UnivExpr.get_noprop (UnivExpr.make y)) eqn:np; [|easy]. + cbn. + replace y with (NoPropLevel.to_level t) in *. + 2: { destruct y; cbn in *; try easy. + all: now inv np. } + unfold global_ext_levels in *. + now apply LevelSet.mem_spec in y_mem. Qed. - Definition eq_termp leq v v' := - match leq with - | Conv => eq_term Σ v v' - | Cumul => leq_term Σ Σ v v' + Definition conv_pb_rel (pb : conv_pb) := + match pb with + | Conv => eq_universe + | Cumul => leq_universe end. + + Definition eq_termp_napp leq napp := + eq_term_upto_univ_napp Σ (eq_universe Σ) (conv_pb_rel leq Σ) napp. + + Notation eq_termp leq := (eq_termp_napp leq 0). + + Lemma eq_termp_mkApps_inv leq v args v' args' : + isApp v = false -> + isApp v' = false -> + eq_termp leq (mkApps v args) (mkApps v' args') -> + eq_termp_napp leq #|args| v v' × All2 (fun x y => eq_term Σ x y) args args'. + Proof. + intros noapp1 noapp2 eq. + now apply eq_term_upto_univ_mkApps_inv in eq as (?&?). + Qed. Lemma conv_cum_alt leq Γ t t' : conv_cum leq Σ Γ t t' -> @@ -1555,7 +1670,37 @@ Section Conversion. now constructor. Qed. - (* Also, is there an invariant of injectivity on the assignment between kernames and inductives *) + Lemma zipp_as_mkApps t π : + zipp t π = mkApps t (decompose_stack π).1. + Proof. + unfold zipp. + now destruct decompose_stack. + Qed. + + Lemma wellformed_nonarity Γ t : + destArity [] t = None -> + wellformed Σ Γ t -> + welltyped Σ Γ t. + Proof. + now intros dar [|[(?&?&?&?)]]. + Qed. + + Lemma wellformed_zipc_tConst_inv Γ c u π : + wellformed Σ Γ (zipc (tConst c u) π) -> + exists cst, + declared_constant Σ c cst + × consistent_instance_ext Σ (cst_universes cst) u. + Proof. + intros h. + destruct hΣ. + zip fold in h. + apply wellformed_context in h; auto. + apply wellformed_nonarity in h; auto. + destruct h as (?&typ). + apply inversion_Const in typ as (?&?&?&wfu&_); auto. + now unfold declared_constant in d. + Qed. + Equations(noeqns) unfold_constants (Γ : context) (leq : conv_pb) (c : kername) (u : Instance.t) (π1 : stack) (h1 : wtp Γ (tConst c u) π1) @@ -1572,13 +1717,25 @@ Section Conversion. isconv_red leq (tConst c u) π1 (subst_instance_constr u' b) π2 aux ; (* Inductive or not found *) | _ with inspect (lookup_env Σ c) := { - | @exist (Some (ConstantDecl {| cst_body := Some b |})) eq1 := + | @exist (Some (ConstantDecl {| cst_body := Some b |})) eq2 := isconv_red leq (subst_instance_constr u b) π1 (tConst c' u') π2 aux ; (* Both Inductive or not found *) | _ := Error (NotFoundConstants c c') _ } }. + Solve All Obligations with + Tactics.program_simplify; + CoreTactics.equations_simpl; + try solve + [match goal with + | [H: wellformed ?Σ ?Γ ?t |- _] => + let id := fresh in + apply wellformed_zipc_tConst_inv in H as id; + destruct id as (?&?&?); + unfold declared_constant in *; + congruence + end]. Next Obligation. eapply red_wellformed ; auto. - exact h2. @@ -1631,90 +1788,22 @@ Section Conversion. Next Obligation. (* Both c and c' are axioms. Either they are different constants or they are not convertible because the universes are different. *) - destruct ne. - - apply conv_cum_alt in H as [(?&?&(r1&r2)&e)]. - unfold zipp in r1, r2. - eapply red_zipp_inv in r1. - Qed. - Next Obligation. - (* inductive vs non-defined constant *) - (* invert H and discriminate between inductive and constant *) - todo "Completeness"%string. - Qed. - Next Obligation. - (* Not found vs non-defined constant *) - (* invert H and discriminate between inductive and constant *) - todo "Completeness"%string. - Qed. - Next Obligation. - eapply red_wellformed ; auto ; [ exact h1 | ]. - constructor. eapply red_zipc. eapply red_const. eassumption. - Qed. - Next Obligation. - eapply R_cored. simpl. eapply cored_zipc. - eapply cored_const. eassumption. - Qed. - Next Obligation. - destruct hΣ as [wΣ]. - etransitivity ; try eassumption. - eapply red_conv_cum_l ; try assumption. - eapply red_zipp. eapply red_const. eassumption. - Qed. - Next Obligation. - (* Contrapositive of previous obligation *) - apply h; clear h. - destruct hΣ as [wΣ]. - etransitivity ; try eassumption. - eapply red_conv_cum_r ; try assumption. - eapply red_zipp. eapply red_const. eassumption. - Qed. - Next Obligation. - (* Non-defined constant vs inductive *) - todo "Completeness"%string. - Qed. - Next Obligation. - (* Two inductives: missing hypothesis to ditch that case ?*) - todo "Completeness"%string. - Qed. - Next Obligation. - (* Not found vs inductive *) - todo "Completeness"%string. - Qed. - Next Obligation. - eapply red_wellformed ; auto ; [ exact h1 | ]. - constructor. eapply red_zipc. eapply red_const. eassumption. - Qed. - Next Obligation. - eapply R_cored. simpl. eapply cored_zipc. - eapply cored_const. eassumption. - Qed. - Next Obligation. - destruct hΣ as [wΣ]. - etransitivity ; try eassumption. - eapply red_conv_cum_l ; try assumption. - eapply red_zipp. eapply red_const. eassumption. - Qed. - Next Obligation. - (* Contrapositive of previous obligation *) - apply h; clear h. - destruct hΣ as [wΣ]. - etransitivity ; try eassumption. - eapply red_conv_cum_r ; try assumption. - eapply red_zipp. eapply red_const. eassumption. - Qed. - Next Obligation. - (* Constant w/o body vs not found *) - todo "Completeness"%string. - Qed. - Next Obligation. - (* Inductive vs not found *) - todo "Completeness"%string. - Qed. - Next Obligation. - (* Not found vs Not found *) - todo "Completeness"%string. + apply conv_cum_alt in H as [(?&?&(r1&r2)&e)]. + rewrite zipp_as_mkApps in r1, r2. + destruct hΣ. + symmetry in e0, unknown3. + eapply PCUICConfluence.red_mkApps_tConst_axiom in r1 as (?&->&?); eauto. + eapply PCUICConfluence.red_mkApps_tConst_axiom in r2 as (?&->&?); eauto. + apply eq_termp_mkApps_inv in e as (?&?); [|easy|easy]. + depelim e. + destruct ne as [|(_&ne)]; [congruence|]. + + clear aux. + apply wellformed_zipc_tConst_inv in h1 as (cst1&decl1&cons1). + apply wellformed_zipc_tConst_inv in h2 as (cst2&decl2&cons2). + eapply PCUICWeakeningEnv.declared_constant_inj in decl1; eauto; subst. + eapply eq_universe_instance_spec in r; eauto. Qed. - (* TODO (RE)MOVE *) Lemma destArity_eq_term_upto_univ : @@ -1743,6 +1832,24 @@ Section Conversion. + constructor. all: assumption. Qed. + Lemma wellformed_zipc_tCase_brs_length Γ p motive discr brs π : + wellformed Σ Γ (zipc (tCase p motive discr brs) π) -> + exists mib oib, declared_inductive Σ mib p.1 oib /\ #|brs| = #|ind_ctors oib|. + Proof. + intros wf. + zip fold in wf. + apply wellformed_context in wf; [|assumption]. + apply wellformed_nonarity in wf as (?&typ); [|easy]. + destruct hΣ. + apply inversion_Case in typ as (?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?); auto. + exists x1, x2. + split; [easy|]. + apply All2_length in a as ->. + apply map_option_out_length in e2. + rewrite PCUICElimination.length_of_btys in e2. + congruence. + Qed. + Equations isconv_branches (Γ : context) (ind : inductive) (par : nat) (p c : term) (brs1 brs2 : list (nat × term)) @@ -1793,20 +1900,24 @@ Section Conversion. Next Obligation. destruct h1 as [h1]. apply All2_length in h1 as e1. - zip fold in h. apply wellformed_context in h. 2: assumption. clear aux. - zip fold in h'. apply wellformed_context in h'. 2: assumption. - simpl in *. - todo "Number of branches"%string. - (* We should be able to conclude that something is wrong - from typing of two cases with different number of branches. - It might be best to wait for the new typing rule of case - to get that the number of branches is only dependent on the inductive. - *) + apply wellformed_zipc_tCase_brs_length in h as (?&?&?&?). + apply wellformed_zipc_tCase_brs_length in h' as (?&?&?&?). + pose proof (PCUICInductiveInversion.declared_inductive_unique_sig H H1) as u; noconf u. + rewrite app_length in *. + cbn in *. + lia. Qed. Next Obligation. - (* Symmetric case of the previous one, so should do a lemma *) - todo "Number of branches"%string. + destruct h1 as [h1]. + apply All2_length in h1 as e1. + clear aux. + apply wellformed_zipc_tCase_brs_length in h as (?&?&?&?). + apply wellformed_zipc_tCase_brs_length in h' as (?&?&?&?). + pose proof (PCUICInductiveInversion.declared_inductive_unique_sig H H1) as u; noconf u. + rewrite app_length in *. + cbn in *. + lia. Qed. Next Obligation. eapply R_positionR. all: simpl. @@ -2348,13 +2459,37 @@ Section Conversion. apply h; clear h. destruct H as [H]; constructor; apply (All2_impl H). intuition. - Qed. + Qed. + + Lemma invert_type_mkApps_tProd Γ na A b args T : + Σ;;; Γ |- mkApps (tProd na A b) args : T -> args = []. + Proof. + intros typ. + destruct hΣ. + apply PCUICValidity.inversion_mkApps in typ as (?&typ_prod&typ_args); [|easy]. + apply inversion_Prod in typ_prod as (?&?&?&?&?); [|easy]. + eapply PCUICSpine.typing_spine_strengthen in typ_args; eauto. + clear -typ_args. + depelim typ_args. + - easy. + - now apply cumul_Sort_Prod_inv in c. + Qed. Lemma wellformed_zipc_tProd_appstack_nil {Γ na A B l ρ} : wellformed Σ Γ (zipc (tProd na A B) (appstack l ρ)) -> l = []. Proof. - (* Proof idea: a sort cannot be applied to anything *) - todo "Completeness"%string. + intros wh. + rewrite zipc_appstack in wh. + zip fold in wh. + apply wellformed_context in wh; [|easy]. + cbn in wh. + destruct l as [|? ? _] using List.rev_ind; [easy|]. + rewrite <- mkApps_nested in wh. + cbn in wh. + apply wellformed_nonarity in wh as (?&typ); auto. + change (tApp ?h ?a) with (mkApps h [a]) in typ. + rewrite mkApps_nested in typ. + now apply invert_type_mkApps_tProd in typ. Qed. (* @@ -2407,24 +2542,17 @@ Section Conversion. todo "Completeness"%string. Qed. *) - Lemma conv_red_full_l : - forall Σ Γ (t u v : term), - red Σ.1 Γ v t -> Σ;;; Γ |- v = u -> Σ;;; Γ |- t = u. - Proof. - admit. - Admitted. - Opaque reduce_stack. Equations(noeqns) _isconv_prog (Γ : context) (leq : conv_pb) (t1 : term) (π1 : stack) (h1 : wtp Γ t1 π1) (t2 : term) (π2 : stack) (h2 : wtp Γ t2 π2) (hx : conv_stack_ctx Γ π1 π2) - (ir1 : isred (t1, π1)) (ir2 : isred (t2, π2)) + (ir1 : isred_full Γ t1 π1) (ir2 : isred_full Γ t2 π2) (aux : Aux Term Γ t1 π1 t2 π2 h2) : ConversionResult (conv_term leq Γ t1 π1 t2 π2) := _isconv_prog Γ leq t1 π1 h1 t2 π2 h2 hx ir1 ir2 aux with prog_viewc t1 t2 := { - | prog_view_App _ _ _ _ := !; + | prog_view_App _ _ _ _ := False_rect _ _; | prog_view_Const c u c' u' with eq_dec c c' := { | left eq1 with inspect (eqb_universe_instance u u') := { @@ -2440,10 +2568,10 @@ Section Conversion. } } ; (* If universes are different, we unfold one of the constants *) - | @exist false _ := unfold_constants Γ leq c u π1 h1 c' u' π2 h2 hx aux + | @exist false uneq_u := unfold_constants Γ leq c u π1 h1 c' u' π2 h2 _ hx aux } ; (* If the two constants are different, we unfold one of them *) - | right _ := unfold_constants Γ leq c u π1 h1 c' u' π2 h2 hx aux + | right uneq_c := unfold_constants Γ leq c u π1 h1 c' u' π2 h2 _ hx aux } ; | prog_view_Lambda na A1 t1 na' A2 t2 @@ -2633,6 +2761,10 @@ Section Conversion. eapply red_zipp. eapply red_const. eassumption. Qed. + Next Obligation. + right; split; [easy|]. + now rewrite <- uneq_u. + Qed. (* tLambda *) Next Obligation. @@ -2666,13 +2798,11 @@ Section Conversion. rewrite stack_context_appstack in h0. rewrite stack_context_appstack in h. - destruct ir1 as [_ hl1]. cbn in hl1. - specialize (hl1 eq_refl). - destruct l1 ; try discriminate hl1. clear hl1. + apply isred_full_nobeta in ir1; [|easy]. + destruct l1 ; try discriminate ir1. clear ir1. - destruct ir2 as [_ hl2]. cbn in hl2. - specialize (hl2 eq_refl). - destruct l2 ; try discriminate hl2. clear hl2. + apply isred_full_nobeta in ir2; [|easy]. + destruct l2 ; try discriminate ir2. clear ir2. simpl in *. destruct hΣ. now eapply conv_cum_Lambda. @@ -2693,13 +2823,11 @@ Section Conversion. rewrite stack_context_appstack in h0. rewrite stack_context_appstack in H. - destruct ir1 as [_ hl1]. cbn in hl1. - specialize (hl1 eq_refl). - destruct l1 ; try discriminate hl1. clear hl1. + apply isred_full_nobeta in ir1; [|easy]. + destruct l1 ; try discriminate ir1. clear ir1. - destruct ir2 as [_ hl2]. cbn in hl2. - specialize (hl2 eq_refl). - destruct l2 ; try discriminate hl2. clear hl2. + apply isred_full_nobeta in ir2; [|easy]. + destruct l2 ; try discriminate ir2. clear ir2. rewrite e1, e2 in H. simpl in *. @@ -2723,13 +2851,11 @@ Section Conversion. pose proof (decompose_stack_eq _ _ _ e2). subst. rewrite stack_context_appstack in H. - destruct ir1 as [_ hl1]. cbn in hl1. - specialize (hl1 eq_refl). - destruct l1 ; try discriminate hl1. clear hl1. + apply isred_full_nobeta in ir1; [|easy]. + destruct l1 ; try discriminate ir1. clear ir1. - destruct ir2 as [_ hl2]. cbn in hl2. - specialize (hl2 eq_refl). - destruct l2 ; try discriminate hl2. clear hl2. + apply isred_full_nobeta in ir2; [|easy]. + destruct l2 ; try discriminate ir2. clear ir2. rewrite e1, e2 in H. simpl in *. @@ -2903,6 +3029,15 @@ Section Conversion. Qed. Next Obligation. apply h; clear h. + destruct ir1 as (_&wh1), ir2 as (_&wh2). + cbn in *. + apply conv_cum_alt in H. + rewrite !zipp_as_mkApps in *. + apply whnf_mkApps_inv in wh1; [|easy]. + apply whnf_mkApps_inv in wh2; [|easy]. + depelim wh1; solve_discr. + depelim wh2; solve_discr. + unfold zipp in wh1, wh2. (* Proof idea: 1) prove that the 2 tCase are in conv_cum relation because they are whne (needs to add thys hypothesis) @@ -3072,19 +3207,25 @@ Section Conversion. by apply reduce_term_sound. destruct hx as [hx]. destruct hΣ as [ wfΣ]. - (* Proof idea: wlog leq = Conv, *) destruct leq. - set (t0 := zipp _ _). destruct H as [Hconv]. constructor. - eapply conv_red_full_l. 1: apply red_zipp, red_case_c, hc. + unfold Aux in aux. + eapply conv_red_l_inv. 1: auto. 2: apply red_zipp, red_case_c, hc. symmetry. eapply conv_conv_ctx. 1: assumption. 2: eapply conv_context_sym; eassumption. - eapply conv_red_full_l. 1: apply red_zipp, red_case_c, hc'. + eapply conv_red_l_inv. 1: auto. 2: apply red_zipp, red_case_c, hc'. eapply conv_conv_ctx; try eassumption. symmetry; assumption. - - todo "Completeness"%string. (* see previous comment *) + - destruct H as [Hconv]. + constructor. + eapply cumul_red_l_inv. 1: auto. 2: apply red_zipp, red_case_c, hc. + eapply cumul_conv_ctx. 1: assumption. + 2: eapply conv_context_sym; eassumption. + eapply cumul_red_r_inv; [auto|eauto|]. 2: apply red_zipp, red_case_c, hc'. + eapply cumul_conv_ctx; eauto. Qed. (* tProj *) @@ -3213,20 +3354,32 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; - pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as ir + pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as ir; + pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w2 end. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. - rewrite <- eq3 in ir. destruct ir as [? hl]. split. - - assumption. - - simpl. intro h. specialize (hl h). cbn in hl. - case_eq (decompose_stack ρ). intros l π e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. - apply isStackApp_false_appstack in hl. subst. cbn. - eapply decompose_stack_noStackApp. symmetry. eassumption. + - rewrite <- eq3 in ir. destruct ir as [? hl]. + assumption. + - rewrite <- eq3 in w2, d2. + rewrite decompose_stack_appstack in d2. + rewrite stack_context_stack_cat. + cbn in *. + rewrite <- stack_context_decompose in w2. + rewrite <- (stack_context_decompose ρ) in w2 |- *. + rewrite d2 in w2 |- *. + erewrite decompose_stack_twice in w2 by eauto. + cbn in *. + assert (decompose_stack ρ = ((decompose_stack ρ).1, (decompose_stack ρ).2)) + by (now destruct (decompose_stack ρ)). + rewrite d2 in H. + apply decompose_stack_eq in H as ->. + unfold zipp. + rewrite stack_cat_appstack, decompose_stack_appstack. + cbn. + erewrite (decompose_stack_twice θ l' θ') by auto. + cbn; rewrite app_nil_r. + unfold zipp in w2. + now destruct (decompose_stack ρ). Qed. Next Obligation. destruct hΣ as [wΣ]. @@ -3369,7 +3522,8 @@ Section Conversion. rewrite <- eq3 in ir. destruct ir as [? hl]. split. - assumption. - - simpl. intro h. specialize (hl h). cbn in hl. + - rewrite stack_context_stack_cat. + simpl. intro h. specialize (hl h). cbn in hl. case_eq (decompose_stack ρ). intros l π e. rewrite e in d2. cbn in d2. subst. pose proof (decompose_stack_eq _ _ _ e). subst. diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index 4c2cc4bbf..96175c18b 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -1658,6 +1658,9 @@ Section Reduce. rewrite eq in haux. cbn in haux. assumption. Qed. + + Lemma stack_context_decompose' π π' : + stack_context (decompose_stack π).2 = stack_context p. Theorem reduce_term_complete Γ t h : whnf flags Σ Γ (reduce_term Γ t h). diff --git a/template-coq/theories/common/uGraph.v b/template-coq/theories/common/uGraph.v index ca58989dd..c85af61f9 100644 --- a/template-coq/theories/common/uGraph.v +++ b/template-coq/theories/common/uGraph.v @@ -1721,7 +1721,6 @@ Section CheckLeq2. exact I. Qed. - (* todo complete *) Lemma check_eqb_universe_spec' u1 u2 : check_eqb_universe G u1 u2 -> eq_universe uctx.2 u1 u2. Proof. @@ -1740,6 +1739,23 @@ Section CheckLeq2. split; eapply gc_leq_universe_n_iff; (destruct (gc_of_constraints uctx.2); [cbn in *|contradiction HG]); tas. Qed. + + Lemma eq_universe_spec' u1 u2 : + levels_declared u1 -> + levels_declared u2 -> + eq_universe uctx.2 u1 u2 -> + check_eqb_universe G u1 u2. + Proof. + intros decl1 decl2. + apply levels_gc_declared_declared in decl1. + apply levels_gc_declared_declared in decl2. + rewrite gc_eq_universe_iff. + unfold is_graph_of_uctx, gc_of_uctx in HG. + destruct gc_of_constraints; [cbn in *|contradiction HG]. + intros eq. + apply <- check_eqb_universe_spec; eauto. + exact eq. + Qed. (* todo complete *) Lemma check_constraints_spec ctrs From 444918b194595eedd94c1bf673a4c113607cd250 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Wed, 21 Oct 2020 17:42:10 +0200 Subject: [PATCH 07/26] Finish tCase case --- pcuic/_CoqProject.in | 1 + pcuic/theories/PCUICConversionInversion.v | 216 +++++++++++++++ pcuic/theories/PCUICNormal.v | 270 +++++++++++++----- pcuic/theories/PCUICSafeLemmata.v | 17 -- safechecker/theories/PCUICSafeConversion.v | 303 ++++++++++++++++----- safechecker/theories/PCUICSafeReduce.v | 3 - 6 files changed, 659 insertions(+), 151 deletions(-) create mode 100644 pcuic/theories/PCUICConversionInversion.v diff --git a/pcuic/_CoqProject.in b/pcuic/_CoqProject.in index 17e9bdb8b..2fbb43b7b 100644 --- a/pcuic/_CoqProject.in +++ b/pcuic/_CoqProject.in @@ -27,6 +27,7 @@ theories/PCUICParallelReductionConfluence.v theories/PCUICConfluence.v theories/PCUICContextConversion.v theories/PCUICConversion.v +theories/PCUICConversionInversion.v theories/PCUICGeneration.v theories/PCUICAlpha.v theories/PCUICCtxShape.v diff --git a/pcuic/theories/PCUICConversionInversion.v b/pcuic/theories/PCUICConversionInversion.v new file mode 100644 index 000000000..6233bda1f --- /dev/null +++ b/pcuic/theories/PCUICConversionInversion.v @@ -0,0 +1,216 @@ +From Equations Require Import Equations. +From MetaCoq.PCUIC Require Import PCUICAst. +From MetaCoq.PCUIC Require Import PCUICAstUtils. +From MetaCoq.PCUIC Require Import PCUICCumulativity. +From MetaCoq.PCUIC Require Import PCUICCumulProp. +From MetaCoq.PCUIC Require Import PCUICEquality. +From MetaCoq.PCUIC Require Import PCUICNormal. +From MetaCoq.PCUIC Require Import PCUICReduction. +From MetaCoq.PCUIC Require Import PCUICTyping. +From MetaCoq.Template Require Import config. +From MetaCoq.Template Require Import utils. + +Local Set Keyed Unification. + +Set Default Goal Selector "!". + +Section fixed. + Context {cf : checker_flags}. + Context (Σ : global_env_ext). + Context (wf : ∥wf_ext Σ∥). + + Definition conv_pb_rel (pb : conv_pb) := + match pb with + | Conv => eq_universe + | Cumul => leq_universe + end. + + Definition eq_termp_napp leq napp := + eq_term_upto_univ_napp Σ (eq_universe Σ) (conv_pb_rel leq Σ) napp. + + Notation eq_termp leq := (eq_termp_napp leq 0). + + Lemma eq_termp_mkApps_inv leq v args v' args' : + isApp v = false -> + isApp v' = false -> + eq_termp leq (mkApps v args) (mkApps v' args') -> + eq_termp_napp leq #|args| v v' × All2 (fun x y => eq_term Σ Σ x y) args args'. + Proof. + intros noapp1 noapp2 eq. + now apply eq_term_upto_univ_mkApps_inv in eq as (?&?). + Qed. + + Definition isIndConstructApp (t : term) : bool := + match (decompose_app t).1 with + | tInd _ _ + | tConstruct _ _ _ => true + | _ => false + end. + + Lemma isIndConstructApp_mkApps hd args : + isIndConstructApp (mkApps hd args) = isIndConstructApp hd. + Proof. + unfold isIndConstructApp. + destruct (mkApps_elim hd args). + now rewrite !decompose_app_mkApps by easy. + Qed. + + Lemma eq_term_upto_univ_napp_nonind Re Rle napp t t' : + eq_term_upto_univ_napp Σ Re Rle napp t t' -> + isIndConstructApp t = false -> + eq_term_upto_univ Σ Re Rle t t'. + Proof. + intros eq not_ind. + generalize 0. + intros k. + induction eq in k, not_ind |- *; eauto using eq_term_upto_univ_napp. + - rewrite (isIndConstructApp_mkApps _ [u]) in not_ind. + constructor; auto. + - discriminate not_ind. + - discriminate not_ind. + Qed. + + Lemma whne_red1_isIndConstructApp Γ t t' : + whne RedFlags.default Σ Γ t -> + red1 Σ Γ t t' -> + isIndConstructApp t' = isIndConstructApp t. + Proof. + apply (whne_red1_ind + RedFlags.default Σ Γ + (fun t t' => isIndConstructApp t' = isIndConstructApp t)); intros; cbn in *; try easy. + - now rewrite !(isIndConstructApp_mkApps _ [v]). + - now rewrite (isIndConstructApp_mkApps _ [v]), (isIndConstructApp_mkApps _ [N2]). + - now rewrite !isIndConstructApp_mkApps. + - now rewrite !isIndConstructApp_mkApps. + - now rewrite !isIndConstructApp_mkApps. + Qed. + + Lemma whne_red_isIndConstructApp Γ t t' : + whne RedFlags.default Σ Γ t -> + red Σ Γ t t' -> + isIndConstructApp t' = isIndConstructApp t. + Proof. + intros wh r. + induction r in wh |- * using red_rect_n1. + - easy. + - apply whne_red1_isIndConstructApp in X as ->; eauto. + eapply whne_pres; eauto. + Qed. + + Lemma conv_cum_alt leq Γ t t' : + conv_cum leq Σ Γ t t' <-> + ∥∑ v v', (red Σ Γ t v × red Σ Γ t' v') × eq_termp leq v v'∥. + Proof. + assert (forall P Q, (P <~> Q) -> (∥P∥ <-> ∥Q∥)) by + (intros P Q H; split; intros [p]; constructor; apply H in p; auto). + destruct leq; cbn; apply H; [apply conv_alt_red|apply cumul_alt]. + Qed. + + Lemma eq_term_eq_termp leq x y : + eq_term Σ Σ x y -> + eq_termp leq x y. + Proof. + destruct leq; [easy|]. + cbn. + now apply PCUICCumulProp.eq_term_upto_univ_napp_leq. + Qed. + + Lemma conv_cum_mkApps_inv leq Γ hd args hd' args' : + conv_cum leq Σ Γ (mkApps hd args) (mkApps hd' args') -> + match hd with + | tApp _ _ + | tInd _ _ + | tConstruct _ _ _ => False + | _ => True + end -> + match hd' with + | tApp _ _ + | tInd _ _ + | tConstruct _ _ _ => False + | _ => True + end -> + whne RedFlags.default Σ Γ hd -> + whne RedFlags.default Σ Γ hd' -> + ∥conv_cum leq Σ Γ hd hd' × All2 (conv Σ Γ) args args'∥. + Proof. + intros conv shape shape' wh wh'. + apply conv_cum_alt in conv as [(?&?&(r1&r2)&e)]. + eapply whne_red_from_mkApps in r1; auto; [|now destruct hd]. + destruct r1 as [(?&?&->&?&?)]. + eapply whne_red_from_mkApps in r2; auto; [|now destruct hd']. + destruct r2 as [(?&?&->&?&?)]. + apply eq_termp_mkApps_inv in e as (?&?). + 2-3: erewrite whne_red_isApp; [| |now eauto]; eauto; try now destruct hd, hd'. + constructor. + split. + - apply conv_cum_alt. + constructor. + exists x1, x. + split; eauto. + eapply eq_term_upto_univ_napp_nonind; eauto. + erewrite whne_red_isIndConstructApp; [| |now eauto]; eauto. + now destruct hd. + - clear -a1 a a0. + induction a1 in args, args', x2, a, x3, a0, a1 |- *; depelim a; depelim a0; [now constructor|]. + constructor. + + apply conv_alt_red. + now exists x, y. + + now apply IHa1. + Qed. + + Lemma conv_cum_tCase_inv leq Γ p motive discr brs p' motive' discr' brs' : + conv_cum leq Σ Γ (tCase p motive discr brs) (tCase p' motive' discr' brs') -> + whnf RedFlags.default Σ Γ (tCase p motive discr brs) -> + whnf RedFlags.default Σ Γ (tCase p' motive' discr' brs') -> + ∥ p = p' × + Σ;;; Γ |- motive = motive' × + Σ;;; Γ |- discr = discr' × + All2 (fun br br' => br.1 = br'.1 × Σ;;; Γ |- br.2 = br'.2) brs brs'∥. + Proof. + intros conv whl whr. + depelim whl; solve_discr. + depelim H; solve_discr; try discriminate. + depelim whr; solve_discr. + depelim H0; solve_discr; try discriminate. + apply conv_cum_alt in conv as [(?&?&(r1&r2)&?)]. + apply whne_red_from_tCase in r1 as [(?&?&?&->&?&?&?)]; auto. + apply whne_red_from_tCase in r2 as [(?&?&?&->&?&?&?)]; auto. + depelim e. + constructor. + split; [easy|]. + split; [apply conv_alt_red; now exists x1, x|]. + split; [apply conv_alt_red; now exists x2, x4|]. + clear -a a0 a1. + induction a1 in brs, brs', x3, a, x5, a0, a1 |- *; depelim a; depelim a0; [now constructor|]. + constructor. + + destruct p, p0, r. + split; [congruence|]. + apply conv_alt_red; now exists x.2, y.2. + + now apply IHa1. + Qed. + + (* TODO: Remove to a place that actually should be depending on typing *) + Lemma conv_cum_red_inv leq Γ t1 t2 t1' t2' : + conv_cum leq Σ Γ t1 t2 -> + red Σ Γ t1 t1' -> + red Σ Γ t2 t2' -> + conv_cum leq Σ Γ t1' t2'. + Proof. + destruct wf. + intros cc r1 r2. + destruct leq; cbn in *. + - destruct cc. + constructor. + eapply PCUICConversion.conv_red_l_inv; [eauto| |eauto]. + apply conv_sym. + eapply PCUICConversion.conv_red_l_inv; [eauto| |eauto]. + apply conv_sym. + auto. + - destruct cc. + constructor. + eapply PCUICConversion.cumul_red_l_inv; [eauto| |eauto]. + eapply PCUICConversion.cumul_red_r_inv; [eauto| |eauto]. + auto. + Qed. + +End fixed. diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index 853b250e1..15e987717 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -3,7 +3,7 @@ From Coq Require Import Bool String List Program BinPos Compare_dec Arith Lia. From MetaCoq.Template Require Import config Universes monad_utils utils BasicAst AstUtils UnivSubst. -From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICLiftSubst PCUICTyping PCUICInduction. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICEquality PCUICLiftSubst PCUICTyping PCUICInduction. Require Import ssreflect. Set Asymmetric Patterns. @@ -980,9 +980,54 @@ Proof. eapply whnf_pres; eauto. Qed. +Lemma notapp_eq_mkApps_inv t hd args : + negb (isApp t) -> + t = mkApps hd args -> + args = [] /\ hd = t. +Proof. + intros noapp ->. + destruct args using List.rev_ind; [|rewrite <- mkApps_nested in noapp; discriminate noapp]. + easy. +Qed. + +Lemma whne_red1_isApp Σ Γ t t' : + whne RedFlags.default Σ Γ t -> + red1 Σ Γ t t' -> + isApp t' = isApp t. +Proof. + intros wh r. + apply (whne_red1_ind + RedFlags.default Σ Γ + (fun t t' => isApp t' = isApp t)) + with (t := t) (t' := t'); intros; cbn in *; try easy. + - apply OnOne2_length in X. + destruct args using List.rev_ind; + destruct x using List.rev_ind; cbn in *; + rewrite ?app_length in X; + cbn in *; try easy. + now rewrite <- !mkApps_nested. + - destruct args using List.rev_ind; cbn in *; [now auto|]. + now rewrite <- !mkApps_nested. + - destruct args using List.rev_ind; cbn in *; [now auto|]. + now rewrite <- !mkApps_nested. +Qed. + +Lemma whne_red_isApp Σ Γ t t' : + whne RedFlags.default Σ Γ t -> + red Σ Γ t t' -> + isApp t' = isApp t. +Proof. + intros wh r. + induction r in wh |- * using red_rect_n1. + - easy. + - apply whne_red1_isApp in X as ->; eauto. + eapply whne_pres; eauto. +Qed. + Lemma whne_red1_from_mkApps f Σ Γ hd args t : RedFlags.beta f = true -> RedFlags.fix_ f = true -> + (* Can we get rid of this precondition? *) isApp hd = false -> whne f Σ Γ (mkApps hd args) -> red1 Σ Γ (mkApps hd args) t -> @@ -1036,87 +1081,176 @@ Proof. eauto with pcuic. Qed. -Lemma whnf_red1_from_mkApps f Σ Γ hd args t : - RedFlags.beta f = true -> - RedFlags.fix_ f = true -> - isApp hd = false -> - whnf f Σ Γ (mkApps hd args) -> - red1 Σ Γ (mkApps hd args) t -> - ∥∑ hd' args', - t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. -Proof. - intros nobeta nofix noapp wh r. - depelim wh; - try solve [ - match goal with - | [H: ?a = mkApps ?h ?args |- _] => - apply (mkApps_notApp_inj a h [] args) in H as (<-&<-); auto; - constructor; eexists _, []; split; [reflexivity|]; - eauto with pcuic - end]. - - now eapply whne_red1_from_mkApps. - - apply mkApps_notApp_inj in x as (<-&<-); auto. - apply red1_mkApps_tConstruct_inv in r as (?&->&?). - constructor; eexists _, _; split; [reflexivity|]. - split; [easy|]. - eapply OnOne2_All2; eauto. - - apply mkApps_notApp_inj in x as (<-&<-); auto. - apply red1_mkApps_tInd_inv in r as (?&->&?). - constructor; eexists _, _; split; [reflexivity|]. - split; [easy|]. - eapply OnOne2_All2; eauto. - - apply mkApps_notApp_inj in x as (<-&<-); auto. - apply red1_mkApps_tFix_inv in r as [[(?&->&?)|(?&->&?)]|(?&->&?)]; auto. - 1-3: constructor; eexists _, _; split; [reflexivity|]; - eauto using OnOne2_All2, All2_same with pcuic. - destruct unfold_fix as [(?&?)|]; [|easy]. - unfold is_constructor. - now rewrite H. - - apply mkApps_notApp_inj in x as (<-&<-); auto. - apply red1_mkApps_tCoFix_inv in r as [[(?&->&?)|(?&->&?)]|(?&->&?)]; auto. - 1-3: constructor; eexists _, _; split; [reflexivity|]; - eauto using OnOne2_All2, All2_same with pcuic. -Qed. - -Lemma whnf_red_from_mkApps f Σ Γ hd args t : - RedFlags.beta f = true -> - RedFlags.fix_ f = true -> +Lemma whne_red_from_mkApps Σ Γ hd args t : + whne RedFlags.default Σ Γ hd -> isApp hd = false -> - whnf f Σ Γ (mkApps hd args) -> red Σ Γ (mkApps hd args) t -> ∥∑ hd' args', t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. Proof. - intros nobeta nofix notapp wh r. remember (mkApps hd args) as from eqn:fromeq. - revert args fromeq. - induction r using red_rect_1n; intros args ->. - - eapply whnf_red1_from_mkApps in r as [(?&?&->&?&?)]; eauto. - now constructor; eexists _, _; split; [reflexivity|]. + intros wh notapp r. + revert hd notapp wh args fromeq. + induction r using red_rect_n1; intros hd notapp wh args ->. - constructor; eexists _, _; split; [reflexivity|]. split; [easy|]. now apply All2_same. - - fold (red Σ Γ (mkApps hd args) y) in r1. - apply whnf_red1_from_mkApps in r1. + - specialize (IHr _ notapp wh _ eq_refl) as [(?&?&->&?&?)]. + eapply whne_red1_from_mkApps in X as [(?&?&->&?&?)]; eauto. + 5: apply whne_mkApps; eapply whne_pres; eauto. + all: auto. + 2: erewrite whne_red_isApp; eauto. + constructor; eexists _, _; split; [reflexivity|]. + split; [now etransitivity; eauto|]. + eapply All2_trans; eauto. + typeclasses eauto. +Qed. -Lemma whnf_red_from_mkApps Σ Γ t - whnf RedFlags.default Σ Γ t -> - red Σ Γ t (mkApps (tInd ind u) args) -> - exists args', t = mkApps (tInd ind u) args'. +Lemma whne_eq_term_upto_univ_napp f Σ Γ t Re Rle napp t' : + whne f Σ Γ t -> + eq_term_upto_univ_napp Σ Re Rle napp t t' -> + whne f Σ Γ t'. +Proof. + intros wh eq. + induction wh in Re, Rle, napp, t, t', eq, wh |- *; depelim eq; + try solve [eauto using whne; depelim wh; solve_discr; eauto using whne]. + - destruct args as [|? ? _] using List.rev_ind; [discriminate x|]. + rewrite <- mkApps_nested in x; cbn in x; inv x. + apply eq_term_upto_univ_napp_mkApps_l_inv in eq1 as (?&?&(?&?)&->). + depelim e. + change (tApp ?h ?a) with (mkApps h [a]); rewrite mkApps_nested. + assert (All2 (eq_term_upto_univ Σ Re Re) (args ++ [x0]) (x1 ++ [u'])) + by (now apply All2_app). + unfold unfold_fix in *. + destruct (nth_error mfix idx) eqn:nth; [|easy]. + eapply All2_nth_error_Some in nth; eauto. + destruct nth as (?&?&?&?). + inv H. + rewrite e0 in H0. + eapply All2_nth_error_Some in H0; eauto. + destruct H0 as (?&?&?). + eapply whne_fixapp. + + unfold unfold_fix. + rewrite e. + reflexivity. + + eassumption. + + eapply IHwh; eauto. + - destruct args using List.rev_ind; [|rewrite <- mkApps_nested in x; discriminate x]. + now rewrite nth_error_nil in H0. +Qed. + +Lemma whne_eq_term {cf:checker_flags} f Σ φ Γ t t' : + whne f Σ Γ t -> + eq_term Σ φ t t' -> + whne f Σ Γ t'. +Proof. + apply whne_eq_term_upto_univ_napp. +Qed. + +Lemma whnf_eq_term_upto_univ_napp f Σ Γ t Re Rle napp t' : + whnf f Σ Γ t -> + eq_term_upto_univ_napp Σ Re Rle napp t t' -> + whnf f Σ Γ t'. +Proof. + intros wh eq. + depelim wh. + - constructor. + eapply whne_eq_term_upto_univ_napp; eauto. + - depelim eq. + apply whnf_sort. + - depelim eq. + apply whnf_prod. + - depelim eq. + apply whnf_lam. + - apply eq_term_upto_univ_napp_mkApps_l_inv in eq as (?&?&(?&?)&->). + depelim e. + apply whnf_cstrapp. + - apply eq_term_upto_univ_napp_mkApps_l_inv in eq as (?&?&(?&?)&->). + depelim e. + apply whnf_indapp. + - apply eq_term_upto_univ_napp_mkApps_l_inv in eq as (?&?&(?&?)&->). + depelim e. + apply whnf_fixapp. + unfold unfold_fix in *. + destruct (nth_error mfix' idx) eqn:nth; [|easy]. + eapply All2_nth_error_Some_r in nth; eauto. + destruct nth as (?&?&(?&?)&?). + rewrite e in H. + rewrite e2 in H. + eapply All2_nth_error_None; eauto. + - apply eq_term_upto_univ_napp_mkApps_l_inv in eq as (?&?&(?&?)&->). + depelim e. + apply whnf_cofixapp. +Qed. + +Lemma whnf_eq_term {cf:checker_flags} f Σ φ Γ t t' : + whnf f Σ Γ t -> + eq_term Σ φ t t' -> + whnf f Σ Γ t'. +Proof. + apply whnf_eq_term_upto_univ_napp. +Qed. + +Hint Resolve All2_same : pcuic. + +Lemma whne_red1_from_tCase Σ Γ p motive discr brs t : + whne RedFlags.default Σ Γ (tCase p motive discr brs) -> + red1 Σ Γ (tCase p motive discr brs) t -> + ∥∑ motive' discr' brs', + t = tCase p motive' discr' brs' × + red Σ Γ motive motive' × + red Σ Γ discr discr' × + All2 (fun br br' => br.1 = br'.1 × red Σ Γ br.2 br'.2) brs brs'∥. Proof. - remember (mkApps (tInd ind u) args) as hd eqn:hdeq. intros wh r. - revert args hdeq. - induction r using red_rect'; intros args ->. - - easy. - - apply whnf_red1_mkApps_tInd in X as (?&->); eauto. - eapply whnf_pres; eauto. + remember (tCase p motive discr brs) as from eqn:fromeq. + revert motive discr brs fromeq. + apply (whne_red1_ind + RedFlags.default Σ Γ + (fun t t' => (forall motive discr brs, t = tCase p motive discr brs -> _ : Prop))) + with (t0 := from) (t' := t); intros; cbn in *; try discriminate; solve_discr; auto. + - inv H0. + constructor; eexists _, _, _; split; [reflexivity|]. + eauto 7 with pcuic. + - inv H1. + constructor; eexists _, _, _; split; [reflexivity|]. + eauto 7 with pcuic. + - inv H0. + constructor; eexists _, _, _; split; [reflexivity|]. + split; [eauto|]. + split; [eauto|]. + eapply OnOne2_All2; eauto. + intros ? ? (?&?). + eauto. +Qed. + +Lemma whne_red_from_tCase Σ Γ p motive discr brs t : + whne RedFlags.default Σ Γ (tCase p motive discr brs) -> + red Σ Γ (tCase p motive discr brs) t -> + ∥∑ motive' discr' brs', + t = tCase p motive' discr' brs' × + red Σ Γ motive motive' × + red Σ Γ discr discr' × + All2 (fun br br' => br.1 = br'.1 × red Σ Γ br.2 br'.2) brs brs'∥. +Proof. + intros wh r. + remember (tCase p motive discr brs) eqn:fromeq. + revert fromeq. + induction r using red_rect_n1; intros ->. + - constructor; eauto 9 with pcuic. + - destruct (IHr eq_refl) as [(?&?&?&->&?&?&?)]. + eapply whne_red1_from_tCase in X as [(?&?&?&->&?&?&?)]. + + constructor; eexists _, _, _; split; [reflexivity|]. + split; [etransitivity; eauto|]. + split; [etransitivity; eauto|]. + eapply All2_trans; eauto. + clear. + intros x y z (?&?) (?&?). + split; [congruence|]. + etransitivity; eauto. + + eapply whne_pres; eauto. Qed. -Lemma whne_red1_from_mkApps_tCase f Σ Γ p motive discr brs : - all_except_delta f -> - whnf f Σ Γ (tCase p motive discr brs) -> - red1 Σ Γ (mkApps (tCase (* Definition head_arg_is_constructor mfix idx args := diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index 86a4671b1..13a5ddcd9 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -1387,23 +1387,6 @@ Section Lemmata. eapply cumul_conv_ctx. all: eauto. Qed. - Lemma conv_cum_zipp : - forall Γ leq u v π1 π2, - conv_cum leq Σ Γ u v -> - conv_context Σ (Γ,,, stack_context π1) (Γ,,, stack_context π2) -> - conv_cum leq Σ Γ (zipp u π1) (zipp v π2). - Proof. - intros Γ leq u v π1 π2 c_uv c_stacks. - unfold zipp. - destruct (decompose_stack π1) eqn:decomp1, (decompose_stack π2) eqn:decomp2. - apply decompose_stack_eq in decomp1 as ->. - apply decompose_stack_eq in decomp2 as ->. - rewrite !stack_context_appstack in c_stacks. - destruct hΣ. - destruct leq. - - destruct c_uv as [c_uv]. - constructor. - apply mkApps_conv_args; auto. End Lemmata. Lemma Case_Construct_ind_eq {cf:checker_flags} Σ (hΣ : ∥ wf Σ.1 ∥) : diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index fd98e55a9..8dbf25009 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -5,7 +5,8 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICReflect PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICCumulativity PCUICEquality PCUICConversion PCUICSafeLemmata PCUICNormal PCUICInversion PCUICReduction PCUICPosition - PCUICContextConversion PCUICSN PCUICUtils PCUICWeakening. + PCUICContextConversion PCUICSN PCUICUtils PCUICWeakening + PCUICConversionInversion. From MetaCoq.SafeChecker Require Import PCUICSafeReduce. Require Import Equations.Prop.DepElim. @@ -1635,41 +1636,7 @@ Section Conversion. unfold global_ext_levels in *. now apply LevelSet.mem_spec in y_mem. Qed. - - Definition conv_pb_rel (pb : conv_pb) := - match pb with - | Conv => eq_universe - | Cumul => leq_universe - end. - - Definition eq_termp_napp leq napp := - eq_term_upto_univ_napp Σ (eq_universe Σ) (conv_pb_rel leq Σ) napp. - - Notation eq_termp leq := (eq_termp_napp leq 0). - - Lemma eq_termp_mkApps_inv leq v args v' args' : - isApp v = false -> - isApp v' = false -> - eq_termp leq (mkApps v args) (mkApps v' args') -> - eq_termp_napp leq #|args| v v' × All2 (fun x y => eq_term Σ x y) args args'. - Proof. - intros noapp1 noapp2 eq. - now apply eq_term_upto_univ_mkApps_inv in eq as (?&?). - Qed. - Lemma conv_cum_alt leq Γ t t' : - conv_cum leq Σ Γ t t' -> - ∥∑ v v', (red Σ Γ t v × red Σ Γ t' v') × eq_termp leq v v'∥. - Proof. - destruct leq; cbn. - - intros [c]. - apply conv_alt_red in c. - now constructor. - - intros [c]. - apply cumul_alt in c. - now constructor. - Qed. - Lemma zipp_as_mkApps t π : zipp t π = mkApps t (decompose_stack π).1. Proof. @@ -2492,6 +2459,72 @@ Section Conversion. now apply invert_type_mkApps_tProd in typ. Qed. + Lemma whnf_whne_upgrade Γ t : + whnf RedFlags.default Σ Γ t -> + whne nodelta_flags Σ Γ t -> + whne RedFlags.default Σ Γ t. + Proof. + intros whn whe. + induction whe; cbn in *; try easy. + - now depelim whn; solve_discr. + - depelim whn. + + easy. + + destruct v0 as [|? ? _] using List.rev_ind; [discriminate H|]. + rewrite <- mkApps_nested in H. + cbn in H; noconf H. + constructor. + eapply IHwhe. + apply whnf_cstrapp. + + destruct v0 as [|? ? _] using List.rev_ind; [discriminate H|]. + rewrite <- mkApps_nested in H. + cbn in H; noconf H. + constructor. + eapply IHwhe. + apply whnf_indapp. + + destruct v0 as [|? ? _] using List.rev_ind; [discriminate H0|]. + rewrite <- mkApps_nested in H0. + cbn in H0; noconf H0. + constructor. + eapply IHwhe. + apply whnf_fixapp. + destruct unfold_fix; [|easy]. + destruct p. + apply nth_error_None. + apply nth_error_None in H. + rewrite app_length in H; cbn in H. + lia. + + destruct v0 as [|? ? _] using List.rev_ind; [discriminate H|]. + rewrite <- mkApps_nested in H. + cbn in H; noconf H. + constructor. + eapply IHwhe. + apply whnf_cofixapp. + - depelim whn; solve_discr; try easy. + rewrite H1 in H. + congruence. + - depelim whn; solve_discr; easy. + - depelim whn; solve_discr; easy. + Qed. + + Lemma whne_conv_context f Γ Γ' t : + whne f Σ Γ t -> + conv_context Σ Γ Γ' -> + whne f Σ Γ' t. + Proof. + intros wh conv. + induction wh; eauto using whne. + destruct nth_error eqn:nth; [|easy]. + cbn in H. + eapply context_relation_nth in nth; eauto. + destruct nth as (?&?&?&?). + constructor. + rewrite e. + cbn. + depelim c1; auto. + cbn in *. + congruence. + Qed. + (* Notation whne := (whne flags Σ). @@ -2541,6 +2574,9 @@ Section Conversion. todo "Completeness"%string. Qed. *) + + Axiom marker : string -> Type. + Axiom unwrap_marker : forall s {m : marker s} {A}, A. Opaque reduce_stack. Equations(noeqns) _isconv_prog (Γ : context) (leq : conv_pb) @@ -3029,33 +3065,92 @@ Section Conversion. Qed. Next Obligation. apply h; clear h. + (* Proof idea: + 1) prove that the 2 tCase are fully whnf because the discriminees are + equal to their full reduction + 2) prove that the 2 tCase are in conv_cum relation because they are whnf + 3) wlog leq = Conv because there is no cumulativity between tCase + 4) invert the conv judgement using that the tCase are whne + (so preserved by wh-reductions) + *) destruct ir1 as (_&wh1), ir2 as (_&wh2). cbn in *. - apply conv_cum_alt in H. - rewrite !zipp_as_mkApps in *. + symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). + apply eqb_term_spec in c_is_red. + apply eqb_term_spec in c'_is_red. + eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. + eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. + rewrite !zipp_as_mkApps in H. + rewrite zipp_as_mkApps in wh1, wh2. apply whnf_mkApps_inv in wh1; [|easy]. apply whnf_mkApps_inv in wh2; [|easy]. depelim wh1; solve_discr. depelim wh2; solve_discr. - unfold zipp in wh1, wh2. - (* Proof idea: - 1) prove that the 2 tCase are in conv_cum relation - because they are whne (needs to add thys hypothesis) - 2) wlog leq = Conv because there is no cumulativity between tCase - 3) invert the conv judgement using that the tCase are whne - (so preserved by wh-reductions) - *) - todo "Completeness"%string. - Qed. - Next Obligation. - apply h ; clear h. - cbn. (* Same idea as previous case *) - todo "Completeness"%string. - Qed. - Next Obligation. - apply h; clear h. - cbn. (* Same idea as previous case *) - todo "Completeness"%string. + depelim H; cbn in *; try easy; solve_discr. + depelim H0; cbn in *; try easy; solve_discr. + apply whnf_whne_upgrade in c_is_red; auto. + apply whnf_whne_upgrade in c'_is_red; auto. + destruct hΣ, hx. + assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). + { eapply whne_conv_context; eauto. + apply conv_context_sym; eauto. } + apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. + apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. + now constructor. + Qed. + Next Obligation. + apply h; cbn; clear h. + destruct ir1 as (_&wh1), ir2 as (_&wh2). + cbn in *. + symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). + apply eqb_term_spec in c_is_red. + apply eqb_term_spec in c'_is_red. + eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. + eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. + rewrite !zipp_as_mkApps in H. + rewrite zipp_as_mkApps in wh1, wh2. + apply whnf_mkApps_inv in wh1; [|easy]. + apply whnf_mkApps_inv in wh2; [|easy]. + depelim wh1; solve_discr. + depelim wh2; solve_discr. + depelim H; cbn in *; try easy; solve_discr. + depelim H0; cbn in *; try easy; solve_discr. + apply whnf_whne_upgrade in c_is_red; auto. + apply whnf_whne_upgrade in c'_is_red; auto. + destruct hΣ, hx. + assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). + { eapply whne_conv_context; eauto. + apply conv_context_sym; eauto. } + apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. + apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. + now constructor. + Qed. + Next Obligation. + apply h; cbn; clear h. + destruct ir1 as (_&wh1), ir2 as (_&wh2). + cbn in *. + symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). + apply eqb_term_spec in c_is_red. + apply eqb_term_spec in c'_is_red. + eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. + eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. + rewrite !zipp_as_mkApps in H. + rewrite zipp_as_mkApps in wh1, wh2. + apply whnf_mkApps_inv in wh1; [|easy]. + apply whnf_mkApps_inv in wh2; [|easy]. + depelim wh1; solve_discr. + depelim wh2; solve_discr. + depelim H; cbn in *; try easy; solve_discr. + depelim H0; cbn in *; try easy; solve_discr. + apply whnf_whne_upgrade in c_is_red; auto. + apply whnf_whne_upgrade in c'_is_red; auto. + destruct hΣ, hx. + assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). + { eapply whne_conv_context; eauto. + apply conv_context_sym; eauto. } + apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. + apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. + now constructor. Qed. Next Obligation. (* Proof idea: @@ -3066,7 +3161,31 @@ Section Conversion. 2) deduce that ind = ind' and par = par' 3) conclude that it contradicts eqDiff *) - todo "Completeness"%string. + destruct ir1 as (_&wh1), ir2 as (_&wh2). + cbn in *. + symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). + apply eqb_term_spec in c_is_red. + apply eqb_term_spec in c'_is_red. + eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. + eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. + rewrite !zipp_as_mkApps in H. + rewrite zipp_as_mkApps in wh1, wh2. + apply whnf_mkApps_inv in wh1; [|easy]. + apply whnf_mkApps_inv in wh2; [|easy]. + depelim wh1; solve_discr. + depelim wh2; solve_discr. + depelim H; cbn in *; try easy; solve_discr. + depelim H0; cbn in *; try easy; solve_discr. + apply whnf_whne_upgrade in c_is_red; auto. + apply whnf_whne_upgrade in c'_is_red; auto. + destruct hΣ, hx. + assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). + { eapply whne_conv_context; eauto. + apply conv_context_sym; eauto. } + apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. + apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. + rewrite eq_inductive_refl, Nat.eqb_refl in eqDiff. + congruence. Qed. Next Obligation. eapply red_wellformed ; auto. @@ -3249,13 +3368,30 @@ Section Conversion. eapply conv_Proj_c. assumption. Qed. Next Obligation. - apply h; clear h. - cbn. - todo "Completeness"%string. + todo "incompleteness". + (*exfalso; apply h; cbn; clear h. + destruct ir1 as (_&wh1), ir2 as (_&wh2). + (*symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). + apply eqb_term_spec in c_is_red. + apply eqb_term_spec in c'_is_red. + eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. + eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete].*) + rewrite !zipp_as_mkApps in H. + rewrite zipp_as_mkApps in wh1, wh2. + apply whnf_mkApps_inv in wh1; [|easy]. + apply whnf_mkApps_inv in wh2; [|easy]. + depelim wh1; solve_discr. + depelim wh2; solve_discr. + depelim H; cbn in *; try easy; solve_discr. + depelim H0; cbn in *; try easy; solve_discr. + apply whnf_whne_upgrade in c_is_red; auto. + apply whnf_whne_upgrade in c'_is_red; auto. + rewrite zipp_as_mkApps in H. + todo "Completeness"%string.*) Qed. Next Obligation. (* Proof idea: c and c' are whne so from H, p = p' contradicting eqDiff *) - todo "Completeness"%string. + todo "Incompleteness"%string. Qed. (* tFix *) @@ -3427,9 +3563,50 @@ Section Conversion. - rewrite e. assumption. Qed. Next Obligation. + destruct hΣ as [wΣ]. + apply unfold_one_fix_red_zipp in eq1 as r1. + destruct r1 as [r1]. + match type of eq3 with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; + pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + end. + rewrite <- eq3 in r2. + rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. + cbn in d2. + apply unfold_one_fix_decompose in eq1 as d1. + assert (r2' : red (fst Σ) (Γ ,,, stack_context π1) (zipp fn θ) (zipp fn' (ρ +++ θ'))). + { unfold zipp. + case_eq (decompose_stack ρ). intros l ξ e. + rewrite e in d2. cbn in d2. subst. + pose proof (decompose_stack_eq _ _ _ e). subst. + rewrite stack_cat_appstack. rewrite decompose_stack_appstack. + rewrite <- eq2. + cbn in r2. rewrite 2!zipc_appstack in r2. cbn in r2. + rewrite <- eq2 in d1. cbn in d1. subst. + case_eq (decompose_stack π1). intros l1 ρ1 e1. + simpl. rewrite e1 in r2. simpl in r2. + pose proof (decompose_stack_eq _ _ _ e1). subst. + rewrite decompose_stack_twice with (1 := e1). cbn. + rewrite app_nil_r. + rewrite stack_context_appstack. assumption. + } + pose proof (red_trans _ _ _ _ _ r1 r2') as r. + assert (e : stack_context π1 = stack_context (ρ +++ θ')). + { case_eq (decompose_stack ρ). intros l ξ e. + rewrite e in d2. cbn in d2. subst. + pose proof (decompose_stack_eq _ _ _ e). subst. + rewrite stack_cat_appstack. + rewrite stack_context_appstack. + rewrite <- eq2 in d1. cbn in d1. subst. + case_eq (decompose_stack π1). intros l1 ρ1 e1. + pose proof (decompose_stack_eq _ _ _ e1). subst. + rewrite stack_context_appstack. reflexivity. + } apply h; clear h. - (* Idea: a fixpoint should be convertible to its unfolding + *) - todo "Completeness"%string. + rewrite <- e. + eapply conv_cum_red_inv; eauto. + exact hΣ'. Qed. Next Obligation. cbn. rewrite zipc_appstack. cbn. diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index 96175c18b..112b2f438 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -1659,9 +1659,6 @@ Section Reduce. assumption. Qed. - Lemma stack_context_decompose' π π' : - stack_context (decompose_stack π).2 = stack_context p. - Theorem reduce_term_complete Γ t h : whnf flags Σ Γ (reduce_term Γ t h). Proof. From 4c42000e026c2fac9e37efd740b549432101ece0 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Thu, 22 Oct 2020 16:22:00 +0200 Subject: [PATCH 08/26] Finish tFix case --- erasure/theories/EArities.v | 2 +- pcuic/theories/PCUICContextConversion.v | 110 ++++++ pcuic/theories/PCUICConversionInversion.v | 201 ++++++++-- pcuic/theories/PCUICNormal.v | 237 ++++++++---- safechecker/theories/PCUICSafeConversion.v | 418 +++++++++++++++------ safechecker/theories/PCUICSafeReduce.v | 52 ++- 6 files changed, 794 insertions(+), 226 deletions(-) diff --git a/erasure/theories/EArities.v b/erasure/theories/EArities.v index fb716d7b5..3f91dcd16 100644 --- a/erasure/theories/EArities.v +++ b/erasure/theories/EArities.v @@ -627,4 +627,4 @@ Proof. exists u; split; auto. eapply cumul_prop2; eauto. eapply PCUICValidity.validity; eauto. eapply cumul_prop1; eauto. eapply PCUICValidity.validity; eauto. -Qed. \ No newline at end of file +Qed. diff --git a/pcuic/theories/PCUICContextConversion.v b/pcuic/theories/PCUICContextConversion.v index a01dc2f79..882ba6add 100644 --- a/pcuic/theories/PCUICContextConversion.v +++ b/pcuic/theories/PCUICContextConversion.v @@ -642,6 +642,116 @@ Notation conv_context Σ Γ Γ' := (context_relation (conv_decls Σ) Γ Γ'). Hint Resolve conv_ctx_refl' : pcuic. +Section ContextChangeTypesReduction. + Context {cf : checker_flags}. + Context (Σ : global_env_ext). + Context (wfΣ : wf Σ). + + Inductive change_decl_type : context_decl -> context_decl -> Type := + | change_vass_type : forall (na na' : name) (T T' : term), + change_decl_type (vass na T) (vass na' T') + | change_vdef_type : forall (na na' : name) (b T T' : term), + change_decl_type (vdef na b T) (vdef na' b T'). + + Lemma context_change_decl_types_red1 Γ Γ' s t : + context_relation (fun _ _ => change_decl_type) Γ Γ' -> red1 Σ Γ s t -> red Σ Γ' s t. + Proof. + intros HT X0. induction X0 using red1_ind_all in Γ', HT |- *; eauto. + all:pcuic. + - econstructor. econstructor. + rewrite <- H. + induction HT in i |- *; destruct i; eauto. + now inv p. + - + eapply PCUICReduction.red_abs. eapply IHX0; eauto. eauto. + - + eapply PCUICReduction.red_abs. eauto. eapply IHX0. eauto. + eauto. econstructor. eauto. econstructor. + - + eapply PCUICReduction.red_letin. eapply IHX0; eauto. + all:eauto. + - + eapply PCUICReduction.red_letin; eauto. + - + eapply PCUICReduction.red_letin; eauto. eapply IHX0; eauto. + econstructor. eauto. econstructor. + - eapply PCUICReduction.red_case; eauto. clear. + eapply All_All2_refl. induction brs; eauto. + - eapply PCUICReduction.red_case; eauto. clear. + eapply All_All2_refl. induction brs; eauto. + - destruct ind. + eapply red_case; eauto. + clear - X HT. + induction X. + + econstructor. destruct p. destruct p. + split; eauto. + eapply All_All2_refl. + induction tl; eauto. + + econstructor. now split. + eassumption. + - + eapply PCUICReduction.red_proj_c. eauto. + - + eapply PCUICReduction.red_app; eauto. + - eapply PCUICReduction.red_app; eauto. + - + eapply PCUICReduction.red_prod; eauto. + - + eapply PCUICReduction.red_prod; eauto. eapply IHX0. eauto. eauto. + econstructor. + eauto. econstructor. + - eapply PCUICReduction.red_evar; eauto. + induction X; eauto. econstructor. eapply p; eauto. + induction tl; eauto. + - eapply PCUICReduction.red_fix_one_ty. + eapply OnOne2_impl ; eauto. + intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. + inversion e. subst. clear e. + split ; auto. + - eapply PCUICReduction.red_fix_one_body. + eapply OnOne2_impl ; eauto. + intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. + inversion e. subst. clear e. + split ; auto. + eapply ih ; auto. + clear - HT. + induction (fix_context mfix0) as [| [na [b|] ty] Δ ihΔ]. + + auto. + + simpl. constructor ; eauto. + constructor. + + simpl. constructor ; eauto. + constructor. + - eapply PCUICReduction.red_cofix_one_ty. + eapply OnOne2_impl ; eauto. + intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. + inversion e. subst. clear e. + split ; auto. + - eapply PCUICReduction.red_cofix_one_body. + eapply OnOne2_impl ; eauto. + intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. + inversion e. subst. clear e. + split ; auto. + eapply ih ; auto. + clear - HT. + induction (fix_context mfix0) as [| [na [b|] ty] Δ ihΔ]. + + auto. + + simpl. constructor ; eauto. + constructor. + + simpl. constructor ; eauto. + constructor. + Qed. + + (* todo: update erasure to use this *) + Lemma context_change_decl_types_red Γ Γ' s t : + context_relation (fun _ _ => change_decl_type) Γ Γ' -> red Σ Γ s t -> red Σ Γ' s t. + Proof. + intros. induction X0 using red_rect'; eauto. + etransitivity. eapply IHX0. + eapply context_change_decl_types_red1; eauto. + Qed. +End ContextChangeTypesReduction. + + (* Lemma wf_local_conv_ctx {cf:checker_flags} Σ Γ Δ (wfΓ : wf_local Σ Γ) : wf Σ -> *) (* All_local_env_over typing *) (* (fun (Σ : global_env_ext) (Γ : context) wfΓ (t T : term) Ht => *) diff --git a/pcuic/theories/PCUICConversionInversion.v b/pcuic/theories/PCUICConversionInversion.v index 6233bda1f..5eb93396b 100644 --- a/pcuic/theories/PCUICConversionInversion.v +++ b/pcuic/theories/PCUICConversionInversion.v @@ -1,9 +1,11 @@ From Equations Require Import Equations. From MetaCoq.PCUIC Require Import PCUICAst. From MetaCoq.PCUIC Require Import PCUICAstUtils. +From MetaCoq.PCUIC Require Import PCUICContextConversion. From MetaCoq.PCUIC Require Import PCUICCumulativity. From MetaCoq.PCUIC Require Import PCUICCumulProp. From MetaCoq.PCUIC Require Import PCUICEquality. +From MetaCoq.PCUIC Require Import PCUICLiftSubst. From MetaCoq.PCUIC Require Import PCUICNormal. From MetaCoq.PCUIC Require Import PCUICReduction. From MetaCoq.PCUIC Require Import PCUICTyping. @@ -30,16 +32,6 @@ Section fixed. Notation eq_termp leq := (eq_termp_napp leq 0). - Lemma eq_termp_mkApps_inv leq v args v' args' : - isApp v = false -> - isApp v' = false -> - eq_termp leq (mkApps v args) (mkApps v' args') -> - eq_termp_napp leq #|args| v v' × All2 (fun x y => eq_term Σ Σ x y) args args'. - Proof. - intros noapp1 noapp2 eq. - now apply eq_term_upto_univ_mkApps_inv in eq as (?&?). - Qed. - Definition isIndConstructApp (t : term) : bool := match (decompose_app t).1 with | tInd _ _ @@ -84,17 +76,42 @@ Section fixed. - now rewrite !isIndConstructApp_mkApps. - now rewrite !isIndConstructApp_mkApps. Qed. - - Lemma whne_red_isIndConstructApp Γ t t' : - whne RedFlags.default Σ Γ t -> + + Lemma whnf_red1_isIndConstructApp Γ t t' : + whnf RedFlags.default Σ Γ t -> + red1 Σ Γ t t' -> + isIndConstructApp t' = isIndConstructApp t. + Proof. + intros wh r. + depelim wh. + - eapply whne_red1_isIndConstructApp; eauto. + - depelim r; solve_discr. + - depelim r; auto; solve_discr. + - depelim r; auto; solve_discr. + - apply red1_mkApps_tConstruct_inv in r as (?&->&?). + now rewrite !isIndConstructApp_mkApps. + - apply red1_mkApps_tInd_inv in r as (?&->&?). + now rewrite !isIndConstructApp_mkApps. + - apply red1_mkApps_tFix_inv in r. + 2: { destruct unfold_fix as [(?&?)|]; [|easy]. + now unfold is_constructor; rewrite H. } + now destruct r as [[(?&->&o)|(?&->&o)]|(?&->&o)]; + rewrite !isIndConstructApp_mkApps. + - apply red1_mkApps_tCoFix_inv in r. + now destruct r as [[(?&->&o)|(?&->&o)]|(?&->&o)]; + rewrite !isIndConstructApp_mkApps. + Qed. + + Lemma whnf_red_isIndConstructApp Γ t t' : + whnf RedFlags.default Σ Γ t -> red Σ Γ t t' -> isIndConstructApp t' = isIndConstructApp t. Proof. intros wh r. induction r in wh |- * using red_rect_n1. - easy. - - apply whne_red1_isIndConstructApp in X as ->; eauto. - eapply whne_pres; eauto. + - apply whnf_red1_isIndConstructApp in X as ->; eauto. + eapply whnf_pres; eauto. Qed. Lemma conv_cum_alt leq Γ t t' : @@ -114,6 +131,16 @@ Section fixed. cbn. now apply PCUICCumulProp.eq_term_upto_univ_napp_leq. Qed. + + Lemma eq_termp_mkApps_inv leq v args v' args' : + isApp v = false -> + isApp v' = false -> + eq_termp leq (mkApps v args) (mkApps v' args') -> + eq_termp_napp leq #|args| v v' × All2 (fun x y => eq_term Σ Σ x y) args args'. + Proof. + intros noapp1 noapp2 eq. + now apply eq_term_upto_univ_mkApps_inv in eq as (?&?). + Qed. Lemma conv_cum_mkApps_inv leq Γ hd args hd' args' : conv_cum leq Σ Γ (mkApps hd args) (mkApps hd' args') -> @@ -129,26 +156,37 @@ Section fixed. | tConstruct _ _ _ => False | _ => True end -> - whne RedFlags.default Σ Γ hd -> - whne RedFlags.default Σ Γ hd' -> + whnf RedFlags.default Σ Γ (mkApps hd args) -> + whnf RedFlags.default Σ Γ (mkApps hd' args') -> ∥conv_cum leq Σ Γ hd hd' × All2 (conv Σ Γ) args args'∥. Proof. intros conv shape shape' wh wh'. apply conv_cum_alt in conv as [(?&?&(r1&r2)&e)]. - eapply whne_red_from_mkApps in r1; auto; [|now destruct hd]. - destruct r1 as [(?&?&->&?&?)]. - eapply whne_red_from_mkApps in r2; auto; [|now destruct hd']. - destruct r2 as [(?&?&->&?&?)]. - apply eq_termp_mkApps_inv in e as (?&?). - 2-3: erewrite whne_red_isApp; [| |now eauto]; eauto; try now destruct hd, hd'. + apply whnf_red_from_mkApps in r1 as [(?&?&->&?&?)]; auto. + apply whnf_red_from_mkApps in r2 as [(?&?&->&?&?)]; auto. + assert (isApp x1 = false). + { erewrite whnf_red_isApp. + 3: eauto. + 1: now destruct hd. + apply whnf_mkApps_inv in wh; auto. + now destruct hd. } + assert (isApp x = false). + { erewrite whnf_red_isApp. + 3: eauto. + 1: now destruct hd'. + apply whnf_mkApps_inv in wh'; auto. + now destruct hd'. } + apply eq_termp_mkApps_inv in e as (?&?); auto. constructor. split. - apply conv_cum_alt. constructor. exists x1, x. - split; eauto. + split; [split|]; eauto with pcuic. eapply eq_term_upto_univ_napp_nonind; eauto. - erewrite whne_red_isIndConstructApp; [| |now eauto]; eauto. + erewrite whnf_red_isIndConstructApp; [| |now eauto]. + 1: now destruct hd. + apply whnf_mkApps_inv in wh; auto. now destruct hd. - clear -a1 a a0. induction a1 in args, args', x2, a, x3, a0, a1 |- *; depelim a; depelim a0; [now constructor|]. @@ -189,7 +227,118 @@ Section fixed. + now apply IHa1. Qed. - (* TODO: Remove to a place that actually should be depending on typing *) + Lemma fix_context_change_decl_types Γ mfix mfix' : + #|mfix| = #|mfix'| -> + context_relation (fun _ _ => change_decl_type) (Γ,,, fix_context mfix) (Γ,,, fix_context mfix'). + Proof. + intros len. + apply context_relation_app_inv. + - now rewrite !fix_context_length. + - apply context_relation_refl. + intros. + destruct x. + destruct decl_body; constructor. + - unfold fix_context, mapi. + generalize 0 at 2 4. + induction mfix in mfix', len |- *; intros n. + + destruct mfix'; [|cbn in *; discriminate len]. + constructor. + + destruct mfix'; cbn in *; [discriminate len|]. + apply context_relation_app_inv. + * now rewrite !List.rev_length, !mapi_rec_length. + * constructor; [constructor|]. + constructor. + * apply IHmfix; lia. + Qed. + + Lemma red_tFix_inv Γ mfix idx t : + red Σ Γ (tFix mfix idx) t -> + ∑ mfix', + t = tFix mfix' idx × + All2 (fun d d' => dname d = dname d' × rarg d = rarg d' × + red Σ Γ (dtype d) (dtype d') × + red Σ (Γ,,, fix_context mfix) (dbody d) (dbody d')) + mfix mfix'. + Proof. + intros r. + remember (tFix mfix idx) eqn:fromeq. + revert mfix fromeq. + induction r using red_rect_n1; intros mfix ->. + - eexists _; split; [reflexivity|]. + apply All2_same. + eauto. + - specialize (IHr _ eq_refl) as (?&->&all). + depelim X. + + exfalso. + destruct args using List.rev_ind; [|rewrite <- mkApps_nested in H; discriminate H]. + unfold is_constructor in e0. + rewrite nth_error_nil in e0; congruence. + + match type of all with + | All2 ?P _ _ => eapply OnOne2_All2 with (Q := P) in o + end; eauto. + 2: intros ? ? (?&[= -> -> ->]); eauto. + eexists _; split; [reflexivity|]. + eapply All2_trans; eauto. + intros ? ? ? (?&?&?&?) (?&?&?&?); eauto. + do 2 (split; [easy|]). + split; [etransitivity; eauto|]. + etransitivity; eauto. + + match type of all with + | All2 ?P _ _ => eapply OnOne2_All2 with (Q := P) in o + end; eauto. + 2: { intros ? ? (?&[= -> -> ->]); eauto. + do 2 (split; [easy|]). + split; [etransitivity; eauto|]. + eapply context_change_decl_types_red; eauto. + apply fix_context_change_decl_types. + apply All2_length in all. + easy. } + eexists _; split; [reflexivity|]. + eapply All2_trans; eauto. + intros ? ? ? (?&?&?&?) (?&?&?&?); eauto. + do 2 (split; [easy|]). + split; [etransitivity; eauto|]. + etransitivity; eauto. + Qed. + + Lemma conv_cum_tFix_inv leq Γ mfix idx mfix' idx' : + conv_cum leq Σ Γ (tFix mfix idx) (tFix mfix' idx') -> + ∥idx = idx' × + All2 (fun d d' => rarg d = rarg d' × + Σ;;; Γ |- dtype d = dtype d' × + Σ;;; Γ,,, fix_context mfix |- dbody d = dbody d') + mfix mfix'∥. + Proof. + intros conv. + apply conv_cum_alt in conv as [(?&?&(r1&r2)&?)]. + apply red_tFix_inv in r1 as (?&->&?). + apply red_tFix_inv in r2 as (?&->&?). + depelim e. + constructor. + split; [easy|]. + clear -a a0 a1. + cut (#|mfix| = #|mfix'|); + [|now apply All2_length in a; apply All2_length in a0; apply All2_length in a1]. + revert a a0 a1. + generalize mfix at 1 3 4. + generalize mfix' at 1 3. + intros ctx_fix ctx_fix'. + intros all1 all2 all len_eq. + induction all in mfix, mfix', x1, x, all1, all2, all |- *; + depelim all1; depelim all2; [constructor|]. + constructor; [|now auto]. + destruct p as (?&?&?&?), p0 as (?&?&?&?), r as ((?&?)&?). + split; [congruence|]. + split; [now apply conv_alt_red; exists (dtype x), (dtype y)|]. + apply conv_alt_red. + exists (dbody x), (dbody y). + split; [|easy]. + split; [easy|]. + eapply context_change_decl_types_red; eauto. + eapply fix_context_change_decl_types; eauto. + Qed. + + (* TODO: Move to a place that actually should be depending on typing *) Lemma conv_cum_red_inv leq Γ t1 t2 t1' t2' : conv_cum leq Σ Γ t1 t2 -> red Σ Γ t1 t1' -> diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index 15e987717..d36a8a8b0 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -980,16 +980,6 @@ Proof. eapply whnf_pres; eauto. Qed. -Lemma notapp_eq_mkApps_inv t hd args : - negb (isApp t) -> - t = mkApps hd args -> - args = [] /\ hd = t. -Proof. - intros noapp ->. - destruct args using List.rev_ind; [|rewrite <- mkApps_nested in noapp; discriminate noapp]. - easy. -Qed. - Lemma whne_red1_isApp Σ Γ t t' : whne RedFlags.default Σ Γ t -> red1 Σ Γ t t' -> @@ -1012,94 +1002,199 @@ Proof. now rewrite <- !mkApps_nested. Qed. -Lemma whne_red_isApp Σ Γ t t' : - whne RedFlags.default Σ Γ t -> +Lemma isApp_mkApps' hd args hd' args' : + isApp hd = isApp hd' -> + #|args| = #|args'| -> + isApp (mkApps hd args) = isApp (mkApps hd' args'). +Proof. + intros hd_app len. + destruct args using List.rev_ind; destruct args' using List.rev_ind; auto. + all: rewrite !app_length in len; cbn in len; try lia. + now rewrite <- !mkApps_nested. +Qed. + +Lemma whnf_red1_isApp Σ Γ t t' : + whnf RedFlags.default Σ Γ t -> + red1 Σ Γ t t' -> + isApp t' = isApp t. +Proof. + intros wh r. + depelim wh. + - eapply whne_red1_isApp; eauto. + - depelim r; solve_discr. + - depelim r; auto; solve_discr. + - depelim r; auto; solve_discr. + - apply red1_mkApps_tConstruct_inv in r as (?&->&?). + apply OnOne2_length in o. + now apply isApp_mkApps'. + - apply red1_mkApps_tInd_inv in r as (?&->&?). + apply OnOne2_length in o. + now apply isApp_mkApps'. + - apply red1_mkApps_tFix_inv in r. + 2: { destruct unfold_fix as [(?&?)|]; [|easy]. + now unfold is_constructor; rewrite H. } + now destruct r as [[(?&->&o)|(?&->&o)]|(?&->&o)]; + apply OnOne2_length in o; + apply isApp_mkApps'. + - apply red1_mkApps_tCoFix_inv in r. + now destruct r as [[(?&->&o)|(?&->&o)]|(?&->&o)]; + apply OnOne2_length in o; + apply isApp_mkApps'. +Qed. + +Lemma whnf_red_isApp Σ Γ t t' : + whnf RedFlags.default Σ Γ t -> red Σ Γ t t' -> isApp t' = isApp t. Proof. intros wh r. induction r in wh |- * using red_rect_n1. - easy. - - apply whne_red1_isApp in X as ->; eauto. - eapply whne_pres; eauto. + - apply whnf_red1_isApp in X as ->; eauto. + eapply whnf_pres; eauto. Qed. -Lemma whne_red1_from_mkApps f Σ Γ hd args t : - RedFlags.beta f = true -> - RedFlags.fix_ f = true -> - (* Can we get rid of this precondition? *) - isApp hd = false -> - whne f Σ Γ (mkApps hd args) -> +Ltac solve_red_from_mkApps := + match goal with + | [H: ?a = mkApps ?h ?args |- _] => + destruct args using List.rev_ind; [|rewrite <- mkApps_nested in H; discriminate H]; + cbn in H; subst; + constructor; eexists _, []; split; [reflexivity|]; + eauto with pcuic + end. + +Hint Resolve All2_same All2_firstn All2_skipn OnOne2_All2 red_mkApps All2_app : pcuic. +Lemma whne_red1_from_mkApps Σ Γ hd args t : + whne RedFlags.default Σ Γ (mkApps hd args) -> red1 Σ Γ (mkApps hd args) t -> ∥∑ hd' args', t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. Proof. - intros redbeta redfix noapp wh r. + intros wh r. remember (mkApps hd args) eqn:eq. revert args eq. apply (whne_red1_ind - f Σ Γ + RedFlags.default Σ Γ (fun t t' => (forall args, t = mkApps hd args -> _ : Prop))) with (t1 := t0) (t' := t); intros; cbn in *; - try solve [congruence]; - try solve [ - match goal with - | [H: ?a = mkApps ?h ?args |- _] => - apply (mkApps_notApp_inj a h [] args) in H as (<-&<-); auto; - constructor; eexists _, []; split; [reflexivity|]; - eauto with pcuic - end]. - - destruct args as [|? ? _] using List.rev_ind; [destruct hd; cbn in *; congruence|]. - rewrite <- mkApps_nested in H1. - inv H1. - destruct (H0 _ eq_refl) as [(?&?&->&?&?)]. - constructor; exists x0, (x1 ++ [x]). - rewrite <- mkApps_nested. + try congruence; try solve [solve_red_from_mkApps]. + - destruct args as [|? ? _] using List.rev_ind. + + cbn in H1; subst. + constructor; eexists _, []; split; [reflexivity|]. + eauto with pcuic. + + rewrite <- mkApps_nested in H1. + inv H1. + specialize (H0 _ eq_refl) as [(?&?&->&?&?)]. + constructor; exists x0, (x1 ++ [x]). + rewrite <- mkApps_nested; split; [reflexivity|]. + split; [easy|]; eauto with pcuic. + - destruct args as [|? ? _] using List.rev_ind. + + cbn in H0; subst. + constructor; eexists _, []; split; [reflexivity|]. + eauto with pcuic. + + rewrite <- mkApps_nested in H0. + inv H0. + constructor; exists hd, (args ++ [N2]). + rewrite <- mkApps_nested; split; [reflexivity|]. + split; [reflexivity|]; eauto with pcuic. + - destruct (mkApps_elim hd args0). + apply mkApps_notApp_inj in H3 as (<-&<-); auto; [|now destruct (isApp f)]. + constructor; exists (mkApps (tFix mfix idx) (firstn n x)), (skipn n x). + rewrite mkApps_nested. + rewrite firstn_skipn. split; [easy|]. - split; [now eauto|]. - apply All2_app; eauto. - - destruct args as [|? ? _] using List.rev_ind; [destruct hd; cbn in *; congruence|]. - rewrite <- mkApps_nested in H0. - inv H0. - constructor; exists hd, (args ++ [N2]). - rewrite <- mkApps_nested. + split; eauto with pcuic. + - destruct (mkApps_elim hd args0). + apply mkApps_notApp_inj in H2 as (<-&<-); auto; [|now destruct (isApp f)]. + constructor; exists (mkApps (tFix x idx) (firstn n args)), (skipn n args). + rewrite mkApps_nested. + rewrite firstn_skipn. split; [easy|]. + split; eauto with pcuic. + - destruct (mkApps_elim hd args0). + apply mkApps_notApp_inj in H2 as (<-&<-); auto; [|now destruct (isApp f)]. + constructor; exists (mkApps (tFix x idx) (firstn n args)), (skipn n args). + rewrite mkApps_nested. + rewrite firstn_skipn. split; [easy|]. - apply All2_app; eauto with pcuic. - apply All2_same; eauto. - - apply mkApps_notApp_inj in H3 as (<-&<-); auto. - constructor; eexists _, _; split; [reflexivity|]. - split; [eauto|]. - eapply OnOne2_All2; eauto. - - apply mkApps_notApp_inj in H2 as (<-&<-); auto. - constructor; eexists _, _; split; [reflexivity|]. - split; [|now apply All2_same]. - eauto with pcuic. - - apply mkApps_notApp_inj in H2 as (<-&<-); auto. - constructor; eexists _, _; split; [reflexivity|]. - split; [|now apply All2_same]. - eauto with pcuic. + split; eauto with pcuic. +Qed. + +Lemma whnf_red1_from_mkApps Σ Γ hd args t : + whnf RedFlags.default Σ Γ (mkApps hd args) -> + red1 Σ Γ (mkApps hd args) t -> + ∥∑ hd' args', + t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. +Proof. + intros wh r. + depelim wh; try solve [solve_red_from_mkApps]. + - eapply whne_red1_from_mkApps; eauto. + - destruct (mkApps_elim hd args). + apply mkApps_notApp_inj in x as (<-&<-); auto; [|now destruct (isApp f)]. + apply red1_mkApps_tConstruct_inv in r as (?&->&?). + constructor; exists (mkApps (tConstruct i n u) (firstn n0 x)), (skipn n0 x). + rewrite mkApps_nested. + rewrite firstn_skipn. + split; [reflexivity|]. + eauto 9 with pcuic. + - destruct (mkApps_elim hd args). + apply mkApps_notApp_inj in x as (<-&<-); auto; [|now destruct (isApp f)]. + apply red1_mkApps_tInd_inv in r as (?&->&?). + constructor; exists (mkApps (tInd i u) (firstn n x)), (skipn n x). + rewrite mkApps_nested. + rewrite firstn_skipn. + split; [reflexivity|]. + split; eauto with pcuic. + - destruct (mkApps_elim hd args). + apply mkApps_notApp_inj in x as (<-&<-); auto; [|now destruct (isApp f)]. + apply red1_mkApps_tFix_inv in r. + 2: { destruct unfold_fix as [(?&?)|]; [|easy]. + now unfold is_constructor; rewrite H. } + destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]. + + constructor; exists (mkApps (tFix mfix idx) (firstn n x)), (skipn n x). + rewrite mkApps_nested firstn_skipn. + split; [easy|]. + eauto 9 with pcuic. + + constructor; exists (mkApps (tFix x idx) (firstn n v)), (skipn n v). + rewrite mkApps_nested firstn_skipn. + split; [easy|]. + split; eauto 9 with pcuic. + + constructor; exists (mkApps (tFix x idx) (firstn n v)), (skipn n v). + rewrite mkApps_nested firstn_skipn. + split; [easy|]. + split; eauto 9 with pcuic. + - destruct (mkApps_elim hd args). + apply mkApps_notApp_inj in x as (<-&<-); auto; [|now destruct (isApp f)]. + apply red1_mkApps_tCoFix_inv in r. + destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]. + + constructor; exists (mkApps (tCoFix mfix idx) (firstn n x)), (skipn n x). + rewrite mkApps_nested firstn_skipn. + split; [easy|]. + eauto 9 with pcuic. + + constructor; exists (mkApps (tCoFix x idx) (firstn n v)), (skipn n v). + rewrite mkApps_nested firstn_skipn. + split; [easy|]. + split; eauto 9 with pcuic. + + constructor; exists (mkApps (tCoFix x idx) (firstn n v)), (skipn n v). + rewrite mkApps_nested firstn_skipn. + split; [easy|]. + split; eauto 9 with pcuic. Qed. -Lemma whne_red_from_mkApps Σ Γ hd args t : - whne RedFlags.default Σ Γ hd -> - isApp hd = false -> +Lemma whnf_red_from_mkApps Σ Γ hd args t : + whnf RedFlags.default Σ Γ (mkApps hd args) -> red Σ Γ (mkApps hd args) t -> ∥∑ hd' args', t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. Proof. remember (mkApps hd args) as from eqn:fromeq. - intros wh notapp r. - revert hd notapp wh args fromeq. - induction r using red_rect_n1; intros hd notapp wh args ->. - - constructor; eexists _, _; split; [reflexivity|]. - split; [easy|]. - now apply All2_same. - - specialize (IHr _ notapp wh _ eq_refl) as [(?&?&->&?&?)]. - eapply whne_red1_from_mkApps in X as [(?&?&->&?&?)]; eauto. - 5: apply whne_mkApps; eapply whne_pres; eauto. - all: auto. - 2: erewrite whne_red_isApp; eauto. + intros wh r. + revert hd wh args fromeq. + induction r using red_rect_n1; intros hd wh args ->. + - constructor; eauto 10 with pcuic. + - specialize (IHr _ wh _ eq_refl) as [(?&?&->&?&?)]. + eapply whnf_red1_from_mkApps in X as [(?&?&->&?&?)]; eauto. + 2: eapply whnf_pres; eauto. constructor; eexists _, _; split; [reflexivity|]. split; [now etransitivity; eauto|]. eapply All2_trans; eauto. diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 8dbf25009..0e4b173fb 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -774,7 +774,7 @@ Section Conversion. isred_full Γ t π -> isLambda t -> isStackApp π = false. - Proof. + Proof. intros (?&isr) islam. destruct t; cbn in *; try easy. unfold zipp in isr. @@ -784,8 +784,49 @@ Section Conversion. apply whne_mkApps_inv in H0; [|easy]. destruct H0 as [|(?&?&?&?&?&?&?&?)]; [|discriminate]. depelim H0; solve_discr; discriminate. + Qed. + + Lemma decompose_stack_stack_cat π π' : + decompose_stack (π +++ π') = + ((decompose_stack π).1 ++ + match (decompose_stack π).2 with + | ε => (decompose_stack π').1 + | _ => [] + end, + (decompose_stack π).2 +++ + match (decompose_stack π).2 with + | ε => (decompose_stack π').2 + | _ => π' + end). + Proof. + induction π in π' |- *; cbn in *; auto. + - now destruct decompose_stack. + - rewrite !IHπ. + now destruct (decompose_stack π). Qed. + Lemma zipp_stack_cat π π' t : + isStackApp π = false -> + zipp t (π' +++ π) = zipp t π'. + Proof. + intros no_stack_app. + unfold zipp. + rewrite decompose_stack_stack_cat. + destruct (decompose_stack π') eqn:decomp. + cbn. + destruct s; try now rewrite app_nil_r. + now destruct π; cbn in *; rewrite ?app_nil_r. + Qed. + + Lemma zipp_appstack t args π : + zipp t (appstack args π) = zipp (mkApps t args) π. + Proof. + unfold zipp. + rewrite decompose_stack_appstack. + rewrite <- mkApps_nested. + now destruct decompose_stack. + Qed. + Definition Ret s Γ t π t' π' := forall (leq : conv_pb), conv_stack_ctx Γ π π' -> @@ -1050,29 +1091,19 @@ Section Conversion. pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as r1; pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w1 end. - split. - - rewrite <- eq1 in r1. destruct r1 as [ha hl]. - assumption. - - rewrite <- eq1 in w1. - apply (f_equal (snd ∘ decompose_stack ∘ snd)) in eq1. - rewrite reduce_stack_decompose, decompose_stack_appstack in eq1. - rewrite stack_context_stack_cat. - cbn in *. - rewrite <- stack_context_decompose in w1. - rewrite <- (stack_context_decompose π1') in w1 |- *. - rewrite eq1 in w1 |- *. - rewrite <- e1 in w1. - cbn in *. - unfold zipp. - assert (decompose_stack π1' = ((decompose_stack π1').1, (decompose_stack π1').2)) - by (now destruct (decompose_stack π1')). - rewrite eq1 in H. - apply decompose_stack_eq in H as ->. - rewrite stack_cat_appstack, decompose_stack_appstack. - erewrite (decompose_stack_twice _ _ ρ1); eauto. - cbn; rewrite app_nil_r. - unfold zipp in w1. - now destruct (decompose_stack π1'). + rewrite <- eq1 in r1, w1. + destruct r1 as (ha&hl). + split; [easy|]. + rewrite stack_context_stack_cat. + apply (f_equal (snd ∘ decompose_stack ∘ snd)) in eq1. + rewrite reduce_stack_decompose, decompose_stack_appstack in eq1. + cbn in eq1. + rewrite <- (stack_context_decompose π1'), eq1 in *. + rewrite <- (stack_context_decompose π1), <- e1 in w1. + cbn in w1 |- *. + rewrite zipp_stack_cat. + 2: eapply decompose_stack_noStackApp; eauto. + auto. Qed. Next Obligation. match type of eq2 with @@ -1080,29 +1111,19 @@ Section Conversion. pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as r2; pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w2 end. - split. - - rewrite <- eq2 in r2. destruct r2 as [ha hl]. - assumption. - - rewrite <- eq2 in w2. - apply (f_equal (snd ∘ decompose_stack ∘ snd)) in eq2. - rewrite reduce_stack_decompose, decompose_stack_appstack in eq2. - rewrite stack_context_stack_cat. - cbn in *. - rewrite <- stack_context_decompose in w2. - rewrite <- (stack_context_decompose π2') in w2 |- *. - rewrite eq2 in w2 |- *. - rewrite <- e2 in w2. - cbn in *. - unfold zipp. - assert (decompose_stack π2' = ((decompose_stack π2').1, (decompose_stack π2').2)) - by (now destruct (decompose_stack π2')). - rewrite eq2 in H. - apply decompose_stack_eq in H as ->. - rewrite stack_cat_appstack, decompose_stack_appstack. - erewrite (decompose_stack_twice _ _ ρ2); eauto. - cbn; rewrite app_nil_r. - unfold zipp in w2. - now destruct (decompose_stack π2'). + rewrite <- eq2 in r2, w2. + destruct r2 as (ha&hl). + split; [easy|]. + rewrite stack_context_stack_cat. + apply (f_equal (snd ∘ decompose_stack ∘ snd)) in eq2. + rewrite reduce_stack_decompose, decompose_stack_appstack in eq2. + cbn in eq2. + rewrite <- (stack_context_decompose π2'), eq2 in *. + rewrite <- (stack_context_decompose π2), <- e2 in w2. + cbn in w2 |- *. + rewrite zipp_stack_cat. + 2: eapply decompose_stack_noStackApp; eauto. + auto. Qed. Next Obligation. destruct hΣ as [wΣ]. @@ -1245,8 +1266,6 @@ Section Conversion. * eapply conv_context_sym. all: auto. Qed. - - Opaque reduce_stack. Equations unfold_one_fix (Γ : context) (mfix : mfixpoint term) (idx : nat) (π : stack) (h : wtp Γ (tFix mfix idx) π) @@ -1481,6 +1500,82 @@ Section Conversion. reflexivity. Qed. + Lemma zipp_as_mkApps t π : + zipp t π = mkApps t (decompose_stack π).1. + Proof. + unfold zipp. + now destruct decompose_stack. + Qed. + + Lemma unfold_one_fix_None Γ mfix idx π wf : + None = unfold_one_fix Γ mfix idx π wf -> + ∥∑args, + All2 (red Σ (Γ,,, stack_context π)) (decompose_stack π).1 args × + whnf RedFlags.default Σ (Γ,,, stack_context π) (mkApps (tFix mfix idx) args)∥. + Proof. + funelim (unfold_one_fix Γ mfix idx π wf). + all: intros [=]. + - constructor; eexists _; split; [apply All2_same; reflexivity|]. + eapply whnf_fixapp. + now rewrite <- e. + - constructor; eexists _; split; [apply All2_same; reflexivity|]. + eapply whnf_fixapp. + rewrite <- e. + destruct (decompose_stack π) eqn:decomp. + apply decompose_stack_noStackApp in decomp as ?. + apply decompose_stack_eq in decomp as ->. + clear H. + symmetry in e0. + now apply decompose_stack_at_appstack_None in e0. + - match type of e1 with + | _ = reduce_stack ?a ?b ?c ?d ?e ?f ?g => + pose proof (reduce_stack_sound a b c d e f g) as [r]; + pose proof (reduce_stack_whnf a b c d e f g) as wh; + pose proof (reduce_stack_decompose a b c d e f g) as decomp; + pose proof (reduce_stack_isred a b c d e f g) as isr + end. + rewrite <- e1 in r, wh, decomp, isr. + specialize (isr eq_refl) as (noapp&_). + cbn in *. + clear H e1 H0. + symmetry in e0. + apply decompose_stack_at_length in e0 as ?. + apply decompose_stack_at_eq in e0 as ->. + rewrite stack_context_appstack. + cbn. + apply wellformed_zipc_zipp in h; auto. + rewrite <- (stack_context_decompose s0), decomp in wh. + change (App t0 s) with (appstack [t0] s) in *. + rewrite !decompose_stack_appstack. + rewrite zipp_as_mkApps, !decompose_stack_appstack in h. + destruct h as [(ty&typ)|[(?&?&dar&?)]]; cycle 1. + { cbn in dar. rewrite destArity_tFix in dar; congruence. } + cbn in *. + rewrite stack_context_appstack in typ. + cbn in *. + destruct (decompose_stack s0) eqn:decomp'. + apply decompose_stack_eq in decomp' as ->. + rewrite zipc_appstack in r. + rewrite zipp_appstack in wh. + cbn in *; subst; cbn in *. + destruct hΣ. + constructor; exists (l ++ (mkApps t2 l0) :: (decompose_stack s).1). + eapply PCUICSR.subject_reduction in typ. + 2: eauto. + 2: apply red_mkApps; [reflexivity|]. + 1: split. + 1,3: apply All2_app; [apply All2_same; reflexivity|]. + 1,2: constructor; [|apply All2_same; reflexivity]. + 1-2: eassumption. + apply whnf_ne. + econstructor. + + eauto. + + rewrite nth_error_snoc by easy. + eauto. + + eapply whnf_fix_arg_whne; eauto. + now destruct t2. + Qed. + (* Tailored view for isconv_prog *) Equations prog_discr (t1 t2 : term) : Prop := prog_discr (tApp _ _) (tApp _ _) := False ; @@ -1637,13 +1732,6 @@ Section Conversion. now apply LevelSet.mem_spec in y_mem. Qed. - Lemma zipp_as_mkApps t π : - zipp t π = mkApps t (decompose_stack π).1. - Proof. - unfold zipp. - now destruct decompose_stack. - Qed. - Lemma wellformed_nonarity Γ t : destArity [] t = None -> wellformed Σ Γ t -> @@ -2506,12 +2594,14 @@ Section Conversion. - depelim whn; solve_discr; easy. Qed. - Lemma whne_conv_context f Γ Γ' t : +(* TODO: move to safe emmata *) + Lemma whne_context_relation f rel Γ Γ' t : + (forall Γ Γ' c c', rel Γ Γ' c c' -> (decl_body c = None <-> decl_body c' = None)) -> whne f Σ Γ t -> - conv_context Σ Γ Γ' -> + context_relation rel Γ Γ' -> whne f Σ Γ' t. Proof. - intros wh conv. + intros behaves wh conv. induction wh; eauto using whne. destruct nth_error eqn:nth; [|easy]. cbn in H. @@ -2520,10 +2610,43 @@ Section Conversion. constructor. rewrite e. cbn. - depelim c1; auto. - cbn in *. + specialize (behaves _ _ _ _ r). + f_equal. + apply behaves. congruence. Qed. + + Lemma whnf_context_relation f rel Γ Γ' t : + (forall Γ Γ' c c', rel Γ Γ' c c' -> (decl_body c = None <-> decl_body c' = None)) -> + whnf f Σ Γ t -> + context_relation rel Γ Γ' -> + whnf f Σ Γ' t. + Proof. + intros behaves wh conv. + destruct wh; eauto using whnf. + apply whnf_ne. + eapply whne_context_relation; eauto. + Qed. + + Lemma whne_conv_context f Γ Γ' t : + whne f Σ Γ t -> + conv_context Σ Γ Γ' -> + whne f Σ Γ' t. + Proof. + apply whne_context_relation. + intros ? ? ? ? r. + now depelim r. + Qed. + + Lemma whnf_conv_context f Γ Γ' t : + whnf f Σ Γ t -> + conv_context Σ Γ Γ' -> + whnf f Σ Γ' t. + Proof. + apply whnf_context_relation. + intros ? ? ? ? r. + now depelim r. + Qed. (* Notation whne := (whne flags Σ). @@ -2709,7 +2832,7 @@ Section Conversion. isconv_prog leq fn' (ρ +++ θ') (tFix mfix' idx') π2 aux } } ; - | _ with inspect (unfold_one_fix Γ mfix' idx' π2 _) := { + | @exist None unfold_fix1 with inspect (unfold_one_fix Γ mfix' idx' π2 _) := { | @exist (Some (fn, θ)) eq1 with inspect (decompose_stack θ) := { | @exist (l', θ') eq2 @@ -2718,7 +2841,7 @@ Section Conversion. isconv_prog leq (tFix mfix idx) π1 fn' (ρ +++ θ') aux } } ; - | _ with inspect (eqb idx idx') := { + | @exist None unfold_fix2 with inspect (eqb idx idx') := { | @exist true eq4 with isconv_fix Γ mfix idx π1 _ mfix' idx' π2 _ _ _ aux := { | Success h1 with isconv_args_raw leq (tFix mfix idx) π1 (tFix mfix' idx') π2 aux := { | Success h2 := yes ; @@ -2726,7 +2849,7 @@ Section Conversion. } ; | Error e h := Error e _ } ; - | @exist false _ := + | @exist false idx_uneq := Error ( CannotUnfoldFix (Γ ,,, stack_context π1) mfix idx @@ -3095,6 +3218,7 @@ Section Conversion. { eapply whne_conv_context; eauto. apply conv_context_sym; eauto. } apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. + 2-3: apply whnf_mkApps; eauto using whne. apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. now constructor. Qed. @@ -3122,6 +3246,7 @@ Section Conversion. { eapply whne_conv_context; eauto. apply conv_context_sym; eauto. } apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. + 2-3: apply whnf_mkApps; eauto using whne. apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. now constructor. Qed. @@ -3149,6 +3274,7 @@ Section Conversion. { eapply whne_conv_context; eauto. apply conv_context_sym; eauto. } apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. + 2-3: apply whnf_mkApps; eauto using whne. apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. now constructor. Qed. @@ -3183,6 +3309,7 @@ Section Conversion. { eapply whne_conv_context; eauto. apply conv_context_sym; eauto. } apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. + 2-3: apply whnf_mkApps; eauto using whne. apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. rewrite eq_inductive_refl, Nat.eqb_refl in eqDiff. congruence. @@ -3493,29 +3620,16 @@ Section Conversion. pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as ir; pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w2 end. - split. - - rewrite <- eq3 in ir. destruct ir as [? hl]. - assumption. - - rewrite <- eq3 in w2, d2. - rewrite decompose_stack_appstack in d2. - rewrite stack_context_stack_cat. - cbn in *. - rewrite <- stack_context_decompose in w2. - rewrite <- (stack_context_decompose ρ) in w2 |- *. - rewrite d2 in w2 |- *. - erewrite decompose_stack_twice in w2 by eauto. - cbn in *. - assert (decompose_stack ρ = ((decompose_stack ρ).1, (decompose_stack ρ).2)) - by (now destruct (decompose_stack ρ)). - rewrite d2 in H. - apply decompose_stack_eq in H as ->. - unfold zipp. - rewrite stack_cat_appstack, decompose_stack_appstack. - cbn. - erewrite (decompose_stack_twice θ l' θ') by auto. - cbn; rewrite app_nil_r. - unfold zipp in w2. - now destruct (decompose_stack ρ). + rewrite <- eq3 in ir. + destruct ir as (?&hl). + split; [easy|]. + rewrite <- eq3 in w2, d2. + rewrite decompose_stack_appstack in d2. + cbn in *. + rewrite zipp_stack_cat. + 2: eapply decompose_stack_noStackApp; eauto. + rewrite stack_context_stack_cat. + now rewrite app_context_assoc. Qed. Next Obligation. destruct hΣ as [wΣ]. @@ -3692,21 +3806,18 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; - pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as ir + pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as ir; + pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w2 end. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. - rewrite <- eq3 in ir. destruct ir as [? hl]. - split. - - assumption. - - rewrite stack_context_stack_cat. - simpl. intro h. specialize (hl h). cbn in hl. - case_eq (decompose_stack ρ). intros l π e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. - apply isStackApp_false_appstack in hl. subst. cbn. - eapply decompose_stack_noStackApp. symmetry. eassumption. + rewrite <- eq3 in d2, ir, w2. + rewrite decompose_stack_appstack in d2. + cbn in *. + destruct ir as (?&hl). + split; [easy|]. + rewrite stack_context_stack_cat. + rewrite zipp_stack_cat. + 2: eapply decompose_stack_noStackApp; eauto. + now rewrite app_context_assoc. Qed. Next Obligation. destruct hΣ as [wΣ]. @@ -3765,8 +3876,55 @@ Section Conversion. Qed. Next Obligation. apply h; clear h. - (* Idea: a fixpoint should be convertible to its unfolding + *) - todo "Completeness"%string. + destruct hΣ as [wΣ]. + apply unfold_one_fix_red_zipp in eq1 as r1. + destruct r1 as [r1]. + match type of eq3 with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; + pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + end. + rewrite <- eq3 in r2. + rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. + cbn in d2. + apply unfold_one_fix_decompose in eq1 as d1. + assert (r2' : ∥ red (fst Σ) (Γ ,,, stack_context π2) (zipp fn θ) (zipp fn' (ρ +++ θ')) ∥). + { unfold zipp. + destruct hx as [hx]. + constructor. + case_eq (decompose_stack ρ). intros l ξ e. + rewrite e in d2. cbn in d2. subst. + pose proof (decompose_stack_eq _ _ _ e). subst. + rewrite stack_cat_appstack. rewrite decompose_stack_appstack. + rewrite <- eq2. + cbn in r2. rewrite 2!zipc_appstack in r2. cbn in r2. + rewrite <- eq2 in d1. cbn in d1. subst. + case_eq (decompose_stack π2). intros l2 ρ2 e2. + simpl. rewrite e2 in r2. simpl in r2. + pose proof (decompose_stack_eq _ _ _ e2). subst. + rewrite decompose_stack_twice with (1 := e2). cbn. + rewrite app_nil_r. + rewrite stack_context_appstack. + assumption. + } + assert (e : stack_context π2 = stack_context (ρ +++ θ')). + { case_eq (decompose_stack ρ). intros l ξ e. + rewrite e in d2. cbn in d2. subst. + pose proof (decompose_stack_eq _ _ _ e). subst. + rewrite stack_cat_appstack. + rewrite stack_context_appstack. + rewrite <- eq2 in d1. cbn in d1. subst. + case_eq (decompose_stack π2). intros l2 ρ2 e2. simpl. + pose proof (decompose_stack_eq _ _ _ e2). subst. + rewrite stack_context_appstack. reflexivity. + } + destruct r2' as [r2']. + destruct hx as [hx]. + pose proof (red_trans _ _ _ _ _ r1 r2') as r. + eapply conv_cum_conv_ctx; eauto. + 2: apply conv_context_sym; eauto. + eapply conv_cum_red_inv; [exact hΣ'| |reflexivity|exact r]. + eapply conv_cum_conv_ctx; eauto. Qed. Next Obligation. change (true = eqb idx idx') in eq4. @@ -3790,15 +3948,55 @@ Section Conversion. Qed. Next Obligation. apply h; clear h. - (* Idea : The fixpoints are neutral (do not unfold) *) - (* so each definitions are convertible *) - todo "Completeness"%string. + (* Idea : Since the fixpoints do not unfold they are convertible to fixpoints + in whnf implying that their definitions are convertible. *) + rewrite !zipp_as_mkApps in H. + apply unfold_one_fix_None in unfold_fix1 as [(?&?&?)]. + apply unfold_one_fix_None in unfold_fix2 as [(?&?&?)]. + destruct hΣ, hx. + eapply conv_cum_conv_ctx in H; eauto. + eapply conv_cum_red_inv in H. + 2: exact hΣ'. + 2: reflexivity. + 2: apply red_mkApps; [reflexivity|exact a0]. + eapply conv_cum_conv_ctx in H; eauto. + 2: apply conv_context_sym; eauto. + eapply conv_cum_red_inv in H. + 2: exact hΣ'. + 2: apply red_mkApps; [reflexivity|exact a]. + 2: reflexivity. + apply conv_cum_mkApps_inv in H as [(?&?)]; auto. + 2: eapply whnf_conv_context; eauto. + 2: eapply conv_context_sym; eauto. + apply conv_cum_tFix_inv in c as [(->&?)]. + constructor. + eapply All2_impl; eauto. + cbn; intros; easy. Qed. Next Obligation. - (* Proof idea: reduce to the case where we have eq_term - between tFix using that the fixpoints do not unfold - and then conlude with eqDiff *) - todo "Completeness"%string. + (* Idea : Since the fixpoints do not unfold they are convertible to fixpoints + in whnf implying that the indices have to be equal. *) + rewrite !zipp_as_mkApps in H. + apply unfold_one_fix_None in unfold_fix1 as [(?&?&?)]. + apply unfold_one_fix_None in unfold_fix2 as [(?&?&?)]. + destruct hΣ, hx. + eapply conv_cum_conv_ctx in H; eauto. + eapply conv_cum_red_inv in H. + 2: exact hΣ'. + 2: reflexivity. + 2: apply red_mkApps; [reflexivity|exact a0]. + eapply conv_cum_conv_ctx in H; eauto. + 2: apply conv_context_sym; eauto. + eapply conv_cum_red_inv in H. + 2: exact hΣ'. + 2: apply red_mkApps; [reflexivity|exact a]. + 2: reflexivity. + apply conv_cum_mkApps_inv in H as [(?&?)]; auto. + 2: eapply whnf_conv_context; eauto. + 2: eapply conv_context_sym; eauto. + apply conv_cum_tFix_inv in c as [(->&?)]. + rewrite Nat.eqb_refl in idx_uneq. + congruence. Qed. (* tCoFix *) diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index 112b2f438..a5f796762 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -1292,6 +1292,38 @@ Section Reduce. - unfold isCoFix_app in cof. now rewrite decompose_app_mkApps in cof. Qed. + + Lemma whnf_fix_arg_whne mfix idx body Γ t before args aftr ty : + unfold_fix mfix idx = Some (#|before|, body) -> + match t with + | tConstruct _ _ _ => False + | tApp _ _ => False + | _ => True + end -> + whnf flags Σ Γ (mkApps t args) -> + Σ;;; Γ |- mkApps (tFix mfix idx) (before ++ mkApps t args :: aftr) : ty -> + whne flags Σ Γ (mkApps t args). + Proof. + destruct hΣ. + intros uf shape wh typ. + apply inversion_mkApps in typ as (fix_ty & typ_fix & typ_args); auto. + apply inversion_Fix in typ_fix as (def&?&?&?&?&?&?); auto. + eapply All_nth_error in a; eauto. + eapply wf_fixpoint_spine in i0; eauto. + 2: { eapply PCUICSpine.typing_spine_strengthen; eauto. } + unfold unfold_fix in uf. + rewrite e in uf. + rewrite nth_error_snoc in i0 by congruence. + destruct i0 as (?&?&?&typ&fin). + eapply whnf_non_ctor_finite_ind_typed; try eassumption. + - unfold isConstruct_app. + rewrite decompose_app_mkApps by (now destruct t). + cbn. + now destruct t. + - destruct check_recursivity_kind eqn:cofin in |- *; [|easy]. + eapply check_recursivity_kind_inj in fin; [|exact cofin]. + congruence. + Qed. Lemma reduce_stack_whnf : forall Γ t π h, @@ -1490,24 +1522,8 @@ Section Reduce. eapply whne_fixapp. + eassumption. + now apply nth_error_snoc. - + destruct hΣ. - apply inversion_mkApps in typ as (fix_ty & typ_fix & typ_args); auto. - apply inversion_Fix in typ_fix as (def&?&?&?&?&?&?); auto. - eapply All_nth_error in a; eauto. - eapply wf_fixpoint_spine in i0; eauto. - 2: { eapply PCUICSpine.typing_spine_strengthen; eauto. } - unfold unfold_fix in e. - rewrite e0 in e. - rewrite nth_error_snoc in i0 by congruence. - destruct i0 as (?&?&?&typ&fin). - eapply whnf_non_ctor_finite_ind_typed; try eassumption. - * unfold isConstruct_app. - rewrite decompose_app_mkApps by (now rewrite noapp). - cbn. - now destruct t2. - * destruct check_recursivity_kind eqn:cofin in |- *; [|easy]. - eapply check_recursivity_kind_inj in fin; [|exact cofin]. - congruence. + + eapply whnf_fix_arg_whne; eauto. + now destruct t2. - unfold zipp. destruct decompose_stack. From bac7ad3836c7054800497ae9a3e354def1416c6b Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Thu, 22 Oct 2020 17:11:46 +0200 Subject: [PATCH 09/26] Fix build error --- safechecker/theories/PCUICSafeChecker.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index 8f6cb32cb..9c09db36b 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -1070,7 +1070,7 @@ Section Typecheck. Next Obligation. symmetry in Heq_anonymous2. - unfold iscumul in Heq_anonymous2. simpl in Heq_anonymous2. destruct wildcard'. + unfold iscumul in Heq_anonymous2. simpl in Heq_anonymous2. apply isconv_term_sound in Heq_anonymous2. red in Heq_anonymous2. noconf Heq_I''. @@ -1113,7 +1113,7 @@ Section Typecheck. Next Obligation. intros. clearbody btyswf. idtac; Program.Tactics.program_simplify. symmetry in Heq_anonymous2. - unfold iscumul in Heq_anonymous2. simpl in Heq_anonymous2. destruct wildcard'. + unfold iscumul in Heq_anonymous2. simpl in Heq_anonymous2. apply isconv_term_sound in Heq_anonymous2. noconf Heq_I''. noconf Heq_I'. noconf Heq_I. noconf Heq_d. noconf Heq_d'. simpl in *. From 891d8370633025e428e0eab895c21f6188046994 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Thu, 22 Oct 2020 18:13:23 +0200 Subject: [PATCH 10/26] Clean up erasure to use context_change_decl_types_red --- erasure/theories/EArities.v | 114 ------------------------ erasure/theories/ErasureFunction.v | 24 ++--- erasure/theories/SafeErasureFunction.v | 34 ++----- pcuic/theories/PCUICContextConversion.v | 1 - 4 files changed, 15 insertions(+), 158 deletions(-) diff --git a/erasure/theories/EArities.v b/erasure/theories/EArities.v index 3f91dcd16..888684215 100644 --- a/erasure/theories/EArities.v +++ b/erasure/theories/EArities.v @@ -310,120 +310,6 @@ Proof. - eexists _, _; intuition eauto. Qed. -Inductive conv_decls (Σ : global_env_ext) Γ (Γ' : context) : forall (x y : context_decl), Type := -| conv_vass : forall (na na' : name) (T T' : term), - (* isWfArity_or_Type Σ Γ' T' -> *) - Σ;;; Γ |- T = T' -> conv_decls Σ Γ Γ' (vass na T) (vass na' T') -| conv_vdef_type : forall (na : name) (b T : term), - (* isWfArity_or_Type Σ Γ' T -> *) - conv_decls Σ Γ Γ' (vdef na b T) (vdef na b T). - -Lemma conv_context_refl (Σ : global_env_ext) Γ : - wf Σ -> wf_local Σ Γ -> - context_relation (@conv_decls Σ) Γ Γ. -Proof. - induction Γ; try econstructor. - intros wfΣ wfΓ; depelim wfΓ; econstructor; eauto; - constructor; auto. -Qed. - -Lemma context_conversion_red1 (Σ : global_env_ext) Γ Γ' s t : wf Σ -> (* Σ ;;; Γ' |- t : T -> *) - context_relation (@conv_decls Σ) Γ Γ' -> red1 Σ Γ s t -> red Σ Γ' s t. -Proof. - intros HΣ HT X0. induction X0 using red1_ind_all in Γ', HΣ, HT |- *; eauto. - all:pcuic. - - econstructor. econstructor. - rewrite <- H. - induction HT in i |- *; destruct i; eauto. - now inv p. - - - eapply PCUICReduction.red_abs. eapply IHX0; eauto. eauto. - - - eapply PCUICReduction.red_abs. eauto. eapply IHX0. eauto. - eauto. econstructor. eauto. econstructor. - eapply PCUICCumulativity.conv_refl'. - - - eapply PCUICReduction.red_letin. eapply IHX0; eauto. - all:eauto. - - - eapply PCUICReduction.red_letin; eauto. - - - eapply PCUICReduction.red_letin; eauto. eapply IHX0; eauto. - econstructor. eauto. econstructor. - - eapply PCUICReduction.red_case; eauto. clear. - eapply All_All2_refl. induction brs; eauto. - - eapply PCUICReduction.red_case; eauto. clear. - eapply All_All2_refl. induction brs; eauto. - - destruct a. - eapply red_case; eauto. - clear - HΣ X HT. - induction X. - + econstructor. destruct p. destruct p. - split; eauto. - eapply All_All2_refl. - induction tl; eauto. - + econstructor. now split. - eassumption. - - - eapply PCUICReduction.red_proj_c. eauto. - - - eapply PCUICReduction.red_app; eauto. - - eapply PCUICReduction.red_app; eauto. - - - eapply PCUICReduction.red_prod; eauto. - - - eapply PCUICReduction.red_prod; eauto. eapply IHX0. eauto. eauto. - econstructor. - eauto. econstructor. eapply PCUICCumulativity.conv_refl'. - - eapply PCUICReduction.red_evar; eauto. - induction X; eauto. econstructor. eapply p; eauto. - induction tl; eauto. - - eapply PCUICReduction.red_fix_one_ty. - eapply OnOne2_impl ; eauto. - intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. - inversion e. subst. clear e. - split ; auto. - - eapply PCUICReduction.red_fix_one_body. - eapply OnOne2_impl ; eauto. - intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. - inversion e. subst. clear e. - split ; auto. - eapply ih ; auto. - clear - HT. - induction (fix_context mfix0) as [| [na [b|] ty] Δ ihΔ]. - + auto. - + simpl. constructor ; eauto. - constructor. - + simpl. constructor ; eauto. - constructor. apply PCUICCumulativity.conv_refl'. - - eapply PCUICReduction.red_cofix_one_ty. - eapply OnOne2_impl ; eauto. - intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. - inversion e. subst. clear e. - split ; auto. - - eapply PCUICReduction.red_cofix_one_body. - eapply OnOne2_impl ; eauto. - intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. - inversion e. subst. clear e. - split ; auto. - eapply ih ; auto. - clear - HT. - induction (fix_context mfix0) as [| [na [b|] ty] Δ ihΔ]. - + auto. - + simpl. constructor ; eauto. - constructor. - + simpl. constructor ; eauto. - constructor. apply PCUICCumulativity.conv_refl'. -Qed. - -Lemma context_conversion_red (Σ : global_env_ext) Γ Γ' s t : wf Σ -> - context_relation (@conv_decls Σ) Γ Γ' -> red Σ Γ s t -> red Σ Γ' s t. -Proof. - intros. induction X1 using red_rect'; eauto. - etransitivity. eapply IHX1. - eapply context_conversion_red1; eauto. -Qed. - Lemma isWfArity_or_Type_red: forall (Σ : global_env_ext) (Γ : context) (T : term), wf Σ -> wf_local Σ Γ -> isWfArity_or_Type Σ Γ T -> forall x5 : term, red Σ Γ T x5 -> isWfArity_or_Type Σ Γ x5. diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 4c50192d9..06f932f09 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -221,23 +221,13 @@ Next Obligation. eapply invert_red_prod in X2 as (? & ? & [] & ?); eauto. subst. cbn in *. exists x4; split; eauto. - destruct HT as [ [] | [] ]. - ++ sq. etransitivity; eauto. - eapply context_conversion_red; eauto. econstructor. - - eapply conv_context_refl; eauto. econstructor. - - eapply conv_sym, red_conv; eauto. - - ++ sq. etransitivity. eassumption. - - eapply context_conversion_red; eauto. econstructor. - - eapply conv_context_refl; eauto. - - econstructor. - - eapply conv_sym, red_conv; eauto. + constructor. + etransitivity; eauto. + eapply PCUICContextConversion.context_change_decl_types_red; eauto. + constructor; [|constructor]. + eapply PCUICContextConversion.context_relation_refl. + intros. + destruct x0 as [? [|]? ]; constructor. Qed. Hint Constructors squash : core. diff --git a/erasure/theories/SafeErasureFunction.v b/erasure/theories/SafeErasureFunction.v index b02807361..f03ad80bc 100644 --- a/erasure/theories/SafeErasureFunction.v +++ b/erasure/theories/SafeErasureFunction.v @@ -185,32 +185,14 @@ Section fix_sigma. eapply invert_red_prod in X2 as (? & ? & [] & ?); eauto. subst. cbn in *. exists x4; split; eauto. - - destruct HT as [ [] | [] ]. - ++ sq. pose proof t. pose proof t. - - eapply subject_reduction in X3. 2:eauto. 2:{ etransitivity. exact X. exact r0. } - eapply inversion_Prod in X3 as (? & ? & ? & ? & ?) ; auto. - - eapply subject_reduction in X2. 2:eauto. 2:{ exact X0. } - eapply inversion_Prod in X2 as (? & ? & ? & ? & ?) ; auto. - - etransitivity. eassumption. - - eapply context_conversion_red; eauto. econstructor. - - eapply conv_context_refl; eauto. econstructor. - - eapply conv_sym, red_conv; eauto. - ++ sq. etransitivity. eassumption. - - eapply context_conversion_red; eauto. econstructor. - - eapply conv_context_refl; eauto. - - econstructor. - - eapply conv_sym, red_conv; eauto. + + constructor. + etransitivity; eauto. + eapply PCUICContextConversion.context_change_decl_types_red; eauto. + constructor; [|constructor]. + eapply PCUICContextConversion.context_relation_refl. + intros. + destruct x0 as [? [|]? ]; constructor. Qed. Next Obligation. diff --git a/pcuic/theories/PCUICContextConversion.v b/pcuic/theories/PCUICContextConversion.v index 882ba6add..70a512a33 100644 --- a/pcuic/theories/PCUICContextConversion.v +++ b/pcuic/theories/PCUICContextConversion.v @@ -741,7 +741,6 @@ Section ContextChangeTypesReduction. constructor. Qed. - (* todo: update erasure to use this *) Lemma context_change_decl_types_red Γ Γ' s t : context_relation (fun _ _ => change_decl_type) Γ Γ' -> red Σ Γ s t -> red Σ Γ' s t. Proof. From de4745fd04b759129c03381d0f29b95eda30e4c3 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Thu, 22 Oct 2020 19:06:45 +0200 Subject: [PATCH 11/26] Finish tProj case --- pcuic/theories/PCUICConversionInversion.v | 20 +++++ pcuic/theories/PCUICNormal.v | 34 ++++++++ safechecker/theories/PCUICSafeConversion.v | 93 ++++++++++++++++++++-- 3 files changed, 142 insertions(+), 5 deletions(-) diff --git a/pcuic/theories/PCUICConversionInversion.v b/pcuic/theories/PCUICConversionInversion.v index 5eb93396b..04390abd4 100644 --- a/pcuic/theories/PCUICConversionInversion.v +++ b/pcuic/theories/PCUICConversionInversion.v @@ -362,4 +362,24 @@ Section fixed. auto. Qed. + Lemma conv_cum_tProj_inv leq Γ p c p' c' : + conv_cum leq Σ Γ (tProj p c) (tProj p' c') -> + whnf RedFlags.default Σ Γ (tProj p c) -> + whnf RedFlags.default Σ Γ (tProj p' c') -> + ∥ p = p' × Σ;;; Γ |- c = c' ∥. + Proof. + intros conv whl whr. + depelim whl; solve_discr. + depelim H; solve_discr; try discriminate. + depelim whr; solve_discr. + depelim H0; solve_discr; try discriminate. + apply conv_cum_alt in conv as [(?&?&(r1&r2)&?)]. + apply whne_red_from_tProj in r1 as [(?&->&?)]; auto. + apply whne_red_from_tProj in r2 as [(?&->&?)]; auto. + depelim e. + constructor. + split; [easy|]. + now apply conv_alt_red; exists x1, x. + Qed. + End fixed. diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index d36a8a8b0..1e6034681 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -1346,6 +1346,40 @@ Proof. + eapply whne_pres; eauto. Qed. +Lemma whne_red1_from_tProj Σ Γ p c t : + whne RedFlags.default Σ Γ (tProj p c) -> + red1 Σ Γ (tProj p c) t -> + ∥∑ c', t = tProj p c' × red Σ Γ c c'∥. +Proof. + intros wh r. + remember (tProj p c) as from eqn:fromeq. + revert c fromeq. + apply (whne_red1_ind + RedFlags.default Σ Γ + (fun t t' => (forall c, t = _ -> _ : Prop))) + with (t0 := from) (t' := t); intros; cbn in *; try discriminate; solve_discr; auto. + inv H1. + constructor; eexists _; split; [reflexivity|]. + eauto. +Qed. + +Lemma whne_red_from_tProj Σ Γ p c t : + whne RedFlags.default Σ Γ (tProj p c) -> + red Σ Γ (tProj p c) t -> + ∥∑ c', t = tProj p c' × red Σ Γ c c'∥. +Proof. + intros wh r. + remember (tProj p c) eqn:fromeq. + revert fromeq. + induction r using red_rect_n1; intros ->. + - constructor; eauto 9 with pcuic. + - destruct (IHr eq_refl) as [(?&->&?)]. + eapply whne_red1_from_tProj in X as [(?&->&?)]. + + constructor; eexists _; split; [reflexivity|]. + etransitivity; eauto. + + eapply whne_pres; eauto. +Qed. + (* Definition head_arg_is_constructor mfix idx args := diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index e3c0818b3..4310392e3 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -774,7 +774,7 @@ Section Conversion. isred_full Γ t π -> isLambda t -> isStackApp π = false. - Proof. + Proof. intros (?&isr) islam. destruct t; cbn in *; try easy. unfold zipp in isr. @@ -3527,11 +3527,67 @@ Section Conversion. eapply conv_Proj_c. assumption. Qed. Next Obligation. - todo "Completeness". + apply h; cbn; clear h. + destruct ir1 as (_&wh1), ir2 as (_&wh2). + cbn in *. + rename eq3 into c_is_red. + rename eq4 into c'_is_red. + symmetry in c'_is_red. + apply eqb_term_spec in c'_is_red. + eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. + symmetry in c_is_red. + apply eqb_term_spec in c_is_red. + eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. + rewrite !zipp_as_mkApps in H. + rewrite zipp_as_mkApps in wh1, wh2. + apply whnf_mkApps_inv in wh1; [|easy]. + apply whnf_mkApps_inv in wh2; [|easy]. + depelim wh1; solve_discr. + depelim wh2; solve_discr. + depelim H; cbn in *; try easy; solve_discr. + depelim H0; cbn in *; try easy; solve_discr. + apply whnf_whne_upgrade in c_is_red; auto. + apply whnf_whne_upgrade in c'_is_red; auto. + destruct hΣ, hx. + assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). + { eapply whne_conv_context; eauto. + apply conv_context_sym; eauto. } + apply conv_cum_mkApps_inv in H1 as [(conv_proj&conv_args)]; auto. + 2-3: apply whnf_mkApps; eauto using whne. + apply conv_cum_tProj_inv in conv_proj as [(->&?)]; auto. + constructor; auto. Qed. Next Obligation. (* Proof idea: c and c' are whne so from H, p = p' contradicting eqDiff *) - todo "Completeness". + destruct ir1 as (_&wh1), ir2 as (_&wh2). + cbn in *. + rename eq3 into c_is_red. + rename eq4 into c'_is_red. + symmetry in c'_is_red. + apply eqb_term_spec in c'_is_red. + eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. + symmetry in c_is_red. + apply eqb_term_spec in c_is_red. + eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. + rewrite !zipp_as_mkApps in H. + rewrite zipp_as_mkApps in wh1, wh2. + apply whnf_mkApps_inv in wh1; [|easy]. + apply whnf_mkApps_inv in wh2; [|easy]. + depelim wh1; solve_discr. + depelim wh2; solve_discr. + depelim H; cbn in *; try easy; solve_discr. + depelim H0; cbn in *; try easy; solve_discr. + apply whnf_whne_upgrade in c_is_red; auto. + apply whnf_whne_upgrade in c'_is_red; auto. + destruct hΣ, hx. + assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). + { eapply whne_conv_context; eauto. + apply conv_context_sym; eauto. } + apply conv_cum_mkApps_inv in H1 as [(conv_proj&conv_args)]; auto. + 2-3: apply whnf_mkApps; eauto using whne. + apply conv_cum_tProj_inv in conv_proj as [(->&?)]; auto. + rewrite eq_prod_refl in eq5; + eauto using eq_prod_refl, Nat.eqb_refl, eq_string_refl, eq_inductive_refl. Qed. Next Obligation. eapply red_wellformed ; auto. @@ -3592,7 +3648,21 @@ Section Conversion. + eapply conv_context_sym. all: auto. Qed. Next Obligation. - todo "Completeness". + apply h; cbn; clear h. + pose proof hΣ as w. destruct w. + destruct hx as [hx]. + match type of eq4 with + | context [ reduce_term ?f ?Σ ?hΣ ?Γ c' ?h ] => + pose proof (reduce_term_sound f Σ hΣ Γ c' h) as hr + end. + destruct hr as [hr]. + etransitivity; [eassumption|]. + eapply conv_cum_context_convp; eauto. + 2: eapply conv_context_sym; eauto. + eapply red_conv_cum_l ; try assumption. + eapply red_zipp. + eapply red_proj_c. + eassumption. Qed. Next Obligation. eapply red_wellformed ; auto. @@ -3647,7 +3717,20 @@ Section Conversion. - assumption. Qed. Next Obligation. - todo "Completeness". + apply h; cbn; clear h. + destruct hΣ as [wΣ]. + destruct hx as [hx]. + match type of eq3 with + | context [ reduce_term ?f ?Σ ?hΣ ?Γ c ?h ] => + pose proof (reduce_term_sound f Σ hΣ Γ c h) as hr + end. + destruct hr as [hr]. + etransitivity. + - eapply red_conv_cum_r ; try assumption. + eapply red_zipp. + eapply red_proj_c. + eassumption. + - assumption. Qed. (* tFix *) From 2abfd51ca8717b60134f1c7e2c55f3d7999e7a8d Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Fri, 23 Oct 2020 14:10:56 +0200 Subject: [PATCH 12/26] Prove a much more useful inversion lemma for reductions from whnf terms --- erasure/theories/ErasureCorrectness.v | 2 +- erasure/theories/ErasureFunction.v | 14 +- erasure/theories/SafeErasureFunction.v | 7 +- pcuic/_CoqProject.in | 1 + pcuic/theories/PCUICContextConversion.v | 209 +------ pcuic/theories/PCUICContextRelation.v | 258 ++++++++ pcuic/theories/PCUICConversionInversion.v | 154 ++--- pcuic/theories/PCUICCumulProp.v | 4 +- pcuic/theories/PCUICNormal.v | 701 +++++++++------------- safechecker/theories/PCUICSafeChecker.v | 20 +- 10 files changed, 623 insertions(+), 747 deletions(-) create mode 100644 pcuic/theories/PCUICContextRelation.v diff --git a/erasure/theories/ErasureCorrectness.v b/erasure/theories/ErasureCorrectness.v index b5f658a34..056b0d9f9 100644 --- a/erasure/theories/ErasureCorrectness.v +++ b/erasure/theories/ErasureCorrectness.v @@ -79,7 +79,7 @@ Lemma wf_local_rel_conv: forall Σ : global_env × universes_decl, wf Σ.1 -> forall Γ Γ' : context, - PCUICContextConversion.context_relation (PCUICContextConversion.conv_decls Σ) Γ Γ' -> + PCUICContextRelation.context_relation (PCUICContextConversion.conv_decls Σ) Γ Γ' -> forall Γ0 : context, wf_local Σ Γ' -> wf_local_rel Σ Γ Γ0 -> wf_local_rel Σ Γ' Γ0. Proof. intros Σ wfΣ Γ Γ' X1 Γ0 ? w0. induction w0. diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index 06f932f09..af336fdf7 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -130,7 +130,9 @@ Next Obligation. destruct HΣ. eapply PCUICConfluence.red_confluence in X0 as [t'' []]. 3:exact X1. 2:eauto. eapply invert_red_sort in r0; eauto. subst. - eapply whnf_red_sort in r. congruence. eauto. + eapply whnf_red_shape in r; auto. + depelim r. + congruence. Qed. Program Definition reduce_to_prod' Γ t (h : wellformed Σ Γ t) @@ -153,7 +155,9 @@ Next Obligation. destruct HΣ. eapply PCUICConfluence.red_confluence in X2 as [t'' []]. 3:exact X3. 2:eauto. eapply invert_red_prod in r0 as (? & ? & [] & ?); eauto. subst. - eapply whnf_red_prod in r as (? & ? & ?). congruence. eauto. + eapply whnf_red_shape in r; auto. + depelim r. + congruence. Qed. Equations is_arity Γ (HΓ : ∥wf_local Σ Γ∥) T (HT : wellformed Σ Γ T) : @@ -223,11 +227,11 @@ Next Obligation. constructor. etransitivity; eauto. - eapply PCUICContextConversion.context_change_decl_types_red; eauto. + eapply PCUICContextRelation.context_change_decl_types_red; eauto. constructor; [|constructor]. - eapply PCUICContextConversion.context_relation_refl. + eapply PCUICContextRelation.context_relation_refl. intros. - destruct x0 as [? [|]? ]; constructor. + reflexivity. Qed. Hint Constructors squash : core. diff --git a/erasure/theories/SafeErasureFunction.v b/erasure/theories/SafeErasureFunction.v index f03ad80bc..8b9736491 100644 --- a/erasure/theories/SafeErasureFunction.v +++ b/erasure/theories/SafeErasureFunction.v @@ -188,11 +188,10 @@ Section fix_sigma. constructor. etransitivity; eauto. - eapply PCUICContextConversion.context_change_decl_types_red; eauto. + eapply PCUICContextRelation.context_change_decl_types_red; eauto. constructor; [|constructor]. - eapply PCUICContextConversion.context_relation_refl. - intros. - destruct x0 as [? [|]? ]; constructor. + eapply PCUICContextRelation.context_relation_refl. + reflexivity. Qed. Next Obligation. diff --git a/pcuic/_CoqProject.in b/pcuic/_CoqProject.in index 2fbb43b7b..18694e2bd 100644 --- a/pcuic/_CoqProject.in +++ b/pcuic/_CoqProject.in @@ -25,6 +25,7 @@ theories/PCUICReduction.v theories/PCUICParallelReduction.v theories/PCUICParallelReductionConfluence.v theories/PCUICConfluence.v +theories/PCUICContextRelation.v theories/PCUICContextConversion.v theories/PCUICConversion.v theories/PCUICConversionInversion.v diff --git a/pcuic/theories/PCUICContextConversion.v b/pcuic/theories/PCUICContextConversion.v index 70a512a33..3ead28726 100644 --- a/pcuic/theories/PCUICContextConversion.v +++ b/pcuic/theories/PCUICContextConversion.v @@ -6,6 +6,8 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICParallelReduction PCUICEquality PCUICUnivSubstitution PCUICParallelReductionConfluence PCUICConfluence. +From MetaCoq.PCUIC Require Export PCUICContextRelation. + From Coq Require Import CRelationClasses ssreflect. From Equations Require Import Equations. @@ -22,75 +24,6 @@ Ltac my_rename_hyp h th := Ltac rename_hyp h ht ::= my_rename_hyp h ht. -Inductive context_relation (P : context -> context -> context_decl -> context_decl -> Type) - : forall (Γ Γ' : context), Type := -| ctx_rel_nil : context_relation P nil nil -| ctx_rel_vass na na' T U Γ Γ' : - context_relation P Γ Γ' -> - P Γ Γ' (vass na T) (vass na' U) -> - context_relation P (vass na T :: Γ) (vass na' U :: Γ') -| ctx_rel_def na na' t T u U Γ Γ' : - context_relation P Γ Γ' -> - P Γ Γ' (vdef na t T) (vdef na' u U) -> - context_relation P (vdef na t T :: Γ) (vdef na' u U :: Γ'). -Derive Signature for context_relation. -Arguments context_relation P Γ Γ' : clear implicits. - -Lemma context_relation_length P Γ Γ' : - context_relation P Γ Γ' -> #|Γ| = #|Γ'|. -Proof. - induction 1; cbn; congruence. -Qed. - -Lemma context_relation_impl {P Q Γ Γ'} : - (forall Γ Γ' d d', P Γ Γ' d d' -> Q Γ Γ' d d') -> - context_relation P Γ Γ' -> context_relation Q Γ Γ'. -Proof. - induction 2; constructor; auto. -Qed. - -Lemma context_relation_refl P : (forall Δ x, P Δ Δ x x) -> - forall Δ, context_relation P Δ Δ. -Proof. - intros HP. - induction Δ. - constructor; auto. - destruct a as [? [?|] ?]; constructor; auto. -Qed. - -Lemma context_relation_nth {P n Γ Γ' d} : - context_relation P Γ Γ' -> nth_error Γ n = Some d -> - { d' & ((nth_error Γ' n = Some d') * - let Γs := skipn (S n) Γ in - let Γs' := skipn (S n) Γ' in - context_relation P Γs Γs' * - P Γs Γs' d d')%type }. -Proof. - induction n in Γ, Γ', d |- *; destruct Γ; intros Hrel H; noconf H. - - depelim Hrel. - simpl. eexists; intuition eauto. - eexists; intuition eauto. - - depelim Hrel. - destruct (IHn _ _ _ Hrel H). - cbn -[skipn] in *. - eexists; intuition eauto. - destruct (IHn _ _ _ Hrel H). - eexists; intuition eauto. -Qed. - -Lemma context_relation_trans P : - (forall Γ Γ' Γ'' x y z, - context_relation P Γ Γ' -> - context_relation P Γ' Γ'' -> - context_relation P Γ Γ'' -> - P Γ Γ' x y -> P Γ' Γ'' y z -> P Γ Γ'' x z) -> - Transitive (context_relation P). -Proof. - intros HP x y z H. induction H in z |- *; auto; - intros H'; unfold context in *; depelim H'; - try constructor; eauto; hnf in H0; noconf H0; eauto. -Qed. - Hint Resolve conv_refl' : pcuic. Arguments skipn : simpl never. @@ -642,115 +575,6 @@ Notation conv_context Σ Γ Γ' := (context_relation (conv_decls Σ) Γ Γ'). Hint Resolve conv_ctx_refl' : pcuic. -Section ContextChangeTypesReduction. - Context {cf : checker_flags}. - Context (Σ : global_env_ext). - Context (wfΣ : wf Σ). - - Inductive change_decl_type : context_decl -> context_decl -> Type := - | change_vass_type : forall (na na' : name) (T T' : term), - change_decl_type (vass na T) (vass na' T') - | change_vdef_type : forall (na na' : name) (b T T' : term), - change_decl_type (vdef na b T) (vdef na' b T'). - - Lemma context_change_decl_types_red1 Γ Γ' s t : - context_relation (fun _ _ => change_decl_type) Γ Γ' -> red1 Σ Γ s t -> red Σ Γ' s t. - Proof. - intros HT X0. induction X0 using red1_ind_all in Γ', HT |- *; eauto. - all:pcuic. - - econstructor. econstructor. - rewrite <- H. - induction HT in i |- *; destruct i; eauto. - now inv p. - - - eapply PCUICReduction.red_abs. eapply IHX0; eauto. eauto. - - - eapply PCUICReduction.red_abs. eauto. eapply IHX0. eauto. - eauto. econstructor. eauto. econstructor. - - - eapply PCUICReduction.red_letin. eapply IHX0; eauto. - all:eauto. - - - eapply PCUICReduction.red_letin; eauto. - - - eapply PCUICReduction.red_letin; eauto. eapply IHX0; eauto. - econstructor. eauto. econstructor. - - eapply PCUICReduction.red_case; eauto. clear. - eapply All_All2_refl. induction brs; eauto. - - eapply PCUICReduction.red_case; eauto. clear. - eapply All_All2_refl. induction brs; eauto. - - destruct ind. - eapply red_case; eauto. - clear - X HT. - induction X. - + econstructor. destruct p. destruct p. - split; eauto. - eapply All_All2_refl. - induction tl; eauto. - + econstructor. now split. - eassumption. - - - eapply PCUICReduction.red_proj_c. eauto. - - - eapply PCUICReduction.red_app; eauto. - - eapply PCUICReduction.red_app; eauto. - - - eapply PCUICReduction.red_prod; eauto. - - - eapply PCUICReduction.red_prod; eauto. eapply IHX0. eauto. eauto. - econstructor. - eauto. econstructor. - - eapply PCUICReduction.red_evar; eauto. - induction X; eauto. econstructor. eapply p; eauto. - induction tl; eauto. - - eapply PCUICReduction.red_fix_one_ty. - eapply OnOne2_impl ; eauto. - intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. - inversion e. subst. clear e. - split ; auto. - - eapply PCUICReduction.red_fix_one_body. - eapply OnOne2_impl ; eauto. - intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. - inversion e. subst. clear e. - split ; auto. - eapply ih ; auto. - clear - HT. - induction (fix_context mfix0) as [| [na [b|] ty] Δ ihΔ]. - + auto. - + simpl. constructor ; eauto. - constructor. - + simpl. constructor ; eauto. - constructor. - - eapply PCUICReduction.red_cofix_one_ty. - eapply OnOne2_impl ; eauto. - intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. - inversion e. subst. clear e. - split ; auto. - - eapply PCUICReduction.red_cofix_one_body. - eapply OnOne2_impl ; eauto. - intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. - inversion e. subst. clear e. - split ; auto. - eapply ih ; auto. - clear - HT. - induction (fix_context mfix0) as [| [na [b|] ty] Δ ihΔ]. - + auto. - + simpl. constructor ; eauto. - constructor. - + simpl. constructor ; eauto. - constructor. - Qed. - - Lemma context_change_decl_types_red Γ Γ' s t : - context_relation (fun _ _ => change_decl_type) Γ Γ' -> red Σ Γ s t -> red Σ Γ' s t. - Proof. - intros. induction X0 using red_rect'; eauto. - etransitivity. eapply IHX0. - eapply context_change_decl_types_red1; eauto. - Qed. -End ContextChangeTypesReduction. - - (* Lemma wf_local_conv_ctx {cf:checker_flags} Σ Γ Δ (wfΓ : wf_local Σ Γ) : wf Σ -> *) (* All_local_env_over typing *) (* (fun (Σ : global_env_ext) (Γ : context) wfΓ (t T : term) Ht => *) @@ -786,35 +610,6 @@ Qed. Hint Constructors conv_decls : pcuic. -Lemma context_relation_app {P} Γ Γ' Δ Δ' : - #|Δ| = #|Δ'| -> - context_relation P (Γ ,,, Δ) (Γ' ,,, Δ') -> - context_relation P Γ Γ' * context_relation (fun Δ Δ' => P (Γ ,,, Δ) (Γ' ,,, Δ')) Δ Δ'. -Proof. - intros H. - induction Δ in H, Δ', Γ, Γ' |- *; - destruct Δ'; try discriminate. - intuition auto. constructor. - intros H'. simpl in H. - specialize (IHΔ Γ Γ' Δ' ltac:(lia)). - depelim H'; specialize (IHΔ H'); intuition auto; - constructor; auto. -Qed. - -Lemma context_relation_app_inv {P} Γ Γ' Δ Δ' : - #|Δ| = #|Δ'| -> - context_relation P Γ Γ' -> context_relation (fun Δ Δ' => P (Γ ,,, Δ) (Γ' ,,, Δ')) Δ Δ' -> - context_relation P (Γ ,,, Δ) (Γ' ,,, Δ'). -Proof. - intros H. - induction 2; simpl; auto. - constructor. apply IHX0. simpl in H. lia. - apply p. - constructor. apply IHX0. simpl in H; lia. - apply p. -Qed. - - Lemma eq_context_upto_conv_context {cf:checker_flags} (Σ : global_env_ext) Re : RelationClasses.subrelation Re (eq_universe Σ) -> subrelation (eq_context_upto Σ Re) (fun Γ Γ' => conv_context Σ Γ Γ'). diff --git a/pcuic/theories/PCUICContextRelation.v b/pcuic/theories/PCUICContextRelation.v new file mode 100644 index 000000000..9c7686471 --- /dev/null +++ b/pcuic/theories/PCUICContextRelation.v @@ -0,0 +1,258 @@ +From Equations Require Import Equations. +From MetaCoq.Template Require Import config utils. +From MetaCoq.PCUIC Require Import PCUICAst PCUICLiftSubst PCUICReduction. + +From Coq Require Import CRelationClasses. + +Ltac pcuic := + try repeat red; cbn in *; + try (solve [ intuition auto; eauto with pcuic || (try lia || congruence) ]). + +Inductive context_relation (P : context -> context -> context_decl -> context_decl -> Type) + : forall (Γ Γ' : context), Type := +| ctx_rel_nil : context_relation P nil nil +| ctx_rel_vass na na' T U Γ Γ' : + context_relation P Γ Γ' -> + P Γ Γ' (vass na T) (vass na' U) -> + context_relation P (vass na T :: Γ) (vass na' U :: Γ') +| ctx_rel_def na na' t T u U Γ Γ' : + context_relation P Γ Γ' -> + P Γ Γ' (vdef na t T) (vdef na' u U) -> + context_relation P (vdef na t T :: Γ) (vdef na' u U :: Γ'). + +Derive Signature for context_relation. +Arguments context_relation P Γ Γ' : clear implicits. + +Lemma context_relation_length P Γ Γ' : + context_relation P Γ Γ' -> #|Γ| = #|Γ'|. +Proof. + induction 1; cbn; congruence. +Qed. + +Lemma context_relation_impl {P Q Γ Γ'} : + (forall Γ Γ' d d', P Γ Γ' d d' -> Q Γ Γ' d d') -> + context_relation P Γ Γ' -> context_relation Q Γ Γ'. +Proof. + induction 2; constructor; auto. +Qed. + +Lemma context_relation_refl P : (forall Δ x, P Δ Δ x x) -> + forall Δ, context_relation P Δ Δ. +Proof. + intros HP. + induction Δ. + constructor; auto. + destruct a as [? [?|] ?]; constructor; auto. +Qed. + +Lemma context_relation_nth {P n Γ Γ' d} : + context_relation P Γ Γ' -> nth_error Γ n = Some d -> + { d' & ((nth_error Γ' n = Some d') * + let Γs := skipn (S n) Γ in + let Γs' := skipn (S n) Γ' in + context_relation P Γs Γs' * + P Γs Γs' d d')%type }. +Proof. + induction n in Γ, Γ', d |- *; destruct Γ; intros Hrel H; noconf H. + - depelim Hrel. + simpl. eexists; intuition eauto. + eexists; intuition eauto. + - depelim Hrel. + destruct (IHn _ _ _ Hrel H). + cbn -[skipn] in *. + eexists; intuition eauto. + destruct (IHn _ _ _ Hrel H). + eexists; intuition eauto. +Qed. + +Lemma context_relation_trans P : + (forall Γ Γ' Γ'' x y z, + context_relation P Γ Γ' -> + context_relation P Γ' Γ'' -> + context_relation P Γ Γ'' -> + P Γ Γ' x y -> P Γ' Γ'' y z -> P Γ Γ'' x z) -> + Transitive (context_relation P). +Proof. + intros HP x y z H. induction H in z |- *; auto; + intros H'; unfold context in *; depelim H'; + try constructor; eauto; hnf in H0; noconf H0; eauto. +Qed. + +Lemma context_relation_app {P} Γ Γ' Δ Δ' : + #|Δ| = #|Δ'| -> + context_relation P (Γ ,,, Δ) (Γ' ,,, Δ') -> + context_relation P Γ Γ' * context_relation (fun Δ Δ' => P (Γ ,,, Δ) (Γ' ,,, Δ')) Δ Δ'. +Proof. + intros H. + induction Δ in H, Δ', Γ, Γ' |- *; + destruct Δ'; try discriminate. + intuition auto. constructor. + intros H'. simpl in H. + specialize (IHΔ Γ Γ' Δ' ltac:(lia)). + depelim H'; specialize (IHΔ H'); intuition auto; + constructor; auto. +Qed. + +Lemma context_relation_app_inv {P} Γ Γ' Δ Δ' : + #|Δ| = #|Δ'| -> + context_relation P Γ Γ' -> context_relation (fun Δ Δ' => P (Γ ,,, Δ) (Γ' ,,, Δ')) Δ Δ' -> + context_relation P (Γ ,,, Δ) (Γ' ,,, Δ'). +Proof. + intros H. + induction 2; simpl; auto. + constructor. apply IHX0. simpl in H. lia. + apply p. + constructor. apply IHX0. simpl in H; lia. + apply p. +Qed. + +Section ContextChangeTypesReduction. + Context {cf : checker_flags}. + Context (Σ : global_env). + + Inductive change_decl_type : context_decl -> context_decl -> Type := + | change_vass_type : forall (na na' : name) (T T' : term), + change_decl_type (vass na T) (vass na' T') + | change_vdef_type : forall (na na' : name) (b T T' : term), + change_decl_type (vdef na b T) (vdef na' b T'). + + Derive Signature for change_decl_type. + + Global Instance change_decl_type_refl : Reflexive change_decl_type. + Proof. intros [? [|]]; constructor. Qed. + + Global Instance change_decl_type_sym : Symmetric change_decl_type. + Proof. + intros x y rel. + depelim rel; constructor. + Qed. + + Global Instance change_decl_type_trans : Transitive change_decl_type. + Proof. + intros x y z xy yz. + depelim xy; depelim yz; constructor. + Qed. + + Global Instance change_decl_type_equiv : Equivalence change_decl_type. + Proof. constructor; typeclasses eauto. Qed. + + Lemma context_change_decl_types_red1 Γ Γ' s t : + context_relation (fun _ _ => change_decl_type) Γ Γ' -> red1 Σ Γ s t -> red Σ Γ' s t. + Proof. + intros HT X0. induction X0 using red1_ind_all in Γ', HT |- *; eauto. + all:pcuic. + - econstructor. econstructor. + rewrite <- H. + induction HT in i |- *; destruct i; eauto. + now inv p. + - + eapply PCUICReduction.red_abs. eapply IHX0; eauto. eauto. + - + eapply PCUICReduction.red_abs. eauto. eapply IHX0. eauto. + eauto. econstructor. eauto. econstructor. + - + eapply PCUICReduction.red_letin. eapply IHX0; eauto. + all:eauto. + - + eapply PCUICReduction.red_letin; eauto. + - + eapply PCUICReduction.red_letin; eauto. eapply IHX0; eauto. + econstructor. eauto. econstructor. + - eapply PCUICReduction.red_case; eauto. clear. + eapply All_All2_refl. induction brs; eauto. + - eapply PCUICReduction.red_case; eauto. clear. + eapply All_All2_refl. induction brs; eauto. + - destruct ind. + eapply red_case; eauto. + clear - X HT. + induction X. + + econstructor. destruct p. destruct p. + split; eauto. + eapply All_All2_refl. + induction tl; eauto. + + econstructor. now split. + eassumption. + - + eapply PCUICReduction.red_proj_c. eauto. + - + eapply PCUICReduction.red_app; eauto. + - eapply PCUICReduction.red_app; eauto. + - + eapply PCUICReduction.red_prod; eauto. + - + eapply PCUICReduction.red_prod; eauto. eapply IHX0. eauto. eauto. + econstructor. + eauto. econstructor. + - eapply PCUICReduction.red_evar; eauto. + induction X; eauto. econstructor. eapply p; eauto. + induction tl; eauto. + - eapply PCUICReduction.red_fix_one_ty. + eapply OnOne2_impl ; eauto. + intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. + inversion e. subst. clear e. + split ; auto. + - eapply PCUICReduction.red_fix_one_body. + eapply OnOne2_impl ; eauto. + intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. + inversion e. subst. clear e. + split ; auto. + eapply ih ; auto. + clear - HT. + induction (fix_context mfix0) as [| [na [b|] ty] Δ ihΔ]. + + auto. + + simpl. constructor ; eauto. + constructor. + + simpl. constructor ; eauto. + constructor. + - eapply PCUICReduction.red_cofix_one_ty. + eapply OnOne2_impl ; eauto. + intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. + inversion e. subst. clear e. + split ; auto. + - eapply PCUICReduction.red_cofix_one_body. + eapply OnOne2_impl ; eauto. + intros [? ? ? ?] [? ? ? ?] [[r ih] e]. simpl in *. + inversion e. subst. clear e. + split ; auto. + eapply ih ; auto. + clear - HT. + induction (fix_context mfix0) as [| [na [b|] ty] Δ ihΔ]. + + auto. + + simpl. constructor ; eauto. + constructor. + + simpl. constructor ; eauto. + constructor. + Qed. + + Lemma context_change_decl_types_red Γ Γ' s t : + context_relation (fun _ _ => change_decl_type) Γ Γ' -> red Σ Γ s t -> red Σ Γ' s t. + Proof. + intros. induction X0 using red_rect'; eauto. + etransitivity. eapply IHX0. + eapply context_change_decl_types_red1; eauto. + Qed. +End ContextChangeTypesReduction. + +Lemma fix_context_change_decl_types Γ mfix mfix' : + #|mfix| = #|mfix'| -> + context_relation (fun _ _ => change_decl_type) (Γ,,, fix_context mfix) (Γ,,, fix_context mfix'). +Proof. + intros len. + apply context_relation_app_inv. + - now rewrite !fix_context_length. + - apply context_relation_refl. + intros. + destruct x. + destruct decl_body; constructor. + - unfold fix_context, mapi. + generalize 0 at 2 4. + induction mfix in mfix', len |- *; intros n. + + destruct mfix'; [|cbn in *; discriminate len]. + constructor. + + destruct mfix'; cbn in *; [discriminate len|]. + apply context_relation_app_inv. + * now rewrite !List.rev_length, !mapi_rec_length. + * constructor; [constructor|]. + constructor. + * apply IHmfix; lia. +Qed. diff --git a/pcuic/theories/PCUICConversionInversion.v b/pcuic/theories/PCUICConversionInversion.v index 04390abd4..8c1d8723f 100644 --- a/pcuic/theories/PCUICConversionInversion.v +++ b/pcuic/theories/PCUICConversionInversion.v @@ -162,8 +162,8 @@ Section fixed. Proof. intros conv shape shape' wh wh'. apply conv_cum_alt in conv as [(?&?&(r1&r2)&e)]. - apply whnf_red_from_mkApps in r1 as [(?&?&->&?&?)]; auto. - apply whnf_red_from_mkApps in r2 as [(?&?&->&?&?)]; auto. + apply whnf_red_mkApps_l in r1 as [(?&?&->&?&?)]; auto. + apply whnf_red_mkApps_l in r2 as [(?&?&->&?&?)]; auto. assert (isApp x1 = false). { erewrite whnf_red_isApp. 3: eauto. @@ -211,94 +211,23 @@ Section fixed. depelim whr; solve_discr. depelim H0; solve_discr; try discriminate. apply conv_cum_alt in conv as [(?&?&(r1&r2)&?)]. - apply whne_red_from_tCase in r1 as [(?&?&?&->&?&?&?)]; auto. - apply whne_red_from_tCase in r2 as [(?&?&?&->&?&?&?)]; auto. + eapply whnf_red_shape in r1; eauto. + eapply whnf_red_shape in r2; eauto. + depelim r1. + depelim r2. depelim e. constructor. split; [easy|]. - split; [apply conv_alt_red; now exists x1, x|]. - split; [apply conv_alt_red; now exists x2, x4|]. - clear -a a0 a1. - induction a1 in brs, brs', x3, a, x5, a0, a1 |- *; depelim a; depelim a0; [now constructor|]. + split; [apply conv_alt_red; now exists motive'0, motive'1|]. + split; [apply conv_alt_red; now exists discr'0, discr'1|]. + clear -a X1 X4. + induction a in brs, brs', brs'0, brs'1, X1, X4, a |- *; + depelim X1; depelim X4; [now constructor|]. constructor. + destruct p, p0, r. split; [congruence|]. apply conv_alt_red; now exists x.2, y.2. - + now apply IHa1. - Qed. - - Lemma fix_context_change_decl_types Γ mfix mfix' : - #|mfix| = #|mfix'| -> - context_relation (fun _ _ => change_decl_type) (Γ,,, fix_context mfix) (Γ,,, fix_context mfix'). - Proof. - intros len. - apply context_relation_app_inv. - - now rewrite !fix_context_length. - - apply context_relation_refl. - intros. - destruct x. - destruct decl_body; constructor. - - unfold fix_context, mapi. - generalize 0 at 2 4. - induction mfix in mfix', len |- *; intros n. - + destruct mfix'; [|cbn in *; discriminate len]. - constructor. - + destruct mfix'; cbn in *; [discriminate len|]. - apply context_relation_app_inv. - * now rewrite !List.rev_length, !mapi_rec_length. - * constructor; [constructor|]. - constructor. - * apply IHmfix; lia. - Qed. - - Lemma red_tFix_inv Γ mfix idx t : - red Σ Γ (tFix mfix idx) t -> - ∑ mfix', - t = tFix mfix' idx × - All2 (fun d d' => dname d = dname d' × rarg d = rarg d' × - red Σ Γ (dtype d) (dtype d') × - red Σ (Γ,,, fix_context mfix) (dbody d) (dbody d')) - mfix mfix'. - Proof. - intros r. - remember (tFix mfix idx) eqn:fromeq. - revert mfix fromeq. - induction r using red_rect_n1; intros mfix ->. - - eexists _; split; [reflexivity|]. - apply All2_same. - eauto. - - specialize (IHr _ eq_refl) as (?&->&all). - depelim X. - + exfalso. - destruct args using List.rev_ind; [|rewrite <- mkApps_nested in H; discriminate H]. - unfold is_constructor in e0. - rewrite nth_error_nil in e0; congruence. - + match type of all with - | All2 ?P _ _ => eapply OnOne2_All2 with (Q := P) in o - end; eauto. - 2: intros ? ? (?&[= -> -> ->]); eauto. - eexists _; split; [reflexivity|]. - eapply All2_trans; eauto. - intros ? ? ? (?&?&?&?) (?&?&?&?); eauto. - do 2 (split; [easy|]). - split; [etransitivity; eauto|]. - etransitivity; eauto. - + match type of all with - | All2 ?P _ _ => eapply OnOne2_All2 with (Q := P) in o - end; eauto. - 2: { intros ? ? (?&[= -> -> ->]); eauto. - do 2 (split; [easy|]). - split; [etransitivity; eauto|]. - eapply context_change_decl_types_red; eauto. - apply fix_context_change_decl_types. - apply All2_length in all. - easy. } - eexists _; split; [reflexivity|]. - eapply All2_trans; eauto. - intros ? ? ? (?&?&?&?) (?&?&?&?); eauto. - do 2 (split; [easy|]). - split; [etransitivity; eauto|]. - etransitivity; eauto. + + now apply IHa. Qed. Lemma conv_cum_tFix_inv leq Γ mfix idx mfix' idx' : @@ -311,23 +240,30 @@ Section fixed. Proof. intros conv. apply conv_cum_alt in conv as [(?&?&(r1&r2)&?)]. - apply red_tFix_inv in r1 as (?&->&?). - apply red_tFix_inv in r2 as (?&->&?). + assert (forall defs i, whnf RedFlags.default Σ Γ (tFix defs i)). + { intros defs i. + apply whnf_fixapp with (v := []). + destruct unfold_fix as [(?&?)|]; [|easy]. + now rewrite nth_error_nil. } + eapply whnf_red_shape in r1; eauto. + eapply whnf_red_shape in r2; eauto. + depelim r1. + depelim r2. depelim e. constructor. split; [easy|]. - clear -a a0 a1. + clear -a X X0. cut (#|mfix| = #|mfix'|); - [|now apply All2_length in a; apply All2_length in a0; apply All2_length in a1]. - revert a a0 a1. + [|now apply All2_length in a; apply All2_length in X; apply All2_length in X0]. + revert a X X0. generalize mfix at 1 3 4. generalize mfix' at 1 3. intros ctx_fix ctx_fix'. intros all1 all2 all len_eq. - induction all in mfix, mfix', x1, x, all1, all2, all |- *; + induction all in mfix, mfix', mfix'0, mfix'1, all1, all2, all |- *; depelim all1; depelim all2; [constructor|]. constructor; [|now auto]. - destruct p as (?&?&?&?), p0 as (?&?&?&?), r as ((?&?)&?). + destruct p as ((?&?)&?), p0 as (?&?&?&?), r as (?&?&?&?). split; [congruence|]. split; [now apply conv_alt_red; exists (dtype x), (dtype y)|]. apply conv_alt_red. @@ -337,6 +273,24 @@ Section fixed. eapply context_change_decl_types_red; eauto. eapply fix_context_change_decl_types; eauto. Qed. + + Lemma conv_cum_tProj_inv leq Γ p c p' c' : + conv_cum leq Σ Γ (tProj p c) (tProj p' c') -> + whnf RedFlags.default Σ Γ (tProj p c) -> + whnf RedFlags.default Σ Γ (tProj p' c') -> + ∥ p = p' × Σ;;; Γ |- c = c' ∥. + Proof. + intros conv whl whr. + apply conv_cum_alt in conv as [(?&?&(r1&r2)&?)]. + eapply whnf_red_shape in r1; eauto. + eapply whnf_red_shape in r2; eauto. + depelim r1. + depelim r2. + depelim e. + constructor. + split; [easy|]. + now apply conv_alt_red; exists c'0, c'1. + Qed. (* TODO: Move to a place that actually should be depending on typing *) Lemma conv_cum_red_inv leq Γ t1 t2 t1' t2' : @@ -362,24 +316,4 @@ Section fixed. auto. Qed. - Lemma conv_cum_tProj_inv leq Γ p c p' c' : - conv_cum leq Σ Γ (tProj p c) (tProj p' c') -> - whnf RedFlags.default Σ Γ (tProj p c) -> - whnf RedFlags.default Σ Γ (tProj p' c') -> - ∥ p = p' × Σ;;; Γ |- c = c' ∥. - Proof. - intros conv whl whr. - depelim whl; solve_discr. - depelim H; solve_discr; try discriminate. - depelim whr; solve_discr. - depelim H0; solve_discr; try discriminate. - apply conv_cum_alt in conv as [(?&?&(r1&r2)&?)]. - apply whne_red_from_tProj in r1 as [(?&->&?)]; auto. - apply whne_red_from_tProj in r2 as [(?&->&?)]; auto. - depelim e. - constructor. - split; [easy|]. - now apply conv_alt_red; exists x1, x. - Qed. - End fixed. diff --git a/pcuic/theories/PCUICCumulProp.v b/pcuic/theories/PCUICCumulProp.v index a1e33bf65..dddd78f1f 100644 --- a/pcuic/theories/PCUICCumulProp.v +++ b/pcuic/theories/PCUICCumulProp.v @@ -932,7 +932,7 @@ Proof. specialize (X1 _ _ H dom _ (eq_term_upto_univ_napp_leq X5_1)). specialize (X3 t0 B H). assert(conv_context Σ (Γ ,, vass na ty) (Γ ,, vass n t)). - { repeat constructor; pcuic. eapply conv_ctx_refl. } + { repeat constructor; pcuic. } forward X3 by eapply context_conversion; eauto; pcuic. specialize (X3 _ X5_2). eapply cumul_cumul_prop in cum; eauto. eapply cumul_prop_trans; eauto. @@ -942,7 +942,7 @@ Proof. specialize (X1 _ _ H dom _ (eq_term_upto_univ_napp_leq X7_2)). specialize (X3 _ _ H bod _ (eq_term_upto_univ_napp_leq X7_1)). assert(conv_context Σ (Γ ,, vdef na t ty) (Γ ,, vdef n b b_ty)). - { repeat constructor; pcuic. eapply conv_ctx_refl. } + { repeat constructor; pcuic. } specialize (X5 u A H). forward X5 by eapply context_conversion; eauto; pcuic. specialize (X5 _ X7_3). diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index 1e6034681..53ba61851 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -3,7 +3,8 @@ From Coq Require Import Bool String List Program BinPos Compare_dec Arith Lia. From MetaCoq.Template Require Import config Universes monad_utils utils BasicAst AstUtils UnivSubst. -From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICEquality PCUICLiftSubst PCUICTyping PCUICInduction. +From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICContextRelation + PCUICEquality PCUICLiftSubst PCUICTyping PCUICInduction. Require Import ssreflect. Set Asymmetric Patterns. @@ -159,6 +160,7 @@ Derive Signature for whnf. Derive Signature for whne. Derive Signature for All. Hint Constructors whnf whne : core. +Hint Constructors whnf whne : pcuic. Local Ltac inv H := inversion H; subst; clear H. @@ -831,215 +833,282 @@ Proof. - eauto. Qed. -Lemma whnf_red1_sort Σ Γ t u : - whnf RedFlags.default Σ Γ t -> - red1 Σ Γ t (tSort u) -> t = tSort u. -Proof. - intros. remember (tSort u) as t'. - induction X using red1_ind_all. - all: repeat match goal with - | [ H : whnf _ _ _ (?f ?a) |- _ ] => depelim H - | [ H : whne _ _ _ (?f ?a)|- _ ] => depelim H - end. - all:try (cbn in *; congruence). - all:do 2 help. - - eapply whne_mkApps_inv in H. destruct H; try firstorder congruence. depelim H. help. firstorder. - - rewrite <- mkApps_nested in Heqt'. inv Heqt'. - - eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. - inv x. rewrite H2 in H. inv H. - destruct args0 using rev_ind; cbn in *. - destruct rarg; inv H0. - rewrite mkApps_snoc in Heqt'. congruence. - - destruct args using rev_ind. - + inv Heqt'. destruct narg; inv H1. - + rewrite <- mkApps_nested in Heqt'. inv Heqt'. - - eapply whne_mkApps_inv in H as [ | (? & ? & ? & ? & ? & ? & ? & ? & ?)]; cbn; eauto. - depelim H. help. congruence. -Qed. - -Lemma whnf_red_sort Σ Γ t u : - whnf RedFlags.default Σ Γ t -> - red Σ Γ t (tSort u) -> t = tSort u. -Proof. - intros. remember (tSort u) as t'. induction X using red_rect'. - - eauto. - - subst. eapply whnf_red1_sort in X0. subst. eauto. - eapply whnf_pres; eauto. -Qed. - -Lemma whnf_red1_prod Σ Γ t na t1 t2 : - whnf RedFlags.default Σ Γ t -> - red1 Σ Γ t (tProd na t1 t2) -> exists t1 t2, t = tProd na t1 t2. -Proof. - intros. remember (tProd na t1 t2) as t'. - induction X using red1_ind_all. - all: repeat match goal with - | [ H : whnf _ _ _ (?f ?a) |- _ ] => depelim H - | [ H : whne _ _ _ (?f ?a)|- _ ] => depelim H - end. - all:try (cbn in *; congruence). - all:do 2 help. - - eapply whne_mkApps_inv in H. destruct H; try firstorder congruence. depelim H. help. firstorder. - - rewrite <- mkApps_nested in Heqt'. inv Heqt'. - - eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. - inv x. rewrite H2 in H. inv H. - destruct args0 using rev_ind; cbn in *. - destruct rarg; inv H0. - rewrite mkApps_snoc in Heqt'. congruence. - - destruct args using rev_ind. - + inv Heqt'. destruct narg; inv H1. - + rewrite <- mkApps_nested in Heqt'. inv Heqt'. - - eapply whne_mkApps_inv in H. destruct H; try firstorder congruence. depelim H. help. firstorder. - - inv Heqt'. eauto. - - inv Heqt'. eauto. -Qed. +Hint Resolve All2_same All2_firstn All2_skipn OnOne2_All2 red_mkApps All2_app : pcuic. -Lemma whnf_red_prod Σ Γ t na t1 t2 : - whnf RedFlags.default Σ Γ t -> - red Σ Γ t (tProd na t1 t2) -> exists t1 t2, t = tProd na t1 t2. +Inductive whnf_red_shape_spec Σ Γ : term -> term -> Prop := +| wr_shape_tRel i : + option_map decl_body (nth_error Γ i) = Some None -> + whnf_red_shape_spec Σ Γ (tRel i) (tRel i) +| wr_shape_tVar id : whnf_red_shape_spec Σ Γ (tVar id) (tVar id) +| wr_shape_tEvar n l l' : + All2 (red Σ Γ) l l' -> + whnf_red_shape_spec Σ Γ (tEvar n l) (tEvar n l') +| wr_shape_tConst kn u decl : + lookup_env Σ kn = Some (ConstantDecl decl) -> + cst_body decl = None -> + whnf_red_shape_spec Σ Γ (tConst kn u) (tConst kn u) +| wr_shape_tApp hd hd' arg arg' : + red Σ Γ hd hd' -> + whnf_red_shape_spec Σ Γ hd hd' -> + red Σ Γ arg arg' -> + whnf_red_shape_spec Σ Γ (tApp hd arg) (tApp hd' arg') +| wr_shape_tFix mfix mfix' idx : + All2 (fun d d' => dname d = dname d' × + rarg d = rarg d' × + red Σ Γ (dtype d) (dtype d') × + red Σ (Γ,,, fix_context mfix) (dbody d) (dbody d')) + mfix mfix' -> + whnf_red_shape_spec Σ Γ (tFix mfix idx) (tFix mfix' idx) +| wr_shape_tCase p motive motive' discr discr' brs brs' : + red Σ Γ motive motive' -> + red Σ Γ discr discr' -> + All2 (fun br br' => br.1 = br'.1 × red Σ Γ br.2 br'.2) brs brs' -> + whnf_red_shape_spec Σ Γ (tCase p motive discr brs) (tCase p motive' discr' brs') +| wr_shape_tProj p c c' : + red Σ Γ c c' -> + whnf_red_shape_spec Σ Γ (tProj p c) (tProj p c') +| wr_shape_tSort s : whnf_red_shape_spec Σ Γ (tSort s) (tSort s) +| wr_shape_tProd na A A' B B' : + red Σ Γ A A' -> + red Σ (Γ,, vass na A) B B' -> + whnf_red_shape_spec Σ Γ (tProd na A B) (tProd na A' B') +| wr_shape_tLambda na A A' B B' : + red Σ Γ A A' -> + red Σ (Γ,, vass na A) B B' -> + whnf_red_shape_spec Σ Γ (tLambda na A B) (tLambda na A' B') +| wr_shape_tConstruct i n u : + whnf_red_shape_spec Σ Γ (tConstruct i n u) (tConstruct i n u) +| wr_shape_tInd i u : + whnf_red_shape_spec Σ Γ (tInd i u) (tInd i u) +| wr_shape_tCoFix mfix mfix' idx : + All2 (fun d d' => dname d = dname d' × + rarg d = rarg d' × + red Σ Γ (dtype d) (dtype d') × + red Σ (Γ,,, fix_context mfix) (dbody d) (dbody d')) + mfix mfix' -> + whnf_red_shape_spec Σ Γ (tCoFix mfix idx) (tCoFix mfix' idx). + +Derive Signature for whnf_red_shape_spec. + +Instance whnf_red_shape_spec_trans Σ Γ : RelationClasses.Transitive (whnf_red_shape_spec Σ Γ). Proof. - intros. remember (tProd na t1 t2) as t'. revert t1 t2 Heqt'. induction X using red_rect'; intros. - - eauto. - - subst. eapply whnf_red1_prod in X0 as (? & ? & ?). subst. eauto. - eapply whnf_pres; eauto. + intros x y z xy yz. + revert z yz. + induction xy; intros z yz; depelim yz; eauto using whnf_red_shape_spec. + - constructor. + eapply All2_trans; eauto. + typeclasses eauto. + - constructor; eauto. + all: etransitivity; eauto. + - constructor. + eapply All2_trans; eauto. + 1: { intros ? ? ? (->&->&?&?) (->&->&?&?). + do 2 (split; auto). + split; etransitivity; eauto. } + eapply All2_impl; eauto. + cbn. + intros ? ? (->&->&?&?). + do 3 (split; auto). + eapply context_change_decl_types_red; eauto. + apply fix_context_change_decl_types. + now apply All2_length in X. + - constructor; try solve [etransitivity; eauto]. + eapply All2_trans; eauto. + intros ? ? ? (->&?) (->&?). + split; auto. + etransitivity; eauto. + - constructor. + etransitivity; eauto. + - constructor; etransitivity; eauto. + eapply context_change_decl_types_red; eauto. + constructor; eauto; [|constructor]. + apply context_relation_refl. + intros; reflexivity. + - constructor; etransitivity; eauto. + eapply context_change_decl_types_red; eauto. + constructor; eauto; [|constructor]. + apply context_relation_refl. + intros; reflexivity. + - constructor. + eapply All2_trans; eauto. + 1: { intros ? ? ? (->&->&?&?) (->&->&?&?). + do 2 (split; auto). + split; etransitivity; eauto. } + eapply All2_impl; eauto. + cbn. + intros ? ? (->&->&?&?). + do 3 (split; auto). + eapply context_change_decl_types_red; eauto. + apply fix_context_change_decl_types. + now apply All2_length in X. Qed. -Lemma whne_mkApps_tInd Σ Γ ind u args : - whne RedFlags.default Σ Γ (mkApps (tInd ind u) args) -> - False. +Lemma whnf_red_shape_spec_mkApps Σ Γ hd hd' args args' : + whnf_red_shape_spec Σ Γ hd hd' -> + red Σ Γ hd hd' -> + All2 (red Σ Γ) args args' -> + whnf_red_shape_spec Σ Γ (mkApps hd args) (mkApps hd' args'). Proof. - intros wh. - apply whne_mkApps_inv in wh; [|easy]. - destruct wh as [|(?&?&?&?&?&?&?&?&?)]. - - depelim H; solve_discr. - - inv H. + intros shape r ra. + induction ra using All2_ind_rev; auto. + rewrite <- !mkApps_nested. + cbn. + constructor; auto. + apply red_mkApps; auto. Qed. -Lemma whne_red1_mkApps_tInd Σ Γ ind u args t : +Lemma whnf_red_shape_spec_refl_whne Σ Γ t : whne RedFlags.default Σ Γ t -> - red1 Σ Γ t (mkApps (tInd ind u) args) -> - False. + whnf_red_shape_spec Σ Γ t t. Proof. - intros wh r. - remember (mkApps (tInd ind u) args) eqn:eq. - revert args eq. - apply (whne_red1_ind - RedFlags.default - Σ Γ - (fun t t' => forall args, t' = mkApps (tInd ind u) args -> False)) - with (t := t) (t' := t0); intros; try easy; solve_discr. - - destruct args using List.rev_ind; [discriminate H1|]. - rewrite <- mkApps_nested in H1. - inv H1. - eauto. - - destruct args using List.rev_ind; [discriminate H0|]. - rewrite <- mkApps_nested in H0. - inv H0. - now apply whne_mkApps_tInd in H. -Qed. - -Lemma whnf_red1_mkApps_tInd Σ Γ ind u args t : - whnf RedFlags.default Σ Γ t -> - red1 Σ Γ t (mkApps (tInd ind u) args) -> - exists args', t = mkApps (tInd ind u) args'. -Proof. - intros wh r. - remember (mkApps (tInd ind u) args) as t' eqn:t'eq. - revert t' r args t'eq. - induction wh; intros t' r args' ->; cbn in *; try easy. - - exfalso; eauto using whne_red1_mkApps_tInd. - - depelim r; solve_discr. - - depelim r; solve_discr. - - depelim r; solve_discr. - - apply red1_mkApps_tConstruct_inv in r as (?&?&?); solve_discr. - - apply red1_mkApps_tInd_inv in r as (?&?&?). - apply (f_equal decompose_app) in e. - rewrite !decompose_app_mkApps in e; try easy. - now inv e. - - apply red1_mkApps_tFix_inv in r. - 2: { destruct unfold_fix as [(?&?)|]; [|easy]. - unfold is_constructor. - now rewrite H. } - destruct r as [[(?&?&?)|(?&?&?)]|(?&?&?)]; solve_discr. - - eapply red1_mkApps_tCoFix_inv in r as [[(?&?&?)|(?&?&?)]|(?&?&?)]; solve_discr. -Qed. - -Lemma whnf_red_mkApps_tInd Σ Γ ind u args t : - whnf RedFlags.default Σ Γ t -> - red Σ Γ t (mkApps (tInd ind u) args) -> - exists args', t = mkApps (tInd ind u) args'. -Proof. - remember (mkApps (tInd ind u) args) as hd eqn:hdeq. - intros wh r. - revert args hdeq. - induction r using red_rect'; intros args ->. - - easy. - - apply whnf_red1_mkApps_tInd in X as (?&->); eauto. - eapply whnf_pres; eauto. + intros wh. + induction wh; cbn in *; try discriminate; eauto using whnf_red_shape_spec with pcuic. + apply whnf_red_shape_spec_mkApps; auto. + 2: apply All2_same; auto. + constructor. + apply All2_same; auto. Qed. -Lemma whne_red1_isApp Σ Γ t t' : +Lemma whne_red1_shape Σ Γ t t' : whne RedFlags.default Σ Γ t -> red1 Σ Γ t t' -> - isApp t' = isApp t. + whnf_red_shape_spec Σ Γ t t'. Proof. intros wh r. apply (whne_red1_ind RedFlags.default Σ Γ - (fun t t' => isApp t' = isApp t)) - with (t := t) (t' := t'); intros; cbn in *; try easy. - - apply OnOne2_length in X. - destruct args using List.rev_ind; - destruct x using List.rev_ind; cbn in *; - rewrite ?app_length in X; - cbn in *; try easy. - now rewrite <- !mkApps_nested. - - destruct args using List.rev_ind; cbn in *; [now auto|]. - now rewrite <- !mkApps_nested. - - destruct args using List.rev_ind; cbn in *; [now auto|]. - now rewrite <- !mkApps_nested. + (whnf_red_shape_spec Σ Γ)); intros; cbn in *; try easy; eauto using whnf_red_shape_spec. + - constructor. + eapply OnOne2_All2; eauto. + - constructor; auto. + now apply whnf_red_shape_spec_refl_whne. + - apply whnf_red_shape_spec_mkApps; eauto. + 1: constructor; apply All2_same; auto. + eapply OnOne2_All2; eauto. + - apply whnf_red_shape_spec_mkApps; eauto. + + constructor. + eapply OnOne2_All2; eauto. + cbn. + intros ? ? (?&[= -> -> ->]). + auto. + + now apply red1_red, fix_red_ty. + + apply All2_same; auto. + - apply whnf_red_shape_spec_mkApps; eauto. + + constructor. + eapply OnOne2_All2; eauto. + cbn. + intros ? ? (?&[= -> -> ->]). + auto. + + now apply red1_red, fix_red_body. + + apply All2_same; auto. + - constructor; auto. + apply All2_same; auto. + - constructor; auto. + apply All2_same; auto. + - constructor; auto. + eapply OnOne2_All2; eauto. + intros ? ? (?&?); auto. Qed. -Lemma isApp_mkApps' hd args hd' args' : - isApp hd = isApp hd' -> - #|args| = #|args'| -> - isApp (mkApps hd args) = isApp (mkApps hd' args'). +Lemma whnf_red_shape_spec_refl Σ Γ t : + whnf RedFlags.default Σ Γ t -> + whnf_red_shape_spec Σ Γ t t. Proof. - intros hd_app len. - destruct args using List.rev_ind; destruct args' using List.rev_ind; auto. - all: rewrite !app_length in len; cbn in len; try lia. - now rewrite <- !mkApps_nested. + intros wh. + induction wh; cbn in *; try discriminate; eauto using whnf_red_shape_spec with pcuic. + - eapply whnf_red_shape_spec_refl_whne; eauto. + - apply whnf_red_shape_spec_mkApps; auto. + 2: apply All2_same; auto. + constructor. + - apply whnf_red_shape_spec_mkApps; auto. + 2: apply All2_same; auto. + constructor. + - apply whnf_red_shape_spec_mkApps; auto. + 2: apply All2_same; auto. + constructor. + apply All2_same; auto. + - apply whnf_red_shape_spec_mkApps; auto. + 2: apply All2_same; auto. + constructor. + apply All2_same; auto. Qed. -Lemma whnf_red1_isApp Σ Γ t t' : +Lemma whnf_red1_shape Σ Γ t t' : whnf RedFlags.default Σ Γ t -> red1 Σ Γ t t' -> - isApp t' = isApp t. + whnf_red_shape_spec Σ Γ t t'. Proof. intros wh r. - depelim wh. - - eapply whne_red1_isApp; eauto. + destruct wh. + - eapply whne_red1_shape; eauto. - depelim r; solve_discr. - - depelim r; auto; solve_discr. - - depelim r; auto; solve_discr. + - depelim r; solve_discr; constructor; eauto using whnf_red_shape_spec. + - depelim r; solve_discr; constructor; eauto using whnf_red_shape_spec. - apply red1_mkApps_tConstruct_inv in r as (?&->&?). - apply OnOne2_length in o. - now apply isApp_mkApps'. + apply whnf_red_shape_spec_mkApps; eauto using whnf_red_shape_spec. + eapply OnOne2_All2; eauto. - apply red1_mkApps_tInd_inv in r as (?&->&?). - apply OnOne2_length in o. - now apply isApp_mkApps'. + apply whnf_red_shape_spec_mkApps; eauto using whnf_red_shape_spec. + eapply OnOne2_All2; eauto. - apply red1_mkApps_tFix_inv in r. 2: { destruct unfold_fix as [(?&?)|]; [|easy]. now unfold is_constructor; rewrite H. } - now destruct r as [[(?&->&o)|(?&->&o)]|(?&->&o)]; - apply OnOne2_length in o; - apply isApp_mkApps'. + destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]. + + apply whnf_red_shape_spec_mkApps; auto. + 2: eapply OnOne2_All2; eauto. + constructor. + apply All2_same; auto. + + apply whnf_red_shape_spec_mkApps. + 3: apply All2_same; auto. + 2: apply red1_red, fix_red_ty; auto. + constructor. + eapply OnOne2_All2; eauto. + cbn. + intros ? ? (?&[= -> -> ->]). + auto. + + apply whnf_red_shape_spec_mkApps. + 3: apply All2_same; auto. + 2: apply red1_red, fix_red_body; auto. + constructor. + eapply OnOne2_All2; eauto. + cbn. + intros ? ? (?&[= -> -> ->]). + auto. - apply red1_mkApps_tCoFix_inv in r. - now destruct r as [[(?&->&o)|(?&->&o)]|(?&->&o)]; - apply OnOne2_length in o; - apply isApp_mkApps'. + destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]. + + apply whnf_red_shape_spec_mkApps; auto. + 2: eapply OnOne2_All2; eauto. + constructor. + apply All2_same; auto. + + apply whnf_red_shape_spec_mkApps. + 3: apply All2_same; auto. + 2: apply red1_red, cofix_red_ty; auto. + constructor. + eapply OnOne2_All2; eauto. + cbn. + intros ? ? (?&[= -> -> ->]). + auto. + + apply whnf_red_shape_spec_mkApps. + 3: apply All2_same; auto. + 2: apply red1_red, cofix_red_body; auto. + constructor. + eapply OnOne2_All2; eauto. + cbn. + intros ? ? (?&[= -> -> ->]). + auto. +Qed. + +Lemma whnf_red_shape Σ Γ t t' : + whnf RedFlags.default Σ Γ t -> + red Σ Γ t t' -> + whnf_red_shape_spec Σ Γ t t'. +Proof. + intros wh r. + induction r using red_rect_n1. + - apply whnf_red_shape_spec_refl; auto. + - eapply whnf_red1_shape in X. + + etransitivity; eauto. + + eapply whnf_pres; eauto. Qed. Lemma whnf_red_isApp Σ Γ t t' : @@ -1048,157 +1117,67 @@ Lemma whnf_red_isApp Σ Γ t t' : isApp t' = isApp t. Proof. intros wh r. - induction r in wh |- * using red_rect_n1. - - easy. - - apply whnf_red1_isApp in X as ->; eauto. - eapply whnf_pres; eauto. + eapply whnf_red_shape in r; eauto. + now depelim r. Qed. -Ltac solve_red_from_mkApps := - match goal with - | [H: ?a = mkApps ?h ?args |- _] => - destruct args using List.rev_ind; [|rewrite <- mkApps_nested in H; discriminate H]; - cbn in H; subst; - constructor; eexists _, []; split; [reflexivity|]; - eauto with pcuic - end. - -Hint Resolve All2_same All2_firstn All2_skipn OnOne2_All2 red_mkApps All2_app : pcuic. -Lemma whne_red1_from_mkApps Σ Γ hd args t : - whne RedFlags.default Σ Γ (mkApps hd args) -> - red1 Σ Γ (mkApps hd args) t -> +Lemma whnf_red_mkApps_l Σ Γ hd args t : + whnf RedFlags.default Σ Γ (mkApps hd args) -> + red Σ Γ (mkApps hd args) t -> ∥∑ hd' args', t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. Proof. intros wh r. - remember (mkApps hd args) eqn:eq. - revert args eq. - apply (whne_red1_ind - RedFlags.default Σ Γ - (fun t t' => (forall args, t = mkApps hd args -> _ : Prop))) - with (t1 := t0) (t' := t); intros; cbn in *; - try congruence; try solve [solve_red_from_mkApps]. - - destruct args as [|? ? _] using List.rev_ind. - + cbn in H1; subst. - constructor; eexists _, []; split; [reflexivity|]. - eauto with pcuic. - + rewrite <- mkApps_nested in H1. - inv H1. - specialize (H0 _ eq_refl) as [(?&?&->&?&?)]. - constructor; exists x0, (x1 ++ [x]). - rewrite <- mkApps_nested; split; [reflexivity|]. - split; [easy|]; eauto with pcuic. - - destruct args as [|? ? _] using List.rev_ind. - + cbn in H0; subst. - constructor; eexists _, []; split; [reflexivity|]. - eauto with pcuic. - + rewrite <- mkApps_nested in H0. - inv H0. - constructor; exists hd, (args ++ [N2]). - rewrite <- mkApps_nested; split; [reflexivity|]. - split; [reflexivity|]; eauto with pcuic. - - destruct (mkApps_elim hd args0). - apply mkApps_notApp_inj in H3 as (<-&<-); auto; [|now destruct (isApp f)]. - constructor; exists (mkApps (tFix mfix idx) (firstn n x)), (skipn n x). - rewrite mkApps_nested. - rewrite firstn_skipn. - split; [easy|]. - split; eauto with pcuic. - - destruct (mkApps_elim hd args0). - apply mkApps_notApp_inj in H2 as (<-&<-); auto; [|now destruct (isApp f)]. - constructor; exists (mkApps (tFix x idx) (firstn n args)), (skipn n args). - rewrite mkApps_nested. - rewrite firstn_skipn. - split; [easy|]. - split; eauto with pcuic. - - destruct (mkApps_elim hd args0). - apply mkApps_notApp_inj in H2 as (<-&<-); auto; [|now destruct (isApp f)]. - constructor; exists (mkApps (tFix x idx) (firstn n args)), (skipn n args). - rewrite mkApps_nested. - rewrite firstn_skipn. - split; [easy|]. - split; eauto with pcuic. + eapply whnf_red_shape in wh; eauto. + induction args in hd, t, r, wh |- * using List.rev_ind. + - cbn in *. + constructor; eexists _, []; split; [reflexivity|]. + eauto with pcuic. + - rewrite <- mkApps_nested in wh. + cbn in wh. + depelim wh. + apply IHargs in X as [(?&?&->&?&?)]; auto. + constructor. + exists x0, (x1 ++ [arg']). + rewrite <- mkApps_nested. + eauto with pcuic. Qed. -Lemma whnf_red1_from_mkApps Σ Γ hd args t : - whnf RedFlags.default Σ Γ (mkApps hd args) -> - red1 Σ Γ (mkApps hd args) t -> - ∥∑ hd' args', - t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. +Lemma whnf_red_mkApps_r Σ Γ t hd' args' : + whnf RedFlags.default Σ Γ t -> + red Σ Γ t (mkApps hd' args') -> + ∥∑ hd args, + t = mkApps hd args × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. Proof. intros wh r. - depelim wh; try solve [solve_red_from_mkApps]. - - eapply whne_red1_from_mkApps; eauto. - - destruct (mkApps_elim hd args). - apply mkApps_notApp_inj in x as (<-&<-); auto; [|now destruct (isApp f)]. - apply red1_mkApps_tConstruct_inv in r as (?&->&?). - constructor; exists (mkApps (tConstruct i n u) (firstn n0 x)), (skipn n0 x). - rewrite mkApps_nested. - rewrite firstn_skipn. - split; [reflexivity|]. - eauto 9 with pcuic. - - destruct (mkApps_elim hd args). - apply mkApps_notApp_inj in x as (<-&<-); auto; [|now destruct (isApp f)]. - apply red1_mkApps_tInd_inv in r as (?&->&?). - constructor; exists (mkApps (tInd i u) (firstn n x)), (skipn n x). - rewrite mkApps_nested. - rewrite firstn_skipn. - split; [reflexivity|]. - split; eauto with pcuic. - - destruct (mkApps_elim hd args). - apply mkApps_notApp_inj in x as (<-&<-); auto; [|now destruct (isApp f)]. - apply red1_mkApps_tFix_inv in r. - 2: { destruct unfold_fix as [(?&?)|]; [|easy]. - now unfold is_constructor; rewrite H. } - destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]. - + constructor; exists (mkApps (tFix mfix idx) (firstn n x)), (skipn n x). - rewrite mkApps_nested firstn_skipn. - split; [easy|]. - eauto 9 with pcuic. - + constructor; exists (mkApps (tFix x idx) (firstn n v)), (skipn n v). - rewrite mkApps_nested firstn_skipn. - split; [easy|]. - split; eauto 9 with pcuic. - + constructor; exists (mkApps (tFix x idx) (firstn n v)), (skipn n v). - rewrite mkApps_nested firstn_skipn. - split; [easy|]. - split; eauto 9 with pcuic. - - destruct (mkApps_elim hd args). - apply mkApps_notApp_inj in x as (<-&<-); auto; [|now destruct (isApp f)]. - apply red1_mkApps_tCoFix_inv in r. - destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]. - + constructor; exists (mkApps (tCoFix mfix idx) (firstn n x)), (skipn n x). - rewrite mkApps_nested firstn_skipn. - split; [easy|]. - eauto 9 with pcuic. - + constructor; exists (mkApps (tCoFix x idx) (firstn n v)), (skipn n v). - rewrite mkApps_nested firstn_skipn. - split; [easy|]. - split; eauto 9 with pcuic. - + constructor; exists (mkApps (tCoFix x idx) (firstn n v)), (skipn n v). - rewrite mkApps_nested firstn_skipn. - split; [easy|]. - split; eauto 9 with pcuic. + eapply whnf_red_shape in wh; eauto. + induction args' in hd', t, r, wh |- * using List.rev_ind. + - cbn in *. + constructor; eexists _, []; split; [reflexivity|]. + eauto with pcuic. + - rewrite <- mkApps_nested in wh. + cbn in wh. + depelim wh. + apply IHargs' in X as [(?&?&->&?&?)]; auto. + constructor. + exists x0, (x1 ++ [arg]). + rewrite <- mkApps_nested. + eauto with pcuic. Qed. -Lemma whnf_red_from_mkApps Σ Γ hd args t : +Lemma whnf_red_mkApps Σ Γ hd args hd' args' : whnf RedFlags.default Σ Γ (mkApps hd args) -> - red Σ Γ (mkApps hd args) t -> - ∥∑ hd' args', - t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. + red Σ Γ (mkApps hd args) (mkApps hd' args') -> + isApp hd = false -> + isApp hd' = false -> + ∥red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. Proof. - remember (mkApps hd args) as from eqn:fromeq. - intros wh r. - revert hd wh args fromeq. - induction r using red_rect_n1; intros hd wh args ->. - - constructor; eauto 10 with pcuic. - - specialize (IHr _ wh _ eq_refl) as [(?&?&->&?&?)]. - eapply whnf_red1_from_mkApps in X as [(?&?&->&?&?)]; eauto. - 2: eapply whnf_pres; eauto. - constructor; eexists _, _; split; [reflexivity|]. - split; [now etransitivity; eauto|]. - eapply All2_trans; eauto. - typeclasses eauto. + intros wh r notapp notapp'. + apply whnf_red_mkApps_l in r as [(?&?&eq&?&?)]; auto. + apply whnf_mkApps_inv in wh; [|now destruct (isApp hd)]. + eapply whnf_red_isApp in wh as ?; eauto. + apply mkApps_notApp_inj in eq as (->&->); auto; [|congruence]. + easy. Qed. Lemma whne_eq_term_upto_univ_napp f Σ Γ t Re Rle napp t' : @@ -1286,100 +1265,6 @@ Proof. apply whnf_eq_term_upto_univ_napp. Qed. -Hint Resolve All2_same : pcuic. - -Lemma whne_red1_from_tCase Σ Γ p motive discr brs t : - whne RedFlags.default Σ Γ (tCase p motive discr brs) -> - red1 Σ Γ (tCase p motive discr brs) t -> - ∥∑ motive' discr' brs', - t = tCase p motive' discr' brs' × - red Σ Γ motive motive' × - red Σ Γ discr discr' × - All2 (fun br br' => br.1 = br'.1 × red Σ Γ br.2 br'.2) brs brs'∥. -Proof. - intros wh r. - remember (tCase p motive discr brs) as from eqn:fromeq. - revert motive discr brs fromeq. - apply (whne_red1_ind - RedFlags.default Σ Γ - (fun t t' => (forall motive discr brs, t = tCase p motive discr brs -> _ : Prop))) - with (t0 := from) (t' := t); intros; cbn in *; try discriminate; solve_discr; auto. - - inv H0. - constructor; eexists _, _, _; split; [reflexivity|]. - eauto 7 with pcuic. - - inv H1. - constructor; eexists _, _, _; split; [reflexivity|]. - eauto 7 with pcuic. - - inv H0. - constructor; eexists _, _, _; split; [reflexivity|]. - split; [eauto|]. - split; [eauto|]. - eapply OnOne2_All2; eauto. - intros ? ? (?&?). - eauto. -Qed. - -Lemma whne_red_from_tCase Σ Γ p motive discr brs t : - whne RedFlags.default Σ Γ (tCase p motive discr brs) -> - red Σ Γ (tCase p motive discr brs) t -> - ∥∑ motive' discr' brs', - t = tCase p motive' discr' brs' × - red Σ Γ motive motive' × - red Σ Γ discr discr' × - All2 (fun br br' => br.1 = br'.1 × red Σ Γ br.2 br'.2) brs brs'∥. -Proof. - intros wh r. - remember (tCase p motive discr brs) eqn:fromeq. - revert fromeq. - induction r using red_rect_n1; intros ->. - - constructor; eauto 9 with pcuic. - - destruct (IHr eq_refl) as [(?&?&?&->&?&?&?)]. - eapply whne_red1_from_tCase in X as [(?&?&?&->&?&?&?)]. - + constructor; eexists _, _, _; split; [reflexivity|]. - split; [etransitivity; eauto|]. - split; [etransitivity; eauto|]. - eapply All2_trans; eauto. - clear. - intros x y z (?&?) (?&?). - split; [congruence|]. - etransitivity; eauto. - + eapply whne_pres; eauto. -Qed. - -Lemma whne_red1_from_tProj Σ Γ p c t : - whne RedFlags.default Σ Γ (tProj p c) -> - red1 Σ Γ (tProj p c) t -> - ∥∑ c', t = tProj p c' × red Σ Γ c c'∥. -Proof. - intros wh r. - remember (tProj p c) as from eqn:fromeq. - revert c fromeq. - apply (whne_red1_ind - RedFlags.default Σ Γ - (fun t t' => (forall c, t = _ -> _ : Prop))) - with (t0 := from) (t' := t); intros; cbn in *; try discriminate; solve_discr; auto. - inv H1. - constructor; eexists _; split; [reflexivity|]. - eauto. -Qed. - -Lemma whne_red_from_tProj Σ Γ p c t : - whne RedFlags.default Σ Γ (tProj p c) -> - red Σ Γ (tProj p c) t -> - ∥∑ c', t = tProj p c' × red Σ Γ c c'∥. -Proof. - intros wh r. - remember (tProj p c) eqn:fromeq. - revert fromeq. - induction r using red_rect_n1; intros ->. - - constructor; eauto 9 with pcuic. - - destruct (IHr eq_refl) as [(?&->&?)]. - eapply whne_red1_from_tProj in X as [(?&->&?)]. - + constructor; eexists _; split; [reflexivity|]. - etransitivity; eauto. - + eapply whne_pres; eauto. -Qed. - (* Definition head_arg_is_constructor mfix idx args := diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index 9c09db36b..81460dc6f 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -445,9 +445,10 @@ Section Typecheck. destruct HΣ. eapply red_confluence in r as (?&r1&r2); eauto. apply invert_red_sort in r2 as ->. - apply whnf_red_sort in r1; auto. + eapply whnf_red_shape in r1; eauto. + depelim r1. clear Heq. - rewrite r1 in n0. + rewrite H0 in n0. now cbn in n0. Qed. @@ -485,9 +486,10 @@ Section Typecheck. destruct HΣ. eapply red_confluence in r as (?&r1&r2); eauto. apply invert_red_prod in r2 as (?&?&(->&?)&?); auto. - apply whnf_red_prod in r1 as (?&?&eq); auto. + eapply whnf_red_shape in r1; auto. + depelim r1. clear Heq. - rewrite eq in n0. + rewrite H0 in n0. now cbn in n0. Qed. @@ -569,13 +571,11 @@ Section Typecheck. cbn in *. rewrite app_nil_r in wh. apply red_mkApps_tInd in r2 as (?&->&?); auto. - apply whnf_red_mkApps_tInd in r1 as (?&eq); auto. - apply (f_equal decompose_app) in eq. - rewrite !decompose_app_mkApps in eq; [now rewrite notapp|easy|]. - noconf eq. + apply whnf_red_mkApps in r1 as [(?&?)]; auto. + apply whnf_red_shape in r; [|now apply whnf_mkApps_inv in wh]. + depelim r. noconf e0. - cbn in *. - easy. + discriminate i0. Qed. Definition iscumul Γ := isconv_term Σ HΣ Hφ G HG Γ Cumul. From 2caed2d5f1cc6e1334b8887305d827ba93e0a760 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Sun, 25 Oct 2020 09:41:06 +0100 Subject: [PATCH 13/26] Finish tApp case and work on fallback case For the tApp case (i.e. isconv_args) we have to refactor slightly. Previously isconv_args was being passed a proof that the heads are convertible and then produced a proof that the entire mkApps will be convertible. The opposite is not true: if the heads are convertible and the args are not convertible, it is still possible for the full mkApps to be convertible (eg. constants with unused args). This also uncovered a source of incompleteness for tCase and tFix where safe conversion rejected two tCase/two tFix that were syntactically equal if their args weren't convertible. Fix this as part of this change as well. To remedy this change isconv_args to return a proof that the arguments are (or aren't) convertible pairwise and then move the handling up to the places where we are using isconv_args. Also start working on the fallback case. It now takes another argument putting a constraint on the shape of the heads. Refactor some universe stuff to make it a little more uniform (in particular, no need to case on whether we are doing Conv/Cumul) anymore. Other than that, some slight refactorings. --- pcuic/theories/PCUICConversionInversion.v | 81 +- pcuic/theories/PCUICNormal.v | 54 + pcuic/theories/PCUICSafeLemmata.v | 2 - safechecker/theories/PCUICSafeConversion.v | 1091 +++++++++++++++----- 4 files changed, 916 insertions(+), 312 deletions(-) diff --git a/pcuic/theories/PCUICConversionInversion.v b/pcuic/theories/PCUICConversionInversion.v index 8c1d8723f..291590be0 100644 --- a/pcuic/theories/PCUICConversionInversion.v +++ b/pcuic/theories/PCUICConversionInversion.v @@ -2,6 +2,7 @@ From Equations Require Import Equations. From MetaCoq.PCUIC Require Import PCUICAst. From MetaCoq.PCUIC Require Import PCUICAstUtils. From MetaCoq.PCUIC Require Import PCUICContextConversion. +From MetaCoq.PCUIC Require Import PCUICConversion. From MetaCoq.PCUIC Require Import PCUICCumulativity. From MetaCoq.PCUIC Require Import PCUICCumulProp. From MetaCoq.PCUIC Require Import PCUICEquality. @@ -141,53 +142,62 @@ Section fixed. intros noapp1 noapp2 eq. now apply eq_term_upto_univ_mkApps_inv in eq as (?&?). Qed. + + Definition conv_cum_napp leq Γ napp t t' := + match t with + | tInd _ _ + | tConstruct _ _ _ => ∥eq_termp_napp leq napp t t'∥ + | _ => conv_cum leq Σ Γ t t' + end. Lemma conv_cum_mkApps_inv leq Γ hd args hd' args' : conv_cum leq Σ Γ (mkApps hd args) (mkApps hd' args') -> - match hd with - | tApp _ _ - | tInd _ _ - | tConstruct _ _ _ => False - | _ => True - end -> - match hd' with - | tApp _ _ - | tInd _ _ - | tConstruct _ _ _ => False - | _ => True - end -> + isApp hd = false -> + isApp hd' = false -> whnf RedFlags.default Σ Γ (mkApps hd args) -> whnf RedFlags.default Σ Γ (mkApps hd' args') -> - ∥conv_cum leq Σ Γ hd hd' × All2 (conv Σ Γ) args args'∥. + ∥conv_cum_napp leq Γ #|args| hd hd' × conv_terms Σ Γ args args'∥. Proof. - intros conv shape shape' wh wh'. + intros conv notapp notapp' wh wh'. apply conv_cum_alt in conv as [(?&?&(r1&r2)&e)]. apply whnf_red_mkApps_l in r1 as [(?&?&->&?&?)]; auto. apply whnf_red_mkApps_l in r2 as [(?&?&->&?&?)]; auto. assert (isApp x1 = false). { erewrite whnf_red_isApp. 3: eauto. - 1: now destruct hd. + 1: assumption. apply whnf_mkApps_inv in wh; auto. - now destruct hd. } + now rewrite notapp. } assert (isApp x = false). { erewrite whnf_red_isApp. 3: eauto. - 1: now destruct hd'. + 1: assumption. apply whnf_mkApps_inv in wh'; auto. - now destruct hd'. } + now rewrite notapp'. } apply eq_termp_mkApps_inv in e as (?&?); auto. constructor. split. - - apply conv_cum_alt. - constructor. - exists x1, x. - split; [split|]; eauto with pcuic. - eapply eq_term_upto_univ_napp_nonind; eauto. - erewrite whnf_red_isIndConstructApp; [| |now eauto]. - 1: now destruct hd. - apply whnf_mkApps_inv in wh; auto. - now destruct hd. + - assert (isIndConstructApp x1 = isIndConstructApp hd). + { eapply whnf_red_isIndConstructApp; eauto. + apply whnf_mkApps_inv in wh; auto. + now rewrite notapp. } + destruct hd. + all: cbn. + 1-9, 12-15: apply conv_cum_alt. + 1-13: constructor. + 1-13: exists x1, x. + 1-13: split; [split|]; eauto with pcuic. + 1-13: (eapply eq_term_upto_univ_napp_nonind; [exact e|try exact H1]). + 1: discriminate notapp. + all: apply whnf_mkApps_inv in wh; auto. + all: eapply whnf_red_shape in r; auto. + all: depelim r. + all: apply whnf_mkApps_inv in wh'; auto; [|now rewrite notapp']. + all: eapply whnf_red_shape in r0; auto. + all: depelim e. + all: depelim r0. + all: apply All2_length in a. + all: constructor; constructor; rewrite a; auto. - clear -a1 a a0. induction a1 in args, args', x2, a, x3, a0, a1 |- *; depelim a; depelim a0; [now constructor|]. constructor. @@ -315,5 +325,20 @@ Section fixed. eapply PCUICConversion.cumul_red_r_inv; [eauto| |eauto]. auto. Qed. - + + Lemma conv_cum_red_inv' leq Γ Γ' t1 t2 t1' t2' : + conv_context Σ Γ Γ' -> + conv_cum leq Σ Γ t1 t2 -> + red Σ Γ t1 t1' -> + red Σ Γ' t2 t2' -> + conv_cum leq Σ Γ t1' t2'. + Proof. + intros cc r1 r2 conv_ctx. + destruct wf. + eapply conv_cum_red_inv; [|now eauto|reflexivity]. + eapply conv_cum_conv_ctx; eauto. + 2: apply conv_context_sym; eauto. + eapply conv_cum_red_inv; [|reflexivity|now eauto]. + eapply conv_cum_conv_ctx; eauto. + Qed. End fixed. diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index 53ba61851..eeb47b8cf 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -22,6 +22,7 @@ Module RedFlags. }. Definition default := mk true true true true true true. + Definition nodelta := mk true true true false true true. End RedFlags. @@ -514,6 +515,59 @@ Proof. solve_discr. Qed. + +Lemma whnf_mkApps_tSort_inv f Σ Γ s args : + whnf f Σ Γ (mkApps (tSort s) args) -> args = []. +Proof. + intros wh. + inversion wh; solve_discr. + clear -H. + remember (mkApps (tSort s) args) eqn:teq. + exfalso. + revert teq. + induction H in args |- *; intros; solve_discr. + destruct args as [|? ? _] using List.rev_ind; [easy|]. + rewrite <- mkApps_nested in teq. + cbn in teq. + inv teq. + eauto. +Qed. + +Lemma whnf_mkApps_tProd_inv f Σ Γ na A B args : + whnf f Σ Γ (mkApps (tProd na A B) args) -> args = []. +Proof. + intros wh. + inversion wh; solve_discr. + clear -H. + remember (mkApps (tProd na A B) args) eqn:teq. + exfalso. + revert teq. + induction H in args |- *; intros; solve_discr. + destruct args as [|? ? _] using List.rev_ind; [easy|]. + rewrite <- mkApps_nested in teq. + cbn in teq. + inv teq. + eauto. +Qed. + +Lemma whnf_mkApps_tLambda_inv f Σ Γ na A b args : + RedFlags.beta f = true -> + whnf f Σ Γ (mkApps (tLambda na A b) args) -> args = []. +Proof. + intros nobeta wh. + inversion wh; solve_discr. + clear -H nobeta. + remember (mkApps (tLambda na A b) args) eqn:teq. + exfalso. + revert teq. + induction H in args |- *; intros; solve_discr. + destruct args as [|? ? _] using List.rev_ind; [easy|]. + rewrite <- mkApps_nested in teq. + cbn in teq. + inv teq. + eauto. +Qed. + Section whne_red1_ind. Context (flags : RedFlags.t). Context (Σ : global_env). diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index db29db8ec..dc16ad281 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -17,8 +17,6 @@ Local Set Keyed Unification. Set Default Goal Selector "!". -Definition nodelta_flags := RedFlags.mk true true true false true true. - (* TODO MOVE *) Lemma All2_app_inv_both : forall A B (P : A -> B -> Type) l1 l2 r1 r2, diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 4310392e3..ea40d7f86 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -5,7 +5,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICReflect PCUICLiftSubst PCUICUnivSubst PCUICTyping PCUICCumulativity PCUICEquality PCUICConversion PCUICSafeLemmata PCUICNormal PCUICInversion PCUICReduction PCUICPosition - PCUICContextConversion PCUICSN PCUICUtils PCUICWeakening + PCUICPrincipality PCUICContextConversion PCUICSN PCUICUtils PCUICWeakening PCUICConversionInversion. From MetaCoq.SafeChecker Require Import PCUICSafeReduce. @@ -518,16 +518,19 @@ Section Conversion. simpl in *. eapply R_aux_stateR. all: simpl. all: auto. Qed. - - Definition leqb_term := - eqb_term_upto_univ Σ (check_eqb_universe G) (check_leqb_universe G). - - Definition eqb_term := - eqb_term_upto_univ Σ (check_eqb_universe G) (check_eqb_universe G). - - Lemma leqb_term_spec t u : - leqb_term t u -> - leq_term Σ (global_ext_constraints Σ) t u. + + Definition conv_pb_relb pb := + match pb with + | Conv => check_eqb_universe G + | Cumul => check_leqb_universe G + end. + + Definition eqb_termp_napp pb napp := + eqb_term_upto_univ_napp Σ (check_eqb_universe G) (conv_pb_relb pb) napp. + + Lemma eqb_termp_napp_spec pb napp t u : + eqb_termp_napp pb napp t u -> + eq_termp_napp Σ pb napp t u. Proof. pose proof hΣ'. apply eqb_term_upto_univ_impl. @@ -537,28 +540,35 @@ Section Conversion. + now eapply global_ext_uctx_consistent. + assumption. - intros u1 u2. - eapply (check_leqb_universe_spec' G (global_ext_uctx Σ)). - + now eapply wf_ext_global_uctx_invariants. - + now eapply global_ext_uctx_consistent. - + assumption. + destruct pb. + + eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). + * now eapply wf_ext_global_uctx_invariants. + * now eapply global_ext_uctx_consistent. + * assumption. + + eapply (check_leqb_universe_spec' G (global_ext_uctx Σ)). + * now eapply wf_ext_global_uctx_invariants. + * now eapply global_ext_uctx_consistent. + * assumption. Qed. - + + Notation eqb_termp pb := (eqb_termp_napp pb 0). + Notation eqb_term := (eqb_termp Conv). + Notation leqb_term := (eqb_termp Cumul). + + Lemma leqb_term_spec t u : + leqb_term t u -> + leq_term Σ (global_ext_constraints Σ) t u. + Proof. + intros. + now apply eqb_termp_napp_spec in H. + Qed. + Lemma eqb_term_spec t u : eqb_term t u -> eq_term Σ t u. Proof. - pose proof hΣ'. - apply eqb_term_upto_univ_impl. - - intros u1 u2. - eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). - + now eapply wf_ext_global_uctx_invariants. - + now eapply global_ext_uctx_consistent. - + assumption. - - intros u1 u2. - eapply (check_eqb_universe_spec' G (global_ext_uctx Σ)). - + now eapply wf_ext_global_uctx_invariants. - + now eapply global_ext_uctx_consistent. - + assumption. + intros. + now apply eqb_termp_napp_spec in H. Qed. Lemma eqb_term_refl : @@ -625,24 +635,7 @@ Section Conversion. Definition leqb_term_stack t1 π1 t2 π2 := eqb_ctx (stack_context π1) (stack_context π2) && leqb_term (zipp t1 π1) (zipp t2 π2). - - Definition eqb_termp leq u v := - match leq with - | Conv => eqb_term u v - | Cumul => leqb_term u v - end. - - Lemma eqb_termp_spec : - forall leq u v Γ, - eqb_termp leq u v -> - conv_cum leq Σ Γ u v. - Proof. - intros leq u v Γ e. - destruct leq. - - simpl. constructor. constructor. eapply eqb_term_spec. assumption. - - simpl. constructor. constructor. eapply leqb_term_spec. assumption. - Qed. - + Lemma leqb_term_stack_spec : forall Γ t1 π1 t2 π2, leqb_term_stack t1 π1 t2 π2 -> @@ -738,7 +731,7 @@ Section Conversion. (Γ1 : context) (t1 : term) (args1 l1 : list term) (Γ2 : context) (t2 : term) (l2 : list term) - | HeadMistmatch + | HeadMismatch (leq : conv_pb) (Γ1 : context) (t1 : term) (Γ2 : context) (t2 : term). @@ -768,7 +761,7 @@ Section Conversion. Definition isred_full Γ t π := isApp t = false /\ - whnf nodelta_flags Σ (Γ,,, stack_context π) (zipp t π). + whnf RedFlags.nodelta Σ (Γ,,, stack_context π) (zipp t π). Lemma isred_full_nobeta Γ t π : isred_full Γ t π -> @@ -826,25 +819,42 @@ Section Conversion. rewrite <- mkApps_nested. now destruct decompose_stack. Qed. + + (* Tailored view for isconv_prog *) + Equations prog_discr (t1 t2 : term) : Prop := + prog_discr (tApp _ _) (tApp _ _) := False ; + prog_discr (tConst _ _) (tConst _ _) := False ; + prog_discr (tLambda _ _ _) (tLambda _ _ _) := False ; + prog_discr (tProd _ _ _) (tProd _ _ _) := False ; + prog_discr (tCase _ _ _ _) (tCase _ _ _ _) := False ; + prog_discr (tProj _ _) (tProj _ _) := False ; + prog_discr (tFix _ _) (tFix _ _) := False ; + prog_discr (tCoFix _ _) (tCoFix _ _) := False ; + prog_discr _ _ := True. Definition Ret s Γ t π t' π' := forall (leq : conv_pb), conv_stack_ctx Γ π π' -> match s with - | Fallback - | Term => isred_full Γ t π - | _ => True - end -> - match s with - | Fallback - | Term => isred_full Γ t' π' - | _ => True - end -> - match s with - | Args => conv_cum leq Σ (Γ ,,, stack_context π) t t' - | _ => True - end -> - ConversionResult (conv_term leq Γ t π t' π'). + | Reduction => + ConversionResult (conv_term leq Γ t π t' π') + + | Term => + isred_full Γ t π -> + isred_full Γ t' π' -> + ConversionResult (conv_term leq Γ t π t' π') + + | Fallback => + isred_full Γ t π -> + isred_full Γ t' π' -> + prog_discr t t' -> + ConversionResult (conv_term leq Γ t π t' π') + + | Args => + ConversionResult (∥conv_terms Σ (Γ,,, stack_context π) + (decompose_stack π).1 + (decompose_stack π').1∥) + end. Definition Aux s Γ t1 π1 t2 π2 h2 := forall s' t1' π1' t2' π2' @@ -869,13 +879,13 @@ Section Conversion. (only parsing). Notation isconv_red_raw leq t1 π1 t2 π2 aux := - (aux Reduction t1 π1 t2 π2 _ _ _ _ leq _ I I I) (only parsing). + (aux Reduction t1 π1 t2 π2 _ _ _ _ leq _) (only parsing). Notation isconv_prog_raw leq t1 π1 t2 π2 aux := - (aux Term t1 π1 t2 π2 _ _ _ _ leq _ _ _ I) (only parsing). + (aux Term t1 π1 t2 π2 _ _ _ _ leq _ _ _) (only parsing). Notation isconv_args_raw leq t1 π1 t2 π2 aux := - (aux Args t1 π1 t2 π2 _ _ _ _ leq _ I I _) (only parsing). + (aux Args t1 π1 t2 π2 _ _ _ _ leq _) (only parsing). Notation isconv_fallback_raw leq t1 π1 t2 π2 aux := - (aux Fallback t1 π1 t2 π2 _ _ _ _ leq _ _ _ I) (only parsing). + (aux Fallback t1 π1 t2 π2 _ _ _ _ leq _ _ _ _) (only parsing). Notation isconv_red leq t1 π1 t2 π2 aux := (repack (isconv_red_raw leq t1 π1 t2 π2 aux)) (only parsing). @@ -897,9 +907,9 @@ Section Conversion. with inspect (decompose_stack π1) := { | @exist (args1, ρ1) e1 with inspect (decompose_stack π2) := { | @exist (args2, ρ2) e2 - with inspect (reduce_stack nodelta_flags Σ hΣ (Γ ,,, stack_context π1) t1 (appstack args1 ε) _) := { + with inspect (reduce_stack RedFlags.nodelta Σ hΣ (Γ ,,, stack_context π1) t1 (appstack args1 ε) _) := { | @exist (t1',π1') eq1 - with inspect (reduce_stack nodelta_flags Σ hΣ (Γ ,,, stack_context π2) t2 (appstack args2 ε) _) := { + with inspect (reduce_stack RedFlags.nodelta Σ hΣ (Γ ,,, stack_context π2) t2 (appstack args2 ε) _) := { | @exist (t2',π2') eq2 => isconv_prog leq t1' (π1' +++ ρ1) t2' (π2' +++ ρ2) aux } } @@ -920,7 +930,7 @@ Section Conversion. match type of eq1 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r1] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d1 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d1 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c1 end. rewrite <- eq1 in r1. @@ -951,7 +961,7 @@ Section Conversion. match type of eq2 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 end. rewrite <- eq2 in r2. @@ -979,7 +989,7 @@ Section Conversion. Next Obligation. match type of eq1 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d1 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d1 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c1 end. rewrite <- eq1 in d1. cbn in d1. @@ -993,7 +1003,7 @@ Section Conversion. - assert (ee1 := eq1). rewrite e in ee1. inversion ee1. subst. match type of eq2 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 end. rewrite <- eq2 in d2. cbn in d2. @@ -1064,7 +1074,7 @@ Section Conversion. Next Obligation. match type of eq1 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d1 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d1 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c1 end. rewrite <- eq1 in d1. cbn in d1. @@ -1073,7 +1083,7 @@ Section Conversion. pose proof (decompose_stack_eq _ _ _ (eq_sym e1)). subst. match type of eq2 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 end. rewrite <- eq2 in d2. cbn in d2. @@ -1132,7 +1142,7 @@ Section Conversion. match type of eq1 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r1] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d1 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d1 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c1 end. rewrite <- eq1 in r1. @@ -1143,7 +1153,7 @@ Section Conversion. match type of eq2 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 end. rewrite <- eq2 in r2. @@ -1204,7 +1214,7 @@ Section Conversion. match type of eq1 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r1] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d1 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d1 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c1 end. rewrite <- eq1 in r1. @@ -1215,7 +1225,7 @@ Section Conversion. match type of eq2 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 end. rewrite <- eq2 in r2. @@ -1576,18 +1586,6 @@ Section Conversion. now destruct t2. Qed. - (* Tailored view for isconv_prog *) - Equations prog_discr (t1 t2 : term) : Prop := - prog_discr (tApp _ _) (tApp _ _) := False ; - prog_discr (tConst _ _) (tConst _ _) := False ; - prog_discr (tLambda _ _ _) (tLambda _ _ _) := False ; - prog_discr (tProd _ _ _) (tProd _ _ _) := False ; - prog_discr (tCase _ _ _ _) (tCase _ _ _ _) := False ; - prog_discr (tProj _ _) (tProj _ _) := False ; - prog_discr (tFix _ _) (tFix _ _) := False ; - prog_discr (tCoFix _ _) (tCoFix _ _) := False ; - prog_discr _ _ := True. - Inductive prog_view : term -> term -> Type := | prog_view_App u1 v1 u2 v2 : prog_view (tApp u1 v1) (tApp u2 v2) @@ -1739,6 +1737,22 @@ Section Conversion. Proof. now intros dar [|[(?&?&?&?)]]. Qed. + + Lemma wellformed_nonarity_mkApps Γ hd args : + destArity [] hd = None -> + wellformed Σ Γ (mkApps hd args) -> + welltyped Σ Γ (mkApps hd args). + Proof. + intros dar wf. + destruct wf as [|[(?&?&dar'&?)]]; auto. + exfalso. + clear -dar dar'. + induction args in args, hd, dar, dar' |- *. + - cbn in *. + congruence. + - cbn in *. + now apply IHargs in dar'; [|easy]. + Qed. Lemma wellformed_zipc_tConst_inv Γ c u π : wellformed Σ Γ (zipc (tConst c u) π) -> @@ -2549,7 +2563,7 @@ Section Conversion. Lemma whnf_whne_upgrade Γ t : whnf RedFlags.default Σ Γ t -> - whne nodelta_flags Σ Γ t -> + whne RedFlags.nodelta Σ Γ t -> whne RedFlags.default Σ Γ t. Proof. intros whn whe. @@ -2699,6 +2713,98 @@ Section Conversion. Qed. *) (* See https://github.com/coq/coq/blob/master/kernel/reduction.ml#L367 *) + + Lemma cumul_mkApps Γ hd args hd' args' : + wf Σ -> + Σ;;; Γ |- hd <= hd' -> + conv_terms Σ Γ args args' -> + Σ;;; Γ |- mkApps hd args <= mkApps hd' args'. + Proof. + intros wf cum cum_args. + revert hd hd' cum. + induction cum_args; intros hd hd' cum; auto. + cbn. + apply IHcum_args. + eapply cumul_trans; auto. + - eapply cumul_App_l; eauto. + - eapply cumul_App_r; eauto. + Qed. + + Lemma conv_cum_zipp leq Γ t t' π π' : + conv_cum leq Σ Γ t t' -> + ∥conv_terms Σ Γ (decompose_stack π).1 (decompose_stack π').1∥ -> + conv_cum leq Σ Γ (zipp t π) (zipp t' π'). + Proof. + intros conv conv_args. + rewrite !zipp_as_mkApps. + destruct leq; cbn in *. + - sq. + apply mkApps_conv_args; auto. + - sq. + apply cumul_mkApps; auto. + Qed. + + Lemma conv_terms_conv_ctx Γ Γ' ts ts' : + wf Σ -> + conv_terms Σ Γ ts ts' -> + conv_context Σ Γ Γ' -> + conv_terms Σ Γ' ts ts'. + Proof. + intros wf conv ctx. + induction conv; [constructor|]. + constructor; auto. + eapply conv_conv_ctx; eauto. + Qed. + + Lemma conv_terms_red Γ ts ts' tsr tsr' : + wf Σ -> + conv_terms Σ Γ tsr tsr' -> + All2 (red Σ Γ) ts tsr -> + All2 (red Σ Γ) ts' tsr' -> + conv_terms Σ Γ ts ts'. + Proof. + intros wf conv all all'. + induction conv in ts, ts', all, all' |- *; depelim all; depelim all'; [constructor|]. + constructor; [|auto]. + eapply red_conv_conv; eauto. + symmetry. + eapply red_conv_conv; eauto. + symmetry. + eauto. + Qed. + + Lemma conv_terms_red' Γ Γ' ts ts' tsr tsr' : + wf Σ -> + conv_context Σ Γ Γ' -> + conv_terms Σ Γ tsr tsr' -> + All2 (red Σ Γ) ts tsr -> + All2 (red Σ Γ') ts' tsr' -> + conv_terms Σ Γ ts ts'. + Proof. + intros wf convctx conv all all2. + eapply conv_terms_red. + 1: eauto. + 2: eauto. + 2: apply All2_same; reflexivity. + eapply conv_terms_conv_ctx; eauto. + 2: eapply conv_context_sym; eauto. + eapply conv_terms_red. + 1: eauto. + 3: eauto. + 2: apply All2_same; reflexivity. + eapply conv_terms_conv_ctx; eauto. + Qed. + + (* Check if the heads of an application are equal, and if so, + if all args are convertible *) + Notation eqb_heads_and_convertible_args leq hd hd' π π' aux := + (if eqb_term hd hd' then + if isconv_args_raw leq hd π hd' π' aux then + true + else + false + else + false). Opaque reduce_stack. Equations(noeqns) _isconv_prog (Γ : context) (leq : conv_pb) @@ -2762,8 +2868,12 @@ Section Conversion. } ; | prog_view_Case ind par p c brs ind' par' p' c' brs' - with inspect (eqb_term (tCase (ind, par) p c brs) (tCase (ind', par') p' c' brs')) := { - | @exist true eq1 := isconv_args leq (tCase (ind, par) p c brs) π1 (tCase (ind', par') p' c' brs') π2 aux ; + with inspect (eqb_heads_and_convertible_args + leq + (tCase (ind, par) p c brs) + (tCase (ind', par') p' c' brs') + π1 π2 aux) := { + | @exist true eq1 := Success _; | @exist false _ with inspect (reduce_term RedFlags.default Σ hΣ (Γ ,,, stack_context π1) c _) := { | @exist cred eq1 with inspect (reduce_term RedFlags.default Σ hΣ (Γ ,,, stack_context π2) c' _) := { | @exist cred' eq2 with inspect (eqb_term cred c && eqb_term cred' c') := { @@ -2779,11 +2889,9 @@ Section Conversion. c' (Case (ind', par') p' brs' π2) aux := { | Success h2 with isconv_branches' Γ ind par p c brs π1 _ ind' par' p' c' brs' π2 _ _ _ _ aux := { - | Success h3 - with isconv_args_raw leq (tCase (ind, par) p c brs) π1 (tCase (ind', par') p' c' brs') π2 aux := { - | Success h4 := yes ; - | Error e h := Error e h - } ; + | Success h3 := isconv_args leq + (tCase (ind, par) p c brs) π1 + (tCase (ind', par') p' c' brs') π2 aux; | Error e h := Error e _ } ; | Error e h := Error e _ @@ -2837,13 +2945,17 @@ Section Conversion. } ; | prog_view_Fix mfix idx mfix' idx' - with inspect (eqb_term (tFix mfix idx) (tFix mfix' idx')) := { - | @exist true eq1 := isconv_args leq (tFix mfix idx) π1 (tFix mfix' idx') π2 aux ; + with inspect (eqb_heads_and_convertible_args + leq + (tFix mfix idx) + (tFix mfix' idx') + π1 π2 aux) := { + | @exist true eq1 := Success _; | @exist false eqDiff with inspect (unfold_one_fix Γ mfix idx π1 _) := { | @exist (Some (fn, θ)) eq1 with inspect (decompose_stack θ) := { | @exist (l', θ') eq2 - with inspect (reduce_stack nodelta_flags Σ hΣ (Γ ,,, stack_context θ') fn (appstack l' ε) _) := { + with inspect (reduce_stack RedFlags.nodelta Σ hΣ (Γ ,,, stack_context θ') fn (appstack l' ε) _) := { | @exist (fn', ρ) eq3 := isconv_prog leq fn' (ρ +++ θ') (tFix mfix' idx') π2 aux } @@ -2852,7 +2964,7 @@ Section Conversion. | @exist (Some (fn, θ)) eq1 with inspect (decompose_stack θ) := { | @exist (l', θ') eq2 - with inspect (reduce_stack nodelta_flags Σ hΣ (Γ ,,, stack_context θ') fn (appstack l' ε) _) := { + with inspect (reduce_stack RedFlags.nodelta Σ hΣ (Γ ,,, stack_context θ') fn (appstack l' ε) _) := { | @exist (fn', ρ) eq3 := isconv_prog leq (tFix mfix idx) π1 fn' (ρ +++ θ') aux } @@ -2902,6 +3014,7 @@ Section Conversion. simpl. constructor. Qed. Next Obligation. + apply conv_cum_zipp; auto. destruct hΣ. eapply conv_conv_cum. constructor. constructor. @@ -2936,6 +3049,36 @@ Section Conversion. eapply red_zipp. eapply red_const. eassumption. Qed. + Next Obligation. + apply h; cbn; clear h. + eapply conv_cum_red_inv. + - exact hΣ'. + - exact H. + - apply red_zipp. + eapply red_const; eauto. + - apply red_zipp. + eapply red_const; eauto. + Qed. + Next Obligation. + apply h; clear h. + rewrite !zipp_as_mkApps in H. + apply conv_cum_mkApps_inv in H as [(?&?)]; eauto. + - now constructor. + - apply whnf_mkApps. + eapply whne_const; eauto. + - apply whnf_mkApps. + eapply whne_const; eauto. + Qed. + Next Obligation. + apply wellformed_zipc_tConst_inv in h1 as (?&?&?). + unfold declared_constant in *. + congruence. + Qed. + Next Obligation. + apply wellformed_zipc_tConst_inv in h1 as (?&?&?). + unfold declared_constant in *. + congruence. + Qed. Next Obligation. right; split; [easy|]. now rewrite <- uneq_u. @@ -3137,12 +3280,16 @@ Section Conversion. simpl. constructor. Qed. Next Obligation. + destruct eqb_term eqn:eq_terms; [|congruence]. + destruct aux eqn:conv_args; [|congruence]. + apply conv_cum_zipp; auto. destruct hΣ. eapply conv_conv_cum. constructor. constructor. eapply eqb_term_spec. auto. Qed. Next Obligation. + clear wildcard7. destruct hΣ as [wΣ]. zip fold in h1. apply wellformed_context in h1 ; auto. simpl in h1. destruct h1 as [[T h1] | [[ctx [s [h1 _]]]]] ; [| discriminate ]. @@ -3153,7 +3300,7 @@ Section Conversion. Qed. Next Obligation. destruct hΣ as [wΣ]. - clear aux. + clear aux wildcard7. zip fold in h2. apply wellformed_context in h2 ; auto. simpl in h2. destruct h2 as [[T h2] | [[ctx [s [h2 _]]]]] ; [| discriminate ]. apply inversion_Case in h2 as hh ; auto. @@ -3193,13 +3340,14 @@ Section Conversion. destruct h1 as [h1], h2 as [h2], h3 as [h3]. unfold zipp in h1, h2. simpl in h1, h2. destruct hΣ as [wΣ]. - eapply conv_conv_cum. - constructor. change (eq_inductive ind ind') with (eqb ind ind') in eq4. destruct (eqb_spec ind ind'). 2: discriminate. change (Nat.eqb par par') with (eqb par par') in eq4. destruct (eqb_spec par par'). 2: discriminate. subst. + eapply conv_cum_zipp; auto. + eapply conv_conv_cum. + constructor. eapply conv_Case. all: assumption. Qed. Next Obligation. @@ -3295,14 +3443,34 @@ Section Conversion. now constructor. Qed. Next Obligation. - (* Proof idea: - 1) from H, derive that some - tCase (ind, par) _ _ _ and tCase (ind',par') _ _ _ - are in eqb_term - maybe using conv_alt_red - 2) deduce that ind = ind' and par = par' - 3) conclude that it contradicts eqDiff - *) + apply h; cbn; clear h. + destruct ir1 as (_&wh1), ir2 as (_&wh2). + cbn in *. + symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). + apply eqb_term_spec in c_is_red. + apply eqb_term_spec in c'_is_red. + eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. + eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. + rewrite !zipp_as_mkApps in H. + rewrite zipp_as_mkApps in wh1, wh2. + apply whnf_mkApps_inv in wh1; [|easy]. + apply whnf_mkApps_inv in wh2; [|easy]. + depelim wh1; solve_discr. + depelim wh2; solve_discr. + depelim H; cbn in *; try easy; solve_discr. + depelim H0; cbn in *; try easy; solve_discr. + apply whnf_whne_upgrade in c_is_red; auto. + apply whnf_whne_upgrade in c'_is_red; auto. + destruct hΣ, hx. + assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). + { eapply whne_conv_context; eauto. + apply conv_context_sym; eauto. } + apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. + 2-3: apply whnf_mkApps; eauto using whne. + apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. + constructor; auto. + Qed. + Next Obligation. destruct ir1 as (_&wh1), ir2 as (_&wh2). cbn in *. symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). @@ -3518,10 +3686,10 @@ Section Conversion. simpl. constructor. Qed. Next Obligation. - (* destruct hΣ. *) - destruct h1 as [h]. + destruct h1 as [h1]. change (true = eqb p p') in eq5. destruct (eqb_spec p p'). 2: discriminate. subst. + apply conv_cum_zipp; auto. eapply conv_conv_cum. constructor. eapply conv_Proj_c. assumption. @@ -3557,6 +3725,37 @@ Section Conversion. apply conv_cum_tProj_inv in conv_proj as [(->&?)]; auto. constructor; auto. Qed. + Next Obligation. + apply h; cbn; clear h. + destruct ir1 as (_&wh1), ir2 as (_&wh2). + cbn in *. + rename eq3 into c_is_red. + rename eq4 into c'_is_red. + symmetry in c'_is_red. + apply eqb_term_spec in c'_is_red. + eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. + symmetry in c_is_red. + apply eqb_term_spec in c_is_red. + eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. + rewrite !zipp_as_mkApps in H. + rewrite zipp_as_mkApps in wh1, wh2. + apply whnf_mkApps_inv in wh1; [|easy]. + apply whnf_mkApps_inv in wh2; [|easy]. + depelim wh1; solve_discr. + depelim wh2; solve_discr. + depelim H; cbn in *; try easy; solve_discr. + depelim H0; cbn in *; try easy; solve_discr. + apply whnf_whne_upgrade in c_is_red; auto. + apply whnf_whne_upgrade in c'_is_red; auto. + destruct hΣ, hx. + assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). + { eapply whne_conv_context; eauto. + apply conv_context_sym; eauto. } + apply conv_cum_mkApps_inv in H1 as [(conv_proj&conv_args)]; auto. + 2-3: apply whnf_mkApps; eauto using whne. + apply conv_cum_tProj_inv in conv_proj as [(->&?)]; auto. + constructor; auto. + Qed. Next Obligation. (* Proof idea: c and c' are whne so from H, p = p' contradicting eqDiff *) destruct ir1 as (_&wh1), ir2 as (_&wh2). @@ -3741,6 +3940,9 @@ Section Conversion. Qed. Next Obligation. destruct hΣ. + destruct eqb_term eqn:eq_term; [|congruence]. + destruct aux eqn:eq_args; [|congruence]. + apply conv_cum_zipp; auto. eapply conv_conv_cum. constructor. constructor. eapply eqb_term_spec. auto. @@ -3768,7 +3970,7 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 end. rewrite <- eq3 in r2. cbn in r2. rewrite zipc_appstack in r2. cbn in r2. @@ -3808,7 +4010,7 @@ Section Conversion. rewrite <- eq2 in d1. simpl in d1. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. cbn in d2. @@ -3821,14 +4023,14 @@ Section Conversion. case_eq (decompose_stack π1). intros args ρ e'. simpl. apply decompose_stack_eq in e'. subst. - clear eq3. + clear eq3 eqDiff. rewrite stack_context_appstack in hx. assumption. Qed. Next Obligation. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as ir; pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w2 end. @@ -3850,7 +4052,7 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in r2. rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. @@ -3895,7 +4097,7 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in r2. rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. @@ -3957,7 +4159,7 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 end. rewrite <- eq3 in r2. cbn in r2. rewrite zipc_appstack in r2. cbn in r2. @@ -3996,7 +4198,7 @@ Section Conversion. apply unfold_one_fix_decompose in eq1 as d1. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. @@ -4010,14 +4212,14 @@ Section Conversion. rewrite stack_cat_appstack. rewrite stack_context_appstack. apply decompose_stack_eq in e2 as ?. subst. - clear eq3. + clear eq3 eqDiff. rewrite stack_context_appstack in hx. assumption. Qed. Next Obligation. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as ir; pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w2 end. @@ -4038,7 +4240,7 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in r2. rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. @@ -4094,7 +4296,7 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in r2. rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. @@ -4152,35 +4354,50 @@ Section Conversion. Next Obligation. destruct h1 as [h1]. destruct hΣ. - eapply conv_conv_cum_l. change (true = eqb idx idx') in eq4. destruct (eqb_spec idx idx'). 2: discriminate. subst. + apply conv_cum_zipp; auto. + eapply conv_conv_cum_l. eapply conv_Fix. all: assumption. Qed. Next Obligation. apply h; clear h. (* Idea : Since the fixpoints do not unfold they are convertible to fixpoints - in whnf implying that their definitions are convertible. *) + in whnf implying that the args are convertible. *) rewrite !zipp_as_mkApps in H. apply unfold_one_fix_None in unfold_fix1 as [(?&?&?)]. apply unfold_one_fix_None in unfold_fix2 as [(?&?&?)]. destruct hΣ, hx. - eapply conv_cum_conv_ctx in H; eauto. - eapply conv_cum_red_inv in H. + eapply conv_cum_red_inv' in H. 2: exact hΣ'. - 2: reflexivity. + 2: apply red_mkApps; [reflexivity|exact a]. 2: apply red_mkApps; [reflexivity|exact a0]. - eapply conv_cum_conv_ctx in H; eauto. - 2: apply conv_context_sym; eauto. - eapply conv_cum_red_inv in H. + 2: eassumption. + apply conv_cum_mkApps_inv in H as [(_&conv_args)]; auto. + 2: eapply whnf_conv_context; eauto. + 2: eapply conv_context_sym; eauto. + constructor. + eapply conv_terms_red'; eauto. + Qed. + Next Obligation. + apply h; clear h. + (* Idea : Since the fixpoints do not unfold they are convertible to fixpoints + in whnf implying that their definitions are convertible. *) + rewrite !zipp_as_mkApps in H. + apply unfold_one_fix_None in unfold_fix1 as [(?&?&?)]. + apply unfold_one_fix_None in unfold_fix2 as [(?&?&?)]. + destruct hΣ, hx. + eapply conv_cum_red_inv' in H. 2: exact hΣ'. 2: apply red_mkApps; [reflexivity|exact a]. - 2: reflexivity. + 2: apply red_mkApps; [reflexivity|exact a0]. + 2: eassumption. apply conv_cum_mkApps_inv in H as [(?&?)]; auto. 2: eapply whnf_conv_context; eauto. 2: eapply conv_context_sym; eauto. - apply conv_cum_tFix_inv in c as [(->&?)]. + cbn in c0. + apply conv_cum_tFix_inv in c0 as [(->&?)]. constructor. eapply All2_impl; eauto. cbn; intros; easy. @@ -4192,21 +4409,15 @@ Section Conversion. apply unfold_one_fix_None in unfold_fix1 as [(?&?&?)]. apply unfold_one_fix_None in unfold_fix2 as [(?&?&?)]. destruct hΣ, hx. - eapply conv_cum_conv_ctx in H; eauto. - eapply conv_cum_red_inv in H. - 2: exact hΣ'. - 2: reflexivity. - 2: apply red_mkApps; [reflexivity|exact a0]. - eapply conv_cum_conv_ctx in H; eauto. - 2: apply conv_context_sym; eauto. - eapply conv_cum_red_inv in H. + eapply conv_cum_red_inv' in H. 2: exact hΣ'. 2: apply red_mkApps; [reflexivity|exact a]. - 2: reflexivity. + 2: apply red_mkApps; [reflexivity|exact a0]. + 2: eauto. apply conv_cum_mkApps_inv in H as [(?&?)]; auto. 2: eapply whnf_conv_context; eauto. 2: eapply conv_context_sym; eauto. - apply conv_cum_tFix_inv in c as [(->&?)]. + apply conv_cum_tFix_inv in c0 as [(->&?)]. rewrite Nat.eqb_refl in idx_uneq. congruence. Qed. @@ -4219,13 +4430,18 @@ Section Conversion. Qed. Next Obligation. destruct hΣ. + apply conv_cum_zipp; auto. eapply conv_conv_cum. constructor. constructor. eapply eqb_term_spec. auto. Qed. (* Beware !! There is a TODO in the code *) Next Obligation. - todo "Completeness"%string. + apply h; clear h. + todo "Cofix". + Qed. + Next Obligation. + todo "Cofix". Qed. (* Fallback *) @@ -4266,7 +4482,16 @@ Section Conversion. pzt x = pzt y /\ positionR (` (pps1 x)) (` (pps1 y)) -> Ret Reduction Γ u1 (coApp (mkApps t1 ca1) (appstack a1 π1)) u2 ρ2. - + + Lemma appstack_cons a args π : + appstack (a :: args) π = App a (appstack args π). + Proof. reflexivity. Qed. + + Lemma fst_decompose_stack_nil π : + isStackApp π = false -> + (decompose_stack π).1 = []. + Proof. now destruct π. Qed. + Equations(noeqns) _isconv_args' (leq : conv_pb) (Γ : context) (t1 : term) (args1 : list term) (l1 : list term) (π1 : stack) @@ -4277,12 +4502,11 @@ Section Conversion. (h2 : wtp Γ t2 (appstack l2 π2)) (hπ2 : isStackApp π2 = false) (hx : conv_stack_ctx Γ π1 π2) - (h : conv_cum leq Σ (Γ ,,, stack_context π1) (mkApps t1 args1) t2) (aux : Aux' Γ t1 args1 l1 π1 t2 (appstack l2 π2) h2) - : ConversionResult (conv_term leq Γ (mkApps t1 args1) (appstack l1 π1) t2 (appstack l2 π2)) by struct l1 := - _isconv_args' leq Γ t1 args1 (u1 :: l1) π1 h1 hπ1 t2 (u2 :: l2) π2 h2 hπ2 hx h aux - with aux u1 u2 args1 l1 (coApp t2 (appstack l2 π2)) _ _ _ _ Conv _ I I I := { - | Success H1 with _isconv_args' leq Γ t1 (args1 ++ [u1]) l1 π1 _ _ (tApp t2 u2) l2 π2 _ _ _ _ _ := { + : ConversionResult (∥conv_terms Σ (Γ,,, stack_context π1) l1 l2∥) by struct l1 := + _isconv_args' leq Γ t1 args1 (u1 :: l1) π1 h1 hπ1 t2 (u2 :: l2) π2 h2 hπ2 hx aux + with aux u1 u2 args1 l1 (coApp t2 (appstack l2 π2)) _ _ _ _ Conv _ := { + | Success H1 with _isconv_args' leq Γ t1 (args1 ++ [u1]) l1 π1 _ _ (tApp t2 u2) l2 π2 _ _ _ _ := { | Success H2 := yes ; | Error e herr := Error ( @@ -4301,31 +4525,22 @@ Section Conversion. ) _ } ; - _isconv_args' leq Γ t1 args1 [] π1 h1 hπ1 t2 [] π2 h2 hπ2 hx h aux := yes ; + _isconv_args' leq Γ t1 args1 [] π1 h1 hπ1 t2 [] π2 h2 hπ2 hx aux := yes ; - _isconv_args' leq Γ t1 args1 l1 π1 h1 hπ1 t2 l2 π2 h2 hπ2 hx h aux := + _isconv_args' leq Γ t1 args1 l1 π1 h1 hπ1 t2 l2 π2 h2 hπ2 hx aux := Error ( StackMismatch (Γ ,,, stack_context π1) t1 args1 l1 (Γ ,,, stack_context π2) t2 l2 ) _. Next Obligation. - unfold zipp. - case_eq (decompose_stack π1). intros l1 ρ1 e1. - apply decompose_stack_eq in e1. subst. - apply isStackApp_false_appstack in hπ1. subst. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - apply decompose_stack_eq in e2. subst. - apply isStackApp_false_appstack in hπ2. subst. - simpl. - rewrite stack_context_appstack in h. - assumption. + constructor; constructor. Defined. Next Obligation. - todo "Completeness"%string. + destruct H as [H]; depelim H. Qed. Next Obligation. - todo "Completeness"%string. + destruct H as [H]; depelim H. Qed. Next Obligation. split. 1: reflexivity. @@ -4338,14 +4553,6 @@ Section Conversion. Next Obligation. rewrite <- mkApps_nested. assumption. Defined. - Next Obligation. - destruct hΣ as [wΣ]. - destruct H1 as [H1]. - unfold zipp in H1. simpl in H1. - rewrite stack_context_appstack in H1. - rewrite <- !mkApps_nested. simpl. - eapply App_conv'. all: assumption. - Defined. Next Obligation. simpl in H0. destruct H0 as [eq hp]. rewrite app_length in H. cbn in H. @@ -4370,47 +4577,34 @@ Section Conversion. destruct hΣ as [wΣ]. destruct H1 as [H1]. unfold zipp. simpl. - rewrite stack_context_appstack. - rewrite 2!decompose_stack_appstack. simpl. unfold zipp in H1. simpl in H1. rewrite stack_context_appstack in H1. - unfold zipp in H2. rewrite 2!decompose_stack_appstack in H2. - rewrite stack_context_appstack in H2. - rewrite <- !mkApps_nested in H2. cbn in H2. - rewrite <- !mkApps_nested. - case_eq (decompose_stack π1). intros a1 ρ1 e1. - rewrite e1 in H2. cbn in H2. cbn. - case_eq (decompose_stack π2). intros a2 ρ2 e2. - rewrite e2 in H2. cbn in H2. cbn. - pose proof (decompose_stack_eq _ _ _ e1). - pose proof (decompose_stack_eq _ _ _ e2). - subst. - rewrite 2!stack_context_appstack in hx. - rewrite stack_context_appstack in H1. - rewrite !stack_context_appstack in H2. - rewrite !stack_context_appstack in h. - rewrite stack_context_appstack. - assumption. + destruct H2 as [H2]. + constructor; constructor; auto. Defined. Next Obligation. apply herr; clear herr. - todo "Completeness"%string. + destruct H as [H]; depelim H. + constructor; auto. Qed. Next Obligation. - apply herr; clear herr. - todo "Completeness"%string. + apply herr; cbn; clear herr. + destruct H as [H]; depelim H. + rewrite stack_context_appstack. + constructor; auto. Qed. Equations(noeqns) _isconv_args (leq : conv_pb) (Γ : context) (t1 : term) (π1 : stack) (h1 : wtp Γ t1 π1) (t2 : term) (π2 : stack) (h2 : wtp Γ t2 π2) (hx : conv_stack_ctx Γ π1 π2) - (he : conv_cum leq Σ (Γ ,,, stack_context π1) t1 t2) (aux : Aux Args Γ t1 π1 t2 π2 h2) - : ConversionResult (conv_term leq Γ t1 π1 t2 π2) := - _isconv_args leq Γ t1 π1 h1 t2 π2 h2 hx he aux with inspect (decompose_stack π1) := { + : ConversionResult (∥conv_terms Σ (Γ,,, stack_context π1) + (decompose_stack π1).1 + (decompose_stack π2).1∥) := + _isconv_args leq Γ t1 π1 h1 t2 π2 h2 hx aux with inspect (decompose_stack π1) := { | @exist (l1, θ1) eq1 with inspect (decompose_stack π2) := { - | @exist (l2, θ2) eq2 with _isconv_args' leq Γ t1 [] l1 θ1 _ _ t2 l2 θ2 _ _ _ _ _ := { + | @exist (l2, θ2) eq2 with _isconv_args' leq Γ t1 [] l1 θ1 _ _ t2 l2 θ2 _ _ _ _ := { | Success h := yes ; | Error e h := Error e _ } @@ -4436,11 +4630,11 @@ Section Conversion. rewrite 2!stack_context_appstack in hx. assumption. Qed. - Next Obligation. + (*Next Obligation. pose proof (decompose_stack_eq _ _ _ (eq_sym eq1)). subst. rewrite stack_context_appstack in he. assumption. - Qed. + Qed.*) Next Obligation. specialize (aux Reduction) as h. cbn in h. eapply h. all: auto. @@ -4450,17 +4644,22 @@ Section Conversion. unshelve eapply R_positionR. 2: assumption. simpl. rewrite eq. reflexivity. Qed. + (* The obligation tactic wipes out a useful hypothesis here? *) + Obligation Tactic := idtac. Next Obligation. - pose proof (decompose_stack_eq _ _ _ (eq_sym eq1)). subst. - pose proof (decompose_stack_eq _ _ _ (eq_sym eq2)). subst. - assumption. + intros; cbn in *. + rewrite <- (stack_context_decompose π1). + rewrite <- eq1, <- eq2. + auto. Qed. Next Obligation. - apply h; clear h. - pose proof (decompose_stack_eq _ _ _ (eq_sym eq1)). subst. - pose proof (decompose_stack_eq _ _ _ (eq_sym eq2)). subst. - assumption. + intros; cbn in *. + rewrite <- (stack_context_decompose π1). + rewrite <- eq1, <- eq2. + auto. Qed. + Obligation Tactic := + Tactics.program_simplify; CoreTactics.equations_simpl; try Tactics.program_solve_wf. Equations unfold_one_case (Γ : context) (ind : inductive) (par : nat) (p c : term) (brs : list (nat × term)) @@ -4744,53 +4943,149 @@ Section Conversion. - reflexivity. - reflexivity. Qed. + + Lemma reducible_head_None Γ t π h : + isApp t = false -> + whnf RedFlags.nodelta Σ (Γ,,, stack_context π) (mkApps t (decompose_stack π).1) -> + None = reducible_head Γ t π h -> + ∥∑ t' args', + whnf_red_shape_spec Σ (Γ,,, stack_context π) t t' × + All2 (red Σ (Γ,,, stack_context π)) (decompose_stack π).1 args' × + whnf RedFlags.default Σ (Γ,,, stack_context π) (mkApps t' args')∥. + Proof. + funelim (reducible_head Γ t π h); intros notapp wh [=]. + - apply whnf_mkApps_inv in wh; auto. + depelim wh; solve_discr. + depelim H; solve_discr; try discriminate. + constructor; eexists _, (decompose_stack π).1. + split; [now constructor|]. + split; eauto with pcuic. + apply whnf_mkApps. + eauto with pcuic. + - apply whnf_mkApps_inv in wh; auto. + depelim wh; solve_discr. + constructor; eexists _, (decompose_stack π).1. + split; [constructor|]. + split; eauto with pcuic. + apply whnf_mkApps. + constructor. + - apply whnf_mkApps_inv in wh; auto. + depelim wh; solve_discr. + constructor; eexists _, (decompose_stack π).1. + split; [constructor; eauto with pcuic|]. + split; eauto with pcuic. + apply whnf_mkApps. + depelim H; solve_discr; eauto with pcuic. + - apply whnf_mkApps_tSort_inv in wh as ->. + constructor; eexists _, []. + eauto using whnf_red_shape_spec with pcuic. + - apply whnf_mkApps_tProd_inv in wh as ->. + constructor; eexists _, []. + eauto using whnf_red_shape_spec with pcuic. + - depelim wh; solve_discr. + + apply whne_mkApps_inv in H as [|(?&?&?&?&?&?&?&?&?)]; auto; try discriminate. + depelim H; solve_discr. + discriminate. + + rewrite H1. + constructor; eexists _, []; eauto using whnf_red_shape_spec with pcuic. + - apply whnf_mkApps_inv in wh; auto. + depelim wh; solve_discr. + depelim H; solve_discr. + discriminate. + - discriminate. + - constructor; eexists _, (decompose_stack π).1; eauto using whnf_red_shape_spec with pcuic. + - constructor; eexists _, (decompose_stack π).1; eauto using whnf_red_shape_spec with pcuic. + - eapply unfold_one_fix_None in H0 as [(?&?&?)]. + constructor; eexists _, x. + split; [constructor; eauto with pcuic|]. + eauto with pcuic. + - constructor; eexists _, (decompose_stack π).1. + split; [constructor; eauto with pcuic|]. + eauto with pcuic. + - constructor; eexists _, (decompose_stack π).1. + split; [econstructor|]; eauto. + split; [eauto with pcuic|]. + apply whnf_mkApps. + eapply whne_const; eauto. + - zip fold in h. + apply wellformed_context in h; auto. + destruct hΣ. + apply wellformed_nonarity in h as (?&typ); auto. + apply inversion_Const in typ as (?&?&?&?); auto. + unfold declared_constant in d; congruence. + - zip fold in h. + apply wellformed_context in h; auto. + destruct hΣ. + apply wellformed_nonarity in h as (?&typ); auto. + apply inversion_Const in typ as (?&?&?&?); auto. + unfold declared_constant in d; congruence. + - admit. + - admit. + Admitted. + + Lemma whnf_red_shape_spec_red Γ t t' : + whnf_red_shape_spec Σ Γ t t' -> + ∥red Σ Γ t t'∥. + Proof. + intros wh. + depelim wh; constructor; eauto with pcuic. + - apply red_evar; auto. + - apply red_app; auto. + - apply red_fix_congr. + eapply All2_impl; eauto. + cbn. + intros ? ? (->&->&r1&r2). + eauto. + - apply red_case; auto. + eapply All2_impl; eauto. + cbn. + intros ? ? (->&?). + eauto. + - apply red_proj_c; auto. + - apply red_prod; auto. + - apply PCUICConfluence.red_abs_alt; auto. + - apply red_cofix_congr. + eapply All2_impl; eauto. + cbn. + intros ? ? (->&->&r1&r2). + eauto. + Qed. (* TODO Factorise *) Equations(noeqns) _isconv_fallback (Γ : context) (leq : conv_pb) (t1 : term) (π1 : stack) (h1 : wtp Γ t1 π1) (t2 : term) (π2 : stack) (h2 : wtp Γ t2 π2) (ir1 : isred_full Γ t1 π1) (ir2 : isred_full Γ t2 π2) + (hdiscr : prog_discr t1 t2) (hx : conv_stack_ctx Γ π1 π2) (aux : Aux Fallback Γ t1 π1 t2 π2 h2) : ConversionResult (conv_term leq Γ t1 π1 t2 π2) := - _isconv_fallback Γ leq t1 π1 h1 t2 π2 h2 ir1 ir2 hx aux + _isconv_fallback Γ leq t1 π1 h1 t2 π2 h2 ir1 ir2 hdiscr hx aux with inspect (reducible_head Γ t1 π1 h1) := { | @exist (Some (rt1, ρ1)) eq1 with inspect (decompose_stack ρ1) := { | @exist (l1, θ1) eq2 - with inspect (reduce_stack nodelta_flags Σ hΣ (Γ ,,, stack_context ρ1) rt1 (appstack l1 ε) _) := { + with inspect (reduce_stack RedFlags.nodelta Σ hΣ (Γ ,,, stack_context ρ1) rt1 (appstack l1 ε) _) := { | @exist (rt1', θ1') eq3 := isconv_prog leq rt1' (θ1' +++ θ1) t2 π2 aux } } ; - | @exist None _ with inspect (reducible_head Γ t2 π2 h2) := { + | @exist None nored1 with inspect (reducible_head Γ t2 π2 h2) := { | @exist (Some (rt2, ρ2)) eq1 with inspect (decompose_stack ρ2) := { | @exist (l2, θ2) eq2 - with inspect (reduce_stack nodelta_flags Σ hΣ (Γ ,,, stack_context ρ2) rt2 (appstack l2 ε) _) := { + with inspect (reduce_stack RedFlags.nodelta Σ hΣ (Γ ,,, stack_context ρ2) rt2 (appstack l2 ε) _) := { | @exist (rt2', θ2') eq3 := isconv_prog leq t1 π1 rt2' (θ2' +++ θ2) aux } } ; - | @exist None _ with inspect leq := { - | @exist Conv eq1 with inspect (eqb_term t1 t2) := { - | @exist true eq2 := isconv_args Conv t1 π1 t2 π2 aux ; - | @exist false _ := - Error ( - HeadMistmatch - Conv + | @exist None nored2 with inspect (eqb_termp_napp leq #|(decompose_stack π1).1| t1 t2) := { + | @exist true eq1 := isconv_args leq t1 π1 t2 π2 aux; + | @exist false noteq := + Error ( + HeadMismatch + leq (Γ ,,, stack_context π1) t1 (Γ ,,, stack_context π2) t2 ) _ - } ; - | @exist Cumul eq1 with inspect (leqb_term t1 t2) := { - | @exist true eq2 := isconv_args Cumul t1 π1 t2 π2 aux ; - | @exist false _ := - Error ( - HeadMistmatch - Cumul - (Γ ,,, stack_context π1) t1 - (Γ ,,, stack_context π2) t2 - ) _ - } } } }. @@ -4820,7 +5115,7 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in r2. cbn in r2. rewrite <- eq3 in d2. cbn in d2. @@ -4852,7 +5147,7 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in r2. cbn in r2. rewrite <- eq3 in d2. cbn in d2. @@ -4877,7 +5172,7 @@ Section Conversion. rewrite e' in d1. cbn in d1. subst. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. cbn in d2. @@ -4894,23 +5189,93 @@ Section Conversion. Next Obligation. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_isred nodelta_flags _ hΣ _ _ _ h eq_refl) as ir + pose proof (reduce_stack_isred RedFlags.nodelta _ hΣ _ _ _ h eq_refl) as ir; + pose proof (reduce_stack_whnf RedFlags.nodelta _ hΣ _ _ _ h) as wh end. - todo "isred_full". - (*rewrite <- eq3 in ir. destruct ir as [ia il]. simpl in ia, il. - split. - - cbn. assumption. - - simpl. intro hl. specialize (il hl). - destruct θ1'. all: simpl. all: try reflexivity. all: try discriminate. - eapply decompose_stack_noStackApp. eauto.*) + rewrite <- eq3 in ir, wh. + destruct ir as (?&_). + split; auto. + rewrite stack_context_stack_cat. + rewrite <- (stack_context_decompose ρ1) in wh. + rewrite <- eq2 in wh. + rewrite zipp_stack_cat; [rewrite app_context_assoc; auto|]. + eapply decompose_stack_noStackApp; eauto. Qed. Next Obligation. - todo "isred_full". + destruct hΣ as [wΣ]. + apply reducible_head_red_zipp in eq1 as r1. destruct r1 as [r1]. + apply reducible_head_decompose in eq1 as d1. + rewrite <- eq2 in d1. cbn in d1. + case_eq (decompose_stack π1). intros l' s' e'. + rewrite e' in d1. cbn in d1. subst. + match type of eq3 with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 + end. + rewrite <- eq3 in r2. cbn in r2. + rewrite <- eq3 in d2. cbn in d2. + rewrite decompose_stack_appstack in d2. cbn in d2. + rewrite zipc_appstack in r2. cbn in r2. + case_eq (decompose_stack θ1'). intros l s e. + rewrite e in d2. cbn in d2. subst. + apply decompose_stack_eq in e as ?. subst. + apply decompose_stack_eq in e' as ?. subst. + rewrite stack_context_appstack. + rewrite stack_cat_appstack in h. + rewrite stack_context_appstack in h. + etransitivity ; try eassumption. + unfold zipp. rewrite e'. + unfold zipp in r1. rewrite e' in r1. rewrite <- eq2 in r1. + rewrite decompose_stack_appstack. + erewrite decompose_stack_twice ; eauto. simpl. + rewrite app_nil_r. + eapply red_conv_cum_l ; try assumption. + rewrite stack_context_appstack in r1. + eapply red_trans ; try eassumption. + clear eq3. symmetry in eq2. apply decompose_stack_eq in eq2. subst. + rewrite stack_context_appstack in r2. + rewrite zipc_appstack in r2. cbn in r2. + assumption. Qed. Next Obligation. apply h; clear h. (* Contrapositive of previous case *) - todo "Completeness"%string. + destruct hΣ as [wΣ]. + apply reducible_head_red_zipp in eq1 as r1. destruct r1 as [r1]. + apply reducible_head_decompose in eq1 as d1. + rewrite <- eq2 in d1. cbn in d1. + case_eq (decompose_stack π1). intros l' s' e'. + rewrite e' in d1. cbn in d1. subst. + match type of eq3 with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 + end. + rewrite <- eq3 in r2. cbn in r2. + rewrite <- eq3 in d2. cbn in d2. + rewrite decompose_stack_appstack in d2. cbn in d2. + rewrite zipc_appstack in r2. cbn in r2. + case_eq (decompose_stack θ1'). intros l s e. + rewrite e in d2. cbn in d2. subst. + apply decompose_stack_eq in e as ?. subst. + apply decompose_stack_eq in e' as ?. subst. + rewrite stack_context_appstack in H. + rewrite stack_cat_appstack. + rewrite stack_context_appstack. + etransitivity ; try eassumption. + unfold zipp. rewrite e'. + unfold zipp in r1. rewrite e' in r1. rewrite <- eq2 in r1. + rewrite decompose_stack_appstack. + erewrite decompose_stack_twice ; eauto. simpl. + rewrite app_nil_r. + eapply red_conv_cum_r ; try assumption. + rewrite stack_context_appstack in r1. + eapply red_trans ; try eassumption. + clear eq3. symmetry in eq2. apply decompose_stack_eq in eq2. subst. + rewrite stack_context_appstack in r2. + rewrite zipc_appstack in r2. cbn in r2. + assumption. Qed. Next Obligation. cbn. rewrite zipc_appstack. cbn. @@ -4938,7 +5303,7 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in r2. cbn in r2. rewrite <- eq3 in d2. cbn in d2. @@ -4970,7 +5335,7 @@ Section Conversion. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in r2. cbn in r2. rewrite <- eq3 in d2. cbn in d2. @@ -4995,7 +5360,7 @@ Section Conversion. rewrite e' in d1. cbn in d1. subst. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose nodelta_flags _ hΣ _ _ _ h) as d2 + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 end. rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. cbn in d2. @@ -5012,47 +5377,191 @@ Section Conversion. Next Obligation. match type of eq3 with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_isred nodelta_flags _ hΣ _ _ _ h eq_refl) as ir + pose proof (reduce_stack_isred RedFlags.nodelta _ hΣ _ _ _ h eq_refl) as ir; + pose proof (reduce_stack_whnf RedFlags.nodelta _ hΣ _ _ _ h) as wh end. - todo "isred_full". - (*rewrite <- eq3 in ir. destruct ir as [ia il]. simpl in ia, il. - split. - - cbn. assumption. - - simpl. intro hl. specialize (il hl). - destruct θ2'. all: simpl. all: try reflexivity. all: try discriminate. - eapply decompose_stack_noStackApp. eauto.*) + rewrite <- eq3 in ir, wh. + destruct ir as (?&_). + split; auto. + rewrite stack_context_stack_cat. + rewrite <- (stack_context_decompose ρ2) in wh. + rewrite <- eq2 in wh. + rewrite zipp_stack_cat; [rewrite app_context_assoc; auto|]. + eapply decompose_stack_noStackApp; eauto. Qed. Next Obligation. - todo "isred_full". + destruct hΣ as [wΣ]. + apply reducible_head_red_zipp in eq1 as r1. destruct r1 as [r1]. + apply reducible_head_decompose in eq1 as d1. + rewrite <- eq2 in d1. cbn in d1. + case_eq (decompose_stack π2). intros l' s' e'. + rewrite e' in d1. cbn in d1. subst. + match type of eq3 with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 + end. + rewrite <- eq3 in r2. cbn in r2. + rewrite <- eq3 in d2. cbn in d2. + rewrite decompose_stack_appstack in d2. cbn in d2. + rewrite zipc_appstack in r2. cbn in r2. + case_eq (decompose_stack θ2'). intros l s e. + rewrite e in d2. cbn in d2. subst. + unfold zipp. rewrite e'. + unfold zipp in r1. rewrite e' in r1. rewrite <- eq2 in r1. + apply decompose_stack_eq in e as ?. subst. + apply decompose_stack_eq in e' as ?. subst. + clear eq3. + rewrite stack_context_appstack in hx. + destruct hx as [hx]. + etransitivity ; try eassumption. + unfold zipp. + rewrite stack_cat_appstack. + rewrite decompose_stack_appstack. + erewrite decompose_stack_twice ; eauto. simpl. + rewrite app_nil_r. + eapply conv_cum_context_convp. + - assumption. + - eapply red_conv_cum_r. 1: assumption. + rewrite stack_context_appstack in r1. + eapply red_trans ; try eassumption. + symmetry in eq2. apply decompose_stack_eq in eq2. subst. + rewrite stack_context_appstack in r2. + rewrite zipc_appstack in r2. cbn in r2. + assumption. + - eapply conv_context_sym. all: auto. Qed. Next Obligation. apply h; clear h. (* Contrapositive of previous case *) - todo "Completeness"%string. + destruct hΣ as [wΣ]. + apply reducible_head_red_zipp in eq1 as r1. destruct r1 as [r1]. + apply reducible_head_decompose in eq1 as d1. + rewrite <- eq2 in d1. cbn in d1. + case_eq (decompose_stack π2). intros l' s' e'. + rewrite e' in d1. cbn in d1. subst. + match type of eq3 with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; + pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 + end. + rewrite <- eq3 in r2. cbn in r2. + rewrite <- eq3 in d2. cbn in d2. + rewrite decompose_stack_appstack in d2. cbn in d2. + rewrite zipc_appstack in r2. cbn in r2. + case_eq (decompose_stack θ2'). intros l s e. + rewrite e in d2. cbn in d2. subst. + unfold zipp in r1. rewrite e' in r1. + unfold zipp in H. rewrite e' in H. rewrite <- eq2 in r1. + apply decompose_stack_eq in e as ?. subst. + apply decompose_stack_eq in e' as ?. subst. + clear eq3. + rewrite stack_context_appstack in hx. + destruct hx as [hx]. + etransitivity ; try eassumption. + unfold zipp. + rewrite stack_cat_appstack. + rewrite decompose_stack_appstack. + erewrite decompose_stack_twice ; eauto. simpl. + rewrite app_nil_r. + eapply conv_cum_context_convp. + - assumption. + - eapply red_conv_cum_l. + rewrite stack_context_appstack in r1. + eapply red_trans ; try eassumption. + symmetry in eq2. apply decompose_stack_eq in eq2. subst. + rewrite stack_context_appstack in r2. + rewrite zipc_appstack in r2. cbn in r2. + assumption. + - eapply conv_context_sym. all: auto. Qed. Next Obligation. eapply R_stateR. all: simpl. all: try reflexivity. constructor. Qed. Next Obligation. - constructor. constructor. - eapply eqb_term_spec. auto. - Qed. - Next Obligation. - (* Proof idea: t1, t2 are un whnf, convertible but distinct ? *) - todo "Completeness"%string. - Qed. - Next Obligation. - eapply R_stateR. all: simpl. all: try reflexivity. + destruct h, hΣ. + todo "variant of conv_cum_mkApps". + (*constructor. + rewrite !zipp_as_mkApps. + apply mkApps_conv_args; auto. constructor. + eapply eqb_term_spec. auto.*) Qed. Next Obligation. - constructor. constructor. - eapply leqb_term_spec. auto. + apply h; clear h. + destruct ir1 as (notapp1&whδ1), ir2 as (notapp2&whδ2). + rewrite !zipp_as_mkApps in *. + apply reducible_head_None in nored1 as [(?&?&s1&r1&wh1)]; auto. + apply reducible_head_None in nored2 as [(?&?&s2&r2&wh2)]; auto. + destruct hΣ, hx. + apply whnf_red_shape_spec_red in s1 as H'. + destruct H'. + apply whnf_red_shape_spec_red in s2 as H'. + destruct H'. + eapply conv_cum_red_inv' in H. + 2: apply hΣ'. + 2: apply red_mkApps; [eassumption|eassumption]. + 2: apply red_mkApps; [eassumption|eassumption]. + 2: eassumption. + apply conv_cum_mkApps_inv in H as [(?&?)]; auto. + 2: now depelim s1. + 2: now depelim s2. + 2: eapply whnf_conv_context; eauto. + 2: eapply conv_context_sym; eauto. + constructor. + eapply conv_terms_red'; eauto. Qed. Next Obligation. - (* Proof idea: t1, t2 are un whnf, in cumulative relation but not in leqb_term relation ? *) - todo "Completeness"%string. + unfold eqb_termp_napp in noteq. + destruct ir1 as (notapp1&whδ1), ir2 as (notapp2&whδ2). + rewrite !zipp_as_mkApps in *. + apply reducible_head_None in nored1 as [(?&?&s1&rargs1&wh1)]; auto. + apply reducible_head_None in nored2 as [(?&?&s2&rargs2&wh2)]; auto. + destruct hΣ, hx. + apply whnf_red_shape_spec_red in s1 as H'. + destruct H'. + apply whnf_red_shape_spec_red in s2 as H'. + destruct H'. + eapply conv_cum_red_inv' in H. + 2: apply hΣ'. + 2: apply red_mkApps; eassumption. + 2: apply red_mkApps; eassumption. + 2: eassumption. + apply conv_cum_mkApps_inv in H as [(conv_hds&_)]; auto. + 2: now depelim s1. + 2: now depelim s2. + 2: eapply whnf_conv_context; eauto. + 2: eapply conv_context_sym; eauto. + apply whnf_mkApps_inv in wh1; [|now depelim s1]. + destruct t1; cbn in *. + all: depelim s1. + 9: { destruct conv_hds as [H]. + depelim H. + depelim s2. + todo "R_global_instance => compare_global_instance". } + 9: { destruct conv_hds as [H]. + depelim H. + depelim s2. + todo "R_global_instance => compare_global_instance". } + all: apply conv_cum_alt in conv_hds as [(?&?&(r1&r2)&?)]. + all: eapply whnf_red_shape in r1; auto. + all: depelim r1. + all: depelim e. + all: apply whnf_mkApps_inv in wh2; [|now depelim s2]. + all: eapply whnf_conv_context in wh2; [|apply conv_context_sym; eauto]. + all: eapply whnf_red_shape in r2; auto. + all: depelim r2. + all: depelim s2. + all: destruct hdiscr. + - now rewrite Nat.eqb_refl in noteq. + - now rewrite eq_string_refl in noteq. + - zip fold in h1. + apply wellformed_context in h1; auto. + destruct h1 as [(?&typ)|[(?&?&?&?)]]. + + now apply inversion_Evar in typ. + + cbn in e; congruence. + - todo "conv_pb_rel -> conv_pb_relb". Qed. Equations _isconv (s : state) (Γ : context) @@ -5061,18 +5570,16 @@ Section Conversion. (aux : Aux s Γ t1 π1 t2 π2 h2) : Ret s Γ t1 π1 t2 π2 := _isconv Reduction Γ t1 π1 h1 t2 π2 h2 aux := - λ { | leq | hx | _ | _ | _ := _isconv_red Γ leq t1 π1 h1 t2 π2 h2 hx aux } ; + λ { | leq | hx := _isconv_red Γ leq t1 π1 h1 t2 π2 h2 hx aux } ; _isconv Term Γ t1 π1 h1 t2 π2 h2 aux := - λ { | leq | hx | r1 | r2 | _ := _isconv_prog Γ leq t1 π1 h1 t2 π2 h2 hx r1 r2 aux } ; + λ { | leq | hx | r1 | r2 := _isconv_prog Γ leq t1 π1 h1 t2 π2 h2 hx r1 r2 aux } ; _isconv Args Γ t1 π1 h1 t2 π2 h2 aux := - λ { | leq | hx | _ | _ | he := _isconv_args leq Γ t1 π1 h1 t2 π2 h2 hx he aux } ; + λ { | leq | hx := _isconv_args leq Γ t1 π1 h1 t2 π2 h2 hx aux } ; _isconv Fallback Γ t1 π1 h1 t2 π2 h2 aux := - λ { | leq | hx | r1 | r2 | _ := _isconv_fallback Γ leq t1 π1 h1 t2 π2 h2 r1 r2 hx aux }. - - Set Printing Coercions. + λ { | leq | hx | r1 | r2 | hd := _isconv_fallback Γ leq t1 π1 h1 t2 π2 h2 r1 r2 hd hx aux }. Equations(noeqns) isconv_full (s : state) (Γ : context) (t1 : term) (π1 : stack) (h1 : wtp Γ t1 π1) @@ -5089,9 +5596,8 @@ Section Conversion. (fun pp f => _) (x := mkpack Γ s t1 π1 t2 π2 _) _ _ _ _. - Obligations. Next Obligation. - unshelve eapply _isconv. 1: exact (st pp). all: try assumption. + unshelve eapply _isconv. all: try assumption. intros s' t1' π1' t2' π2' h1' h2' hx' hR. apply wellformed_zipc_zipp in h1. 2: auto. destruct pp. @@ -5108,9 +5614,9 @@ Section Conversion. | ConvError : ConversionError -> ConversionResultSummary. Definition isconv Γ leq t1 π1 h1 t2 π2 h2 hx := - match isconv_full Reduction Γ t1 π1 h1 t2 π2 h2 leq hx I I I with + match isconv_full Reduction Γ t1 π1 h1 t2 π2 h2 leq hx with | Success _ => ConvSuccess - | Error e h => ConvError e + | Error e _ => ConvError e end. Theorem isconv_sound : @@ -5124,6 +5630,17 @@ Section Conversion. - auto. - discriminate. Qed. + + Theorem isconv_complete : + forall Γ leq t1 π1 h1 t2 π2 h2 hx e, + isconv Γ leq t1 π1 h1 t2 π2 h2 hx = ConvError e -> + ~conv_cum leq Σ (Γ ,,, stack_context π1) (zipp t1 π1) (zipp t2 π2). + Proof. + unfold isconv. + intros Γ leq t1 π1 h1 t2 π2 h2 hx. + destruct isconv_full as []; auto. + intros ? [=]. + Qed. Definition isconv_term Γ leq t1 (h1 : wellformed Σ Γ t1) t2 (h2 : wellformed Σ Γ t2) := isconv Γ leq t1 ε h1 t2 ε h2 (sq (conv_ctx_refl _ Γ)). @@ -5137,6 +5654,16 @@ Section Conversion. unfold isconv_term. intro h. apply isconv_sound in h. apply h. Qed. + + Theorem isconv_term_complete : + forall Γ leq t1 h1 t2 h2 e, + isconv_term Γ leq t1 h1 t2 h2 = ConvError e -> + ~conv_cum leq Σ Γ t1 t2. + Proof. + intros Γ leq t1 h1 t2 h2 e. + unfold isconv_term. intro h. + apply isconv_complete in h. apply h. + Qed. Transparent reduce_stack. End Conversion. From e39550c55d3750e1afbb8c9433a7781e193bd073 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Sun, 25 Oct 2020 11:09:11 +0100 Subject: [PATCH 14/26] Rename whnf_red_shape_spec and some small fixes whnf_red_shape_spec -> whnf_red whnf_red_shape -> whnf_red_inv --- erasure/theories/ErasureFunction.v | 4 +- pcuic/theories/PCUICConversionInversion.v | 16 +- pcuic/theories/PCUICNormal.v | 563 ++++++--------------- safechecker/theories/PCUICSafeChecker.v | 10 +- safechecker/theories/PCUICSafeConversion.v | 212 +++----- 5 files changed, 215 insertions(+), 590 deletions(-) diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index af336fdf7..cad9def3a 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -130,7 +130,7 @@ Next Obligation. destruct HΣ. eapply PCUICConfluence.red_confluence in X0 as [t'' []]. 3:exact X1. 2:eauto. eapply invert_red_sort in r0; eauto. subst. - eapply whnf_red_shape in r; auto. + eapply whnf_red_inv in r; auto. depelim r. congruence. Qed. @@ -155,7 +155,7 @@ Next Obligation. destruct HΣ. eapply PCUICConfluence.red_confluence in X2 as [t'' []]. 3:exact X3. 2:eauto. eapply invert_red_prod in r0 as (? & ? & [] & ?); eauto. subst. - eapply whnf_red_shape in r; auto. + eapply whnf_red_inv in r; auto. depelim r. congruence. Qed. diff --git a/pcuic/theories/PCUICConversionInversion.v b/pcuic/theories/PCUICConversionInversion.v index 291590be0..478c87dea 100644 --- a/pcuic/theories/PCUICConversionInversion.v +++ b/pcuic/theories/PCUICConversionInversion.v @@ -190,10 +190,10 @@ Section fixed. 1-13: (eapply eq_term_upto_univ_napp_nonind; [exact e|try exact H1]). 1: discriminate notapp. all: apply whnf_mkApps_inv in wh; auto. - all: eapply whnf_red_shape in r; auto. + all: eapply whnf_red_inv in r; auto. all: depelim r. all: apply whnf_mkApps_inv in wh'; auto; [|now rewrite notapp']. - all: eapply whnf_red_shape in r0; auto. + all: eapply whnf_red_inv in r0; auto. all: depelim e. all: depelim r0. all: apply All2_length in a. @@ -221,8 +221,8 @@ Section fixed. depelim whr; solve_discr. depelim H0; solve_discr; try discriminate. apply conv_cum_alt in conv as [(?&?&(r1&r2)&?)]. - eapply whnf_red_shape in r1; eauto. - eapply whnf_red_shape in r2; eauto. + eapply whnf_red_inv in r1; eauto. + eapply whnf_red_inv in r2; eauto. depelim r1. depelim r2. depelim e. @@ -255,8 +255,8 @@ Section fixed. apply whnf_fixapp with (v := []). destruct unfold_fix as [(?&?)|]; [|easy]. now rewrite nth_error_nil. } - eapply whnf_red_shape in r1; eauto. - eapply whnf_red_shape in r2; eauto. + eapply whnf_red_inv in r1; eauto. + eapply whnf_red_inv in r2; eauto. depelim r1. depelim r2. depelim e. @@ -292,8 +292,8 @@ Section fixed. Proof. intros conv whl whr. apply conv_cum_alt in conv as [(?&?&(r1&r2)&?)]. - eapply whnf_red_shape in r1; eauto. - eapply whnf_red_shape in r2; eauto. + eapply whnf_red_inv in r1; eauto. + eapply whnf_red_inv in r2; eauto. depelim r1. depelim r2. depelim e. diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index eeb47b8cf..783528f87 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -889,66 +889,149 @@ Qed. Hint Resolve All2_same All2_firstn All2_skipn OnOne2_All2 red_mkApps All2_app : pcuic. -Inductive whnf_red_shape_spec Σ Γ : term -> term -> Prop := -| wr_shape_tRel i : +(* For terms in whnf we have a very useful inversion lemma for reductions. + This next relation is a subrelation of red that classifies how whnf terms reduce. *) +Inductive whnf_red Σ Γ : term -> term -> Prop := +| whnf_red_tRel i : option_map decl_body (nth_error Γ i) = Some None -> - whnf_red_shape_spec Σ Γ (tRel i) (tRel i) -| wr_shape_tVar id : whnf_red_shape_spec Σ Γ (tVar id) (tVar id) -| wr_shape_tEvar n l l' : + whnf_red Σ Γ (tRel i) (tRel i) +| whnf_red_tVar id : whnf_red Σ Γ (tVar id) (tVar id) +| whnf_red_tEvar n l l' : All2 (red Σ Γ) l l' -> - whnf_red_shape_spec Σ Γ (tEvar n l) (tEvar n l') -| wr_shape_tConst kn u decl : + whnf_red Σ Γ (tEvar n l) (tEvar n l') +| whnf_red_tConst kn u decl : lookup_env Σ kn = Some (ConstantDecl decl) -> cst_body decl = None -> - whnf_red_shape_spec Σ Γ (tConst kn u) (tConst kn u) -| wr_shape_tApp hd hd' arg arg' : + whnf_red Σ Γ (tConst kn u) (tConst kn u) +| whnf_red_tApp hd hd' arg arg' : red Σ Γ hd hd' -> - whnf_red_shape_spec Σ Γ hd hd' -> + whnf_red Σ Γ hd hd' -> red Σ Γ arg arg' -> - whnf_red_shape_spec Σ Γ (tApp hd arg) (tApp hd' arg') -| wr_shape_tFix mfix mfix' idx : + whnf_red Σ Γ (tApp hd arg) (tApp hd' arg') +| whnf_red_tFix mfix mfix' idx : All2 (fun d d' => dname d = dname d' × rarg d = rarg d' × red Σ Γ (dtype d) (dtype d') × red Σ (Γ,,, fix_context mfix) (dbody d) (dbody d')) mfix mfix' -> - whnf_red_shape_spec Σ Γ (tFix mfix idx) (tFix mfix' idx) -| wr_shape_tCase p motive motive' discr discr' brs brs' : + whnf_red Σ Γ (tFix mfix idx) (tFix mfix' idx) +| whnf_red_tCase p motive motive' discr discr' brs brs' : red Σ Γ motive motive' -> red Σ Γ discr discr' -> All2 (fun br br' => br.1 = br'.1 × red Σ Γ br.2 br'.2) brs brs' -> - whnf_red_shape_spec Σ Γ (tCase p motive discr brs) (tCase p motive' discr' brs') -| wr_shape_tProj p c c' : + whnf_red Σ Γ (tCase p motive discr brs) (tCase p motive' discr' brs') +| whnf_red_tProj p c c' : red Σ Γ c c' -> - whnf_red_shape_spec Σ Γ (tProj p c) (tProj p c') -| wr_shape_tSort s : whnf_red_shape_spec Σ Γ (tSort s) (tSort s) -| wr_shape_tProd na A A' B B' : + whnf_red Σ Γ (tProj p c) (tProj p c') +| whnf_red_tSort s : whnf_red Σ Γ (tSort s) (tSort s) +| whnf_red_tProd na A A' B B' : red Σ Γ A A' -> red Σ (Γ,, vass na A) B B' -> - whnf_red_shape_spec Σ Γ (tProd na A B) (tProd na A' B') -| wr_shape_tLambda na A A' B B' : + whnf_red Σ Γ (tProd na A B) (tProd na A' B') +| whnf_red_tLambda na A A' B B' : red Σ Γ A A' -> red Σ (Γ,, vass na A) B B' -> - whnf_red_shape_spec Σ Γ (tLambda na A B) (tLambda na A' B') -| wr_shape_tConstruct i n u : - whnf_red_shape_spec Σ Γ (tConstruct i n u) (tConstruct i n u) -| wr_shape_tInd i u : - whnf_red_shape_spec Σ Γ (tInd i u) (tInd i u) -| wr_shape_tCoFix mfix mfix' idx : + whnf_red Σ Γ (tLambda na A B) (tLambda na A' B') +| whnf_red_tConstruct i n u : + whnf_red Σ Γ (tConstruct i n u) (tConstruct i n u) +| whnf_red_tInd i u : + whnf_red Σ Γ (tInd i u) (tInd i u) +| whnf_red_tCoFix mfix mfix' idx : All2 (fun d d' => dname d = dname d' × rarg d = rarg d' × red Σ Γ (dtype d) (dtype d') × red Σ (Γ,,, fix_context mfix) (dbody d) (dbody d')) mfix mfix' -> - whnf_red_shape_spec Σ Γ (tCoFix mfix idx) (tCoFix mfix' idx). + whnf_red Σ Γ (tCoFix mfix idx) (tCoFix mfix' idx). -Derive Signature for whnf_red_shape_spec. +Derive Signature for whnf_red. -Instance whnf_red_shape_spec_trans Σ Γ : RelationClasses.Transitive (whnf_red_shape_spec Σ Γ). +Lemma whnf_red_red Σ Γ t t' : + whnf_red Σ Γ t t' -> + ∥red Σ Γ t t'∥. +Proof. + intros wh. + depelim wh; constructor; eauto with pcuic. + - apply red_evar; auto. + - apply red_app; auto. + - apply red_fix_congr. + eapply All2_impl; eauto. + cbn. + intros ? ? (->&->&r1&r2). + eauto. + - apply red_case; auto. + eapply All2_impl; eauto. + cbn. + intros ? ? (->&?). + eauto. + - apply red_proj_c; auto. + - apply red_prod; auto. + - apply red_abs; auto. + eapply context_change_decl_types_red; eauto. + constructor; [|constructor]. + apply context_relation_refl. + reflexivity. + - apply red_cofix_congr. + eapply All2_impl; eauto. + cbn. + intros ? ? (->&->&r1&r2). + eauto. +Qed. + +Lemma whnf_red_mkApps Σ Γ hd hd' args args' : + whnf_red Σ Γ hd hd' -> + red Σ Γ hd hd' -> + All2 (red Σ Γ) args args' -> + whnf_red Σ Γ (mkApps hd args) (mkApps hd' args'). +Proof. + intros shape r ra. + induction ra using All2_ind_rev; auto. + rewrite <- !mkApps_nested. + cbn. + constructor; auto. + apply red_mkApps; auto. +Qed. + +Lemma whnf_red_refl_whne Σ Γ t : + whne RedFlags.default Σ Γ t -> + whnf_red Σ Γ t t. +Proof. + intros wh. + induction wh; cbn in *; try discriminate; eauto using whnf_red with pcuic. + apply whnf_red_mkApps; auto. + 2: apply All2_same; auto. + constructor. + apply All2_same; auto. +Qed. + +Lemma whnf_red_refl Σ Γ t : + whnf RedFlags.default Σ Γ t -> + whnf_red Σ Γ t t. +Proof. + intros wh. + induction wh; cbn in *; try discriminate; eauto using whnf_red with pcuic. + - eapply whnf_red_refl_whne; eauto. + - apply whnf_red_mkApps; auto. + 2: apply All2_same; auto. + constructor. + - apply whnf_red_mkApps; auto. + 2: apply All2_same; auto. + constructor. + - apply whnf_red_mkApps; auto. + 2: apply All2_same; auto. + constructor. + apply All2_same; auto. + - apply whnf_red_mkApps; auto. + 2: apply All2_same; auto. + constructor. + apply All2_same; auto. +Qed. + +Instance whnf_red_trans Σ Γ : RelationClasses.Transitive (whnf_red Σ Γ). Proof. intros x y z xy yz. revert z yz. - induction xy; intros z yz; depelim yz; eauto using whnf_red_shape_spec. + induction xy; intros z yz; depelim yz; eauto using whnf_red. - constructor. eapply All2_trans; eauto. typeclasses eauto. @@ -997,49 +1080,23 @@ Proof. now apply All2_length in X. Qed. -Lemma whnf_red_shape_spec_mkApps Σ Γ hd hd' args args' : - whnf_red_shape_spec Σ Γ hd hd' -> - red Σ Γ hd hd' -> - All2 (red Σ Γ) args args' -> - whnf_red_shape_spec Σ Γ (mkApps hd args) (mkApps hd' args'). -Proof. - intros shape r ra. - induction ra using All2_ind_rev; auto. - rewrite <- !mkApps_nested. - cbn. - constructor; auto. - apply red_mkApps; auto. -Qed. - -Lemma whnf_red_shape_spec_refl_whne Σ Γ t : - whne RedFlags.default Σ Γ t -> - whnf_red_shape_spec Σ Γ t t. -Proof. - intros wh. - induction wh; cbn in *; try discriminate; eauto using whnf_red_shape_spec with pcuic. - apply whnf_red_shape_spec_mkApps; auto. - 2: apply All2_same; auto. - constructor. - apply All2_same; auto. -Qed. - -Lemma whne_red1_shape Σ Γ t t' : +Lemma whne_red1_inv Σ Γ t t' : whne RedFlags.default Σ Γ t -> red1 Σ Γ t t' -> - whnf_red_shape_spec Σ Γ t t'. + whnf_red Σ Γ t t'. Proof. intros wh r. apply (whne_red1_ind RedFlags.default Σ Γ - (whnf_red_shape_spec Σ Γ)); intros; cbn in *; try easy; eauto using whnf_red_shape_spec. + (whnf_red Σ Γ)); intros; cbn in *; try easy; eauto using whnf_red. - constructor. eapply OnOne2_All2; eauto. - constructor; auto. - now apply whnf_red_shape_spec_refl_whne. - - apply whnf_red_shape_spec_mkApps; eauto. + now apply whnf_red_refl_whne. + - apply whnf_red_mkApps; eauto. 1: constructor; apply All2_same; auto. eapply OnOne2_All2; eauto. - - apply whnf_red_shape_spec_mkApps; eauto. + - apply whnf_red_mkApps; eauto. + constructor. eapply OnOne2_All2; eauto. cbn. @@ -1047,7 +1104,7 @@ Proof. auto. + now apply red1_red, fix_red_ty. + apply All2_same; auto. - - apply whnf_red_shape_spec_mkApps; eauto. + - apply whnf_red_mkApps; eauto. + constructor. eapply OnOne2_All2; eauto. cbn. @@ -1064,55 +1121,32 @@ Proof. intros ? ? (?&?); auto. Qed. -Lemma whnf_red_shape_spec_refl Σ Γ t : - whnf RedFlags.default Σ Γ t -> - whnf_red_shape_spec Σ Γ t t. -Proof. - intros wh. - induction wh; cbn in *; try discriminate; eauto using whnf_red_shape_spec with pcuic. - - eapply whnf_red_shape_spec_refl_whne; eauto. - - apply whnf_red_shape_spec_mkApps; auto. - 2: apply All2_same; auto. - constructor. - - apply whnf_red_shape_spec_mkApps; auto. - 2: apply All2_same; auto. - constructor. - - apply whnf_red_shape_spec_mkApps; auto. - 2: apply All2_same; auto. - constructor. - apply All2_same; auto. - - apply whnf_red_shape_spec_mkApps; auto. - 2: apply All2_same; auto. - constructor. - apply All2_same; auto. -Qed. - -Lemma whnf_red1_shape Σ Γ t t' : +Lemma whnf_red1_inv Σ Γ t t' : whnf RedFlags.default Σ Γ t -> red1 Σ Γ t t' -> - whnf_red_shape_spec Σ Γ t t'. + whnf_red Σ Γ t t'. Proof. intros wh r. destruct wh. - - eapply whne_red1_shape; eauto. + - eapply whne_red1_inv; eauto. - depelim r; solve_discr. - - depelim r; solve_discr; constructor; eauto using whnf_red_shape_spec. - - depelim r; solve_discr; constructor; eauto using whnf_red_shape_spec. + - depelim r; solve_discr; constructor; eauto using whnf_red. + - depelim r; solve_discr; constructor; eauto using whnf_red. - apply red1_mkApps_tConstruct_inv in r as (?&->&?). - apply whnf_red_shape_spec_mkApps; eauto using whnf_red_shape_spec. + apply whnf_red_mkApps; eauto using whnf_red. eapply OnOne2_All2; eauto. - apply red1_mkApps_tInd_inv in r as (?&->&?). - apply whnf_red_shape_spec_mkApps; eauto using whnf_red_shape_spec. + apply whnf_red_mkApps; eauto using whnf_red. eapply OnOne2_All2; eauto. - apply red1_mkApps_tFix_inv in r. 2: { destruct unfold_fix as [(?&?)|]; [|easy]. now unfold is_constructor; rewrite H. } destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]. - + apply whnf_red_shape_spec_mkApps; auto. + + apply whnf_red_mkApps; auto. 2: eapply OnOne2_All2; eauto. constructor. apply All2_same; auto. - + apply whnf_red_shape_spec_mkApps. + + apply whnf_red_mkApps. 3: apply All2_same; auto. 2: apply red1_red, fix_red_ty; auto. constructor. @@ -1120,7 +1154,7 @@ Proof. cbn. intros ? ? (?&[= -> -> ->]). auto. - + apply whnf_red_shape_spec_mkApps. + + apply whnf_red_mkApps. 3: apply All2_same; auto. 2: apply red1_red, fix_red_body; auto. constructor. @@ -1130,11 +1164,11 @@ Proof. auto. - apply red1_mkApps_tCoFix_inv in r. destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]. - + apply whnf_red_shape_spec_mkApps; auto. + + apply whnf_red_mkApps; auto. 2: eapply OnOne2_All2; eauto. constructor. apply All2_same; auto. - + apply whnf_red_shape_spec_mkApps. + + apply whnf_red_mkApps. 3: apply All2_same; auto. 2: apply red1_red, cofix_red_ty; auto. constructor. @@ -1142,7 +1176,7 @@ Proof. cbn. intros ? ? (?&[= -> -> ->]). auto. - + apply whnf_red_shape_spec_mkApps. + + apply whnf_red_mkApps. 3: apply All2_same; auto. 2: apply red1_red, cofix_red_body; auto. constructor. @@ -1152,15 +1186,15 @@ Proof. auto. Qed. -Lemma whnf_red_shape Σ Γ t t' : +Lemma whnf_red_inv Σ Γ t t' : whnf RedFlags.default Σ Γ t -> red Σ Γ t t' -> - whnf_red_shape_spec Σ Γ t t'. + whnf_red Σ Γ t t'. Proof. intros wh r. induction r using red_rect_n1. - - apply whnf_red_shape_spec_refl; auto. - - eapply whnf_red1_shape in X. + - apply whnf_red_refl; auto. + - eapply whnf_red1_inv in X. + etransitivity; eauto. + eapply whnf_pres; eauto. Qed. @@ -1171,7 +1205,7 @@ Lemma whnf_red_isApp Σ Γ t t' : isApp t' = isApp t. Proof. intros wh r. - eapply whnf_red_shape in r; eauto. + eapply whnf_red_inv in r; eauto. now depelim r. Qed. @@ -1182,7 +1216,7 @@ Lemma whnf_red_mkApps_l Σ Γ hd args t : t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. Proof. intros wh r. - eapply whnf_red_shape in wh; eauto. + eapply whnf_red_inv in wh; eauto. induction args in hd, t, r, wh |- * using List.rev_ind. - cbn in *. constructor; eexists _, []; split; [reflexivity|]. @@ -1204,7 +1238,7 @@ Lemma whnf_red_mkApps_r Σ Γ t hd' args' : t = mkApps hd args × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. Proof. intros wh r. - eapply whnf_red_shape in wh; eauto. + eapply whnf_red_inv in wh; eauto. induction args' in hd', t, r, wh |- * using List.rev_ind. - cbn in *. constructor; eexists _, []; split; [reflexivity|]. @@ -1219,7 +1253,7 @@ Proof. eauto with pcuic. Qed. -Lemma whnf_red_mkApps Σ Γ hd args hd' args' : +Lemma whnf_red_mkApps_inv Σ Γ hd args hd' args' : whnf RedFlags.default Σ Γ (mkApps hd args) -> red Σ Γ (mkApps hd args) (mkApps hd' args') -> isApp hd = false -> @@ -1267,14 +1301,6 @@ Proof. now rewrite nth_error_nil in H0. Qed. -Lemma whne_eq_term {cf:checker_flags} f Σ φ Γ t t' : - whne f Σ Γ t -> - eq_term Σ φ t t' -> - whne f Σ Γ t'. -Proof. - apply whne_eq_term_upto_univ_napp. -Qed. - Lemma whnf_eq_term_upto_univ_napp f Σ Γ t Re Rle napp t' : whnf f Σ Γ t -> eq_term_upto_univ_napp Σ Re Rle napp t t' -> @@ -1318,308 +1344,3 @@ Lemma whnf_eq_term {cf:checker_flags} f Σ φ Γ t t' : Proof. apply whnf_eq_term_upto_univ_napp. Qed. - -(* - - Definition head_arg_is_constructor mfix idx args := - match unfold_fix mfix idx with Some (narg, body) => is_constructor narg args | None => false end. - - - Inductive normal (Γ : context) : term -> Prop := - | nf_ne t : neutral Γ t -> normal Γ t - | nf_lam na A B : normal Γ A -> normal (Γ ,, vass na A) B -> - normal Γ (tLambda na A B) - | nf_cstrapp i n u v : All (normal Γ) v -> normal Γ (mkApps (tConstruct i n u) v) - | nf_indapp i u v : All (normal Γ) v -> normal Γ (mkApps (tInd i u) v) - | nf_fix mfix idx args : All (compose (normal (Γ ,,, PCUICLiftSubst.fix_context mfix)) dbody) mfix -> - All (normal Γ) args -> - head_arg_is_constructor mfix idx args = false -> - All (compose (normal Γ) dtype) mfix -> - normal Γ (mkApps (tFix mfix idx) args) - | nf_cofix mfix idx : All (compose (normal (Γ ,,, PCUICLiftSubst.fix_context mfix)) dbody) mfix -> - All (compose (normal Γ) dtype) mfix -> - normal Γ (tCoFix mfix idx) - - with neutral (Γ : context) : term -> Prop := - | ne_rel i : - (forall b, option_map decl_body (nth_error Γ i) <> Some (Some b)) -> - neutral Γ (tRel i) - | ne_var v : neutral Γ (tVar v) - | ne_evar n l : neutral Γ (tEvar n l) - | ne_const c u : - (forall decl b, lookup_env Σ c = Some (ConstantDecl decl) -> decl.(cst_body) = Some b -> False) -> - neutral Γ (tConst c u) - | ne_sort s : neutral Γ (tSort s) - | ne_prod na A B : normal Γ A -> normal (Γ ,, vass na A) B -> - neutral Γ (tProd na A B) - | ne_app f v : neutral Γ f -> normal Γ v -> neutral Γ (tApp f v) - | ne_case i p c brs : neutral Γ c -> normal Γ p -> All (compose (normal Γ) snd) brs -> - neutral Γ (tCase i p c brs) - | ne_proj p c : neutral Γ c -> neutral Γ (tProj p c). - -Definition normal_neutral_dec Σ Γ t : ({normal Σ Γ t} + {~ (normal Σ Γ t)}) * ({neutral Σ Γ t} + {~ (neutral Σ Γ t)}). -Proof. - induction t using term_forall_mkApps_ind in Γ |- *; split; eauto. - all: try now (right; intros H; depelim H). - - destruct (option_map decl_body (nth_error Γ n)) as [ [ | ] | ] eqn:E. - + right. intros H. depelim H. depelim H. congruence. help. help. help. - + eauto. - + right. intros H. depelim H. depelim H. congruence. help. help. help. - - destruct (option_map decl_body (nth_error Γ n)) as [ [ | ] | ] eqn:E. - + right. intros H. depelim H. congruence. - + eauto. - + right. intros H. depelim H. congruence. - - todo "evar". - - todo "evar". - - destruct (IHt1 Γ) as [[] _]; - [destruct (IHt2 (Γ,, vass n t1)) as [[] _]|]; eauto. - + right. intros H. depelim H. depelim H. eauto. help. help. help. - + right. intros H. depelim H. depelim H. eauto. help. help. help. - - destruct (IHt1 Γ) as [[] _]; - [destruct (IHt2 (Γ,, vass n t1)) as [[] _]|]; eauto. - + right. intros H. depelim H. eauto. - + right. intros H. depelim H. eauto. - - destruct (IHt1 Γ) as [[] _]; - [destruct (IHt2 (Γ,, vass n t1)) as [[] _]|]; eauto. - + right. intros H. depelim H. depelim H. eauto. help. help. help. - + right. intros H. depelim H. depelim H. eauto. help. help. help. - - right. intros H. depelim H. depelim H. help. help. help. - - destruct (IHt Γ) as [[] _]. - + destruct dec_All with (P := (normal Σ Γ)) (L := v). - -- eapply All_impl. eassumption. intros ? ?. apply X0. - -- destruct t. all:eauto using normal_mkApps, All_Forall. - all: try now (left; depelim n; help; eauto). - ++ destruct v as [ | ? v]. - ** eauto. - ** right. todo "admit". (* intros ?. depelim X. depelim X. all:help. clear IHt. *) - (* eapply neutral_mk_Apps_inv in H0. as []; eauto. *) - (* depelim H1. *) - ++ destruct (head_arg_is_constructor mfix idx v) eqn:E. - ** right. intros ?. todo "admit". (* depelim H1. depelim H1. all:help. clear IHv. *) - (* eapply neutral_mk_Apps_inv in H1 as []; eauto. depelim H1. *) - ** left. todo "admit". (* depelim n. all:help. depelim H1. *) - (* eapply (f_equal decompose_app) in x; *) - (* rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. inv x. *) - (* eauto. *) - ++ todo "cofix". - -- right. intros ?. eapply f. eapply Forall_All. - now eapply normal_mk_Apps_inv. - + right. intros ?. eapply n. now eapply normal_mk_Apps_inv. - - destruct v using rev_ind. - + cbn. eapply IHt. - + rewrite <- mkApps_nested. cbn. - eapply All_app in X as []. eapply IHv in a. inv a0. clear X0. - rename X into IHt2. - revert a. - generalize (mkApps t v) as t1. intros t1 IHt1. - destruct (IHt1) as []; - [destruct (IHt2 Γ) as [[] _]|]; eauto. - * right. intros HH. depelim HH. eauto. - * right. intros HH. depelim HH. eauto. - - destruct (lookup_env Σ s) as [[] | ] eqn:E. - + destruct (cst_body c) eqn:E2. - * right. intros H. depelim H. depelim H. congruence. help. help. help. - * eauto. - + right. intros H. depelim H. depelim H. congruence. help. help. help. - + right. intros H. depelim H. depelim H. congruence. help. help. help. - - destruct (lookup_env Σ s) as [[] | ] eqn:E. - + destruct (cst_body c) eqn:E2. - * right. intros H. depelim H. congruence. - * eauto. - + right. intros H. depelim H. congruence. - + right. intros H. depelim H. congruence. - - left. eapply nf_indapp with (v := []). econstructor. - - left. eapply nf_cstrapp with (v := []). econstructor. - - destruct (IHt2 Γ) as [_ []]. - + destruct (IHt1 Γ) as [[] _]. - * destruct dec_All with(L := l) (P := (normal Σ Γ ∘ @snd nat term)). - -- eapply All_impl. eassumption. intros ? ?. eapply X0. - -- eauto. - -- right. intros ?. depelim H. depelim H. all:help. eauto. - * right. intros ?. depelim H. depelim H. all:help. eauto. - + right. intros ?. depelim H. depelim H. all:help. eauto. - - destruct (IHt2 Γ) as [_ []]. - + destruct (IHt1 Γ) as [[] _]. - * destruct dec_All with(L := l) (P := (normal Σ Γ ∘ @snd nat term)). - -- eapply All_impl. eassumption. intros ? ?. eapply X0. - -- eauto. - -- right. intros ?. depelim H. all:help. eauto. - * right. intros ?. depelim H. all:help. eauto. - + right. intros ?. depelim H. all:help. eauto. - - destruct (IHt Γ) as [_ []]. - + eauto. - + right. intros H. depelim H. depelim H. eauto. help. help. help. - - destruct (IHt Γ) as [_ []]. - + eauto. - + right. intros H. depelim H. eauto. - - hnf in X. - - destruct dec_All with (P := (normal Σ Γ ∘ dtype)) (L := m). - eapply All_impl. eassumption. cbn. intros. eapply X0. - - destruct dec_All with (P := normal Σ (Γ ,,, PCUICLiftSubst.fix_context m) ∘ dbody) (L := m). - eapply All_impl. exact X. cbn. intros. eapply X0. - - + left. eapply nf_fix with (args := []). eauto. eauto. unfold head_arg_is_constructor. - destruct unfold_fix; eauto. destruct p. - unfold is_constructor. destruct n0; eauto. eauto. - + right. intros H. depelim H. depelim H. help. help. eapply f. - eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. inv x. - eauto. - + right. intros H. depelim H. depelim H. help. help. - eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. inv x. - eauto. - - hnf in X. - - destruct dec_All with (P := (normal Σ Γ ∘ dtype)) (L := m). - eapply All_impl. eassumption. cbn. intros. eapply X0. - - destruct dec_All with (P := normal Σ (Γ ,,, PCUICLiftSubst.fix_context m) ∘ dbody) (L := m). - eapply All_impl. exact X. cbn. intros. eapply X0. - - + left. eapply nf_cofix. eauto. eauto. - + right. intros H. depelim H. depelim H. help. help. help. eauto. - + right. intros H. depelim H. depelim H. help. help. help. eauto. -Defined. - -Definition normal_dec Σ Γ t := fst (normal_neutral_dec Σ Γ t). -Definition neutral_dec Σ Γ t := snd (normal_neutral_dec Σ Γ t). - *) - - -(* - - -Lemma OnOn2_contra A (P : A -> A -> Type) l1 l2 : (forall x y, P x y -> False) -> OnOne2 P l1 l2 -> False. -Proof. - intros. induction X; eauto. -Qed. - -Lemma normal_nf Σ Γ t t' : normal Σ Γ t \/ neutral Σ Γ t -> red1 Σ Γ t t' -> False. -Proof. - intros. induction X using red1_ind_all; destruct H. - all: repeat match goal with - | [ H : normal _ _ (?f ?a) |- _ ] => depelim H - | [ H : neutral _ _ (?f ?a)|- _ ] => depelim H - end. - all: try congruence. - all: help. - all: try tauto. - all: try now (clear - H; depind H; help; eauto). - - eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. - inv x. unfold head_arg_is_constructor in H. - rewrite H0 in H. congruence. - - eapply OnOne2_All_mix_left in X; try eassumption. - eapply OnOn2_contra; try eassumption. - firstorder. - - eapply OnOne2_All_mix_left in X; try eassumption. - eapply OnOn2_contra; try eassumption. - firstorder. - - eapply IHX. left. - eapply nf_cstrapp. now eapply All_app in X as [X _]. - - eapply IHX. left. - eapply nf_indapp. now eapply All_app in X as [X _]. - - clear IHargs. - eapply IHX. left. - eapply nf_fix. - + eauto. - + eapply All_app in X0. eapply X0. - + unfold head_arg_is_constructor in *. - destruct unfold_fix; eauto. destruct p. - unfold is_constructor in *. - destruct (nth_error args n) eqn:E; eauto. - erewrite nth_error_app_left in H; eauto. - + eauto. - - eapply IHX. left. - eapply All_app in X as [_ X]. now inv X. - - eapply IHX. left. - eapply All_app in X as [_ X]. now inv X. - - eapply IHX. left. - eapply All_app in X0 as [_ X0]. now inv X0. - - eapply OnOne2_All_mix_left in X; try eassumption. - eapply OnOn2_contra; try eassumption. - firstorder. - - eapply OnOne2_All_mix_left in X; try eassumption. - eapply OnOn2_contra; try eassumption. - firstorder. - - eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. inv x. - eapply OnOne2_All_mix_left in X; try eassumption. - eapply OnOn2_contra; try eassumption. -(* firstorder. *) -(* - eapply (f_equal decompose_app) in x; *) -(* rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. inv x. *) -(* eapply OnOne2_All_mix_left in X. 2:exact H. *) -(* eapply OnOn2_contra; try eassumption. *) -(* firstorder. *) -(* - eapply OnOne2_All_mix_left in X. 2:exact H0. *) -(* eapply OnOn2_contra; try eassumption. *) -(* firstorder. *) -(* - eapply OnOne2_All_mix_left in X. 2:exact H. *) -(* eapply OnOn2_contra; try eassumption. *) -(* firstorder. *) -(* Qed. *) -Admitted. - -Hint Constructors normal neutral : core. - -Lemma normal_mk_Apps_inv: - forall (Σ : global_env) (Γ : context) (t : term) (v : list term), ~ isApp t -> normal Σ Γ (mkApps t v) -> normal Σ Γ t /\ Forall (normal Σ Γ) v. -Proof. -(* intros Σ Γ t v H H1. *) -(* induction v using rev_ind. *) -(* - split. eapply H1. econstructor. *) -(* - rewrite <- mkApps_nested in H1. cbn in H1. depelim H1. depelim H0. *) -(* + split. *) -(* * firstorder. *) -(* * eapply app_Forall. firstorder. firstorder. *) -(* + change (tApp (mkApps t v) x0) with (mkApps (mkApps t v) [x0]) in *. *) -(* rewrite mkApps_nested in x. *) -(* eapply (f_equal decompose_app) in x. *) -(* rewrite !decompose_app_mkApps in x; cbn in *; try congruence. firstorder. inv x. *) -(* split. eapply nf_cstrapp with (v := []). econstructor. *) -(* now eapply All_Forall. *) -(* + change (tApp (mkApps t v) x0) with (mkApps (mkApps t v) [x0]) in *. *) -(* rewrite mkApps_nested in x. *) -(* eapply (f_equal decompose_app) in x. *) -(* rewrite !decompose_app_mkApps in x; cbn in *; try congruence. firstorder. inv x. *) -(* split. eapply nf_indapp with (v := []). econstructor. *) -(* now eapply All_Forall. *) -(* + change (tApp (mkApps t v) x0) with (mkApps (mkApps t v) [x0]) in *. *) -(* rewrite mkApps_nested in x. *) -(* eapply (f_equal decompose_app) in x. *) -(* rewrite !decompose_app_mkApps in x; cbn in *; try congruence. firstorder. inv x. *) -(* split. eapply nf_fix with (args := []). *) -(* * eauto. *) -(* * econstructor. *) -(* * unfold head_arg_is_constructor in *. *) -(* destruct unfold_fix; eauto. destruct p. *) -(* unfold is_constructor in *. destruct n; reflexivity. *) -(* * eauto. *) -(* * now eapply All_Forall. *) -(* Qed. *) -Admitted. - -Lemma neutral_mk_Apps_inv: - forall (Σ : global_env) (Γ : context) (t : term) (v : list term), ~ isApp t -> neutral Σ Γ (mkApps t v) -> neutral Σ Γ t /\ Forall (normal Σ Γ) v. -Proof. - intros Σ Γ t v H H1. - induction v using rev_ind. - - split. eapply H1. econstructor. - - rewrite <- mkApps_nested in H1. cbn in H1. depelim H1. - split. - + firstorder. - + eapply app_Forall. firstorder. firstorder. -Qed. - -Lemma normal_mkApps Σ Γ t v : - neutral Σ Γ t -> Forall (normal Σ Γ) v -> normal Σ Γ (mkApps t v). -Proof. - intros. induction H0 in t, H |- *; cbn in *. - - eauto. - - eapply IHForall. eauto. -Qed. - *) diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index 81460dc6f..3d70ae03c 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -223,7 +223,7 @@ Fixpoint string_of_conv_error Σ (e : ConversionError) : string := "\nand\n" ^ print_term Σ Γ2 t2 ^ "are convertible/convertible (TODO) but applied to a different number" ^ " of arguments." - | HeadMistmatch leq Γ1 t1 Γ2 t2 => + | HeadMismatch leq Γ1 t1 Γ2 t2 => "Terms\n" ^ print_term Σ Γ1 t1 ^ "\nand\n" ^ print_term Σ Γ2 t2 ^ @@ -445,7 +445,7 @@ Section Typecheck. destruct HΣ. eapply red_confluence in r as (?&r1&r2); eauto. apply invert_red_sort in r2 as ->. - eapply whnf_red_shape in r1; eauto. + eapply whnf_red_inv in r1; eauto. depelim r1. clear Heq. rewrite H0 in n0. @@ -486,7 +486,7 @@ Section Typecheck. destruct HΣ. eapply red_confluence in r as (?&r1&r2); eauto. apply invert_red_prod in r2 as (?&?&(->&?)&?); auto. - eapply whnf_red_shape in r1; auto. + eapply whnf_red_inv in r1; auto. depelim r1. clear Heq. rewrite H0 in n0. @@ -571,8 +571,8 @@ Section Typecheck. cbn in *. rewrite app_nil_r in wh. apply red_mkApps_tInd in r2 as (?&->&?); auto. - apply whnf_red_mkApps in r1 as [(?&?)]; auto. - apply whnf_red_shape in r; [|now apply whnf_mkApps_inv in wh]. + apply whnf_red_mkApps_inv in r1 as [(?&?)]; auto. + apply whnf_red_inv in r; [|now apply whnf_mkApps_inv in wh]. depelim r. noconf e0. discriminate i0. diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index ea40d7f86..2c41c9841 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -551,9 +551,9 @@ Section Conversion. * assumption. Qed. - Notation eqb_termp pb := (eqb_termp_napp pb 0). - Notation eqb_term := (eqb_termp Conv). - Notation leqb_term := (eqb_termp Cumul). + Definition eqb_termp pb := (eqb_termp_napp pb 0). + Definition eqb_term := (eqb_termp Conv). + Definition leqb_term := (eqb_termp Cumul). Lemma leqb_term_spec t u : leqb_term t u -> @@ -743,22 +743,6 @@ Section Conversion. Arguments Success {_} _. Arguments Error {_} _. - (* Definition Ret s Γ t π t' π' := - forall leq, - conv_stack_ctx Γ π π' -> - match s with - | Reduction => - { b : bool | if b then conv_term leq Γ t π t' π' else True } - | Fallback - | Term => - isred (t, π) -> - isred (t', π') -> - { b : bool | if b then conv_term leq Γ t π t' π' else True } - | Args => - conv leq Σ (Γ ,,, stack_context π) t t' -> - { b : bool | if b then conv_term leq Γ t π t' π' else True } - end. *) - Definition isred_full Γ t π := isApp t = false /\ whnf RedFlags.nodelta Σ (Γ,,, stack_context π) (zipp t π). @@ -820,7 +804,7 @@ Section Conversion. now destruct decompose_stack. Qed. - (* Tailored view for isconv_prog *) + (* Tailored view for isconv_prog and precondition for fallback case *) Equations prog_discr (t1 t2 : term) : Prop := prog_discr (tApp _ _) (tApp _ _) := False ; prog_discr (tConst _ _) (tConst _ _) := False ; @@ -832,28 +816,32 @@ Section Conversion. prog_discr (tCoFix _ _) (tCoFix _ _) := False ; prog_discr _ _ := True. + (* Note that the arity of this should be the same for all s as otherwise + the extracted code is not correct *) Definition Ret s Γ t π t' π' := forall (leq : conv_pb), conv_stack_ctx Γ π π' -> match s with - | Reduction => - ConversionResult (conv_term leq Γ t π t' π') - - | Term => - isred_full Γ t π -> - isred_full Γ t' π' -> - ConversionResult (conv_term leq Γ t π t' π') - - | Fallback => - isred_full Γ t π -> - isred_full Γ t' π' -> - prog_discr t t' -> - ConversionResult (conv_term leq Γ t π t' π') - - | Args => - ConversionResult (∥conv_terms Σ (Γ,,, stack_context π) - (decompose_stack π).1 - (decompose_stack π').1∥) + | Term + | Fallback => isred_full Γ t π + | _ => True + end -> + match s with + | Term + | Fallback => isred_full Γ t' π' + | _ => True + end -> + match s with + | Fallback => prog_discr t t' + | _ => True + end -> + match s with + | Reduction + | Term + | Fallback => ConversionResult (conv_term leq Γ t π t' π') + | Args => ConversionResult (∥conv_terms Σ (Γ,,, stack_context π) + (decompose_stack π).1 + (decompose_stack π').1∥) end. Definition Aux s Γ t1 π1 t2 π2 h2 := @@ -879,11 +867,11 @@ Section Conversion. (only parsing). Notation isconv_red_raw leq t1 π1 t2 π2 aux := - (aux Reduction t1 π1 t2 π2 _ _ _ _ leq _) (only parsing). + (aux Reduction t1 π1 t2 π2 _ _ _ _ leq _ I I I) (only parsing). Notation isconv_prog_raw leq t1 π1 t2 π2 aux := - (aux Term t1 π1 t2 π2 _ _ _ _ leq _ _ _) (only parsing). + (aux Term t1 π1 t2 π2 _ _ _ _ leq _ _ _ I) (only parsing). Notation isconv_args_raw leq t1 π1 t2 π2 aux := - (aux Args t1 π1 t2 π2 _ _ _ _ leq _) (only parsing). + (aux Args t1 π1 t2 π2 _ _ _ _ leq _ I I I) (only parsing). Notation isconv_fallback_raw leq t1 π1 t2 π2 aux := (aux Fallback t1 π1 t2 π2 _ _ _ _ leq _ _ _ _) (only parsing). @@ -2662,58 +2650,6 @@ Section Conversion. now depelim r. Qed. -(* - Notation whne := (whne flags Σ). - - Lemma app_conv_inv : forall Γ t t' u u', - (forall napp, eq_term_upto_univ_napp Σ.1 (eq_universe Σ) (eq_universe Σ) napp t t' -> - eq_term Σ t t') -> - whne Γ t -> - whne Γ t' -> - Σ ;;; Γ |- tApp t u = tApp t' u' -> - Σ ;;; Γ |- t = t'. - Proof. - intros Γ t t' u u' cum_irr wh wh' hconv. - depind hconv. - + intros **. depelim e. eapply conv_refl, cum_irr; eassumption. - + (* Proof idea: invert the reduction r using thw whne hyp to apply the IH *) - todo "Completeness"%string. - + (* Proof idea: invert the reduction r using thw whne hyp to apply the IH *) - todo "Completeness"%string. - Qed. - - - - Lemma zipp_Case_conv_inv : - forall Γ indn p p' c c' brs brs' π π', - let t := tCase indn p c brs in - let t' := tCase indn p' c' brs' in - whne Γ t -> whne Γ t' -> - Σ ;;; Γ |- zipp t π = zipp t' π' -> - Σ ;;; Γ |- t = t'. - Proof. - (* Proof idea: decompose the stacks, prove by induction that the - each side is applied to the same number of arguments. *) - intros Γ indn p p' c c' brs brs' π π' t t' wht wht'. -(* - destruct (decompose_stack π) as [l s] eqn:Hpi. - destruct (decompose_stack π') as [l' s'] eqn: Hpi'. - unfold zipp; rewrite Hpi, Hpi'; clear s s' Hpi Hpi'. - induction l as [|x l IH] in l' |- *; destruct l' as [|x' l']. - + trivial. - + simpl. intro H. depelim H. - - - assert (cum_irr : forall napp, eq_term_upto_univ_napp Σ.1 (eq_universe Σ) (eq_universe Σ) napp t t' -> eq_term Σ t t'). - - intros n heq; depelim heq; now constructor. - - - *) - - todo "Completeness"%string. - Qed. *) - - (* See https://github.com/coq/coq/blob/master/kernel/reduction.ml#L367 *) - Lemma cumul_mkApps Γ hd args hd' args' : wf Σ -> Σ;;; Γ |- hd <= hd' -> @@ -2806,6 +2742,7 @@ Section Conversion. else false). + (* See https://github.com/coq/coq/blob/master/kernel/reduction.ml#L367 *) Opaque reduce_stack. Equations(noeqns) _isconv_prog (Γ : context) (leq : conv_pb) (t1 : term) (π1 : stack) (h1 : wtp Γ t1 π1) @@ -4371,9 +4308,9 @@ Section Conversion. destruct hΣ, hx. eapply conv_cum_red_inv' in H. 2: exact hΣ'. + 2: eassumption. 2: apply red_mkApps; [reflexivity|exact a]. 2: apply red_mkApps; [reflexivity|exact a0]. - 2: eassumption. apply conv_cum_mkApps_inv in H as [(_&conv_args)]; auto. 2: eapply whnf_conv_context; eauto. 2: eapply conv_context_sym; eauto. @@ -4390,9 +4327,9 @@ Section Conversion. destruct hΣ, hx. eapply conv_cum_red_inv' in H. 2: exact hΣ'. + 2: eassumption. 2: apply red_mkApps; [reflexivity|exact a]. 2: apply red_mkApps; [reflexivity|exact a0]. - 2: eassumption. apply conv_cum_mkApps_inv in H as [(?&?)]; auto. 2: eapply whnf_conv_context; eauto. 2: eapply conv_context_sym; eauto. @@ -4411,9 +4348,9 @@ Section Conversion. destruct hΣ, hx. eapply conv_cum_red_inv' in H. 2: exact hΣ'. + 2: eassumption. 2: apply red_mkApps; [reflexivity|exact a]. 2: apply red_mkApps; [reflexivity|exact a0]. - 2: eauto. apply conv_cum_mkApps_inv in H as [(?&?)]; auto. 2: eapply whnf_conv_context; eauto. 2: eapply conv_context_sym; eauto. @@ -4505,7 +4442,7 @@ Section Conversion. (aux : Aux' Γ t1 args1 l1 π1 t2 (appstack l2 π2) h2) : ConversionResult (∥conv_terms Σ (Γ,,, stack_context π1) l1 l2∥) by struct l1 := _isconv_args' leq Γ t1 args1 (u1 :: l1) π1 h1 hπ1 t2 (u2 :: l2) π2 h2 hπ2 hx aux - with aux u1 u2 args1 l1 (coApp t2 (appstack l2 π2)) _ _ _ _ Conv _ := { + with aux u1 u2 args1 l1 (coApp t2 (appstack l2 π2)) _ _ _ _ Conv _ I I I := { | Success H1 with _isconv_args' leq Γ t1 (args1 ++ [u1]) l1 π1 _ _ (tApp t2 u2) l2 π2 _ _ _ _ := { | Success H2 := yes ; | Error e herr := @@ -4630,11 +4567,6 @@ Section Conversion. rewrite 2!stack_context_appstack in hx. assumption. Qed. - (*Next Obligation. - pose proof (decompose_stack_eq _ _ _ (eq_sym eq1)). subst. - rewrite stack_context_appstack in he. - assumption. - Qed.*) Next Obligation. specialize (aux Reduction) as h. cbn in h. eapply h. all: auto. @@ -4949,7 +4881,7 @@ Section Conversion. whnf RedFlags.nodelta Σ (Γ,,, stack_context π) (mkApps t (decompose_stack π).1) -> None = reducible_head Γ t π h -> ∥∑ t' args', - whnf_red_shape_spec Σ (Γ,,, stack_context π) t t' × + whnf_red Σ (Γ,,, stack_context π) t t' × All2 (red Σ (Γ,,, stack_context π)) (decompose_stack π).1 args' × whnf RedFlags.default Σ (Γ,,, stack_context π) (mkApps t' args')∥. Proof. @@ -4978,23 +4910,23 @@ Section Conversion. depelim H; solve_discr; eauto with pcuic. - apply whnf_mkApps_tSort_inv in wh as ->. constructor; eexists _, []. - eauto using whnf_red_shape_spec with pcuic. + eauto using whnf_red with pcuic. - apply whnf_mkApps_tProd_inv in wh as ->. constructor; eexists _, []. - eauto using whnf_red_shape_spec with pcuic. + eauto using whnf_red with pcuic. - depelim wh; solve_discr. + apply whne_mkApps_inv in H as [|(?&?&?&?&?&?&?&?&?)]; auto; try discriminate. depelim H; solve_discr. discriminate. + rewrite H1. - constructor; eexists _, []; eauto using whnf_red_shape_spec with pcuic. + constructor; eexists _, []; eauto using whnf_red with pcuic. - apply whnf_mkApps_inv in wh; auto. depelim wh; solve_discr. depelim H; solve_discr. discriminate. - discriminate. - - constructor; eexists _, (decompose_stack π).1; eauto using whnf_red_shape_spec with pcuic. - - constructor; eexists _, (decompose_stack π).1; eauto using whnf_red_shape_spec with pcuic. + - constructor; eexists _, (decompose_stack π).1; eauto using whnf_red with pcuic. + - constructor; eexists _, (decompose_stack π).1; eauto using whnf_red with pcuic. - eapply unfold_one_fix_None in H0 as [(?&?&?)]. constructor; eexists _, x. split; [constructor; eauto with pcuic|]. @@ -5023,34 +4955,6 @@ Section Conversion. - admit. Admitted. - Lemma whnf_red_shape_spec_red Γ t t' : - whnf_red_shape_spec Σ Γ t t' -> - ∥red Σ Γ t t'∥. - Proof. - intros wh. - depelim wh; constructor; eauto with pcuic. - - apply red_evar; auto. - - apply red_app; auto. - - apply red_fix_congr. - eapply All2_impl; eauto. - cbn. - intros ? ? (->&->&r1&r2). - eauto. - - apply red_case; auto. - eapply All2_impl; eauto. - cbn. - intros ? ? (->&?). - eauto. - - apply red_proj_c; auto. - - apply red_prod; auto. - - apply PCUICConfluence.red_abs_alt; auto. - - apply red_cofix_congr. - eapply All2_impl; eauto. - cbn. - intros ? ? (->&->&r1&r2). - eauto. - Qed. - (* TODO Factorise *) Equations(noeqns) _isconv_fallback (Γ : context) (leq : conv_pb) (t1 : term) (π1 : stack) (h1 : wtp Γ t1 π1) @@ -5495,15 +5399,15 @@ Section Conversion. apply reducible_head_None in nored1 as [(?&?&s1&r1&wh1)]; auto. apply reducible_head_None in nored2 as [(?&?&s2&r2&wh2)]; auto. destruct hΣ, hx. - apply whnf_red_shape_spec_red in s1 as H'. + apply whnf_red_red in s1 as H'. destruct H'. - apply whnf_red_shape_spec_red in s2 as H'. + apply whnf_red_red in s2 as H'. destruct H'. eapply conv_cum_red_inv' in H. - 2: apply hΣ'. + 2: exact hΣ'. + 2: eassumption. 2: apply red_mkApps; [eassumption|eassumption]. 2: apply red_mkApps; [eassumption|eassumption]. - 2: eassumption. apply conv_cum_mkApps_inv in H as [(?&?)]; auto. 2: now depelim s1. 2: now depelim s2. @@ -5519,15 +5423,15 @@ Section Conversion. apply reducible_head_None in nored1 as [(?&?&s1&rargs1&wh1)]; auto. apply reducible_head_None in nored2 as [(?&?&s2&rargs2&wh2)]; auto. destruct hΣ, hx. - apply whnf_red_shape_spec_red in s1 as H'. + apply whnf_red_red in s1 as H'. destruct H'. - apply whnf_red_shape_spec_red in s2 as H'. + apply whnf_red_red in s2 as H'. destruct H'. eapply conv_cum_red_inv' in H. - 2: apply hΣ'. + 2: exact hΣ'. + 2: eassumption. 2: apply red_mkApps; eassumption. 2: apply red_mkApps; eassumption. - 2: eassumption. apply conv_cum_mkApps_inv in H as [(conv_hds&_)]; auto. 2: now depelim s1. 2: now depelim s2. @@ -5545,14 +5449,14 @@ Section Conversion. depelim s2. todo "R_global_instance => compare_global_instance". } all: apply conv_cum_alt in conv_hds as [(?&?&(r1&r2)&?)]. - all: eapply whnf_red_shape in r1; auto. - all: depelim r1. - all: depelim e. + all: eapply whnf_red_inv in r1; auto. + all: inversion r1; subst; clear r1. + all: inversion e; subst; clear e. all: apply whnf_mkApps_inv in wh2; [|now depelim s2]. all: eapply whnf_conv_context in wh2; [|apply conv_context_sym; eauto]. - all: eapply whnf_red_shape in r2; auto. - all: depelim r2. - all: depelim s2. + all: eapply whnf_red_inv in r2; auto. + all: inversion r2; subst; clear r2. + all: inversion s2; subst; clear s2. all: destruct hdiscr. - now rewrite Nat.eqb_refl in noteq. - now rewrite eq_string_refl in noteq. @@ -5570,13 +5474,13 @@ Section Conversion. (aux : Aux s Γ t1 π1 t2 π2 h2) : Ret s Γ t1 π1 t2 π2 := _isconv Reduction Γ t1 π1 h1 t2 π2 h2 aux := - λ { | leq | hx := _isconv_red Γ leq t1 π1 h1 t2 π2 h2 hx aux } ; + λ { | leq | hx | _ | _ | _ := _isconv_red Γ leq t1 π1 h1 t2 π2 h2 hx aux } ; _isconv Term Γ t1 π1 h1 t2 π2 h2 aux := - λ { | leq | hx | r1 | r2 := _isconv_prog Γ leq t1 π1 h1 t2 π2 h2 hx r1 r2 aux } ; + λ { | leq | hx | r1 | r2 | _ := _isconv_prog Γ leq t1 π1 h1 t2 π2 h2 hx r1 r2 aux } ; _isconv Args Γ t1 π1 h1 t2 π2 h2 aux := - λ { | leq | hx := _isconv_args leq Γ t1 π1 h1 t2 π2 h2 hx aux } ; + λ { | leq | hx | _ | _ | _ := _isconv_args leq Γ t1 π1 h1 t2 π2 h2 hx aux } ; _isconv Fallback Γ t1 π1 h1 t2 π2 h2 aux := λ { | leq | hx | r1 | r2 | hd := _isconv_fallback Γ leq t1 π1 h1 t2 π2 h2 r1 r2 hd hx aux }. @@ -5614,7 +5518,7 @@ Section Conversion. | ConvError : ConversionError -> ConversionResultSummary. Definition isconv Γ leq t1 π1 h1 t2 π2 h2 hx := - match isconv_full Reduction Γ t1 π1 h1 t2 π2 h2 leq hx with + match isconv_full Reduction Γ t1 π1 h1 t2 π2 h2 leq hx I I I with | Success _ => ConvSuccess | Error e _ => ConvError e end. From 990a5365a6c0cc71f50b7d33027414ae03bced53 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Sun, 25 Oct 2020 16:56:46 +0100 Subject: [PATCH 15/26] Finish last cases except sorts --- pcuic/theories/PCUICConversion.v | 27 -- pcuic/theories/PCUICConversionInversion.v | 21 + safechecker/theories/PCUICSafeConversion.v | 435 ++++++++++++++++++--- safechecker/theories/PCUICSafeReduce.v | 90 +++-- template-coq/theories/common/uGraph.v | 17 + 5 files changed, 474 insertions(+), 116 deletions(-) diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index 4b7a0b15d..f8707fc3c 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -1461,33 +1461,6 @@ Section Inversions. + apply conv_Case_brs. all: assumption. Qed. - Lemma Case_conv_inv : - forall Γ indn p p' c c' brs brs', - Σ ;;; Γ |- tCase indn p c brs = tCase indn p' c' brs' -> - Σ ;;; Γ |- p = p' × Σ ;;; Γ |- c = c' × - All2 (fun u v => u.1 = v.1 × Σ ;;; Γ |- u.2 = v.2) brs brs'. - Proof. - (* intros Γ indn p p' c c' brs brs' H. - depind H. - + depelim e; split; [| split]. - 1,2: by apply conv_refl. - apply (All2_impl a). - intros x y [? ?] ; split. 1: by []. - by apply conv_refl. *) - (* TODO needs to add a whne hypothesis *) - todo "Completeness"%string. - Qed. - -(* - Lemma Case_conv_cum_inv : - forall leq Γ indn p p' c c' brs brs', - conv_cum leq (tCase indn p c brs) (tCase indn p' c' brs') -> - Σ ;;; Γ |- p = p' /\ Σ ;;; Γ |- c = c' /\ - ∥ All2 (fun u v => u.1 = v.1 × Σ ;;; Γ |- u.2 = v.2) brs brs' ∥. - Proof. - todo "Completeness"%string. - Qed. *) - Lemma conv_Proj_c : forall Γ p u v, Σ ;;; Γ |- u = v -> diff --git a/pcuic/theories/PCUICConversionInversion.v b/pcuic/theories/PCUICConversionInversion.v index 478c87dea..a71ea5002 100644 --- a/pcuic/theories/PCUICConversionInversion.v +++ b/pcuic/theories/PCUICConversionInversion.v @@ -123,6 +123,27 @@ Section fixed. (intros P Q H; split; intros [p]; constructor; apply H in p; auto). destruct leq; cbn; apply H; [apply conv_alt_red|apply cumul_alt]. Qed. + + Lemma conv_terms_alt Γ args args' : + conv_terms Σ Γ args args' <~> + ∑ argsr argsr', + All2 (red Σ Γ) args argsr × + All2 (red Σ Γ) args' argsr' × + All2 (eq_term Σ Σ) argsr argsr'. + Proof. + split. + - intros conv. + induction conv. + + exists [], []; eauto with pcuic. + + apply conv_alt_red in r as (xr&yr&(xred&yred)&xy). + specialize IHconv as (argsr&argsr'&?&?&?). + exists (xr :: argsr), (yr :: argsr'). + eauto 7 with pcuic. + - intros (argsr&argsr'&r&r'&eqs). + induction eqs in args, args', r, r' |- *; depelim r; depelim r'; [constructor|]. + constructor; auto. + apply conv_alt_red; eauto. + Qed. Lemma eq_term_eq_termp leq x y : eq_term Σ Σ x y -> diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 2c41c9841..d9adce766 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -1672,52 +1672,125 @@ Section Conversion. + eapply global_ext_uctx_consistent. eapply hΣ'. Qed. + + Lemma mem_level_declared_level x : + LevelSet.mem x (global_ext_levels Σ) -> + UnivExprSet.For_all + (fun e => + on_Some_or_None (fun l => LevelSet.In (NoPropLevel.to_level l) (global_ext_uctx Σ).1) + (UnivExpr.get_noprop e)) (Universe.make x). + Proof. + intros mem ? ->%UnivExprSet.singleton_spec. + destruct (UnivExpr.get_noprop (UnivExpr.make x)) eqn:np; [|easy]. + cbn. + replace x with (NoPropLevel.to_level t) in *. + 2: { destruct x; cbn in *; try easy. + all: now inv np. } + unfold global_ext_levels in *. + now apply LevelSet.mem_spec in mem. + Qed. - Lemma eq_universe_instance_spec udecl u v : - consistent_instance_ext Σ udecl u -> - consistent_instance_ext Σ udecl v -> - R_universe_instance (eq_universe (global_ext_constraints Σ)) u v -> - eqb_universe_instance u v. + Lemma conv_pb_rel_spec leq x y : + LevelSet.mem x (global_ext_levels Σ) -> + LevelSet.mem y (global_ext_levels Σ) -> + conv_pb_rel leq (global_ext_constraints Σ) (Universe.make x) (Universe.make y) -> + conv_pb_relb leq (Universe.make x) (Universe.make y). Proof. - intros consu consv r. - apply Forall2_forallb2. - unfold consistent_instance_ext, consistent_instance, R_universe_instance in *. - destruct udecl; [destruct u,v; cbn in *; try congruence; constructor|]. - destruct consu as (_&all_mem_u&_&_). - destruct consv as (_&all_mem_v&_&_). - apply forallb_Forall in all_mem_u. - apply forallb_Forall in all_mem_v. - apply Forall2_map. - apply Forall2_map_inv in r. - eapply Forall_Forall2_and in r; [|exact all_mem_u]. - eapply Forall_Forall2_and' in r; [|exact all_mem_v]. - clear all_mem_u all_mem_v. - cbn beta in r. - eapply Forall2_impl; [eassumption|]; cbn beta in *. - intros x y ((x_mem&x_eq_y)&y_mem). - eapply eq_universe_spec'; eauto. - - eapply wf_ext_global_uctx_invariants. - eapply hΣ'. - - eapply global_ext_uctx_consistent. - eapply hΣ'. - - intros ? ->%UnivExprSet.singleton_spec. - destruct (UnivExpr.get_noprop (UnivExpr.make x)) eqn:np; [|easy]. - cbn. - replace x with (NoPropLevel.to_level t) in *. - 2: { destruct x; cbn in *; try easy. - all: now inv np. } - unfold global_ext_levels in *. - now apply LevelSet.mem_spec in x_mem. - - intros ? ->%UnivExprSet.singleton_spec. - destruct (UnivExpr.get_noprop (UnivExpr.make y)) eqn:np; [|easy]. - cbn. - replace y with (NoPropLevel.to_level t) in *. - 2: { destruct y; cbn in *; try easy. - all: now inv np. } - unfold global_ext_levels in *. - now apply LevelSet.mem_spec in y_mem. + intros memx memy r. + destruct leq. + - eapply eq_universe_spec'; eauto. + + eapply wf_ext_global_uctx_invariants. + eapply hΣ'. + + eapply global_ext_uctx_consistent. + eapply hΣ'. + + now apply mem_level_declared_level. + + now apply mem_level_declared_level. + - eapply leq_universe_spec'; eauto. + + eapply wf_ext_global_uctx_invariants. + eapply hΣ'. + + eapply global_ext_uctx_consistent. + eapply hΣ'. + + now apply mem_level_declared_level. + + now apply mem_level_declared_level. + Qed. + + Lemma R_universe_instance_spec u u' : + Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) u -> + Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) u' -> + R_universe_instance (eq_universe (global_ext_constraints Σ)) u u' -> + eqb_universe_instance u u'. + Proof. + intros memu memu' r. + induction u in u', memu, memu', r |- *. + - now destruct u'. + - destruct u'; [easy|]. + depelim memu. + depelim memu'. + depelim r. + cbn in *. + apply Bool.andb_true_iff. + split. + + apply (conv_pb_rel_spec Conv); auto. + + now apply IHu. + Qed. + + Lemma R_universe_variance_spec leq v u u' : + LevelSet.mem u (global_ext_levels Σ) -> + LevelSet.mem u' (global_ext_levels Σ) -> + R_universe_variance (eq_universe Σ) (conv_pb_rel leq Σ) v u u' -> + compare_universe_variance (check_eqb_universe G) (conv_pb_relb leq) v u u'. + Proof. + intros memu memu' r. + destruct v; cbn in *; auto. + - apply conv_pb_rel_spec; auto. + - apply (conv_pb_rel_spec Conv); auto. Qed. + Lemma R_universe_instance_variance_spec leq v u u' : + Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) u -> + Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) u' -> + R_universe_instance_variance (eq_universe Σ) (conv_pb_rel leq Σ) v u u' -> + compare_universe_instance_variance (check_eqb_universe G) (conv_pb_relb leq) v u u'. + Proof. + intros memu memu' r. + induction u in v, u', memu, memu', r |- *. + - now destruct u'. + - destruct u'; [easy|]. + depelim memu. + depelim memu'. + cbn in *. + destruct v; auto. + apply Bool.andb_true_iff. + destruct r. + split. + + apply R_universe_variance_spec; auto. + + now apply IHu. + Qed. + + Lemma R_global_instance_spec u v leq gr napp : + Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) u -> + Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) v-> + R_global_instance Σ (eq_universe Σ) (conv_pb_rel leq Σ) gr napp u v -> + compare_global_instance Σ (check_eqb_universe G) (conv_pb_relb leq) gr napp u v. + Proof. + intros consu consv r. + unfold compare_global_instance, R_global_instance, R_opt_variance in *. + destruct global_variance. + - apply R_universe_instance_variance_spec; auto. + - apply R_universe_instance_spec; auto. + Qed. + + Lemma consistent_instance_ext_all_mem udecl u : + consistent_instance_ext Σ udecl u -> + Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) u. + Proof. + intros cons. + unfold consistent_instance_ext, consistent_instance in *. + destruct udecl; [now destruct u|]. + destruct cons as (_&mems&_). + now apply forallb_Forall. + Qed. + Lemma wellformed_nonarity Γ t : destArity [] t = None -> wellformed Σ Γ t -> @@ -1859,7 +1932,9 @@ Section Conversion. apply wellformed_zipc_tConst_inv in h1 as (cst1&decl1&cons1). apply wellformed_zipc_tConst_inv in h2 as (cst2&decl2&cons2). eapply PCUICWeakeningEnv.declared_constant_inj in decl1; eauto; subst. - eapply eq_universe_instance_spec in r; eauto. + apply consistent_instance_ext_all_mem in cons1. + apply consistent_instance_ext_all_mem in cons2. + eapply R_universe_instance_spec in r; auto. Qed. (* TODO (RE)MOVE *) @@ -4606,7 +4681,7 @@ Section Conversion. | @exist (Some (narg, fn)) eq2 with inspect (decompose_stack ρ) := { | @exist (args, ξ) eq' := Some (tCase (ind, par) p (mkApps fn args) brs) } ; - | @exist None eq2 := None + | @exist None eq2 := False_rect _ _ (* why does ! not work in this file? *) } ; | ccview_other t _ := None } @@ -4620,6 +4695,31 @@ Section Conversion. [? [? [? [? [? [? [ht0 [? ?]]]]]]]]]]]]]]]. left; eexists. eassumption. Qed. + Next Obligation. + exfalso. + match type of eq with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; + pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as decomp + end. + rewrite <- eq in r, decomp. + cbn in *. + destruct (decompose_stack ρ) eqn:decomp'. + apply decompose_stack_eq in decomp' as ->. + cbn in *; subst. + rewrite zipc_appstack in r. + cbn in *. + clear eq. + apply wellformed_nonarity in h as (?&typ); auto. + destruct hΣ. + apply inversion_Case in typ as (?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?); auto. + eapply PCUICSR.subject_reduction in t0; eauto. + apply PCUICValidity.inversion_mkApps in t0 as (?&?&?); auto. + apply inversion_CoFix in t0 as (?&?&?&?&?&?&?); auto. + unfold unfold_cofix in eq2. + rewrite e3 in eq2. + congruence. + Qed. Lemma unfold_one_case_cored : forall Γ ind par p c brs h t, @@ -4647,6 +4747,9 @@ Section Conversion. eapply cored_red_cored. + constructor. eapply red_iota. + eapply red_case_c. eassumption. + - match type of eq with + | _ = False_rect _ ?f => destruct f + end. - match type of e with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; @@ -4664,6 +4767,42 @@ Section Conversion. + constructor. eapply red_cofix_case. eauto. + eapply red_case_c. eassumption. Qed. + + Lemma unfold_one_case_None Γ ind par p c brs h : + None = unfold_one_case Γ ind par p c brs h -> + ∥∑ c', red Σ Γ c c' × whne RedFlags.default Σ Γ c'∥. + Proof. + funelim (unfold_one_case Γ ind par p c brs h); intros [=]. + - match type of e with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; + pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as wh; + pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as isr; + pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as decomp + end. + rewrite <- e in r, decomp, wh, isr. + cbn in *. + constructor; exists (zipc t0 s). + split; [easy|]. + destruct (decompose_stack s) eqn:decomp'. + apply decompose_stack_eq in decomp' as ->. + cbn in *; subst. + rewrite zipc_appstack in r |- *. + rewrite zipp_appstack, stack_context_appstack in wh. + cbn in *. + destruct isr as (noapp&_). + clear e H. + apply wellformed_nonarity in h as (?&typ); auto. + destruct hΣ. + eapply PCUICSR.subject_reduction in typ. + 2: eauto. + 2: eapply red_case_c; eauto. + eapply whnf_case_arg_whne; eauto. + now destruct t0. + - match type of H2 with + | _ = False_rect _ ?f => destruct f + end. + Qed. Equations unfold_one_proj (Γ : context) (p : projection) (c : term) (h : wellformed Σ Γ (tProj p c)) : option term := @@ -4674,14 +4813,14 @@ Section Conversion. | cc0view_construct ind' ui with inspect (decompose_stack ρ) := { | @exist (args, ξ) eq' with inspect (nth_error args (pars + narg)) := { | @exist (Some arg) eq2 := Some arg ; - | @exist None _ := None + | @exist None _ := False_rect _ _ } } ; | cc0view_cofix mfix idx with inspect (decompose_stack ρ) := { | @exist (args, ξ) eq' with inspect (unfold_cofix mfix idx) := { | @exist (Some (narg, fn)) eq2 := Some (tProj (i, pars, narg) (mkApps fn args)) ; - | @exist None eq2 := None + | @exist None eq2 := False_rect _ _ } } ; | cc0view_other t _ := None @@ -4695,6 +4834,76 @@ Section Conversion. destruct h as [uni [mdecl [idecl [pdecl [args' [? [? [? ?]]]]]]]]. left. eexists. eassumption. Qed. + Next Obligation. + match type of eq with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; + pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as decomp + end. + rewrite <- eq in r, decomp. + cbn in *. + destruct (decompose_stack ρ) eqn:decomp'. + apply decompose_stack_eq in decomp' as ->. + noconf eq'. + cbn in *; subst. + rewrite zipc_appstack in r. + cbn in *. + clear eq. + apply wellformed_nonarity in h as (?&typ); auto. + destruct hΣ. + apply inversion_Proj in typ as (?&?&?&?&?&decl&?&?&?); auto. + assert (d := decl). + eapply PCUICSR.subject_reduction in t; eauto. + apply PCUICValidity.inversion_mkApps in t as spine; auto. + destruct spine as (?&typ_ctor&_). + apply inversion_Construct in typ_ctor as (?&?&?&?&decl_ctor&?&?); auto. + destruct hΣ as [wfΣ]. + eapply PCUICInductiveInversion.Construct_Ind_ind_eq with (wfΣ0 := wfΣ) (Hdecl := decl_ctor) in t. + cbn in t. + destruct decl_ctor as (?&?); cbn in *. + destruct All2_nth_error_Some; cbn in t. + destruct p0; cbn in t. + destruct t as ((((<-&?)&?)&?)&?). + clear s. + destruct decl as (?&?&?). + pose proof (PCUICInductiveInversion.declared_inductive_unique_sig d0 H) as H''; noconf H''. + cbn in *. + pose proof (PCUICWeakeningEnv.on_declared_projection wfΣ d) as (_&proj'). + destruct d. + pose proof (PCUICInductiveInversion.declared_inductive_unique_sig H d) as H''; noconf H''. + cbn in *. + destruct ind_cshapes; [easy|]. + destruct l; [|easy]. + noconf e1. + destruct proj' as ((_&?)&_). + symmetry in wildcard. + apply nth_error_None in wildcard. + lia. + Qed. + Next Obligation. + match type of eq with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; + pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as decomp + end. + rewrite <- eq in r, decomp. + cbn in *. + destruct (decompose_stack ρ) eqn:decomp'. + apply decompose_stack_eq in decomp' as ->. + cbn in *; subst. + rewrite zipc_appstack in r. + cbn in *. + clear eq. + apply wellformed_nonarity in h as (?&typ); auto. + destruct hΣ. + apply inversion_Proj in typ as (?&?&?&?&?&?&?&?&?); auto. + eapply PCUICSR.subject_reduction in t; eauto. + apply PCUICValidity.inversion_mkApps in t as (?&?&?); auto. + apply inversion_CoFix in t as (?&?&?&?&?&?&?); auto. + unfold unfold_cofix in eq2. + rewrite e0 in eq2. + congruence. + Qed. Lemma unfold_one_proj_cored : forall Γ p c h t, @@ -4721,6 +4930,9 @@ Section Conversion. eapply cored_red_cored. + constructor. eapply red_proj. eauto. + eapply red_proj_c. eassumption. + - match type of eq with + | _ = False_rect _ ?f => destruct f + end. - match type of e with | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; @@ -4736,6 +4948,48 @@ Section Conversion. eapply cored_red_cored. + constructor. eapply red_cofix_proj. eauto. + eapply red_proj_c. eassumption. + - match type of eq with + | _ = False_rect _ ?f => destruct f + end. + Qed. + + Lemma unfold_one_proj_None Γ p c h : + None = unfold_one_proj Γ p c h -> + ∥∑ c', red Σ Γ c c' × whne RedFlags.default Σ Γ c'∥. + Proof. + funelim (unfold_one_proj Γ p c h); intros [=]. + - match type of e with + | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => + pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; + pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as wh; + pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as isr; + pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as decomp + end. + rewrite <- e in r, decomp, wh, isr. + cbn in *. + constructor; exists (zipc t0 s). + split; [easy|]. + destruct (decompose_stack s) eqn:decomp'. + apply decompose_stack_eq in decomp' as ->. + cbn in *; subst. + rewrite zipc_appstack in r |- *. + rewrite zipp_appstack, stack_context_appstack in wh. + cbn in *. + destruct isr as (noapp&_). + clear e H. + apply wellformed_nonarity in h as (?&typ); auto. + destruct hΣ. + eapply PCUICSR.subject_reduction in typ. + 2: eauto. + 2: eapply red_proj_c; eauto. + eapply whnf_proj_arg_whne; eauto. + now destruct t0. + - match type of H3 with + | _ = False_rect _ ?f => destruct f + end. + - match type of H3 with + | _ = False_rect _ ?f => destruct f + end. Qed. Equations reducible_head (Γ : context) (t : term) (π : stack) @@ -4951,9 +5205,23 @@ Section Conversion. apply wellformed_nonarity in h as (?&typ); auto. apply inversion_Const in typ as (?&?&?&?); auto. unfold declared_constant in d; congruence. - - admit. - - admit. - Admitted. + - clear H. + apply unfold_one_case_None in e as [(c'&r&whcase)]. + constructor; exists (tCase (i, n) p c' brs), (decompose_stack π).1. + split. + + constructor; eauto with pcuic. + + split; [eauto with pcuic|]. + apply whnf_mkApps. + auto. + - clear H. + apply unfold_one_proj_None in e as [(c'&r&whproj)]. + constructor; exists (tProj p0 c'), (decompose_stack π).1. + split. + + constructor; eauto with pcuic. + + split; [eauto with pcuic|]. + apply whnf_mkApps. + auto. + Qed. (* TODO Factorise *) Equations(noeqns) _isconv_fallback (Γ : context) (leq : conv_pb) @@ -5385,12 +5653,18 @@ Section Conversion. Qed. Next Obligation. destruct h, hΣ. - todo "variant of conv_cum_mkApps". - (*constructor. + apply conv_terms_alt in X as (argsr&argsr'&?&?&?). rewrite !zipp_as_mkApps. - apply mkApps_conv_args; auto. - constructor. - eapply eqb_term_spec. auto.*) + apply conv_cum_alt. + constructor; eexists _, _. + split; [split|]. + - apply red_mkApps; [reflexivity|eassumption]. + - apply red_mkApps; [reflexivity|eassumption]. + - apply eq_term_upto_univ_napp_mkApps; auto. + rewrite Nat.add_0_r. + apply All2_length in a. + rewrite a in eq1. + apply eqb_termp_napp_spec; eauto. Qed. Next Obligation. apply h; clear h. @@ -5416,6 +5690,7 @@ Section Conversion. constructor. eapply conv_terms_red'; eauto. Qed. + Next Obligation. unfold eqb_termp_napp in noteq. destruct ir1 as (notapp1&whδ1), ir2 as (notapp2&whδ2). @@ -5443,11 +5718,43 @@ Section Conversion. 9: { destruct conv_hds as [H]. depelim H. depelim s2. - todo "R_global_instance => compare_global_instance". } + zip fold in h1. + zip fold in h2. + apply wellformed_context in h1; auto. + clear aux. + apply wellformed_context in h2; auto. + apply wellformed_nonarity in h1 as (?&typ1); auto. + apply wellformed_nonarity in h2 as (?&typ2); auto. + apply inversion_Ind in typ1 as (?&?&?&?&?&?); auto. + apply inversion_Ind in typ2 as (?&?&?&?&?&?); auto. + apply consistent_instance_ext_all_mem in c1. + apply consistent_instance_ext_all_mem in c. + apply R_global_instance_spec in r; auto. + rewrite eq_inductive_refl in noteq. + apply All2_length in rargs1. + rewrite <- rargs1 in r. + cbn in *. + easy. } 9: { destruct conv_hds as [H]. depelim H. depelim s2. - todo "R_global_instance => compare_global_instance". } + zip fold in h1. + zip fold in h2. + apply wellformed_context in h1; auto. + clear aux. + apply wellformed_context in h2; auto. + apply wellformed_nonarity in h1 as (?&typ1); auto. + apply wellformed_nonarity in h2 as (?&typ2); auto. + apply inversion_Construct in typ1 as (?&?&?&?&?&?&?); auto. + apply inversion_Construct in typ2 as (?&?&?&?&?&?&?); auto. + apply consistent_instance_ext_all_mem in c1. + apply consistent_instance_ext_all_mem in c. + apply R_global_instance_spec in r; auto. + rewrite eq_inductive_refl, Nat.eqb_refl in noteq. + apply All2_length in rargs1. + rewrite <- rargs1 in r. + cbn in *. + easy. } all: apply conv_cum_alt in conv_hds as [(?&?&(r1&r2)&?)]. all: eapply whnf_red_inv in r1; auto. all: inversion r1; subst; clear r1. @@ -5465,7 +5772,21 @@ Section Conversion. destruct h1 as [(?&typ)|[(?&?&?&?)]]. + now apply inversion_Evar in typ. + cbn in e; congruence. - - todo "conv_pb_rel -> conv_pb_relb". + - zip fold in h1. + zip fold in h2. + apply wellformed_context in h1; auto. + clear aux. + apply wellformed_context in h2; auto. + eapply conv_pb_rel_spec in H0. + assert (forall Σ Γ s, wellformed Σ Γ (tSort s) -> + Forall (fun s => LevelSet.In s (global_ext_levels Σ)) s). + destruct h1 as [(?&typ)|]. + + apply inversion_Sort in typ as (?&?&?&?&?); auto. + apply conv_pb_rel_spec in H0: + admit. + + destruct H as [(?&?&?&?)]. + cbn in *. + apply conv_pb_rel_spec in H0. Qed. Equations _isconv (s : state) (Γ : context) diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index a5f796762..8fcb72d4f 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -1325,6 +1325,58 @@ Section Reduce. congruence. Qed. + Lemma whnf_case_arg_whne Γ hd args ind par p brs T : + match hd with + | tApp _ _ + | tConstruct _ _ _ + | tCoFix _ _ => False + | _ => True + end -> + whnf flags Σ Γ (mkApps hd args) -> + Σ;;; Γ |- tCase (ind, par) p (mkApps hd args) brs : T -> + whne flags Σ Γ (mkApps hd args). + Proof. + intros shape wh typ. + destruct hΣ. + apply inversion_Case in typ as (?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?); auto. + eapply whnf_non_ctor_finite_ind_typed; try eassumption. + - unfold isConstruct_app. + now rewrite decompose_app_mkApps; destruct hd. + - unfold isCoFinite in e1. + unfold check_recursivity_kind. + cbn. + unfold declared_inductive, declared_minductive in d. + cbn in d. + rewrite (proj1 d). + now destruct ind_finite. + Qed. + + Lemma whnf_proj_arg_whne Γ hd args p T : + match hd with + | tApp _ _ + | tCoFix _ _ => False + | tConstruct _ c _ => c <> 0 + | _ => True + end -> + whnf flags Σ Γ (mkApps hd args) -> + Σ;;; Γ |- tProj p (mkApps hd args) : T -> + whne flags Σ Γ (mkApps hd args). + Proof. + intros shape wh typ. + destruct hΣ as [wfΣ]. + apply inversion_Proj in typ as (?&?&?&?&?&?&?&?&?); auto. + cbn in *. + eapply whnf_non_ctor_non_cofix_ind_typed; try eassumption. + - unfold isConstruct_app. + rewrite decompose_app_mkApps by (now destruct hd). + cbn. + destruct hd; try easy. + eapply PCUICInductiveInversion.projected_constructor_eq in d; [|eassumption|eassumption]. + congruence. + - unfold isCoFix_app. + now rewrite decompose_app_mkApps; destruct hd. + Qed. + Lemma reduce_stack_whnf : forall Γ t π h, let '(u, ρ) := reduce_stack Γ t π h in @@ -1575,22 +1627,10 @@ Section Reduce. try rewrite stack_context_appstack in H; cbn in *). rewrite app_nil_r in *. - destruct hΣ. - apply inversion_Case in typ as (?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?); auto. apply whnf_mkApps, whne_case. - eapply whnf_non_ctor_finite_ind_typed; try eassumption. - + unfold isConstruct_app. - destruct H as (noapp&_). - rewrite decompose_app_mkApps by (now rewrite noapp). - cbn. - now destruct t0. - + unfold isCoFinite in e1. - unfold check_recursivity_kind. - cbn. - unfold declared_inductive, declared_minductive in d0. - cbn in d0. - rewrite (proj1 d0). - now destruct ind_finite. + eapply whnf_case_arg_whne; eauto. + destruct H as (noapp&_). + now destruct t0. - match goal with | |- context [ reduce ?x ?y ?z ] => case_eq (reduce x y z) ; @@ -1639,24 +1679,10 @@ Section Reduce. try rewrite stack_context_appstack in H; cbn in *). rewrite app_nil_r in *. - destruct hΣ as [wfΣ]. - apply inversion_Proj in typ as (?&?&?&?&?&?&?&?&?); auto. - cbn in *. apply whnf_mkApps, whne_proj. - eapply whnf_non_ctor_non_cofix_ind_typed; try eassumption. - + unfold isConstruct_app. - destruct H as (noapp&_). - rewrite decompose_app_mkApps by (now rewrite noapp). - cbn. - destruct t0; try easy. - cbn in *. - eapply PCUICInductiveInversion.projected_constructor_eq in t; [|eassumption|eassumption]. - congruence. - + unfold isCoFix_app. - destruct H as (noapp&_). - rewrite decompose_app_mkApps by (now rewrite noapp). - cbn. - now destruct t0. + eapply whnf_proj_arg_whne; eauto. + destruct H as (noapp&_). + now destruct t0. - match goal with | |- context [ reduce ?x ?y ?z ] => case_eq (reduce x y z) ; diff --git a/template-coq/theories/common/uGraph.v b/template-coq/theories/common/uGraph.v index c85af61f9..2814d1601 100644 --- a/template-coq/theories/common/uGraph.v +++ b/template-coq/theories/common/uGraph.v @@ -1720,6 +1720,23 @@ Section CheckLeq2. destruct (gc_of_constraints uctx.2). cbn in *. exact H. exact I. Qed. + + Lemma leq_universe_spec' u1 u2 : + levels_declared u1 -> + levels_declared u2 -> + leq_universe uctx.2 u1 u2 -> + check_leqb_universe G u1 u2. + Proof. + intros decl1 decl2. + apply levels_gc_declared_declared in decl1. + apply levels_gc_declared_declared in decl2. + rewrite gc_leq_universe_iff. + unfold is_graph_of_uctx, gc_of_uctx in HG. + destruct gc_of_constraints; [cbn in *|contradiction HG]. + intros eq. + apply <- check_leqb_universe_spec; eauto. + exact eq. + Qed. Lemma check_eqb_universe_spec' u1 u2 : check_eqb_universe G u1 u2 -> eq_universe uctx.2 u1 u2. From a12ee2fee90f81acbcb785e4bc4dd4bb0ace2131 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Sun, 25 Oct 2020 19:32:42 +0100 Subject: [PATCH 16/26] Finish necessary universe stuff; add axiom for wf sorts --- safechecker/theories/PCUICSafeConversion.v | 120 ++++++++++++--------- 1 file changed, 67 insertions(+), 53 deletions(-) diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index d9adce766..5c2bb8db5 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -1672,49 +1672,62 @@ Section Conversion. + eapply global_ext_uctx_consistent. eapply hΣ'. Qed. - - Lemma mem_level_declared_level x : - LevelSet.mem x (global_ext_levels Σ) -> + + Arguments LevelSet.mem : simpl never. + + Lemma For_all_to_spec u : + UnivExprSet.For_all + (fun e => LevelSet.mem (UnivExpr.get_level e) (global_ext_levels Σ)) (Universe.t_set u) -> UnivExprSet.For_all (fun e => - on_Some_or_None (fun l => LevelSet.In (NoPropLevel.to_level l) (global_ext_uctx Σ).1) - (UnivExpr.get_noprop e)) (Universe.make x). + on_Some_or_None (fun l => LevelSet.In (NoPropLevel.to_level l) (global_ext_levels Σ)) + (UnivExpr.get_noprop e)) u. Proof. - intros mem ? ->%UnivExprSet.singleton_spec. - destruct (UnivExpr.get_noprop (UnivExpr.make x)) eqn:np; [|easy]. - cbn. - replace x with (NoPropLevel.to_level t) in *. - 2: { destruct x; cbn in *; try easy. - all: now inv np. } - unfold global_ext_levels in *. - now apply LevelSet.mem_spec in mem. + intros all ? isin. + specialize (all _ isin). + destruct x; [easy|]. + destruct e. + cbn in *. + now apply LevelSet.mem_spec. Qed. - Lemma conv_pb_rel_spec leq x y : + Lemma conv_pb_relb_complete leq u u' : + UnivExprSet.For_all + (fun e => LevelSet.mem (UnivExpr.get_level e) (global_ext_levels Σ)) (Universe.t_set u) -> + UnivExprSet.For_all + (fun e => LevelSet.mem (UnivExpr.get_level e) (global_ext_levels Σ)) (Universe.t_set u') -> + conv_pb_rel leq (global_ext_constraints Σ) u u' -> + conv_pb_relb leq u u'. + Proof. + intros all1%For_all_to_spec all2%For_all_to_spec conv. + destruct leq. + - eapply eq_universe_spec'; eauto. + + apply wf_ext_global_uctx_invariants, hΣ'. + + apply global_ext_uctx_consistent, hΣ'. + - eapply leq_universe_spec'; eauto. + + apply wf_ext_global_uctx_invariants, hΣ'. + + apply global_ext_uctx_consistent, hΣ'. + Qed. + + Lemma get_level_make l : + UnivExpr.get_level (UnivExpr.make l) = l. + Proof. now destruct l. Qed. + + Lemma conv_pb_relb_make_complete leq x y : LevelSet.mem x (global_ext_levels Σ) -> LevelSet.mem y (global_ext_levels Σ) -> conv_pb_rel leq (global_ext_constraints Σ) (Universe.make x) (Universe.make y) -> conv_pb_relb leq (Universe.make x) (Universe.make y). Proof. intros memx memy r. - destruct leq. - - eapply eq_universe_spec'; eauto. - + eapply wf_ext_global_uctx_invariants. - eapply hΣ'. - + eapply global_ext_uctx_consistent. - eapply hΣ'. - + now apply mem_level_declared_level. - + now apply mem_level_declared_level. - - eapply leq_universe_spec'; eauto. - + eapply wf_ext_global_uctx_invariants. - eapply hΣ'. - + eapply global_ext_uctx_consistent. - eapply hΣ'. - + now apply mem_level_declared_level. - + now apply mem_level_declared_level. + apply conv_pb_relb_complete; auto. + - intros ? ->%UnivExprSet.singleton_spec. + now rewrite get_level_make. + - intros ? ->%UnivExprSet.singleton_spec. + now rewrite get_level_make. Qed. - Lemma R_universe_instance_spec u u' : + Lemma eqb_universe_instance_complete u u' : Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) u -> Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) u' -> R_universe_instance (eq_universe (global_ext_constraints Σ)) u u' -> @@ -1730,11 +1743,11 @@ Section Conversion. cbn in *. apply Bool.andb_true_iff. split. - + apply (conv_pb_rel_spec Conv); auto. + + apply (conv_pb_relb_make_complete Conv); auto. + now apply IHu. Qed. - Lemma R_universe_variance_spec leq v u u' : + Lemma compare_universe_variance_complete leq v u u' : LevelSet.mem u (global_ext_levels Σ) -> LevelSet.mem u' (global_ext_levels Σ) -> R_universe_variance (eq_universe Σ) (conv_pb_rel leq Σ) v u u' -> @@ -1742,11 +1755,11 @@ Section Conversion. Proof. intros memu memu' r. destruct v; cbn in *; auto. - - apply conv_pb_rel_spec; auto. - - apply (conv_pb_rel_spec Conv); auto. + - apply conv_pb_relb_make_complete; auto. + - apply (conv_pb_relb_make_complete Conv); auto. Qed. - Lemma R_universe_instance_variance_spec leq v u u' : + Lemma compare_universe_instance_variance_complete leq v u u' : Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) u -> Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) u' -> R_universe_instance_variance (eq_universe Σ) (conv_pb_rel leq Σ) v u u' -> @@ -1763,11 +1776,11 @@ Section Conversion. apply Bool.andb_true_iff. destruct r. split. - + apply R_universe_variance_spec; auto. + + apply compare_universe_variance_complete; auto. + now apply IHu. Qed. - Lemma R_global_instance_spec u v leq gr napp : + Lemma compare_global_instance_complete u v leq gr napp : Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) u -> Forall (fun u => LevelSet.mem u (global_ext_levels Σ)) v-> R_global_instance Σ (eq_universe Σ) (conv_pb_rel leq Σ) gr napp u v -> @@ -1776,8 +1789,8 @@ Section Conversion. intros consu consv r. unfold compare_global_instance, R_global_instance, R_opt_variance in *. destruct global_variance. - - apply R_universe_instance_variance_spec; auto. - - apply R_universe_instance_spec; auto. + - apply compare_universe_instance_variance_complete; auto. + - apply eqb_universe_instance_complete; auto. Qed. Lemma consistent_instance_ext_all_mem udecl u : @@ -1934,7 +1947,7 @@ Section Conversion. eapply PCUICWeakeningEnv.declared_constant_inj in decl1; eauto; subst. apply consistent_instance_ext_all_mem in cons1. apply consistent_instance_ext_all_mem in cons2. - eapply R_universe_instance_spec in r; auto. + eapply eqb_universe_instance_complete in r; auto. Qed. (* TODO (RE)MOVE *) @@ -4448,12 +4461,14 @@ Section Conversion. eapply eqb_term_spec. auto. Qed. (* Beware !! There is a TODO in the code *) + Axiom todo_cofix_conversion : forall {A}, A. Next Obligation. apply h; clear h. - todo "Cofix". + exact todo_cofix_conversion. Qed. Next Obligation. todo "Cofix". + exact todo_cofix_conversion. Qed. (* Fallback *) @@ -5130,6 +5145,12 @@ Section Conversion. - reflexivity. Qed. + Axiom wellformed_sort_declared : + forall Σ Γ s, + wellformed Σ Γ (tSort s) -> + UnivExprSet.For_all + (fun e => LevelSet.mem (UnivExpr.get_level e) (global_ext_levels Σ)) (Universe.t_set s). + Lemma reducible_head_None Γ t π h : isApp t = false -> whnf RedFlags.nodelta Σ (Γ,,, stack_context π) (mkApps t (decompose_stack π).1) -> @@ -5729,7 +5750,7 @@ Section Conversion. apply inversion_Ind in typ2 as (?&?&?&?&?&?); auto. apply consistent_instance_ext_all_mem in c1. apply consistent_instance_ext_all_mem in c. - apply R_global_instance_spec in r; auto. + apply compare_global_instance_complete in r; auto. rewrite eq_inductive_refl in noteq. apply All2_length in rargs1. rewrite <- rargs1 in r. @@ -5749,7 +5770,7 @@ Section Conversion. apply inversion_Construct in typ2 as (?&?&?&?&?&?&?); auto. apply consistent_instance_ext_all_mem in c1. apply consistent_instance_ext_all_mem in c. - apply R_global_instance_spec in r; auto. + apply compare_global_instance_complete in r; auto. rewrite eq_inductive_refl, Nat.eqb_refl in noteq. apply All2_length in rargs1. rewrite <- rargs1 in r. @@ -5777,16 +5798,9 @@ Section Conversion. apply wellformed_context in h1; auto. clear aux. apply wellformed_context in h2; auto. - eapply conv_pb_rel_spec in H0. - assert (forall Σ Γ s, wellformed Σ Γ (tSort s) -> - Forall (fun s => LevelSet.In s (global_ext_levels Σ)) s). - destruct h1 as [(?&typ)|]. - + apply inversion_Sort in typ as (?&?&?&?&?); auto. - apply conv_pb_rel_spec in H0: - admit. - + destruct H as [(?&?&?&?)]. - cbn in *. - apply conv_pb_rel_spec in H0. + apply wellformed_sort_declared in h1. + apply wellformed_sort_declared in h2. + eapply conv_pb_relb_complete in H0; eauto. Qed. Equations _isconv (s : state) (Γ : context) From ce8fd2d0f4aa0f87a745a83e66e7647b983bc335 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Mon, 26 Oct 2020 23:02:57 +0100 Subject: [PATCH 17/26] Clean up and reorganize proof - Factorize some lemmas - Introduce ltac for simplifying stacks - Reorganize things into correct files --- pcuic/_CoqProject.in | 2 +- ...ionInversion.v => PCUICConvCumInversion.v} | 157 +- pcuic/theories/PCUICConversion.v | 218 +- pcuic/theories/PCUICCumulativity.v | 39 +- pcuic/theories/PCUICInductiveInversion.v | 111 +- pcuic/theories/PCUICNormal.v | 122 +- pcuic/theories/PCUICSafeLemmata.v | 196 +- safechecker/theories/PCUICSafeConversion.v | 2085 ++++------------- safechecker/theories/PCUICSafeReduce.v | 6 +- template-coq/theories/common/uGraph.v | 4 +- 10 files changed, 1048 insertions(+), 1892 deletions(-) rename pcuic/theories/{PCUICConversionInversion.v => PCUICConvCumInversion.v} (59%) diff --git a/pcuic/_CoqProject.in b/pcuic/_CoqProject.in index 18694e2bd..5d6529337 100644 --- a/pcuic/_CoqProject.in +++ b/pcuic/_CoqProject.in @@ -28,7 +28,7 @@ theories/PCUICConfluence.v theories/PCUICContextRelation.v theories/PCUICContextConversion.v theories/PCUICConversion.v -theories/PCUICConversionInversion.v +theories/PCUICConvCumInversion.v theories/PCUICGeneration.v theories/PCUICAlpha.v theories/PCUICCtxShape.v diff --git a/pcuic/theories/PCUICConversionInversion.v b/pcuic/theories/PCUICConvCumInversion.v similarity index 59% rename from pcuic/theories/PCUICConversionInversion.v rename to pcuic/theories/PCUICConvCumInversion.v index a71ea5002..e0f66a799 100644 --- a/pcuic/theories/PCUICConversionInversion.v +++ b/pcuic/theories/PCUICConvCumInversion.v @@ -20,18 +20,6 @@ Set Default Goal Selector "!". Section fixed. Context {cf : checker_flags}. Context (Σ : global_env_ext). - Context (wf : ∥wf_ext Σ∥). - - Definition conv_pb_rel (pb : conv_pb) := - match pb with - | Conv => eq_universe - | Cumul => leq_universe - end. - - Definition eq_termp_napp leq napp := - eq_term_upto_univ_napp Σ (eq_universe Σ) (conv_pb_rel leq Σ) napp. - - Notation eq_termp leq := (eq_termp_napp leq 0). Definition isIndConstructApp (t : term) : bool := match (decompose_app t).1 with @@ -62,46 +50,6 @@ Section fixed. - discriminate not_ind. - discriminate not_ind. Qed. - - Lemma whne_red1_isIndConstructApp Γ t t' : - whne RedFlags.default Σ Γ t -> - red1 Σ Γ t t' -> - isIndConstructApp t' = isIndConstructApp t. - Proof. - apply (whne_red1_ind - RedFlags.default Σ Γ - (fun t t' => isIndConstructApp t' = isIndConstructApp t)); intros; cbn in *; try easy. - - now rewrite !(isIndConstructApp_mkApps _ [v]). - - now rewrite (isIndConstructApp_mkApps _ [v]), (isIndConstructApp_mkApps _ [N2]). - - now rewrite !isIndConstructApp_mkApps. - - now rewrite !isIndConstructApp_mkApps. - - now rewrite !isIndConstructApp_mkApps. - Qed. - - Lemma whnf_red1_isIndConstructApp Γ t t' : - whnf RedFlags.default Σ Γ t -> - red1 Σ Γ t t' -> - isIndConstructApp t' = isIndConstructApp t. - Proof. - intros wh r. - depelim wh. - - eapply whne_red1_isIndConstructApp; eauto. - - depelim r; solve_discr. - - depelim r; auto; solve_discr. - - depelim r; auto; solve_discr. - - apply red1_mkApps_tConstruct_inv in r as (?&->&?). - now rewrite !isIndConstructApp_mkApps. - - apply red1_mkApps_tInd_inv in r as (?&->&?). - now rewrite !isIndConstructApp_mkApps. - - apply red1_mkApps_tFix_inv in r. - 2: { destruct unfold_fix as [(?&?)|]; [|easy]. - now unfold is_constructor; rewrite H. } - now destruct r as [[(?&->&o)|(?&->&o)]|(?&->&o)]; - rewrite !isIndConstructApp_mkApps. - - apply red1_mkApps_tCoFix_inv in r. - now destruct r as [[(?&->&o)|(?&->&o)]|(?&->&o)]; - rewrite !isIndConstructApp_mkApps. - Qed. Lemma whnf_red_isIndConstructApp Γ t t' : whnf RedFlags.default Σ Γ t -> @@ -109,56 +57,18 @@ Section fixed. isIndConstructApp t' = isIndConstructApp t. Proof. intros wh r. - induction r in wh |- * using red_rect_n1. - - easy. - - apply whnf_red1_isIndConstructApp in X as ->; eauto. - eapply whnf_pres; eauto. - Qed. - - Lemma conv_cum_alt leq Γ t t' : - conv_cum leq Σ Γ t t' <-> - ∥∑ v v', (red Σ Γ t v × red Σ Γ t' v') × eq_termp leq v v'∥. - Proof. - assert (forall P Q, (P <~> Q) -> (∥P∥ <-> ∥Q∥)) by - (intros P Q H; split; intros [p]; constructor; apply H in p; auto). - destruct leq; cbn; apply H; [apply conv_alt_red|apply cumul_alt]. - Qed. - - Lemma conv_terms_alt Γ args args' : - conv_terms Σ Γ args args' <~> - ∑ argsr argsr', - All2 (red Σ Γ) args argsr × - All2 (red Σ Γ) args' argsr' × - All2 (eq_term Σ Σ) argsr argsr'. - Proof. - split. - - intros conv. - induction conv. - + exists [], []; eauto with pcuic. - + apply conv_alt_red in r as (xr&yr&(xred&yred)&xy). - specialize IHconv as (argsr&argsr'&?&?&?). - exists (xr :: argsr), (yr :: argsr'). - eauto 7 with pcuic. - - intros (argsr&argsr'&r&r'&eqs). - induction eqs in args, args', r, r' |- *; depelim r; depelim r'; [constructor|]. - constructor; auto. - apply conv_alt_red; eauto. - Qed. - - Lemma eq_term_eq_termp leq x y : - eq_term Σ Σ x y -> - eq_termp leq x y. - Proof. - destruct leq; [easy|]. - cbn. - now apply PCUICCumulProp.eq_term_upto_univ_napp_leq. + eapply whnf_red_inv in r; eauto. + induction r; auto. + rewrite (isIndConstructApp_mkApps _ [arg']), (isIndConstructApp_mkApps _ [arg]). + apply IHr. + now apply whnf_mkApps_inv with (l := [arg]) in wh. Qed. Lemma eq_termp_mkApps_inv leq v args v' args' : isApp v = false -> isApp v' = false -> - eq_termp leq (mkApps v args) (mkApps v' args') -> - eq_termp_napp leq #|args| v v' × All2 (fun x y => eq_term Σ Σ x y) args args'. + eq_termp leq Σ (mkApps v args) (mkApps v' args') -> + eq_termp_napp leq Σ #|args| v v' × All2 (fun x y => eq_term Σ Σ x y) args args'. Proof. intros noapp1 noapp2 eq. now apply eq_term_upto_univ_mkApps_inv in eq as (?&?). @@ -167,7 +77,7 @@ Section fixed. Definition conv_cum_napp leq Γ napp t t' := match t with | tInd _ _ - | tConstruct _ _ _ => ∥eq_termp_napp leq napp t t'∥ + | tConstruct _ _ _ => ∥eq_termp_napp leq Σ napp t t'∥ | _ => conv_cum leq Σ Γ t t' end. @@ -187,21 +97,18 @@ Section fixed. { erewrite whnf_red_isApp. 3: eauto. 1: assumption. - apply whnf_mkApps_inv in wh; auto. - now rewrite notapp. } + apply whnf_mkApps_inv in wh; auto. } assert (isApp x = false). { erewrite whnf_red_isApp. 3: eauto. 1: assumption. - apply whnf_mkApps_inv in wh'; auto. - now rewrite notapp'. } + apply whnf_mkApps_inv in wh'; auto. } apply eq_termp_mkApps_inv in e as (?&?); auto. constructor. split. - assert (isIndConstructApp x1 = isIndConstructApp hd). { eapply whnf_red_isIndConstructApp; eauto. - apply whnf_mkApps_inv in wh; auto. - now rewrite notapp. } + apply whnf_mkApps_inv in wh; auto. } destruct hd. all: cbn. 1-9, 12-15: apply conv_cum_alt. @@ -213,7 +120,7 @@ Section fixed. all: apply whnf_mkApps_inv in wh; auto. all: eapply whnf_red_inv in r; auto. all: depelim r. - all: apply whnf_mkApps_inv in wh'; auto; [|now rewrite notapp']. + all: apply whnf_mkApps_inv in wh'; auto. all: eapply whnf_red_inv in r0; auto. all: depelim e. all: depelim r0. @@ -322,44 +229,4 @@ Section fixed. split; [easy|]. now apply conv_alt_red; exists c'0, c'1. Qed. - - (* TODO: Move to a place that actually should be depending on typing *) - Lemma conv_cum_red_inv leq Γ t1 t2 t1' t2' : - conv_cum leq Σ Γ t1 t2 -> - red Σ Γ t1 t1' -> - red Σ Γ t2 t2' -> - conv_cum leq Σ Γ t1' t2'. - Proof. - destruct wf. - intros cc r1 r2. - destruct leq; cbn in *. - - destruct cc. - constructor. - eapply PCUICConversion.conv_red_l_inv; [eauto| |eauto]. - apply conv_sym. - eapply PCUICConversion.conv_red_l_inv; [eauto| |eauto]. - apply conv_sym. - auto. - - destruct cc. - constructor. - eapply PCUICConversion.cumul_red_l_inv; [eauto| |eauto]. - eapply PCUICConversion.cumul_red_r_inv; [eauto| |eauto]. - auto. - Qed. - - Lemma conv_cum_red_inv' leq Γ Γ' t1 t2 t1' t2' : - conv_context Σ Γ Γ' -> - conv_cum leq Σ Γ t1 t2 -> - red Σ Γ t1 t1' -> - red Σ Γ' t2 t2' -> - conv_cum leq Σ Γ t1' t2'. - Proof. - intros cc r1 r2 conv_ctx. - destruct wf. - eapply conv_cum_red_inv; [|now eauto|reflexivity]. - eapply conv_cum_conv_ctx; eauto. - 2: apply conv_context_sym; eauto. - eapply conv_cum_red_inv; [|reflexivity|now eauto]. - eapply conv_cum_conv_ctx; eauto. - Qed. End fixed. diff --git a/pcuic/theories/PCUICConversion.v b/pcuic/theories/PCUICConversion.v index f8707fc3c..1e2883b92 100644 --- a/pcuic/theories/PCUICConversion.v +++ b/pcuic/theories/PCUICConversion.v @@ -337,7 +337,6 @@ Proof. repeat constructor. Qed. - Lemma conv_Prod_l_inv {cf:checker_flags} (Σ : global_env_ext) Γ na dom codom T : wf Σ -> Σ ;;; Γ |- tProd na dom codom = T -> @@ -1137,6 +1136,21 @@ Section Inversions. - eapply cumul_red_r ; try eassumption. econstructor. assumption. Qed. + + Lemma cumul_mkApps Γ hd args hd' args' : + Σ;;; Γ |- hd <= hd' -> + All2 (conv Σ Γ) args args' -> + Σ;;; Γ |- mkApps hd args <= mkApps hd' args'. + Proof. + intros cum cum_args. + revert hd hd' cum. + induction cum_args; intros hd hd' cum; auto. + cbn. + apply IHcum_args. + eapply cumul_trans; auto. + - eapply cumul_App_l; eauto. + - eapply cumul_App_r; eauto. + Qed. Lemma conv_App_r {Γ f x y} : Σ ;;; Γ |- x = y -> @@ -1298,7 +1312,112 @@ Section Inversions. apply ctx_rel_vass. 1: reflexivity. now constructor. Qed. + + Lemma conv_cum_conv_ctx leq Γ Γ' T U : + conv_cum leq Σ Γ T U -> + conv_context Σ Γ Γ' -> + conv_cum leq Σ Γ' T U. + Proof. + destruct leq; cbn; intros; sq. + - eapply conv_conv_ctx; eassumption. + - eapply cumul_conv_ctx; eassumption. + Qed. + + Lemma conv_cum_red leq Γ t1 t2 t1' t2' : + red Σ Γ t1' t1 -> + red Σ Γ t2' t2 -> + conv_cum leq Σ Γ t1 t2 -> + conv_cum leq Σ Γ t1' t2'. + Proof. + intros r1 r2 cc. + destruct leq; cbn in *. + - destruct cc. + constructor. + eapply red_conv_conv; eauto. + apply conv_sym. + eapply red_conv_conv; eauto. + apply conv_sym. + auto. + - destruct cc. + constructor. + eapply red_cumul_cumul; eauto. + eapply red_cumul_cumul_inv; eauto. + Qed. + + Lemma conv_cum_red_conv leq Γ Γ' t1 t2 t1' t2' : + conv_context Σ Γ Γ' -> + red Σ Γ t1' t1 -> + red Σ Γ' t2' t2 -> + conv_cum leq Σ Γ t1 t2 -> + conv_cum leq Σ Γ t1' t2'. + Proof. + intros conv_ctx r1 r2 cc. + eapply conv_cum_red; [now eauto|reflexivity|]. + eapply conv_cum_conv_ctx; eauto. + 2: apply conv_context_sym; eauto. + eapply conv_cum_red; [reflexivity|now eauto|]. + eapply conv_cum_conv_ctx; eauto. + Qed. + + Lemma conv_cum_red_inv leq Γ t1 t2 t1' t2' : + red Σ Γ t1 t1' -> + red Σ Γ t2 t2' -> + conv_cum leq Σ Γ t1 t2 -> + conv_cum leq Σ Γ t1' t2'. + Proof. + intros r1 r2 cc. + destruct leq; cbn in *. + - destruct cc. + constructor. + eapply conv_red_l_inv; [eauto| |eauto]. + apply conv_sym. + eapply conv_red_l_inv; [eauto| |eauto]. + apply conv_sym. + auto. + - destruct cc. + constructor. + eapply cumul_red_l_inv; [eauto| |eauto]. + eapply cumul_red_r_inv; [eauto| |eauto]. + auto. + Qed. + + Lemma conv_cum_red_conv_inv leq Γ Γ' t1 t2 t1' t2' : + conv_context Σ Γ Γ' -> + red Σ Γ t1 t1' -> + red Σ Γ' t2 t2' -> + conv_cum leq Σ Γ t1 t2 -> + conv_cum leq Σ Γ t1' t2'. + Proof. + intros conv_ctx r1 r2 cc. + eapply conv_cum_red_inv; [now eauto|reflexivity|]. + eapply conv_cum_conv_ctx; eauto. + 2: apply conv_context_sym; eauto. + eapply conv_cum_red_inv; [reflexivity|now eauto|]. + eapply conv_cum_conv_ctx; eauto. + Qed. + + Lemma conv_cum_red_iff leq Γ t1 t2 t1' t2' : + red Σ Γ t1' t1 -> + red Σ Γ t2' t2 -> + conv_cum leq Σ Γ t1 t2 <-> conv_cum leq Σ Γ t1' t2'. + Proof. + intros r1 r2. + split; intros cc. + - eapply conv_cum_red; eauto. + - eapply conv_cum_red_inv; eauto. + Qed. + Lemma conv_cum_red_conv_iff leq Γ Γ' t1 t2 t1' t2' : + conv_context Σ Γ Γ' -> + red Σ Γ t1' t1 -> + red Σ Γ' t2' t2 -> + conv_cum leq Σ Γ t1 t2 <-> conv_cum leq Σ Γ t1' t2'. + Proof. + intros conv_ctx r1 r2. + split; intros cc. + - eapply conv_cum_red_conv; eauto. + - eapply conv_cum_red_conv_inv; eauto. + Qed. Lemma cumul_Case_c : forall Γ indn p brs u v, @@ -1868,17 +1987,6 @@ Section Inversions. * eapply conv_cumul, conv_LetIn_ty with (na := na1); tea. Qed. - Lemma conv_cum_conv_ctx leq Γ Γ' T U : - conv_cum leq Σ Γ T U -> - conv_context Σ Γ Γ' -> - conv_cum leq Σ Γ' T U. - Proof. - destruct leq; cbn; intros; sq. - - eapply conv_conv_ctx; eassumption. - - eapply cumul_conv_ctx; eassumption. - Qed. - - Lemma it_mkLambda_or_LetIn_conv_cum leq Γ Δ1 Δ2 t1 t2 : conv_context Σ (Γ ,,, Δ1) (Γ ,,, Δ2) -> conv_cum leq Σ (Γ ,,, Δ1) t1 t2 -> @@ -2387,6 +2495,92 @@ Notation conv_terms Σ Γ := (All2 (conv Σ Γ)). Instance conv_terms_Proper {cf:checker_flags} Σ Γ : CMorphisms.Proper (eq ==> eq ==> arrow)%signature (conv_terms Σ Γ). Proof. intros x y -> x' y' -> f. exact f. Qed. +Lemma conv_terms_alt {cf:checker_flags} Σ Γ args args' : + conv_terms Σ Γ args args' <~> + ∑ argsr argsr', + All2 (red Σ Γ) args argsr × + All2 (red Σ Γ) args' argsr' × + All2 (eq_term Σ Σ) argsr argsr'. +Proof. + split. + - intros conv. + induction conv. + + exists [], []; eauto with pcuic. + + apply conv_alt_red in r as (xr&yr&(xred&yred)&xy). + specialize IHconv as (argsr&argsr'&?&?&?). + exists (xr :: argsr), (yr :: argsr'). + eauto 7 with pcuic. + - intros (argsr&argsr'&r&r'&eqs). + induction eqs in args, args', r, r' |- *; depelim r; depelim r'; [constructor|]. + constructor; auto. + apply conv_alt_red; eauto. +Qed. + +Lemma conv_terms_conv_ctx {cf:checker_flags} (Σ : global_env_ext) Γ Γ' ts ts' : + wf Σ -> + conv_context Σ Γ Γ' -> + conv_terms Σ Γ ts ts' -> + conv_terms Σ Γ' ts ts'. +Proof. + intros wf ctx conv. + induction conv; [constructor|]. + constructor; auto. + eapply conv_conv_ctx; eauto. +Qed. + +Lemma conv_terms_red {cf:checker_flags} (Σ : global_env_ext) Γ ts ts' tsr tsr' : + All2 (red Σ Γ) ts tsr -> + All2 (red Σ Γ) ts' tsr' -> + conv_terms Σ Γ tsr tsr' -> + conv_terms Σ Γ ts ts'. +Proof. + intros all all' conv. + induction conv in ts, ts', all, all' |- *; depelim all; depelim all'; [constructor|]. + constructor; [|auto]. + eapply red_conv_conv; eauto. + symmetry. + eapply red_conv_conv; eauto. + symmetry. + eauto. +Qed. + +Lemma conv_terms_red_inv {cf:checker_flags} (Σ : global_env_ext) Γ ts ts' tsr tsr' : + wf Σ -> + All2 (red Σ Γ) ts tsr -> + All2 (red Σ Γ) ts' tsr' -> + conv_terms Σ Γ ts ts' -> + conv_terms Σ Γ tsr tsr'. +Proof. + intros wf all all' conv. + induction conv in tsr, tsr', all, all' |- *; depelim all; depelim all'; [constructor|]. + constructor; [|auto]. + eapply conv_red_l_inv; [eauto| |eauto]. + symmetry. + eapply conv_red_l_inv; [eauto| |eauto]. + symmetry. + eauto. +Qed. + +Lemma conv_terms_red_conv {cf:checker_flags} (Σ : global_env_ext) Γ Γ' ts ts' tsr tsr' : + wf Σ -> + conv_context Σ Γ Γ' -> + All2 (red Σ Γ) ts tsr -> + All2 (red Σ Γ') ts' tsr' -> + conv_terms Σ Γ tsr tsr' -> + conv_terms Σ Γ ts ts'. +Proof. + intros wf convctx all all2 conv. + eapply conv_terms_red. + 1: eassumption. + 1: apply All2_same; reflexivity. + eapply conv_terms_conv_ctx; eauto. + 1: eapply conv_context_sym; eauto. + eapply conv_terms_red. + 1: apply All2_same; reflexivity. + 1: eauto. + eapply conv_terms_conv_ctx; eauto. +Qed. + Lemma cumul_subst_conv {cf:checker_flags} (Σ : global_env_ext) Γ Δ Δ' Γ' s s' b : wf Σ -> All2 (conv Σ Γ) s s' -> untyped_subslet Γ s Δ -> diff --git a/pcuic/theories/PCUICCumulativity.v b/pcuic/theories/PCUICCumulativity.v index 743fa4b23..30de075ac 100644 --- a/pcuic/theories/PCUICCumulativity.v +++ b/pcuic/theories/PCUICCumulativity.v @@ -249,7 +249,6 @@ Proof. econstructor 2; eauto. Qed. -(* This should be called sym instead of inv, no ? *) Lemma red_conv_conv_inv `{cf : checker_flags} Σ Γ t u v : red (fst Σ) Γ t u -> Σ ;;; Γ |- v = u -> Σ ;;; Γ |- v = t. Proof. @@ -291,13 +290,43 @@ Inductive conv_pb := | Conv | Cumul. -Definition conv_cum `{cf : checker_flags} leq Σ Γ u v := +Definition conv_cum {cf:checker_flags} leq Σ Γ u v := match leq with | Conv => ∥ Σ ;;; Γ |- u = v ∥ | Cumul => ∥ Σ ;;; Γ |- u <= v ∥ end. -Lemma conv_conv_cum_l `{cf : checker_flags} : +Definition conv_pb_rel {cf:checker_flags} (pb : conv_pb) := + match pb with + | Conv => eq_universe + | Cumul => leq_universe + end. + +Definition eq_termp_napp {cf:checker_flags} leq (Σ : global_env_ext) napp := + eq_term_upto_univ_napp Σ (eq_universe Σ) (conv_pb_rel leq Σ) napp. + +Definition eq_termp {cf:checker_flags} leq Σ := (eq_termp_napp leq Σ 0). + +Lemma eq_term_eq_termp {cf:checker_flags} leq (Σ : global_env_ext) x y : + eq_term Σ Σ x y -> + eq_termp leq Σ x y. +Proof. + destruct leq; [easy|]. + cbn. + apply eq_term_upto_univ_leq; auto. + typeclasses eauto. +Qed. + +Lemma conv_cum_alt {cf:checker_flags} leq Σ Γ t t' : + conv_cum leq Σ Γ t t' <-> + ∥∑ v v', (red Σ Γ t v × red Σ Γ t' v') × eq_termp leq Σ v v'∥. +Proof. + assert (forall P Q, (P <~> Q) -> (∥P∥ <-> ∥Q∥)) by + (intros P Q H; split; intros [p]; constructor; apply H in p; auto). + destruct leq; cbn; apply H; [apply conv_alt_red|apply cumul_alt]. +Qed. + +Lemma conv_conv_cum_l {cf:checker_flags} : forall (Σ : global_env_ext) leq Γ u v, Σ ;;; Γ |- u = v -> conv_cum leq Σ Γ u v. @@ -307,7 +336,7 @@ Proof. - cbn. constructor. now apply conv_cumul. Qed. -Lemma conv_conv_cum_r `{cf : checker_flags} : +Lemma conv_conv_cum_r {cf:checker_flags} : forall (Σ : global_env_ext) leq Γ u v, Σ ;;; Γ |- u = v -> conv_cum leq Σ Γ v u. @@ -318,7 +347,7 @@ Proof. now apply conv_sym. Qed. -Lemma cumul_App_l `{cf : checker_flags} : +Lemma cumul_App_l {cf:checker_flags} : forall {Σ Γ f g x}, Σ ;;; Γ |- f <= g -> Σ ;;; Γ |- tApp f x <= tApp g x. diff --git a/pcuic/theories/PCUICInductiveInversion.v b/pcuic/theories/PCUICInductiveInversion.v index 5f6cbab45..5d0652118 100644 --- a/pcuic/theories/PCUICInductiveInversion.v +++ b/pcuic/theories/PCUICInductiveInversion.v @@ -1634,31 +1634,50 @@ Proof. constructor. constructor. Qed. -Lemma Case_Construct_ind_eq {cf:checker_flags} Σ (hΣ : ∥ wf Σ.1 ∥) - {Γ ind ind' npar pred i u brs args} : - (∑ T, Σ ;;; Γ |- tCase (ind, npar) pred (mkApps (tConstruct ind' i u) args) brs : T) -> +Lemma declared_inductive_unique {Σ mdecl idecl p} (q r : declared_inductive Σ mdecl p idecl) : q = r. +Proof. + unfold declared_inductive in q, r. + destruct q, r. + now rewrite (uip e e0) (uip d d0). +Qed. + +Lemma declared_inductive_unique_sig {cf:checker_flags} {Σ ind mib decl mib' decl'} + (decl1 : declared_inductive Σ mib ind decl) + (decl2 : declared_inductive Σ mib' ind decl') : + @sigmaI _ (fun '(m, d) => declared_inductive Σ m ind d) + (mib, decl) decl1 = + @sigmaI _ _ (mib', decl') decl2. +Proof. + pose proof (declared_inductive_inj decl1 decl2) as (<-&<-). + pose proof (declared_inductive_unique decl1 decl2) as ->. + reflexivity. +Qed. + +Lemma invert_Case_Construct {cf:checker_flags} Σ (hΣ : ∥ wf Σ.1 ∥) + {Γ ind ind' npar pred i u brs args T} : + Σ ;;; Γ |- tCase (ind, npar) pred (mkApps (tConstruct ind' i u) args) brs : T -> ind = ind'. Proof. destruct hΣ as [wΣ]. - intros [A h]. + intros h. apply inversion_Case in h as ih ; auto. destruct ih as [uni [args' [mdecl [idecl [pty [indctx [pctx [ps [btys [? [? [? [? [ht0 [? ?]]]]]]]]]]]]]]]. - pose proof ht0 as typec. - eapply inversion_mkApps in typec as [A' [tyc tyargs]]; auto. - eapply (inversion_Construct Σ wΣ) in tyc as [mdecl' [idecl' [cdecl' [wfl [declc [Hu tyc]]]]]]. - epose proof (PCUICInductiveInversion.Construct_Ind_ind_eq _ ht0 declc); eauto. - destruct on_declared_constructor as [[onmind oib] [cs [? ?]]]. - simpl in *. - intuition auto. + pose proof ht0 as typec. + eapply inversion_mkApps in typec as [A' [tyc tyargs]]; auto. + eapply (inversion_Construct Σ wΣ) in tyc as [mdecl' [idecl' [cdecl' [wfl [declc [Hu tyc]]]]]]. + epose proof (PCUICInductiveInversion.Construct_Ind_ind_eq _ ht0 declc); eauto. + destruct on_declared_constructor as [[onmind oib] [cs [? ?]]]. + simpl in *. + intuition auto. Qed. -Lemma Proj_Constuct_ind_eq {cf:checker_flags} Σ (hΣ : ∥ wf Σ.1 ∥) {Γ i i' pars narg c u l} : - (∑ T, Σ ;;; Γ |- tProj (i, pars, narg) (mkApps (tConstruct i' c u) l) : T) -> +Lemma Proj_Construct_ind_eq {cf:checker_flags} Σ (hΣ : ∥ wf Σ.1 ∥) {Γ i i' pars narg c u l T} : + Σ ;;; Γ |- tProj (i, pars, narg) (mkApps (tConstruct i' c u) l) : T -> i = i'. Proof. destruct hΣ as [wΣ]. - intros [T h]. + intros h. apply inversion_Proj in h ; auto. destruct h as [uni [mdecl [idecl [pdecl [args' [? [hc [? ?]]]]]]]]. pose proof hc as typec. @@ -1670,19 +1689,22 @@ Proof. intuition auto. Qed. -Lemma Proj_Constuct_projargs {cf:checker_flags} Σ (hΣ : ∥ wf Σ.1 ∥) {Γ i pars narg c u l} : - (∑ T, Σ ;;; Γ |- tProj (i, pars, narg) (mkApps (tConstruct i c u) l) : T) -> - pars + narg < #|l|. +Lemma invert_Proj_Construct {cf:checker_flags} Σ (hΣ : ∥ wf Σ.1 ∥) {Γ i i' pars narg c u l T} : + Σ ;;; Γ |- tProj (i, pars, narg) (mkApps (tConstruct i' c u) l) : T -> + i = i' /\ c = 0 /\ pars + narg < #|l|. Proof. + intros h. + assert (h' := h). + apply Proj_Construct_ind_eq in h' as <-; auto. destruct hΣ as [wΣ]. - intros [T h]. + split; [reflexivity|]. apply inversion_Proj in h ; auto. destruct h as [uni [mdecl [idecl [pdecl [args' [? [hc [? ?]]]]]]]]. clear c0. pose proof hc as typec. eapply inversion_mkApps in typec as [A' [tyc tyargs]]; auto. eapply (inversion_Construct Σ wΣ) in tyc as [mdecl' [idecl' [cdecl' [wfl [declc [Hu tyc]]]]]]. - pose proof (declared_inductive_inj d.p1 declc.p1) as [? ?]; subst mdecl' idecl'. + pose proof (declared_inductive_unique_sig d.p1 declc.p1) as H; noconf H. set (declc' := (conj (let (x, _) := d in x) declc.p2) : declared_constructor Σ.1 mdecl idecl (i, c) cdecl'). epose proof (PCUICInductiveInversion.Construct_Ind_ind_eq _ hc declc'); eauto. @@ -1703,57 +1725,6 @@ Proof. destruct y as [[_ onps] onp]. lia. Qed. -Lemma declared_inductive_unique {Σ mdecl idecl p} (q r : declared_inductive Σ mdecl p idecl) : q = r. -Proof. - unfold declared_inductive in q, r. - destruct q, r. - now rewrite (uip e e0) (uip d d0). -Qed. - -Lemma declared_inductive_unique_sig {cf:checker_flags} {Σ ind mib decl mib' decl'} - (decl1 : declared_inductive Σ mib ind decl) - (decl2 : declared_inductive Σ mib' ind decl') : - @sigmaI _ (fun '(m, d) => declared_inductive Σ m ind d) - (mib, decl) decl1 = - @sigmaI _ _ (mib', decl') decl2. -Proof. - pose proof (declared_inductive_inj decl1 decl2) as (<-&<-). - pose proof (declared_inductive_unique decl1 decl2) as ->. - reflexivity. -Qed. - -Lemma projected_constructor_eq - {cf:checker_flags} Σ (hΣ : ∥wf Σ.1∥) Γ ind c u args ind' u' args' mib oib npars idx proj : - Σ;;; Γ |- mkApps (tConstruct ind c u) args : mkApps (tInd ind' u') args' -> - declared_projection Σ.1 mib oib (ind', npars, idx) proj -> - c = 0. -Proof. - intros typ decl. - destruct hΣ as [wfΣ]. - pose proof (inversion_mkApps wfΣ typ) as (?&ctor_typ&_); auto. - apply inversion_Construct in ctor_typ as (?&?&?&?&?&?&?); auto. - unshelve eapply Construct_Ind_ind_eq in typ. - 5: eassumption. - 1: now auto. - cbn in typ. - destruct d; cbn in typ. - destruct All2_nth_error_Some; cbn in typ. - destruct p; cbn in typ. - destruct typ as ((((<-&?)&?)&?)&?). - pose proof (on_declared_projection wfΣ decl) as (_&proj'). - clear -e0 proj'. - cbn in *. - destruct decl as (decl_oib&?&?). - pose proof (declared_inductive_unique_sig decl_oib d) as H'. - noconf H'. - destruct ind_cshapes; [easy|]. - destruct l; [|easy]. - destruct c; [easy|]. - cbn in *. - now rewrite nth_error_nil in e0. -Qed. - - Ltac unf_env := change PCUICEnvironment.it_mkProd_or_LetIn with it_mkProd_or_LetIn in *; change PCUICEnvironment.to_extended_list_k with to_extended_list_k in *; diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index 783528f87..3e74ee0e0 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -199,40 +199,51 @@ Hint Resolve negb_is_true : core. Lemma whnf_mkApps_inv flags : forall Σ Γ t l, - ~ isApp t -> whnf flags Σ Γ (mkApps t l) -> whnf flags Σ Γ t. Proof. - intros Σ Γ t l Hr h. - induction l using rev_ind. - - assumption. - - rewrite <- mkApps_nested in h. cbn in h. depelim h. depelim H. eauto. - + change (tApp (mkApps t l) x0) with (mkApps (mkApps t l) [x0]) in *. - rewrite mkApps_nested in x. - eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; cbn in *; try congruence. eauto. inv x. - eapply whnf_fixapp with (v := []); eauto. rewrite H. now destruct rarg. - + change (tApp (mkApps t l) x0) with (mkApps (mkApps t l) [x0]) in *. - rewrite mkApps_nested in x. - eapply (f_equal decompose_app) in x. - rewrite !decompose_app_mkApps in x; cbn in *; try congruence. firstorder. inv x. - eapply whnf_cstrapp with (v := []). - + change (tApp (mkApps t l) x0) with (mkApps (mkApps t l) [x0]) in *. - rewrite mkApps_nested in x. - eapply (f_equal decompose_app) in x. - rewrite !decompose_app_mkApps in x; cbn in *; try congruence. firstorder. inv x. - eapply whnf_indapp with (v := []). - + change (tApp (mkApps t l) x0) with (mkApps (mkApps t l) [x0]) in *. - rewrite mkApps_nested in x. - eapply (f_equal decompose_app) in x. - rewrite !decompose_app_mkApps in x; cbn in *; try congruence. firstorder. inv x. - eapply whnf_fixapp with (v := []). - destruct unfold_fix as [[rarg arg] | ]; eauto. now destruct rarg. - + change (tApp (mkApps t l) x0) with (mkApps (mkApps t l) [x0]) in *. - rewrite mkApps_nested in x. - eapply (f_equal decompose_app) in x. - rewrite !decompose_app_mkApps in x; cbn in *; try congruence. firstorder. inv x. - eauto. + intros Σ Γ t l h. + remember (mkApps t l) eqn:hdeq. + revert hdeq. + induction h in h, t, l, t0 |- *; intros eq; subst. + - destruct (mkApps_elim t l). + apply whne_mkApps_inv in H; auto. + destruct H as [wh|(?&?&?&?&?&->&?&?&?)]. + + now apply whnf_mkApps. + + destruct (Nat.leb_spec n x1). + * apply whnf_fixapp. + rewrite H. + apply nth_error_Some_length in H0. + apply nth_error_None. + rewrite firstn_length. + lia. + * apply whnf_ne. + eapply whne_fixapp; eauto. + assert (x1 < #|l|) by (now apply nth_error_Some_length in H0). + rewrite <- (firstn_skipn n l) in H0. + rewrite nth_error_app1 in H0; auto. + rewrite firstn_length. + lia. + - destruct l using List.rev_ind; [|now rewrite <- mkApps_nested in eq]. + cbn in *; subst; auto. + - destruct l using List.rev_ind; [|now rewrite <- mkApps_nested in eq]. + cbn in *; subst; auto. + - destruct l using List.rev_ind; [|now rewrite <- mkApps_nested in eq]. + cbn in *; subst; auto. + - destruct (mkApps_elim t l). + apply mkApps_eq_inj in eq as (<-&<-); auto. + - destruct (mkApps_elim t l). + apply mkApps_eq_inj in eq as (<-&<-); auto. + - destruct (mkApps_elim t l). + apply mkApps_eq_inj in eq as (<-&<-); auto. + apply whnf_fixapp. + destruct unfold_fix as [(?&?)|]; [|easy]. + apply nth_error_None. + apply nth_error_None in H. + rewrite firstn_length. + lia. + - destruct (mkApps_elim t l). + apply mkApps_eq_inj in eq as (<-&<-); auto. Qed. Lemma whnf_fixapp' {flags} Σ Γ mfix idx narg body v : @@ -244,6 +255,53 @@ Proof. Qed. Hint Resolve whnf_fixapp' : core. +Lemma whnf_whne_nodelta_upgrade Σ Γ t : + whnf RedFlags.default Σ Γ t -> + whne RedFlags.nodelta Σ Γ t -> + whne RedFlags.default Σ Γ t. +Proof. + intros whn whe. + induction whe; cbn in *; try easy. + - now inv whn; solve_discr. + - inv whn. + + easy. + + destruct v0 as [|? ? _] using List.rev_ind; [discriminate|]. + rewrite <- mkApps_nested in *. + cbn in *; inv H0. + constructor. + eapply IHwhe. + apply whnf_cstrapp. + + destruct v0 as [|? ? _] using List.rev_ind; [discriminate|]. + rewrite <- mkApps_nested in *. + cbn in *; inv H0. + constructor. + eapply IHwhe. + apply whnf_indapp. + + destruct v0 as [|? ? _] using List.rev_ind; [discriminate|]. + rewrite <- mkApps_nested in *. + cbn in *; inv H. + constructor. + eapply IHwhe. + apply whnf_fixapp. + destruct unfold_fix; [|easy]. + destruct p. + apply nth_error_None. + apply nth_error_None in H0. + rewrite app_length in H0; cbn in *. + lia. + + destruct v0 as [|? ? _] using List.rev_ind; [discriminate|]. + rewrite <- mkApps_nested in *. + cbn in *; inv H0. + constructor. + eapply IHwhe. + apply whnf_cofixapp. + - inv whn; solve_discr; try easy. + rewrite H in H2. + congruence. + - inv whn; solve_discr; easy. + - inv whn; solve_discr; easy. +Qed. + Definition whnf_whne_dec flags Σ Γ t : ({whnf flags Σ Γ t} + {~ (whnf flags Σ Γ t)}) * ({whne flags Σ Γ t} + {~ (whne flags Σ Γ t)}). Proof. induction t using term_forall_mkApps_ind in Γ |- *; split; eauto. @@ -1262,7 +1320,7 @@ Lemma whnf_red_mkApps_inv Σ Γ hd args hd' args' : Proof. intros wh r notapp notapp'. apply whnf_red_mkApps_l in r as [(?&?&eq&?&?)]; auto. - apply whnf_mkApps_inv in wh; [|now destruct (isApp hd)]. + apply whnf_mkApps_inv in wh. eapply whnf_red_isApp in wh as ?; eauto. apply mkApps_notApp_inj in eq as (->&->); auto; [|congruence]. easy. diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index dc16ad281..6985179e7 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -1164,7 +1164,7 @@ Section Lemmata. destruct ρ. all: auto. exfalso. eapply decompose_stack_not_app. eassumption. Qed. - + (* TODO MOVE *) Lemma stack_context_decompose : forall π, @@ -1175,7 +1175,72 @@ Section Lemmata. cbn. pose proof (decompose_stack_eq _ _ _ e). subst. rewrite stack_context_appstack. reflexivity. Qed. + + Lemma decompose_stack_stack_cat π π' : + decompose_stack (π +++ π') = + ((decompose_stack π).1 ++ + match (decompose_stack π).2 with + | ε => (decompose_stack π').1 + | _ => [] + end, + (decompose_stack π).2 +++ + match (decompose_stack π).2 with + | ε => (decompose_stack π').2 + | _ => π' + end). + Proof. + induction π in π' |- *; cbn in *; auto. + - now destruct decompose_stack. + - rewrite !IHπ. + now destruct (decompose_stack π). + Qed. + + Lemma zipp_stack_cat π π' t : + isStackApp π = false -> + zipp t (π' +++ π) = zipp t π'. + Proof. + intros no_stack_app. + unfold zipp. + rewrite decompose_stack_stack_cat. + destruct (decompose_stack π') eqn:decomp. + cbn. + destruct s; try now rewrite app_nil_r. + now destruct π; cbn in *; rewrite ?app_nil_r. + Qed. + + Lemma zipp_appstack t args π : + zipp t (appstack args π) = zipp (mkApps t args) π. + Proof. + unfold zipp. + rewrite decompose_stack_appstack. + rewrite <- mkApps_nested. + now destruct decompose_stack. + Qed. + Lemma appstack_cons a args π : + appstack (a :: args) π = App a (appstack args π). + Proof. reflexivity. Qed. + + Lemma fst_decompose_stack_nil π : + isStackApp π = false -> + (decompose_stack π).1 = []. + Proof. now destruct π. Qed. + + Lemma zipp_as_mkApps t π : + zipp t π = mkApps t (decompose_stack π).1. + Proof. + unfold zipp. + now destruct decompose_stack. + Qed. + + Lemma zipp_noStackApp t π : + isStackApp π = false -> + zipp t π = t. + Proof. + intros. + now rewrite zipp_as_mkApps fst_decompose_stack_nil. + Qed. + Lemma it_mkLambda_or_LetIn_inj : forall Γ u v, it_mkLambda_or_LetIn Γ u = @@ -1342,11 +1407,8 @@ Section Lemmata. Proof. intros Γ i pars narg i' c u l [[T h]|[[ctx [s [e _]]]]]; [|discriminate]. - epose proof (PCUICInductiveInversion.Proj_Constuct_ind_eq _ hΣ). - forward H by (exists T; eauto). subst i'. - epose proof (PCUICInductiveInversion.Proj_Constuct_projargs _ hΣ). - forward H by (exists T; eauto). - now apply (nth_error_Some). + apply PCUICInductiveInversion.invert_Proj_Construct in h as (<-&->&?); auto. + now apply nth_error_Some. Qed. Lemma cored_zipc : @@ -1382,44 +1444,94 @@ Section Lemmata. rewrite stack_context_appstack. assumption. Qed. + + Lemma conv_cum_zipp leq Γ t t' π π' : + conv_cum leq Σ Γ t t' -> + ∥conv_terms Σ Γ (decompose_stack π).1 (decompose_stack π').1∥ -> + conv_cum leq Σ Γ (zipp t π) (zipp t' π'). + Proof. + intros conv conv_args. + rewrite !zipp_as_mkApps. + destruct leq; cbn in *. + - sq. + apply mkApps_conv_args; auto. + - sq. + apply cumul_mkApps; auto. + Qed. + + Lemma whne_context_relation f rel Γ Γ' t : + (forall Γ Γ' c c', rel Γ Γ' c c' -> (decl_body c = None <-> decl_body c' = None)) -> + whne f Σ Γ t -> + context_relation rel Γ Γ' -> + whne f Σ Γ' t. + Proof. + intros behaves wh conv. + induction wh; eauto using whne. + destruct nth_error eqn:nth; [|easy]. + cbn in H. + eapply context_relation_nth in nth; eauto. + destruct nth as (?&?&?&?). + constructor. + rewrite e. + cbn. + specialize (behaves _ _ _ _ r). + f_equal. + apply behaves. + congruence. + Qed. + + Lemma whnf_context_relation f rel Γ Γ' t : + (forall Γ Γ' c c', rel Γ Γ' c c' -> (decl_body c = None <-> decl_body c' = None)) -> + whnf f Σ Γ t -> + context_relation rel Γ Γ' -> + whnf f Σ Γ' t. + Proof. + intros behaves wh conv. + destruct wh; eauto using whnf. + apply whnf_ne. + eapply whne_context_relation; eauto. + Qed. + + Lemma whne_conv_context f Γ Γ' t : + whne f Σ Γ t -> + conv_context Σ Γ Γ' -> + whne f Σ Γ' t. + Proof. + apply whne_context_relation. + intros ? ? ? ? r. + now depelim r. + Qed. + + Lemma whnf_conv_context f Γ Γ' t : + whnf f Σ Γ t -> + conv_context Σ Γ Γ' -> + whnf f Σ Γ' t. + Proof. + apply whnf_context_relation. + intros ? ? ? ? r. + now depelim r. + Qed. + + Lemma Case_Construct_ind_eq : + forall {Γ ind ind' npar pred i u brs args}, + wellformed Σ Γ (tCase (ind, npar) pred (mkApps (tConstruct ind' i u) args) brs) -> + ind = ind'. + Proof. + destruct hΣ as [wΣ]. + intros Γ ind ind' npar pred i u brs args [[A h]|[[ctx [s [e _]]]]]; + [|discriminate]. + apply PCUICInductiveInversion.invert_Case_Construct in h; auto. + Qed. - Lemma conv_cum_context_convp : - forall Γ Γ' leq u v, - conv_cum leq Σ Γ u v -> - conv_context Σ Γ Γ' -> - conv_cum leq Σ Γ' u v. + Lemma Proj_Construct_ind_eq : + forall Γ i i' pars narg c u l, + wellformed Σ Γ (tProj (i, pars, narg) (mkApps (tConstruct i' c u) l)) -> + i = i'. Proof. - intros Γ Γ' leq u v h hx. - destruct hΣ. - destruct leq. - - simpl. destruct h. constructor. - eapply conv_conv_ctx. all: eauto. - - simpl in *. destruct h. constructor. - eapply cumul_conv_ctx. all: eauto. + destruct hΣ as [wΣ]. + intros Γ i i' pars narg c u l [[T h]|[[ctx [s [e _]]]]]; + [|discriminate]. + now apply PCUICInductiveInversion.invert_Proj_Construct in h. Qed. End Lemmata. - -Lemma Case_Construct_ind_eq {cf:checker_flags} Σ (hΣ : ∥ wf Σ.1 ∥) : - forall {Γ ind ind' npar pred i u brs args}, - wellformed Σ Γ (tCase (ind, npar) pred (mkApps (tConstruct ind' i u) args) brs) -> - ind = ind'. -Proof. -destruct hΣ as [wΣ]. -intros Γ ind ind' npar pred i u brs args [[A h]|[[ctx [s [e _]]]]]; - [|discriminate]. - eapply PCUICInductiveInversion.Case_Construct_ind_eq; eauto. - sq; auto. -Qed. - -Lemma Proj_Constuct_ind_eq {cf:checker_flags} Σ (hΣ : ∥ wf Σ.1 ∥): -forall Γ i i' pars narg c u l, - wellformed Σ Γ (tProj (i, pars, narg) (mkApps (tConstruct i' c u) l)) -> - i = i'. -Proof. - destruct hΣ as [wΣ]. - intros Γ i i' pars narg c u l [[T h]|[[ctx [s [e _]]]]]; - [|discriminate]. - eapply PCUICInductiveInversion.Proj_Constuct_ind_eq; eauto. - sq; auto. -Qed. diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 5c2bb8db5..c2c7c6375 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -6,7 +6,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils PCUICCumulativity PCUICEquality PCUICConversion PCUICSafeLemmata PCUICNormal PCUICInversion PCUICReduction PCUICPosition PCUICPrincipality PCUICContextConversion PCUICSN PCUICUtils PCUICWeakening - PCUICConversionInversion. + PCUICConvCumInversion. From MetaCoq.SafeChecker Require Import PCUICSafeReduce. Require Import Equations.Prop.DepElim. @@ -530,7 +530,7 @@ Section Conversion. Lemma eqb_termp_napp_spec pb napp t u : eqb_termp_napp pb napp t u -> - eq_termp_napp Σ pb napp t u. + eq_termp_napp pb Σ napp t u. Proof. pose proof hΣ'. apply eqb_term_upto_univ_impl. @@ -762,48 +762,133 @@ Section Conversion. destruct H0 as [|(?&?&?&?&?&?&?&?)]; [|discriminate]. depelim H0; solve_discr; discriminate. Qed. - - Lemma decompose_stack_stack_cat π π' : - decompose_stack (π +++ π') = - ((decompose_stack π).1 ++ - match (decompose_stack π).2 with - | ε => (decompose_stack π').1 - | _ => [] - end, - (decompose_stack π).2 +++ - match (decompose_stack π).2 with - | ε => (decompose_stack π').2 - | _ => π' - end). + + Lemma eta_pair {A B} (p q : A * B) : + p = q -> + (p.1, p.2) = (q.1, q.2). + Proof. now destruct p, q. Qed. + + Ltac is_var t := + match goal with + | v : _ |- _ => + match t with + | v => idtac + end + end. + + Lemma zipp_stack_cat_decompose_stack t π π' : + zipp t (π +++ (decompose_stack π').2) = zipp t π. Proof. - induction π in π' |- *; cbn in *; auto. - - now destruct decompose_stack. - - rewrite !IHπ. - now destruct (decompose_stack π). + rewrite zipp_stack_cat; auto. + destruct decompose_stack eqn:decomp. + now apply decompose_stack_noStackApp in decomp. Qed. - Lemma zipp_stack_cat π π' t : - isStackApp π = false -> - zipp t (π' +++ π) = zipp t π'. + Lemma zipc_decompose_stack_empty t π : + (decompose_stack π).2 = ε -> + zipc t π = zipp t π. Proof. - intros no_stack_app. - unfold zipp. - rewrite decompose_stack_stack_cat. - destruct (decompose_stack π') eqn:decomp. + destruct decompose_stack eqn:decomp. + apply decompose_stack_eq in decomp as ->. + cbn; intros ->. + rewrite zipc_appstack, zipp_as_mkApps, decompose_stack_appstack. cbn. - destruct s; try now rewrite app_nil_r. - now destruct π; cbn in *; rewrite ?app_nil_r. + now rewrite app_nil_r. Qed. + + Ltac reduce_stack_facts := + repeat + match goal with + | [H: (?a, ?b) = reduce_stack ?f ?Σ ?wf ?Γ ?t ?π ?h |- _] => + let rid := fresh "r" in + let decompid := fresh "d" in + let whid := fresh "w" in + let isr := fresh "isr" in + pose proof (reduce_stack_sound f Σ wf Γ t π h) as [rid]; + pose proof (reduce_stack_decompose f Σ wf Γ t π h) as decompid; + pose proof (reduce_stack_whnf f Σ wf Γ t π h) as whid; + pose proof (reduce_stack_isred f Σ wf Γ t π h) as isr; + rewrite <- H in rid, decompid, whid, isr; cbn in rid, decompid, whid, isr; + clear H + end. - Lemma zipp_appstack t args π : - zipp t (appstack args π) = zipp (mkApps t args) π. + Lemma zipc_unfold_decompose_stack t π : + zipc t π = zipc (mkApps t (decompose_stack π).1) (decompose_stack π).2. Proof. - unfold zipp. - rewrite decompose_stack_appstack. - rewrite <- mkApps_nested. - now destruct decompose_stack. - Qed. - + rewrite <- zipc_appstack. + destruct (decompose_stack π) eqn:decomp. + now apply decompose_stack_eq in decomp as ->. + Qed. + + Ltac simpl_stacks := + (repeat + match goal with + | [H: (?a, ?b) = decompose_stack ?π |- _] => + is_var a; + is_var b; + apply eta_pair in H; cbn in H; noconf H + end); + (repeat + match goal with + | [H: context[decompose_stack (appstack ?l ?ρ)] |- _] => + (rewrite (decompose_stack_appstack l ρ) in H; cbn in H) || fail 2 + | [H: context[stack_context (?π +++ ?π')] |- _] => + (rewrite (stack_context_stack_cat π' π) in H; cbn in H) || fail 2 + | [H: (decompose_stack ?π).2 = ε, H': context[stack_context ?π] |- _] => + (rewrite <- (stack_context_decompose π), H in H'; cbn in H') || fail 2 + | [H: (decompose_stack ?π).2 = ε, H': context[zipc ?t ?π] |- _] => + (rewrite (zipc_decompose_stack_empty t π H) in H'; cbn in H') || fail 2 + | [H: context[stack_context (decompose_stack ?π).2] |- _] => + (rewrite (stack_context_decompose π) in H; cbn in H) || fail 2 + | [H: context[zipp ?t (?π +++ (decompose_stack ?π').2)] |- _] => + (rewrite (zipp_stack_cat_decompose_stack t π π') in H; cbn in H) || fail 2 + | [H: context[zipc ?t (appstack ?args ?π)] |- _] => + (rewrite (@zipc_appstack t args π) in H; cbn in H) || fail 2 + | [H: context[zipc ?t (?π +++ ?π')] |- _] => + (rewrite (zipc_stack_cat t π π') in H; cbn in H) || fail 2 + | [H: context[zip (mkApps ?t (decompose_stack ?π).1, decompose_stack ?π).2] |- _] => + unfold zip in H + | [H: context[zipc (mkApps ?t (decompose_stack ?π).1) (decompose_stack ?π).2] |- _] => + (rewrite <- (zipc_unfold_decompose_stack t π) in H; cbn in H) || fail 2 + | [H: isStackApp ?π = false, H': context[zipp ?t ?π] |- _] => + (rewrite (zipp_noStackApp t π H) in H'; cbn in H') || fail 2 + | [H: (decompose_stack ?π).2 = (decompose_stack ?π').2, H': context[stack_context ?π] |- _] => + (rewrite <- (stack_context_decompose π), H, (stack_context_decompose π') in H'; cbn in H') + || fail 2 + + | [|- context[decompose_stack (appstack ?l ?ρ)]] => + (rewrite (decompose_stack_appstack l ρ); cbn) || fail 2 + | [|- context[stack_context (?π +++ ?π')]] => + (rewrite (stack_context_stack_cat π' π); cbn) || fail 2 + | [H: (decompose_stack ?π).2 = ε |- context[stack_context ?π]] => + (rewrite <- (stack_context_decompose π), H; cbn) || fail 2 + | [H: (decompose_stack ?π).2 = ε |- context[zipc ?t ?π]] => + (rewrite (zipc_decompose_stack_empty t π H); cbn) || fail 2 + | [|- context[stack_context (decompose_stack ?π).2]] => + (rewrite (stack_context_decompose π); cbn) || fail 2 + | [|- context[zipp ?t (?π +++ (decompose_stack ?π').2)]] => + (rewrite (zipp_stack_cat_decompose_stack t π π'); cbn) || fail 2 + | [|- context[zipc ?t (appstack ?args ?π)]] => + (rewrite (@zipc_appstack t args π); cbn) || fail 2 + | [|- context[zipc ?t (?π +++ ?π')]] => + (rewrite (zipc_stack_cat t π π'); cbn) || fail 2 + | [|- context[zip (mkApps ?t (decompose_stack ?π).1, decompose_stack ?π).2]] => + unfold zip + | [|- context[zipc (mkApps ?t (decompose_stack ?π).1) (decompose_stack ?π).2]] => + (rewrite <- (zipc_unfold_decompose_stack t π); cbn) || fail 2 + | [H: isStackApp ?π = false |- context[zipp ?t ?π]] => + (rewrite (zipp_noStackApp t π H); cbn) || fail 2 + | [H: (decompose_stack ?π).2 = (decompose_stack ?π').2 |- context[stack_context ?π]] => + (rewrite <- (stack_context_decompose π), H, (stack_context_decompose π'); cbn) || fail 2 + end); + repeat + match goal with + | [H: context[zipp ?t ?π] |- _] => rewrite (zipp_as_mkApps t π) in H + | [|- context[zipp ?t ?π]] => rewrite (zipp_as_mkApps t π) + end. + + Ltac simpl_reduce_stack := reduce_stack_facts; simpl_stacks. + (* Tailored view for isconv_prog and precondition for fallback case *) Equations prog_discr (t1 t2 : term) : Prop := prog_discr (tApp _ _) (tApp _ _) := False ; @@ -821,20 +906,9 @@ Section Conversion. Definition Ret s Γ t π t' π' := forall (leq : conv_pb), conv_stack_ctx Γ π π' -> - match s with - | Term - | Fallback => isred_full Γ t π - | _ => True - end -> - match s with - | Term - | Fallback => isred_full Γ t' π' - | _ => True - end -> - match s with - | Fallback => prog_discr t t' - | _ => True - end -> + (match s with Fallback | Term => isred_full Γ t π | _ => True end) -> + (match s with Fallback | Term => isred_full Γ t' π' | _ => True end) -> + (match s with | Fallback => prog_discr t t' | _ => True end) -> match s with | Reduction | Term @@ -913,66 +987,16 @@ Section Conversion. eapply wellformed_zipc_stack_context. all: eassumption. Qed. Next Obligation. - pose proof hΣ as hΣ'. - destruct hΣ' as [wΣ]. - match type of eq1 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r1] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d1 ; - pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c1 - end. - rewrite <- eq1 in r1. - rewrite <- eq1 in d1. cbn in d1. - rewrite <- eq1 in c1. cbn in c1. - rewrite stack_context_appstack in c1. cbn in c1. - - pose proof (decompose_stack_eq _ _ _ (eq_sym e1)). subst. - clear eq1 eq2. - rewrite zipc_appstack in h1. - case_eq (decompose_stack π1'). intros args1' ρ1' e1'. - rewrite e1' in d1. cbn in d1. - rewrite decompose_stack_appstack in d1. cbn in d1. subst. - pose proof (decompose_stack_eq _ _ _ e1'). subst. - rewrite stack_cat_appstack. - rewrite zipc_appstack. - - rewrite stack_context_appstack in r1. cbn in r1. - rewrite 2!zipc_appstack in r1. cbn in r1. - + simpl_reduce_stack. eapply red_wellformed ; try assumption ; revgoals. - - constructor. zip fold. eapply red_context. eassumption. - - cbn. assumption. + - constructor. zip fold. eapply red_context. simpl_stacks. eassumption. + - cbn. simpl_stacks. assumption. Qed. Next Obligation. - pose proof hΣ as hΣ'. - destruct hΣ' as [wΣ]. - match type of eq2 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; - pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 - end. - rewrite <- eq2 in r2. - rewrite <- eq2 in d2. cbn in d2. - rewrite <- eq2 in c2. cbn in c2. - rewrite stack_context_appstack in c2. cbn in c2. - - pose proof (decompose_stack_eq _ _ _ (eq_sym e2)). subst. - clear eq1 eq2 aux. - rewrite zipc_appstack in h2. - case_eq (decompose_stack π2'). intros args2' ρ2' e2'. - rewrite e2' in d2. cbn in d2. - rewrite decompose_stack_appstack in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e2'). subst. - rewrite stack_cat_appstack. - rewrite zipc_appstack. - - rewrite stack_context_appstack in r2. cbn in r2. - rewrite 2!zipc_appstack in r2. cbn in r2. - + simpl_reduce_stack. eapply red_wellformed ; try assumption ; revgoals. - - constructor. zip fold. eapply red_context. eassumption. - - cbn. assumption. + - constructor. zip fold. eapply red_context. simpl_stacks. eassumption. + - cbn. simpl_stacks. assumption. Qed. Next Obligation. match type of eq1 with @@ -1060,208 +1084,33 @@ Section Conversion. assumption. Qed. Next Obligation. - match type of eq1 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d1 ; - pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c1 - end. - rewrite <- eq1 in d1. cbn in d1. - rewrite <- eq1 in c1. cbn in c1. - rewrite stack_context_appstack in c1. cbn in c1. - pose proof (decompose_stack_eq _ _ _ (eq_sym e1)). subst. - match type of eq2 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; - pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 - end. - rewrite <- eq2 in d2. cbn in d2. - rewrite <- eq2 in c2. cbn in c2. - rewrite stack_context_appstack in c2. cbn in c2. - pose proof (decompose_stack_eq _ _ _ (eq_sym e2)). subst. - rewrite 2!stack_context_stack_cat. - rewrite c1. rewrite c2. simpl. - rewrite 2!stack_context_appstack in hx. + simpl_reduce_stack. assumption. Qed. Next Obligation. - match type of eq1 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as r1; - pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w1 - end. - rewrite <- eq1 in r1, w1. - destruct r1 as (ha&hl). + simpl_reduce_stack. + specialize (isr0 eq_refl) as (?&?). split; [easy|]. - rewrite stack_context_stack_cat. - apply (f_equal (snd ∘ decompose_stack ∘ snd)) in eq1. - rewrite reduce_stack_decompose, decompose_stack_appstack in eq1. - cbn in eq1. - rewrite <- (stack_context_decompose π1'), eq1 in *. - rewrite <- (stack_context_decompose π1), <- e1 in w1. - cbn in w1 |- *. - rewrite zipp_stack_cat. - 2: eapply decompose_stack_noStackApp; eauto. - auto. + simpl_stacks. + easy. Qed. Next Obligation. - match type of eq2 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as r2; - pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w2 - end. - rewrite <- eq2 in r2, w2. - destruct r2 as (ha&hl). + simpl_reduce_stack. + specialize (isr eq_refl) as (?&?). split; [easy|]. - rewrite stack_context_stack_cat. - apply (f_equal (snd ∘ decompose_stack ∘ snd)) in eq2. - rewrite reduce_stack_decompose, decompose_stack_appstack in eq2. - cbn in eq2. - rewrite <- (stack_context_decompose π2'), eq2 in *. - rewrite <- (stack_context_decompose π2), <- e2 in w2. - cbn in w2 |- *. - rewrite zipp_stack_cat. - 2: eapply decompose_stack_noStackApp; eauto. - auto. + simpl_stacks. + easy. Qed. Next Obligation. - destruct hΣ as [wΣ]. - unfold zipp. rewrite <- e1. rewrite <- e2. - - match type of eq1 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r1] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d1 ; - pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c1 - end. - rewrite <- eq1 in r1. - rewrite <- eq1 in d1. cbn in d1. - rewrite <- eq1 in c1. cbn in c1. - rewrite stack_context_appstack in c1. cbn in c1. - - match type of eq2 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; - pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 - end. - rewrite <- eq2 in r2. - rewrite <- eq2 in d2. cbn in d2. - rewrite <- eq2 in c2. cbn in c2. - rewrite stack_context_appstack in c2. cbn in c2. - - clear eq1 eq2. - - pose proof (decompose_stack_eq _ _ _ (eq_sym e1)). subst. - case_eq (decompose_stack π1'). intros args1' ρ1' e1'. - rewrite e1' in d1. cbn in d1. - rewrite decompose_stack_appstack in d1. cbn in d1. subst. - pose proof (decompose_stack_eq _ _ _ e1'). subst. - - pose proof (decompose_stack_eq _ _ _ (eq_sym e2)). subst. - case_eq (decompose_stack π2'). intros args2' ρ2' e2'. - rewrite e2' in d2. cbn in d2. - rewrite decompose_stack_appstack in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e2'). subst. - - rewrite stack_context_appstack in r1. cbn in r1. - rewrite 2!zipc_appstack in r1. cbn in r1. - - rewrite stack_context_appstack in r2. cbn in r2. - rewrite 2!zipc_appstack in r2. cbn in r2. - - rewrite 2!stack_cat_appstack in h. - unfold zipp in h. - rewrite 2!decompose_stack_appstack in h. - rewrite decompose_stack_twice with (1 := eq_sym e1) in h. - rewrite decompose_stack_twice with (1 := eq_sym e2) in h. - simpl in h. - rewrite 2!app_nil_r in h. - rewrite 2!stack_context_appstack in hx. - rewrite stack_context_appstack in h. - - rewrite stack_context_appstack. - - destruct hx as [hx]. - etransitivity. - - eapply red_conv_cum_l; eassumption. - - etransitivity. - + eassumption. - + eapply conv_cum_context_convp. - * assumption. - * eapply red_conv_cum_r. all: eauto. - * eapply conv_context_sym. all: auto. + simpl_reduce_stack. + destruct hΣ, hx. + apply -> conv_cum_red_conv_iff; eauto. Qed. Next Obligation. - rename H into e; apply h; clear h. - - (* FIXME: the script below is essentially a copy-paste of the previous obligation - with only the hypothesis and the goal swapped... *) - destruct hΣ as [wΣ]. - unfold zipp in e. rewrite <- e1, <- e2 in e. - - match type of eq1 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r1] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d1 ; - pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c1 - end. - rewrite <- eq1 in r1. - rewrite <- eq1 in d1. cbn in d1. - rewrite <- eq1 in c1. cbn in c1. - rewrite stack_context_appstack in c1. cbn in c1. - - match type of eq2 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; - pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 - end. - rewrite <- eq2 in r2. - rewrite <- eq2 in d2. cbn in d2. - rewrite <- eq2 in c2. cbn in c2. - rewrite stack_context_appstack in c2. cbn in c2. - - clear eq1 eq2. - - pose proof (decompose_stack_eq _ _ _ (eq_sym e1)). subst. - case_eq (decompose_stack π1'). intros args1' ρ1' e1'. - rewrite e1' in d1. cbn in d1. - rewrite decompose_stack_appstack in d1. cbn in d1. subst. - pose proof (decompose_stack_eq _ _ _ e1'). subst. - - pose proof (decompose_stack_eq _ _ _ (eq_sym e2)). subst. - case_eq (decompose_stack π2'). intros args2' ρ2' e2'. - rewrite e2' in d2. cbn in d2. - rewrite decompose_stack_appstack in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e2'). subst. - - rewrite stack_context_appstack in r1. cbn in r1. - rewrite 2!zipc_appstack in r1. cbn in r1. - - rewrite stack_context_appstack in r2. cbn in r2. - rewrite 2!zipc_appstack in r2. cbn in r2. - - rewrite 2!stack_cat_appstack. - unfold zipp. - rewrite 2!decompose_stack_appstack. - rewrite decompose_stack_twice with (1 := eq_sym e1). - rewrite decompose_stack_twice with (1 := eq_sym e2). - simpl. - rewrite 2!app_nil_r. - rewrite 2!stack_context_appstack in hx. - rewrite stack_context_appstack. - - rewrite stack_context_appstack in e. - - destruct hx as [hx]. - etransitivity. - - eapply red_conv_cum_r; eassumption. - - etransitivity. - + eassumption. - + eapply conv_cum_context_convp. - * assumption. - * eapply red_conv_cum_l. all: eauto. - * eapply conv_context_sym. all: auto. + apply h; clear h. + simpl_reduce_stack. + destruct hΣ, hx. + apply <- conv_cum_red_conv_iff; eauto. Qed. Opaque reduce_stack. @@ -1498,13 +1347,6 @@ Section Conversion. reflexivity. Qed. - Lemma zipp_as_mkApps t π : - zipp t π = mkApps t (decompose_stack π).1. - Proof. - unfold zipp. - now destruct decompose_stack. - Qed. - Lemma unfold_one_fix_None Γ mfix idx π wf : None = unfold_one_fix Γ mfix idx π wf -> ∥∑args, @@ -1700,11 +1542,11 @@ Section Conversion. conv_pb_relb leq u u'. Proof. intros all1%For_all_to_spec all2%For_all_to_spec conv. - destruct leq. - - eapply eq_universe_spec'; eauto. + destruct leq; cbn. + - eapply check_eqb_universe_complete; eauto. + apply wf_ext_global_uctx_invariants, hΣ'. + apply global_ext_uctx_consistent, hΣ'. - - eapply leq_universe_spec'; eauto. + - eapply check_leqb_universe_complete; eauto. + apply wf_ext_global_uctx_invariants, hΣ'. + apply global_ext_uctx_consistent, hΣ'. Qed. @@ -1812,22 +1654,6 @@ Section Conversion. now intros dar [|[(?&?&?&?)]]. Qed. - Lemma wellformed_nonarity_mkApps Γ hd args : - destArity [] hd = None -> - wellformed Σ Γ (mkApps hd args) -> - welltyped Σ Γ (mkApps hd args). - Proof. - intros dar wf. - destruct wf as [|[(?&?&dar'&?)]]; auto. - exfalso. - clear -dar dar'. - induction args in args, hd, dar, dar' |- *. - - cbn in *. - congruence. - - cbn in *. - now apply IHargs in dar'; [|easy]. - Qed. - Lemma wellformed_zipc_tConst_inv Γ c u π : wellformed Σ Γ (zipc (tConst c u) π) -> exists cst, @@ -2112,7 +1938,7 @@ Section Conversion. constructor. constructor. all: intuition eauto. Qed. Next Obligation. - (* COntrapositive of previous obligation *) + (* Contrapositive of previous obligation *) apply h; clear h. destruct H as [H]; inversion H; now constructor. Qed. @@ -2193,11 +2019,11 @@ Section Conversion. } ; | Error e h := no e } ; - | @exist false eq1 := Error ( + | @exist false eq1 := no ( FixRargMismatch idx (Γ ,,, stack_context π) u mfix1 mfix2 (Γ ,,, stack_context π') v mfix1' mfix2' - ) _ + ) } ; isconv_fix_types Γ idx mfix1 [] π h mfix1' [] π' h' hx h1 aux := yes ; @@ -2206,11 +2032,11 @@ Section Conversion. and then conclude this case is not possible. *) isconv_fix_types Γ idx mfix1 mfix2 π h mfix1' mfix2' π' h' hx h1 aux := - Error ( + no ( FixMfixMismatch idx (Γ ,,, stack_context π) (mfix1 ++ mfix2) (Γ ,,, stack_context π') (mfix1' ++ mfix2') - ) _. + ). Next Obligation. constructor. constructor. @@ -2637,188 +2463,135 @@ Section Conversion. now apply invert_type_mkApps_tProd in typ. Qed. - Lemma whnf_whne_upgrade Γ t : - whnf RedFlags.default Σ Γ t -> - whne RedFlags.nodelta Σ Γ t -> - whne RedFlags.default Σ Γ t. - Proof. - intros whn whe. - induction whe; cbn in *; try easy. - - now depelim whn; solve_discr. - - depelim whn. - + easy. - + destruct v0 as [|? ? _] using List.rev_ind; [discriminate H|]. - rewrite <- mkApps_nested in H. - cbn in H; noconf H. - constructor. - eapply IHwhe. - apply whnf_cstrapp. - + destruct v0 as [|? ? _] using List.rev_ind; [discriminate H|]. - rewrite <- mkApps_nested in H. - cbn in H; noconf H. - constructor. - eapply IHwhe. - apply whnf_indapp. - + destruct v0 as [|? ? _] using List.rev_ind; [discriminate H0|]. - rewrite <- mkApps_nested in H0. - cbn in H0; noconf H0. - constructor. - eapply IHwhe. - apply whnf_fixapp. - destruct unfold_fix; [|easy]. - destruct p. - apply nth_error_None. - apply nth_error_None in H. - rewrite app_length in H; cbn in H. - lia. - + destruct v0 as [|? ? _] using List.rev_ind; [discriminate H|]. - rewrite <- mkApps_nested in H. - cbn in H; noconf H. - constructor. - eapply IHwhe. - apply whnf_cofixapp. - - depelim whn; solve_discr; try easy. - rewrite H1 in H. - congruence. - - depelim whn; solve_discr; easy. - - depelim whn; solve_discr; easy. - Qed. - -(* TODO: move to safe emmata *) - Lemma whne_context_relation f rel Γ Γ' t : - (forall Γ Γ' c c', rel Γ Γ' c c' -> (decl_body c = None <-> decl_body c' = None)) -> - whne f Σ Γ t -> - context_relation rel Γ Γ' -> - whne f Σ Γ' t. + Lemma reduced_case_discriminee_whne Γ π ind par p c brs h : + eqb_term (reduce_term + RedFlags.default + Σ hΣ (Γ,,, stack_context π) c h) c = true -> + isred_full Γ (tCase (ind, par) p c brs) π -> + whne RedFlags.default Σ (Γ,,, stack_context π) c. Proof. - intros behaves wh conv. - induction wh; eauto using whne. - destruct nth_error eqn:nth; [|easy]. - cbn in H. - eapply context_relation_nth in nth; eauto. - destruct nth as (?&?&?&?). - constructor. - rewrite e. - cbn. - specialize (behaves _ _ _ _ r). - f_equal. - apply behaves. - congruence. + intros eq ir. + destruct ir as (_&wh). + apply eqb_term_spec in eq. + eapply whnf_eq_term in eq; [|now apply reduce_term_complete]. + rewrite zipp_as_mkApps in wh. + depelim wh; solve_discr. + apply whne_mkApps_inv in H as [|(?&?&?&?&?&?&?&?&?)]; [|easy|easy]. + depelim H; cbn in *; try easy; solve_discr. + apply whnf_whne_nodelta_upgrade in eq; auto. Qed. - Lemma whnf_context_relation f rel Γ Γ' t : - (forall Γ Γ' c c', rel Γ Γ' c c' -> (decl_body c = None <-> decl_body c' = None)) -> - whnf f Σ Γ t -> - context_relation rel Γ Γ' -> - whnf f Σ Γ' t. + Lemma inv_reduced_discriminees_case leq Γ π π' ind ind' par par' p p' c c' brs brs' h h' : + conv_stack_ctx Γ π π' -> + true = eqb_term (reduce_term + RedFlags.default + Σ hΣ (Γ,,, stack_context π) c h) c && + eqb_term (reduce_term + RedFlags.default + Σ hΣ (Γ,,, stack_context π') c' h') c' -> + isred_full Γ (tCase (ind, par) p c brs) π -> + isred_full Γ (tCase (ind', par') p' c' brs') π' -> + conv_cum + leq Σ (Γ,,, stack_context π) + (zipp (tCase (ind, par) p c brs) π) + (zipp (tCase (ind', par') p' c' brs') π') -> + ∥(ind, par) = (ind', par') × + Σ;;; Γ,,, stack_context π |- p = p' × + Σ;;; Γ,,, stack_context π |- c = c' × + All2 (fun br br' => br.1 = br'.1 × (Σ;;; Γ,,, stack_context π |- br.2 = br'.2)) brs brs' × + conv_terms Σ (Γ,,, stack_context π) (decompose_stack π).1 (decompose_stack π').1∥. Proof. - intros behaves wh conv. - destruct wh; eauto using whnf. - apply whnf_ne. - eapply whne_context_relation; eauto. + intros [] eq isr1 isr2 cc. + symmetry in eq; apply Bool.andb_true_iff in eq as (c_is_red&c'_is_red). + eapply reduced_case_discriminee_whne in c_is_red as wh1; eauto. + eapply reduced_case_discriminee_whne in c'_is_red as wh2; eauto. + destruct hΣ. + rewrite !zipp_as_mkApps in cc. + eapply whne_conv_context in wh2. + 2: eapply conv_context_sym; eauto. + apply conv_cum_mkApps_inv in cc as [(conv_case&conv_args)]; eauto using whnf_mkApps. + eapply conv_cum_tCase_inv in conv_case; eauto. + destruct conv_case as [([= <- <-]&?&?&?)]. + constructor; auto. Qed. - - Lemma whne_conv_context f Γ Γ' t : - whne f Σ Γ t -> - conv_context Σ Γ Γ' -> - whne f Σ Γ' t. + + Lemma reduced_proj_body_whne Γ π p c h : + true = eqb_term (reduce_term + RedFlags.default + Σ hΣ (Γ,,, stack_context π) c h) c -> + isred_full Γ (tProj p c) π -> + whne RedFlags.default Σ (Γ,,, stack_context π) c. Proof. - apply whne_context_relation. - intros ? ? ? ? r. - now depelim r. + intros eq%eq_sym ir. + destruct ir as (_&wh). + apply eqb_term_spec in eq. + eapply whnf_eq_term in eq; [|now apply reduce_term_complete]. + rewrite zipp_as_mkApps in wh. + depelim wh; solve_discr. + apply whne_mkApps_inv in H as [|(?&?&?&?&?&?&?&?&?)]; [|easy|easy]. + depelim H; cbn in *; try easy; solve_discr. + apply whnf_whne_nodelta_upgrade in eq; auto. Qed. - Lemma whnf_conv_context f Γ Γ' t : - whnf f Σ Γ t -> - conv_context Σ Γ Γ' -> - whnf f Σ Γ' t. - Proof. - apply whnf_context_relation. - intros ? ? ? ? r. - now depelim r. - Qed. - - Lemma cumul_mkApps Γ hd args hd' args' : - wf Σ -> - Σ;;; Γ |- hd <= hd' -> - conv_terms Σ Γ args args' -> - Σ;;; Γ |- mkApps hd args <= mkApps hd' args'. - Proof. - intros wf cum cum_args. - revert hd hd' cum. - induction cum_args; intros hd hd' cum; auto. - cbn. - apply IHcum_args. - eapply cumul_trans; auto. - - eapply cumul_App_l; eauto. - - eapply cumul_App_r; eauto. - Qed. - - Lemma conv_cum_zipp leq Γ t t' π π' : - conv_cum leq Σ Γ t t' -> - ∥conv_terms Σ Γ (decompose_stack π).1 (decompose_stack π').1∥ -> - conv_cum leq Σ Γ (zipp t π) (zipp t' π'). - Proof. - intros conv conv_args. - rewrite !zipp_as_mkApps. - destruct leq; cbn in *. - - sq. - apply mkApps_conv_args; auto. - - sq. - apply cumul_mkApps; auto. - Qed. - - Lemma conv_terms_conv_ctx Γ Γ' ts ts' : - wf Σ -> - conv_terms Σ Γ ts ts' -> - conv_context Σ Γ Γ' -> - conv_terms Σ Γ' ts ts'. + Lemma inv_reduced_body_proj leq Γ π π' p p' c c' h h' : + conv_stack_ctx Γ π π' -> + true = eqb_term (reduce_term + RedFlags.default + Σ hΣ (Γ,,, stack_context π) c h) c -> + true = eqb_term (reduce_term + RedFlags.default + Σ hΣ (Γ,,, stack_context π') c' h') c' -> + isred_full Γ (tProj p c) π -> + isred_full Γ (tProj p' c') π' -> + conv_cum leq Σ (Γ,,, stack_context π) (zipp (tProj p c) π) (zipp (tProj p' c') π') -> + ∥p = p' × + Σ;;; Γ,,, stack_context π |- c = c' × + conv_terms Σ (Γ,,, stack_context π) (decompose_stack π).1 (decompose_stack π').1∥. Proof. - intros wf conv ctx. - induction conv; [constructor|]. + intros [] c_is_red c'_is_red isr1 isr2 cc. + eapply reduced_proj_body_whne in c_is_red as wh1; eauto. + eapply reduced_proj_body_whne in c'_is_red as wh2; eauto. + destruct hΣ. + rewrite !zipp_as_mkApps in cc. + eapply whne_conv_context in wh2. + 2: eapply conv_context_sym; eauto. + apply conv_cum_mkApps_inv in cc as [(conv_proj&conv_args)]; eauto using whnf_mkApps. + eapply conv_cum_tProj_inv in conv_proj; eauto. + destruct conv_proj as [(<-&?)]. constructor; auto. - eapply conv_conv_ctx; eauto. Qed. - Lemma conv_terms_red Γ ts ts' tsr tsr' : - wf Σ -> - conv_terms Σ Γ tsr tsr' -> - All2 (red Σ Γ) ts tsr -> - All2 (red Σ Γ) ts' tsr' -> - conv_terms Σ Γ ts ts'. - Proof. - intros wf conv all all'. - induction conv in ts, ts', all, all' |- *; depelim all; depelim all'; [constructor|]. - constructor; [|auto]. - eapply red_conv_conv; eauto. - symmetry. - eapply red_conv_conv; eauto. - symmetry. - eauto. - Qed. - - Lemma conv_terms_red' Γ Γ' ts ts' tsr tsr' : - wf Σ -> - conv_context Σ Γ Γ' -> - conv_terms Σ Γ tsr tsr' -> - All2 (red Σ Γ) ts tsr -> - All2 (red Σ Γ') ts' tsr' -> - conv_terms Σ Γ ts ts'. + Lemma inv_stuck_fixes leq Γ mfix idx π mfix' idx' π' h h' : + conv_stack_ctx Γ π π' -> + None = unfold_one_fix Γ mfix idx π h -> + None = unfold_one_fix Γ mfix' idx' π' h' -> + conv_cum leq Σ (Γ,,, stack_context π) (zipp (tFix mfix idx) π) (zipp (tFix mfix' idx') π') -> + ∥idx = idx' × + All2 + (fun d d' => + rarg d = rarg d' × + Σ;;; Γ,,, stack_context π |- dtype d = dtype d' × + Σ;;; Γ,,, stack_context π,,, fix_context mfix |- dbody d = dbody d') mfix mfix' × + conv_terms Σ (Γ,,, stack_context π) (decompose_stack π).1 (decompose_stack π').1∥. Proof. - intros wf convctx conv all all2. - eapply conv_terms_red. - 1: eauto. - 2: eauto. - 2: apply All2_same; reflexivity. - eapply conv_terms_conv_ctx; eauto. + intros [?] uf1 uf2 cc. + rewrite !zipp_as_mkApps in cc. + apply unfold_one_fix_None in uf1 as [(?&?&?)]. + apply unfold_one_fix_None in uf2 as [(?&?&?)]. + destruct hΣ. + eapply conv_cum_red_conv_inv in cc. + 2: assumption. + 2: eassumption. + 2: apply red_mkApps; [reflexivity|exact a]. + 2: apply red_mkApps; [reflexivity|exact a0]. + apply conv_cum_mkApps_inv in cc as [(conv_fix&conv_args)]; auto. + 2: eapply whnf_conv_context; eauto. 2: eapply conv_context_sym; eauto. - eapply conv_terms_red. - 1: eauto. - 3: eauto. - 2: apply All2_same; reflexivity. - eapply conv_terms_conv_ctx; eauto. + apply conv_cum_tFix_inv in conv_fix as [(<-&?)]. + constructor; split; [|split]; auto. + eapply conv_terms_red_conv; eauto. Qed. - + (* Check if the heads of an application are equal, and if so, if all args are convertible *) Notation eqb_heads_and_convertible_args leq hd hd' π π' aux := @@ -2853,7 +2626,7 @@ Section Conversion. isconv_red leq (subst_instance_constr u body) π1 (subst_instance_constr u' body) π2 aux ; (* Inductive or not found *) - | @exist _ _ := Error (NotFoundConstant c) _ + | @exist _ _ := no (NotFoundConstant c) } } ; (* If universes are different, we unfold one of the constants *) @@ -2871,11 +2644,11 @@ Section Conversion. t1 (Lambda_tm na A1 π1) t2 (Lambda_tm na' A2 π2) aux ; | Error e h := - Error ( + no ( LambdaNotConvertibleTypes (Γ ,,, stack_context π1) na A1 t1 (Γ ,,, stack_context π2) na' A2 t2 e - ) _ + ) } ; | prog_view_Prod na A1 B1 na' A2 B2 @@ -2885,11 +2658,11 @@ Section Conversion. B1 (Prod_r na A1 π1) B2 (Prod_r na' A2 π2) aux ; | Error e h := - Error ( + no ( ProdNotConvertibleDomains (Γ ,,, stack_context π1) na A1 B1 (Γ ,,, stack_context π2) na' A2 B2 e - ) _ + ) } ; | prog_view_Case ind par p c brs ind' par' p' c' brs' @@ -2898,8 +2671,8 @@ Section Conversion. (tCase (ind, par) p c brs) (tCase (ind', par') p' c' brs') π1 π2 aux) := { - | @exist true eq1 := Success _; - | @exist false _ with inspect (reduce_term RedFlags.default Σ hΣ (Γ ,,, stack_context π1) c _) := { + | @exist true eq1 := yes; + | @exist false not_eqb with inspect (reduce_term RedFlags.default Σ hΣ (Γ ,,, stack_context π1) c _) := { | @exist cred eq1 with inspect (reduce_term RedFlags.default Σ hΣ (Γ ,,, stack_context π2) c' _) := { | @exist cred' eq2 with inspect (eqb_term cred c && eqb_term cred' c') := { | @exist true eq3 with inspect (eqb (ind, par) (ind', par')) := { @@ -2917,18 +2690,18 @@ Section Conversion. | Success h3 := isconv_args leq (tCase (ind, par) p c brs) π1 (tCase (ind', par') p' c' brs') π2 aux; - | Error e h := Error e _ + | Error e h := no e } ; - | Error e h := Error e _ + | Error e h := no e } ; - | Error e h := Error e _ + | Error e h := no e } ; | @exist false eqDiff := - Error ( + no ( CaseOnDifferentInd (Γ ,,, stack_context π1) ind par p c brs (Γ ,,, stack_context π2) ind' par' p' c' brs' - ) _ + ) } ; | @exist false eq3 := isconv_red leq (tCase (ind, par) p cred brs) π1 @@ -2945,14 +2718,14 @@ Section Conversion. | @exist true eq4 with inspect (eqb p p') := { | @exist true eq5 with isconv_red_raw Conv c (Proj p π1) c' (Proj p' π2) aux := { | Success h1 := isconv_args leq (tProj p c) π1 (tProj p' c') π2 aux ; - | Error e h := Error e _ + | Error e h := no e } ; | @exist false eq5 := - Error ( + no ( DistinctStuckProj (Γ ,,, stack_context π1) p c (Γ ,,, stack_context π2) p' c' - ) _ + ) } ; | @exist false eq4 := isconv_red leq @@ -2975,7 +2748,7 @@ Section Conversion. (tFix mfix idx) (tFix mfix' idx') π1 π2 aux) := { - | @exist true eq1 := Success _; + | @exist true eq1 := yes; | @exist false eqDiff with inspect (unfold_one_fix Γ mfix idx π1 _) := { | @exist (Some (fn, θ)) eq1 with inspect (decompose_stack θ) := { @@ -2998,16 +2771,16 @@ Section Conversion. | @exist true eq4 with isconv_fix Γ mfix idx π1 _ mfix' idx' π2 _ _ _ aux := { | Success h1 with isconv_args_raw leq (tFix mfix idx) π1 (tFix mfix' idx') π2 aux := { | Success h2 := yes ; - | Error e h := Error e _ + | Error e h := no e } ; - | Error e h := Error e _ + | Error e h := no e } ; | @exist false idx_uneq := - Error ( + no ( CannotUnfoldFix (Γ ,,, stack_context π1) mfix idx (Γ ,,, stack_context π2) mfix' idx' - )_ + ) } } } @@ -3017,11 +2790,11 @@ Section Conversion. with inspect (eqb_term (tCoFix mfix idx) (tCoFix mfix' idx')) := { | @exist true eq1 := isconv_args leq (tCoFix mfix idx) π1 (tCoFix mfix' idx') π2 aux ; | @exist false _ := - Error ( + no ( DistinctCoFix (Γ ,,, stack_context π1) mfix idx (Γ ,,, stack_context π2) mfix' idx' - ) _ (* TODO Incomplete *) + ) (* TODO Incomplete *) } ; | prog_view_other t1 t2 h := @@ -3076,13 +2849,14 @@ Section Conversion. Qed. Next Obligation. apply h; cbn; clear h. + destruct hΣ. eapply conv_cum_red_inv. - - exact hΣ'. - - exact H. + - assumption. - apply red_zipp. eapply red_const; eauto. - apply red_zipp. eapply red_const; eauto. + - exact H. Qed. Next Obligation. apply h; clear h. @@ -3130,23 +2904,10 @@ Section Conversion. Next Obligation. destruct h0 as [h0]. destruct hx as [hx]. - - unfold zipp in h0. simpl in h0. - unfold zipp in h. simpl in h. - unfold zipp. - case_eq (decompose_stack π1). intros l1 ρ1 e1. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - pose proof (decompose_stack_eq _ _ _ e1). subst. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite stack_context_appstack in h0. - rewrite stack_context_appstack in h. - apply isred_full_nobeta in ir1; [|easy]. - destruct l1 ; try discriminate ir1. clear ir1. - apply isred_full_nobeta in ir2; [|easy]. - destruct l2 ; try discriminate ir2. clear ir2. - simpl in *. + cbn in *. + simpl_stacks. destruct hΣ. now eapply conv_cum_Lambda. Qed. @@ -3155,59 +2916,23 @@ Section Conversion. apply h; clear h. destruct h0 as [h0]. destruct hx as [hx]. - - unfold zipp in h0. simpl in h0. - unfold zipp in H. simpl in H. - unfold zipp. - case_eq (decompose_stack π1). intros l1 ρ1 e1. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - pose proof (decompose_stack_eq _ _ _ e1). subst. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite stack_context_appstack in h0. - rewrite stack_context_appstack in H. - apply isred_full_nobeta in ir1; [|easy]. - destruct l1 ; try discriminate ir1. clear ir1. - apply isred_full_nobeta in ir2; [|easy]. - destruct l2 ; try discriminate ir2. clear ir2. - rewrite e1, e2 in H. - simpl in *. - - pose proof (wellformed_zipc_stack_context _ hΣ _ _ _ _ _ e1 h1) as - [wf]%wellformed_wf_local. + cbn in *. + simpl_stacks. destruct hΣ. - - apply Lambda_conv_cum_inv in H. 2: assumption. - destruct H as [_ ?]; assumption. + apply Lambda_conv_cum_inv in H as (_&?); auto. Qed. Next Obligation. (* Contrapositive of previous obligation *) apply h; clear h. destruct hx as [hx]. - - unfold zipp in H. simpl in H. - unfold zipp. - case_eq (decompose_stack π1). intros l1 ρ1 e1. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - pose proof (decompose_stack_eq _ _ _ e1). subst. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite stack_context_appstack in H. - apply isred_full_nobeta in ir1; [|easy]. - destruct l1 ; try discriminate ir1. clear ir1. - apply isred_full_nobeta in ir2; [|easy]. - destruct l2 ; try discriminate ir2. clear ir2. - rewrite e1, e2 in H. - simpl in *. - - pose proof (wellformed_zipc_stack_context _ hΣ _ _ _ _ _ e1 h1) as - [wf]%wellformed_wf_local. + cbn in *. + simpl_stacks. destruct hΣ. - - apply Lambda_conv_cum_inv in H. 2: assumption. - destruct H as [[?] _]; now constructor. + apply Lambda_conv_cum_inv in H as (?&_); auto. Qed. (* tProd *) @@ -3232,69 +2957,42 @@ Section Conversion. destruct hΣ as [wΣ]. destruct h0 as [h0]. destruct hx as [hx]. - unfold zipp in h0. simpl in h0. - unfold zipp in h. simpl in h. - unfold zipp. - case_eq (decompose_stack π1). intros l1 ρ1 e1. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - pose proof (decompose_stack_eq _ _ _ e1). subst. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite stack_context_appstack in h0. - rewrite stack_context_appstack in h. - - eapply wellformed_zipc_stack_context in h1 ; tea. - rewrite stack_context_appstack in h1. - rewrite zipc_appstack in h1. simpl in h1. - pose proof (wellformed_wf_local _ _ h1) as hw1. - apply mkApps_Prod_nil' in h1 ; auto. subst. - destruct hw1 as [hw1]. + cbn in *. + apply wellformed_zipc_zipp in h1; auto. clear aux. - eapply wellformed_zipc_stack_context in h2 ; tea. - rewrite stack_context_appstack in h2. - pose proof (wellformed_wf_local _ _ h2) as hw2. - rewrite zipc_appstack in h2. - apply mkApps_Prod_nil' in h2 ; auto. subst. - - simpl. - eapply conv_cum_Prod. all: auto. + apply wellformed_zipc_zipp in h2; auto. + rewrite !zipp_as_mkApps in *. + apply mkApps_Prod_nil' in h1 as ->; auto. + apply mkApps_Prod_nil' in h2 as ->; auto. + eapply conv_cum_Prod; auto. Qed. Next Obligation. (* Codomains are not convertible *) apply h; clear h. destruct hΣ as [wΣ], h0 as [h0], hx as [hx]. - unfold zipp in h0, H |- *. simpl in *. - - case_eq (decompose_stack π1). intros l1 ρ1 e1. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - pose proof (decompose_stack_eq _ _ _ e1). subst. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite stack_context_appstack in h0, H |- *. - rewrite e1, e2 in H. - pose proof (wellformed_zipc_tProd_appstack_nil h1). - pose proof (wellformed_zipc_tProd_appstack_nil h2). - subst; simpl in *. - apply Prod_conv_cum_inv in H. 2: assumption. - now destruct H as [_ H]. + cbn in *. + apply wellformed_zipc_zipp in h1; auto. + clear aux. + apply wellformed_zipc_zipp in h2; auto. + rewrite !zipp_as_mkApps in *. + apply mkApps_Prod_nil' in h1; auto. + apply mkApps_Prod_nil' in h2; auto. + rewrite h1, h2 in H. + apply Prod_conv_cum_inv in H as (_&?); auto. Qed. Next Obligation. (* Domains are not convertible *) apply h; clear h. - destruct hΣ as [wΣ], hx as [hx]. - unfold zipp in H |- *. simpl in *. - - case_eq (decompose_stack π1). intros l1 ρ1 e1. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - pose proof (decompose_stack_eq _ _ _ e1). subst. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite stack_context_appstack in H |- *. - rewrite e1, e2 in H. - pose proof (wellformed_zipc_tProd_appstack_nil h1). - pose proof (wellformed_zipc_tProd_appstack_nil h2). - subst; simpl in *. - - apply Prod_conv_cum_inv in H. 2: assumption. - now destruct H. + cbn in *. + apply wellformed_zipc_zipp in h1; auto. + clear aux. + apply wellformed_zipc_zipp in h2; auto. + rewrite !zipp_as_mkApps in *. + apply mkApps_Prod_nil' in h1; auto. + apply mkApps_Prod_nil' in h2; auto. + rewrite h1, h2 in H. + apply Prod_conv_cum_inv in H as (?&_); auto. Qed. @@ -3314,7 +3012,7 @@ Section Conversion. eapply eqb_term_spec. auto. Qed. Next Obligation. - clear wildcard7. + clear not_eqb. destruct hΣ as [wΣ]. zip fold in h1. apply wellformed_context in h1 ; auto. simpl in h1. destruct h1 as [[T h1] | [[ctx [s [h1 _]]]]] ; [| discriminate ]. @@ -3325,7 +3023,7 @@ Section Conversion. Qed. Next Obligation. destruct hΣ as [wΣ]. - clear aux wildcard7. + clear aux not_eqb. zip fold in h2. apply wellformed_context in h2 ; auto. simpl in h2. destruct h2 as [[T h2] | [[ctx [s [h2 _]]]]] ; [| discriminate ]. apply inversion_Case in h2 as hh ; auto. @@ -3377,149 +3075,26 @@ Section Conversion. Qed. Next Obligation. apply h; clear h. - (* Proof idea: - 1) prove that the 2 tCase are fully whnf because the discriminees are - equal to their full reduction - 2) prove that the 2 tCase are in conv_cum relation because they are whnf - 3) wlog leq = Conv because there is no cumulativity between tCase - 4) invert the conv judgement using that the tCase are whne - (so preserved by wh-reductions) - *) - destruct ir1 as (_&wh1), ir2 as (_&wh2). - cbn in *. - symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). - apply eqb_term_spec in c_is_red. - apply eqb_term_spec in c'_is_red. - eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. - eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. - rewrite !zipp_as_mkApps in H. - rewrite zipp_as_mkApps in wh1, wh2. - apply whnf_mkApps_inv in wh1; [|easy]. - apply whnf_mkApps_inv in wh2; [|easy]. - depelim wh1; solve_discr. - depelim wh2; solve_discr. - depelim H; cbn in *; try easy; solve_discr. - depelim H0; cbn in *; try easy; solve_discr. - apply whnf_whne_upgrade in c_is_red; auto. - apply whnf_whne_upgrade in c'_is_red; auto. - destruct hΣ, hx. - assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). - { eapply whne_conv_context; eauto. - apply conv_context_sym; eauto. } - apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. - 2-3: apply whnf_mkApps; eauto using whne. - apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. - now constructor. + eapply inv_reduced_discriminees_case in H as [([= <- <-]&?&?&?&?)]; eauto. + constructor; auto. Qed. Next Obligation. apply h; cbn; clear h. - destruct ir1 as (_&wh1), ir2 as (_&wh2). - cbn in *. - symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). - apply eqb_term_spec in c_is_red. - apply eqb_term_spec in c'_is_red. - eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. - eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. - rewrite !zipp_as_mkApps in H. - rewrite zipp_as_mkApps in wh1, wh2. - apply whnf_mkApps_inv in wh1; [|easy]. - apply whnf_mkApps_inv in wh2; [|easy]. - depelim wh1; solve_discr. - depelim wh2; solve_discr. - depelim H; cbn in *; try easy; solve_discr. - depelim H0; cbn in *; try easy; solve_discr. - apply whnf_whne_upgrade in c_is_red; auto. - apply whnf_whne_upgrade in c'_is_red; auto. - destruct hΣ, hx. - assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). - { eapply whne_conv_context; eauto. - apply conv_context_sym; eauto. } - apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. - 2-3: apply whnf_mkApps; eauto using whne. - apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. - now constructor. + eapply inv_reduced_discriminees_case in H as [([= <- <-]&?&?&?&?)]; eauto. + constructor; auto. Qed. Next Obligation. apply h; cbn; clear h. - destruct ir1 as (_&wh1), ir2 as (_&wh2). - cbn in *. - symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). - apply eqb_term_spec in c_is_red. - apply eqb_term_spec in c'_is_red. - eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. - eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. - rewrite !zipp_as_mkApps in H. - rewrite zipp_as_mkApps in wh1, wh2. - apply whnf_mkApps_inv in wh1; [|easy]. - apply whnf_mkApps_inv in wh2; [|easy]. - depelim wh1; solve_discr. - depelim wh2; solve_discr. - depelim H; cbn in *; try easy; solve_discr. - depelim H0; cbn in *; try easy; solve_discr. - apply whnf_whne_upgrade in c_is_red; auto. - apply whnf_whne_upgrade in c'_is_red; auto. - destruct hΣ, hx. - assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). - { eapply whne_conv_context; eauto. - apply conv_context_sym; eauto. } - apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. - 2-3: apply whnf_mkApps; eauto using whne. - apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. - now constructor. + eapply inv_reduced_discriminees_case in H as [([= <- <-]&?&?&?&?)]; eauto. + constructor; auto. Qed. Next Obligation. apply h; cbn; clear h. - destruct ir1 as (_&wh1), ir2 as (_&wh2). - cbn in *. - symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). - apply eqb_term_spec in c_is_red. - apply eqb_term_spec in c'_is_red. - eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. - eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. - rewrite !zipp_as_mkApps in H. - rewrite zipp_as_mkApps in wh1, wh2. - apply whnf_mkApps_inv in wh1; [|easy]. - apply whnf_mkApps_inv in wh2; [|easy]. - depelim wh1; solve_discr. - depelim wh2; solve_discr. - depelim H; cbn in *; try easy; solve_discr. - depelim H0; cbn in *; try easy; solve_discr. - apply whnf_whne_upgrade in c_is_red; auto. - apply whnf_whne_upgrade in c'_is_red; auto. - destruct hΣ, hx. - assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). - { eapply whne_conv_context; eauto. - apply conv_context_sym; eauto. } - apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. - 2-3: apply whnf_mkApps; eauto using whne. - apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. + eapply inv_reduced_discriminees_case in H as [([= <- <-]&?&?&?&?)]; eauto. constructor; auto. Qed. Next Obligation. - destruct ir1 as (_&wh1), ir2 as (_&wh2). - cbn in *. - symmetry in eq3; apply Bool.andb_true_iff in eq3 as (c_is_red&c'_is_red). - apply eqb_term_spec in c_is_red. - apply eqb_term_spec in c'_is_red. - eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. - eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. - rewrite !zipp_as_mkApps in H. - rewrite zipp_as_mkApps in wh1, wh2. - apply whnf_mkApps_inv in wh1; [|easy]. - apply whnf_mkApps_inv in wh2; [|easy]. - depelim wh1; solve_discr. - depelim wh2; solve_discr. - depelim H; cbn in *; try easy; solve_discr. - depelim H0; cbn in *; try easy; solve_discr. - apply whnf_whne_upgrade in c_is_red; auto. - apply whnf_whne_upgrade in c'_is_red; auto. - destruct hΣ, hx. - assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). - { eapply whne_conv_context; eauto. - apply conv_context_sym; eauto. } - apply conv_cum_mkApps_inv in H1 as [(conv_case&conv_args)]; auto. - 2-3: apply whnf_mkApps; eauto using whne. - apply conv_cum_tCase_inv in conv_case as [([= <- <-]&conv_p&conv_c&conv_brs)]; eauto. + eapply inv_reduced_discriminees_case in H as [([= <- <-]&?&?&?&?)]; eauto. rewrite eq_inductive_refl, Nat.eqb_refl in eqDiff. congruence. Qed. @@ -3640,7 +3215,7 @@ Section Conversion. clear. induction brs ; eauto. - etransitivity ; try eassumption. - eapply conv_cum_context_convp. + eapply conv_cum_conv_ctx. + assumption. + eapply red_conv_cum_r. 1: assumption. eapply red_zipp. @@ -3721,95 +3296,19 @@ Section Conversion. Qed. Next Obligation. apply h; cbn; clear h. - destruct ir1 as (_&wh1), ir2 as (_&wh2). - cbn in *. - rename eq3 into c_is_red. - rename eq4 into c'_is_red. - symmetry in c'_is_red. - apply eqb_term_spec in c'_is_red. - eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. - symmetry in c_is_red. - apply eqb_term_spec in c_is_red. - eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. - rewrite !zipp_as_mkApps in H. - rewrite zipp_as_mkApps in wh1, wh2. - apply whnf_mkApps_inv in wh1; [|easy]. - apply whnf_mkApps_inv in wh2; [|easy]. - depelim wh1; solve_discr. - depelim wh2; solve_discr. - depelim H; cbn in *; try easy; solve_discr. - depelim H0; cbn in *; try easy; solve_discr. - apply whnf_whne_upgrade in c_is_red; auto. - apply whnf_whne_upgrade in c'_is_red; auto. - destruct hΣ, hx. - assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). - { eapply whne_conv_context; eauto. - apply conv_context_sym; eauto. } - apply conv_cum_mkApps_inv in H1 as [(conv_proj&conv_args)]; auto. - 2-3: apply whnf_mkApps; eauto using whne. - apply conv_cum_tProj_inv in conv_proj as [(->&?)]; auto. + eapply inv_reduced_body_proj in H; eauto. + destruct H as [(<-&?&?)]. constructor; auto. Qed. Next Obligation. apply h; cbn; clear h. - destruct ir1 as (_&wh1), ir2 as (_&wh2). - cbn in *. - rename eq3 into c_is_red. - rename eq4 into c'_is_red. - symmetry in c'_is_red. - apply eqb_term_spec in c'_is_red. - eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. - symmetry in c_is_red. - apply eqb_term_spec in c_is_red. - eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. - rewrite !zipp_as_mkApps in H. - rewrite zipp_as_mkApps in wh1, wh2. - apply whnf_mkApps_inv in wh1; [|easy]. - apply whnf_mkApps_inv in wh2; [|easy]. - depelim wh1; solve_discr. - depelim wh2; solve_discr. - depelim H; cbn in *; try easy; solve_discr. - depelim H0; cbn in *; try easy; solve_discr. - apply whnf_whne_upgrade in c_is_red; auto. - apply whnf_whne_upgrade in c'_is_red; auto. - destruct hΣ, hx. - assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). - { eapply whne_conv_context; eauto. - apply conv_context_sym; eauto. } - apply conv_cum_mkApps_inv in H1 as [(conv_proj&conv_args)]; auto. - 2-3: apply whnf_mkApps; eauto using whne. - apply conv_cum_tProj_inv in conv_proj as [(->&?)]; auto. + eapply inv_reduced_body_proj in H; eauto. + destruct H as [(<-&?&?)]. constructor; auto. Qed. Next Obligation. - (* Proof idea: c and c' are whne so from H, p = p' contradicting eqDiff *) - destruct ir1 as (_&wh1), ir2 as (_&wh2). - cbn in *. - rename eq3 into c_is_red. - rename eq4 into c'_is_red. - symmetry in c'_is_red. - apply eqb_term_spec in c'_is_red. - eapply whnf_eq_term in c'_is_red; [|now apply reduce_term_complete]. - symmetry in c_is_red. - apply eqb_term_spec in c_is_red. - eapply whnf_eq_term in c_is_red; [|now apply reduce_term_complete]. - rewrite !zipp_as_mkApps in H. - rewrite zipp_as_mkApps in wh1, wh2. - apply whnf_mkApps_inv in wh1; [|easy]. - apply whnf_mkApps_inv in wh2; [|easy]. - depelim wh1; solve_discr. - depelim wh2; solve_discr. - depelim H; cbn in *; try easy; solve_discr. - depelim H0; cbn in *; try easy; solve_discr. - apply whnf_whne_upgrade in c_is_red; auto. - apply whnf_whne_upgrade in c'_is_red; auto. - destruct hΣ, hx. - assert (whne RedFlags.default Σ (Γ,,, stack_context π1) c'). - { eapply whne_conv_context; eauto. - apply conv_context_sym; eauto. } - apply conv_cum_mkApps_inv in H1 as [(conv_proj&conv_args)]; auto. - 2-3: apply whnf_mkApps; eauto using whne. - apply conv_cum_tProj_inv in conv_proj as [(->&?)]; auto. + eapply inv_reduced_body_proj in H; eauto. + destruct H as [(<-&?&?)]. rewrite eq_prod_refl in eq5; eauto using eq_prod_refl, Nat.eqb_refl, eq_string_refl, eq_inductive_refl. Qed. @@ -3863,7 +3362,7 @@ Section Conversion. destruct hr as [hr]. etransitivity. - eassumption. - - eapply conv_cum_context_convp. + - eapply conv_cum_conv_ctx. + assumption. + eapply red_conv_cum_r ; try assumption. eapply red_zipp. @@ -3881,7 +3380,7 @@ Section Conversion. end. destruct hr as [hr]. etransitivity; [eassumption|]. - eapply conv_cum_context_convp; eauto. + eapply conv_cum_conv_ctx; eauto. 2: eapply conv_context_sym; eauto. eapply red_conv_cum_l ; try assumption. eapply red_zipp. @@ -3992,26 +3491,20 @@ Section Conversion. Next Obligation. apply unfold_one_fix_red in eq1 as r1. apply unfold_one_fix_decompose in eq1 as d1. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; - pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 - end. - rewrite <- eq3 in r2. cbn in r2. rewrite zipc_appstack in r2. cbn in r2. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. - rewrite <- eq3 in c2. cbn in c2. rewrite stack_context_appstack in c2. - cbn in c2. - case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (red_wellformed _ hΣ h1 r1) as hh. - apply red_context in r2. - pose proof (decompose_stack_eq _ _ _ (eq_sym eq2)). subst. - rewrite zipc_appstack in hh. cbn in r2. - pose proof (red_wellformed _ hΣ hh (sq r2)) as hh2. - rewrite zipc_stack_cat. - assumption. + reduce_stack_facts. + clear eqDiff. + simpl_stacks. + destruct r1 as [r1]. + eapply red_wellformed. + 1: eauto. + 1: exact h1. + constructor. + etransitivity. + 2: { apply red_zipc. + rewrite stack_context_decompose. + exact r. } + simpl_stacks. + eassumption. Qed. Next Obligation. apply unfold_one_fix_cored in eq1 as r1. @@ -4032,151 +3525,51 @@ Section Conversion. Qed. Next Obligation. apply unfold_one_fix_decompose in eq1 as d1. - rewrite <- eq2 in d1. simpl in d1. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. - case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ (eq_sym eq2)). subst. - apply decompose_stack_eq in e as ?. subst. - rewrite stack_cat_appstack. - rewrite stack_context_appstack. - case_eq (decompose_stack π1). intros args ρ e'. - simpl. - apply decompose_stack_eq in e'. subst. - clear eq3 eqDiff. - rewrite stack_context_appstack in hx. - assumption. + reduce_stack_facts. + clear eqDiff. + now simpl_stacks. Qed. Next Obligation. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; - pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as ir; - pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w2 - end. - rewrite <- eq3 in ir. - destruct ir as (?&hl). + simpl_reduce_stack. + specialize (isr eq_refl) as (?&_). split; [easy|]. - rewrite <- eq3 in w2, d2. - rewrite decompose_stack_appstack in d2. cbn in *. - rewrite zipp_stack_cat. - 2: eapply decompose_stack_noStackApp; eauto. - rewrite stack_context_stack_cat. - now rewrite app_context_assoc. + now simpl_stacks. Qed. Next Obligation. - destruct hΣ as [wΣ]. apply unfold_one_fix_red_zipp in eq1 as r1. - destruct r1 as [r1]. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in r2. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. apply unfold_one_fix_decompose in eq1 as d1. - assert (r2' : red (fst Σ) (Γ ,,, stack_context π1) (zipp fn θ) (zipp fn' (ρ +++ θ'))). - { unfold zipp. - case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. rewrite decompose_stack_appstack. - rewrite <- eq2. - cbn in r2. rewrite 2!zipc_appstack in r2. cbn in r2. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π1). intros l1 ρ1 e1. - simpl. rewrite e1 in r2. simpl in r2. - pose proof (decompose_stack_eq _ _ _ e1). subst. - rewrite decompose_stack_twice with (1 := e1). cbn. - rewrite app_nil_r. - rewrite stack_context_appstack. assumption. - } - pose proof (red_trans _ _ _ _ _ r1 r2') as r. - assert (e : stack_context π1 = stack_context (ρ +++ θ')). - { case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. - rewrite stack_context_appstack. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π1). intros l1 ρ1 e1. - pose proof (decompose_stack_eq _ _ _ e1). subst. - rewrite stack_context_appstack. reflexivity. - } - etransitivity. - - eapply red_conv_cum_l. all: eassumption. - - rewrite e. assumption. + reduce_stack_facts. + clear eqDiff. + simpl_reduce_stack. + destruct hΣ, r1 as [r1]. + eapply conv_cum_trans; auto. + 2: exact h. + eapply red_conv_cum_l. + etransitivity; eauto. Qed. Next Obligation. - destruct hΣ as [wΣ]. + apply h; clear h. apply unfold_one_fix_red_zipp in eq1 as r1. - destruct r1 as [r1]. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in r2. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. apply unfold_one_fix_decompose in eq1 as d1. - assert (r2' : red (fst Σ) (Γ ,,, stack_context π1) (zipp fn θ) (zipp fn' (ρ +++ θ'))). - { unfold zipp. - case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. rewrite decompose_stack_appstack. - rewrite <- eq2. - cbn in r2. rewrite 2!zipc_appstack in r2. cbn in r2. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π1). intros l1 ρ1 e1. - simpl. rewrite e1 in r2. simpl in r2. - pose proof (decompose_stack_eq _ _ _ e1). subst. - rewrite decompose_stack_twice with (1 := e1). cbn. - rewrite app_nil_r. - rewrite stack_context_appstack. assumption. - } - pose proof (red_trans _ _ _ _ _ r1 r2') as r. - assert (e : stack_context π1 = stack_context (ρ +++ θ')). - { case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. - rewrite stack_context_appstack. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π1). intros l1 ρ1 e1. - pose proof (decompose_stack_eq _ _ _ e1). subst. - rewrite stack_context_appstack. reflexivity. - } - apply h; clear h. - rewrite <- e. - eapply conv_cum_red_inv; eauto. - exact hΣ'. + reduce_stack_facts. + clear eqDiff. + simpl_reduce_stack. + destruct hΣ, r1 as [r1]. + eapply conv_cum_red_inv. + 1: auto. + 2: reflexivity. + 2: exact H. + etransitivity; eauto. Qed. Next Obligation. - cbn. rewrite zipc_appstack. cbn. apply unfold_one_fix_red_zipp in eq1 as r. apply unfold_one_fix_decompose in eq1 as d. - rewrite <- eq2 in d. simpl in d. - unfold zipp in r. - rewrite <- eq2 in r. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - rewrite e2 in r. - rewrite e2 in d. simpl in d. subst. - apply wellformed_zipc_zipp in h2 as hh2 ; auto. - pose proof (decompose_stack_eq _ _ _ e2). subst. - unfold zipp in hh2. rewrite e2 in hh2. - pose proof (red_wellformed _ hΣ hh2 r) as hh. - rewrite stack_context_appstack in hh. - assumption. + clear eqDiff aux eq1. + apply wellformed_zipc_zipp in h2; auto. + cbn in *. + simpl_stacks. + eapply red_wellformed; eauto. Qed. Next Obligation. apply unfold_one_fix_red in eq1 as r1. @@ -4242,128 +3635,38 @@ Section Conversion. assumption. Qed. Next Obligation. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; - pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as ir; - pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as w2 - end. - rewrite <- eq3 in d2, ir, w2. - rewrite decompose_stack_appstack in d2. - cbn in *. - destruct ir as (?&hl). + simpl_reduce_stack. + specialize (isr eq_refl) as (?&_). split; [easy|]. - rewrite stack_context_stack_cat. - rewrite zipp_stack_cat. - 2: eapply decompose_stack_noStackApp; eauto. - now rewrite app_context_assoc. + simpl_stacks. + eauto. Qed. Next Obligation. - destruct hΣ as [wΣ]. apply unfold_one_fix_red_zipp in eq1 as r1. - destruct r1 as [r1]. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in r2. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. apply unfold_one_fix_decompose in eq1 as d1. - assert (r2' : ∥ red (fst Σ) (Γ ,,, stack_context π2) (zipp fn θ) (zipp fn' (ρ +++ θ')) ∥). - { unfold zipp. - destruct hx as [hx]. - constructor. - case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. rewrite decompose_stack_appstack. - rewrite <- eq2. - cbn in r2. rewrite 2!zipc_appstack in r2. cbn in r2. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - simpl. rewrite e2 in r2. simpl in r2. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite decompose_stack_twice with (1 := e2). cbn. - rewrite app_nil_r. - rewrite stack_context_appstack. - assumption. - } - assert (e : stack_context π2 = stack_context (ρ +++ θ')). - { case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. - rewrite stack_context_appstack. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π2). intros l2 ρ2 e2. simpl. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite stack_context_appstack. reflexivity. - } - destruct r2' as [r2']. - destruct hx as [hx]. - pose proof (red_trans _ _ _ _ _ r1 r2') as r. - etransitivity ; revgoals. - - eapply conv_cum_context_convp. - + assumption. - + eapply red_conv_cum_r. 1: assumption. - eassumption. - + eapply conv_context_sym. 1: auto. - assumption. - - assumption. + reduce_stack_facts. + clear eqDiff. + simpl_reduce_stack. + destruct r1 as [r1]. + destruct hΣ, hx. + eapply conv_cum_red_conv. + 1,2,5: eauto. + 1: reflexivity. + etransitivity; eauto. Qed. Next Obligation. apply h; clear h. - destruct hΣ as [wΣ]. apply unfold_one_fix_red_zipp in eq1 as r1. - destruct r1 as [r1]. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in r2. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. apply unfold_one_fix_decompose in eq1 as d1. - assert (r2' : ∥ red (fst Σ) (Γ ,,, stack_context π2) (zipp fn θ) (zipp fn' (ρ +++ θ')) ∥). - { unfold zipp. - destruct hx as [hx]. - constructor. - case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. rewrite decompose_stack_appstack. - rewrite <- eq2. - cbn in r2. rewrite 2!zipc_appstack in r2. cbn in r2. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - simpl. rewrite e2 in r2. simpl in r2. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite decompose_stack_twice with (1 := e2). cbn. - rewrite app_nil_r. - rewrite stack_context_appstack. - assumption. - } - assert (e : stack_context π2 = stack_context (ρ +++ θ')). - { case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. - rewrite stack_context_appstack. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π2). intros l2 ρ2 e2. simpl. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite stack_context_appstack. reflexivity. - } - destruct r2' as [r2']. - destruct hx as [hx]. - pose proof (red_trans _ _ _ _ _ r1 r2') as r. - eapply conv_cum_conv_ctx; eauto. - 2: apply conv_context_sym; eauto. - eapply conv_cum_red_inv; [exact hΣ'| |reflexivity|exact r]. - eapply conv_cum_conv_ctx; eauto. + reduce_stack_facts. + clear eqDiff. + simpl_reduce_stack. + destruct r1 as [r1]. + destruct hΣ, hx. + eapply conv_cum_red_conv_inv. + 1,2,5: eauto. + 1: reflexivity. + etransitivity; eauto. Qed. Next Obligation. change (true = eqb idx idx') in eq4. @@ -4388,61 +3691,18 @@ Section Conversion. Qed. Next Obligation. apply h; clear h. - (* Idea : Since the fixpoints do not unfold they are convertible to fixpoints - in whnf implying that the args are convertible. *) - rewrite !zipp_as_mkApps in H. - apply unfold_one_fix_None in unfold_fix1 as [(?&?&?)]. - apply unfold_one_fix_None in unfold_fix2 as [(?&?&?)]. - destruct hΣ, hx. - eapply conv_cum_red_inv' in H. - 2: exact hΣ'. - 2: eassumption. - 2: apply red_mkApps; [reflexivity|exact a]. - 2: apply red_mkApps; [reflexivity|exact a0]. - apply conv_cum_mkApps_inv in H as [(_&conv_args)]; auto. - 2: eapply whnf_conv_context; eauto. - 2: eapply conv_context_sym; eauto. - constructor. - eapply conv_terms_red'; eauto. + eapply inv_stuck_fixes in H as [(<-&?&?)]; eauto. + constructor; auto. Qed. Next Obligation. apply h; clear h. - (* Idea : Since the fixpoints do not unfold they are convertible to fixpoints - in whnf implying that their definitions are convertible. *) - rewrite !zipp_as_mkApps in H. - apply unfold_one_fix_None in unfold_fix1 as [(?&?&?)]. - apply unfold_one_fix_None in unfold_fix2 as [(?&?&?)]. - destruct hΣ, hx. - eapply conv_cum_red_inv' in H. - 2: exact hΣ'. - 2: eassumption. - 2: apply red_mkApps; [reflexivity|exact a]. - 2: apply red_mkApps; [reflexivity|exact a0]. - apply conv_cum_mkApps_inv in H as [(?&?)]; auto. - 2: eapply whnf_conv_context; eauto. - 2: eapply conv_context_sym; eauto. - cbn in c0. - apply conv_cum_tFix_inv in c0 as [(->&?)]. - constructor. + eapply inv_stuck_fixes in H as [(<-&?&?)]; eauto. + constructor; auto. eapply All2_impl; eauto. cbn; intros; easy. Qed. Next Obligation. - (* Idea : Since the fixpoints do not unfold they are convertible to fixpoints - in whnf implying that the indices have to be equal. *) - rewrite !zipp_as_mkApps in H. - apply unfold_one_fix_None in unfold_fix1 as [(?&?&?)]. - apply unfold_one_fix_None in unfold_fix2 as [(?&?&?)]. - destruct hΣ, hx. - eapply conv_cum_red_inv' in H. - 2: exact hΣ'. - 2: eassumption. - 2: apply red_mkApps; [reflexivity|exact a]. - 2: apply red_mkApps; [reflexivity|exact a0]. - apply conv_cum_mkApps_inv in H as [(?&?)]; auto. - 2: eapply whnf_conv_context; eauto. - 2: eapply conv_context_sym; eauto. - apply conv_cum_tFix_inv in c0 as [(->&?)]. + eapply inv_stuck_fixes in H as [(<-&?&?)]; eauto. rewrite Nat.eqb_refl in idx_uneq. congruence. Qed. @@ -4467,7 +3727,6 @@ Section Conversion. exact todo_cofix_conversion. Qed. Next Obligation. - todo "Cofix". exact todo_cofix_conversion. Qed. @@ -4510,15 +3769,6 @@ Section Conversion. positionR (` (pps1 x)) (` (pps1 y)) -> Ret Reduction Γ u1 (coApp (mkApps t1 ca1) (appstack a1 π1)) u2 ρ2. - Lemma appstack_cons a args π : - appstack (a :: args) π = App a (appstack args π). - Proof. reflexivity. Qed. - - Lemma fst_decompose_stack_nil π : - isStackApp π = false -> - (decompose_stack π).1 = []. - Proof. now destruct π. Qed. - Equations(noeqns) _isconv_args' (leq : conv_pb) (Γ : context) (t1 : term) (args1 : list term) (l1 : list term) (π1 : stack) @@ -4536,30 +3786,30 @@ Section Conversion. | Success H1 with _isconv_args' leq Γ t1 (args1 ++ [u1]) l1 π1 _ _ (tApp t2 u2) l2 π2 _ _ _ _ := { | Success H2 := yes ; | Error e herr := - Error ( + no ( StackTailError leq (Γ ,,, stack_context π1) t1 args1 u1 l1 (Γ ,,, stack_context π2) t2 u2 l2 e - ) _ + ) } ; | Error e herr := - Error ( + no ( StackHeadError leq (Γ ,,, stack_context π1) t1 args1 u1 l1 (Γ ,,, stack_context π2) t2 u2 l2 e - ) _ + ) } ; _isconv_args' leq Γ t1 args1 [] π1 h1 hπ1 t2 [] π2 h2 hπ2 hx aux := yes ; _isconv_args' leq Γ t1 args1 l1 π1 h1 hπ1 t2 l2 π2 h2 hπ2 hx aux := - Error ( + no ( StackMismatch (Γ ,,, stack_context π1) t1 args1 l1 (Γ ,,, stack_context π2) t2 l2 - ) _. + ). Next Obligation. constructor; constructor. Defined. @@ -4633,7 +3883,7 @@ Section Conversion. | @exist (l1, θ1) eq1 with inspect (decompose_stack π2) := { | @exist (l2, θ2) eq2 with _isconv_args' leq Γ t1 [] l1 θ1 _ _ t2 l2 θ2 _ _ _ _ := { | Success h := yes ; - | Error e h := Error e _ + | Error e h := no e } } }. @@ -4712,19 +3962,7 @@ Section Conversion. Qed. Next Obligation. exfalso. - match type of eq with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; - pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as decomp - end. - rewrite <- eq in r, decomp. - cbn in *. - destruct (decompose_stack ρ) eqn:decomp'. - apply decompose_stack_eq in decomp' as ->. - cbn in *; subst. - rewrite zipc_appstack in r. - cbn in *. - clear eq. + simpl_reduce_stack. apply wellformed_nonarity in h as (?&typ); auto. destruct hΣ. apply inversion_Case in typ as (?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?); auto. @@ -4745,17 +3983,10 @@ Section Conversion. revert e. funelim (unfold_one_case Γ ind par p c brs h). all: intros eq ; noconf eq. - - match type of e with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; - pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as d - end. - rewrite <- e in r. - rewrite <- e in d. cbn in d. rewrite <- e0 in d. cbn in d. subst. - cbn in r. - clear H. symmetry in e0. apply decompose_stack_eq in e0. subst. - rewrite zipc_appstack in r. cbn in r. - assert (r' : ∥ red Σ Γ (tCase (ind, par) p c brs) (tCase (ind, par) p (mkApps (tConstruct ind0 n ui) l) brs) ∥). + - clear H H0. + simpl_reduce_stack. + assert (r' : ∥ red Σ Γ (tCase (ind, par) p c brs) + (tCase (ind, par) p (mkApps (tConstruct ind0 n ui) (decompose_stack s).1) brs) ∥). { constructor. eapply red_case_c. eassumption. } pose proof (red_wellformed _ hΣ h r') as h'. eapply Case_Construct_ind_eq in h' ; eauto. subst. @@ -4765,17 +3996,10 @@ Section Conversion. - match type of eq with | _ = False_rect _ ?f => destruct f end. - - match type of e with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; - pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as d - end. - rewrite <- e in r. - rewrite <- e in d. cbn in d. rewrite <- e1 in d. cbn in d. subst. - cbn in r. - clear H. symmetry in e1. apply decompose_stack_eq in e1. subst. - rewrite zipc_appstack in r. cbn in r. - assert (r' : ∥ red Σ Γ (tCase (ind, par) p c brs) (tCase (ind, par) p (mkApps (tCoFix mfix idx) l) brs) ∥). + - clear H H0 H1. + simpl_reduce_stack. + assert (r' : ∥ red Σ Γ (tCase (ind, par) p c brs) + (tCase (ind, par) p (mkApps (tCoFix mfix idx) (decompose_stack s).1) brs) ∥). { constructor. eapply red_case_c. eassumption. } pose proof (red_wellformed _ hΣ h r') as h'. eapply cored_red_cored. @@ -4788,25 +4012,11 @@ Section Conversion. ∥∑ c', red Σ Γ c c' × whne RedFlags.default Σ Γ c'∥. Proof. funelim (unfold_one_case Γ ind par p c brs h); intros [=]. - - match type of e with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; - pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as wh; - pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as isr; - pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as decomp - end. - rewrite <- e in r, decomp, wh, isr. - cbn in *. - constructor; exists (zipc t0 s). + - clear H. + simpl_reduce_stack. + constructor; exists (mkApps t0 (decompose_stack s).1). split; [easy|]. - destruct (decompose_stack s) eqn:decomp'. - apply decompose_stack_eq in decomp' as ->. - cbn in *; subst. - rewrite zipc_appstack in r |- *. - rewrite zipp_appstack, stack_context_appstack in wh. - cbn in *. - destruct isr as (noapp&_). - clear e H. + specialize (isr eq_refl) as (noapp&_). apply wellformed_nonarity in h as (?&typ); auto. destruct hΣ. eapply PCUICSR.subject_reduction in typ. @@ -4828,7 +4038,7 @@ Section Conversion. | cc0view_construct ind' ui with inspect (decompose_stack ρ) := { | @exist (args, ξ) eq' with inspect (nth_error args (pars + narg)) := { | @exist (Some arg) eq2 := Some arg ; - | @exist None _ := False_rect _ _ + | @exist None eq2 := False_rect _ _ } } ; | cc0view_cofix mfix idx with inspect (decompose_stack ρ) := { @@ -4850,65 +4060,18 @@ Section Conversion. left. eexists. eassumption. Qed. Next Obligation. - match type of eq with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; - pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as decomp - end. - rewrite <- eq in r, decomp. - cbn in *. - destruct (decompose_stack ρ) eqn:decomp'. - apply decompose_stack_eq in decomp' as ->. - noconf eq'. - cbn in *; subst. - rewrite zipc_appstack in r. - cbn in *. - clear eq. + simpl_reduce_stack. apply wellformed_nonarity in h as (?&typ); auto. destruct hΣ. - apply inversion_Proj in typ as (?&?&?&?&?&decl&?&?&?); auto. - assert (d := decl). - eapply PCUICSR.subject_reduction in t; eauto. - apply PCUICValidity.inversion_mkApps in t as spine; auto. - destruct spine as (?&typ_ctor&_). - apply inversion_Construct in typ_ctor as (?&?&?&?&decl_ctor&?&?); auto. - destruct hΣ as [wfΣ]. - eapply PCUICInductiveInversion.Construct_Ind_ind_eq with (wfΣ0 := wfΣ) (Hdecl := decl_ctor) in t. - cbn in t. - destruct decl_ctor as (?&?); cbn in *. - destruct All2_nth_error_Some; cbn in t. - destruct p0; cbn in t. - destruct t as ((((<-&?)&?)&?)&?). - clear s. - destruct decl as (?&?&?). - pose proof (PCUICInductiveInversion.declared_inductive_unique_sig d0 H) as H''; noconf H''. - cbn in *. - pose proof (PCUICWeakeningEnv.on_declared_projection wfΣ d) as (_&proj'). - destruct d. - pose proof (PCUICInductiveInversion.declared_inductive_unique_sig H d) as H''; noconf H''. - cbn in *. - destruct ind_cshapes; [easy|]. - destruct l; [|easy]. - noconf e1. - destruct proj' as ((_&?)&_). - symmetry in wildcard. - apply nth_error_None in wildcard. + eapply PCUICSR.subject_reduction in typ. + 2: auto. + 2: apply red_proj_c; eassumption. + apply PCUICInductiveInversion.invert_Proj_Construct in typ as (<-&_&?); auto. + apply eq_sym, nth_error_None in eq2. lia. Qed. Next Obligation. - match type of eq with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; - pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as decomp - end. - rewrite <- eq in r, decomp. - cbn in *. - destruct (decompose_stack ρ) eqn:decomp'. - apply decompose_stack_eq in decomp' as ->. - cbn in *; subst. - rewrite zipc_appstack in r. - cbn in *. - clear eq. + simpl_reduce_stack. apply wellformed_nonarity in h as (?&typ); auto. destruct hΣ. apply inversion_Proj in typ as (?&?&?&?&?&?&?&?&?); auto. @@ -4929,35 +4092,19 @@ Section Conversion. revert e. funelim (unfold_one_proj Γ p c h). all: intros eq ; noconf eq. - - match type of e with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; - pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as d - end. - rewrite <- e in r. - rewrite <- e in d. cbn in d. rewrite <- e0 in d. cbn in d. subst. - cbn in r. - clear H0. symmetry in e0. apply decompose_stack_eq in e0. subst. - rewrite zipc_appstack in r. cbn in r. + - clear H H0 H1. + simpl_reduce_stack. pose proof (red_proj_c (i, n0, n) _ _ r) as r'. pose proof (red_wellformed _ hΣ h (sq r')) as h'. - apply Proj_Constuct_ind_eq in h' ; auto. subst. + apply Proj_Construct_ind_eq in h' ; auto. subst. eapply cored_red_cored. + constructor. eapply red_proj. eauto. + eapply red_proj_c. eassumption. - match type of eq with | _ = False_rect _ ?f => destruct f end. - - match type of e with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; - pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as d - end. - rewrite <- e in r. - rewrite <- e in d. cbn in d. rewrite <- e0 in d. cbn in d. subst. - cbn in r. - clear H0. symmetry in e0. apply decompose_stack_eq in e0. subst. - rewrite zipc_appstack in r. cbn in r. + - clear H H0 H1. + simpl_reduce_stack. pose proof (red_proj_c (i, n0, n) _ _ r) as r'. pose proof (red_wellformed _ hΣ h (sq r')) as h'. eapply cored_red_cored. @@ -4973,25 +4120,11 @@ Section Conversion. ∥∑ c', red Σ Γ c c' × whne RedFlags.default Σ Γ c'∥. Proof. funelim (unfold_one_proj Γ p c h); intros [=]. - - match type of e with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_sound f Σ hΣ Γ t π h) as [r] ; - pose proof (reduce_stack_whnf f Σ hΣ Γ t π h) as wh; - pose proof (reduce_stack_isred f Σ hΣ Γ t π h eq_refl) as isr; - pose proof (reduce_stack_decompose f Σ hΣ Γ t π h) as decomp - end. - rewrite <- e in r, decomp, wh, isr. - cbn in *. - constructor; exists (zipc t0 s). + - clear H. + simpl_reduce_stack. + constructor; exists (mkApps t0 (decompose_stack s).1). split; [easy|]. - destruct (decompose_stack s) eqn:decomp'. - apply decompose_stack_eq in decomp' as ->. - cbn in *; subst. - rewrite zipc_appstack in r |- *. - rewrite zipp_appstack, stack_context_appstack in wh. - cbn in *. - destruct isr as (noapp&_). - clear e H. + specialize (isr eq_refl) as (noapp&_). apply wellformed_nonarity in h as (?&typ); auto. destruct hΣ. eapply PCUICSR.subject_reduction in typ. @@ -5052,8 +4185,7 @@ Section Conversion. funelim (reducible_head Γ t π h). all: intro ee ; noconf ee. - eapply unfold_one_fix_red_zipp. eassumption. - - constructor. unfold zipp. - case_eq (decompose_stack π). intros l s eq. + - constructor. simpl_stacks. eapply red_mkApps_f. eapply trans_red. + reflexivity. @@ -5062,17 +4194,13 @@ Section Conversion. * reflexivity. - apply unfold_one_case_cored in e as r. apply cored_red in r. destruct r as [r]. - constructor. unfold zipp. - case_eq (decompose_stack π). intros l s e'. + constructor. simpl_stacks. eapply red_mkApps_f. - apply decompose_stack_eq in e'. subst. assumption. - apply unfold_one_proj_cored in e as r. apply cored_red in r. destruct r as [r]. - constructor. unfold zipp. - case_eq (decompose_stack π). intros l s e'. + constructor. simpl_stacks. eapply red_mkApps_f. - apply decompose_stack_eq in e'. subst. assumption. Qed. @@ -5273,59 +4401,31 @@ Section Conversion. | @exist None nored2 with inspect (eqb_termp_napp leq #|(decompose_stack π1).1| t1 t2) := { | @exist true eq1 := isconv_args leq t1 π1 t2 π2 aux; | @exist false noteq := - Error ( + no ( HeadMismatch leq (Γ ,,, stack_context π1) t1 (Γ ,,, stack_context π2) t2 - ) _ + ) } } }. Next Obligation. - cbn. rewrite zipc_appstack. cbn. apply reducible_head_red_zipp in eq1 as r. apply reducible_head_decompose in eq1 as d. - rewrite <- eq2 in d. simpl in d. - unfold zipp in r. - rewrite <- eq2 in r. - case_eq (decompose_stack π1). intros l1' ρ1' e1. - rewrite e1 in r. - rewrite e1 in d. simpl in d. subst. apply wellformed_zipc_zipp in h1 as hh1. 2: auto. - apply decompose_stack_eq in e1 as ?. subst. - unfold zipp in hh1. rewrite e1 in hh1. + cbn in *. + simpl_stacks. pose proof (red_wellformed _ hΣ hh1 r) as hh. - symmetry in eq2. - apply decompose_stack_eq in eq2. subst. - rewrite stack_context_appstack. - rewrite stack_context_appstack in hh. assumption. Qed. Next Obligation. apply reducible_head_cored in eq1 as r1. apply cored_red in r1. destruct r1 as [r1]. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in r2. cbn in r2. - rewrite <- eq3 in d2. cbn in d2. - rewrite decompose_stack_appstack in d2. cbn in d2. - rewrite zipc_appstack in r2. cbn in r2. - case_eq (decompose_stack θ1'). intros l s e. - rewrite e in d2. cbn in d2. subst. - apply decompose_stack_eq in e as ?. subst. - rewrite stack_cat_appstack. - rewrite zipc_appstack in r2. cbn in r2. - rewrite zipc_appstack. - pose proof (eq_sym eq2) as eq2'. - apply decompose_stack_eq in eq2'. subst. - rewrite stack_context_appstack in r2. + simpl_reduce_stack. eapply red_wellformed ; auto ; revgoals. - - constructor. zip fold. eapply red_context. eassumption. - - rewrite zipc_appstack in r1. cbn. + - constructor. zip fold. eapply red_context. simpl_stacks. eassumption. + - cbn. simpl_stacks. eapply red_wellformed ; auto ; revgoals. + constructor. eassumption. + assumption. @@ -5360,134 +4460,42 @@ Section Conversion. Qed. Next Obligation. apply reducible_head_decompose in eq1 as d1. - rewrite <- eq2 in d1. cbn in d1. - case_eq (decompose_stack π1). intros l' s' e'. - rewrite e' in d1. cbn in d1. subst. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in d2. cbn in d2. - rewrite decompose_stack_appstack in d2. cbn in d2. - case_eq (decompose_stack θ1'). intros l s e. - rewrite e in d2. cbn in d2. subst. - apply decompose_stack_eq in e as ?. subst. - apply decompose_stack_eq in e' as ?. subst. - rewrite stack_cat_appstack. - rewrite stack_context_appstack. - clear eq3. - rewrite stack_context_appstack in hx. + simpl_reduce_stack. assumption. Qed. Next Obligation. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_isred RedFlags.nodelta _ hΣ _ _ _ h eq_refl) as ir; - pose proof (reduce_stack_whnf RedFlags.nodelta _ hΣ _ _ _ h) as wh - end. - rewrite <- eq3 in ir, wh. - destruct ir as (?&_). + simpl_reduce_stack. + specialize (isr eq_refl) as (?&_). split; auto. - rewrite stack_context_stack_cat. - rewrite <- (stack_context_decompose ρ1) in wh. - rewrite <- eq2 in wh. - rewrite zipp_stack_cat; [rewrite app_context_assoc; auto|]. - eapply decompose_stack_noStackApp; eauto. + now simpl_stacks. Qed. Next Obligation. - destruct hΣ as [wΣ]. apply reducible_head_red_zipp in eq1 as r1. destruct r1 as [r1]. apply reducible_head_decompose in eq1 as d1. - rewrite <- eq2 in d1. cbn in d1. - case_eq (decompose_stack π1). intros l' s' e'. - rewrite e' in d1. cbn in d1. subst. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in r2. cbn in r2. - rewrite <- eq3 in d2. cbn in d2. - rewrite decompose_stack_appstack in d2. cbn in d2. - rewrite zipc_appstack in r2. cbn in r2. - case_eq (decompose_stack θ1'). intros l s e. - rewrite e in d2. cbn in d2. subst. - apply decompose_stack_eq in e as ?. subst. - apply decompose_stack_eq in e' as ?. subst. - rewrite stack_context_appstack. - rewrite stack_cat_appstack in h. - rewrite stack_context_appstack in h. + simpl_reduce_stack. + destruct hΣ as [wΣ]. etransitivity ; try eassumption. - unfold zipp. rewrite e'. - unfold zipp in r1. rewrite e' in r1. rewrite <- eq2 in r1. - rewrite decompose_stack_appstack. - erewrite decompose_stack_twice ; eauto. simpl. - rewrite app_nil_r. eapply red_conv_cum_l ; try assumption. - rewrite stack_context_appstack in r1. eapply red_trans ; try eassumption. - clear eq3. symmetry in eq2. apply decompose_stack_eq in eq2. subst. - rewrite stack_context_appstack in r2. - rewrite zipc_appstack in r2. cbn in r2. - assumption. Qed. Next Obligation. apply h; clear h. - (* Contrapositive of previous case *) - destruct hΣ as [wΣ]. apply reducible_head_red_zipp in eq1 as r1. destruct r1 as [r1]. apply reducible_head_decompose in eq1 as d1. - rewrite <- eq2 in d1. cbn in d1. - case_eq (decompose_stack π1). intros l' s' e'. - rewrite e' in d1. cbn in d1. subst. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in r2. cbn in r2. - rewrite <- eq3 in d2. cbn in d2. - rewrite decompose_stack_appstack in d2. cbn in d2. - rewrite zipc_appstack in r2. cbn in r2. - case_eq (decompose_stack θ1'). intros l s e. - rewrite e in d2. cbn in d2. subst. - apply decompose_stack_eq in e as ?. subst. - apply decompose_stack_eq in e' as ?. subst. - rewrite stack_context_appstack in H. - rewrite stack_cat_appstack. - rewrite stack_context_appstack. + simpl_reduce_stack. + destruct hΣ as [wΣ]. etransitivity ; try eassumption. - unfold zipp. rewrite e'. - unfold zipp in r1. rewrite e' in r1. rewrite <- eq2 in r1. - rewrite decompose_stack_appstack. - erewrite decompose_stack_twice ; eauto. simpl. - rewrite app_nil_r. eapply red_conv_cum_r ; try assumption. - rewrite stack_context_appstack in r1. - eapply red_trans ; try eassumption. - clear eq3. symmetry in eq2. apply decompose_stack_eq in eq2. subst. - rewrite stack_context_appstack in r2. - rewrite zipc_appstack in r2. cbn in r2. - assumption. + eapply red_trans ; eassumption. Qed. Next Obligation. - cbn. rewrite zipc_appstack. cbn. apply reducible_head_red_zipp in eq1 as r. apply reducible_head_decompose in eq1 as d. - rewrite <- eq2 in d. - unfold zipp in r. - rewrite <- eq2 in r. - case_eq (decompose_stack π2). intros l2' ρ2' e2. - rewrite e2 in r. - rewrite e2 in d. simpl in d. subst. + cbn. + simpl_stacks. apply wellformed_zipc_zipp in h2 as hh2. 2: auto. - apply decompose_stack_eq in e2 as ?. subst. - unfold zipp in hh2. rewrite e2 in hh2. + simpl_stacks. pose proof (red_wellformed _ hΣ hh2 r) as hh. - symmetry in eq2. - apply decompose_stack_eq in eq2. subst. - rewrite stack_context_appstack. - rewrite stack_context_appstack in hh. assumption. Qed. Next Obligation. @@ -5548,80 +4556,25 @@ Section Conversion. Qed. Next Obligation. apply reducible_head_decompose in eq1 as d1. - rewrite <- eq2 in d1. cbn in d1. - case_eq (decompose_stack π2). intros l' s' e'. - rewrite e' in d1. cbn in d1. subst. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in d2. cbn in d2. - rewrite decompose_stack_appstack in d2. cbn in d2. - case_eq (decompose_stack θ2'). intros l s e. - rewrite e in d2. cbn in d2. subst. - apply decompose_stack_eq in e as ?. subst. - apply decompose_stack_eq in e' as ?. subst. - rewrite stack_cat_appstack. - rewrite stack_context_appstack. - clear eq3. - rewrite stack_context_appstack in hx. - assumption. + now simpl_reduce_stack. Qed. Next Obligation. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_isred RedFlags.nodelta _ hΣ _ _ _ h eq_refl) as ir; - pose proof (reduce_stack_whnf RedFlags.nodelta _ hΣ _ _ _ h) as wh - end. - rewrite <- eq3 in ir, wh. - destruct ir as (?&_). + simpl_reduce_stack. + specialize (isr eq_refl) as (?&_). split; auto. - rewrite stack_context_stack_cat. - rewrite <- (stack_context_decompose ρ2) in wh. - rewrite <- eq2 in wh. - rewrite zipp_stack_cat; [rewrite app_context_assoc; auto|]. - eapply decompose_stack_noStackApp; eauto. + now simpl_stacks. Qed. Next Obligation. destruct hΣ as [wΣ]. apply reducible_head_red_zipp in eq1 as r1. destruct r1 as [r1]. apply reducible_head_decompose in eq1 as d1. - rewrite <- eq2 in d1. cbn in d1. - case_eq (decompose_stack π2). intros l' s' e'. - rewrite e' in d1. cbn in d1. subst. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in r2. cbn in r2. - rewrite <- eq3 in d2. cbn in d2. - rewrite decompose_stack_appstack in d2. cbn in d2. - rewrite zipc_appstack in r2. cbn in r2. - case_eq (decompose_stack θ2'). intros l s e. - rewrite e in d2. cbn in d2. subst. - unfold zipp. rewrite e'. - unfold zipp in r1. rewrite e' in r1. rewrite <- eq2 in r1. - apply decompose_stack_eq in e as ?. subst. - apply decompose_stack_eq in e' as ?. subst. - clear eq3. - rewrite stack_context_appstack in hx. + simpl_reduce_stack. destruct hx as [hx]. etransitivity ; try eassumption. - unfold zipp. - rewrite stack_cat_appstack. - rewrite decompose_stack_appstack. - erewrite decompose_stack_twice ; eauto. simpl. - rewrite app_nil_r. - eapply conv_cum_context_convp. + eapply conv_cum_conv_ctx. - assumption. - eapply red_conv_cum_r. 1: assumption. - rewrite stack_context_appstack in r1. - eapply red_trans ; try eassumption. - symmetry in eq2. apply decompose_stack_eq in eq2. subst. - rewrite stack_context_appstack in r2. - rewrite zipc_appstack in r2. cbn in r2. - assumption. + eapply red_trans ; eassumption. - eapply conv_context_sym. all: auto. Qed. Next Obligation. @@ -5630,42 +4583,13 @@ Section Conversion. destruct hΣ as [wΣ]. apply reducible_head_red_zipp in eq1 as r1. destruct r1 as [r1]. apply reducible_head_decompose in eq1 as d1. - rewrite <- eq2 in d1. cbn in d1. - case_eq (decompose_stack π2). intros l' s' e'. - rewrite e' in d1. cbn in d1. subst. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in r2. cbn in r2. - rewrite <- eq3 in d2. cbn in d2. - rewrite decompose_stack_appstack in d2. cbn in d2. - rewrite zipc_appstack in r2. cbn in r2. - case_eq (decompose_stack θ2'). intros l s e. - rewrite e in d2. cbn in d2. subst. - unfold zipp in r1. rewrite e' in r1. - unfold zipp in H. rewrite e' in H. rewrite <- eq2 in r1. - apply decompose_stack_eq in e as ?. subst. - apply decompose_stack_eq in e' as ?. subst. - clear eq3. - rewrite stack_context_appstack in hx. + simpl_reduce_stack. destruct hx as [hx]. etransitivity ; try eassumption. - unfold zipp. - rewrite stack_cat_appstack. - rewrite decompose_stack_appstack. - erewrite decompose_stack_twice ; eauto. simpl. - rewrite app_nil_r. - eapply conv_cum_context_convp. + eapply conv_cum_conv_ctx. - assumption. - eapply red_conv_cum_l. - rewrite stack_context_appstack in r1. - eapply red_trans ; try eassumption. - symmetry in eq2. apply decompose_stack_eq in eq2. subst. - rewrite stack_context_appstack in r2. - rewrite zipc_appstack in r2. cbn in r2. - assumption. + eapply red_trans ; eassumption. - eapply conv_context_sym. all: auto. Qed. Next Obligation. @@ -5698,8 +4622,8 @@ Section Conversion. destruct H'. apply whnf_red_red in s2 as H'. destruct H'. - eapply conv_cum_red_inv' in H. - 2: exact hΣ'. + eapply conv_cum_red_conv_inv in H. + 2: assumption. 2: eassumption. 2: apply red_mkApps; [eassumption|eassumption]. 2: apply red_mkApps; [eassumption|eassumption]. @@ -5709,9 +4633,8 @@ Section Conversion. 2: eapply whnf_conv_context; eauto. 2: eapply conv_context_sym; eauto. constructor. - eapply conv_terms_red'; eauto. + eapply conv_terms_red_conv; eauto. Qed. - Next Obligation. unfold eqb_termp_napp in noteq. destruct ir1 as (notapp1&whδ1), ir2 as (notapp2&whδ2). @@ -5723,22 +4646,22 @@ Section Conversion. destruct H'. apply whnf_red_red in s2 as H'. destruct H'. - eapply conv_cum_red_inv' in H. - 2: exact hΣ'. + eapply conv_cum_red_conv_inv in H. + 2: assumption. 2: eassumption. 2: apply red_mkApps; eassumption. 2: apply red_mkApps; eassumption. apply conv_cum_mkApps_inv in H as [(conv_hds&_)]; auto. - 2: now depelim s1. - 2: now depelim s2. + 2: now inversion s1; subst. + 2: now inversion s2; subst. 2: eapply whnf_conv_context; eauto. 2: eapply conv_context_sym; eauto. - apply whnf_mkApps_inv in wh1; [|now depelim s1]. + apply whnf_mkApps_inv in wh1. destruct t1; cbn in *. - all: depelim s1. + all: inversion s1; subst; clear s1. 9: { destruct conv_hds as [H]. - depelim H. - depelim s2. + inversion H; subst; clear H. + inversion s2; subst; clear s2. zip fold in h1. zip fold in h2. apply wellformed_context in h1; auto. @@ -5750,15 +4673,15 @@ Section Conversion. apply inversion_Ind in typ2 as (?&?&?&?&?&?); auto. apply consistent_instance_ext_all_mem in c1. apply consistent_instance_ext_all_mem in c. - apply compare_global_instance_complete in r; auto. + apply compare_global_instance_complete in H3; auto. rewrite eq_inductive_refl in noteq. apply All2_length in rargs1. - rewrite <- rargs1 in r. + rewrite <- rargs1 in H3. cbn in *. easy. } 9: { destruct conv_hds as [H]. - depelim H. - depelim s2. + inversion H; subst; clear H. + inversion s2; subst; clear s2. zip fold in h1. zip fold in h2. apply wellformed_context in h1; auto. @@ -5770,17 +4693,17 @@ Section Conversion. apply inversion_Construct in typ2 as (?&?&?&?&?&?&?); auto. apply consistent_instance_ext_all_mem in c1. apply consistent_instance_ext_all_mem in c. - apply compare_global_instance_complete in r; auto. + apply compare_global_instance_complete in H4; auto. rewrite eq_inductive_refl, Nat.eqb_refl in noteq. apply All2_length in rargs1. - rewrite <- rargs1 in r. + rewrite <- rargs1 in H4. cbn in *. easy. } all: apply conv_cum_alt in conv_hds as [(?&?&(r1&r2)&?)]. all: eapply whnf_red_inv in r1; auto. all: inversion r1; subst; clear r1. - all: inversion e; subst; clear e. - all: apply whnf_mkApps_inv in wh2; [|now depelim s2]. + all: inversion e0; subst; clear e0. + all: apply whnf_mkApps_inv in wh2. all: eapply whnf_conv_context in wh2; [|apply conv_context_sym; eauto]. all: eapply whnf_red_inv in r2; auto. all: inversion r2; subst; clear r2. @@ -5792,7 +4715,7 @@ Section Conversion. apply wellformed_context in h1; auto. destruct h1 as [(?&typ)|[(?&?&?&?)]]. + now apply inversion_Evar in typ. - + cbn in e; congruence. + + cbn in e0; congruence. - zip fold in h1. zip fold in h2. apply wellformed_context in h1; auto. diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index 8fcb72d4f..745310780 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -865,7 +865,7 @@ Section Reduce. constructor. cbn in hh. rewrite zipc_appstack in hh. cbn in hh. zip fold in hh. apply wellformed_context in hh. 2: assumption. - simpl in hh. apply Proj_Constuct_ind_eq in hh. all: eauto. + simpl in hh. apply Proj_Construct_ind_eq in hh. all: eauto. subst. constructor. eauto. Qed. Next Obligation. @@ -1364,6 +1364,7 @@ Section Reduce. Proof. intros shape wh typ. destruct hΣ as [wfΣ]. + assert (typ' := typ). apply inversion_Proj in typ as (?&?&?&?&?&?&?&?&?); auto. cbn in *. eapply whnf_non_ctor_non_cofix_ind_typed; try eassumption. @@ -1371,7 +1372,8 @@ Section Reduce. rewrite decompose_app_mkApps by (now destruct hd). cbn. destruct hd; try easy. - eapply PCUICInductiveInversion.projected_constructor_eq in d; [|eassumption|eassumption]. + destruct p as ((?&?)&?). + eapply PCUICInductiveInversion.invert_Proj_Construct in typ' as (->&->&?); auto. congruence. - unfold isCoFix_app. now rewrite decompose_app_mkApps; destruct hd. diff --git a/template-coq/theories/common/uGraph.v b/template-coq/theories/common/uGraph.v index 2814d1601..9badf1385 100644 --- a/template-coq/theories/common/uGraph.v +++ b/template-coq/theories/common/uGraph.v @@ -1721,7 +1721,7 @@ Section CheckLeq2. exact I. Qed. - Lemma leq_universe_spec' u1 u2 : + Lemma check_leqb_universe_complete u1 u2 : levels_declared u1 -> levels_declared u2 -> leq_universe uctx.2 u1 u2 -> @@ -1757,7 +1757,7 @@ Section CheckLeq2. (destruct (gc_of_constraints uctx.2); [cbn in *|contradiction HG]); tas. Qed. - Lemma eq_universe_spec' u1 u2 : + Lemma check_eqb_universe_complete u1 u2 : levels_declared u1 -> levels_declared u2 -> eq_universe uctx.2 u1 u2 -> From 1aedda6a8284a5e767c8ec2e5b589dd570e39495 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Tue, 27 Oct 2020 17:17:34 +0100 Subject: [PATCH 18/26] Move whnf, whne, whnf_red into Type and strengthen some lemmas whnf_red_mkApps_inv is now based on whnf_red instead of whnf + red and shows whnf_red for the heads instead of plain (and weak) red. --- erasure/theories/ErasureFunction.v | 4 +- pcuic/theories/PCUICCanonicity.v | 8 +- pcuic/theories/PCUICConvCumInversion.v | 77 ++- pcuic/theories/PCUICNormal.v | 565 ++++++++++----------- pcuic/theories/PCUICSafeLemmata.v | 6 +- safechecker/theories/PCUICSafeChecker.v | 17 +- safechecker/theories/PCUICSafeConversion.v | 158 +++--- safechecker/theories/PCUICSafeReduce.v | 74 +-- template-coq/theories/utils/All_Forall.v | 13 +- 9 files changed, 431 insertions(+), 491 deletions(-) diff --git a/erasure/theories/ErasureFunction.v b/erasure/theories/ErasureFunction.v index cad9def3a..28b820296 100644 --- a/erasure/theories/ErasureFunction.v +++ b/erasure/theories/ErasureFunction.v @@ -127,7 +127,7 @@ Defined. Next Obligation. pose proof (hnf_sound HΣ (h := h)). repeat match goal with [H : squash (red _ _ _ _ ) |- _ ] => destruct H end. - destruct HΣ. + destruct H, HΣ. eapply PCUICConfluence.red_confluence in X0 as [t'' []]. 3:exact X1. 2:eauto. eapply invert_red_sort in r0; eauto. subst. eapply whnf_red_inv in r; auto. @@ -152,7 +152,7 @@ Defined. Next Obligation. pose proof (hnf_sound HΣ (h := h)). repeat match goal with [H : squash (red _ _ _ _ ) |- _ ] => destruct H end. - destruct HΣ. + destruct H, HΣ. eapply PCUICConfluence.red_confluence in X2 as [t'' []]. 3:exact X3. 2:eauto. eapply invert_red_prod in r0 as (? & ? & [] & ?); eauto. subst. eapply whnf_red_inv in r; auto. diff --git a/pcuic/theories/PCUICCanonicity.v b/pcuic/theories/PCUICCanonicity.v index e837b143a..df2c1f1d0 100644 --- a/pcuic/theories/PCUICCanonicity.v +++ b/pcuic/theories/PCUICCanonicity.v @@ -1030,7 +1030,7 @@ Section WeakNormalization. - clear wh_neutral_empty_gen wh_normal_empty_gen. subst. apply inversion_Const in typed as [decl' [wfd [declc [cu cum]]]]; eauto. specialize (axfree _ _ declc). - red in declc. rewrite declc in H. noconf H. congruence. + red in declc. rewrite declc in e. noconf e. congruence. - simpl in cl; move/andP: cl => [clf cla]. eapply inversion_App in typed as [na [A [B [Hf _]]]]; eauto. - specialize (wh_neutral_empty_gen _ _ axfree ne). subst. @@ -1039,15 +1039,15 @@ Section WeakNormalization. eapply inversion_Fix in t as (? & ? & ? & ? & ? & ? & ?); auto. eapply typing_spine_strengthen in t0; eauto. eapply nth_error_all in a; eauto. simpl in a. - rewrite /unfold_fix in H. rewrite e in H. noconf H. + rewrite /unfold_fix in e. rewrite e1 in e. noconf e. eapply (wf_fixpoint_spine wfΣ) in t0; eauto. - rewrite H0 in t0. destruct t0 as [ind [u [indargs [tyarg ckind]]]]. + rewrite e0 in t0. destruct t0 as [ind [u [indargs [tyarg ckind]]]]. clear wh_normal_empty_gen. now specialize (wh_neutral_empty_gen _ tyarg eq_refl). - move/andP: cl => [/andP[_ clc] _]. eapply inversion_Case in typed; firstorder eauto. - eapply inversion_Proj in typed; firstorder auto. - - eapply wh_neutral_empty_gen in H; eauto. + - eapply wh_neutral_empty_gen in w; eauto. - eapply inversion_Sort in typed as (? & ? & ? & ? & ?); auto. eapply invert_cumul_sort_l in c as (? & ? & ?); auto. eapply red_mkApps_tInd in r as (? & eq & ?); eauto; eauto. diff --git a/pcuic/theories/PCUICConvCumInversion.v b/pcuic/theories/PCUICConvCumInversion.v index e0f66a799..ecc8c5a69 100644 --- a/pcuic/theories/PCUICConvCumInversion.v +++ b/pcuic/theories/PCUICConvCumInversion.v @@ -52,16 +52,13 @@ Section fixed. Qed. Lemma whnf_red_isIndConstructApp Γ t t' : - whnf RedFlags.default Σ Γ t -> - red Σ Γ t t' -> + whnf_red Σ Γ t t' -> isIndConstructApp t' = isIndConstructApp t. Proof. - intros wh r. - eapply whnf_red_inv in r; eauto. + intros r. induction r; auto. rewrite (isIndConstructApp_mkApps _ [arg']), (isIndConstructApp_mkApps _ [arg]). apply IHr. - now apply whnf_mkApps_inv with (l := [arg]) in wh. Qed. Lemma eq_termp_mkApps_inv leq v args v' args' : @@ -91,24 +88,14 @@ Section fixed. Proof. intros conv notapp notapp' wh wh'. apply conv_cum_alt in conv as [(?&?&(r1&r2)&e)]. - apply whnf_red_mkApps_l in r1 as [(?&?&->&?&?)]; auto. - apply whnf_red_mkApps_l in r2 as [(?&?&->&?&?)]; auto. - assert (isApp x1 = false). - { erewrite whnf_red_isApp. - 3: eauto. - 1: assumption. - apply whnf_mkApps_inv in wh; auto. } - assert (isApp x = false). - { erewrite whnf_red_isApp. - 3: eauto. - 1: assumption. - apply whnf_mkApps_inv in wh'; auto. } - apply eq_termp_mkApps_inv in e as (?&?); auto. + apply whnf_red_inv, whnf_red_mkApps_l_inv in r1 as (?&?&->&?&?); auto. + apply whnf_red_inv, whnf_red_mkApps_l_inv in r2 as (?&?&->&?&?); auto. + apply whnf_red_isApp in w as ?. + apply whnf_red_isApp in w0 as ?. + apply eq_termp_mkApps_inv in e as (?&?); try congruence. constructor. split. - - assert (isIndConstructApp x1 = isIndConstructApp hd). - { eapply whnf_red_isIndConstructApp; eauto. - apply whnf_mkApps_inv in wh; auto. } + - apply whnf_red_isIndConstructApp in w as ?. destruct hd. all: cbn. 1-9, 12-15: apply conv_cum_alt. @@ -116,14 +103,10 @@ Section fixed. 1-13: exists x1, x. 1-13: split; [split|]; eauto with pcuic. 1-13: (eapply eq_term_upto_univ_napp_nonind; [exact e|try exact H1]). - 1: discriminate notapp. - all: apply whnf_mkApps_inv in wh; auto. - all: eapply whnf_red_inv in r; auto. - all: depelim r. - all: apply whnf_mkApps_inv in wh'; auto. - all: eapply whnf_red_inv in r0; auto. + 1-13: cbn in *; congruence. + all: depelim w. all: depelim e. - all: depelim r0. + all: depelim w0. all: apply All2_length in a. all: constructor; constructor; rewrite a; auto. - clear -a1 a a0. @@ -145,27 +128,27 @@ Section fixed. Proof. intros conv whl whr. depelim whl; solve_discr. - depelim H; solve_discr; try discriminate. + depelim w; solve_discr; try discriminate. depelim whr; solve_discr. - depelim H0; solve_discr; try discriminate. - apply conv_cum_alt in conv as [(?&?&(r1&r2)&?)]. + depelim w0; solve_discr; try discriminate. + apply conv_cum_alt in conv as [(?&?&(r1&r2)&eq)]. eapply whnf_red_inv in r1; eauto. eapply whnf_red_inv in r2; eauto. depelim r1. depelim r2. - depelim e. + depelim eq. constructor. split; [easy|]. split; [apply conv_alt_red; now exists motive'0, motive'1|]. split; [apply conv_alt_red; now exists discr'0, discr'1|]. - clear -a X1 X4. - induction a in brs, brs', brs'0, brs'1, X1, X4, a |- *; - depelim X1; depelim X4; [now constructor|]. - constructor. - + destruct p, p0, r. - split; [congruence|]. - apply conv_alt_red; now exists x.2, y.2. - + now apply IHa. + clear -a a0 a1. + induction a in brs, brs', brs'0, brs'1, a0, a1, a |- *; + depelim a0; depelim a1; [now constructor|]. + constructor; eauto. + destruct p, p0, r. + split; [congruence|]. + apply conv_alt_red. + eauto. Qed. Lemma conv_cum_tFix_inv leq Γ mfix idx mfix' idx' : @@ -177,7 +160,7 @@ Section fixed. mfix mfix'∥. Proof. intros conv. - apply conv_cum_alt in conv as [(?&?&(r1&r2)&?)]. + apply conv_cum_alt in conv as [(?&?&(r1&r2)&eq)]. assert (forall defs i, whnf RedFlags.default Σ Γ (tFix defs i)). { intros defs i. apply whnf_fixapp with (v := []). @@ -187,13 +170,13 @@ Section fixed. eapply whnf_red_inv in r2; eauto. depelim r1. depelim r2. - depelim e. + depelim eq. constructor. split; [easy|]. - clear -a X X0. + clear -a a0 a1. cut (#|mfix| = #|mfix'|); - [|now apply All2_length in a; apply All2_length in X; apply All2_length in X0]. - revert a X X0. + [|now apply All2_length in a; apply All2_length in a0; apply All2_length in a1]. + revert a a0 a1. generalize mfix at 1 3 4. generalize mfix' at 1 3. intros ctx_fix ctx_fix'. @@ -201,7 +184,7 @@ Section fixed. induction all in mfix, mfix', mfix'0, mfix'1, all1, all2, all |- *; depelim all1; depelim all2; [constructor|]. constructor; [|now auto]. - destruct p as ((?&?)&?), p0 as (?&?&?&?), r as (?&?&?&?). + destruct r as ((?&?)&?), p as (?&?&?&?), p0 as (?&?&?&?). split; [congruence|]. split; [now apply conv_alt_red; exists (dtype x), (dtype y)|]. apply conv_alt_red. @@ -227,6 +210,6 @@ Section fixed. depelim e. constructor. split; [easy|]. - now apply conv_alt_red; exists c'0, c'1. + apply conv_alt_red; eauto. Qed. End fixed. diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index 3e74ee0e0..8666fd8ec 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -38,7 +38,7 @@ Section Normal. Context (Σ : global_env). (* Relative to reduction flags *) - Inductive whnf (Γ : context) : term -> Prop := + Inductive whnf (Γ : context) : term -> Type := | whnf_ne t : whne Γ t -> whnf Γ t | whnf_sort s : whnf Γ (tSort s) | whnf_prod na A B : whnf Γ (tProd na A B) @@ -53,7 +53,7 @@ Section Normal. whnf Γ (mkApps (tFix mfix idx) v) | whnf_cofixapp mfix idx v : whnf Γ (mkApps (tCoFix mfix idx) v) - with whne (Γ : context) : term -> Prop := + with whne (Γ : context) : term -> Type := | whne_rel i : option_map decl_body (nth_error Γ i) = Some None -> whne Γ (tRel i) @@ -139,7 +139,12 @@ Section Normal. forall Γ t l, negb (isApp t) -> whne Γ (mkApps t l) -> - whne Γ t \/ exists mfix idx narg body a, t = tFix mfix idx /\ unfold_fix mfix idx = Some (narg, body) /\ nth_error l narg = Some a /\ whne Γ a. + whne Γ t + + ∑ mfix idx narg body a, + t = tFix mfix idx × + unfold_fix mfix idx = Some (narg, body) × + nth_error l narg = Some a × + whne Γ a. Proof. intros Γ t l Ha h. revert t Ha h. induction l using rev_ind; intros. @@ -207,28 +212,28 @@ Proof. revert hdeq. induction h in h, t, l, t0 |- *; intros eq; subst. - destruct (mkApps_elim t l). - apply whne_mkApps_inv in H; auto. - destruct H as [wh|(?&?&?&?&?&->&?&?&?)]. + apply whne_mkApps_inv in w; auto. + destruct w as [wh|(?&?&?&?&?&->&?&?&?)]. + now apply whnf_mkApps. - + destruct (Nat.leb_spec n x1). + + destruct (leb_spec_Set n x1). * apply whnf_fixapp. - rewrite H. - apply nth_error_Some_length in H0. + rewrite e. + apply nth_error_Some_length in e0. apply nth_error_None. rewrite firstn_length. lia. * apply whnf_ne. eapply whne_fixapp; eauto. - assert (x1 < #|l|) by (now apply nth_error_Some_length in H0). - rewrite <- (firstn_skipn n l) in H0. - rewrite nth_error_app1 in H0; auto. + assert (x1 < #|l|) by (now apply nth_error_Some_length in e0). + rewrite <- (firstn_skipn n l) in e0. + rewrite nth_error_app1 in e0; auto. rewrite firstn_length. lia. - - destruct l using List.rev_ind; [|now rewrite <- mkApps_nested in eq]. + - destruct l using MCList.rev_ind; [|now rewrite <- mkApps_nested in eq]. cbn in *; subst; auto. - - destruct l using List.rev_ind; [|now rewrite <- mkApps_nested in eq]. + - destruct l using MCList.rev_ind; [|now rewrite <- mkApps_nested in eq]. cbn in *; subst; auto. - - destruct l using List.rev_ind; [|now rewrite <- mkApps_nested in eq]. + - destruct l using MCList.rev_ind; [|now rewrite <- mkApps_nested in eq]. cbn in *; subst; auto. - destruct (mkApps_elim t l). apply mkApps_eq_inj in eq as (<-&<-); auto. @@ -239,7 +244,7 @@ Proof. apply whnf_fixapp. destruct unfold_fix as [(?&?)|]; [|easy]. apply nth_error_None. - apply nth_error_None in H. + apply nth_error_None in y. rewrite firstn_length. lia. - destruct (mkApps_elim t l). @@ -247,9 +252,9 @@ Proof. Qed. Lemma whnf_fixapp' {flags} Σ Γ mfix idx narg body v : -unfold_fix mfix idx = Some (narg, body) -> -nth_error v narg = None -> -whnf flags Σ Γ (mkApps (tFix mfix idx) v). + unfold_fix mfix idx = Some (narg, body) -> + nth_error v narg = None -> + whnf flags Σ Γ (mkApps (tFix mfix idx) v). Proof. intros E1 H. eapply whnf_fixapp. rewrite E1. eauto. Qed. @@ -265,19 +270,19 @@ Proof. - now inv whn; solve_discr. - inv whn. + easy. - + destruct v0 as [|? ? _] using List.rev_ind; [discriminate|]. + + destruct v0 as [|? ? _] using MCList.rev_ind; [discriminate|]. rewrite <- mkApps_nested in *. cbn in *; inv H0. constructor. eapply IHwhe. apply whnf_cstrapp. - + destruct v0 as [|? ? _] using List.rev_ind; [discriminate|]. + + destruct v0 as [|? ? _] using MCList.rev_ind; [discriminate|]. rewrite <- mkApps_nested in *. cbn in *; inv H0. constructor. eapply IHwhe. apply whnf_indapp. - + destruct v0 as [|? ? _] using List.rev_ind; [discriminate|]. + + destruct v0 as [|? ? _] using MCList.rev_ind; [discriminate|]. rewrite <- mkApps_nested in *. cbn in *; inv H. constructor. @@ -289,162 +294,158 @@ Proof. apply nth_error_None in H0. rewrite app_length in H0; cbn in *. lia. - + destruct v0 as [|? ? _] using List.rev_ind; [discriminate|]. + + destruct v0 as [|? ? _] using MCList.rev_ind; [discriminate|]. rewrite <- mkApps_nested in *. cbn in *; inv H0. constructor. eapply IHwhe. apply whnf_cofixapp. - inv whn; solve_discr; try easy. - rewrite H in H2. + rewrite e in H0. congruence. - inv whn; solve_discr; easy. - inv whn; solve_discr; easy. Qed. -Definition whnf_whne_dec flags Σ Γ t : ({whnf flags Σ Γ t} + {~ (whnf flags Σ Γ t)}) * ({whne flags Σ Γ t} + {~ (whne flags Σ Γ t)}). -Proof. - induction t using term_forall_mkApps_ind in Γ |- *; split; eauto. - all: try now (right; intros H; depelim H;help). - - destruct (RedFlags.zeta flags) eqn:Er. - + destruct (option_map decl_body (nth_error Γ n)) as [ [ | ] | ] eqn:E. - * right. intros H. depelim H. depelim H. congruence. congruence. all:help. - * eauto. - * right. intros H. depelim H. depelim H. congruence. congruence. all:help. - + eauto. +Definition whnf_whne_dec flags Σ Γ t : + ({∥whnf flags Σ Γ t∥} + {~∥whnf flags Σ Γ t∥}) * + ({∥whne flags Σ Γ t∥} + {~∥whne flags Σ Γ t∥}). +Proof with eauto using sq with pcuic. + induction t using term_forall_mkApps_ind in Γ |- *; split... + all: try now (right; intros [H]; depelim H;help). - destruct (RedFlags.zeta flags) eqn:Er. - + destruct (option_map decl_body (nth_error Γ n)) as [ [ | ] | ] eqn:E. - * right. intros H. depelim H. congruence. congruence. help. - * eauto. - * right. intros H. depelim H. congruence. congruence. help. - + eauto. - - destruct (RedFlags.beta flags) eqn:Er; eauto. - right. intros ?. depelim H. congruence. help. - - destruct (RedFlags.zeta flags) eqn:Er; eauto. - right. intros ?. depelim H. depelim H. all:help. congruence. - - destruct (RedFlags.zeta flags) eqn:Er; eauto. - right. intros ?. depelim H. congruence. help. + + destruct (option_map decl_body (nth_error Γ n)) as [ [ | ] | ] eqn:E... + * right. intros [H]. depelim H. depelim w. congruence. congruence. all:help. + * right. intros [H]. depelim H. depelim w. congruence. congruence. all:help. + + eauto using sq. + - destruct (RedFlags.zeta flags) eqn:Er... + destruct (option_map decl_body (nth_error Γ n)) as [ [ | ] | ] eqn:E... + + right. intros [H]. depelim H. congruence. congruence. help. + + right. intros [H]. depelim H. congruence. congruence. help. + - destruct (RedFlags.beta flags) eqn:Er... + right. intros [H]. depelim H. congruence. help. + - destruct (RedFlags.zeta flags) eqn:Er... + right. intros [H]. depelim H. depelim w. all:help. congruence. + - destruct (RedFlags.zeta flags) eqn:Er... + right. intros [H]. depelim H. congruence. help. - destruct (IHt Γ) as [[] _]. + destruct t. all:try now (left; eauto using whnf_mkApps, All_Forall). - all: try now left; eapply whnf_mkApps; depelim w; eauto; help. - * destruct v as [ | ? v]. - -- eauto. - -- right. intros ?. depelim H0. depelim H0. all:help. clear IHl. - eapply whne_mkApps_inv in H0 as []; eauto. - ++ depelim H0. help. + all: try now left; destruct s as [w]; constructor; eapply whnf_mkApps; depelim w; eauto; help. + * destruct v as [ | ? v]... + right. intros [w]. depelim w. depelim w. all:help. clear IHt. + eapply whne_mkApps_inv in w as []... + -- depelim w. help. + -- firstorder congruence. + * destruct v as [ | ? v]... + right. intros [w]. depelim w. depelim w. all:help. clear IHt. + eapply whne_mkApps_inv in w as []... + -- depelim w. help. + -- firstorder congruence. + * destruct v as [ | ? v]... + destruct (RedFlags.beta flags) eqn:?. + -- right. intros [w]. depelim w. depelim w. all:help. clear IHl. + eapply whne_mkApps_inv in w as []... + ++ depelim w. congruence. help. ++ firstorder congruence. - * destruct v as [ | ? v]. - -- eauto. - -- right. intros ?. depelim H0. depelim H0. all:help. clear IHl. - eapply whne_mkApps_inv in H0 as []; eauto. - ++ depelim H0. help. - ++ firstorder congruence. - * destruct v as [ | ? v]. - -- eauto. - -- destruct (RedFlags.beta flags) eqn:?. - ++ right. intros ?. depelim H0. depelim H0. all:help. clear IHl. - eapply whne_mkApps_inv in H0 as []; eauto. - ** depelim H0. congruence. help. - ** firstorder congruence. - ++ left. - apply whnf_mkApps. - constructor. - assumption. + -- left; constructor. + apply whnf_mkApps. + constructor. + assumption. * destruct (unfold_fix mfix idx) as [(narg, body) |] eqn:E1. - -- destruct (nth_error v narg) as [a | ] eqn:E2. - ++ destruct (nth_error_all E2 X Γ) as [_ []]. - ** left. eauto. - ** destruct (RedFlags.fix_ flags) eqn:?. - --- right. intros ?. depelim H0. depelim H0. all:help. clear IHv. - eapply whne_mkApps_inv in H0 as []; eauto. - +++ depelim H0; [|congruence]. help. - eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. - inv x. destruct rarg; inv H1. - +++ destruct H0 as (? & ? & ? & ? & ? & ? & ? & ? & ?). inv H0. - rewrite E1 in H1. inv H1. - eapply (nth_error_app_left v [x0]) in H2. - rewrite E2 in H2. inv H2. eauto. - +++ help. - eapply (f_equal decompose_app) in x; - rewrite !decompose_app_mkApps in x; - cbn in *; try firstorder congruence. - inv x. rewrite E1 in H0. congruence. - --- left. - apply whnf_mkApps. - constructor. - assumption. - ++ left. eauto. - -- left. eapply whnf_fixapp. rewrite E1. eauto. - + right. intros ?. eapply n. now eapply whnf_mkApps_inv. + -- destruct (nth_error v narg) as [a | ] eqn:E2... + destruct (nth_error_all E2 X Γ) as [_ []]. + ++ left. destruct s0... + ++ destruct (RedFlags.fix_ flags) eqn:?. + ** right. intros [w]. depelim w. depelim w. all:help. clear IHv. + eapply whne_mkApps_inv in w as []; eauto. + --- depelim w; [|congruence]. help. + eapply (f_equal decompose_app) in x; + rewrite !decompose_app_mkApps in x; cbn in *; try firstorder congruence. + inv x. destruct rarg; inv e0. + --- destruct s0 as (? & ? & ? & ? & ? & ? & ? & ? & ?). inv e. + rewrite E1 in e0. inv e0. + eapply (nth_error_app_left v [x0]) in e1. + rewrite E2 in e1. inv e1... + --- help. + eapply (f_equal decompose_app) in x; + rewrite !decompose_app_mkApps in x; + cbn in *; try firstorder congruence. + inv x. rewrite E1 in y. congruence. + ** left. + constructor. + apply whnf_mkApps. + constructor. + assumption. + -- left. constructor. eapply whnf_fixapp. rewrite E1. eauto. + + right. intros [w]. eapply n. constructor. now eapply whnf_mkApps_inv. - destruct (IHt Γ) as [_ []]. - + left. now eapply whne_mkApps. + + left. destruct s as [w]. constructor. now eapply whne_mkApps. + destruct t. - all: try now (right; intros [ | (mfix & idx & narg & body & a & [=] & ? & ? & ?) ] % whne_mkApps_inv; subst; cbn; eauto). + all: try (right; intros [[ | (mfix & idx & narg & body & a & [=] & ? & ? & ?) ] % whne_mkApps_inv]; subst)... * destruct (unfold_fix mfix idx) as [(narg, body) |] eqn:E1. -- destruct (nth_error v narg) as [a | ] eqn:E2. ++ destruct (nth_error_all E2 X Γ) as [_ []]. - ** left. eauto. + ** left. destruct s. constructor. eauto. ** destruct (RedFlags.fix_ flags) eqn:?. - --- right. intros ?. depelim H0. depelim H0. all:help. clear IHv. - eapply whne_mkApps_inv in H0 as []; eauto. - destruct H0 as (? & ? & ? & ? & ? & ? & ? & ? & ?). inv H0. - rewrite E1 in H1. inv H1. - eapply (nth_error_app_left v [x0]) in H2. - rewrite E2 in H2. inv H2. eauto. - --- left. apply whne_mkApps. constructor. assumption. - ++ right. intros [ | (mfix' & idx' & narg' & body' & a' & [=] & ? & ? & ?) ] % whne_mkApps_inv; subst; cbn; eauto. + --- right. intros ?. depelim H0. depelim X0. all:help. clear IHv. + eapply whne_mkApps_inv in X0 as []... + destruct s as (? & ? & ? & ? & ? & ? & ? & ? & ?). inv e. + rewrite E1 in e0. inv e0. + eapply (nth_error_app_left v [x0]) in e1. + rewrite E2 in e1. inv e1... + --- left. constructor. apply whne_mkApps. constructor. assumption. + ++ right. intros [[ | (mfix' & idx' & narg' & body' & a' & [=] & ? & ? & ?) ] % whne_mkApps_inv]; subst; cbn... congruence. - -- right. intros [ | (mfix' & idx' & narg' & body' & a' & [=] & ? & ? & ?) ] % whne_mkApps_inv; subst; cbn; eauto. + -- right. intros [[ | (mfix' & idx' & narg' & body' & a' & [=] & ? & ? & ?) ] % whne_mkApps_inv]; subst; cbn... congruence. - * right. intros [ | (mfix' & idx' & narg' & body' & a' & [=] & ? & ? & ?) ] % whne_mkApps_inv; subst; cbn; eauto. - - destruct (RedFlags.delta flags) eqn:Er; eauto. + * right. intros [[ | (mfix' & idx' & narg' & body' & a' & [=] & ? & ? & ?) ] % whne_mkApps_inv]; subst; cbn... + - destruct (RedFlags.delta flags) eqn:Er... destruct (lookup_env Σ s) as [[] | ] eqn:E. - + destruct (cst_body c) eqn:E2. - * right. intros H. depelim H. depelim H. congruence. congruence. all:help. - * eauto. - + right. intros H. depelim H. depelim H. congruence. congruence. all:help. - + right. intros H. depelim H. depelim H. congruence. congruence. all:help. - - destruct (RedFlags.delta flags) eqn:Er; eauto. + + destruct (cst_body c) eqn:E2... + right. intros [w]. depelim w. depelim w. congruence. congruence. all:help. + + right. intros [w]. depelim w. depelim w. congruence. congruence. all:help. + + right. intros [w]. depelim w. depelim w. congruence. congruence. all:help. + - destruct (RedFlags.delta flags) eqn:Er... destruct (lookup_env Σ s) as [[] | ] eqn:E. - + destruct (cst_body c) eqn:E2. - * right. intros H. depelim H. congruence. congruence. help. - * eauto. - + right. intros H. depelim H. congruence. congruence. help. - + right. intros H. depelim H. congruence. congruence. help. - - left. eapply whnf_indapp with (v := []). - - left. eapply whnf_cstrapp with (v := []). - - destruct (RedFlags.iota flags) eqn:Eq; eauto. + + destruct (cst_body c) eqn:E2... + right. intros [w]. depelim w. congruence. congruence. help. + + right. intros [w]. depelim w. congruence. congruence. help. + + right. intros [w]. depelim w. congruence. congruence. help. + - left. constructor. eapply whnf_indapp with (v := []). + - left. constructor. eapply whnf_cstrapp with (v := []). + - destruct (RedFlags.iota flags) eqn:Eq... destruct (IHt2 Γ) as [_ []]. - + eauto. - + right. intros ?. depelim H. depelim H. all:help. eauto. congruence. - - destruct (RedFlags.iota flags) eqn:Eq; eauto. + + left. destruct s... + + right. intros [w]. depelim w. depelim w. all:help... + - destruct (RedFlags.iota flags) eqn:Eq... destruct (IHt2 Γ) as [_ []]. - + eauto. - + right. intros ?. depelim H. all:help. eauto. congruence. - - destruct (RedFlags.iota flags) eqn:Eq; eauto. + + left. destruct s... + + right. intros [w]. depelim w. all:help... + - destruct (RedFlags.iota flags) eqn:Eq... destruct (IHt Γ) as [_ []]. - + eauto. - + right. intros H. depelim H. depelim H. all:eauto. all:help. - - destruct (RedFlags.iota flags) eqn:Eq; eauto. + + left. destruct s0... + + right. intros [w]. depelim w. depelim w. all:help... + - destruct (RedFlags.iota flags) eqn:Eq... destruct (IHt Γ) as [_ []]. - + eauto. - + right. intros H. depelim H. all:help. eauto. congruence. + + left. destruct s0... + + right. intros [w]. depelim w. all:help... - destruct (unfold_fix m n) as [(narg, body) |] eqn:E1. - + left. eapply whnf_fixapp with (v := []). rewrite E1. now destruct narg. - + left. eapply whnf_fixapp with (v := []). now rewrite E1. - - destruct (RedFlags.fix_ flags) eqn:?; eauto. - right. intros ?. depelim H; [|congruence]. - destruct args using rev_ind; try rewrite mkApps_snoc in H3; inv x. - destruct rarg; inv H0. rewrite mkApps_snoc in H3. inv H3. - - left. eapply whnf_cofixapp with (v := []). + + left. constructor. eapply whnf_fixapp with (v := []). rewrite E1. now destruct narg. + + left. constructor. eapply whnf_fixapp with (v := []). now rewrite E1. + - destruct (RedFlags.fix_ flags) eqn:?... + right. intros [w]. depelim w; [|congruence]. + destruct args using rev_ind. + + destruct rarg; inv e0. + + rewrite mkApps_snoc in x. inv x. + - left. constructor. eapply whnf_cofixapp with (v := []). Defined. Definition whnf_dec flags Σ Γ t := fst (whnf_whne_dec flags Σ Γ t). Definition whne_dec flags Σ Γ t := snd (whnf_whne_dec flags Σ Γ t). Lemma red1_mkApps_tConstruct_inv Σ Γ i n u v t' : - red1 Σ Γ (mkApps (tConstruct i n u) v) t' -> ∑ v', (t' = mkApps (tConstruct i n u) v') * (OnOne2 (red1 Σ Γ) v v'). + red1 Σ Γ (mkApps (tConstruct i n u) v) t' -> + ∑ v', (t' = mkApps (tConstruct i n u) v') * (OnOne2 (red1 Σ Γ) v v'). Proof. revert t'. induction v using rev_ind; intros. - cbn in *. depelim X. help. @@ -459,7 +460,8 @@ Proof. Qed. Lemma red1_mkApps_tInd_inv Σ Γ i u v t' : - red1 Σ Γ (mkApps (tInd i u) v) t' -> ∑ v', (t' = mkApps (tInd i u) v') * (OnOne2 (red1 Σ Γ) v v'). + red1 Σ Γ (mkApps (tInd i u) v) t' -> + ∑ v', (t' = mkApps (tInd i u) v') * (OnOne2 (red1 Σ Γ) v v'). Proof. revert t'. induction v using rev_ind; intros. - cbn in *. depelim X. help. @@ -473,20 +475,6 @@ Proof. eapply OnOne2_app. econstructor. eauto. Qed. -Ltac hhelp' := - try - repeat match goal with - | [ H0 : _ = mkApps _ _ |- _ ] => - eapply (f_equal decompose_app) in H0; - rewrite !decompose_app_mkApps in H0; cbn in *; firstorder congruence - | [ H1 : tApp _ _ = mkApps _ ?args |- _ ] => - destruct args eqn:E using rev_ind ; cbn in *; - [ inv H1 | rewrite <- mkApps_nested in H1; cbn in *; inv H1] - end. - -Ltac hhelp := hhelp'; - try match goal with | [ H0 : mkApps _ _ = _ |- _ ] => symmetry in H0 end; hhelp'. - Lemma is_constructor_app_false i x y : is_constructor i (x ++ y) = false -> is_constructor i x = false. @@ -572,19 +560,18 @@ Proof. depelim wh. solve_discr. Qed. - Lemma whnf_mkApps_tSort_inv f Σ Γ s args : whnf f Σ Γ (mkApps (tSort s) args) -> args = []. Proof. intros wh. inversion wh; solve_discr. - clear -H. + clear -X. remember (mkApps (tSort s) args) eqn:teq. exfalso. revert teq. - induction H in args |- *; intros; solve_discr. - destruct args as [|? ? _] using List.rev_ind; [easy|]. + induction X in args |- *; intros; solve_discr. + destruct args as [|? ? _] using MCList.rev_ind; [easy|]. rewrite <- mkApps_nested in teq. cbn in teq. inv teq. @@ -596,11 +583,11 @@ Lemma whnf_mkApps_tProd_inv f Σ Γ na A B args : Proof. intros wh. inversion wh; solve_discr. - clear -H. + clear -X. remember (mkApps (tProd na A B) args) eqn:teq. exfalso. revert teq. - induction H in args |- *; intros; solve_discr. + induction X in args |- *; intros; solve_discr. destruct args as [|? ? _] using List.rev_ind; [easy|]. rewrite <- mkApps_nested in teq. cbn in teq. @@ -614,11 +601,11 @@ Lemma whnf_mkApps_tLambda_inv f Σ Γ na A b args : Proof. intros nobeta wh. inversion wh; solve_discr. - clear -H nobeta. + clear -X nobeta. remember (mkApps (tLambda na A b) args) eqn:teq. exfalso. revert teq. - induction H in args |- *; intros; solve_discr. + induction X in args |- *; intros; solve_discr. destruct args as [|? ? _] using List.rev_ind; [easy|]. rewrite <- mkApps_nested in teq. cbn in teq. @@ -630,7 +617,7 @@ Section whne_red1_ind. Context (flags : RedFlags.t). Context (Σ : global_env). Context (Γ : context). -Context (P : term -> term -> Prop). +Context (P : term -> term -> Type). Lemma whne_red1_ind (Hrel : forall i body, @@ -765,66 +752,63 @@ Proof. - eauto. - depelim r; solve_discr. unfold declared_constant in isdecl. - rewrite H in isdecl. + rewrite e in isdecl. inv isdecl. congruence. - eauto. - depelim r; eauto. + depelim wh; solve_discr. now apply Hbeta_nobeta. - + destruct args as [|? ? _] using List.rev_ind; [now solve_discr|]. + + destruct args as [|? ? _] using MCList.rev_ind; [now solve_discr|]. rewrite <- mkApps_nested in x. cbn in *. inv x. apply whne_mkApps_inv in wh; [|easy]. destruct wh as [|]. - * depelim H. + * depelim w. -- solve_discr. - rewrite H in e. - inv e. - rewrite nth_error_nil in H0; congruence. - -- change (tApp ?hd ?a) with (mkApps hd [a]). - rewrite mkApps_nested. + rewrite nth_error_nil in e0; congruence. + -- rewrite <- mkApps_snoc. now eapply Hfix_nofix. - * destruct H as (?&?&?&?&?&?&?&?&?). - inv H. - rewrite H0 in e. + * destruct s as (?&?&?&?&?&?&?&?&?). + inv e1. + rewrite e2 in e. inv e. unfold is_constructor in e0. erewrite nth_error_app_left in e0 by eassumption. - now apply whne_isConstruct_app in H2; auto. + now apply whne_isConstruct_app in w; auto. - eapply red1_mkApps_tFix_inv in r; eauto. - 2: { rewrite H. + 2: { rewrite e. unfold is_constructor. - rewrite H0. + rewrite e0. destruct isConstruct_app eqn:ctor; [|easy]. now eapply whne_isConstruct_app in ctor. } destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]; eauto. - depelim r; eauto. - destruct args using List.rev_ind; [|rewrite <- mkApps_nested in x; cbn in x; discriminate]. + destruct args using MCList.rev_ind; [|rewrite <- mkApps_nested in x; cbn in x; discriminate]. cbn in *. - unfold is_constructor in e0. - rewrite nth_error_nil in e0; discriminate. + unfold is_constructor in e1. + rewrite nth_error_nil in e1; discriminate. - depelim r; eauto. + apply whne_mkApps_inv in wh; [|easy]. destruct wh as [|(?&?&?&?&?&?&?)]; [|discriminate]. - depelim H. + depelim w. solve_discr. + solve_discr. + apply whne_mkApps_inv in wh; [|easy]. destruct wh as [|(?&?&?&?&?&?&?)]; [|discriminate]. - depelim H. + depelim w. solve_discr. - eauto. - depelim r; eauto. + solve_discr. + apply whne_mkApps_inv in wh; [|easy]. destruct wh as [|(?&?&?&?&?&?&?)]; [|discriminate]. - depelim H. + depelim w. solve_discr. + apply whne_mkApps_inv in wh; [|easy]. destruct wh as [|(?&?&?&?&?&?&?)]; [|discriminate]. - depelim H. + depelim w. solve_discr. - depelim r; eauto. solve_discr. @@ -842,7 +826,7 @@ Proof. - eapply OnOne2_nth_error in H0; eauto. destruct H0 as (?&?&[->|]). + eapply whne_fixapp; eauto. - + apply H2 in r0. + + apply X1 in r0. eapply whne_fixapp; eauto. - unfold unfold_fix in *. destruct (nth_error mfix idx) eqn:nth; [|easy]. @@ -872,7 +856,7 @@ Proof. unfold unfold_fix. rewrite e. inv e0. - rewrite H4. + rewrite H3. reflexivity. Qed. @@ -880,10 +864,7 @@ Lemma whne_pres Σ Γ t t' : red Σ Γ t t' -> whne RedFlags.default Σ Γ t -> whne RedFlags.default Σ Γ t'. Proof. - induction 1; intros. - - eapply whne_pres1; eauto. - - eauto. - - eauto. + induction 1 using red_rect_n1; eauto using whne_pres1. Qed. Lemma whnf_pres1 Σ Γ t t' : @@ -908,20 +889,20 @@ Proof. - eapply red1_mkApps_tFix_inv in r; eauto. 2: { destruct unfold_fix as [(?&?)|]; [|easy]. unfold is_constructor. - now rewrite H. } + now rewrite y. } destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]. + apply whnf_fixapp. apply OnOne2_length in o. destruct unfold_fix as [(?&?)|]; [|easy]. apply nth_error_None. - apply nth_error_None in H. + apply nth_error_None in y. lia. + apply whnf_fixapp. unfold unfold_fix in *. destruct (nth_error x idx) eqn:nth; [|easy]. eapply OnOne2_nth_error_r in nth; eauto. destruct nth as (?&nth&?). - rewrite nth in H. + rewrite nth in y. destruct s as [->|(?&?)]; [easy|]. now inv e. + apply whnf_fixapp. @@ -929,7 +910,7 @@ Proof. destruct (nth_error x idx) eqn:nth; [|easy]. eapply OnOne2_nth_error_r in nth; eauto. destruct nth as (?&nth&?). - rewrite nth in H. + rewrite nth in y. destruct s as [->|(?&?)]; [easy|]. now inv e. - eapply red1_mkApps_tCoFix_inv in r as [[(?&->&?)|(?&->&?)]|(?&->&?)]; eauto. @@ -939,17 +920,14 @@ Lemma whnf_pres Σ Γ t t' : red Σ Γ t t' -> whnf RedFlags.default Σ Γ t -> whnf RedFlags.default Σ Γ t'. Proof. - induction 1; intros. - - eapply whnf_pres1; eauto. - - eauto. - - eauto. + induction 1 using red_rect_n1; eauto using whnf_pres1. Qed. Hint Resolve All2_same All2_firstn All2_skipn OnOne2_All2 red_mkApps All2_app : pcuic. (* For terms in whnf we have a very useful inversion lemma for reductions. This next relation is a subrelation of red that classifies how whnf terms reduce. *) -Inductive whnf_red Σ Γ : term -> term -> Prop := +Inductive whnf_red Σ Γ : term -> term -> Type := | whnf_red_tRel i : option_map decl_body (nth_error Γ i) = Some None -> whnf_red Σ Γ (tRel i) (tRel i) @@ -962,7 +940,6 @@ Inductive whnf_red Σ Γ : term -> term -> Prop := cst_body decl = None -> whnf_red Σ Γ (tConst kn u) (tConst kn u) | whnf_red_tApp hd hd' arg arg' : - red Σ Γ hd hd' -> whnf_red Σ Γ hd hd' -> red Σ Γ arg arg' -> whnf_red Σ Γ (tApp hd arg) (tApp hd' arg') @@ -1004,12 +981,14 @@ Inductive whnf_red Σ Γ : term -> term -> Prop := Derive Signature for whnf_red. +Hint Constructors whnf_red : pcuic. + Lemma whnf_red_red Σ Γ t t' : whnf_red Σ Γ t t' -> - ∥red Σ Γ t t'∥. + red Σ Γ t t'. Proof. intros wh. - depelim wh; constructor; eauto with pcuic. + induction wh; eauto with pcuic. - apply red_evar; auto. - apply red_app; auto. - apply red_fix_congr. @@ -1036,6 +1015,8 @@ Proof. eauto. Qed. +Hint Resolve whnf_red_red : pcuic. + Lemma whnf_red_mkApps Σ Γ hd hd' args args' : whnf_red Σ Γ hd hd' -> red Σ Γ hd hd' -> @@ -1043,11 +1024,64 @@ Lemma whnf_red_mkApps Σ Γ hd hd' args args' : whnf_red Σ Γ (mkApps hd args) (mkApps hd' args'). Proof. intros shape r ra. - induction ra using All2_ind_rev; auto. + induction ra using All2_rect_rev; auto. rewrite <- !mkApps_nested. cbn. constructor; auto. - apply red_mkApps; auto. +Qed. + +Hint Resolve whnf_red_mkApps : pcuic. + +Lemma whnf_red_mkApps_l_inv Σ Γ hd args t : + whnf_red Σ Γ (mkApps hd args) t -> + ∑ hd' args', + t = mkApps hd' args' × whnf_red Σ Γ hd hd' × All2 (red Σ Γ) args args'. +Proof. + intros r. + revert hd t r. + induction args using MCList.rev_ind; intros hd t r. + - cbn in *. + eexists _, []; split; [reflexivity|]. + eauto with pcuic. + - rewrite <- mkApps_nested in r. + cbn in r. + depelim r. + apply IHargs in r as (?&?&->&?&?); auto. + exists x0, (x1 ++ [arg']). + rewrite <- mkApps_nested. + eauto with pcuic. +Qed. + +Lemma whnf_red_mkApps_r_inv Σ Γ t hd' args' : + whnf_red Σ Γ t (mkApps hd' args') -> + ∑ hd args, + t = mkApps hd args × whnf_red Σ Γ hd hd' × All2 (red Σ Γ) args args'. +Proof. + intros r. + revert hd' t r. + induction args' using MCList.rev_ind; intros hd' t r. + - cbn in *. + eexists _, []; split; [reflexivity|]. + eauto with pcuic. + - rewrite <- mkApps_nested in r. + cbn in r. + depelim r. + apply IHargs' in r as (?&?&->&?&?); auto. + exists x0, (x1 ++ [arg]). + rewrite <- mkApps_nested. + eauto with pcuic. +Qed. + +Lemma whnf_red_mkApps_inv Σ Γ hd args hd' args' : + whnf_red Σ Γ (mkApps hd args) (mkApps hd' args') -> + isApp hd = false -> + isApp hd' = false -> + whnf_red Σ Γ hd hd' × All2 (red Σ Γ) args args'. +Proof. + intros r notapp notapp'. + apply whnf_red_mkApps_l_inv in r as (?&?&eq&?&?); auto. + apply mkApps_notApp_inj in eq as (->&->); auto. + now depelim w. Qed. Lemma whnf_red_refl_whne Σ Γ t : @@ -1067,14 +1101,8 @@ Lemma whnf_red_refl Σ Γ t : whnf_red Σ Γ t t. Proof. intros wh. - induction wh; cbn in *; try discriminate; eauto using whnf_red with pcuic. + induction wh; cbn in *; try discriminate; eauto with pcuic. - eapply whnf_red_refl_whne; eauto. - - apply whnf_red_mkApps; auto. - 2: apply All2_same; auto. - constructor. - - apply whnf_red_mkApps; auto. - 2: apply All2_same; auto. - constructor. - apply whnf_red_mkApps; auto. 2: apply All2_same; auto. constructor. @@ -1085,7 +1113,7 @@ Proof. apply All2_same; auto. Qed. -Instance whnf_red_trans Σ Γ : RelationClasses.Transitive (whnf_red Σ Γ). +Instance whnf_red_trans Σ Γ : CRelationClasses.Transitive (whnf_red Σ Γ). Proof. intros x y z xy yz. revert z yz. @@ -1106,7 +1134,7 @@ Proof. do 3 (split; auto). eapply context_change_decl_types_red; eauto. apply fix_context_change_decl_types. - now apply All2_length in X. + now apply All2_length in a. - constructor; try solve [etransitivity; eauto]. eapply All2_trans; eauto. intros ? ? ? (->&?) (->&?). @@ -1135,7 +1163,7 @@ Proof. do 3 (split; auto). eapply context_change_decl_types_red; eauto. apply fix_context_change_decl_types. - now apply All2_length in X. + now apply All2_length in a. Qed. Lemma whne_red1_inv Σ Γ t t' : @@ -1179,6 +1207,19 @@ Proof. intros ? ? (?&?); auto. Qed. +Lemma whne_red_inv Σ Γ t t' : + whne RedFlags.default Σ Γ t -> + red Σ Γ t t' -> + whnf_red Σ Γ t t'. +Proof. + intros wh r. + induction r using red_rect_n1. + - now apply whnf_red_refl. + - eapply whne_red1_inv in X; eauto. + + etransitivity; eassumption. + + eapply whne_pres; eauto. +Qed. + Lemma whnf_red1_inv Σ Γ t t' : whnf RedFlags.default Σ Γ t -> red1 Σ Γ t t' -> @@ -1198,7 +1239,7 @@ Proof. eapply OnOne2_All2; eauto. - apply red1_mkApps_tFix_inv in r. 2: { destruct unfold_fix as [(?&?)|]; [|easy]. - now unfold is_constructor; rewrite H. } + now unfold is_constructor; rewrite y. } destruct r as [[(?&->&?)|(?&->&?)]|(?&->&?)]. + apply whnf_red_mkApps; auto. 2: eapply OnOne2_All2; eauto. @@ -1258,74 +1299,13 @@ Proof. Qed. Lemma whnf_red_isApp Σ Γ t t' : - whnf RedFlags.default Σ Γ t -> - red Σ Γ t t' -> + whnf_red Σ Γ t t' -> isApp t' = isApp t. Proof. - intros wh r. - eapply whnf_red_inv in r; eauto. + intros r. now depelim r. Qed. -Lemma whnf_red_mkApps_l Σ Γ hd args t : - whnf RedFlags.default Σ Γ (mkApps hd args) -> - red Σ Γ (mkApps hd args) t -> - ∥∑ hd' args', - t = mkApps hd' args' × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. -Proof. - intros wh r. - eapply whnf_red_inv in wh; eauto. - induction args in hd, t, r, wh |- * using List.rev_ind. - - cbn in *. - constructor; eexists _, []; split; [reflexivity|]. - eauto with pcuic. - - rewrite <- mkApps_nested in wh. - cbn in wh. - depelim wh. - apply IHargs in X as [(?&?&->&?&?)]; auto. - constructor. - exists x0, (x1 ++ [arg']). - rewrite <- mkApps_nested. - eauto with pcuic. -Qed. - -Lemma whnf_red_mkApps_r Σ Γ t hd' args' : - whnf RedFlags.default Σ Γ t -> - red Σ Γ t (mkApps hd' args') -> - ∥∑ hd args, - t = mkApps hd args × red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. -Proof. - intros wh r. - eapply whnf_red_inv in wh; eauto. - induction args' in hd', t, r, wh |- * using List.rev_ind. - - cbn in *. - constructor; eexists _, []; split; [reflexivity|]. - eauto with pcuic. - - rewrite <- mkApps_nested in wh. - cbn in wh. - depelim wh. - apply IHargs' in X as [(?&?&->&?&?)]; auto. - constructor. - exists x0, (x1 ++ [arg]). - rewrite <- mkApps_nested. - eauto with pcuic. -Qed. - -Lemma whnf_red_mkApps_inv Σ Γ hd args hd' args' : - whnf RedFlags.default Σ Γ (mkApps hd args) -> - red Σ Γ (mkApps hd args) (mkApps hd' args') -> - isApp hd = false -> - isApp hd' = false -> - ∥red Σ Γ hd hd' × All2 (red Σ Γ) args args'∥. -Proof. - intros wh r notapp notapp'. - apply whnf_red_mkApps_l in r as [(?&?&eq&?&?)]; auto. - apply whnf_mkApps_inv in wh. - eapply whnf_red_isApp in wh as ?; eauto. - apply mkApps_notApp_inj in eq as (->&->); auto; [|congruence]. - easy. -Qed. - Lemma whne_eq_term_upto_univ_napp f Σ Γ t Re Rle napp t' : whne f Σ Γ t -> eq_term_upto_univ_napp Σ Re Rle napp t t' -> @@ -1334,29 +1314,29 @@ Proof. intros wh eq. induction wh in Re, Rle, napp, t, t', eq, wh |- *; depelim eq; try solve [eauto using whne; depelim wh; solve_discr; eauto using whne]. - - destruct args as [|? ? _] using List.rev_ind; [discriminate x|]. + - destruct args as [|? ? _] using MCList.rev_ind; [discriminate x|]. rewrite <- mkApps_nested in x; cbn in x; inv x. - apply eq_term_upto_univ_napp_mkApps_l_inv in eq1 as (?&?&(?&?)&->). - depelim e. - change (tApp ?h ?a) with (mkApps h [a]); rewrite mkApps_nested. + apply eq_term_upto_univ_napp_mkApps_l_inv in eq1 as (?&?&(eq_hds&?)&->). + depelim eq_hds. + rewrite <- mkApps_snoc. assert (All2 (eq_term_upto_univ Σ Re Re) (args ++ [x0]) (x1 ++ [u'])) by (now apply All2_app). unfold unfold_fix in *. destruct (nth_error mfix idx) eqn:nth; [|easy]. eapply All2_nth_error_Some in nth; eauto. destruct nth as (?&?&?&?). - inv H. - rewrite e0 in H0. - eapply All2_nth_error_Some in H0; eauto. - destruct H0 as (?&?&?). + eapply All2_nth_error_Some in e0; eauto. + inv e. + destruct e0 as (?&?&?). eapply whne_fixapp. + unfold unfold_fix. - rewrite e. + rewrite e1. reflexivity. - + eassumption. + + rewrite <- e2. + eassumption. + eapply IHwh; eauto. - - destruct args using List.rev_ind; [|rewrite <- mkApps_nested in x; discriminate x]. - now rewrite nth_error_nil in H0. + - destruct args using MCList.rev_ind; [|rewrite <- mkApps_nested in x; discriminate x]. + now rewrite nth_error_nil in e0. Qed. Lemma whnf_eq_term_upto_univ_napp f Σ Γ t Re Rle napp t' : @@ -1387,8 +1367,7 @@ Proof. destruct (nth_error mfix' idx) eqn:nth; [|easy]. eapply All2_nth_error_Some_r in nth; eauto. destruct nth as (?&?&(?&?)&?). - rewrite e in H. - rewrite e2 in H. + rewrite e e2 in y. eapply All2_nth_error_None; eauto. - apply eq_term_upto_univ_napp_mkApps_l_inv in eq as (?&?&(?&?)&->). depelim e. diff --git a/pcuic/theories/PCUICSafeLemmata.v b/pcuic/theories/PCUICSafeLemmata.v index 6985179e7..a55fc8964 100644 --- a/pcuic/theories/PCUICSafeLemmata.v +++ b/pcuic/theories/PCUICSafeLemmata.v @@ -1468,11 +1468,11 @@ Section Lemmata. intros behaves wh conv. induction wh; eauto using whne. destruct nth_error eqn:nth; [|easy]. - cbn in H. + cbn in *. eapply context_relation_nth in nth; eauto. - destruct nth as (?&?&?&?). + destruct nth as (?&eq&?&?). constructor. - rewrite e. + rewrite eq. cbn. specialize (behaves _ _ _ _ r). f_equal. diff --git a/safechecker/theories/PCUICSafeChecker.v b/safechecker/theories/PCUICSafeChecker.v index 3d70ae03c..6aae493e5 100644 --- a/safechecker/theories/PCUICSafeChecker.v +++ b/safechecker/theories/PCUICSafeChecker.v @@ -406,7 +406,7 @@ Section Typecheck. apply reduce_term_sound. Defined. - Theorem hnf_complete {Γ t h} : whnf RedFlags.default Σ Γ (hnf Γ t h). + Theorem hnf_complete {Γ t h} : ∥whnf RedFlags.default Σ Γ (hnf Γ t h)∥. Proof. apply reduce_term_complete. Qed. @@ -440,7 +440,7 @@ Section Typecheck. Proof. funelim (reduce_to_sort Γ t wt); try congruence. intros _ s r. - pose proof (@hnf_complete Γ t0 h). + pose proof (@hnf_complete Γ t0 h) as [wh]. pose proof (@hnf_sound Γ t0 h) as [r']. destruct HΣ. eapply red_confluence in r as (?&r1&r2); eauto. @@ -448,7 +448,7 @@ Section Typecheck. eapply whnf_red_inv in r1; eauto. depelim r1. clear Heq. - rewrite H0 in n0. + rewrite H in n0. now cbn in n0. Qed. @@ -481,7 +481,7 @@ Section Typecheck. Proof. funelim (reduce_to_prod Γ t wt); try congruence. intros _ na a b r. - pose proof (@hnf_complete Γ t0 h). + pose proof (@hnf_complete Γ t0 h) as [wh]. pose proof (@hnf_sound Γ t0 h) as [r']. destruct HΣ. eapply red_confluence in r as (?&r1&r2); eauto. @@ -489,7 +489,7 @@ Section Typecheck. eapply whnf_red_inv in r1; auto. depelim r1. clear Heq. - rewrite H0 in n0. + rewrite H in n0. now cbn in n0. Qed. @@ -554,6 +554,7 @@ Section Typecheck. pose proof (reduce_stack_whnf RedFlags.default Σ HΣ Γ t ε h) as wh. unfold reduce_stack in *. destruct reduce_stack_full as ((hd&π)&r'&stack_valid&(notapp&_)). + destruct wh as [wh]. apply Req_red in r' as [r']. unfold Pr in stack_valid. cbn in *. @@ -571,9 +572,9 @@ Section Typecheck. cbn in *. rewrite app_nil_r in wh. apply red_mkApps_tInd in r2 as (?&->&?); auto. - apply whnf_red_mkApps_inv in r1 as [(?&?)]; auto. - apply whnf_red_inv in r; [|now apply whnf_mkApps_inv in wh]. - depelim r. + eapply whnf_red_inv in r1; eauto. + apply whnf_red_mkApps_inv in r1 as (?&?); auto. + depelim w. noconf e0. discriminate i0. Qed. diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index c2c7c6375..e738a9ddd 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -16,22 +16,13 @@ Local Set Keyed Unification. Set Default Goal Selector "!". -Module PSR := PCUICSafeReduce. +Set Equations Transparent. -(* Removing the definitions generated by equations *) -Add Search Blacklist "obligation". -Add Search Blacklist "equation". -Add Search Blacklist "clause". -Add Search Blacklist "Functional". -Add Search Blacklist "graph". -Add Search Blacklist "elim". -(* Removing the eliminators *) -Add Search Blacklist "ind". -Add Search Blacklist "rec". +Module PSR := PCUICSafeReduce. (** * Conversion for PCUIC without fuel - Following PCUICSafereduce, we derive a fuel-free implementation of + Following PCUICSafeReduce, we derive a fuel-free implementation of conversion (and cumulativity) checking for PCUIC. *) @@ -122,8 +113,6 @@ Qed. Section Conversion. Context {cf : checker_flags}. - (* Unused *) - (* Context (flags : RedFlags.t). *) Context (Σ : global_env_ext). Context (hΣ : ∥ wf Σ ∥) (Hφ : ∥ on_udecl Σ.1 Σ.2 ∥). Context (G : universes_graph) (HG : is_graph_of_uctx G (global_ext_uctx Σ)). @@ -221,7 +210,6 @@ Section Conversion. Definition eqt u v := ∥ eq_term Σ Σ u v ∥. - (* TODO REMOVE? *) Lemma eq_term_valid_pos : forall {u v p}, validpos u p -> @@ -281,9 +269,8 @@ Section Conversion. -- intro. eapply posR_Acc. -- intro. eapply stateR_Acc. * intros x x' y [e] [y' [x'' [r [[e1] [e2]]]]]. - eexists _,_. intuition eauto. - -- constructor. assumption. - -- constructor. eapply eq_term_trans. all: eauto. + eexists _,_. intuition eauto using sq. + constructor. eapply eq_term_trans. all: eauto. * intros x. exists (sq (eq_term_refl _ _ _)). intros [[q'' h] ?]. unfold R_aux_obligations_obligation_2. simpl. f_equal. f_equal. @@ -307,9 +294,8 @@ Section Conversion. rewrite ee. right. clear ee. assumption. * eapply wcored_wf. - intros x x' y [e] [y' [x'' [r [[e1] [e2]]]]]. - eexists _,_. intuition eauto. - + constructor. assumption. - + constructor. eapply eq_term_trans. all: eauto. + eexists _,_. intuition eauto using sq. + constructor. eapply eq_term_trans. all: eauto. - intros x. exists (sq (eq_term_refl _ _ _)). intros [[q' h] [? [? ?]]]. unfold R_aux_obligations_obligation_1. simpl. f_equal. f_equal. @@ -745,7 +731,7 @@ Section Conversion. Definition isred_full Γ t π := isApp t = false /\ - whnf RedFlags.nodelta Σ (Γ,,, stack_context π) (zipp t π). + ∥whnf RedFlags.nodelta Σ (Γ,,, stack_context π) (zipp t π)∥. Lemma isred_full_nobeta Γ t π : isred_full Γ t π -> @@ -757,10 +743,11 @@ Section Conversion. unfold zipp in isr. destruct π; cbn in *; try easy. destruct (decompose_stack π) in isr. + destruct isr as [isr]. depelim isr; rewrite mkApps_tApp in *; try solve [solve_discr]. - apply whne_mkApps_inv in H0; [|easy]. - destruct H0 as [|(?&?&?&?&?&?&?&?)]; [|discriminate]. - depelim H0; solve_discr; discriminate. + apply whne_mkApps_inv in w; [|easy]. + destruct w as [|(?&?&?&?&?&?&?&?)]; [|discriminate]. + depelim w; solve_discr; discriminate. Qed. Lemma eta_pair {A B} (p q : A * B) : @@ -1398,7 +1385,7 @@ Section Conversion. rewrite zipc_appstack in r. rewrite zipp_appstack in wh. cbn in *; subst; cbn in *. - destruct hΣ. + destruct hΣ, wh. constructor; exists (l ++ (mkApps t2 l0) :: (decompose_stack s).1). eapply PCUICSR.subject_reduction in typ. 2: eauto. @@ -1493,11 +1480,9 @@ Section Conversion. now eapply All_local_env_app in wf. Qed. - (* TODO MOVE *) Definition eqb_universe_instance u v := forallb2 (check_eqb_universe G) (map Universe.make u) (map Universe.make v). - (* TODO MOVE *) Lemma eqb_universe_instance_spec : forall u v, eqb_universe_instance u v -> @@ -1685,12 +1670,12 @@ Section Conversion. | @exist (Some (ConstantDecl {| cst_body := Some b |})) eq1 := isconv_red leq (tConst c u) π1 (subst_instance_constr u' b) π2 aux ; (* Inductive or not found *) - | _ with inspect (lookup_env Σ c) := { + | @exist _ eq1 with inspect (lookup_env Σ c) := { | @exist (Some (ConstantDecl {| cst_body := Some b |})) eq2 := isconv_red leq (subst_instance_constr u b) π1 (tConst c' u') π2 aux ; (* Both Inductive or not found *) - | _ := Error (NotFoundConstants c c') _ + | @exist _ eq2 := no (NotFoundConstants c c') } }. Solve All Obligations with @@ -1757,14 +1742,14 @@ Section Conversion. Next Obligation. (* Both c and c' are axioms. Either they are different constants or they are not convertible because the universes are different. *) - apply conv_cum_alt in H as [(?&?&(r1&r2)&e)]. + apply conv_cum_alt in H as [(?&?&(r1&r2)&eq)]. rewrite zipp_as_mkApps in r1, r2. destruct hΣ. - symmetry in e0, unknown3. + symmetry in eq1, eq2. eapply PCUICConfluence.red_mkApps_tConst_axiom in r1 as (?&->&?); eauto. eapply PCUICConfluence.red_mkApps_tConst_axiom in r2 as (?&->&?); eauto. - apply eq_termp_mkApps_inv in e as (?&?); [|easy|easy]. - depelim e. + apply eq_termp_mkApps_inv in eq as (eq&?); [|easy|easy]. + depelim eq. destruct ne as [|(_&ne)]; [congruence|]. clear aux. @@ -2468,17 +2453,18 @@ Section Conversion. RedFlags.default Σ hΣ (Γ,,, stack_context π) c h) c = true -> isred_full Γ (tCase (ind, par) p c brs) π -> - whne RedFlags.default Σ (Γ,,, stack_context π) c. + ∥whne RedFlags.default Σ (Γ,,, stack_context π) c∥. Proof. intros eq ir. - destruct ir as (_&wh). + destruct ir as (_&[wh]). apply eqb_term_spec in eq. - eapply whnf_eq_term in eq; [|now apply reduce_term_complete]. + epose proof (reduce_term_complete _ _ _ _ _ _) as [wh']. + eapply whnf_eq_term in eq; [|exact wh']. rewrite zipp_as_mkApps in wh. depelim wh; solve_discr. - apply whne_mkApps_inv in H as [|(?&?&?&?&?&?&?&?&?)]; [|easy|easy]. - depelim H; cbn in *; try easy; solve_discr. - apply whnf_whne_nodelta_upgrade in eq; auto. + apply whne_mkApps_inv in w as [|(?&?&?&?&?&?&?&?&?)]; [|easy|easy]. + depelim w; cbn in *; try easy; solve_discr. + apply whnf_whne_nodelta_upgrade in eq; auto using sq. Qed. Lemma inv_reduced_discriminees_case leq Γ π π' ind ind' par par' p p' c c' brs brs' h h' : @@ -2505,7 +2491,7 @@ Section Conversion. symmetry in eq; apply Bool.andb_true_iff in eq as (c_is_red&c'_is_red). eapply reduced_case_discriminee_whne in c_is_red as wh1; eauto. eapply reduced_case_discriminee_whne in c'_is_red as wh2; eauto. - destruct hΣ. + destruct hΣ, wh1 as [wh1], wh2 as [wh2]. rewrite !zipp_as_mkApps in cc. eapply whne_conv_context in wh2. 2: eapply conv_context_sym; eauto. @@ -2520,17 +2506,18 @@ Section Conversion. RedFlags.default Σ hΣ (Γ,,, stack_context π) c h) c -> isred_full Γ (tProj p c) π -> - whne RedFlags.default Σ (Γ,,, stack_context π) c. + ∥whne RedFlags.default Σ (Γ,,, stack_context π) c∥. Proof. intros eq%eq_sym ir. - destruct ir as (_&wh). + destruct ir as (_&[wh]). apply eqb_term_spec in eq. - eapply whnf_eq_term in eq; [|now apply reduce_term_complete]. + epose proof (reduce_term_complete _ _ _ _ _ _) as [wh']. + eapply whnf_eq_term in eq; [|exact wh']. rewrite zipp_as_mkApps in wh. depelim wh; solve_discr. - apply whne_mkApps_inv in H as [|(?&?&?&?&?&?&?&?&?)]; [|easy|easy]. - depelim H; cbn in *; try easy; solve_discr. - apply whnf_whne_nodelta_upgrade in eq; auto. + apply whne_mkApps_inv in w as [|(?&?&?&?&?&?&?&?&?)]; [|easy|easy]. + depelim w; cbn in *; try easy; solve_discr. + apply whnf_whne_nodelta_upgrade in eq; auto using sq. Qed. Lemma inv_reduced_body_proj leq Γ π π' p p' c c' h h' : @@ -2551,7 +2538,7 @@ Section Conversion. intros [] c_is_red c'_is_red isr1 isr2 cc. eapply reduced_proj_body_whne in c_is_red as wh1; eauto. eapply reduced_proj_body_whne in c'_is_red as wh2; eauto. - destruct hΣ. + destruct hΣ, wh1 as [wh1], wh2 as [wh2]. rewrite !zipp_as_mkApps in cc. eapply whne_conv_context in wh2. 2: eapply conv_context_sym; eauto. @@ -2567,12 +2554,12 @@ Section Conversion. None = unfold_one_fix Γ mfix' idx' π' h' -> conv_cum leq Σ (Γ,,, stack_context π) (zipp (tFix mfix idx) π) (zipp (tFix mfix' idx') π') -> ∥idx = idx' × - All2 - (fun d d' => - rarg d = rarg d' × - Σ;;; Γ,,, stack_context π |- dtype d = dtype d' × - Σ;;; Γ,,, stack_context π,,, fix_context mfix |- dbody d = dbody d') mfix mfix' × - conv_terms Σ (Γ,,, stack_context π) (decompose_stack π).1 (decompose_stack π').1∥. + All2 (fun d d' => + rarg d = rarg d' × + Σ;;; Γ,,, stack_context π |- dtype d = dtype d' × + Σ;;; Γ,,, stack_context π,,, fix_context mfix |- dbody d = dbody d') + mfix mfix' × + conv_terms Σ (Γ,,, stack_context π) (decompose_stack π).1 (decompose_stack π').1∥. Proof. intros [?] uf1 uf2 cc. rewrite !zipp_as_mkApps in cc. @@ -2995,7 +2982,6 @@ Section Conversion. apply Prod_conv_cum_inv in H as (?&_); auto. Qed. - (* tCase *) Next Obligation. unshelve eapply R_stateR. @@ -3736,25 +3722,6 @@ Section Conversion. all: try reflexivity. simpl. constructor. Qed. - - (* TODO MOVE *) - Lemma App_conv' : - forall leq Γ t1 t2 u1 u2, - conv_cum leq Σ Γ t1 t2 -> - Σ ;;; Γ |- u1 = u2 -> - conv_cum leq Σ Γ (tApp t1 u1) (tApp t2 u2). - Proof. - intros leq Γ t1 t2 u1 u2 ht hu. - destruct hΣ. - destruct leq. - - destruct ht as [ht]. - constructor. apply App_conv. all: assumption. - - destruct ht as [ht]. constructor. - eapply cumul_trans. - + assumption. - + eapply cumul_App_l. eassumption. - + eapply cumul_App_r. assumption. - Qed. Definition Aux' Γ t1 args1 l1 π1 t2 π2 h2 := forall u1 u2 ca1 a1 ρ2 @@ -3768,7 +3735,7 @@ Section Conversion. pzt x = pzt y /\ positionR (` (pps1 x)) (` (pps1 y)) -> Ret Reduction Γ u1 (coApp (mkApps t1 ca1) (appstack a1 π1)) u2 ρ2. - + Equations(noeqns) _isconv_args' (leq : conv_pb) (Γ : context) (t1 : term) (args1 : list term) (l1 : list term) (π1 : stack) @@ -3951,7 +3918,6 @@ Section Conversion. | ccview_other t _ := None } }. - Next Obligation. destruct hΣ as [wΣ]. cbn. destruct h as [[T h] | [[ctx [s [h1 _]]]]]; [| discriminate ]. @@ -4014,11 +3980,11 @@ Section Conversion. funelim (unfold_one_case Γ ind par p c brs h); intros [=]. - clear H. simpl_reduce_stack. + apply wellformed_nonarity in h as (?&typ); auto. + destruct hΣ, w. constructor; exists (mkApps t0 (decompose_stack s).1). split; [easy|]. specialize (isr eq_refl) as (noapp&_). - apply wellformed_nonarity in h as (?&typ); auto. - destruct hΣ. eapply PCUICSR.subject_reduction in typ. 2: eauto. 2: eapply red_case_c; eauto. @@ -4122,11 +4088,11 @@ Section Conversion. funelim (unfold_one_proj Γ p c h); intros [=]. - clear H. simpl_reduce_stack. + destruct hΣ, w. + apply wellformed_nonarity in h as (?&typ); auto. constructor; exists (mkApps t0 (decompose_stack s).1). split; [easy|]. specialize (isr eq_refl) as (noapp&_). - apply wellformed_nonarity in h as (?&typ); auto. - destruct hΣ. eapply PCUICSR.subject_reduction in typ. 2: eauto. 2: eapply red_proj_c; eauto. @@ -4289,9 +4255,9 @@ Section Conversion. whnf RedFlags.default Σ (Γ,,, stack_context π) (mkApps t' args')∥. Proof. funelim (reducible_head Γ t π h); intros notapp wh [=]. - - apply whnf_mkApps_inv in wh; auto. + - apply whnf_mkApps_inv in wh. depelim wh; solve_discr. - depelim H; solve_discr; try discriminate. + depelim w; solve_discr; try discriminate. constructor; eexists _, (decompose_stack π).1. split; [now constructor|]. split; eauto with pcuic. @@ -4310,7 +4276,7 @@ Section Conversion. split; [constructor; eauto with pcuic|]. split; eauto with pcuic. apply whnf_mkApps. - depelim H; solve_discr; eauto with pcuic. + depelim w; solve_discr; eauto with pcuic. - apply whnf_mkApps_tSort_inv in wh as ->. constructor; eexists _, []. eauto using whnf_red with pcuic. @@ -4318,14 +4284,14 @@ Section Conversion. constructor; eexists _, []. eauto using whnf_red with pcuic. - depelim wh; solve_discr. - + apply whne_mkApps_inv in H as [|(?&?&?&?&?&?&?&?&?)]; auto; try discriminate. - depelim H; solve_discr. + + apply whne_mkApps_inv in w as [|(?&?&?&?&?&?&?&?&?)]; auto; try discriminate. + depelim w; solve_discr. discriminate. + rewrite H1. constructor; eexists _, []; eauto using whnf_red with pcuic. - apply whnf_mkApps_inv in wh; auto. depelim wh; solve_discr. - depelim H; solve_discr. + depelim w; solve_discr. discriminate. - discriminate. - constructor; eexists _, (decompose_stack π).1; eauto using whnf_red with pcuic. @@ -4613,20 +4579,18 @@ Section Conversion. Qed. Next Obligation. apply h; clear h. - destruct ir1 as (notapp1&whδ1), ir2 as (notapp2&whδ2). + destruct ir1 as (notapp1&[whδ1]), ir2 as (notapp2&[whδ2]). rewrite !zipp_as_mkApps in *. apply reducible_head_None in nored1 as [(?&?&s1&r1&wh1)]; auto. apply reducible_head_None in nored2 as [(?&?&s2&r2&wh2)]; auto. destruct hΣ, hx. - apply whnf_red_red in s1 as H'. - destruct H'. - apply whnf_red_red in s2 as H'. - destruct H'. + apply whnf_red_red in s1 as ?. + apply whnf_red_red in s2 as ?. eapply conv_cum_red_conv_inv in H. 2: assumption. 2: eassumption. - 2: apply red_mkApps; [eassumption|eassumption]. - 2: apply red_mkApps; [eassumption|eassumption]. + 2: apply red_mkApps; eassumption. + 2: apply red_mkApps; eassumption. apply conv_cum_mkApps_inv in H as [(?&?)]; auto. 2: now depelim s1. 2: now depelim s2. @@ -4637,15 +4601,13 @@ Section Conversion. Qed. Next Obligation. unfold eqb_termp_napp in noteq. - destruct ir1 as (notapp1&whδ1), ir2 as (notapp2&whδ2). + destruct ir1 as (notapp1&[whδ1]), ir2 as (notapp2&[whδ2]). rewrite !zipp_as_mkApps in *. apply reducible_head_None in nored1 as [(?&?&s1&rargs1&wh1)]; auto. apply reducible_head_None in nored2 as [(?&?&s2&rargs2&wh2)]; auto. destruct hΣ, hx. - apply whnf_red_red in s1 as H'. - destruct H'. - apply whnf_red_red in s2 as H'. - destruct H'. + apply whnf_red_red in s1 as ?. + apply whnf_red_red in s2 as ?. eapply conv_cum_red_conv_inv in H. 2: assumption. 2: eassumption. diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index 745310780..8206997a0 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -1158,10 +1158,10 @@ Section Reduce. (* *) *) (* Abort. *) - Scheme Acc_ind' := Induction for Acc Sort Prop. + Scheme Acc_ind' := Induction for Acc Sort Type. Lemma Fix_F_prop : - forall A R P f (pred : forall x : A, P x -> Prop) x hx, + forall A R P f (pred : forall x : A, P x -> Type) x hx, (forall x aux, (forall y hy, pred y (aux y hy)) -> pred x (f x aux)) -> pred x (@Fix_F A R P f x hx). Proof. @@ -1171,7 +1171,7 @@ Section Reduce. Qed. Lemma reduce_stack_prop : - forall Γ t π h (P : term × stack -> term × stack -> Prop), + forall Γ t π h (P : term × stack -> term × stack -> Type), (forall t π h aux, (forall t' π' hR, P (t', π') (` (aux t' π' hR))) -> P (t, π) (` (_reduce_stack Γ t π h aux))) -> @@ -1236,14 +1236,14 @@ Section Reduce. Qed. Lemma whnf_non_ctor_finite_ind_typed Γ t ind u args : + wf Σ -> whnf flags Σ Γ t -> isConstruct_app t = false -> check_recursivity_kind Σ (inductive_mind ind) CoFinite = false -> Σ;;; Γ |- t : mkApps (tInd ind u) args -> whne flags Σ Γ t. Proof. - destruct hΣ. - intros wh ctor fin typ. + intros wf wh ctor fin typ. destruct wh. - easy. - apply inversion_Sort in typ as (?&?&?&?&?); auto. @@ -1269,14 +1269,14 @@ Section Reduce. end. Lemma whnf_non_ctor_non_cofix_ind_typed Γ t ind u args : + wf Σ -> whnf flags Σ Γ t -> isConstruct_app t = false -> isCoFix_app t = false -> Σ;;; Γ |- t : mkApps (tInd ind u) args -> whne flags Σ Γ t. Proof. - destruct hΣ. - intros wh ctor cof typ. + intros wf wh ctor cof typ. destruct wh. - easy. - apply inversion_Sort in typ as (?&?&?&?&?); auto. @@ -1294,6 +1294,7 @@ Section Reduce. Qed. Lemma whnf_fix_arg_whne mfix idx body Γ t before args aftr ty : + wf Σ -> unfold_fix mfix idx = Some (#|before|, body) -> match t with | tConstruct _ _ _ => False @@ -1304,8 +1305,7 @@ Section Reduce. Σ;;; Γ |- mkApps (tFix mfix idx) (before ++ mkApps t args :: aftr) : ty -> whne flags Σ Γ (mkApps t args). Proof. - destruct hΣ. - intros uf shape wh typ. + intros wf uf shape wh typ. apply inversion_mkApps in typ as (fix_ty & typ_fix & typ_args); auto. apply inversion_Fix in typ_fix as (def&?&?&?&?&?&?); auto. eapply All_nth_error in a; eauto. @@ -1326,6 +1326,7 @@ Section Reduce. Qed. Lemma whnf_case_arg_whne Γ hd args ind par p brs T : + wf Σ -> match hd with | tApp _ _ | tConstruct _ _ _ @@ -1336,8 +1337,7 @@ Section Reduce. Σ;;; Γ |- tCase (ind, par) p (mkApps hd args) brs : T -> whne flags Σ Γ (mkApps hd args). Proof. - intros shape wh typ. - destruct hΣ. + intros wf shape wh typ. apply inversion_Case in typ as (?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?&?); auto. eapply whnf_non_ctor_finite_ind_typed; try eassumption. - unfold isConstruct_app. @@ -1352,6 +1352,7 @@ Section Reduce. Qed. Lemma whnf_proj_arg_whne Γ hd args p T : + wf Σ -> match hd with | tApp _ _ | tCoFix _ _ => False @@ -1362,8 +1363,7 @@ Section Reduce. Σ;;; Γ |- tProj p (mkApps hd args) : T -> whne flags Σ Γ (mkApps hd args). Proof. - intros shape wh typ. - destruct hΣ as [wfΣ]. + intros wf shape wh typ. assert (typ' := typ). apply inversion_Proj in typ as (?&?&?&?&?&?&?&?&?); auto. cbn in *. @@ -1382,13 +1382,13 @@ Section Reduce. Lemma reduce_stack_whnf : forall Γ t π h, let '(u, ρ) := reduce_stack Γ t π h in - whnf flags Σ (Γ ,,, stack_context ρ) (zipp u ρ). + ∥whnf flags Σ (Γ ,,, stack_context ρ) (zipp u ρ)∥. Proof. intros Γ t π h. eapply reduce_stack_prop with (P := fun x y => let '(u, ρ) := y in - whnf flags Σ (Γ ,,, stack_context ρ) (zipp u ρ) + ∥whnf flags Σ (Γ ,,, stack_context ρ) (zipp u ρ)∥ ). clear. intros t π h aux haux. @@ -1411,13 +1411,13 @@ Section Reduce. all: try solve [ constructor ; constructor ]. all: try solve [ unfold zipp ; case_eq (decompose_stack π) ; intros ; - constructor ; eapply whne_mkApps ; constructor + do 2 constructor ; eapply whne_mkApps ; constructor ]. + unfold zipp. case_eq (decompose_stack π). intros l ρ e. apply decompose_stack_eq in e. subst. destruct l. - * simpl. eapply whnf_sort. + * simpl. eauto with pcuic. * exfalso. destruct hΣ. cbn in h. zip fold in h. apply wellformed_context in h; auto. @@ -1432,7 +1432,7 @@ Section Reduce. case_eq (decompose_stack π). intros l ρ e. apply decompose_stack_eq in e. subst. destruct l. - * simpl. eapply whnf_prod. + * simpl. eauto with pcuic. * exfalso. destruct hΣ. cbn in h. zip fold in h. apply wellformed_context in h; auto. @@ -1444,7 +1444,7 @@ Section Reduce. -- destruct H as [(?&?&?&?)]. cbn in *; easy. - unfold zipp. case_eq (decompose_stack π0). intros l ρ e. - constructor. eapply whne_mkApps. + constructor. constructor. eapply whne_mkApps. eapply whne_rel_nozeta. assumption. - match goal with | |- context [ reduce ?x ?y ?z ] => @@ -1455,7 +1455,7 @@ Section Reduce. rewrite eq in haux. cbn in haux. assumption. - unfold zipp. case_eq (decompose_stack π0). intros. - constructor. eapply whne_mkApps. econstructor. + constructor. constructor. eapply whne_mkApps. econstructor. rewrite <- e. cbn. cbn in H. inversion H. reflexivity. - match goal with @@ -1467,9 +1467,9 @@ Section Reduce. rewrite eq in haux. cbn in haux. assumption. - unfold zipp. case_eq (decompose_stack π1). intros. - constructor. eapply whne_mkApps. eapply whne_letin_nozeta. assumption. + constructor. constructor. eapply whne_mkApps. eapply whne_letin_nozeta. assumption. - unfold zipp. case_eq (decompose_stack π2). intros. - constructor. eapply whne_mkApps. eapply whne_const_nodelta. assumption. + constructor. constructor. eapply whne_mkApps. eapply whne_const_nodelta. assumption. - match goal with | |- context [ reduce ?x ?y ?z ] => case_eq (reduce x y z) ; @@ -1479,7 +1479,7 @@ Section Reduce. rewrite eq in haux. cbn in haux. assumption. - unfold zipp. case_eq (decompose_stack π2). intros. - constructor. eapply whne_mkApps. econstructor. + constructor. constructor. eapply whne_mkApps. econstructor. + symmetry. exact e. + reflexivity. - match goal with @@ -1497,15 +1497,18 @@ Section Reduce. + cbn in *; subst. cbn in *; congruence. + cbn. + constructor. apply whnf_mkApps. apply whne_app. now apply whne_lam_nobeta. - unfold zipp. destruct decompose_stack. + constructor. apply whnf_mkApps. now apply whne_fix_nofix. - unfold zipp. destruct decompose_stack. + constructor. apply whnf_fixapp. now rewrite <- e. - unfold zipp. @@ -1513,6 +1516,7 @@ Section Reduce. apply decompose_stack_noStackApp in eq as ?. apply decompose_stack_eq in eq. subst. + constructor. apply whnf_fixapp. rewrite <- e. apply <- decompose_stack_at_appstack_None; eauto. @@ -1545,8 +1549,7 @@ Section Reduce. rewrite zipc_appstack in h. apply Req_red in r as [r]. cbn in *. - eapply red_wellformed in h; eauto. - 2: constructor; exact r. + eapply red_wellformed in h; eauto using sq. clear r. rewrite zipc_appstack in h. cbn in h. @@ -1571,16 +1574,18 @@ Section Reduce. rewrite app_nil_r in *. rewrite <- app_assoc in *. cbn in *. - constructor. + destruct hΣ, haux. symmetry in e. + constructor. + constructor. eapply whne_fixapp. + eassumption. + now apply nth_error_snoc. + eapply whnf_fix_arg_whne; eauto. now destruct t2. - - unfold zipp. destruct decompose_stack. + constructor. apply whnf_mkApps. now apply whne_case_noiota. - match goal with @@ -1608,8 +1613,7 @@ Section Reduce. apply decompose_stack_eq in e0 as ->. apply Req_red in r as [r]. cbn in *. - eapply red_wellformed in h; eauto. - 2: constructor; exact r. + eapply red_wellformed in h; eauto using sq. clear r. rewrite zipc_appstack in h. cbn in h. @@ -1629,6 +1633,8 @@ Section Reduce. try rewrite stack_context_appstack in H; cbn in *). rewrite app_nil_r in *. + destruct hΣ, haux. + constructor. apply whnf_mkApps, whne_case. eapply whnf_case_arg_whne; eauto. destruct H as (noapp&_). @@ -1642,7 +1648,7 @@ Section Reduce. rewrite eq in haux. cbn in haux. assumption. - unfold zipp. case_eq (decompose_stack π6). intros. - constructor. eapply whne_mkApps. eapply whne_proj_noiota. assumption. + constructor. constructor. eapply whne_mkApps. eapply whne_proj_noiota. assumption. - match type of e with | _ = reduce ?x ?y ?z => specialize (haux x y z); @@ -1660,8 +1666,7 @@ Section Reduce. apply decompose_stack_eq in e0 as ->. apply Req_red in r as [r]. cbn in *. - eapply red_wellformed in h; eauto. - 2: constructor; exact r. + eapply red_wellformed in h; eauto using sq. clear r. rewrite zipc_appstack in h. cbn in h. @@ -1681,7 +1686,8 @@ Section Reduce. try rewrite stack_context_appstack in H; cbn in *). rewrite app_nil_r in *. - apply whnf_mkApps, whne_proj. + destruct hΣ, haux. + constructor. apply whnf_mkApps, whne_proj. eapply whnf_proj_arg_whne; eauto. destruct H as (noapp&_). now destruct t0. @@ -1704,7 +1710,7 @@ Section Reduce. Qed. Theorem reduce_term_complete Γ t h : - whnf flags Σ Γ (reduce_term Γ t h). + ∥whnf flags Σ Γ (reduce_term Γ t h)∥. Proof. pose proof (reduce_stack_whnf Γ t ε h) as H. unfold reduce_term. diff --git a/template-coq/theories/utils/All_Forall.v b/template-coq/theories/utils/All_Forall.v index 1ae01240c..35add2c54 100644 --- a/template-coq/theories/utils/All_Forall.v +++ b/template-coq/theories/utils/All_Forall.v @@ -1058,7 +1058,7 @@ Proof. eexists (y :: l, l0). repeat split; eauto. Qed. -Lemma All2_ind_rev : forall (A B : Type) (R : A -> B -> Type) (P : forall (l : list A) (l0 : list B), Prop), +Lemma All2_rect_rev : forall (A B : Type) (R : A -> B -> Type) (P : forall (l : list A) (l0 : list B), Type), P [] [] -> (forall (x : A) (y : B) (l : list A) (l' : list B) (r : R x y) (a : All2 R l l'), P l l' -> P (l ++ [x])%list (l' ++ [y]))%list -> @@ -1067,9 +1067,18 @@ Proof. intros. revert l0 a. induction l using rev_ind; cbn; intros. - inv a. eauto. - eapply All2_app_inv in a as ([] & [[]]). subst. - inv a0. inv X0. eauto. + inv a0. inv X2. eauto. Qed. +Lemma All2_ind_rev : forall (A B : Type) (R : A -> B -> Type) (P : forall (l : list A) (l0 : list B), Prop), + P [] [] -> + (forall (x : A) (y : B) (l : list A) (l' : list B) (r : R x y) (a : All2 R l l'), + P l l' -> P (l ++ [x])%list (l' ++ [y]))%list -> + forall (l : list A) (l0 : list B) (a : All2 R l l0), P l l0. +Proof. + intros. + eapply All2_rect_rev; eauto. +Qed. Lemma All2_app : forall (A B : Type) (R : A -> B -> Type), From ebd5c42066dc5f27d0b40e78d13439df2d44c959 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Wed, 28 Oct 2020 12:17:00 +0100 Subject: [PATCH 19/26] Minor cleanups, rerun CI --- pcuic/theories/PCUICNormal.v | 13 ------------- pcuic/theories/PCUICReduction.v | 2 -- 2 files changed, 15 deletions(-) diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index 8666fd8ec..58424d6e8 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -1207,19 +1207,6 @@ Proof. intros ? ? (?&?); auto. Qed. -Lemma whne_red_inv Σ Γ t t' : - whne RedFlags.default Σ Γ t -> - red Σ Γ t t' -> - whnf_red Σ Γ t t'. -Proof. - intros wh r. - induction r using red_rect_n1. - - now apply whnf_red_refl. - - eapply whne_red1_inv in X; eauto. - + etransitivity; eassumption. - + eapply whne_pres; eauto. -Qed. - Lemma whnf_red1_inv Σ Γ t t' : whnf RedFlags.default Σ Γ t -> red1 Σ Γ t t' -> diff --git a/pcuic/theories/PCUICReduction.v b/pcuic/theories/PCUICReduction.v index 615909708..e268f244d 100644 --- a/pcuic/theories/PCUICReduction.v +++ b/pcuic/theories/PCUICReduction.v @@ -353,8 +353,6 @@ Proof. eapply Hstep; eauto. now apply clos_rt_rt1n_iff. Defined. - - (** Simple lemmas about reduction *) From b644d751458b0c8a1dc0c008e1ea7f6a444166d7 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Thu, 29 Oct 2020 23:41:27 +0100 Subject: [PATCH 20/26] Try removing packages before install --- Makefile | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/Makefile b/Makefile index 30bc6e689..f0fa1e7ee 100644 --- a/Makefile +++ b/Makefile @@ -93,6 +93,8 @@ ci-local: $(MAKE) all test-suite TIMED=pretty-timed ci-opam: + for pkg in coq-metacoq coq-metacoq-template coq-metacoq-checker coq-metacoq-translations coq-metacoq-pcuic coq-metacoq-safechecker coq-metacoq-erasure; do + opam remove -y $pkg || true + done # Use -v so that regular output is produced opam install -v -y . - opam remove -y coq-metacoq coq-metacoq-template From 4739f90ed0d904a328f08fa07136fd3670287005 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Fri, 30 Oct 2020 11:06:24 +0100 Subject: [PATCH 21/26] Fix build We no longer need PCUICSafeLemmata since the nodelta flags were moved to PCUICNormal. --- Makefile | 4 +--- erasure/_PluginProject.in | 6 +++--- erasure/src/metacoq_erasure_plugin.mlpack | 1 - safechecker/_PluginProject.in | 4 ++-- safechecker/src/metacoq_safechecker_plugin.mlpack | 1 - 5 files changed, 6 insertions(+), 10 deletions(-) diff --git a/Makefile b/Makefile index f0fa1e7ee..30bc6e689 100644 --- a/Makefile +++ b/Makefile @@ -93,8 +93,6 @@ ci-local: $(MAKE) all test-suite TIMED=pretty-timed ci-opam: - for pkg in coq-metacoq coq-metacoq-template coq-metacoq-checker coq-metacoq-translations coq-metacoq-pcuic coq-metacoq-safechecker coq-metacoq-erasure; do - opam remove -y $pkg || true - done # Use -v so that regular output is produced opam install -v -y . + opam remove -y coq-metacoq coq-metacoq-template diff --git a/erasure/_PluginProject.in b/erasure/_PluginProject.in index 2c3c72488..644924b94 100644 --- a/erasure/_PluginProject.in +++ b/erasure/_PluginProject.in @@ -81,8 +81,8 @@ src/pCUICChecker.ml src/pCUICChecker.mli src/pCUICPretty.mli src/pCUICPretty.ml -src/pCUICSafeLemmata.mli -src/pCUICSafeLemmata.ml +#src/pCUICSafeLemmata.mli +#src/pCUICSafeLemmata.ml src/templateToPCUIC.mli src/templateToPCUIC.ml @@ -126,4 +126,4 @@ src/safeTemplateErasure.ml src/g_metacoq_erasure.mlg src/metacoq_erasure_plugin.mlpack -theories/Loader.v \ No newline at end of file +theories/Loader.v diff --git a/erasure/src/metacoq_erasure_plugin.mlpack b/erasure/src/metacoq_erasure_plugin.mlpack index 761a66a55..593fc7b42 100644 --- a/erasure/src/metacoq_erasure_plugin.mlpack +++ b/erasure/src/metacoq_erasure_plugin.mlpack @@ -31,7 +31,6 @@ PCUICChecker PCUICPretty PCUICRetyping TemplateToPCUIC -PCUICSafeLemmata PCUICCumulativity PCUICUnivSubst diff --git a/safechecker/_PluginProject.in b/safechecker/_PluginProject.in index 330268674..a9d3d6fe8 100644 --- a/safechecker/_PluginProject.in +++ b/safechecker/_PluginProject.in @@ -71,8 +71,8 @@ src/pCUICNormal.ml # src/pCUICContextConversion.ml # src/pCUICConversion.mli # src/pCUICConversion.ml -src/pCUICSafeLemmata.mli -src/pCUICSafeLemmata.ml +# src/pCUICSafeLemmata.mli +# src/pCUICSafeLemmata.ml src/templateToPCUIC.mli src/templateToPCUIC.ml src/pCUICChecker.ml diff --git a/safechecker/src/metacoq_safechecker_plugin.mlpack b/safechecker/src/metacoq_safechecker_plugin.mlpack index 3f4b9bc15..a60c314d0 100644 --- a/safechecker/src/metacoq_safechecker_plugin.mlpack +++ b/safechecker/src/metacoq_safechecker_plugin.mlpack @@ -29,7 +29,6 @@ PCUICChecker PCUICPretty PCUICRetyping TemplateToPCUIC -PCUICSafeLemmata PCUICCumulativity PCUICUnivSubst From f418ae244e53ce334f35b5546f960ecb7ae075b4 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Thu, 5 Nov 2020 15:46:36 +0100 Subject: [PATCH 22/26] Finish completeness for new tCase --- safechecker/theories/PCUICSafeConversion.v | 57 ++++++++++++++++------ 1 file changed, 41 insertions(+), 16 deletions(-) diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 5c2b45dd0..459ca4f67 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -2539,8 +2539,8 @@ Section Conversion. conv_stack_ctx Γ π π' -> true = eqb_term (reduce_term RedFlags.default - Σ hΣ (Γ,,, stack_context π) c h) c && - eqb_term (reduce_term + Σ hΣ (Γ,,, stack_context π) c h) c -> + true = eqb_term (reduce_term RedFlags.default Σ hΣ (Γ,,, stack_context π') c' h') c' -> isred_full Γ (tCase (ind, par) p c brs) π -> @@ -2555,8 +2555,7 @@ Section Conversion. All2 (fun br br' => br.1 = br'.1 × (Σ;;; Γ,,, stack_context π |- br.2 = br'.2)) brs brs' × conv_terms Σ (Γ,,, stack_context π) (decompose_stack π).1 (decompose_stack π').1∥. Proof. - intros [] eq isr1 isr2 cc. - symmetry in eq; apply Bool.andb_true_iff in eq as (c_is_red&c'_is_red). + intros [] c_is_red%eq_sym c'_is_red%eq_sym isr1 isr2 cc. eapply reduced_case_discriminee_whne in c_is_red as wh1; eauto. eapply reduced_case_discriminee_whne in c'_is_red as wh2; eauto. destruct hΣ, wh1 as [wh1], wh2 as [wh2]. @@ -3079,11 +3078,11 @@ Section Conversion. assumption. Qed. Next Obligation. - change (eq_inductive ind ind') with (eqb ind ind') in eq5. - destruct (eqb_spec ind ind'). 2: discriminate. - change (Nat.eqb par par') with (eqb par par') in eq5. - destruct (eqb_spec par par'). 2: discriminate. - assumption. + change (eq_inductive ind ind') with (eqb ind ind') in eq5. + destruct (eqb_spec ind ind'). 2: discriminate. + change (Nat.eqb par par') with (eqb par par') in eq5. + destruct (eqb_spec par par'). 2: discriminate. + assumption. Qed. Next Obligation. unshelve eapply R_stateR. @@ -3105,19 +3104,29 @@ Section Conversion. eapply conv_Case. all: assumption. Qed. Next Obligation. - todo "Completeness". + apply h; clear h. + eapply inv_reduced_discriminees_case in H as [([= <- <-]&?&?&?&?)]; eauto. + constructor; auto. Qed. Next Obligation. - todo "Completeness". + apply h; cbn; clear h. + eapply inv_reduced_discriminees_case in H as [([= <- <-]&?&?&?&?)]; eauto. + constructor; auto. Qed. Next Obligation. - todo "Completeness". + apply h; cbn; clear h. + eapply inv_reduced_discriminees_case in H as [([= <- <-]&?&?&?&?)]; eauto. + constructor; auto. Qed. Next Obligation. - todo "Completeness". + apply h; cbn; clear h. + eapply inv_reduced_discriminees_case in H as [([= <- <-]&?&?&?&?)]; eauto. + constructor; auto. Qed. Next Obligation. - todo "Completeness". + eapply inv_reduced_discriminees_case in H as [([= <- <-]&?&?&?&?)]; eauto. + rewrite eq_inductive_refl, Nat.eqb_refl in eq5. + congruence. Qed. Next Obligation. eapply red_wellformed ; auto. @@ -3179,7 +3188,16 @@ Section Conversion. + eapply conv_context_sym. all: auto. Qed. Next Obligation. - todo "Completeness". + apply h; clear h. + pose proof hΣ as []. + destruct hx. + epose proof (reduce_term_sound _ _ _ _ _ _) as [r]. + eapply conv_cum_red_conv_inv. + 1,2,5: eassumption. + 1: reflexivity. + apply red_zipp. + apply red_case_c. + exact r. Qed. Next Obligation. eapply red_wellformed ; auto. @@ -3237,7 +3255,14 @@ Section Conversion. - assumption. Qed. Next Obligation. - todo "Completeness". + apply h; clear h. + pose proof hΣ as []. + destruct hx. + epose proof (reduce_term_sound _ _ _ _ _ _) as [r]. + eapply conv_cum_red_inv. + 1,4: eassumption. + 2: reflexivity. + apply red_zipp, red_case_c, r. Qed. (* tProj *) From b93103cfb4984b74f55d318a8e1219afb51fa54e Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Thu, 5 Nov 2020 15:49:13 +0100 Subject: [PATCH 23/26] Finish completeness for tFix --- safechecker/theories/PCUICSafeConversion.v | 233 +++++++-------------- 1 file changed, 79 insertions(+), 154 deletions(-) diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 459ca4f67..0529fbcc0 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -2816,7 +2816,7 @@ Section Conversion. } ; | Error e h := no e } ; - | @exist false _ := + | @exist false idx_uneq := no ( CannotUnfoldFix (Γ ,,, stack_context π1) mfix idx @@ -3358,8 +3358,6 @@ Section Conversion. simpl. eapply cored_zipc. eapply cored_proj. assumption. Qed. Next Obligation. - (* destruct hΣ as [wΣ]. *) - (* Why doesn't the above work anymore? *) pose proof hΣ as w. destruct w. destruct hx as [hx]. match type of h with @@ -3484,26 +3482,19 @@ Section Conversion. Next Obligation. apply unfold_one_fix_red in eq1 as r1. apply unfold_one_fix_decompose in eq1 as d1. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 ; - pose proof (reduce_stack_context f Σ hΣ Γ t π h) as c2 - end. - rewrite <- eq3 in r2. cbn in r2. rewrite zipc_appstack in r2. cbn in r2. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. - rewrite <- eq3 in c2. cbn in c2. rewrite stack_context_appstack in c2. - cbn in c2. - case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (red_wellformed _ hΣ h1 r1) as hh. - apply red_context in r2. - pose proof (decompose_stack_eq _ _ _ (eq_sym eq2)). subst. - rewrite zipc_appstack in hh. cbn in r2. - pose proof (red_wellformed _ hΣ hh (sq r2)) as hh2. - rewrite zipc_stack_cat. - assumption. + reduce_stack_facts. + simpl_stacks. + destruct r1 as [r1]. + eapply red_wellformed. + 1: eauto. + 1: exact h1. + constructor. + etransitivity. + 2: { apply red_zipc. + rewrite stack_context_decompose. + exact r. } + simpl_stacks. + eassumption. Qed. Next Obligation. apply unfold_one_fix_cored in eq1 as r1. @@ -3524,93 +3515,48 @@ Section Conversion. Qed. Next Obligation. apply unfold_one_fix_decompose in eq1 as d1. - rewrite <- eq2 in d1. simpl in d1. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. - case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ (eq_sym eq2)). subst. - apply decompose_stack_eq in e as ?. subst. - rewrite stack_cat_appstack. - rewrite stack_context_appstack. - case_eq (decompose_stack π1). intros args ρ e'. - simpl. - apply decompose_stack_eq in e'. subst. - clear eq3. - rewrite stack_context_appstack in hx. - assumption. + reduce_stack_facts. + now simpl_stacks. Qed. Next Obligation. - todo "Completeness". + simpl_reduce_stack. + specialize (isr eq_refl) as (?&_). + split; [easy|]. + cbn in *. + now simpl_stacks. Qed. Next Obligation. - destruct hΣ as [wΣ]. apply unfold_one_fix_red_zipp in eq1 as r1. - destruct r1 as [r1]. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in r2. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. apply unfold_one_fix_decompose in eq1 as d1. - assert (r2' : red (fst Σ) (Γ ,,, stack_context π1) (zipp fn θ) (zipp fn' (ρ +++ θ'))). - { unfold zipp. - case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. rewrite decompose_stack_appstack. - rewrite <- eq2. - cbn in r2. rewrite 2!zipc_appstack in r2. cbn in r2. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π1). intros l1 ρ1 e1. - simpl. rewrite e1 in r2. simpl in r2. - pose proof (decompose_stack_eq _ _ _ e1). subst. - rewrite decompose_stack_twice with (1 := e1). cbn. - rewrite app_nil_r. - rewrite stack_context_appstack. assumption. - } - pose proof (red_trans _ _ _ _ _ r1 r2') as r. - assert (e : stack_context π1 = stack_context (ρ +++ θ')). - { case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. - rewrite stack_context_appstack. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π1). intros l1 ρ1 e1. - pose proof (decompose_stack_eq _ _ _ e1). subst. - rewrite stack_context_appstack. reflexivity. - } - etransitivity. - - eapply red_conv_cum_l. all: eassumption. - - rewrite e. assumption. + reduce_stack_facts. + simpl_reduce_stack. + destruct hΣ, r1 as [r1]. + eapply conv_cum_trans; auto. + 2: exact h. + eapply red_conv_cum_l. + etransitivity; eauto. Qed. Next Obligation. - todo "Completeness". + apply h; clear h. + apply unfold_one_fix_red_zipp in eq1 as r1. + apply unfold_one_fix_decompose in eq1 as d1. + reduce_stack_facts. + simpl_reduce_stack. + destruct hΣ, r1 as [r1]. + eapply conv_cum_red_inv. + 1: auto. + 2: reflexivity. + 2: exact H. + etransitivity; eauto. Qed. Next Obligation. - cbn. rewrite zipc_appstack. cbn. apply unfold_one_fix_red_zipp in eq1 as r. apply unfold_one_fix_decompose in eq1 as d. - rewrite <- eq2 in d. simpl in d. - unfold zipp in r. - rewrite <- eq2 in r. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - rewrite e2 in r. - rewrite e2 in d. simpl in d. subst. - apply wellformed_zipc_zipp in h2 as hh2 ; auto. - pose proof (decompose_stack_eq _ _ _ e2). subst. - unfold zipp in hh2. rewrite e2 in hh2. - pose proof (red_wellformed _ hΣ hh2 r) as hh. - rewrite stack_context_appstack in hh. - assumption. + clear aux eq1. + apply wellformed_zipc_zipp in h2; auto. + cbn in *. + simpl_stacks. + eapply red_wellformed; eauto. Qed. Next Obligation. apply unfold_one_fix_red in eq1 as r1. @@ -3676,65 +3622,36 @@ Section Conversion. assumption. Qed. Next Obligation. - todo "Completeness". + simpl_reduce_stack. + specialize (isr eq_refl) as (?&_). + split; [easy|]. + simpl_stacks. + eauto. Qed. Next Obligation. - destruct hΣ as [wΣ]. apply unfold_one_fix_red_zipp in eq1 as r1. - destruct r1 as [r1]. - match type of eq3 with - | _ = reduce_stack ?f ?Σ ?hΣ ?Γ ?t ?π ?h => - destruct (reduce_stack_sound f Σ hΣ Γ t π h) as [r2] ; - pose proof (reduce_stack_decompose RedFlags.nodelta _ hΣ _ _ _ h) as d2 - end. - rewrite <- eq3 in r2. - rewrite <- eq3 in d2. cbn in d2. rewrite decompose_stack_appstack in d2. - cbn in d2. apply unfold_one_fix_decompose in eq1 as d1. - assert (r2' : ∥ red (fst Σ) (Γ ,,, stack_context π2) (zipp fn θ) (zipp fn' (ρ +++ θ')) ∥). - { unfold zipp. - destruct hx as [hx]. - constructor. - case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. rewrite decompose_stack_appstack. - rewrite <- eq2. - cbn in r2. rewrite 2!zipc_appstack in r2. cbn in r2. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π2). intros l2 ρ2 e2. - simpl. rewrite e2 in r2. simpl in r2. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite decompose_stack_twice with (1 := e2). cbn. - rewrite app_nil_r. - rewrite stack_context_appstack. - assumption. - } - assert (e : stack_context π2 = stack_context (ρ +++ θ')). - { case_eq (decompose_stack ρ). intros l ξ e. - rewrite e in d2. cbn in d2. subst. - pose proof (decompose_stack_eq _ _ _ e). subst. - rewrite stack_cat_appstack. - rewrite stack_context_appstack. - rewrite <- eq2 in d1. cbn in d1. subst. - case_eq (decompose_stack π2). intros l2 ρ2 e2. simpl. - pose proof (decompose_stack_eq _ _ _ e2). subst. - rewrite stack_context_appstack. reflexivity. - } - destruct r2' as [r2']. - destruct hx as [hx]. - pose proof (red_trans _ _ _ _ _ r1 r2') as r. - etransitivity ; revgoals. - - eapply conv_cum_conv_ctx. - + assumption. - + eapply red_conv_cum_r. 1: assumption. - eassumption. - + eapply conv_context_sym. 1: auto. - assumption. - - assumption. + reduce_stack_facts. + simpl_reduce_stack. + destruct r1 as [r1]. + destruct hΣ, hx. + eapply conv_cum_red_conv. + 1,2,5: eauto. + 1: reflexivity. + etransitivity; eauto. Qed. Next Obligation. - todo "Completeness". + apply h; clear h. + apply unfold_one_fix_red_zipp in eq1 as r1. + apply unfold_one_fix_decompose in eq1 as d1. + reduce_stack_facts. + simpl_reduce_stack. + destruct r1 as [r1]. + destruct hΣ, hx. + eapply conv_cum_red_conv_inv. + 1,2,5: eauto. + 1: reflexivity. + etransitivity; eauto. Qed. Next Obligation. change (true = eqb idx idx') in eq4. @@ -3758,13 +3675,21 @@ Section Conversion. eapply conv_Fix. all: assumption. Qed. Next Obligation. - todo "Completeness". + apply h; clear h. + eapply inv_stuck_fixes in H as [(<-&?&?)]; eauto. + constructor; auto. Qed. Next Obligation. - todo "Completeness". + apply h; clear h. + eapply inv_stuck_fixes in H as [(<-&?&?)]; eauto. + constructor; auto. + eapply All2_impl; eauto. + cbn; intros; easy. Qed. Next Obligation. - todo "Completeness". + eapply inv_stuck_fixes in H as [(<-&?&?)]; eauto. + rewrite Nat.eqb_refl in idx_uneq. + congruence. Qed. (* tCoFix *) From f2d2915f49b21d25031d55022d60415dbad831db Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Thu, 5 Nov 2020 16:10:12 +0100 Subject: [PATCH 24/26] Finish completeness of tCoFix --- pcuic/theories/PCUICConvCumInversion.v | 41 ++++++++++++++++++++++ safechecker/theories/PCUICSafeConversion.v | 38 ++++++++++++++++---- 2 files changed, 73 insertions(+), 6 deletions(-) diff --git a/pcuic/theories/PCUICConvCumInversion.v b/pcuic/theories/PCUICConvCumInversion.v index ecc8c5a69..13100d6d5 100644 --- a/pcuic/theories/PCUICConvCumInversion.v +++ b/pcuic/theories/PCUICConvCumInversion.v @@ -194,6 +194,47 @@ Section fixed. eapply context_change_decl_types_red; eauto. eapply fix_context_change_decl_types; eauto. Qed. + + Lemma conv_cum_tCoFix_inv leq Γ mfix idx mfix' idx' : + conv_cum leq Σ Γ (tCoFix mfix idx) (tCoFix mfix' idx') -> + ∥idx = idx' × + All2 (fun d d' => rarg d = rarg d' × + Σ;;; Γ |- dtype d = dtype d' × + Σ;;; Γ,,, fix_context mfix |- dbody d = dbody d') + mfix mfix'∥. + Proof. + intros conv. + apply conv_cum_alt in conv as [(?&?&(r1&r2)&eq)]. + assert (forall defs i, whnf RedFlags.default Σ Γ (tCoFix defs i)) + by (intros; apply whnf_cofixapp with (v := [])). + eapply whnf_red_inv in r1; eauto. + eapply whnf_red_inv in r2; eauto. + depelim r1. + depelim r2. + depelim eq. + constructor. + split; [easy|]. + clear -a a0 a1. + cut (#|mfix| = #|mfix'|); + [|now apply All2_length in a; apply All2_length in a0; apply All2_length in a1]. + revert a a0 a1. + generalize mfix at 1 3 4. + generalize mfix' at 1 3. + intros ctx_fix ctx_fix'. + intros all1 all2 all len_eq. + induction all in mfix, mfix', mfix'0, mfix'1, all1, all2, all |- *; + depelim all1; depelim all2; [constructor|]. + constructor; [|now auto]. + destruct r as ((?&?)&?), p as (?&?&?&?), p0 as (?&?&?&?). + split; [congruence|]. + split; [now apply conv_alt_red; exists (dtype x), (dtype y)|]. + apply conv_alt_red. + exists (dbody x), (dbody y). + split; [|easy]. + split; [easy|]. + eapply context_change_decl_types_red; eauto. + eapply fix_context_change_decl_types; eauto. + Qed. Lemma conv_cum_tProj_inv leq Γ p c p' c' : conv_cum leq Σ Γ (tProj p c) (tProj p' c') -> diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index 0529fbcc0..d69cfe779 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -2646,6 +2646,24 @@ Section Conversion. eapply conv_terms_red_conv; eauto. Qed. + Lemma inv_stuck_cofixes leq Γ mfix idx π mfix' idx' π' : + conv_stack_ctx Γ π π' -> + conv_cum leq Σ (Γ,,, stack_context π) (zipp (tCoFix mfix idx) π) (zipp (tCoFix mfix' idx') π') -> + ∥idx = idx' × + All2 (fun d d' => + rarg d = rarg d' × + Σ;;; Γ,,, stack_context π |- dtype d = dtype d' × + Σ;;; Γ,,, stack_context π,,, fix_context mfix |- dbody d = dbody d') + mfix mfix' × + conv_terms Σ (Γ,,, stack_context π) (decompose_stack π).1 (decompose_stack π').1∥. + Proof. + intros [?] cc. + rewrite !zipp_as_mkApps in cc. + apply conv_cum_mkApps_inv in cc as [(conv_cofix&conv_args)]; auto. + apply conv_cum_tCoFix_inv in conv_cofix as [(<-&?)]. + constructor; split; [|split]; auto. + Qed. + (* See https://github.com/coq/coq/blob/master/kernel/reduction.ml#L367 *) Opaque reduce_stack. Equations(noeqns) _isconv_prog (Γ : context) (leq : conv_pb) @@ -2835,7 +2853,7 @@ Section Conversion. } ; | Error e h := no e } ; - | @exist false _ := + | @exist false idx_uneq := no ( DistinctCoFix (Γ ,,, stack_context π1) mfix idx @@ -3715,13 +3733,21 @@ Section Conversion. eapply conv_CoFix. all: assumption. Qed. Next Obligation. - todo "Completeness". + apply h; clear h. + eapply inv_stuck_cofixes in H as [(<-&?&?)]; eauto. + constructor; auto. Qed. Next Obligation. - todo "Completeness". + apply h; clear h. + eapply inv_stuck_cofixes in H as [(<-&?&?)]; eauto. + constructor; auto. + eapply All2_impl; eauto. + cbn; intros; easy. Qed. Next Obligation. - todo "Completeness". + eapply inv_stuck_cofixes in H as [(<-&?&?)]; eauto. + rewrite Nat.eqb_refl in idx_uneq. + congruence. Qed. (* Fallback *) @@ -3921,7 +3947,7 @@ Section Conversion. | @exist (Some (narg, fn)) eq2 with inspect (decompose_stack ρ) := { | @exist (args, ξ) eq' := Some (tCase (ind, par) p (mkApps fn args) brs) } ; - | @exist None eq2 := False_rect _ _ (* why does ! not work in this file? *) + | @exist None eq2 := False_rect _ _ } ; | ccview_other t _ := None } @@ -4797,5 +4823,5 @@ Section Conversion. apply isconv_complete in h. apply h. Qed. Transparent reduce_stack. - + End Conversion. From b5f35d53b26b96d93ee4202557e354f6901e7ed6 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Thu, 5 Nov 2020 16:59:49 +0100 Subject: [PATCH 25/26] Revert some accidental/leftover changes --- safechecker/theories/PCUICSafeConversion.v | 12 +++++------- safechecker/theories/PCUICSafeReduce.v | 6 +++--- 2 files changed, 8 insertions(+), 10 deletions(-) diff --git a/safechecker/theories/PCUICSafeConversion.v b/safechecker/theories/PCUICSafeConversion.v index d69cfe779..1af1e1691 100644 --- a/safechecker/theories/PCUICSafeConversion.v +++ b/safechecker/theories/PCUICSafeConversion.v @@ -16,8 +16,6 @@ Local Set Keyed Unification. Set Default Goal Selector "!". -Set Equations Transparent. - Module PSR := PCUICSafeReduce. (** * Conversion for PCUIC without fuel @@ -3096,11 +3094,11 @@ Section Conversion. assumption. Qed. Next Obligation. - change (eq_inductive ind ind') with (eqb ind ind') in eq5. - destruct (eqb_spec ind ind'). 2: discriminate. - change (Nat.eqb par par') with (eqb par par') in eq5. - destruct (eqb_spec par par'). 2: discriminate. - assumption. + change (eq_inductive ind ind') with (eqb ind ind') in eq5. + destruct (eqb_spec ind ind'). 2: discriminate. + change (Nat.eqb par par') with (eqb par par') in eq5. + destruct (eqb_spec par par'). 2: discriminate. + assumption. Qed. Next Obligation. unshelve eapply R_stateR. diff --git a/safechecker/theories/PCUICSafeReduce.v b/safechecker/theories/PCUICSafeReduce.v index 8206997a0..f73df3b2a 100644 --- a/safechecker/theories/PCUICSafeReduce.v +++ b/safechecker/theories/PCUICSafeReduce.v @@ -1158,10 +1158,10 @@ Section Reduce. (* *) *) (* Abort. *) - Scheme Acc_ind' := Induction for Acc Sort Type. + Scheme Acc_ind' := Induction for Acc Sort Prop. Lemma Fix_F_prop : - forall A R P f (pred : forall x : A, P x -> Type) x hx, + forall A R P f (pred : forall x : A, P x -> Prop) x hx, (forall x aux, (forall y hy, pred y (aux y hy)) -> pred x (f x aux)) -> pred x (@Fix_F A R P f x hx). Proof. @@ -1171,7 +1171,7 @@ Section Reduce. Qed. Lemma reduce_stack_prop : - forall Γ t π h (P : term × stack -> term × stack -> Type), + forall Γ t π h (P : term × stack -> term × stack -> Prop), (forall t π h aux, (forall t' π' hR, P (t', π') (` (aux t' π' hR))) -> P (t, π) (` (_reduce_stack Γ t π h aux))) -> From b644e649606780e7650c429f348a9d31547f8081 Mon Sep 17 00:00:00 2001 From: Jakob Botsch Nielsen Date: Thu, 5 Nov 2020 21:35:16 +0100 Subject: [PATCH 26/26] Clean up unnecessary premise --- erasure/_PluginProject.in | 2 -- pcuic/theories/PCUICNormal.v | 17 +++++------------ 2 files changed, 5 insertions(+), 14 deletions(-) diff --git a/erasure/_PluginProject.in b/erasure/_PluginProject.in index 644924b94..0eeae9d92 100644 --- a/erasure/_PluginProject.in +++ b/erasure/_PluginProject.in @@ -81,8 +81,6 @@ src/pCUICChecker.ml src/pCUICChecker.mli src/pCUICPretty.mli src/pCUICPretty.ml -#src/pCUICSafeLemmata.mli -#src/pCUICSafeLemmata.ml src/templateToPCUIC.mli src/templateToPCUIC.ml diff --git a/pcuic/theories/PCUICNormal.v b/pcuic/theories/PCUICNormal.v index 58424d6e8..e9a8fa340 100644 --- a/pcuic/theories/PCUICNormal.v +++ b/pcuic/theories/PCUICNormal.v @@ -1019,11 +1019,10 @@ Hint Resolve whnf_red_red : pcuic. Lemma whnf_red_mkApps Σ Γ hd hd' args args' : whnf_red Σ Γ hd hd' -> - red Σ Γ hd hd' -> All2 (red Σ Γ) args args' -> whnf_red Σ Γ (mkApps hd args) (mkApps hd' args'). Proof. - intros shape r ra. + intros r ra. induction ra using All2_rect_rev; auto. rewrite <- !mkApps_nested. cbn. @@ -1188,7 +1187,6 @@ Proof. cbn. intros ? ? (?&[= -> -> ->]). auto. - + now apply red1_red, fix_red_ty. + apply All2_same; auto. - apply whnf_red_mkApps; eauto. + constructor. @@ -1196,7 +1194,6 @@ Proof. cbn. intros ? ? (?&[= -> -> ->]). auto. - + now apply red1_red, fix_red_body. + apply All2_same; auto. - constructor; auto. apply All2_same; auto. @@ -1233,16 +1230,14 @@ Proof. constructor. apply All2_same; auto. + apply whnf_red_mkApps. - 3: apply All2_same; auto. - 2: apply red1_red, fix_red_ty; auto. + 2: apply All2_same; auto. constructor. eapply OnOne2_All2; eauto. cbn. intros ? ? (?&[= -> -> ->]). auto. + apply whnf_red_mkApps. - 3: apply All2_same; auto. - 2: apply red1_red, fix_red_body; auto. + 2: apply All2_same; auto. constructor. eapply OnOne2_All2; eauto. cbn. @@ -1255,16 +1250,14 @@ Proof. constructor. apply All2_same; auto. + apply whnf_red_mkApps. - 3: apply All2_same; auto. - 2: apply red1_red, cofix_red_ty; auto. + 2: apply All2_same; auto. constructor. eapply OnOne2_All2; eauto. cbn. intros ? ? (?&[= -> -> ->]). auto. + apply whnf_red_mkApps. - 3: apply All2_same; auto. - 2: apply red1_red, cofix_red_body; auto. + 2: apply All2_same; auto. constructor. eapply OnOne2_All2; eauto. cbn.