Theory Hoare_MonadSE
theory Hoare_MonadSE
imports Symbex_MonadSE
begin
section‹Hoare›
definition hoare⇩3 :: "('σ ⇒ bool) ⇒ ('α, 'σ)MON⇩S⇩E ⇒ ('α ⇒ 'σ ⇒ bool) ⇒ bool" ("(⦃(1_)⦄/ (_)/ ⦃(1_)⦄)" 50)
where "⦃P⦄ M ⦃Q⦄ ≡ (∀σ. P σ ⟶ (case M σ of None => False | Some(x, σ') => Q x σ'))"
definition hoare⇩3' :: "('σ ⇒ bool) ⇒ ('α, 'σ)MON⇩S⇩E ⇒ bool" ("(⦃(1_)⦄/ (_)/†)" 50)
where "⦃P⦄ M † ≡ (∀σ. P σ ⟶ (case M σ of None => True | _ => False))"
subsection‹Basic rules›
lemma skip: " ⦃P⦄ skip⇩S⇩E ⦃λ_. P⦄"
unfolding hoare⇩3_def skip⇩S⇩E_def unit_SE_def
by auto
lemma fail: "⦃P⦄ fail⇩S⇩E †"
unfolding hoare⇩3'_def fail_SE_def unit_SE_def by auto
lemma assert: "⦃P⦄ assert⇩S⇩E P ⦃λ _ _. True⦄"
unfolding hoare⇩3_def assert_SE_def unit_SE_def
by auto
lemma assert_conseq: "Collect P ⊆ Collect Q ⟹ ⦃P⦄ assert⇩S⇩E Q ⦃λ _ _. True⦄"
unfolding hoare⇩3_def assert_SE_def unit_SE_def
by auto
lemma assume_conseq:
assumes "∃ σ. Q σ"
shows "⦃P⦄ assume⇩S⇩E Q ⦃λ _ . Q⦄"
unfolding hoare⇩3_def assume_SE_def unit_SE_def
apply (auto simp : someI2)
using assms by auto
text ‹assignment missing in the calculus because this is viewed as a state specific
operation, definable for concrete instances of @{typ "'σ"}.›
subsection ‹Generalized and special sequence rules›
text‹The decisive idea is to factor out the post-condition on the results of @{term M} :›
lemma sequence :
" ⦃P⦄ M ⦃λx σ. x∈A ∧ Q x σ⦄
⟹ ∀x∈A. ⦃Q x⦄ M' x ⦃R⦄
⟹ ⦃P⦄ x ← M; M' x ⦃R⦄"
unfolding hoare⇩3_def bind_SE_def
by(auto,erule_tac x="σ" in allE, auto split: Option.option.split_asm Option.option.split)
lemma sequence_irpt_l : "⦃P⦄ M † ⟹ ⦃P⦄ x ← M; M' x †"
unfolding hoare⇩3'_def bind_SE_def
by(auto,erule_tac x="σ" in allE, auto split: Option.option.split_asm Option.option.split)
lemma sequence_irpt_r : "⦃P⦄ M ⦃λx σ. x∈A ∧ Q x σ⦄ ⟹ ∀x∈A. ⦃Q x⦄ M' x † ⟹ ⦃P⦄ x ← M; M' x †"
unfolding hoare⇩3'_def hoare⇩3_def bind_SE_def
by(auto,erule_tac x="σ" in allE, auto split: Option.option.split_asm Option.option.split)
lemma sequence' : "⦃P⦄ M ⦃λ_. Q ⦄ ⟹ ⦃Q⦄ M' ⦃R⦄ ⟹ ⦃P⦄ M;- M' ⦃R⦄"
unfolding hoare⇩3_def hoare⇩3_def bind_SE_def bind_SE'_def
by(auto,erule_tac x="σ" in allE, auto split: Option.option.split_asm Option.option.split)
lemma sequence_irpt_l' : "⦃P⦄ M † ⟹ ⦃P⦄ M;- M' †"
unfolding hoare⇩3'_def bind_SE_def bind_SE'_def
by(auto,erule_tac x="σ" in allE, auto split: Option.option.split_asm Option.option.split)
lemma sequence_irpt_r' : "⦃P⦄ M ⦃λ_. Q ⦄ ⟹ ⦃Q⦄ M' † ⟹ ⦃P⦄ M;- M' †"
unfolding hoare⇩3'_def hoare⇩3_def bind_SE_def bind_SE'_def
by(auto,erule_tac x="σ" in allE, auto split: Option.option.split_asm Option.option.split)
subsection‹Generalized and special consequence rules›
lemma consequence :
" Collect P ⊆ Collect P'
⟹ ⦃P'⦄ M ⦃λx σ. x∈A ∧ Q' x σ⦄
⟹ ∀ x∈A. Collect(Q' x) ⊆ Collect (Q x)
⟹ ⦃P⦄ M ⦃λx σ. x∈A ∧ Q x σ⦄"
unfolding hoare⇩3_def bind_SE_def
by(auto,erule_tac x="σ" in allE,auto split: Option.option.split_asm Option.option.split)
lemma consequence_unit :
assumes "(⋀ σ. P σ ⟶ P' σ)"
and "⦃P'⦄ M ⦃λx::unit. λ σ. Q' σ⦄"
and " (⋀ σ. Q' σ ⟶ Q σ)"
shows "⦃P⦄ M ⦃λx σ. Q σ⦄"
proof -
have * : "(λx σ. Q σ) = (λx::unit. λ σ. x∈UNIV ∧ Q σ) " by auto
show ?thesis
apply(subst *)
apply(rule_tac P' = "P'" and Q' = "%_. Q'" in consequence)
apply (simp add: Collect_mono assms(1))
using assms(2) apply auto[1]
by (simp add: Collect_mono assms(3))
qed
lemma consequence_irpt :
" Collect P ⊆ Collect P'
⟹ ⦃P'⦄ M †
⟹ ⦃P⦄ M †"
unfolding hoare⇩3_def hoare⇩3'_def bind_SE_def
by(auto)
lemma consequence_mt_swap :
"(⦃λ_. False⦄ M †) = (⦃λ_. False⦄ M ⦃P⦄)"
unfolding hoare⇩3_def hoare⇩3'_def bind_SE_def
by auto
subsection‹Condition rules›
lemma cond :
" ⦃λσ. P σ ∧ cond σ⦄ M ⦃Q⦄
⟹ ⦃λσ. P σ ∧ ¬ cond σ⦄ M' ⦃Q⦄
⟹ ⦃P⦄if⇩S⇩E cond then M else M' fi⦃Q⦄"
unfolding hoare⇩3_def hoare⇩3'_def bind_SE_def if_SE_def
by auto
lemma cond_irpt :
" ⦃λσ. P σ ∧ cond σ⦄ M †
⟹ ⦃λσ. P σ ∧ ¬ cond σ⦄ M' †
⟹ ⦃P⦄if⇩S⇩E cond then M else M' fi †"
unfolding hoare⇩3_def hoare⇩3'_def bind_SE_def if_SE_def
by auto
text‹ Note that the other four combinations can be directly derived via
the @{thm consequence_mt_swap} rule.›
subsection‹While rules›
text‹The only non-trivial proof is, of course, the while loop rule. Note
that non-terminating loops were mapped to @{term None} following the principle
that our monadic state-transformers represent partial functions in the mathematical
sense.›
lemma while :
assumes * : "⦃λσ. cond σ ∧ P σ⦄ M ⦃λ_. P⦄"
and measure: "∀σ. cond σ ∧ P σ ⟶ M σ ≠ None ∧ f(snd(the(M σ))) < ((f σ)::nat) "
shows "⦃P⦄while⇩S⇩E cond do M od ⦃λ_ σ. ¬cond σ ∧ P σ⦄"
unfolding hoare⇩3_def hoare⇩3'_def bind_SE_def if_SE_def
proof auto
have * : "∀n. ∀ σ. P σ ∧ f σ ≤ n ⟶
(case (while⇩S⇩E cond do M od) σ of
None ⇒ False
| Some (x, σ') ⇒ ¬ cond σ' ∧ P σ')" (is "∀n. ?P n")
proof (rule allI, rename_tac n, induct_tac n)
fix n show "?P 0"
apply(auto)
apply(subst while_SE_unfold)
by (metis (no_types, lifting) gr_implies_not0 if_SE_def measure option.case_eq_if
option.sel option.simps(3) prod.sel(2) split_def unit_SE_def)
next
fix n show " ?P n ⟹ ?P (Suc n)"
apply(auto,subst while_SE_unfold)
apply(case_tac "¬cond σ")
apply (simp add: if_SE_def unit_SE_def)
apply(simp add: if_SE_def)
apply(case_tac "M σ = None")
using measure apply blast
proof (auto simp: bind_SE'_def bind_SE_def)
fix σ σ'
assume 1 : "cond σ"
and 2 : "M σ = Some ((), σ')"
and 3 : " P σ"
and 4 : " f σ ≤ Suc n"
and hyp : "?P n"
have 5 : "P σ'"
by (metis (no_types, lifting) * 1 2 3 case_prodD hoare⇩3_def option.simps(5))
have 6 : "snd(the(M σ)) = σ'"
by (simp add: 2)
have 7 : "cond σ' ⟹ f σ' ≤ n"
using 1 3 4 6 leD measure by auto
show "case (while⇩S⇩E cond do M od) σ' of None ⇒ False
| Some (xa, σ') ⇒ ¬ cond σ' ∧ P σ'"
using 1 3 4 5 6 hyp measure by auto
qed
qed
show "⋀σ. P σ ⟹
case (while⇩S⇩E cond do M od) σ of None ⇒ False
| Some (x, σ') ⇒ ¬ cond σ' ∧ P σ'"
using "*" by blast
qed
lemma while_irpt :
assumes * : "⦃λσ. cond σ ∧ P σ⦄ M ⦃λ_. P⦄ ∨ ⦃λσ. cond σ ∧ P σ⦄ M †"
and measure: "∀σ. cond σ ∧ P σ ⟶ M σ = None ∨ f(snd(the(M σ))) < ((f σ)::nat)"
and enabled: "∀σ. P σ ⟶ cond σ"
shows "⦃P⦄while⇩S⇩E cond do M od †"
unfolding hoare⇩3_def hoare⇩3'_def bind_SE_def if_SE_def
proof auto
have * : "∀n. ∀ σ. P σ ∧ f σ ≤ n ⟶
(case (while⇩S⇩E cond do M od) σ of None ⇒ True | Some a ⇒ False)"
(is "∀n. ?P n ")
proof (rule allI, rename_tac n, induct_tac n)
fix n
have 1 : "⋀σ. P σ ⟹ cond σ"
by (simp add: enabled * )
show "?P 0 "
apply(auto,frule 1)
by (metis assms(2) bind_SE'_def bind_SE_def gr_implies_not0 if_SE_def option.case(1)
option.case_eq_if while_SE_unfold)
next
fix k n
assume hyp : "?P n"
have 1 : "⋀σ. P σ ⟹ cond σ"
by (simp add: enabled * )
show "?P (Suc n) "
apply(auto, frule 1)
apply(subst while_SE_unfold, auto simp: if_SE_def)
proof(insert *,simp_all add: hoare⇩3_def hoare⇩3'_def, erule disjE)
fix σ
assume "P σ"
and "f σ ≤ Suc n"
and "cond σ"
and ** : "∀σ. cond σ ∧ P σ ⟶ (case M σ of None ⇒ False | Some (x, σ') ⇒ P σ')"
obtain "(case M σ of None ⇒ False | Some (x, σ') ⇒ P σ')"
by (simp add: "**" ‹P σ› ‹cond σ›)
then
show "case (M ;- (while⇩S⇩E cond do M od)) σ of None ⇒ True | Some a ⇒ False"
apply(case_tac "M σ", auto, rename_tac σ', simp add: bind_SE'_def bind_SE_def)
proof -
fix σ'
assume "P σ'"
and "M σ = Some ((), σ')"
have "cond σ'" by (simp add: ‹P σ'› enabled)
have "f σ' ≤ n"
using ‹M σ = Some ((), σ')› ‹P σ› ‹cond σ› ‹f σ ≤ Suc n› measure by fastforce
show "case (while⇩S⇩E cond do M od) σ' of None ⇒ True | Some a ⇒ False"
using hyp by (simp add: ‹P σ'› ‹f σ' ≤ n›)
qed
next
fix σ
assume "P σ"
and "f σ ≤ Suc n"
and "cond σ"
and * : "∀σ. cond σ ∧ P σ ⟶ (case M σ of None ⇒ True | Some a ⇒ False)"
obtain ** : "(case M σ of None ⇒ True | Some a ⇒ False)"
by (simp add: "*" ‹P σ› ‹cond σ›)
have "M σ = None"
by (simp add: "**" option.disc_eq_case(1))
show "case (M ;- (while⇩S⇩E cond do M od)) σ of None ⇒ True | Some a ⇒ False"
by (simp add: ‹M σ = None› bind_SE'_def bind_SE_def)
qed
qed
show "⋀σ. P σ ⟹ case (while⇩S⇩E cond do M od) σ of None ⇒ True | Some a ⇒ False" using * by blast
qed
subsection‹Experimental Alternative Definitions (Transformer-Style Rely-Guarantee)›
definition hoare⇩1 :: "('σ ⇒ bool) ⇒ ('α, 'σ)MON⇩S⇩E ⇒ ('α ⇒ 'σ ⇒ bool) ⇒ bool" ("⊢⇩1 ({(1_)}/ (_)/ {(1_)})" 50)
where "(⊢⇩1{P} M {Q} ) = (∀σ. (σ ⊨ (_ ← assume⇩S⇩E P ; x ← M; assert⇩S⇩E (Q x))))"
definition hoare⇩2 :: "('σ ⇒ bool) ⇒ ('α, 'σ)MON⇩S⇩E ⇒ ('α ⇒ 'σ ⇒ bool) ⇒ bool" ("⊢⇩2 ({(1_)}/ (_)/ {(1_)})" 50)
where "(⊢⇩2{P} M {Q} ) = (∀σ. P σ ⟶ (σ ⊨ (x ← M; assert⇩S⇩E (Q x))))"
end