Skip to content

Commit 583efac

Browse files
committed
cleanup
1 parent 0373e50 commit 583efac

File tree

9 files changed

+130
-170
lines changed

9 files changed

+130
-170
lines changed

src/examples/gmap_option.v

Lines changed: 23 additions & 47 deletions
Original file line numberDiff line numberDiff line change
@@ -29,13 +29,13 @@ Definition gmap_mov (a b: gmap K (option V)) : Prop :=
2929
Lemma gmap_dot_comm x y
3030
: gmap_dot x y = gmap_dot y x.
3131
Proof.
32-
intros. unfold gmap_dot, gmerge. apply map_eq. intro. rewrite lookup_merge.
32+
intros. unfold gmap_dot, gmerge. apply map_eq. intro i. rewrite lookup_merge.
3333
rewrite lookup_merge. unfold diag_None. destruct (x !! i), (y !! i); trivial.
3434
Qed.
3535

3636
Lemma gmap_dot_assoc x y z
3737
: gmap_dot x (gmap_dot y z) = gmap_dot (gmap_dot x y) z.
38-
Proof. intros. unfold gmap_dot, gmerge. apply map_eq. intro. rewrite lookup_merge.
38+
Proof. intros. unfold gmap_dot, gmerge. apply map_eq. intro i. rewrite lookup_merge.
3939
rewrite lookup_merge.
4040
rewrite lookup_merge.
4141
unfold diag_None.
@@ -47,30 +47,30 @@ Qed.
4747
Lemma gmap_dot_empty
4848
: ∀ x : gmap K (option V), gmap_dot x ∅ = x.
4949
Proof.
50-
intros. unfold gmap_dot. apply map_eq. intro. rewrite lookup_merge. rewrite lookup_empty.
50+
intros x. unfold gmap_dot. apply map_eq. intro i. rewrite lookup_merge. rewrite lookup_empty.
5151
unfold diag_None, gmerge. destruct (x !! i); trivial.
5252
Qed.
5353

5454
Lemma gmap_dot_empty_left
5555
: ∀ x : gmap K (option V), gmap_dot ∅ x = x.
5656
Proof.
57-
intros. unfold gmap_dot. apply map_eq. intro. rewrite lookup_merge. rewrite lookup_empty.
57+
intros x. unfold gmap_dot. apply map_eq. intro i. rewrite lookup_merge. rewrite lookup_empty.
5858
unfold diag_None, gmerge. destruct (x !! i); trivial.
5959
Qed.
6060

6161
Lemma lookup_gmap_dot_left a b k
6262
: gmap_valid (gmap_dot a b) -> (a !! k ≠ None) -> (gmap_dot a b) !! k = a !! k.
6363
Proof.
64-
unfold gmap_valid, gmap_dot. intros.
65-
have j := H k. rewrite lookup_merge.
64+
unfold gmap_valid, gmap_dot. intros Q R.
65+
have j := Q k. rewrite lookup_merge.
6666
rewrite lookup_merge in j. unfold diag_None, gmerge in *. destruct (a !! k), (b !! k);
6767
trivial; contradiction.
6868
Qed.
6969

