@@ -8,7 +8,7 @@ From MetaCoq.PCUIC Require Import PCUICAst PCUICAstUtils
88 PCUICWeakeningEnv PCUICWeakeningEnvTyp
99 PCUICWellScopedCumulativity
1010 PCUICContextConversion PCUICConversion PCUICCanonicity
11- PCUICSpine PCUICInductives PCUICInductiveInversion PCUICConfluence
11+ PCUICSpine PCUICInductives PCUICInductiveInversion PCUICConfluence
1212 PCUICArities PCUICPrincipality.
1313
1414From MetaCoq.Erasure Require Import Extract .
@@ -60,7 +60,7 @@ Lemma isArity_ind_type (Σ : global_env_ext) mind ind idecl :
6060 declared_inductive (fst Σ) ind mind idecl ->
6161 isArity (ind_type idecl).
6262Proof .
63- intros.
63+ intros.
6464 eapply (declared_inductive_inv weaken_env_prop_typing) in H; eauto.
6565 - inv H. rewrite ind_arity_eq.
6666 change PCUICEnvironment.it_mkProd_or_LetIn with it_mkProd_or_LetIn.
@@ -108,14 +108,14 @@ Proof.
108108 - eapply IHL in H. cbn in H. tauto.
109109Qed .
110110
111- Lemma typing_spine_red (Σ : global_env_ext) Γ (args args' : list PCUICAst.term)
111+ Lemma typing_spine_red (Σ : global_env_ext) Γ (args args' : list PCUICAst.term)
112112 (X : All2 (red Σ Γ) args args') (wfΣ : wf Σ)
113- (T x x0 : PCUICAst.term)
114- (t0 : typing_spine Σ Γ x args x0)
113+ (T x x0 : PCUICAst.term)
114+ (t0 : typing_spine Σ Γ x args x0)
115115 (c : Σ;;; Γ ⊢ x0 ≤ T) x1
116116 (c0 : Σ;;; Γ ⊢ x1 ≤ x) :
117117 isType Σ Γ x1 ->
118- isType Σ Γ T ->
118+ isType Σ Γ T ->
119119 typing_spine Σ Γ x1 args' T.
120120Proof .
121121 intros ? ?. revert args' X.
@@ -128,7 +128,7 @@ Proof.
128128 + eapply IHt0; eauto.
129129 eapply red_ws_cumul_pb_inv.
130130 unfold subst1.
131- eapply isType_tProd in i0 as [dom codom].
131+ eapply isType_tProd in i0 as [dom codom].
132132 eapply (closed_red_red_subst (Δ := [vass na A]) (Γ' := [])); auto.
133133 simpl. eapply isType_wf_local in codom. fvs.
134134 constructor; auto. eapply into_closed_red; auto. fvs. fvs.
@@ -137,11 +137,11 @@ Proof.
137137 eapply subject_reduction; tea.
138138Qed .
139139
140- Lemma it_mkProd_red_Arity {Σ : global_env_ext} {Γ c0 i u l} {wfΣ : wf Σ} :
140+ Lemma it_mkProd_red_Arity {Σ : global_env_ext} {Γ c0 i u l} {wfΣ : wf Σ} :
141141 ~ Is_conv_to_Arity Σ Γ (it_mkProd_or_LetIn c0 (mkApps (tInd i u) l)).
142142Proof .
143143 intros (? & [] & ?). eapply red_it_mkProd_or_LetIn_mkApps_Ind in X as (? & ? & ?). subst.
144- eapply it_mkProd_arity in H. eapply isArity_mkApps in H as [[] ].
144+ eapply it_mkProd_arity in H. eapply isArity_mkApps in H as [[] ].
145145Qed .
146146
147147Lemma invert_it_Ind_eq_prod:
@@ -203,7 +203,7 @@ Proof.
203203 * rewrite /mkProd_or_LetIn /=. simpl => /= sp.
204204 simpl.
205205 dependent elimination sp as [spnil i i' e|spcons i i' e e' sp].
206- { exists (Γ0 ++ [vass na ty]).
206+ { exists (Γ0 ++ [vass na ty]).
207207 exists args. now rewrite it_mkProd_or_LetIn_app. }
208208 eapply ws_cumul_pb_Prod_Prod_inv in e as [eqna dom codom]; pcuic.
209209 eapply (substitution0_ws_cumul_pb (t:=hd0)) in codom; eauto.
@@ -255,7 +255,7 @@ Proof.
255255 eapply PCUICValidity.validity. econstructor; eauto.
256256Qed .
257257
258- Lemma nIs_conv_to_Arity_nArity {Σ : global_env_ext} {wfΣ : wf Σ} {Γ T} :
258+ Lemma nIs_conv_to_Arity_nArity {Σ : global_env_ext} {wfΣ : wf Σ} {Γ T} :
259259 isType Σ Γ T ->
260260 ~ Is_conv_to_Arity Σ Γ T -> ~ isArity T.
261261Proof .
@@ -270,12 +270,12 @@ Lemma tConstruct_no_Type (Σ : global_env_ext) Γ ind c u x1 : wf Σ ->
270270 Is_proof Σ Γ (mkApps (tConstruct ind c u) x1).
271271Proof .
272272 intros wfΣ (? & ? & [ | (? & ? & ?)]).
273- - exfalso.
273+ - exfalso.
274274 eapply nIs_conv_to_Arity_nArity; tea.
275275 eapply PCUICValidity.validity; tea.
276276 eapply type_mkApps_tConstruct_n_conv_arity in t; auto.
277277 - exists x, x0. eauto.
278- Qed .
278+ Qed .
279279
280280(* if a cofixpoint is a type or proof, it is a proof *)
281281
316316Lemma typing_spine_wat (Σ : global_env_ext) (Γ : context) (L : list term)
317317 (x x0 : term) :
318318 wf Σ ->
319- typing_spine Σ Γ x L x0 ->
319+ typing_spine Σ Γ x L x0 ->
320320 isType Σ Γ x0.
321321Proof .
322322 intros wfΣ; induction 1; auto.
@@ -394,8 +394,8 @@ Lemma sort_typing_spine:
394394 forall (Σ : global_env_ext) (Γ : context) (L : list term) (u : Universe.t) (x x0 : term),
395395 wf_ext Σ ->
396396 is_propositional u ->
397- typing_spine Σ Γ x L x0 ->
398- Σ;;; Γ |- x : tSort u ->
397+ typing_spine Σ Γ x L x0 ->
398+ Σ;;; Γ |- x : tSort u ->
399399 ∑ u', Σ;;; Γ |- x0 : tSort u' × is_propositional u'.
400400Proof .
401401 intros Σ Γ L u x x0 HΣ ? t1 c0.
422422Lemma arity_type_inv (Σ : global_env_ext) Γ t T1 T2 : wf_ext Σ -> wf_local Σ Γ ->
423423 Σ ;;; Γ |- t : T1 -> isArity T1 -> Σ ;;; Γ |- t : T2 -> Is_conv_to_Arity Σ Γ T2.
424424Proof .
425- intros wfΣ wfΓ. intros.
426- destruct (common_typing _ _ X X0) as (? & e & ? & ?).
425+ intros wfΣ wfΓ. intros.
426+ destruct (common_typing _ _ X X0) as (? & e & ? & ?).
427427 eapply invert_cumul_arity_l_gen; tea.
428428 eapply invert_cumul_arity_r_gen. 2:exact e.
429429 exists T1. split; auto. sq.
@@ -467,7 +467,7 @@ Lemma leq_term_propositional_sorted_l {Σ Γ v v' u u'} :
467467 wf_ext Σ ->
468468 PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' ->
469469 Σ;;; Γ |- v : tSort u ->
470- Σ;;; Γ |- v' : tSort u' -> is_propositional u ->
470+ Σ;;; Γ |- v' : tSort u' -> is_propositional u ->
471471 leq_universe (global_ext_constraints Σ) u' u.
472472Proof .
473473 intros wf leq Hv Hv' isp.
@@ -481,7 +481,7 @@ Lemma leq_term_propopositional_sorted_r {Σ Γ v v' u u'} :
481481 wf_ext Σ ->
482482 PCUICEquality.leq_term Σ (global_ext_constraints Σ) v v' ->
483483 Σ;;; Γ |- v : tSort u ->
484- Σ;;; Γ |- v' : tSort u' -> is_propositional u' ->
484+ Σ;;; Γ |- v' : tSort u' -> is_propositional u' ->
485485 leq_universe (global_ext_constraints Σ) u u'.
486486Proof .
487487 intros wfΣ leq hv hv' isp.
@@ -579,15 +579,15 @@ Proof.
579579 destruct s as [ | (u & ? & ?)].
580580 - eapply invert_cumul_arity_r in e; eauto. destruct e as (? & [] & ?).
581581 eapply invert_red_prod in X1 as (? & ? & []); eauto; subst. cbn in H.
582- econstructor. exists x3. econstructor.
582+ econstructor. exists x3. econstructor.
583583 eapply type_reduction_closed; eauto. econstructor; eauto.
584584 - sq. eapply cumul_prop1' in e; eauto.
585585 eapply inversion_Prod in e as (? & ? & ? & ? & e) ; auto.
586586 eapply ws_cumul_pb_Sort_inv in e.
587587 eapply leq_universe_propositional_r in e as H0; cbn; eauto.
588588 eexists. split. eassumption. right. eexists. split. eassumption.
589589 eapply is_propositional_sort_prod in H0; eauto.
590- eapply type_Lambda in t1; eauto.
590+ eapply type_Lambda in t1; eauto.
591591 now apply PCUICValidity.validity in t1.
592592Qed .
593593
@@ -614,7 +614,7 @@ Proof.
614614 eapply wcbeval_red; eauto. assumption.
615615Qed .
616616
617- (* Thanks to the restriction to Prop </= Type, erasability is also closed by expansion
617+ (* Thanks to the restriction to Prop </= Type, erasability is also closed by expansion
618618 on well-typed terms. *)
619619
620620Lemma Is_type_eval_inv (Σ : global_env_ext) t v:
@@ -646,7 +646,7 @@ Proof.
646646 intros [s hs]; eapply wt_closed_red_refl; tea.
647647Qed .
648648
649- Lemma nIs_conv_to_Arity_isWfArity_elim {Σ} {wfΣ : wf Σ} {Γ x} :
649+ Lemma nIs_conv_to_Arity_isWfArity_elim {Σ} {wfΣ : wf Σ} {Γ x} :
650650 ~ Is_conv_to_Arity Σ Γ x ->
651651 isWfArity Σ Γ x ->
652652 False.
@@ -659,11 +659,11 @@ Proof.
659659 now eapply it_mkProd_isArity.
660660Qed .
661661
662- Definition isErasable_Type (Σ : global_env_ext) Γ T :=
662+ Definition isErasable_Type (Σ : global_env_ext) Γ T :=
663663 (Is_conv_to_Arity Σ Γ T +
664664 (∑ u : Universe.t, Σ;;; Γ |- T : tSort u × is_propositional u))%type.
665665
666- Lemma isErasable_any_type {Σ} {wfΣ : wf_ext Σ} {Γ t T} :
666+ Lemma isErasable_any_type {Σ} {wfΣ : wf_ext Σ} {Γ t T} :
667667 isErasable Σ Γ t ->
668668 Σ ;;; Γ |- t : T ->
669669 isErasable_Type Σ Γ T.
@@ -680,12 +680,12 @@ Proof.
680680 eapply cumul_prop1'; eauto. eapply PCUICValidity.validity; eauto.
681681Qed .
682682
683- Lemma Is_proof_ty Σ Γ t :
683+ Lemma Is_proof_ty Σ Γ t :
684684 wf_ext Σ ->
685- Is_proof Σ Γ t ->
686- forall t' ty,
685+ Is_proof Σ Γ t ->
686+ forall t' ty,
687687 Σ ;;; Γ |- t : ty ->
688- Σ ;;; Γ |- t' : ty ->
688+ Σ ;;; Γ |- t' : ty ->
689689 Is_proof Σ Γ t'.
690690Proof .
691691 intros wfΣ [ty [u [Hty isp]]].
@@ -750,7 +750,7 @@ Proof.
750750 eapply leq_universe_sprop_l in leu; tea => //.
751751Qed .
752752
753- Lemma typing_spine_inj {Σ Γ Δ s args args' u u'} :
753+ Lemma typing_spine_inj {Σ Γ Δ s args args' u u'} :
754754 wf_ext Σ ->
755755 check_univs ->
756756 prop_sub_type = false ->
@@ -765,12 +765,12 @@ Proof.
765765 eapply is_propositional_lower; tea. apply wf.
766766Qed .
767767
768- Lemma Is_proof_ind Σ Γ t :
768+ Lemma Is_proof_ind Σ Γ t :
769769 wf_ext Σ ->
770- Is_proof Σ Γ t ->
771- forall t' ind u args args',
770+ Is_proof Σ Γ t ->
771+ forall t' ind u args args',
772772 Σ ;;; Γ |- t : mkApps (tInd ind u) args ->
773- Σ ;;; Γ |- t' : mkApps (tInd ind u) args' ->
773+ Σ ;;; Γ |- t' : mkApps (tInd ind u) args' ->
774774 Is_proof Σ Γ t'.
775775Proof .
776776 intros wfΣ [ty [u [Hty isp]]].
@@ -806,21 +806,21 @@ Proof.
806806 eapply inversion_Case in hc as [mdecl [idecl [isdecl [indices ?]]]]; eauto.
807807 eapply inversion_Case in hr as [mdecl' [idecl' [isdecl' [indices' ?]]]]; eauto.
808808 destruct (declared_inductive_inj isdecl isdecl'). subst mdecl' idecl'.
809- intros hp.
809+ intros hp.
810810 epose proof (Is_proof_ind _ _ _ wfΣ hp).
811811 destruct p0 as [[] ?]. destruct p1 as [[] ?].
812812 exact (X _ _ _ _ _ scrut_ty scrut_ty0).
813813Qed .
814814
815815Lemma Is_proof_app {Σ Γ t args ty} {wfΣ : wf_ext Σ} :
816- Is_proof Σ Γ t ->
816+ Is_proof Σ Γ t ->
817817 Σ ;;; Γ |- mkApps t args : ty ->
818818 Is_proof Σ Γ (mkApps t args).
819819Proof .
820820 intros [ty' [u [Hty [isp pu]]]] Htargs.
821821 eapply PCUICValidity.inversion_mkApps in Htargs as [A [Ht sp]].
822- pose proof (PCUICValidity.validity Hty).
823- pose proof (PCUICValidity.validity Ht).
822+ pose proof (PCUICValidity.validity Hty).
823+ pose proof (PCUICValidity.validity Ht).
824824 epose proof (PCUICPrincipality.common_typing _ wfΣ Hty Ht) as [C [Cty [Cty' Ht'']]].
825825 eapply PCUICSpine.typing_spine_strengthen in sp. 3:tea.
826826 edestruct (sort_typing_spine _ _ _ u _ _ _ pu sp) as [u' [Hty' isp']].
@@ -831,7 +831,7 @@ Proof.
831831 now eapply validity.
832832Qed .
833833
834- Lemma isErasable_Propositional {Σ : global_env_ext} {Γ ind n u args} :
834+ Lemma isErasable_Propositional {Σ : global_env_ext} {Γ ind n u args} :
835835 wf_ext Σ ->
836836 isErasable Σ Γ (mkApps (tConstruct ind n u) args) -> isPropositional Σ ind true.
837837Proof .
@@ -863,7 +863,7 @@ Proof.
863863 eapply validity. econstructor; tea.
864864Qed .
865865
866- Lemma nisErasable_Propositional {Σ : global_env_ext} {Γ ind n u} :
866+ Lemma nisErasable_Propositional {Σ : global_env_ext} {Γ ind n u} :
867867 wf_ext Σ ->
868868 welltyped Σ Γ (tConstruct ind n u) ->
869869 (isErasable Σ Γ (tConstruct ind n u) -> False) -> isPropositional Σ ind false.
@@ -887,7 +887,7 @@ Proof.
887887 rewrite onc.(cstr_eq) in e, X.
888888 rewrite !subst_instance_it_mkProd_or_LetIn !PCUICLiftSubst.subst_it_mkProd_or_LetIn in e, X.
889889 len in e; len in X.
890- rewrite subst_cstr_concl_head in e, X.
890+ rewrite subst_cstr_concl_head in e, X.
891891 destruct decli. eapply nth_error_Some_length in H1; eauto.
892892 rewrite -it_mkProd_or_LetIn_app in e, X.
893893 exists (subst_instance_univ u (ind_sort x0)).
@@ -910,9 +910,9 @@ Proof.
910910 do 2 constructor.
911911 rewrite is_propositional_subst_instance in sorts, sorts' |- *.
912912 specialize (sorts' isp). rewrite -sorts'. reflexivity.
913- Qed .
913+ Qed .
914914
915- Lemma isPropositional_propositional Σ (Σ' : E.global_context) ind mdecl idecl mdecl' idecl' :
915+ Lemma isPropositional_propositional Σ (Σ' : E.global_context) ind mdecl idecl mdecl' idecl' :
916916 PCUICAst.declared_inductive Σ ind mdecl idecl ->
917917 EGlobalEnv.declared_inductive Σ' ind mdecl' idecl' ->
918918 erases_mutual_inductive_body mdecl mdecl' ->
@@ -929,14 +929,14 @@ Proof.
929929 rewrite isP. intros ->. f_equal. f_equal. now rewrite indp.
930930Qed .
931931
932- Lemma isPropositional_propositional_cstr Σ (Σ' : E.global_context) ind c mdecl idecl cdecl mdecl' idecl' :
932+ Lemma isPropositional_propositional_cstr Σ (Σ' : E.global_context) ind c mdecl idecl cdecl mdecl' idecl' :
933933 wf Σ ->
934934 PCUICAst.declared_constructor Σ (ind, c) mdecl idecl cdecl ->
935935 EGlobalEnv.declared_inductive Σ' ind mdecl' idecl' ->
936936 erases_mutual_inductive_body mdecl mdecl' ->
937937 erases_one_inductive_body idecl idecl' ->
938- forall b, isPropositional Σ ind b ->
939- EGlobalEnv.constructor_isprop_pars_decl Σ' ind c =
938+ forall b, isPropositional Σ ind b ->
939+ EGlobalEnv.constructor_isprop_pars_decl Σ' ind c =
940940 Some (b, mdecl.(ind_npars), EAst.mkConstructor cdecl.(cstr_name) (context_assumptions cdecl.(cstr_args))).
941941Proof .
942942 intros wfΣ declc decli' em ei b isp.
@@ -959,15 +959,15 @@ Qed.
959959Lemma eval_tCase {cf : checker_flags} {Σ : global_env_ext} ci p discr brs res T :
960960 wf Σ ->
961961 Σ ;;; [] |- tCase ci p discr brs : T ->
962- eval Σ (tCase ci p discr brs) res ->
962+ eval Σ (tCase ci p discr brs) res ->
963963 ∑ c u args, PCUICReduction.red Σ [] (tCase ci p discr brs) (tCase ci p ((mkApps (tConstruct ci.(ci_ind) c u) args)) brs).
964964Proof .
965965 intros wf wt H. depind H; try now (cbn in *; congruence).
966966 - eapply inversion_Case in wt as (? & ? & ? & ? & cinv & ?); eauto.
967967 eexists _, _, _. eapply PCUICReduction.red_case_c. eapply wcbeval_red. 2: eauto. eapply cinv.
968968 - eapply inversion_Case in wt as wt'; eauto. destruct wt' as (? & ? & ? & ? & cinv & ?).
969969 assert (Hred1 : PCUICReduction.red Σ [] (tCase ip p discr brs) (tCase ip p (mkApps fn args) brs)). {
970- etransitivity. { eapply PCUICReduction.red_case_c. eapply wcbeval_red. 2: eauto. eapply cinv. }
970+ etransitivity. { eapply PCUICReduction.red_case_c. eapply wcbeval_red. 2: eauto. eapply cinv. }
971971 econstructor. econstructor.
972972 rewrite closed_unfold_cofix_cunfold_eq. eauto.
973973 enough (closed (mkApps (tCoFix mfix idx) args)) as Hcl by (rewrite closedn_mkApps in Hcl; solve_all).
@@ -1000,7 +1000,7 @@ Proof.
10001000Qed .
10011001
10021002Lemma isErasable_unfold_cofix {Σ : global_env_ext} {Γ mfix idx} {wfΣ : wf Σ} decl :
1003- isErasable Σ Γ (tCoFix mfix idx) ->
1003+ isErasable Σ Γ (tCoFix mfix idx) ->
10041004 nth_error mfix idx = Some decl ->
10051005 isErasable Σ Γ (subst0 (cofix_subst mfix) (dbody decl)).
10061006Proof .
0 commit comments