-
Notifications
You must be signed in to change notification settings - Fork 1
/
Copy pathhoare.v
159 lines (141 loc) · 6.24 KB
/
hoare.v
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
From iris.program_logic Require Export weakestpre.
From iris.base_logic.lib Require Export viewshifts.
From iris.proofmode Require Import tactics.
Set Default Proof Using "Type".
Definition ht `{irisG Λ Σ} (s : stuckness) (E : coPset) (P : iProp Σ)
(e : expr Λ) (Φ : val Λ → iProp Σ) : iProp Σ :=
(□ (P -∗ WP e @ s; E {{ Φ }}))%I.
Instance: Params (@ht) 5.
Notation "{{ P } } e @ s ; E {{ Φ } }" := (ht s E P%I e%E Φ%I)
(at level 20, P, e, Φ at level 200,
format "{{ P } } e @ s ; E {{ Φ } }") : stdpp_scope.
Notation "{{ P } } e @ E {{ Φ } }" := (ht NotStuck E P%I e%E Φ%I)
(at level 20, P, e, Φ at level 200,
format "{{ P } } e @ E {{ Φ } }") : stdpp_scope.
Notation "{{ P } } e @ E ? {{ Φ } }" := (ht MaybeStuck E P%I e%E Φ%I)
(at level 20, P, e, Φ at level 200,
format "{{ P } } e @ E ? {{ Φ } }") : stdpp_scope.
Notation "{{ P } } e {{ Φ } }" := (ht NotStuck ⊤ P%I e%E Φ%I)
(at level 20, P, e, Φ at level 200,
format "{{ P } } e {{ Φ } }") : stdpp_scope.
Notation "{{ P } } e ? {{ Φ } }" := (ht MaybeStuck ⊤ P%I e%E Φ%I)
(at level 20, P, e, Φ at level 200,
format "{{ P } } e ? {{ Φ } }") : stdpp_scope.
Notation "{{ P } } e @ s ; E {{ v , Q } }" := (ht s E P%I e%E (λ v, Q)%I)
(at level 20, P, e, Q at level 200,
format "{{ P } } e @ s ; E {{ v , Q } }") : stdpp_scope.
Notation "{{ P } } e @ E {{ v , Q } }" := (ht NotStuck E P%I e%E (λ v, Q)%I)
(at level 20, P, e, Q at level 200,
format "{{ P } } e @ E {{ v , Q } }") : stdpp_scope.
Notation "{{ P } } e @ E ? {{ v , Q } }" := (ht MaybeStuck E P%I e%E (λ v, Q)%I)
(at level 20, P, e, Q at level 200,
format "{{ P } } e @ E ? {{ v , Q } }") : stdpp_scope.
Notation "{{ P } } e {{ v , Q } }" := (ht NotStuck ⊤ P%I e%E (λ v, Q)%I)
(at level 20, P, e, Q at level 200,
format "{{ P } } e {{ v , Q } }") : stdpp_scope.
Notation "{{ P } } e ? {{ v , Q } }" := (ht MaybeStuck ⊤ P%I e%E (λ v, Q)%I)
(at level 20, P, e, Q at level 200,
format "{{ P } } e ? {{ v , Q } }") : stdpp_scope.
Section hoare.
Context `{irisG Λ Σ}.
Implicit Types s : stuckness.
Implicit Types P Q : iProp Σ.
Implicit Types Φ Ψ : val Λ → iProp Σ.
Implicit Types v : val Λ.
Import uPred.
Global Instance ht_ne s E n :
Proper (dist n ==> eq ==> pointwise_relation _ (dist n) ==> dist n) (ht s E).
Proof. solve_proper. Qed.
Global Instance ht_proper s E :
Proper ((≡) ==> eq ==> pointwise_relation _ (≡) ==> (≡)) (ht s E).
Proof. solve_proper. Qed.
Lemma ht_mono s E P P' Φ Φ' e :
(P ⊢ P') → (∀ v, Φ' v ⊢ Φ v) → {{ P' }} e @ s; E {{ Φ' }} ⊢ {{ P }} e @ s; E {{ Φ }}.
Proof. by intros; apply persistently_mono, wand_mono, wp_mono. Qed.
Lemma ht_stuck_mono s1 s2 E P Φ e :
s1 ⊑ s2 → {{ P }} e @ s1; E {{ Φ }} ⊢ {{ P }} e @ s2; E {{ Φ }}.
Proof. by intros; apply persistently_mono, wand_mono, wp_stuck_mono. Qed.
Global Instance ht_mono' s E :
Proper (flip (⊢) ==> eq ==> pointwise_relation _ (⊢) ==> (⊢)) (ht s E).
Proof. solve_proper. Qed.
Lemma ht_alt s E P Φ e : (P ⊢ WP e @ s; E {{ Φ }}) → {{ P }} e @ s; E {{ Φ }}.
Proof. iIntros (Hwp) "!# HP". by iApply Hwp. Qed.
Lemma ht_val s E v : {{ True }} of_val v @ s; E {{ v', ⌜v = v'⌝ }}.
Proof. iIntros "!# _". by iApply wp_value'. Qed.
Lemma ht_vs s E P P' Φ Φ' e :
(P ={E}=> P') ∧ {{ P' }} e @ s; E {{ Φ' }} ∧ (∀ v, Φ' v ={E}=> Φ v)
⊢ {{ P }} e @ s; E {{ Φ }}.
Proof.
iIntros "(#Hvs & #Hwp & #HΦ) !# HP". iMod ("Hvs" with "HP") as "HP".
iApply wp_fupd. iApply (wp_wand with "[HP]"); [by iApply "Hwp"|].
iIntros (v) "Hv". by iApply "HΦ".
Qed.
Lemma ht_atomic s E1 E2 P P' Φ Φ' e `{!Atomic (stuckness_to_atomicity s) e} :
(P ={E1,E2}=> P') ∧ {{ P' }} e @ s; E2 {{ Φ' }} ∧ (∀ v, Φ' v ={E2,E1}=> Φ v)
⊢ {{ P }} e @ s; E1 {{ Φ }}.
Proof.
iIntros "(#Hvs & #Hwp & #HΦ) !# HP". iApply (wp_atomic _ _ E2); auto.
iMod ("Hvs" with "HP") as "HP". iModIntro.
iApply (wp_wand with "[HP]"); [by iApply "Hwp"|].
iIntros (v) "Hv". by iApply "HΦ".
Qed.
(*
Lemma ht_bind `{LanguageCtx Λ K} s E P Φ Φ' e :
{{ P }} e @ s; E {{ Φ }} ∧ (∀ v, {{ Φ v }} K (of_val v) @ s; E {{ Φ' }})
⊢ {{ P }} K e @ s; E {{ Φ' }}.
Proof.
iIntros "[#Hwpe #HwpK] !# HP". iApply wp_bind.
iApply (wp_wand with "[HP]"); [by iApply "Hwpe"|].
iIntros (v) "Hv". by iApply "HwpK".
Qed.
Lemma ht_stuck_weaken s E P Φ e :
{{ P }} e @ s; E {{ Φ }} ⊢ {{ P }} e @ E ?{{ Φ }}.
Proof.
by iIntros "#Hwp !# ?"; iApply wp_stuck_weaken; iApply "Hwp".
Qed.
Lemma ht_mask_weaken s E1 E2 P Φ e :
E1 ⊆ E2 → {{ P }} e @ s; E1 {{ Φ }} ⊢ {{ P }} e @ s; E2 {{ Φ }}.
Proof.
iIntros (?) "#Hwp !# HP". iApply (wp_mask_mono _ E1 E2); try done.
by iApply "Hwp".
Qed.
Lemma ht_frame_l s E P Φ R e :
{{ P }} e @ s; E {{ Φ }} ⊢ {{ R ∗ P }} e @ s; E {{ v, R ∗ Φ v }}.
Proof. iIntros "#Hwp !# [$ HP]". by iApply "Hwp". Qed.
Lemma ht_frame_r s E P Φ R e :
{{ P }} e @ s; E {{ Φ }} ⊢ {{ P ∗ R }} e @ s; E {{ v, Φ v ∗ R }}.
Proof. iIntros "#Hwp !# [HP $]". by iApply "Hwp". Qed.
Lemma ht_frame_step_l s E1 E2 P R1 R2 e Φ :
to_val e = None → E2 ⊆ E1 →
(R1 ={E1,E2}=> ▷ |={E2,E1}=> R2) ∧ {{ P }} e @ s; E2 {{ Φ }}
⊢ {{ R1 ∗ P }} e @ s; E1 {{ λ v, R2 ∗ Φ v }}.
Proof.
iIntros (??) "[#Hvs #Hwp] !# [HR HP]".
iApply (wp_frame_step_l _ E1 E2); try done.
iSplitL "HR"; [by iApply "Hvs"|by iApply "Hwp"].
Qed.
Lemma ht_frame_step_r s E1 E2 P R1 R2 e Φ :
to_val e = None → E2 ⊆ E1 →
(R1 ={E1,E2}=> ▷ |={E2,E1}=> R2) ∧ {{ P }} e @ s; E2 {{ Φ }}
⊢ {{ P ∗ R1 }} e @ s; E1 {{ λ v, Φ v ∗ R2 }}.
Proof.
iIntros (??) "[#Hvs #Hwp] !# [HP HR]".
iApply (wp_frame_step_r _ E1 E2); try done.
iSplitR "HR"; [by iApply "Hwp"|by iApply "Hvs"].
Qed.
Lemma ht_frame_step_l' s E P R e Φ :
to_val e = None →
{{ P }} e @ s; E {{ Φ }} ⊢ {{ ▷ R ∗ P }} e @ s; E {{ v, R ∗ Φ v }}.
Proof.
iIntros (?) "#Hwp !# [HR HP]".
iApply wp_frame_step_l'; try done. iFrame "HR". by iApply "Hwp".
Qed.
Lemma ht_frame_step_r' s E P Φ R e :
to_val e = None →
{{ P }} e @ s; E {{ Φ }} ⊢ {{ P ∗ ▷ R }} e @ s; E {{ v, Φ v ∗ R }}.
Proof.
iIntros (?) "#Hwp !# [HP HR]".
iApply wp_frame_step_r'; try done. iFrame "HR". by iApply "Hwp".
Qed.
*)
End hoare.