7070
Lemma lookup_gmap_dot_right a b k
7171
: gmap_valid (gmap_dot a b) -> (b !! k ≠ None) -> (gmap_dot a b) !! k = b !! k.
7272
Proof.
73-
unfold gmap_valid, gmap_dot. intros. have j := H k. rewrite lookup_merge.
73+
unfold gmap_valid, gmap_dot. intros Q R. have j := Q k. rewrite lookup_merge.
7474
rewrite lookup_merge in j. unfold diag_None, gmerge in *. destruct (a !! k), (b !! k);
7575
trivial; contradiction.
7676
Qed.
@@ -79,9 +79,9 @@ Lemma lookup_gmap_dot_3mid a b c k
7979
: gmap_valid (gmap_dot (gmap_dot a b) c) -> (b !! k ≠ None) ->
8080
gmap_dot (gmap_dot a b) c !! k = b !! k.
8181
Proof.
82-
intros.
83-
rewrite gmap_dot_comm in H.
84-
rewrite gmap_dot_assoc in H.
82+
intros Q R.
83+
rewrite gmap_dot_comm in Q.
84+
rewrite gmap_dot_assoc in Q.
8585
rewrite gmap_dot_comm.
8686
rewrite gmap_dot_assoc.
8787
apply lookup_gmap_dot_right; trivial.
@@ -91,8 +91,8 @@ Lemma lookup_gmap_dot_3left a b c k
9191
: gmap_valid (gmap_dot (gmap_dot a b) c) -> (a !! k ≠ None) ->
9292
gmap_dot (gmap_dot a b) c !! k = a !! k.
9393
Proof.
94-
intros.
95-
rewrite <- gmap_dot_assoc in H.
94+
intros Q R.
95+
rewrite <- gmap_dot_assoc in Q.
9696
rewrite <- gmap_dot_assoc.
9797
apply lookup_gmap_dot_left; trivial.
9898
Qed.
@@ -107,8 +107,8 @@ Qed.
107107
Lemma gmap_valid_left
108108
: ∀ x y : gmap K (option V), gmap_valid (gmap_dot x y) → gmap_valid x.
109109
Proof.
110-
intros. unfold gmap_valid, gmap_dot in *.
111-
intro. have h := H k. clear H. rewrite lookup_merge in h. unfold diag_None in h.
110+
intros x y. unfold gmap_valid, gmap_dot in *.
111+
intros Q k. have h := Q k. clear Q. rewrite lookup_merge in h. unfold diag_None in h.
112112
unfold gmerge in h.
113113
destruct (x !! k); trivial.
114114
destruct (y !! k); trivial. contradiction.
@@ -124,8 +124,8 @@ Lemma gmap_valid_update_singleton j x y (m: gmap K (option V)) :
124124
gmap_valid (gmap_dot {[j := Some x]} m) ->
125125
gmap_valid (gmap_dot {[j := Some y]} m).
126126
Proof.
127-
intros. unfold gmap_valid, gmap_dot in *. intro.
128-
have r := H k. rewrite lookup_merge. rewrite lookup_merge in r.
127+
intros Q. unfold gmap_valid, gmap_dot in *. intro k.
128+
have r := Q k. rewrite lookup_merge. rewrite lookup_merge in r.
129129
unfold gmerge, diag_None in *.
130130
have h : Decision (j = k) by solve_decision. destruct h.
131131
- subst k. rewrite lookup_singleton. rewrite lookup_singleton in r.
@@ -134,30 +134,6 @@ Proof.
134134
rewrite lookup_singleton_ne in r; trivial.
135135
Qed.
136136

137-
(*
138-
#[refine]
139-
Global Instance gmap_tpcm : TPCM (gmap K (option V)) := {
140-
m_valid p := gmap_valid p ;
141-
dot a b := gmap_dot a b ;
142-
mov a b := gmap_mov a b ;
143-
unit := empty ;
144-
}.
145-
- apply gmap_valid_left.
146-
- unfold gmap_valid. intros. rewrite lookup_empty. trivial.
147-
- apply gmap_dot_empty.
148-
- intros. apply gmap_dot_comm.
149-
- intros. apply gmap_dot_assoc.
150-
- intros. unfold gmap_mov. intro. trivial.
151-
- intros. unfold gmap_mov in *. intros. apply H0. apply H. trivial.
152-
- intros. split.
153-
* unfold gmap_mov in H. apply H. apply H0.
154-
* unfold gmap_mov in H. unfold gmap_mov. intro.
155-
rewrite <- gmap_dot_assoc.
156-
rewrite <- gmap_dot_assoc.
157-
apply H.
158-
Defined.
159-
*)
160-
161137
Definition gmap_le (a b : gmap K (option V)) := ∃ c , gmap_dot a c = b.
162138

