Skip to content

Commit c9f0f30

Browse files
authored
Merge pull request #758 from MetaCoq/record-for-on_global_decls-data
define and use on_global_decls_data everywhere
2 parents 267f0a5 + 43b4bc2 commit c9f0f30

18 files changed

+110
-100
lines changed

erasure/theories/EDeps.v

Lines changed: 17 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -473,11 +473,11 @@ Proof.
473473
apply PCUICWeakeningEnv.lookup_env_Some_fresh in H as not_fresh.
474474
econstructor.
475475
- unfold PCUICAst.declared_constant in *; cbn.
476-
inversion wfΣ; subst.
476+
inversion wfΣ; subst. destruct X0.
477477
destruct (eqb_spec kn0 kn) as [<-|]; [congruence|].
478478
eassumption.
479479
- unfold EGlobalEnv.declared_constant in *. cbn -[ReflectEq.eqb].
480-
inversion wfΣ; subst.
480+
inversion wfΣ; subst. destruct X0.
481481
destruct (ReflectEq.eqb_spec kn0 kn); [congruence|].
482482
eassumption.
483483
- unfold erases_constant_body in *.
@@ -486,10 +486,10 @@ Proof.
486486
assert (PCUICAst.declared_constant (add_global_decl Σ (kn, decl)) kn0 cb).
487487
{ unfold PCUICAst.declared_constant.
488488
cbn.
489-
inversion wfΣ; subst.
489+
inversion wfΣ; subst. destruct X0.
490490
destruct (eqb_spec kn0 kn) as [<-|]; [congruence|].
491491
easy. }
492-
inversion wfΣ; subst.
492+
inversion wfΣ; subst. destruct X0.
493493
eapply declared_constant_inv in H4; eauto.
494494
2:eapply weaken_env_prop_typing.
495495
red in H4.
@@ -512,35 +512,35 @@ Proof.
512512
invs wfΣ.
513513
destruct H0. split. 2: eauto.
514514
destruct d. split; eauto.
515-
red. cbn. cbn in *.
515+
red. cbn. cbn in *. destruct X0.
516516
destruct (eqb_spec (inductive_mind ind) kn). cbn in *.
517517
subst.
518-
eapply PCUICWeakeningEnv.lookup_env_Some_fresh in H5. eauto. eapply H. exact H0.
518+
eapply PCUICWeakeningEnv.lookup_env_Some_fresh in kn_fresh. eauto. eapply H. exact H0.
519519
- econstructor; eauto.
520520
destruct H as [H H'].
521521
split; eauto. red in H |- *.
522-
inv wfΣ.
522+
inv wfΣ. destruct X0.
523523
unfold PCUICEnvironment.lookup_env.
524524
simpl. destruct (eqb_spec (inductive_mind p.1) kn); auto. subst.
525525
eapply PCUICWeakeningEnv.lookup_env_Some_fresh in H; eauto. contradiction.
526526
destruct H0 as [H0 H0'].
527527
split; eauto. red in H0 |- *.
528-
inv wfΣ. cbn. change (eq_kername (inductive_mind p.1) kn) with (ReflectEq.eqb (inductive_mind p.1) kn).
528+
inv wfΣ. destruct X0. cbn. change (eq_kername (inductive_mind p.1) kn) with (ReflectEq.eqb (inductive_mind p.1) kn).
529529
destruct (ReflectEq.eqb_spec (inductive_mind p.1) kn); auto. subst.
530530
destruct H as [H _].
531531
eapply PCUICWeakeningEnv.lookup_env_Some_fresh in H. eauto. contradiction.
532532
- econstructor; eauto.
533533
destruct H as [[[declm decli] declc] [declp hp]].
534534
repeat split; eauto.
535-
inv wfΣ. unfold PCUICAst.declared_minductive in *.
535+
inv wfΣ. destruct X0. unfold PCUICAst.declared_minductive in *.
536536
unfold PCUICEnvironment.lookup_env.
537537
simpl in *.
538538
destruct (ReflectEq.eqb_spec (inductive_mind p.(proj_ind)) kn). subst.
539539
eapply PCUICWeakeningEnv.lookup_env_Some_fresh in declm; eauto. contradiction.
540540
apply declm.
541541
destruct H0 as [[[]]]. destruct a.
542542
repeat split; eauto.
543-
inv wfΣ. simpl. unfold declared_minductive. cbn.
543+
inv wfΣ. destruct X0. simpl. unfold declared_minductive. cbn.
544544
destruct (ReflectEq.eqb_spec (inductive_mind p.(proj_ind)) kn); auto. subst.
545545
destruct H as [[[]]].
546546
eapply PCUICWeakeningEnv.lookup_env_Some_fresh in H. eauto. contradiction.
@@ -703,10 +703,10 @@ Proof.
703703
* unfold erases_constant_body, on_constant_decl in *.
704704
destruct ?; [|easy].
705705
destruct ?; [|easy].
706-
depelim wf. depelim o0. cbn in *.
706+
depelim wf. depelim o0. destruct o1. cbn in *.
707707
eapply (erases_extends ({| universes := univs; declarations := Σ |}, cst_universes cst')); eauto.
708708
cbn. 4:{ split; eauto; cbn; try reflexivity. eexists [_]; cbn; reflexivity. }
709-
constructor; auto. cbn. red in o2. rewrite E in o2. exact o2.
709+
constructor; auto. cbn. red in on_global_decl_d. rewrite E in on_global_decl_d. exact on_global_decl_d.
710710
split; auto.
711711
* intros.
712712
eapply (erases_deps_cons {| universes := univs; declarations := Σ |} _ kn (PCUICEnvironment.ConstantDecl cst')); auto.
@@ -716,9 +716,9 @@ Proof.
716716
unfold on_constant_decl in *.
717717
cbn in *.
718718
eapply (erases_deps_single (_, _)). 3:eauto.
719-
depelim wf. depelim o0.
719+
depelim wf. depelim o0. destruct o1.
720720
now split; cbn; eauto.
721-
depelim wf. depelim o0. do 2 red in o2. now rewrite E in o2.
721+
depelim wf. depelim o0. destruct o1. do 2 red in on_global_decl_d. now rewrite E in on_global_decl_d.
722722
apply IH; eauto. depelim wf. now depelim o0.
723723
+ set (Σu := {| universes := univs; declarations := Σ; retroknowledge := retro |}).
724724
assert (wfΣu : PCUICTyping.wf Σu).
@@ -767,17 +767,17 @@ Proof.
767767
unfold PCUICAst.declared_minductive in decli.
768768
unfold PCUICEnvironment.lookup_env in decli.
769769
simpl in decli. rewrite eq_kername_refl in decli. intuition discriminate.
770-
* inv wf. inv X.
770+
* inv wf. inv X. destruct X1.
771771
specialize (IH _ (H0, X0) erg).
772772
destruct decli as [decli ?].
773773
simpl in decli |- *.
774774
unfold PCUICAst.declared_minductive, PCUICEnvironment.lookup_env in decli.
775775
simpl in decli.
776776
destruct (eqb_specT (inductive_mind k) kn). simpl in *. subst. noconf decli.
777-
destruct (Forall2_nth_error_left (proj1 H) _ _ H3); eauto.
777+
destruct (Forall2_nth_error_left (proj1 H) _ _ H1); eauto.
778778
eexists _, _; intuition eauto. split; eauto. red.
779779
simpl. rewrite eqb_refl. congruence.
780-
destruct (proj2 IH _ _ _ (conj decli H3)) as [m' [i' [decli' ei]]].
780+
destruct (proj2 IH _ _ _ (conj decli H1)) as [m' [i' [decli' ei]]].
781781
eexists _, _; intuition eauto.
782782
destruct decli'; red; split; eauto.
783783
red in d |- *. simpl.

erasure/theories/ErasureCorrectness.v

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -1136,10 +1136,10 @@ Proof.
11361136
induction er; intros wf.
11371137
- constructor.
11381138
- cbn. destruct cb' as [[]].
1139-
cbn in *. depelim wf.
1139+
cbn in *. depelim wf. destruct o.
11401140
rewrite [forallb _ _](IHer wf) andb_true_r.
11411141
red in H. destruct cb as [ty []]; cbn in *.
1142-
unshelve eapply PCUICClosedTyp.subject_closed in o0. cbn. split; auto.
1142+
unshelve eapply PCUICClosedTyp.subject_closed in on_global_decl_d. cbn. split; auto.
11431143
eapply erases_closed in H; tea. elim H.
11441144
cbn. apply IHer. now depelim wf.
11451145
- depelim wf.
@@ -1194,11 +1194,11 @@ Proof.
11941194
move: wf. red in er; cbn in er.
11951195
induction er; intros wf.
11961196
- constructor.
1197-
- cbn. depelim wf.
1197+
- cbn. depelim wf. destruct o.
11981198
constructor; eauto.
11991199
2:eapply erases_global_decls_fresh; tea.
12001200
cbn. red in H.
1201-
do 2 red in o0.
1201+
do 2 red in on_global_decl_d.
12021202
destruct (cst_body cb).
12031203
destruct (E.cst_body cb') => //. cbn.
12041204
set (Σ'' := ({| universes := _ |}, _)) in *.
@@ -1209,7 +1209,7 @@ Proof.
12091209
specialize (H0 H Σ'). eapply H0.
12101210
eapply erases_global_all_deps; tea. split => //.
12111211
destruct (E.cst_body cb') => //.
1212-
- depelim wf.
1212+
- depelim wf. destruct o.
12131213
constructor; eauto.
12141214
now eapply erases_mutual_inductive_body_wf.
12151215
now eapply erases_global_decls_fresh; tea.

erasure/theories/ErasureFunction.v

Lines changed: 25 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -772,7 +772,7 @@ Qed.
772772
case: eqb_spec.
773773
intros ->.
774774
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 ? ? ? ].
776776
cbn in *. eapply Forall_app in f as []. contradiction.
777777
intros _.
778778
eapply IHx; trea.
@@ -908,7 +908,7 @@ Next Obligation.
908908
{ now eexists. }
909909
destruct (abstract_pop_decls_correct X decls prop' _ _ HX H) as [? []].
910910
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 => //.
912912
Qed.
913913
Next Obligation.
914914
pose proof (abstract_env_ext_wf _ H) as [wf].
@@ -922,7 +922,7 @@ Next Obligation.
922922
pose proof (abstract_make_wf_env_ext_correct _ _ _ _ _ HX' H).
923923
clear H HX'. specialize (prop _ HX). destruct x, Σ as [[] u], H0; cbn in *.
924924
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.
926926
Qed.
927927
Next Obligation.
928928
pose proof (abstract_env_exists X) as [[? HX]].
@@ -1218,7 +1218,7 @@ Proof.
12181218
induction er using erases_deps_forall_ind; try solve [now constructor].
12191219
- assert (kn <> kn0).
12201220
{ 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. }
12221222
eapply erases_deps_tConst with cb cb'; eauto.
12231223
red. rewrite /lookup_env lookup_env_cons_fresh //.
12241224
red.
@@ -1240,15 +1240,15 @@ Proof.
12401240
red in H. red.
12411241
inv wfΣ. inv X.
12421242
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.
12441244
- econstructor; eauto.
12451245
red. destruct H. split; eauto.
12461246
destruct d0; split; eauto.
12471247
destruct d0; split; eauto.
12481248
inv wfΣ. inv X.
12491249
red in H |- *.
12501250
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.
12521252
Qed.
12531253

12541254
Lemma lookup_env_ext {Σ kn kn' d d'} :
@@ -1259,7 +1259,7 @@ Proof.
12591259
intros wf hl.
12601260
eapply lookup_env_Some_fresh in hl.
12611261
inv wf. inv X.
1262-
destruct (eqb_spec kn kn'); subst; congruence.
1262+
destruct (eqb_spec kn kn'); subst; destruct X1; congruence.
12631263
Qed.
12641264

12651265
Lemma lookup_env_cons_disc {Σ kn kn' d} :
@@ -1288,11 +1288,11 @@ Proof.
12881288
exists cst. split.
12891289
red in declc |- *. unfold lookup_env in *.
12901290
rewrite lookup_env_cons_fresh //.
1291-
{ eapply lookup_env_Some_fresh in declc.
1291+
{ eapply lookup_env_Some_fresh in declc. destruct X1.
12921292
intros <-; contradiction. }
12931293
exists cst'.
12941294
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.
12961296
intros <-; contradiction. }
12971297
red in ebody. unfold erases_constant_body.
12981298
destruct (cst_body cst) eqn:bod; destruct (E.cst_body cst') eqn:bod' => //.
@@ -1327,7 +1327,7 @@ Proof.
13271327
unfold lookup_env in *.
13281328
rewrite lookup_env_cons_fresh //.
13291329
{ eapply lookup_env_Some_fresh in declc.
1330-
intros <-. contradiction. }
1330+
intros <-. destruct X1. contradiction. }
13311331
exists cst'.
13321332
unfold EGlobalEnv.declared_constant.
13331333
red in ebody. unfold erases_constant_body.
@@ -1462,7 +1462,9 @@ Proof.
14621462
cbn in o0, o. rewrite prf' in o0. rewrite <- Hpop in o0. rewrite Hpop' in o. clear -o o0.
14631463
now depelim o0.
14641464
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+
}
14661468
all: eauto.
14671469
apply IHdecls; eauto.
14681470
{ intros. pose proof (abstract_env_wf _ wfpop) as [wf'].
@@ -1603,11 +1605,12 @@ Lemma on_global_env_ind (P : forall Σ : global_env, wf Σ -> Type)
16031605
(pd : on_global_decl cumulSpec0 (lift_typing typing)
16041606
({| universes := Σ.(universes); declarations := Σ.(declarations); retroknowledge := Σ.(retroknowledge) |}, udecl) kn d),
16051607
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 |}))
16071610
(Σ : global_env) (wfΣ : wf Σ) : P Σ wfΣ.
16081611
Proof.
16091612
destruct Σ as [univs Σ]. destruct wfΣ; cbn in *.
1610-
induction o0. apply Pnil.
1613+
induction o0. apply Pnil. destruct o1.
16111614
apply (Pcons {| universes := univs; declarations := Σ |} kn d (o, o0)).
16121615
exact IHo0.
16131616
Qed.
@@ -2122,19 +2125,19 @@ Proof.
21222125
+ constructor. eapply IHdecls => //; eauto. eapply erase_global_cst_decl_wf_glob; auto.
21232126
eapply erase_global_decls_fresh; auto.
21242127
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.
21262129
+ cbn. eapply IHdecls; eauto.
21272130
+ constructor. eapply IHdecls; eauto.
21282131
destruct wfΣ as [[onu ond]].
2129-
rewrite (heq _ wf) in o. depelim o.
2132+
rewrite (heq _ wf) in o. depelim o. destruct o0.
21302133
eapply (erase_global_ind_decl_wf_glob (kn:=kn)); tea.
21312134
intros. rewrite (abstract_env_irr _ H wfpop).
21322135
unshelve epose proof (abstract_pop_decls_correct X decls _ _ _ wf wfpop) as [? ?].
21332136
{intros; now eexists. }
21342137
destruct Σpop, Σ; cbn in *. now subst.
21352138
eapply erase_global_decls_fresh.
21362139
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.
21382141
+ eapply IHdecls; eauto.
21392142
Qed.
21402143

@@ -2167,7 +2170,7 @@ Proof.
21672170
eapply erase_global_cst_decl_wf_glob.
21682171
eapply erase_global_decls_fresh => //.
21692172
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.
21712174
unfold eg', eg'', hidebody.
21722175
erewrite erase_global_decls_irr.
21732176
eapply IHdecls.
@@ -2188,12 +2191,12 @@ Proof.
21882191
{ now eexists. }
21892192
destruct Σ, Σ0. cbn in *. rewrite prf' in wfΣ.
21902193
depelim wfΣ. cbn in *. rewrite <- H1, H0, <- H2.
2191-
now depelim o0.
2194+
depelim o0. now destruct o1.
21922195
eapply erase_global_decls_fresh => //.
21932196
pose proof (abstract_env_wf _ wf) as [wfΣ].
21942197
pose proof (prf _ wf) as prf'.
21952198
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.
21972200
unfold eg'', hidebody. erewrite erase_global_decls_irr.
21982201
eapply IHdecls. intros x hin. now eapply sub.
21992202
eapply IHdecls => //. }
@@ -2526,8 +2529,8 @@ Next Obligation.
25262529
pose proof (abstract_env_wf _ H) as [?].
25272530
specialize_Σ H. sq. split. cbn. apply X3. cbn.
25282531
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.
25312534
Qed.
25322535
Next Obligation.
25332536
pose proof (abstract_env_ext_wf _ H) as [?].
@@ -2569,7 +2572,7 @@ Proof.
25692572
induction suffix.
25702573
- cbn. rewrite eqb_refl //.
25712574
- cbn. intros ond.
2572-
depelim ond. cbn. red in f.
2575+
depelim ond. destruct o as [f ? ? ? ]. cbn. red in f.
25732576
eapply Forall_app in f as []. depelim H0. cbn in H0.
25742577
destruct (eqb_spec kn kn0); [contradiction|].
25752578
now apply IHsuffix.

pcuic/theories/PCUICExpandLetsCorrectness.v

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -4680,15 +4680,15 @@ Lemma trans_wf {cf} {Σ : global_env_ext} : wf Σ -> wf_trans Σ.
46804680
Proof.
46814681
rewrite /PCUICEnvironment.fst_ctx.
46824682
destruct Σ as [[gunivs Σ retro] udecl]; cbn. intros [onu wfΣ]; cbn in *.
4683-
induction wfΣ as [|Σ0 kn d X IHX f udecl' onu' ond]. constructor; auto. constructor.
4683+
induction wfΣ as [|Σ0 kn d X IHX [f udecl' onu' ond]]. constructor; auto. constructor.
46844684
have onud : on_udecl gunivs (PCUICLookup.universes_decl_of_decl (trans_global_decl d)).
46854685
{ apply (trans_on_udecl (Σ:= {| universes := gunivs; declarations := Σ0; retroknowledge := retro |})) in onu'. destruct d => //. }
46864686
cbn; constructor; eauto.
46874687
rename Σ0 into Σd.
46884688
set (Σ0 := {| universes := gunivs; declarations := Σd; retroknowledge := retro |}).
46894689
rename X into Xd.
46904690
set (X := (onu, Xd) : wf Σ0).
4691-
constructor; auto; try apply IHX.
4691+
constructor; try constructor; auto; try apply IHX.
46924692
{ now apply (fresh_global_map (Σ := Σ0)). }
46934693
destruct d; cbn in *.
46944694
* cbn. red. move: ond; rewrite /on_constant_decl.

pcuic/theories/PCUICFirstorder.v

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -393,11 +393,12 @@ Proof.
393393
- cbn => //.
394394
- cbn. destruct a => //. intros gs ong.
395395
depelim ong. specialize (IHΣ'' _ ong).
396+
destruct o as [f ? ? ?].
396397
destruct g => //.
397398
* intros hl. specialize (IHΣ'' hl).
398399
eapply plookup_env_Some_not_fresh in hl.
399400
cbn. case: eqb_spec.
400-
+ intros <-. apply fresh_global_app in f as [].
401+
+ intros <-. apply fresh_global_app in f as [].
401402
contradiction.
402403
+ now intros neq.
403404
* intros hl. specialize (IHΣ'' hl).

pcuic/theories/PCUICTyping.v

Lines changed: 5 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -776,10 +776,12 @@ Proof.
776776
unfold Σg in o |- *; cbn in o.
777777
rename o into ongu. rename o0 into o. cbn in o |- *.
778778
destruct o. { constructor. }
779-
rename o1 into Xg.
779+
rename o0 into Xg.
780780
set (wfΣ := (ongu, o) : on_global_env cumulSpec0 (lift_typing typing) {| universes := univs; declarations := Σ |}).
781781
set (Σ':= {| universes := univs; declarations := Σ |}) in *.
782-
constructor; auto.
782+
destruct Xg.
783+
rename on_global_decl_d into Xg.
784+
constructor; auto; try constructor; auto.
783785
* unshelve eset (IH' := IH ((Σ', udecl); (wfΣ; []; (tSort Universe.lProp); _; _))).
784786
shelve. simpl. apply type_Prop.
785787
forward IH'. constructor 1; cbn. lia.
@@ -1252,7 +1254,7 @@ Section All_local_env.
12521254
exists [(kn, decl)] => //.
12531255
apply Retroknowledge.extends_refl.
12541256
* split => //.
1255-
* apply o0.
1257+
* destruct o; assumption.
12561258
- intros hl. destruct (IHΣp hl) as [Σ' []].
12571259
exists Σ'.
12581260
split=> //.

0 commit comments

Comments
 (0)