Skip to content

Commit 553263d

Browse files
committed
Strenghten In_size lemma
1 parent 4dc7c47 commit 553263d

File tree

7 files changed

+48
-21
lines changed

7 files changed

+48
-21
lines changed

erasure/theories/EConstructorsAsBlocks.v

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -67,7 +67,7 @@ Section transform_blocks.
6767
all:try lia.
6868
all:try apply (In_size); tea.
6969
all:try lia.
70-
- now apply (In_size id size).
70+
- setoid_rewrite <- (In_size id size H); unfold id; lia.
7171
- change (fun x => size (id x)) with size in H.
7272
eapply (In_size id size) in H. unfold id in H.
7373
change (fun x => size x) with size in H.
@@ -82,6 +82,8 @@ Section transform_blocks.
8282
change (fun x => size x) with size in H.
8383
pose proof (size_mkApps_l napp nnil). lia.
8484
- eapply (In_size snd size) in H. cbn in *. lia.
85+
- eapply (In_size dbody size) in H; lia.
86+
- eapply (In_size dbody size) in H; lia.
8587
- now eapply InPrim_size in H.
8688
Qed.
8789

erasure/theories/EEtaExpanded.v

Lines changed: 9 additions & 14 deletions
Original file line numberDiff line numberDiff line change
@@ -79,23 +79,18 @@ Section isEtaExp.
7979
| tConstruct ind i block_args => isEtaExp_app ind i 0 && is_nil block_args }.
8080
Proof.
8181
all:try lia.
82-
all:try apply (In_size); tea.
83-
all:try lia.
84-
- now apply (In_size id size).
85-
- rewrite size_mkApps.
86-
eapply (In_size id size) in H.
87-
change (fun x => size (id x)) with size in H. unfold id in *; cbn.
88-
lia.
89-
- now eapply size_mkApps_f.
90-
- change (fun x => size (id x)) with size in H.
91-
eapply (In_size id size) in H. unfold id in H.
92-
change (fun x => size x) with size in H.
93-
pose proof (size_mkApps_l napp nnil). lia.
82+
3:now eapply size_mkApps_f.
83+
3:{ pose proof (size_mkApps_l napp nnil).
84+
setoid_rewrite <- (In_size id size H) in H0. unfold id in H0. lia. }
85+
all:try rewrite size_mkApps.
86+
all:try setoid_rewrite <- (In_size id size H); unfold id.
87+
all: try (cbn; lia).
9488
- eapply (In_size snd size) in H. cbn in H; lia.
89+
- eapply (In_size dbody size) in H. cbn in H; lia.
90+
- eapply (In_size dbody size) in H. cbn in H; lia.
9591
- destruct p as [? []]; cbn in *; intuition eauto.
9692
subst. lia.
97-
eapply (In_size id size) in H0. unfold id in H0.
98-
change (fun x => size x) with size in H0. lia.
93+
setoid_rewrite <- (In_size id size H0); unfold id. lia.
9994
Qed.
10095

10196
End isEtaExp.

erasure/theories/EEtaExpandedFix.v

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -301,7 +301,8 @@ Section isEtaExp.
301301
all:try lia.
302302
all:try apply (In_size); tea.
303303
all:try lia.
304-
- now apply (In_size id size).
304+
- setoid_rewrite (In_size id size); tea. unfold id.
305+
now change (fun x => size x) with size.
305306
- rewrite size_mkApps. cbn.
306307
apply (In_size id size) in H.
307308
unfold id in H. change (fun x => size x) with size in H. lia.
@@ -320,6 +321,7 @@ Section isEtaExp.
320321
change (fun x => size x) with size in H.
321322
pose proof (size_mkApps_l napp nnil). lia.
322323
- eapply (In_size snd size) in H. cbn in H; lia.
324+
- eapply (In_size dbody size) in H. cbn in H; lia.
323325
- destruct p as [? []]; cbn in *; eauto. destruct H; subst; try lia.
324326
eapply (In_size id size) in H. unfold id in H.
325327
change (fun x => size x) with size in H. lia.

erasure/theories/EImplementBox.v