163139
Lemma le_of_subset (a b : gmap K (option V))
@@ -166,19 +142,19 @@ Proof.
166142
assert (∀ x : K * option V, Decision (match x with (k,v) => a !! k = None end)) as X
167143
by solve_decision.
168144
exists (map_filter (λ x , match x with (k,v) => a !! k = None end) X b).
169-
unfold gmap_dot. apply map_eq. intro.
145+
unfold gmap_dot. apply map_eq. intro i.
170146
have ff := f i.
171147
rewrite lookup_merge. unfold diag_None, gmerge.
172-
destruct (a !! i) eqn:ai.
148+
destruct (a !! i) as [o|] eqn:ai.
173149
- rewrite map_lookup_filter.
174-
unfold "≫=", option_bind. destruct (b!!i) eqn:bi.
175-
+ unfold guard. have fff := ff o. intuition. inversion H.
150+
unfold "≫=", option_bind. destruct (b!!i) as [o0|] eqn:bi.
151+
+ unfold guard. have fff := ff o. have ffff := fff eq_refl. inversion ffff. subst o0.
176152
destruct (X (i, o)) as [e|e].
177153
* rewrite e in ai. discriminate.
178154
* trivial.
179155
+ have fff := ff o. intuition.
180156
- rewrite map_lookup_filter.
181-
unfold "≫=", option_bind. destruct (b!!i) eqn:bi; trivial.
157+
unfold "≫=", option_bind. destruct (b!!i) as [o|] eqn:bi; trivial.
182158
destruct (X (i, o)); trivial.
183159
contradiction.
184160
Qed.
@@ -194,7 +170,7 @@ Lemma conjunct_and_gmap
194170
Proof.
195171
apply le_of_subset. intros k v e1.
196172
unfold gmap_dot in e1. rewrite lookup_merge in e1. unfold diag_None in e1.
197-
destruct (a1 !! k) eqn:a1k; destruct (a2 !! k) eqn:a2k.
173+
destruct (a1 !! k) as [o|] eqn:a1k; destruct (a2 !! k) as [o0|] eqn:a2k.
198174
- have l := a_disj _ _ _ a1k a2k. contradiction.
199175
- unfold gmerge in e1. inversion e1. subst o. unfold gmap_le in la1.
200176
destruct la1 as [d la]. subst c.
@@ -204,7 +180,7 @@ Proof.
204180
have gvk := gv k.
205181
rewrite lookup_merge in gvk. rewrite a1k in gvk. unfold diag_None in gvk.
206182
unfold gmerge. unfold gmerge in gvk. destruct (d !! k); trivial. contradiction.
207-
- unfold gmerge in e1. inversion e1. subst o. unfold gmap_le in la2.
183+
- unfold gmerge in e1. inversion e1. subst o0. unfold gmap_le in la2.
208184
destruct la2 as [d la]. subst c.
209185
unfold gmap_dot.
210186
unfold gmap_valid in gv. unfold gmap_dot in gv.

src/examples/hash_table_raw.v

