772
772
case: eqb_spec.
773
773
intros ->.
774
774
eapply lookup_global_Some_fresh in hl.
775
- rewrite e in wf. destruct wf as [_ ond]; depelim ond.
775
+ rewrite e in wf. destruct wf as [_ ond]; depelim ond. destruct o as [f ? ? ? ].
776
776
cbn in *. eapply Forall_app in f as []. contradiction.
777
777
intros _.
778
778
eapply IHx; trea.
@@ -908,7 +908,7 @@ Next Obligation.
908
908
{ now eexists. }
909
909
destruct (abstract_pop_decls_correct X decls prop' _ _ HX H) as [? []].
910
910
clear H. specialize (prop _ HX). destruct x, Σ, H0; cbn in *.
911
- subst. sq. destruct wfX. depelim o0. split => //.
911
+ subst. sq. destruct wfX. depelim o0. destruct o1. split => //.
912
912
Qed .
913
913
Next Obligation .
914
914
pose proof (abstract_env_ext_wf _ H) as [wf].
@@ -922,7 +922,7 @@ Next Obligation.
922
922
pose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ HX' H).
923
923
clear H HX'. specialize (prop _ HX). destruct x, Σ as [[] u], H0; cbn in *.
924
924
subst. sq. inversion H3. subst. clear H3. destruct wfX. cbn in *.
925
- rewrite prop in o0. depelim o0. cbn in o2. apply o2 .
925
+ rewrite prop in o0. depelim o0. destruct o1. exact on_global_decl_d .
926
926
Qed .
927
927
Next Obligation .
928
928
pose proof (abstract_env_exists X) as [[? HX]].
@@ -1218,7 +1218,7 @@ Proof.
1218
1218
induction er using erases_deps_forall_ind; try solve [now constructor].
1219
1219
- assert (kn <> kn0).
1220
1220
{ inv wfΣ. inv X. intros <-.
1221
- eapply lookup_env_Some_fresh in H. contradiction. }
1221
+ eapply lookup_env_Some_fresh in H. destruct X1. contradiction. }
1222
1222
eapply erases_deps_tConst with cb cb'; eauto.
1223
1223
red. rewrite /lookup_env lookup_env_cons_fresh //.
1224
1224
red.
@@ -1240,15 +1240,15 @@ Proof.
1240
1240
red in H. red.
1241
1241
inv wfΣ. inv X.
1242
1242
rewrite -H. simpl. unfold lookup_env; simpl. destruct (eqb_spec (inductive_mind p.1) kn); try congruence.
1243
- eapply lookup_env_Some_fresh in H. subst kn; contradiction.
1243
+ eapply lookup_env_Some_fresh in H. destruct X1. subst kn; contradiction.
1244
1244
- econstructor; eauto.
1245
1245
red. destruct H. split; eauto.
1246
1246
destruct d0; split; eauto.
1247
1247
destruct d0; split; eauto.
1248
1248
inv wfΣ. inv X.
1249
1249
red in H |- *.
1250
1250
rewrite -H. simpl. unfold lookup_env; simpl; destruct (eqb_spec (inductive_mind p.(proj_ind)) kn); try congruence.
1251
- eapply lookup_env_Some_fresh in H. subst kn. contradiction.
1251
+ eapply lookup_env_Some_fresh in H. subst kn. destruct X1. contradiction.
1252
1252
Qed .
1253
1253
1254
1254
Lemma lookup_env_ext {Σ kn kn' d d'} :
@@ -1259,7 +1259,7 @@ Proof.
1259
1259
intros wf hl.
1260
1260
eapply lookup_env_Some_fresh in hl.
1261
1261
inv wf. inv X.
1262
- destruct (eqb_spec kn kn'); subst; congruence.
1262
+ destruct (eqb_spec kn kn'); subst; destruct X1; congruence.
1263
1263
Qed .
1264
1264
1265
1265
Lemma lookup_env_cons_disc {Σ kn kn' d} :
@@ -1288,11 +1288,11 @@ Proof.
1288
1288
exists cst. split.
1289
1289
red in declc |- *. unfold lookup_env in *.
1290
1290
rewrite lookup_env_cons_fresh //.
1291
- { eapply lookup_env_Some_fresh in declc.
1291
+ { eapply lookup_env_Some_fresh in declc. destruct X1.
1292
1292
intros <-; contradiction. }
1293
1293
exists cst'.
1294
1294
unfold EGlobalEnv.declared_constant. rewrite EGlobalEnv.elookup_env_cons_fresh //.
1295
- { eapply lookup_env_Some_fresh in declc.
1295
+ { eapply lookup_env_Some_fresh in declc. destruct X1.
1296
1296
intros <-; contradiction. }
1297
1297
red in ebody. unfold erases_constant_body.
1298
1298
destruct (cst_body cst) eqn:bod; destruct (E.cst_body cst') eqn:bod' => //.
@@ -1327,7 +1327,7 @@ Proof.
1327
1327
unfold lookup_env in *.
1328
1328
rewrite lookup_env_cons_fresh //.
1329
1329
{ eapply lookup_env_Some_fresh in declc.
1330
- intros <-. contradiction. }
1330
+ intros <-. destruct X1. contradiction. }
1331
1331
exists cst'.
1332
1332
unfold EGlobalEnv.declared_constant.
1333
1333
red in ebody. unfold erases_constant_body.
@@ -1462,7 +1462,9 @@ Proof.
1462
1462
cbn in o0, o. rewrite prf' in o0. rewrite <- Hpop in o0. rewrite Hpop' in o. clear -o o0.
1463
1463
now depelim o0.
1464
1464
depelim wf. rewrite Hpop' in o0.
1465
- cbn in o0, o. rewrite prf' in o0. rewrite <- Hpop in o0. clear -o0. now depelim o0. }
1465
+ cbn in o0, o. rewrite prf' in o0. rewrite <- Hpop in o0. clear -o0. depelim o0.
1466
+ now destruct o.
1467
+ }
1466
1468
all: eauto.
1467
1469
apply IHdecls; eauto.
1468
1470
{ intros. pose proof (abstract_env_wf _ wfpop) as [wf'].
@@ -1603,11 +1605,12 @@ Lemma on_global_env_ind (P : forall Σ : global_env, wf Σ -> Type)
1603
1605
(pd : on_global_decl cumulSpec0 (lift_typing typing)
1604
1606
({| universes := Σ.(universes); declarations := Σ.(declarations); retroknowledge := Σ.(retroknowledge) |}, udecl) kn d),
1605
1607
P Σ wf -> P (add_global_decl Σ (kn, d))
1606
- (fst wf, globenv_decl _ _ Σ.(universes) Σ.(retroknowledge) Σ.(declarations) kn d (snd wf) Hfresh onud pd))
1608
+ (fst wf, globenv_decl _ _ Σ.(universes) Σ.(retroknowledge) Σ.(declarations) kn d (snd wf)
1609
+ {| kn_fresh := Hfresh ; on_udecl_udecl := onud ; on_global_decl_d := pd |}))
1607
1610
(Σ : global_env) (wfΣ : wf Σ) : P Σ wfΣ.
1608
1611
Proof .
1609
1612
destruct Σ as [univs Σ]. destruct wfΣ; cbn in *.
1610
- induction o0. apply Pnil.
1613
+ induction o0. apply Pnil. destruct o1.
1611
1614
apply (Pcons {| universes := univs; declarations := Σ |} kn d (o, o0)).
1612
1615
exact IHo0.
1613
1616
Qed .
@@ -2122,19 +2125,19 @@ Proof.
2122
2125
+ constructor. eapply IHdecls => //; eauto. eapply erase_global_cst_decl_wf_glob; auto.
2123
2126
eapply erase_global_decls_fresh; auto.
2124
2127
destruct wfΣ. destruct wfΣpop.
2125
- rewrite (heq _ wf) in o0. now depelim o0.
2128
+ rewrite (heq _ wf) in o0. depelim o0. now destruct o3 .
2126
2129
+ cbn. eapply IHdecls; eauto.
2127
2130
+ constructor. eapply IHdecls; eauto.
2128
2131
destruct wfΣ as [[onu ond]].
2129
- rewrite (heq _ wf) in o. depelim o.
2132
+ rewrite (heq _ wf) in o. depelim o. destruct o0.
2130
2133
eapply (erase_global_ind_decl_wf_glob (kn:=kn)); tea.
2131
2134
intros. rewrite (abstract_env_irr _ H wfpop).
2132
2135
unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wf wfpop) as [? ?].
2133
2136
{intros; now eexists. }
2134
2137
destruct Σpop, Σ; cbn in *. now subst.
2135
2138
eapply erase_global_decls_fresh.
2136
2139
destruct wfΣ as [[onu ond]].
2137
- rewrite (heq _ wf) in o. now depelim o.
2140
+ rewrite (heq _ wf) in o. depelim o. now destruct o0.
2138
2141
+ eapply IHdecls; eauto.
2139
2142
Qed .
2140
2143
@@ -2167,7 +2170,7 @@ Proof.
2167
2170
eapply erase_global_cst_decl_wf_glob.
2168
2171
eapply erase_global_decls_fresh => //.
2169
2172
pose proof (abstract_env_wf _ wf) as [wfΣ].
2170
- depelim wfΣ. rewrite (prf _ wf) in o0. clear - o0. now depelim o0.
2173
+ depelim wfΣ. rewrite (prf _ wf) in o0. clear - o0. depelim o0. now destruct o.
2171
2174
unfold eg', eg'', hidebody.
2172
2175
erewrite erase_global_decls_irr.
2173
2176
eapply IHdecls.
@@ -2188,12 +2191,12 @@ Proof.
2188
2191
{ now eexists. }
2189
2192
destruct Σ, Σ0. cbn in *. rewrite prf' in wfΣ.
2190
2193
depelim wfΣ. cbn in *. rewrite <- H1, H0, <- H2.
2191
- now depelim o0.
2194
+ depelim o0. now destruct o1.
2192
2195
eapply erase_global_decls_fresh => //.
2193
2196
pose proof (abstract_env_wf _ wf) as [wfΣ].
2194
2197
pose proof (prf _ wf) as prf'.
2195
2198
destruct Σ. cbn in *. rewrite prf' in wfΣ.
2196
- clear -wfΣ. destruct wfΣ. cbn in *. now depelim o0.
2199
+ clear -wfΣ. destruct wfΣ. cbn in *. depelim o0. now destruct o1.
2197
2200
unfold eg'', hidebody. erewrite erase_global_decls_irr.
2198
2201
eapply IHdecls. intros x hin. now eapply sub.
2199
2202
eapply IHdecls => //. }
@@ -2526,8 +2529,8 @@ Next Obligation.
2526
2529
pose proof (abstract_env_wf _ H) as [?].
2527
2530
specialize_Σ H. sq. split. cbn. apply X3. cbn.
2528
2531
eapply decls_prefix_wf in X3; tea.
2529
- destruct X3 as [onu ond]. cbn in ond.
2530
- now depelim ond.
2532
+ destruct X3 as [onu ond]. cbn in ond.
2533
+ depelim ond. now destruct o.
2531
2534
Qed .
2532
2535
Next Obligation .
2533
2536
pose proof (abstract_env_ext_wf _ H) as [?].
@@ -2569,7 +2572,7 @@ Proof.
2569
2572
induction suffix.
2570
2573
- cbn. rewrite eqb_refl //.
2571
2574
- cbn. intros ond.
2572
- depelim ond. cbn. red in f.
2575
+ depelim ond. destruct o as [f ? ? ? ]. cbn. red in f.
2573
2576
eapply Forall_app in f as []. depelim H0. cbn in H0.
2574
2577
destruct (eqb_spec kn kn0); [contradiction|].
2575
2578
now apply IHsuffix.
0 commit comments