Lines changed: 5 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -59,9 +59,11 @@ Section implement_box.
5959
all:try lia.
6060
all:try apply (In_size); tea.
6161
all:try lia.
62-
- now apply (In_size id size).
63-
- now eapply (In_size id size).
64-
- eapply (In_size snd size) in H. cbn in *. lia.
62+
- setoid_rewrite <- (In_size id size H); unfold id; lia.
63+
- setoid_rewrite <- (In_size id size H); unfold id; lia.
64+
- setoid_rewrite <- (In_size snd size H); cbn; lia.
65+
- setoid_rewrite <- (In_size dbody size H); cbn; lia.
66+
- setoid_rewrite <- (In_size dbody size H); cbn; lia.
6567
- now eapply InPrim_size in H.
6668
Qed.
6769

erasure/theories/ERemoveParams.v

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -63,7 +63,8 @@ Section strip.
6363
Proof.
6464
all:try lia.
6565
all:try apply (In_size); tea.
66-
- now eapply (In_size id size).
66+
- eapply (In_size id size) in H. unfold id in H.
67+
change (fun x => size x) with size in *. lia.
6768
- rewrite size_mkApps.
6869
eapply (In_size id size) in H. change (fun x => size (id x)) with size in H. unfold id in H. cbn. lia.
6970
- rewrite size_mkApps.
@@ -72,6 +73,8 @@ Section strip.
7273
- pose proof (size_mkApps_l napp nnil).
7374
eapply (In_size id size) in H. change (fun x => size (id x)) with size in H. unfold id in H. lia.
7475
- eapply (In_size snd size) in H. cbn in H; lia.
76+
- eapply (In_size dbody size) in H; cbn in H; lia.
77+
- eapply (In_size dbody size) in H; cbn in H; lia.
7578
- destruct p as [? []]; cbn in H; eauto.
7679
intuition auto; subst; cbn; try lia.
7780
eapply (In_size id size) in H0. unfold id in H0.

utils/theories/MCArith.v

Lines changed: 23 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -58,3 +58,26 @@ Proof.
5858
rewrite Pos2Nat.inj_1. reflexivity.
5959
rewrite Pos2Nat.inj_succ. cbn. f_equal. lia.
6060
Qed.
61+
62+
(* Missing rewriting theory on natural number orders *)
63+
64+
Require Import Morphisms Morphisms_Prop.
65+
#[export] Instance proper_S_lt : Morphisms.Proper (lt ==> lt)%signature S.
66+
Proof. red. intros x y h. lia. Qed.
67+
#[export] Instance proper_add_lt_r : Morphisms.Proper (eq ==> lt ==> lt)%signature Nat.add.
68+
Proof. red. intros ??????. lia. Qed.
69+
#[export] Instance proper_add_lt_l : Morphisms.Proper (lt ==> eq ==> lt)%signature Nat.add.
70+
Proof. red. intros ??????. lia. Qed.
71+
72+
#[export] Instance proper_S_le : Morphisms.Proper (le ==> le)%signature S.
73+
Proof. red. intros x y h. lia. Qed.
74+
#[export] Instance proper_add_le_r : Morphisms.Proper (eq ==> le ==> le)%signature Nat.add.
75+
Proof. red. intros ??????. lia. Qed.
76+
#[export] Instance proper_add_le_l : Morphisms.Proper (le ==> eq ==> le)%signature Nat.add.
77+
Proof. red. intros ??????. lia. Qed.
78+
79+
#[export] Instance subrel_eq_le : subrelation eq le.
80+
Proof. red. now intros ?? ->. Qed.
81+
82+
#[export] Instance subrel_lt_le : subrelation lt le.
83+
Proof. red. intros ???. lia. Qed.

utils/theories/MCList.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1406,7 +1406,7 @@ Proof.
14061406
Qed.
14071407

14081408
Lemma In_size {A B} {x : A} {l : list A} (proj : A -> B) (size : B -> nat) :
1409-
In x l -> size (proj x) < S (list_size (size ∘ proj) l).
1409+
In x l -> size (proj x) < (list_size (size ∘ proj) l).
14101410
Proof.
14111411
induction l; cbn => //.
14121412
intros [->|hin]. lia. specialize (IHl hin); lia.

0 commit comments

Comments
 (0)