Lines changed: 16 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -88,7 +88,7 @@ Definition m (k: Key) (v: option Value) :=
8888
Lemma ht_valid_QueryFound j k v v0
8989
: V (ht_dot (s j (Some (k, v0))) (m k v)) -> v = Some v0.
9090
Proof.
91-
unfold V, s, m. intros [z H]. destruct z. unfold ht_dot in *. unfold P in *.
91+
unfold V, s, m. intros [z H]. destruct z as [ms slots]. unfold ht_dot in *. unfold P in *.
9292
repeat (rewrite gmap_dot_empty in H).
9393
repeat (rewrite gmap_dot_empty_left in H).
9494
destruct H as [H [H0 [H1 [H2 [H3 H4]]]]].
@@ -122,7 +122,7 @@ Lemma ht_valid_QueryReachedEnd k a v
122122
(is_full: full a k (hash k) ht_fixed_size)
123123
: V (ht_dot a (m k v)) -> v = None.
124124
Proof.
125-
unfold V, s, m. intros [z H]. destruct z, a. unfold ht_dot in *. unfold P in *.
125+
unfold V, s, m. intros [z H]. destruct z as [ms slots], a as [ms0 slots0]. unfold ht_dot in *. unfold P in *.
126126
unfold full in is_full.
127127
destruct H as [H [H0 [H1 [H2 [H3 H4]]]]].
128128
destruct is_full as [H5 [H6 [H7 H8]]].
@@ -139,7 +139,7 @@ Proof.
139139
repeat (rewrite gmap_dot_empty_left in H3).
140140
repeat (rewrite gmap_dot_empty in H4).
141141
repeat (rewrite gmap_dot_empty_left in H4).
142-
destruct v; trivial. exfalso.
142+
destruct v as [v|]; trivial. exfalso.
143143
have h := H3 k v.
144144
assert (gmap_dot {[k := Some (Some v)]} ms !! k = Some (Some (Some v))) as Q.
145145
{
@@ -167,7 +167,7 @@ Lemma ht_valid_QueryNotFound k a v j
167167
(is_full: full a k (hash k) j)
168168
: V (ht_dot (ht_dot a (m k v)) (s j None)) -> v = None.
169169
Proof.
170-
unfold V, s, m. intros [z H]. destruct z, a. unfold ht_dot in *. unfold P in *.
170+
unfold V, s, m. intros [z H]. destruct z as [ms slots], a as [ms0 slots0]. unfold ht_dot in *. unfold P in *.
171171
unfold full in is_full.
172172
destruct H as [H [H0 [H1 [H2 [H3 H4]]]]].
173173
destruct is_full as [H5 [H6 [H7 H8]]].
@@ -184,7 +184,7 @@ Proof.
184184
repeat (rewrite gmap_dot_empty_left in H3).
185185
repeat (rewrite gmap_dot_empty in H4).
186186
repeat (rewrite gmap_dot_empty_left in H4).
187-
destruct v; trivial. exfalso.
187+
destruct v as [v|]; trivial. exfalso.
188188
have h := H3 k v.
189189
assert (gmap_dot {[k := Some (Some v)]} ms !! k = Some (Some (Some v))) as Q.
190190
{
@@ -200,8 +200,8 @@ Proof.
200200
assert (i < j) as ihfs.
201201
{
202202
have hij : Decision (i < j) by solve_decision. destruct hij; trivial.
203-
assert (j ≤ i) by lia.
204-
have mm := H10 j H5 H11. destruct mm as [k1 [v1 mm]].
203+
assert (j ≤ i) as Hle by lia.
204+
have mm := H10 j H5 Hle. destruct mm as [k1 [v1 mm]].
205205
rewrite lookup_gmap_dot_3mid in mm.
206206
- rewrite lookup_singleton in mm. inversion mm.
207207
- trivial.
@@ -220,19 +220,19 @@ Lemma preserves_lt_fixed_size j a b slots
220220
: ∀ (i : nat) (e : option (option (Key * Value))),
221221
gmap_dot {[j := Some b]} slots !! i = Some e → i < ht_fixed_size.
222222
Proof.
223-
intros.
223+
intros i e Heq.
224224
(*have h : Decision (i = j) by solve_decision. destruct h.*)
225-
destruct (gmap_dot {[j := Some a]} slots !! i) eqn:gd.
225+
destruct (gmap_dot {[j := Some a]} slots !! i) as [o|] eqn:gd.
226226
- exact (ltfs i o gd).
227-
- exfalso. unfold gmap_dot in H, gd.
228-
rewrite lookup_merge in H.
227+
- exfalso. unfold gmap_dot in Heq, gd.
228+
rewrite lookup_merge in Heq.
229229
rewrite lookup_merge in gd.
230230
unfold diag_None, gmerge in *.
231231
have h : Decision (j = i) by solve_decision. destruct h.
232232
+ subst j.
233-
rewrite lookup_singleton in H.
233+
rewrite lookup_singleton in Heq.
234234
rewrite lookup_singleton in gd. destruct (slots !! i); discriminate.
235-
+ rewrite lookup_singleton_ne in H; trivial.
235+
+ rewrite lookup_singleton_ne in Heq; trivial.
236236
rewrite lookup_singleton_ne in gd; trivial.
237237
destruct (slots !! i); discriminate.
238238
Qed.
@@ -296,8 +296,7 @@ Lemma preserves_unique_keys j k v v1 slots
296296
∧ gmap_dot {[j := Some (Some (k, v))]} slots !! i2 = Some (Some (Some (k0, v3)))
297297
→ i1 = i2.
298298
Proof.
299-
intros.
300-
destruct H as [H H0].
299+
intros i1 i2 k0 v2 v3 [H H0].
301300
have ed : Decision (i1 = j) by solve_decision. destruct ed.
302301
- have ed : Decision (i2 = j) by solve_decision. destruct ed.
303302
+ subst. trivial.
@@ -344,8 +343,8 @@ Proof.
344343
rewrite lookup_singleton in ac.
345344
destruct (slots !! j), (slots0 !! j); discriminate.
346345
- rewrite lookup_singleton_ne in ac; trivial.
347-
destruct (slots !! i2), (slots0 !! i2); try inversion ac; try discriminate.
348-
+ subst o. inversion zz. inversion zz. subst k. contradiction.
346+
destruct (slots !! i2) as [a1|], (slots0 !! i2) as [b1|]; try inversion ac; try discriminate.
347+
+ subst b1. inversion zz. inversion zz. subst k. contradiction.
349348
Qed.
350349

351350
Lemma pukn_helper j k0 k v v1 slots v0 ms ms0 slots0 i2 v2

src/guarding/internal/auth_frag_util.v

Lines changed: 14 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -1,9 +1,7 @@
1+
From iris.prelude Require Import options.
12
From iris.algebra Require Export cmra updates.
23
From iris.algebra Require Import proofmode_classes.
34
From iris.algebra Require Import auth.
4-
From iris.algebra Require Import functions.
5-
From iris.algebra Require Import gmap.
6-
From iris.prelude Require Import options.
75

86
From iris.base_logic Require Import upred.
97
From iris.base_logic.lib Require Export own iprop.
@@ -12,8 +10,6 @@ From iris.proofmode Require Import ltac_tactics.
1210
From iris.proofmode Require Import tactics.
1311
From iris.proofmode Require Import coq_tactics.
1412

15-
From stdpp Require Import numbers.
16-
1713
Require Import guarding.own_and.
1814

1915
Section AuthFragUtil.
@@ -25,9 +21,9 @@ Context `{Disc : CmraDiscrete C}.
2521

2622
Lemma auth_op_rhs_is_frag (p: C) z (val : ✓ (● p ⋅ z)) : ∃ q , z = ◯ q.
2723
Proof.
28-
destruct z. exists view_frag_proj. rename view_frag_proj into f.
24+
destruct z as [a f]. exists f.
2925
unfold "◯", "◯V". f_equal.
30-
destruct view_auth_proj; trivial.
26+
destruct a as [p0|]; trivial.
3127
exfalso.
3228
unfold "●", "●V" in val.
3329

@@ -42,7 +38,7 @@ Proof.
4238

4339
rewrite /op /view_op_instance /view_auth_proj /view_frag_proj in val2.
4440
rewrite /valid /view_valid_instance /view_auth_proj in val2.
45-
destruct p0.
41+
destruct p0 as [d a].
4642
assert (Some (DfracOwn 1, to_agree p) ⋅ Some (d, a)
4743
= Some (DfracOwn 1 ⋅ d, to_agree p ⋅ a)) as Q.
4844
{ trivial. }
@@ -66,8 +62,8 @@ Lemma rhs_has_auth (x y : C) (q1 q2: auth C)
6662
Proof.
6763
have ao := auth_op_rhs_is_frag x q1 v.
6864
destruct ao as [q ao]. subst q1.
69-
destruct q2.
70-
exists (◯ view_frag_proj).
65+
destruct q2 as [a f].
66+
exists (◯ f).
7167
unfold "●", "●V", "◯", "◯V".
7268

7369
rewrite view_op_rw.
@@ -78,10 +74,10 @@ Proof.
7874
rewrite view_op_rw in eq.
7975
rewrite view_op_rw in eq.
8076

81-
inversion eq.
82-
unfold view.view_auth_proj in H.
83-
setoid_rewrite H.
84-
+ destruct view_auth_proj; trivial.
77+
inversion eq as [Q R].
78+
unfold view.view_auth_proj in Q.
79+
setoid_rewrite Q.
80+
+ destruct a; trivial.
8581
- rewrite left_id. trivial.
8682
Qed.
8783

@@ -258,11 +254,7 @@ Qed.
258254
Lemma auth_frag_disjointness_helper (q r : C) (a b : auth C)
259255
(eq1 : ◯ q ⋅ a ≡ ● r ⋅ b) : ◯ q ≼ b.
260256
Proof.
261-
destruct a, b.
262-
rename view_frag_proj into f.
263-
rename view_auth_proj into g.
264-
rename view_frag_proj0 into f0.
265-
rename view_auth_proj0 into g0.
257+
destruct a as [g f], b as [g0 f0].
266258
unfold "◯", "◯V" in eq1.
267259
unfold "●", "●V" in eq1.
268260
unfold "◯", "◯V".
@@ -273,9 +265,9 @@ Proof.
273265
replace (@View C C (@auth_view_rel C) (Some (DfracOwn 1, to_agree r)) ε ⋅ @View C C (@auth_view_rel C) g0 f0)
274266
with (@View C C (@auth_view_rel C) (Some (DfracOwn 1, to_agree r) ⋅ g0) (ε ⋅ f0)) in eq1 by trivial.
275267

276-
inversion eq1.
277-
unfold view_frag_proj in H0.
278-
generalize H0.
268+
inversion eq1 as [Q R].
269+
unfold view_frag_proj in R.
270+
generalize R.
279271
rewrite ucmra_unit_left_id.
280272
intro X.
281273

src/guarding/internal/wsat_util.v

Lines changed: 4 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -1,10 +1,9 @@
1-
From stdpp Require Export coPset.
1+
From iris.prelude Require Import options.
22
From iris.algebra Require Import gmap_view gset coPset.
33
From iris.proofmode Require Import proofmode.
4-
From iris.base_logic.lib Require Export own.
5-
From iris.base_logic.lib Require Export wsat.
6-
From iris.prelude Require Import options.
4+
From iris.base_logic.lib Require Export own wsat.
75
Import uPred.
6+
From stdpp Require Export coPset.
87
From iris.algebra Require Import functions.
98

109
Section wsat_util.
@@ -212,7 +211,7 @@ Lemma own_updateP_extra `{i : !inG Σ A} `{j : !inG Σ B} (P : A → gname → P
212211
(∀ (n : nat) (mz : option A) (N : gname → Prop) , ✓{n} (a ⋅? mz) → pred_finite N → ∃ (y: A) (γ': gname), P y γ' ∧ ✓{n} (y ⋅? mz) ∧ ¬ N γ') →
213212
(✓ b) →
214213
own γ a ⊢ |==> ∃ a' γ', ⌜P a' γ'⌝ ∗ own γ a' ∗ own γ' b.
215-
Proof using H Σ.
214+
Proof.
216215
intros Hupd Hvalb. rewrite !own.own_eq.
217216
rewrite -(bupd_mono (∃ m,
218217
⌜ ∃ a' γ', m = (own.iRes_singleton γ a' ⋅ own.iRes_singleton γ' b) ∧ P a' γ' ⌝ ∧ uPred_ownM m)%I).

0 commit comments

Comments
 (0)