@@ -2679,22 +2679,14 @@ Section ConvRedConv.
26792679 Σ ;;; Γ ⊢ tCoFix mfix idx = tCoFix mfix' idx.
26802680 Proof using wfΣ. eapply (ws_cumul_pb_fix_or_cofix (b:=false)). Qed .
26812681
2682- Lemma ws_cumul_pb_eq_le_gen {pb Γ T U} :
2683- Σ ;;; Γ ⊢ T = U ->
2684- Σ ;;; Γ ⊢ T ≤[pb] U.
2685- Proof using Type .
2686- destruct pb => //.
2687- eapply ws_cumul_pb_eq_le.
2688- Qed .
2689-
26902682 Lemma ws_cumul_pb_Lambda_l {Γ na A b na' A' pb} :
26912683 eq_binder_annot na na' ->
26922684 is_open_term (Γ ,, vass na A) b ->
26932685 Σ ;;; Γ ⊢ A = A' ->
26942686 Σ ;;; Γ ⊢ tLambda na A b ≤[pb] tLambda na' A' b.
26952687 Proof using wfΣ.
26962688 intros hna hb h.
2697- eapply ws_cumul_pb_eq_le_gen .
2689+ eapply ws_cumul_eq_pb .
26982690 eapply into_ws_cumul_pb.
26992691 { clear -h hna; induction h.
27002692 - constructor; constructor; auto; reflexivity.
@@ -2705,7 +2697,7 @@ Section ConvRedConv.
27052697 Qed .
27062698
27072699 Lemma ws_cumul_pb_Lambda_r {pb Γ na A b b'} :
2708- Σ ;;; Γ,, vass na A ⊢ b ≤[pb ] b' ->
2700+ Σ ;;; Γ,, vass na A ⊢ b ≤[Conv ] b' ->
27092701 Σ ;;; Γ ⊢ tLambda na A b ≤[pb] tLambda na A b'.
27102702 Proof using wfΣ.
27112703 intros h.
@@ -2731,9 +2723,10 @@ Section ConvRedConv.
27312723 rewrite /on_free_vars_decl /test_decl => /andP[] /= onty ont onu onu'.
27322724 eapply into_ws_cumul_pb => //.
27332725 { clear -h. induction h.
2734- - destruct pb;
2735- eapply cumul_refl; constructor.
2736- all: try reflexivity; auto.
2726+ - eapply cumul_red_l; pcuic.
2727+ eapply cumul_red_r; [|pcuic].
2728+ eapply cumul_refl.
2729+ apply eq_term_upto_univ_subst; trea; tc.
27372730 - destruct pb;
27382731 eapply cumul_red_l; tea; pcuic.
27392732 - destruct pb;
@@ -2743,12 +2736,12 @@ Section ConvRedConv.
27432736 Qed .
27442737
27452738 Lemma ws_cumul_pb_it_mkLambda_or_LetIn_codom {Δ Γ u v pb} :
2746- Σ ;;; (Δ ,,, Γ) ⊢ u ≤[pb ] v ->
2739+ Σ ;;; (Δ ,,, Γ) ⊢ u ≤[Conv ] v ->
27472740 Σ ;;; Δ ⊢ it_mkLambda_or_LetIn Γ u ≤[pb] it_mkLambda_or_LetIn Γ v.
27482741 Proof using wfΣ.
27492742 intros h. revert Δ u v h.
27502743 induction Γ as [| [na [b|] A] Γ ih ] ; intros Δ u v h.
2751- - assumption .
2744+ - by apply ws_cumul_eq_pb .
27522745 - simpl. cbn. eapply ih.
27532746 eapply ws_cumul_pb_LetIn_bo. assumption.
27542747 - simpl. cbn. eapply ih.
@@ -2786,7 +2779,7 @@ Section ConvRedConv.
27862779 Lemma ws_cumul_pb_Lambda {pb Γ na1 na2 A1 A2 t1 t2} :
27872780 eq_binder_annot na1 na2 ->
27882781 Σ ;;; Γ ⊢ A1 = A2 ->
2789- Σ ;;; Γ ,, vass na1 A1 ⊢ t1 ≤[pb ] t2 ->
2782+ Σ ;;; Γ ,, vass na1 A1 ⊢ t1 ≤[Conv ] t2 ->
27902783 Σ ;;; Γ ⊢ tLambda na1 A1 t1 ≤[pb] tLambda na2 A2 t2.
27912784 Proof using wfΣ.
27922785 intros eqna X.
@@ -2800,7 +2793,7 @@ Section ConvRedConv.
28002793 Lemma conv_cum_Lambda leq Γ na1 na2 A1 A2 t1 t2 :
28012794 eq_binder_annot na1 na2 ->
28022795 Σ ;;; Γ ⊢ A1 = A2 ->
2803- sq_ws_cumul_pb leq Σ (Γ ,, vass na1 A1) t1 t2 ->
2796+ sq_ws_cumul_pb Conv Σ (Γ ,, vass na1 A1) t1 t2 ->
28042797 sq_ws_cumul_pb leq Σ Γ (tLambda na1 A1 t1) (tLambda na2 A2 t2).
28052798 Proof using wfΣ.
28062799 intros eqna X []; sq. now apply ws_cumul_pb_Lambda.
@@ -2852,7 +2845,7 @@ Section ConvRedConv.
28522845 intros hna ont ona onu.
28532846 etransitivity.
28542847 { eapply ws_cumul_pb_LetIn_bo; tea. }
2855- eapply ws_cumul_pb_eq_le_gen .
2848+ eapply ws_cumul_eq_pb .
28562849 etransitivity.
28572850 { eapply ws_cumul_pb_LetIn_ty; tea; eauto with fvs. }
28582851 eapply ws_cumul_pb_LetIn_tm; tea; eauto with fvs.
@@ -2861,12 +2854,13 @@ Section ConvRedConv.
28612854
28622855 Lemma ws_cumul_pb_it_mkLambda_or_LetIn {pb Γ Δ1 Δ2 t1 t2} :
28632856 Σ ⊢ Γ ,,, Δ1 = Γ ,,, Δ2 ->
2864- Σ ;;; Γ ,,, Δ1 ⊢ t1 ≤[pb] t2 ->
2857+ Σ ;;; Γ ,,, Δ1 ⊢ t1 = t2 ->
28652858 Σ ;;; Γ ⊢ it_mkLambda_or_LetIn Δ1 t1 ≤[pb] it_mkLambda_or_LetIn Δ2 t2.
28662859 Proof using wfΣ.
28672860 induction Δ1 in Δ2, t1, t2 |- *; intros X Y.
28682861 - apply All2_fold_length in X.
28692862 destruct Δ2; cbn in *; [trivial|].
2863+ 1: by apply ws_cumul_eq_pb.
28702864 rewrite app_length in X; lia.
28712865 - apply All2_fold_length in X as X'.
28722866 destruct Δ2 as [|c Δ2]; simpl in *; [rewrite app_length in X'; lia|].
@@ -2911,7 +2905,7 @@ Section ConvRedConv.
29112905 Lemma ws_cumul_pb_Lambda_inv :
29122906 forall pb Γ na1 na2 A1 A2 b1 b2,
29132907 Σ ;;; Γ ⊢ tLambda na1 A1 b1 ≤[pb] tLambda na2 A2 b2 ->
2914- [× eq_binder_annot na1 na2, Σ ;;; Γ ⊢ A1 = A2 & Σ ;;; Γ ,, vass na1 A1 ⊢ b1 ≤[pb] b2].
2908+ [× eq_binder_annot na1 na2, Σ ;;; Γ ⊢ A1 = A2 & Σ ;;; Γ ,, vass na1 A1 ⊢ b1 = b2].
29152909 Proof using wfΣ.
29162910 intros *.
29172911 move/ws_cumul_pb_red; intros (v & v' & [redv redv' eq]).
@@ -2946,7 +2940,7 @@ Section ConvRedConv.
29462940 Lemma Lambda_conv_cum_inv :
29472941 forall leq Γ na1 na2 A1 A2 b1 b2,
29482942 sq_ws_cumul_pb leq Σ Γ (tLambda na1 A1 b1) (tLambda na2 A2 b2) ->
2949- eq_binder_annot na1 na2 /\ ∥ Σ ;;; Γ ⊢ A1 = A2 ∥ /\ sq_ws_cumul_pb leq Σ (Γ ,, vass na1 A1) b1 b2.
2943+ eq_binder_annot na1 na2 /\ ∥ Σ ;;; Γ ⊢ A1 = A2 ∥ /\ sq_ws_cumul_pb Conv Σ (Γ ,, vass na1 A1) b1 b2.
29502944 Proof using wfΣ.
29512945 intros * []. eapply ws_cumul_pb_Lambda_inv in X as [].
29522946 intuition auto. all:sq; auto.
@@ -3252,7 +3246,7 @@ Section ConvSubst.
32523246 move: clctx. rewrite on_free_vars_ctx_app !app_context_length H => /andP[] //. }
32533247 etransitivity.
32543248 * eapply untyped_substitution_ws_cumul_pb; tea. fvs.
3255- * eapply ws_cumul_pb_eq_le_gen .
3249+ * eapply ws_cumul_eq_pb .
32563250 eapply (untyped_substitution_ws_cumul_pb_subst_conv (Δ := Γ0) (Δ' := Γ1)); tea; eauto.
32573251 Qed .
32583252
@@ -3552,7 +3546,7 @@ Qed.
35523546
35533547Lemma ws_cumul_pb_rel_it_mkLambda_or_LetIn {cf pb Σ} {wfΣ : wf Σ} (Δ Γ Γ' : context) (B B' : term) :
35543548 ws_cumul_ctx_pb_rel Conv Σ Δ Γ Γ' ->
3555- Σ ;;; Δ ,,, Γ ⊢ B ≤[pb] B' ->
3549+ Σ ;;; Δ ,,, Γ ⊢ B = B' ->
35563550 Σ ;;; Δ ⊢ it_mkLambda_or_LetIn Γ B ≤[pb] it_mkLambda_or_LetIn Γ' B'.
35573551Proof .
35583552 move/ws_cumul_ctx_pb_rel_app => hc hb.
@@ -3682,7 +3676,7 @@ Section CumulSubst.
36823676 Σ ;;; Γ ,,, subst_context s 0 Γ' ⊢ subst s #|Γ'| b ≤[pb] subst s' #|Γ'| b.
36833677 Proof using wfΣ.
36843678 move=> cl cl' clb eqsub subs subs'.
3685- eapply ws_cumul_pb_eq_le_gen .
3679+ eapply ws_cumul_eq_pb .
36863680 eapply substitution_ws_cumul_pb_subst_conv; tea; eauto with pcuic.
36873681 Qed .
36883682
@@ -3952,7 +3946,7 @@ Proof.
39523946 eapply (ws_cumul_pb_refl' (exist Γ _) (exist t _)) end.
39533947 all: intros Γ pb; revert Γ.
39543948 - intros; etransitivity; eauto.
3955- - intros. apply ws_cumul_pb_eq_le_gen . apply symmetry.
3949+ - intros. apply ws_cumul_eq_pb . apply symmetry.
39563950 eauto.
39573951 - intros Γ t; intros. unshelve eapply (ws_cumul_pb_refl' (exist Γ _) (exist t _)); eauto.
39583952 - intros Γ ev args args' Hargsargs' Hargsargs'_dep HΓ Hargs Hargs'. cbn in *. eapply ws_cumul_pb_Evar; eauto.
@@ -3974,7 +3968,7 @@ Proof.
39743968 - intros Γ na na' t t' ty ty' u u' Hna _ Heqtt' _ Heqtyty' _ Hequu' HΓ HM HN.
39753969 cbn in *. apply andb_andI in HM; apply andb_andI in HN; destruct HM as [Ht Htyu]; destruct HN as [Ht' Htyu'].
39763970 apply andb_andI in Htyu; apply andb_andI in Htyu'; destruct Htyu as [Hty Hu]; destruct Htyu' as [Hty' Hu'].
3977- eapply ws_cumul_pb_LetIn; eauto. eapply Hequu'; eauto.
3971+ eapply ws_cumul_pb_LetIn; eauto. apply ws_cumul_eq_pb. eapply Hequu'; eauto.
39783972 * change (is_closed_context (Γ,, vdef na t ty)). rewrite on_free_vars_ctx_snoc. apply andb_and. split; eauto.
39793973 rewrite /on_free_vars_decl /test_decl. apply andb_and. split; eauto.
39803974 * rewrite shiftnP_S; eauto.
@@ -3984,7 +3978,7 @@ Proof.
39843978 apply andb_andI in H; apply andb_andI in H'; destruct H as [Hreturn H]; destruct H' as [Hreturn' H'].
39853979 apply andb_andI in H; apply andb_andI in H'; destruct H as [Hcontext H]; destruct H' as [Hcontext' H'].
39863980 apply andb_andI in H; apply andb_andI in H'; destruct H as [Hc Hbrs]; destruct H' as [Hc' Hbrs'].
3987- eapply ws_cumul_pb_eq_le_gen . eapply ws_cumul_pb_Case; eauto.
3981+ eapply ws_cumul_eq_pb . eapply ws_cumul_pb_Case; eauto.
39883982 * rewrite is_open_case_split. repeat (apply andb_and; split); eauto.
39893983 * rewrite is_open_case_split. repeat (apply andb_and; split); eauto.
39903984 * unfold cumul_predicate in Hpp'. unfold ws_cumul_pb_predicate. destruct Hpp' as [Hpp' [Hinst [Hpcon Hpret]]].
@@ -4013,7 +4007,7 @@ Proof.
40134007 rewrite <- app_length in *. tea.
40144008 - intros; eapply ws_cumul_pb_Proj_c; eauto.
40154009 - intros Γ mfix mfix' idx Hmfixmfix' Hmfixmfix'_dep HΓ H H'. cbn in *.
4016- eapply ws_cumul_pb_eq_le_gen . eapply ws_cumul_pb_Fix; eauto. repeat toAll.
4010+ eapply ws_cumul_eq_pb . eapply ws_cumul_pb_Fix; eauto. repeat toAll.
40174011 eapply All2_impl. 1: tea. cbn; intros; destruct_head'_prod.
40184012 pose proof (Hfix := All2_length ltac:(eassumption)).
40194013 unfold test_def in *; repeat toProp; destruct_head'_and.
@@ -4023,7 +4017,7 @@ Proof.
40234017 * rewrite -> shiftnP_add, <- fix_context_length, <- app_length in *; tea.
40244018 * rewrite -> shiftnP_add, <- Hfix, <- fix_context_length, <- app_length in *; tea.
40254019 - intros Γ mfix mfix' idx Hmfixmfix' Hmfixmfix'_dep HΓ H H'. cbn in *.
4026- eapply ws_cumul_pb_eq_le_gen . eapply ws_cumul_pb_CoFix; eauto. repeat toAll.
4020+ eapply ws_cumul_eq_pb . eapply ws_cumul_pb_CoFix; eauto. repeat toAll.
40274021 eapply All2_impl. 1: tea. pose proof (Hfix := All2_length ltac:(eassumption)); cbn; intros. destruct_head'_prod.
40284022 unfold test_def in *.
40294023 repeat toProp; destruct_head'_and.
0 commit comments