@@ -202,98 +202,22 @@ Proof.
202
202
iFrame "p".
203
203
Qed .
204
204
205
- Local Lemma lfguards_and_with_lhs (P Q S : iProp Σ) F n
206
- (point: point_prop S)
207
- (qrx: (P ∧ Q ⊢ S))
208
- : (
209
- (P &&{F; n}&&$> Q)
210
- ⊢
211
- (P &&{F; n}&&$> S)
212
- ).
213
- Proof .
214
- iIntros "[pq kf]".
215
- unfold lfguards, lguards_with.
216
- iFrame "kf".
217
- iIntros (T).
218
- iDestruct ("pq" $! T) as "pq".
219
- iIntros "k".
220
- iAssert (P ∗ (P -∗ storage_bulk_inv F ∗ T) -∗ ▷^n ◇ S)%I with "[pq]" as "X".
221
- {
222
- iIntros "l".
223
- iAssert (▷^n ◇ (P ∧ Q))%I with "[pq l]" as "J". {
224
- rewrite bi.except_0_and.
225
- rewrite bi.laterN_and.
226
- iSplit.
227
- { iDestruct "l" as "[l l']". iModIntro. iFrame. }
228
- { iDestruct ("pq" with "l") as "[r j]". iNext. iMod "j". iFrame "r". }
229
- }
230
- iNext. iMod "J" as "J". iModIntro. iApply qrx. iFrame "J".
231
- }
232
- iDestruct (own_separates_out_except0_point_later _ S with "[X k]") as "J".
233
- { trivial. }
234
- { iFrame "X". iFrame "k". }
235
- iDestruct "J" as "[J T]".
236
- iNext. iMod "J" as "J". iModIntro. iFrame "J". iIntros "o".
237
- iDestruct ("T" with "o") as "[p m]". iApply "m". iFrame "p".
238
- Qed .
239
-
240
- Local Lemma lfguards_or_with_lhs (P R S : iProp Σ) F n
241
- (qrx: (P ∧ R ⊢ False))
242
- : (
243
- (P &&{F; n}&&$> (R ∨ S))
244
- ⊢
245
- (P &&{F; n}&&$> S)
246
- ).
247
- Proof .
248
- iIntros "[prs kf]".
249
- unfold lfguards, lguards_with.
250
- iFrame "kf".
251
- iIntros (T).
252
- iDestruct ("prs" $! T) as "prs".
253
- iIntros "k".
254
- iAssert (▷^n ((◇ P) ∧ (◇ ((R ∨ S) ∗ (R ∨ S -∗ storage_bulk_inv F ∗ T)))))%I with "[prs k]" as "X".
255
- { iSplit.
256
- { iModIntro. iDestruct "k" as "[p p2]". iFrame. }
257
- { iDestruct ("prs" with "k") as "rs". iFrame "rs". }
258
- }
259
- rewrite <- bi.except_0_and. iNext. iMod "X" as "X". iModIntro.
260
- iAssert
261
- (P ∧ (
262
- (R ∗ (R ∨ S -∗ storage_bulk_inv F ∗ T))
263
- ∨
264
- (S ∗ (R ∨ S -∗ storage_bulk_inv F ∗ T))
265
- ))%I with "[X]" as "X".
266
- { iSplit. { iDestruct "X" as "[Q _]". iFrame "Q". }
267
- iDestruct "X" as "[_ [[r|s] rest]]".
268
- { iLeft. iFrame. } { iRight. iFrame. }
269
- }
270
- rewrite bi.and_or_l.
271
- iDestruct "X" as "[a|a]".
272
- { iExFalso. iApply qrx. iSplit.
273
- { iDestruct "a" as "[a _]". iFrame. }
274
- { iDestruct "a" as "[_ [r _]]". iFrame. }
275
- }
276
- {
277
- iDestruct "a" as "[_ [s t]]".
278
- iFrame "s". iIntros "s". iApply "t". iRight. iFrame "s".
279
- }
280
- Qed .
281
-
282
- Local Lemma lfguards_exists_with_lhs X (P : iProp Σ) (S R : X -> iProp Σ) F n
283
- (pr_impl_s: (∀ x, P ∧ R x ⊢ S x))
205
+ Local Lemma lfguards_exists_strengthen X (P Q : iProp Σ) (S R : X -> iProp Σ) F n
206
+ (pr_impl_s: (∀ x, Q ∧ R x ⊢ S x))
284
207
(pers: ∀ x, Persistent (S x))
285
208
: (
209
+ (P &&{F; n}&&$> Q) ∗
286
210
(P &&{F; n}&&$> (∃ (x: X), R x))
287
211
⊢
288
212
(P &&{F; n}&&$> (∃ (x: X), R x ∗ S x))
289
213
).
290
214
Proof .
291
- iIntros "[prs kf]".
215
+ iIntros "[[pq _] [ prs kf] ]".
292
216
unfold lfguards, lguards_with.
293
217
iFrame "kf". iIntros (T). iDestruct ("prs" $! T) as "prs". iIntros "k".
294
- iAssert (▷^n ((◇ P ) ∧ (◇ ((∃ x, R x) ∗ ((∃ x, R x) -∗ storage_bulk_inv F ∗ T)))))%I with "[prs k]" as "X".
218
+ iAssert (▷^n ((◇ Q ) ∧ (◇ ((∃ x, R x) ∗ ((∃ x, R x) -∗ storage_bulk_inv F ∗ T)))))%I with "[pq prs k]" as "X".
295
219
{ iSplit.
296
- { iModIntro. iDestruct "k" as "[p p2 ]". iFrame. }
220
+ { iDestruct ("pq" with "k") as "[dq _ ]". iFrame "dq" . }
297
221
{ iDestruct ("prs" with "k") as "rs". iFrame "rs". } }
298
222
rewrite <- bi.except_0_and. iNext. iMod "X" as "X". iModIntro.
299
223
rewrite bi.sep_exist_r. rewrite bi.and_exist_l. iDestruct "X" as (x) "X".
@@ -304,49 +228,6 @@ Proof.
304
228
iSplitL "r". { iExists x. iFrame "r". iFrame "s". }
305
229
iIntros "j". iDestruct "j" as (x0) "[r s2]". iApply "back". iExists x0. iFrame "r".
306
230
Qed .
307
-
308
- Local Lemma lfguards_or (P Q R S : iProp Σ) F (n: nat)
309
- (qrx: (Q ∧ R ⊢ False))
310
- : (
311
- (P &&{F; n}&&$> Q) ∗ (P &&{F; n}&&$> (R ∨ S))
312
- ⊢
313
- (P &&{F; n}&&$> S)
314
- ).
315
- Proof .
316
- iIntros "[[pq kf] [prs _]]".
317
- unfold lfguards, lguards_with.
318
- iFrame "kf".
319
- iIntros (T).
320
- iDestruct ("pq" $! T) as "pq".
321
- iDestruct ("prs" $! T) as "prs".
322
- iIntros "k".
323
- iAssert (▷^n ((◇ Q) ∧ (◇ ((R ∨ S) ∗ (R ∨ S -∗ storage_bulk_inv F ∗ T)))))%I with "[pq prs k]" as "X".
324
- { iSplit.
325
- { iDestruct ("pq" with "k") as "[dq _]". iFrame "dq". }
326
- { iDestruct ("prs" with "k") as "rs". iFrame "rs". }
327
- }
328
- rewrite <- bi.except_0_and. iNext. iMod "X" as "X". iModIntro.
329
- iAssert
330
- (Q ∧ (
331
- (R ∗ (R ∨ S -∗ storage_bulk_inv F ∗ T))
332
- ∨
333
- (S ∗ (R ∨ S -∗ storage_bulk_inv F ∗ T))
334
- ))%I with "[X]" as "X".
335
- { iSplit. { iDestruct "X" as "[Q _]". iFrame "Q". }
336
- iDestruct "X" as "[_ [[r|s] rest]]".
337
- { iLeft. iFrame. } { iRight. iFrame. }
338
- }
339
- rewrite bi.and_or_l.
340
- iDestruct "X" as "[a|a]".
341
- { iExFalso. iApply qrx. iSplit.
342
- { iDestruct "a" as "[a _]". iFrame. }
343
- { iDestruct "a" as "[_ [r _]]". iFrame. }
344
- }
345
- {
346
- iDestruct "a" as "[_ [s t]]".
347
- iFrame "s". iIntros "s". iApply "t". iRight. iFrame "s".
348
- }
349
- Qed .
350
231
351
232
Local Lemma lfguards_weaken_except0 P Q n
352
233
: □(P -∗ ▷^n ◇ (Q ∗ (Q -∗ P))) ⊢ P &&{ ∅ ; n }&&$> Q.
@@ -1332,18 +1213,27 @@ Qed.
1332
1213
1333
1214
(** Guarding and disjunction * *)
1334
1215
1335
- Lemma guards_cancel_or_with_lhs (P R S : iProp Σ) F n
1336
- (pr_impl_false: (P ∧ R ⊢ False))
1216
+ Lemma guards_strengthen_exists X (P Q : iProp Σ) (S R : X -> iProp Σ) F n
1217
+ (pr_impl_s: (∀ x, Q ∧ R x ⊢ S x))
1218
+ (pers: ∀ x, Persistent (S x))
1337
1219
: (
1338
- (P &&{F; n}&&> (R ∨ S))
1220
+ (P &&{F; n}&&> Q) ∗
1221
+ (P &&{F; n}&&> (∃ (x: X), R x))
1339
1222
⊢
1340
- (P &&{F; n}&&> S )
1223
+ (P &&{F; n}&&> (∃ (x: X), R x ∗ S x) )
1341
1224
).
1342
- Proof .
1225
+ Proof .
1343
1226
rewrite guards_unseal. unfold guards_def.
1344
- iIntros "#a". iDestruct "a" as (m1) "[%cond1 q]".
1345
- iExists m1. iModIntro. iSplit. { iPureIntro. set_solver. }
1346
- iApply lfguards_or_with_lhs; trivial. trivial.
1227
+ iIntros "[#a #b]".
1228
+ iDestruct "a" as (m1) "[%cond1 q]".
1229
+ iDestruct "b" as (m2) "[%cond2 r]".
1230
+ iExists (m1 ∪ m2).
1231
+ iModIntro. iSplit.
1232
+ { iPureIntro. set_solver. }
1233
+ iDestruct (lfguards_weaken_mask_2 P Q P (∃ x, R x) m1 m2 with "[q r]") as "[q1 r1]".
1234
+ { iFrame "q". iFrame "r". }
1235
+ iApply (lfguards_exists_strengthen X P Q S R (m1 ∪ m2)); trivial.
1236
+ iFrame "#".
1347
1237
Qed .
1348
1238
1349
1239
Lemma guards_strengthen_exists_with_lhs X (P : iProp Σ) (S R : X -> iProp Σ) F n
@@ -1355,10 +1245,9 @@ Lemma guards_strengthen_exists_with_lhs X (P : iProp Σ) (S R : X -> iProp Σ) F
1355
1245
(P &&{F; n}&&> (∃ (x: X), R x ∗ S x))
1356
1246
).
1357
1247
Proof .
1358
- rewrite guards_unseal. unfold guards_def.
1359
- iIntros "#a". iDestruct "a" as (m1) "[%cond1 q]".
1360
- iExists m1. iModIntro. iSplit. { iPureIntro. set_solver. }
1361
- iApply lfguards_exists_with_lhs; trivial.
1248
+ iIntros "#G".
1249
+ iApply (guards_strengthen_exists X P P S R F n); trivial. iFrame "G".
1250
+ iApply guards_refl.
1362
1251
Qed .
1363
1252
1364
1253
Lemma guards_cancel_or (P Q R S : iProp Σ) F1 F2 n
@@ -1368,20 +1257,46 @@ Lemma guards_cancel_or (P Q R S : iProp Σ) F1 F2 n
1368
1257
⊢
1369
1258
(P &&{F1 ∪ F2; n}&&> S)
1370
1259
).
1371
- Proof .
1372
- rewrite guards_unseal. unfold guards_def.
1373
- iIntros "[#a #b]".
1374
- iDestruct "a" as (m1) "[%cond1 q]".
1375
- iDestruct "b" as (m2) "[%cond2 r]".
1376
- iExists (m1 ∪ m2).
1377
- iModIntro. iSplit.
1378
- { iPureIntro. set_solver. }
1379
- iDestruct (lfguards_weaken_mask_2 P Q P (R ∨ S) m1 m2 with "[q r]") as "[q1 r1]".
1380
- { iFrame "q". iFrame "r". }
1381
- iApply (lfguards_or P Q R S (m1 ∪ m2)); trivial.
1382
- iFrame "#".
1260
+ Proof .
1261
+ assert (R ∨ S ⊣⊢ (∃ (b: bool), if b then R else S)) as He1.
1262
+ { iIntros. iSplit.
1263
+ - iIntros "RS". iDestruct "RS" as "[R|S]".
1264
+ + iExists true. iFrame. + iExists false. iFrame.
1265
+ - iIntros "RS". iDestruct "RS" as (b) "RS". destruct b.
1266
+ + iLeft. iFrame. + iRight. iFrame.
1267
+ }
1268
+ iIntros "[#G1 #G2]".
1269
+ iDestruct (guards_weaken_mask F1 (F1 ∪ F2) P Q n with "G1") as "#H1". { set_solver. }
1270
+ iDestruct (guards_weaken_mask F2 (F1 ∪ F2) P (R ∨ S) n with "G2") as "#H2". { set_solver. }
1271
+ setoid_rewrite He1.
1272
+ iDestruct (guards_strengthen_exists bool P Q
1273
+ (λ b, ⌜b = false⌝)%I (λ b, if b then R else S) (F1 ∪ F2) n with "[]") as "A".
1274
+ { intros b. iIntros "H". destruct b.
1275
+ - iExFalso. iApply qrx. iFrame "H".
1276
+ - iPureIntro. trivial.
1277
+ }
1278
+ { iFrame "#". }
1279
+ assert (S ⊣⊢ (∃ x : bool, (if x then R else S) ∗ ⌜x = false⌝)) as He2.
1280
+ { iIntros. iSplit.
1281
+ - iIntros "S". iExists false. iFrame. iPureIntro. trivial.
1282
+ - iIntros "S". iDestruct "S" as (b) "[S %xf]". rewrite xf. iFrame.
1283
+ }
1284
+ setoid_rewrite <- He2. iFrame "A".
1383
1285
Qed .
1384
1286
1287
+ Lemma guards_cancel_or_with_lhs (P R S : iProp Σ) F n
1288
+ (pr_impl_false: (P ∧ R ⊢ False))
1289
+ : (
1290
+ (P &&{F; n}&&> (R ∨ S))
1291
+ ⊢
1292
+ (P &&{F; n}&&> S)
1293
+ ).
1294
+ Proof .
1295
+ iIntros "#G".
1296
+ iDestruct (guards_cancel_or P P R S F F n with "[]") as "J"; trivial.
1297
+ { iFrame "G". iApply guards_refl. }
1298
+ replace (F ∪ F) with F by set_solver. iFrame "J".
1299
+ Qed .
1385
1300
1386
1301
(*
1387
1302
Lemma guards_exists {X} (x0: X) (F: X -> iProp Σ) (P: iProp Σ) E
0 commit comments