diff --git a/iris/base_logic/lib/fancy_updates.v b/iris/base_logic/lib/fancy_updates.v index e09517399f6173e5d83a4a1d65dae3dc7b7b8a1f..510986781c20b5e1e5d102ebc7eacb1e26e80469 100644 --- a/iris/base_logic/lib/fancy_updates.v +++ b/iris/base_logic/lib/fancy_updates.v @@ -161,7 +161,7 @@ Qed. where ["Hcredit"] is a credit available in the context and ["HP"] is the assumption from which a later should be stripped. *) Lemma lc_fupd_elim_later `{!invGS Σ} `{!HasLc Σ} E P : - £1 -∗ (▷ P) -∗ |={E}=> P. + £ 1 -∗ (▷ P) -∗ |={E}=> P. Proof. iIntros "Hf Hupd". rewrite uPred_fupd_unseal /uPred_fupd_def has_credits. @@ -174,7 +174,7 @@ Qed. This is typically used as [iApply (lc_fupd_add_later with "Hcredit")], where ["Hcredit"] is a credit available in the context. *) Lemma lc_fupd_add_later `{!invGS Σ} `{!HasLc Σ} E1 E2 P : - £1 -∗ (▷ |={E1, E2}=> P) -∗ |={E1, E2}=> P. + £ 1 -∗ (▷ |={E1, E2}=> P) -∗ |={E1, E2}=> P. Proof. iIntros "Hf Hupd". iApply (fupd_trans E1 E1). iApply (lc_fupd_elim_later with "Hf Hupd"). diff --git a/iris/program_logic/ectx_lifting.v b/iris/program_logic/ectx_lifting.v index a6ecf16adf3246a26ec564ce22c3d063b0a9b27c..5c78293c0dbe2e8057672405d0d539684414673e 100644 --- a/iris/program_logic/ectx_lifting.v +++ b/iris/program_logic/ectx_lifting.v @@ -19,7 +19,7 @@ Lemma wp_lift_head_step_fupd {s E Φ} e1 : to_val e1 = None → (∀ σ1 ns κ κs nt, state_interp σ1 ns (κ ++ κs) nt ={E,∅}=∗ ⌜head_reducible e1 σ1⌠∗ - ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £1 ={∅}=∗ ▷ |={∅,E}=> + ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £ 1 ={∅}=∗ ▷ |={∅,E}=> state_interp σ2 (S ns) κs (length efs + nt) ∗ WP e2 @ s; E {{ Φ }} ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ fork_post }}) @@ -35,7 +35,7 @@ Lemma wp_lift_head_step {s E Φ} e1 : to_val e1 = None → (∀ σ1 ns κ κs nt, state_interp σ1 ns (κ ++ κs) nt ={E,∅}=∗ ⌜head_reducible e1 σ1⌠∗ - ▷ ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £1 ={∅,E}=∗ + ▷ ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £ 1 ={∅,E}=∗ state_interp σ2 (S ns) κs (length efs + nt) ∗ WP e2 @ s; E {{ Φ }} ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ fork_post }}) @@ -69,7 +69,7 @@ Lemma wp_lift_atomic_head_step_fupd {s E1 E2 Φ} e1 : to_val e1 = None → (∀ σ1 ns κ κs nt, state_interp σ1 ns (κ ++ κs) nt ={E1}=∗ ⌜head_reducible e1 σ1⌠∗ - ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £1 ={E1}[E2]▷=∗ + ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £ 1 ={E1}[E2]▷=∗ state_interp σ2 (S ns) κs (length efs + nt) ∗ from_option Φ False (to_val e2) ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ fork_post }}) @@ -85,7 +85,7 @@ Lemma wp_lift_atomic_head_step {s E Φ} e1 : to_val e1 = None → (∀ σ1 ns κ κs nt, state_interp σ1 ns (κ ++ κs) nt ={E}=∗ ⌜head_reducible e1 σ1⌠∗ - ▷ ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £1 ={E}=∗ + ▷ ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £ 1 ={E}=∗ state_interp σ2 (S ns) κs (length efs + nt) ∗ from_option Φ False (to_val e2) ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ fork_post }}) @@ -101,7 +101,7 @@ Lemma wp_lift_atomic_head_step_no_fork_fupd {s E1 E2 Φ} e1 : to_val e1 = None → (∀ σ1 ns κ κs nt, state_interp σ1 ns (κ ++ κs) nt ={E1}=∗ ⌜head_reducible e1 σ1⌠∗ - ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £1 ={E1}[E2]▷=∗ + ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £ 1 ={E1}[E2]▷=∗ ⌜efs = []⌠∗ state_interp σ2 (S ns) κs nt ∗ from_option Φ False (to_val e2)) ⊢ WP e1 @ s; E1 {{ Φ }}. Proof. @@ -116,7 +116,7 @@ Lemma wp_lift_atomic_head_step_no_fork {s E Φ} e1 : to_val e1 = None → (∀ σ1 ns κ κs nt, state_interp σ1 ns (κ ++ κs) nt ={E}=∗ ⌜head_reducible e1 σ1⌠∗ - ▷ ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £1 ={E}=∗ + ▷ ∀ e2 σ2 efs, ⌜head_step e1 σ1 κ e2 σ2 efs⌠-∗ £ 1 ={E}=∗ ⌜efs = []⌠∗ state_interp σ2 (S ns) κs nt ∗ from_option Φ False (to_val e2)) ⊢ WP e1 @ s; E {{ Φ }}. Proof. @@ -131,7 +131,7 @@ Lemma wp_lift_pure_det_head_step_no_fork {s E E' Φ} e1 e2 : (∀ σ1, head_reducible e1 σ1) → (∀ σ1 κ e2' σ2 efs', head_step e1 σ1 κ e2' σ2 efs' → κ = [] ∧ σ2 = σ1 ∧ e2' = e2 ∧ efs' = []) → - (|={E}[E']▷=> £1 -∗ WP e2 @ s; E {{ Φ }}) ⊢ WP e1 @ s; E {{ Φ }}. + (|={E}[E']▷=> £ 1 -∗ WP e2 @ s; E {{ Φ }}) ⊢ WP e1 @ s; E {{ Φ }}. Proof using Hinh. intros. rewrite -(wp_lift_pure_det_step_no_fork e1 e2); eauto. destruct s; by auto. @@ -142,7 +142,7 @@ Lemma wp_lift_pure_det_head_step_no_fork' {s E Φ} e1 e2 : (∀ σ1, head_reducible e1 σ1) → (∀ σ1 κ e2' σ2 efs', head_step e1 σ1 κ e2' σ2 efs' → κ = [] ∧ σ2 = σ1 ∧ e2' = e2 ∧ efs' = []) → - ▷ (£1 -∗ WP e2 @ s; E {{ Φ }}) ⊢ WP e1 @ s; E {{ Φ }}. + ▷ (£ 1 -∗ WP e2 @ s; E {{ Φ }}) ⊢ WP e1 @ s; E {{ Φ }}. Proof using Hinh. intros. rewrite -[(WP e1 @ s; _ {{ _ }})%I]wp_lift_pure_det_head_step_no_fork //. rewrite -step_fupd_intro //. diff --git a/iris/program_logic/lifting.v b/iris/program_logic/lifting.v index 85a7d3662e0d8b0f7d374a97fead374de46ef955..f90f71e20ed34e5651572e4468e8933cc922bf9e 100644 --- a/iris/program_logic/lifting.v +++ b/iris/program_logic/lifting.v @@ -33,7 +33,7 @@ Lemma wp_lift_step_fupd s E Φ e1 : to_val e1 = None → (∀ σ1 ns κ κs nt, state_interp σ1 ns (κ ++ κs) nt ={E,∅}=∗ ⌜if s is NotStuck then reducible e1 σ1 else True⌠∗ - ∀ e2 σ2 efs, ⌜prim_step e1 σ1 κ e2 σ2 efs⌠-∗ £1 ={∅}=∗ ▷ |={∅,E}=> + ∀ e2 σ2 efs, ⌜prim_step e1 σ1 κ e2 σ2 efs⌠-∗ £ 1 ={∅}=∗ ▷ |={∅,E}=> state_interp σ2 (S ns) κs (length efs + nt) ∗ WP e2 @ s; E {{ Φ }} ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ fork_post }}) @@ -76,7 +76,7 @@ Qed. Lemma wp_lift_pure_step_no_fork `{!Inhabited (state Λ)} s E E' Φ e1 : (∀ σ1, if s is NotStuck then reducible e1 σ1 else to_val e1 = None) → (∀ κ σ1 e2 σ2 efs, prim_step e1 σ1 κ e2 σ2 efs → κ = [] ∧ σ2 = σ1 ∧ efs = []) → - (|={E}[E']▷=> ∀ κ e2 efs σ, ⌜prim_step e1 σ κ e2 σ efs⌠-∗ £1 -∗ WP e2 @ s; E {{ Φ }}) + (|={E}[E']▷=> ∀ κ e2 efs σ, ⌜prim_step e1 σ κ e2 σ efs⌠-∗ £ 1 -∗ WP e2 @ s; E {{ Φ }}) ⊢ WP e1 @ s; E {{ Φ }}. Proof. iIntros (Hsafe Hstep) "H". iApply wp_lift_step. @@ -107,7 +107,7 @@ Lemma wp_lift_atomic_step_fupd {s E1 E2 Φ} e1 : to_val e1 = None → (∀ σ1 ns κ κs nt, state_interp σ1 ns (κ ++ κs) nt ={E1}=∗ ⌜if s is NotStuck then reducible e1 σ1 else True⌠∗ - ∀ e2 σ2 efs, ⌜prim_step e1 σ1 κ e2 σ2 efs⌠-∗ £1 ={E1}[E2]▷=∗ + ∀ e2 σ2 efs, ⌜prim_step e1 σ1 κ e2 σ2 efs⌠-∗ £ 1 ={E1}[E2]▷=∗ state_interp σ2 (S ns) κs (length efs + nt) ∗ from_option Φ False (to_val e2) ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ fork_post }}) @@ -129,7 +129,7 @@ Lemma wp_lift_atomic_step {s E Φ} e1 : to_val e1 = None → (∀ σ1 ns κ κs nt, state_interp σ1 ns (κ ++ κs) nt ={E}=∗ ⌜if s is NotStuck then reducible e1 σ1 else True⌠∗ - ▷ ∀ e2 σ2 efs, ⌜prim_step e1 σ1 κ e2 σ2 efs⌠-∗ £1 ={E}=∗ + ▷ ∀ e2 σ2 efs, ⌜prim_step e1 σ1 κ e2 σ2 efs⌠-∗ £ 1 ={E}=∗ state_interp σ2 (S ns) κs (length efs + nt) ∗ from_option Φ False (to_val e2) ∗ [∗ list] ef ∈ efs, WP ef @ s; ⊤ {{ fork_post }}) @@ -145,7 +145,7 @@ Lemma wp_lift_pure_det_step_no_fork `{!Inhabited (state Λ)} {s E E' Φ} e1 e2 : (∀ σ1, if s is NotStuck then reducible e1 σ1 else to_val e1 = None) → (∀ σ1 κ e2' σ2 efs', prim_step e1 σ1 κ e2' σ2 efs' → κ = [] ∧ σ2 = σ1 ∧ e2' = e2 ∧ efs' = []) → - (|={E}[E']▷=> £1 -∗ WP e2 @ s; E {{ Φ }}) ⊢ WP e1 @ s; E {{ Φ }}. + (|={E}[E']▷=> £ 1 -∗ WP e2 @ s; E {{ Φ }}) ⊢ WP e1 @ s; E {{ Φ }}. Proof. iIntros (? Hpuredet) "H". iApply (wp_lift_pure_step_no_fork s E E'); try done. { naive_solver. } @@ -156,7 +156,7 @@ Qed. Lemma wp_pure_step_fupd `{!Inhabited (state Λ)} s E E' e1 e2 φ n Φ : PureExec φ n e1 e2 → φ → - (|={E}[E']▷=>^n £n -∗ WP e2 @ s; E {{ Φ }}) ⊢ WP e1 @ s; E {{ Φ }}. + (|={E}[E']▷=>^n £ n -∗ WP e2 @ s; E {{ Φ }}) ⊢ WP e1 @ s; E {{ Φ }}. Proof. iIntros (Hexec Hφ) "Hwp". specialize (Hexec Hφ). iInduction Hexec as [e|n e1 e2 e3 [Hsafe ?]] "IH"; simpl. @@ -174,7 +174,7 @@ Qed. Lemma wp_pure_step_later `{!Inhabited (state Λ)} s E e1 e2 φ n Φ : PureExec φ n e1 e2 → φ → - ▷^n (£n -∗ WP e2 @ s; E {{ Φ }}) ⊢ WP e1 @ s; E {{ Φ }}. + ▷^n (£ n -∗ WP e2 @ s; E {{ Φ }}) ⊢ WP e1 @ s; E {{ Φ }}. Proof. intros Hexec ?. rewrite -wp_pure_step_fupd //. clear Hexec. enough (∀ P, ▷^n P -∗ |={E}▷=>^n P) as Hwp by apply Hwp. iIntros (?).