header {* \section{The Proof System} *}
theory OG_Hoare imports OG_Tran begin
consts assertions :: "'a ann_com => ('a assn) set"
primrec
"assertions (AnnBasic r f) = {r}"
"assertions (AnnSeq c1 c2) = assertions c1 ∪ assertions c2"
"assertions (AnnCond1 r b c1 c2) = {r} ∪ assertions c1 ∪ assertions c2"
"assertions (AnnCond2 r b c) = {r} ∪ assertions c"
"assertions (AnnWhile r b i c) = {r, i} ∪ assertions c"
"assertions (AnnAwait r b c) = {r}"
consts atomics :: "'a ann_com => ('a assn × 'a com) set"
primrec
"atomics (AnnBasic r f) = {(r, Basic f)}"
"atomics (AnnSeq c1 c2) = atomics c1 ∪ atomics c2"
"atomics (AnnCond1 r b c1 c2) = atomics c1 ∪ atomics c2"
"atomics (AnnCond2 r b c) = atomics c"
"atomics (AnnWhile r b i c) = atomics c"
"atomics (AnnAwait r b c) = {(r ∩ b, c)}"
consts com :: "'a ann_triple_op => 'a ann_com_op"
primrec "com (c, q) = c"
consts post :: "'a ann_triple_op => 'a assn"
primrec "post (c, q) = q"
constdefs interfree_aux :: "('a ann_com_op × 'a assn × 'a ann_com_op) => bool"
"interfree_aux ≡ λ(co, q, co'). co'= None ∨
(∀(r,a) ∈ atomics (the co'). \<parallel>= (q ∩ r) a q ∧
(co = None ∨ (∀p ∈ assertions (the co). \<parallel>= (p ∩ r) a p)))"
constdefs interfree :: "(('a ann_triple_op) list) => bool"
"interfree Ts ≡ ∀i j. i < length Ts ∧ j < length Ts ∧ i ≠ j -->
interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) "
consts ann_hoare :: "('a ann_com × 'a assn) set"
syntax "_ann_hoare" :: "'a ann_com => 'a assn => bool" ("(2\<turnstile> _// _)" [60,90] 45)
translations "\<turnstile> c q" \<rightleftharpoons> "(c, q) ∈ ann_hoare"
consts oghoare :: "('a assn × 'a com × 'a assn) set"
syntax "_oghoare" :: "'a assn => 'a com => 'a assn => bool" ("(3\<parallel>- _//_//_)" [90,55,90] 50)
translations "\<parallel>- p c q" \<rightleftharpoons> "(p, c, q) ∈ oghoare"
inductive oghoare ann_hoare
intros
AnnBasic: "r ⊆ {s. f s ∈ q} ==> \<turnstile> (AnnBasic r f) q"
AnnSeq: "[| \<turnstile> c0 pre c1; \<turnstile> c1 q |] ==> \<turnstile> (AnnSeq c0 c1) q"
AnnCond1: "[| r ∩ b ⊆ pre c1; \<turnstile> c1 q; r ∩ -b ⊆ pre c2; \<turnstile> c2 q|]
==> \<turnstile> (AnnCond1 r b c1 c2) q"
AnnCond2: "[| r ∩ b ⊆ pre c; \<turnstile> c q; r ∩ -b ⊆ q |] ==> \<turnstile> (AnnCond2 r b c) q"
AnnWhile: "[| r ⊆ i; i ∩ b ⊆ pre c; \<turnstile> c i; i ∩ -b ⊆ q |]
==> \<turnstile> (AnnWhile r b i c) q"
AnnAwait: "[| atom_com c; \<parallel>- (r ∩ b) c q |] ==> \<turnstile> (AnnAwait r b c) q"
AnnConseq: "[|\<turnstile> c q; q ⊆ q' |] ==> \<turnstile> c q'"
Parallel: "[| ∀i<length Ts. ∃c q. Ts!i = (Some c, q) ∧ \<turnstile> c q; interfree Ts |]
==> \<parallel>- (\<Inter>i∈{i. i<length Ts}. pre(the(com(Ts!i))))
Parallel Ts
(\<Inter>i∈{i. i<length Ts}. post(Ts!i))"
Basic: "\<parallel>- {s. f s ∈q} (Basic f) q"
Seq: "[| \<parallel>- p c1 r; \<parallel>- r c2 q |] ==> \<parallel>- p (Seq c1 c2) q "
Cond: "[| \<parallel>- (p ∩ b) c1 q; \<parallel>- (p ∩ -b) c2 q |] ==> \<parallel>- p (Cond b c1 c2) q"
While: "[| \<parallel>- (p ∩ b) c p |] ==> \<parallel>- p (While b i c) (p ∩ -b)"
Conseq: "[| p' ⊆ p; \<parallel>- p c q ; q ⊆ q' |] ==> \<parallel>- p' c q'"
section {* Soundness *}
(* In the version Isabelle-10-Sep-1999: HOL: The THEN and ELSE
parts of conditional expressions (if P then x else y) are no longer
simplified. (This allows the simplifier to unfold recursive
functional programs.) To restore the old behaviour, we declare
@{text "lemmas [cong del] = if_weak_cong"}. *)
lemmas [cong del] = if_weak_cong
lemmas ann_hoare_induct = oghoare_ann_hoare.induct [THEN conjunct2]
lemmas oghoare_induct = oghoare_ann_hoare.induct [THEN conjunct1]
lemmas AnnBasic = oghoare_ann_hoare.AnnBasic
lemmas AnnSeq = oghoare_ann_hoare.AnnSeq
lemmas AnnCond1 = oghoare_ann_hoare.AnnCond1
lemmas AnnCond2 = oghoare_ann_hoare.AnnCond2
lemmas AnnWhile = oghoare_ann_hoare.AnnWhile
lemmas AnnAwait = oghoare_ann_hoare.AnnAwait
lemmas AnnConseq = oghoare_ann_hoare.AnnConseq
lemmas Parallel = oghoare_ann_hoare.Parallel
lemmas Basic = oghoare_ann_hoare.Basic
lemmas Seq = oghoare_ann_hoare.Seq
lemmas Cond = oghoare_ann_hoare.Cond
lemmas While = oghoare_ann_hoare.While
lemmas Conseq = oghoare_ann_hoare.Conseq
subsection {* Soundness of the System for Atomic Programs *}
lemma Basic_ntran [rule_format]:
"(Basic f, s) -Pn-> (Parallel Ts, t) --> All_None Ts --> t = f s"
apply(induct "n")
apply(simp (no_asm))
apply(fast dest: rel_pow_Suc_D2 Parallel_empty_lemma elim: transition_cases)
done
lemma SEM_fwhile: "SEM S (p ∩ b) ⊆ p ==> SEM (fwhile b S k) p ⊆ (p ∩ -b)"
apply (induct "k")
apply(simp (no_asm) add: L3_5v_lemma3)
apply(simp (no_asm) add: L3_5iv L3_5ii Parallel_empty)
apply(rule conjI)
apply (blast dest: L3_5i)
apply(simp add: SEM_def sem_def id_def)
apply (blast dest: Basic_ntran rtrancl_imp_UN_rel_pow)
done
lemma atom_hoare_sound [rule_format]:
" \<parallel>- p c q --> atom_com(c) --> \<parallel>= p c q"
apply (unfold com_validity_def)
apply(rule oghoare_induct)
apply simp_all
--{*Basic*}
apply(simp add: SEM_def sem_def)
apply(fast dest: rtrancl_imp_UN_rel_pow Basic_ntran)
--{* Seq *}
apply(rule impI)
apply(rule subset_trans)
prefer 2 apply simp
apply(simp add: L3_5ii L3_5i)
--{* Cond *}
apply(simp add: L3_5iv)
--{* While *}
apply (force simp add: L3_5v dest: SEM_fwhile)
--{* Conseq *}
apply(force simp add: SEM_def sem_def)
done
subsection {* Soundness of the System for Component Programs *}
inductive_cases ann_transition_cases:
"(None,s) -1-> t"
"(Some (AnnBasic r f),s) -1-> t"
"(Some (AnnSeq c1 c2), s) -1-> t"
"(Some (AnnCond1 r b c1 c2), s) -1-> t"
"(Some (AnnCond2 r b c), s) -1-> t"
"(Some (AnnWhile r b I c), s) -1-> t"
"(Some (AnnAwait r b c),s) -1-> t"
text {* Strong Soundness for Component Programs:*}
lemma ann_hoare_case_analysis [rule_format]: "\<turnstile> C q' -->
((∀r f. C = AnnBasic r f --> (∃q. r ⊆ {s. f s ∈ q} ∧ q ⊆ q')) ∧
(∀c0 c1. C = AnnSeq c0 c1 --> (∃q. q ⊆ q' ∧ \<turnstile> c0 pre c1 ∧ \<turnstile> c1 q)) ∧
(∀r b c1 c2. C = AnnCond1 r b c1 c2 --> (∃q. q ⊆ q' ∧
r ∩ b ⊆ pre c1 ∧ \<turnstile> c1 q ∧ r ∩ -b ⊆ pre c2 ∧ \<turnstile> c2 q)) ∧
(∀r b c. C = AnnCond2 r b c -->
(∃q. q ⊆ q' ∧ r ∩ b ⊆ pre c ∧ \<turnstile> c q ∧ r ∩ -b ⊆ q)) ∧
(∀r i b c. C = AnnWhile r b i c -->
(∃q. q ⊆ q' ∧ r ⊆ i ∧ i ∩ b ⊆ pre c ∧ \<turnstile> c i ∧ i ∩ -b ⊆ q)) ∧
(∀r b c. C = AnnAwait r b c --> (∃q. q ⊆ q' ∧ \<parallel>- (r ∩ b) c q)))"
apply(rule ann_hoare_induct)
apply simp_all
apply(rule_tac x=q in exI,simp)+
apply(rule conjI,clarify,simp,clarify,rule_tac x=qa in exI,fast)+
apply(clarify,simp,clarify,rule_tac x=qa in exI,fast)
done
lemma Help: "(transition ∩ {(v,v,u). True}) = (transition)"
apply force
done
lemma Strong_Soundness_aux_aux [rule_format]:
"(co, s) -1-> (co', t) --> (∀c. co = Some c --> s∈ pre c -->
(∀q. \<turnstile> c q --> (if co' = None then t∈q else t ∈ pre(the co') ∧ \<turnstile> (the co') q )))"
apply(rule ann_transition_transition.induct [THEN conjunct1])
apply simp_all
--{* Basic *}
apply clarify
apply(frule ann_hoare_case_analysis)
apply force
--{* Seq *}
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply(fast intro: AnnConseq)
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply clarify
apply(rule conjI)
apply force
apply(rule AnnSeq,simp)
apply(fast intro: AnnConseq)
--{* Cond1 *}
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply(fast intro: AnnConseq)
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply(fast intro: AnnConseq)
--{* Cond2 *}
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply(fast intro: AnnConseq)
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply(fast intro: AnnConseq)
--{* While *}
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply force
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply auto
apply(rule AnnSeq)
apply simp
apply(rule AnnWhile)
apply simp_all
--{* Await *}
apply(frule ann_hoare_case_analysis,simp)
apply clarify
apply(drule atom_hoare_sound)
apply simp
apply(simp add: com_validity_def SEM_def sem_def)
apply(simp add: Help All_None_def)
apply force
done
lemma Strong_Soundness_aux: "[| (Some c, s) -*-> (co, t); s ∈ pre c; \<turnstile> c q |]
==> if co = None then t ∈ q else t ∈ pre (the co) ∧ \<turnstile> (the co) q"
apply(erule rtrancl_induct2)
apply simp
apply(case_tac "a")
apply(fast elim: ann_transition_cases)
apply(erule Strong_Soundness_aux_aux)
apply simp
apply simp_all
done
lemma Strong_Soundness: "[| (Some c, s)-*->(co, t); s ∈ pre c; \<turnstile> c q |]
==> if co = None then t∈q else t ∈ pre (the co)"
apply(force dest:Strong_Soundness_aux)
done
lemma ann_hoare_sound: "\<turnstile> c q ==> \<Turnstile> c q"
apply (unfold ann_com_validity_def ann_SEM_def ann_sem_def)
apply clarify
apply(drule Strong_Soundness)
apply simp_all
done
subsection {* Soundness of the System for Parallel Programs *}
lemma Parallel_length_post_P1: "(Parallel Ts,s) -P1-> (R', t) ==>
(∃Rs. R' = (Parallel Rs) ∧ (length Rs) = (length Ts) ∧
(∀i. i<length Ts --> post(Rs ! i) = post(Ts ! i)))"
apply(erule transition_cases)
apply simp
apply clarify
apply(case_tac "i=ia")
apply simp+
done
lemma Parallel_length_post_PStar: "(Parallel Ts,s) -P*-> (R',t) ==>
(∃Rs. R' = (Parallel Rs) ∧ (length Rs) = (length Ts) ∧
(∀i. i<length Ts --> post(Ts ! i) = post(Rs ! i)))"
apply(erule rtrancl_induct2)
apply(simp_all)
apply clarify
apply simp
apply(drule Parallel_length_post_P1)
apply auto
done
lemma assertions_lemma: "pre c ∈ assertions c"
apply(rule ann_com_com.induct [THEN conjunct1])
apply auto
done
lemma interfree_aux1 [rule_format]:
"(c,s) -1-> (r,t) --> (interfree_aux(c1, q1, c) --> interfree_aux(c1, q1, r))"
apply (rule ann_transition_transition.induct [THEN conjunct1])
apply(safe)
prefer 13
apply (rule TrueI)
apply (simp_all add:interfree_aux_def)
apply force+
done
lemma interfree_aux2 [rule_format]:
"(c,s) -1-> (r,t) --> (interfree_aux(c, q, a) --> interfree_aux(r, q, a) )"
apply (rule ann_transition_transition.induct [THEN conjunct1])
apply(force simp add:interfree_aux_def)+
done
lemma interfree_lemma: "[| (Some c, s) -1-> (r, t);interfree Ts ; i<length Ts;
Ts!i = (Some c, q) |] ==> interfree (Ts[i:= (r, q)])"
apply(simp add: interfree_def)
apply clarify
apply(case_tac "i=j")
apply(drule_tac t = "ia" in not_sym)
apply simp_all
apply(force elim: interfree_aux1)
apply(force elim: interfree_aux2 simp add:nth_list_update)
done
text {* Strong Soundness Theorem for Parallel Programs:*}
lemma Parallel_Strong_Soundness_Seq_aux:
"[|interfree Ts; i<length Ts; com(Ts ! i) = Some(AnnSeq c0 c1) |]
==> interfree (Ts[i:=(Some c0, pre c1)])"
apply(simp add: interfree_def)
apply clarify
apply(case_tac "i=j")
apply(force simp add: nth_list_update interfree_aux_def)
apply(case_tac "i=ia")
apply(erule_tac x=ia in allE)
apply(force simp add:interfree_aux_def assertions_lemma)
apply simp
done
lemma Parallel_Strong_Soundness_Seq [rule_format (no_asm)]:
"[| ∀i<length Ts. (if com(Ts!i) = None then b ∈ post(Ts!i)
else b ∈ pre(the(com(Ts!i))) ∧ \<turnstile> the(com(Ts!i)) post(Ts!i));
com(Ts ! i) = Some(AnnSeq c0 c1); i<length Ts; interfree Ts |] ==>
(∀ia<length Ts. (if com(Ts[i:=(Some c0, pre c1)]! ia) = None
then b ∈ post(Ts[i:=(Some c0, pre c1)]! ia)
else b ∈ pre(the(com(Ts[i:=(Some c0, pre c1)]! ia))) ∧
\<turnstile> the(com(Ts[i:=(Some c0, pre c1)]! ia)) post(Ts[i:=(Some c0, pre c1)]! ia)))
∧ interfree (Ts[i:= (Some c0, pre c1)])"
apply(rule conjI)
apply safe
apply(case_tac "i=ia")
apply simp
apply(force dest: ann_hoare_case_analysis)
apply simp
apply(fast elim: Parallel_Strong_Soundness_Seq_aux)
done
lemma Parallel_Strong_Soundness_aux_aux [rule_format]:
"(Some c, b) -1-> (co, t) -->
(∀Ts. i<length Ts --> com(Ts ! i) = Some c -->
(∀i<length Ts. (if com(Ts ! i) = None then b∈post(Ts!i)
else b∈pre(the(com(Ts!i))) ∧ \<turnstile> the(com(Ts!i)) post(Ts!i))) -->
interfree Ts -->
(∀j. j<length Ts ∧ i≠j --> (if com(Ts!j) = None then t∈post(Ts!j)
else t∈pre(the(com(Ts!j))) ∧ \<turnstile> the(com(Ts!j)) post(Ts!j))) )"
apply(rule ann_transition_transition.induct [THEN conjunct1])
apply safe
prefer 11
apply(rule TrueI)
apply simp_all
--{* Basic *}
apply(erule_tac x = "i" in all_dupE, erule (1) notE impE)
apply(erule_tac x = "j" in allE , erule (1) notE impE)
apply(simp add: interfree_def)
apply(erule_tac x = "j" in allE,simp)
apply(erule_tac x = "i" in allE,simp)
apply(drule_tac t = "i" in not_sym)
apply(case_tac "com(Ts ! j)=None")
apply(force intro: converse_rtrancl_into_rtrancl
simp add: interfree_aux_def com_validity_def SEM_def sem_def All_None_def)
apply(simp add:interfree_aux_def)
apply clarify
apply simp
apply(erule_tac x="pre y" in ballE)
apply(force intro: converse_rtrancl_into_rtrancl
simp add: com_validity_def SEM_def sem_def All_None_def)
apply(simp add:assertions_lemma)
--{* Seqs *}
apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE)
apply(drule Parallel_Strong_Soundness_Seq,simp+)
apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE)
apply(drule Parallel_Strong_Soundness_Seq,simp+)
--{* Await *}
apply(rule_tac x = "i" in allE , assumption , erule (1) notE impE)
apply(erule_tac x = "j" in allE , erule (1) notE impE)
apply(simp add: interfree_def)
apply(erule_tac x = "j" in allE,simp)
apply(erule_tac x = "i" in allE,simp)
apply(drule_tac t = "i" in not_sym)
apply(case_tac "com(Ts ! j)=None")
apply(force intro: converse_rtrancl_into_rtrancl simp add: interfree_aux_def
com_validity_def SEM_def sem_def All_None_def Help)
apply(simp add:interfree_aux_def)
apply clarify
apply simp
apply(erule_tac x="pre y" in ballE)
apply(force intro: converse_rtrancl_into_rtrancl
simp add: com_validity_def SEM_def sem_def All_None_def Help)
apply(simp add:assertions_lemma)
done
lemma Parallel_Strong_Soundness_aux [rule_format]:
"[|(Ts',s) -P*-> (Rs',t); Ts' = (Parallel Ts); interfree Ts;
∀i. i<length Ts --> (∃c q. (Ts ! i) = (Some c, q) ∧ s∈(pre c) ∧ \<turnstile> c q ) |] ==>
∀Rs. Rs' = (Parallel Rs) --> (∀j. j<length Rs -->
(if com(Rs ! j) = None then t∈post(Ts ! j)
else t∈pre(the(com(Rs ! j))) ∧ \<turnstile> the(com(Rs ! j)) post(Ts ! j))) ∧ interfree Rs"
apply(erule rtrancl_induct2)
apply clarify
--{* Base *}
apply force
--{* Induction step *}
apply clarify
apply(drule Parallel_length_post_PStar)
apply clarify
apply (ind_cases "(Parallel Ts, s) -P1-> (Parallel Rs, t)")
apply(rule conjI)
apply clarify
apply(case_tac "i=j")
apply(simp split del:split_if)
apply(erule Strong_Soundness_aux_aux,simp+)
apply force
apply force
apply(simp split del: split_if)
apply(erule Parallel_Strong_Soundness_aux_aux)
apply(simp_all add: split del:split_if)
apply force
apply(rule interfree_lemma)
apply simp_all
done
lemma Parallel_Strong_Soundness:
"[|(Parallel Ts, s) -P*-> (Parallel Rs, t); interfree Ts; j<length Rs;
∀i. i<length Ts --> (∃c q. Ts ! i = (Some c, q) ∧ s∈pre c ∧ \<turnstile> c q) |] ==>
if com(Rs ! j) = None then t∈post(Ts ! j) else t∈pre (the(com(Rs ! j)))"
apply(drule Parallel_Strong_Soundness_aux)
apply simp+
done
lemma oghoare_sound [rule_format]: "\<parallel>- p c q --> \<parallel>= p c q"
apply (unfold com_validity_def)
apply(rule oghoare_induct)
apply(rule TrueI)+
--{* Parallel *}
apply(simp add: SEM_def sem_def)
apply clarify
apply(frule Parallel_length_post_PStar)
apply clarify
apply(drule_tac j=i in Parallel_Strong_Soundness)
apply clarify
apply simp
apply force
apply simp
apply(erule_tac V = "∀i. ?P i" in thin_rl)
apply(drule_tac s = "length Rs" in sym)
apply(erule allE, erule impE, assumption)
apply(force dest: nth_mem simp add: All_None_def)
--{* Basic *}
apply(simp add: SEM_def sem_def)
apply(force dest: rtrancl_imp_UN_rel_pow Basic_ntran)
--{* Seq *}
apply(rule subset_trans)
prefer 2 apply assumption
apply(simp add: L3_5ii L3_5i)
--{* Cond *}
apply(simp add: L3_5iv)
--{* While *}
apply(simp add: L3_5v)
apply (blast dest: SEM_fwhile)
--{* Conseq *}
apply(auto simp add: SEM_def sem_def)
done
end
lemmas
b = c ==> (if b then x else y) = (if c then x else y)
lemmas
b = c ==> (if b then x else y) = (if c then x else y)
lemmas ann_hoare_induct:
[| !!f q r. r ⊆ {s. f s ∈ q} ==> P2.1 (AnnBasic r f) q; !!c0 c1 q. [| \<turnstile> c0 pre c1; P2.1 c0 (pre c1); \<turnstile> c1 q; P2.1 c1 q |] ==> P2.1 (AnnSeq c0 c1) q; !!b c1 c2 q r. [| r ∩ b ⊆ pre c1; \<turnstile> c1 q; P2.1 c1 q; r ∩ - b ⊆ pre c2; \<turnstile> c2 q; P2.1 c2 q |] ==> P2.1 (AnnCond1 r b c1 c2) q; !!b c q r. [| r ∩ b ⊆ pre c; \<turnstile> c q; P2.1 c q; r ∩ - b ⊆ q |] ==> P2.1 (AnnCond2 r b c) q; !!b c i q r. [| r ⊆ i; i ∩ b ⊆ pre c; \<turnstile> c i; P2.1 c i; i ∩ - b ⊆ q |] ==> P2.1 (AnnWhile r b i c) q; !!b c q r. [| atom_com c; \<parallel>- (r ∩ b) c q; P1.1 (r ∩ b) c q |] ==> P2.1 (AnnAwait r b c) q; !!c q q'. [| \<turnstile> c q; P2.1 c q; q ⊆ q' |] ==> P2.1 c q'; !!Ts. [| ∀i<length Ts. ∃c q. Ts ! i = (Some c, q) ∧ \<turnstile> c q ∧ P2.1 c q; interfree Ts |] ==> P1.1 (INT i:{i. i < length Ts}. pre (the (OG_Hoare.com (Ts ! i)))) (Parallel Ts) (INT i:{i. i < length Ts}. OG_Hoare.post (Ts ! i)); !!f q. P1.1 {s. f s ∈ q} (Basic f) q; !!c1 c2 p q r. [| \<parallel>- p c1 r; P1.1 p c1 r; \<parallel>- r c2 q; P1.1 r c2 q |] ==> P1.1 p (Seq c1 c2) q; !!b c1 c2 p q. [| \<parallel>- (p ∩ b) c1 q; P1.1 (p ∩ b) c1 q; \<parallel>- (p ∩ - b) c2 q; P1.1 (p ∩ - b) c2 q |] ==> P1.1 p (Cond b c1 c2) q; !!b c i p. [| \<parallel>- (p ∩ b) c p; P1.1 (p ∩ b) c p |] ==> P1.1 p (While b i c) (p ∩ - b); !!c p p' q q'. [| p' ⊆ p; \<parallel>- p c q; P1.1 p c q; q ⊆ q' |] ==> P1.1 p' c q' |] ==> \<turnstile> xb1 xa1 --> P2.1 xb1 xa1
lemmas ann_hoare_induct:
[| !!f q r. r ⊆ {s. f s ∈ q} ==> P2.1 (AnnBasic r f) q; !!c0 c1 q. [| \<turnstile> c0 pre c1; P2.1 c0 (pre c1); \<turnstile> c1 q; P2.1 c1 q |] ==> P2.1 (AnnSeq c0 c1) q; !!b c1 c2 q r. [| r ∩ b ⊆ pre c1; \<turnstile> c1 q; P2.1 c1 q; r ∩ - b ⊆ pre c2; \<turnstile> c2 q; P2.1 c2 q |] ==> P2.1 (AnnCond1 r b c1 c2) q; !!b c q r. [| r ∩ b ⊆ pre c; \<turnstile> c q; P2.1 c q; r ∩ - b ⊆ q |] ==> P2.1 (AnnCond2 r b c) q; !!b c i q r. [| r ⊆ i; i ∩ b ⊆ pre c; \<turnstile> c i; P2.1 c i; i ∩ - b ⊆ q |] ==> P2.1 (AnnWhile r b i c) q; !!b c q r. [| atom_com c; \<parallel>- (r ∩ b) c q; P1.1 (r ∩ b) c q |] ==> P2.1 (AnnAwait r b c) q; !!c q q'. [| \<turnstile> c q; P2.1 c q; q ⊆ q' |] ==> P2.1 c q'; !!Ts. [| ∀i<length Ts. ∃c q. Ts ! i = (Some c, q) ∧ \<turnstile> c q ∧ P2.1 c q; interfree Ts |] ==> P1.1 (INT i:{i. i < length Ts}. pre (the (OG_Hoare.com (Ts ! i)))) (Parallel Ts) (INT i:{i. i < length Ts}. OG_Hoare.post (Ts ! i)); !!f q. P1.1 {s. f s ∈ q} (Basic f) q; !!c1 c2 p q r. [| \<parallel>- p c1 r; P1.1 p c1 r; \<parallel>- r c2 q; P1.1 r c2 q |] ==> P1.1 p (Seq c1 c2) q; !!b c1 c2 p q. [| \<parallel>- (p ∩ b) c1 q; P1.1 (p ∩ b) c1 q; \<parallel>- (p ∩ - b) c2 q; P1.1 (p ∩ - b) c2 q |] ==> P1.1 p (Cond b c1 c2) q; !!b c i p. [| \<parallel>- (p ∩ b) c p; P1.1 (p ∩ b) c p |] ==> P1.1 p (While b i c) (p ∩ - b); !!c p p' q q'. [| p' ⊆ p; \<parallel>- p c q; P1.1 p c q; q ⊆ q' |] ==> P1.1 p' c q' |] ==> \<turnstile> xb1 xa1 --> P2.1 xb1 xa1
lemmas oghoare_induct:
[| !!f q r. r ⊆ {s. f s ∈ q} ==> P2.1 (AnnBasic r f) q; !!c0 c1 q. [| \<turnstile> c0 pre c1; P2.1 c0 (pre c1); \<turnstile> c1 q; P2.1 c1 q |] ==> P2.1 (AnnSeq c0 c1) q; !!b c1 c2 q r. [| r ∩ b ⊆ pre c1; \<turnstile> c1 q; P2.1 c1 q; r ∩ - b ⊆ pre c2; \<turnstile> c2 q; P2.1 c2 q |] ==> P2.1 (AnnCond1 r b c1 c2) q; !!b c q r. [| r ∩ b ⊆ pre c; \<turnstile> c q; P2.1 c q; r ∩ - b ⊆ q |] ==> P2.1 (AnnCond2 r b c) q; !!b c i q r. [| r ⊆ i; i ∩ b ⊆ pre c; \<turnstile> c i; P2.1 c i; i ∩ - b ⊆ q |] ==> P2.1 (AnnWhile r b i c) q; !!b c q r. [| atom_com c; \<parallel>- (r ∩ b) c q; P1.1 (r ∩ b) c q |] ==> P2.1 (AnnAwait r b c) q; !!c q q'. [| \<turnstile> c q; P2.1 c q; q ⊆ q' |] ==> P2.1 c q'; !!Ts. [| ∀i<length Ts. ∃c q. Ts ! i = (Some c, q) ∧ \<turnstile> c q ∧ P2.1 c q; interfree Ts |] ==> P1.1 (INT i:{i. i < length Ts}. pre (the (OG_Hoare.com (Ts ! i)))) (Parallel Ts) (INT i:{i. i < length Ts}. OG_Hoare.post (Ts ! i)); !!f q. P1.1 {s. f s ∈ q} (Basic f) q; !!c1 c2 p q r. [| \<parallel>- p c1 r; P1.1 p c1 r; \<parallel>- r c2 q; P1.1 r c2 q |] ==> P1.1 p (Seq c1 c2) q; !!b c1 c2 p q. [| \<parallel>- (p ∩ b) c1 q; P1.1 (p ∩ b) c1 q; \<parallel>- (p ∩ - b) c2 q; P1.1 (p ∩ - b) c2 q |] ==> P1.1 p (Cond b c1 c2) q; !!b c i p. [| \<parallel>- (p ∩ b) c p; P1.1 (p ∩ b) c p |] ==> P1.1 p (While b i c) (p ∩ - b); !!c p p' q q'. [| p' ⊆ p; \<parallel>- p c q; P1.1 p c q; q ⊆ q' |] ==> P1.1 p' c q' |] ==> \<parallel>- xe1 xd1 xc1 --> P1.1 xe1 xd1 xc1
lemmas oghoare_induct:
[| !!f q r. r ⊆ {s. f s ∈ q} ==> P2.1 (AnnBasic r f) q; !!c0 c1 q. [| \<turnstile> c0 pre c1; P2.1 c0 (pre c1); \<turnstile> c1 q; P2.1 c1 q |] ==> P2.1 (AnnSeq c0 c1) q; !!b c1 c2 q r. [| r ∩ b ⊆ pre c1; \<turnstile> c1 q; P2.1 c1 q; r ∩ - b ⊆ pre c2; \<turnstile> c2 q; P2.1 c2 q |] ==> P2.1 (AnnCond1 r b c1 c2) q; !!b c q r. [| r ∩ b ⊆ pre c; \<turnstile> c q; P2.1 c q; r ∩ - b ⊆ q |] ==> P2.1 (AnnCond2 r b c) q; !!b c i q r. [| r ⊆ i; i ∩ b ⊆ pre c; \<turnstile> c i; P2.1 c i; i ∩ - b ⊆ q |] ==> P2.1 (AnnWhile r b i c) q; !!b c q r. [| atom_com c; \<parallel>- (r ∩ b) c q; P1.1 (r ∩ b) c q |] ==> P2.1 (AnnAwait r b c) q; !!c q q'. [| \<turnstile> c q; P2.1 c q; q ⊆ q' |] ==> P2.1 c q'; !!Ts. [| ∀i<length Ts. ∃c q. Ts ! i = (Some c, q) ∧ \<turnstile> c q ∧ P2.1 c q; interfree Ts |] ==> P1.1 (INT i:{i. i < length Ts}. pre (the (OG_Hoare.com (Ts ! i)))) (Parallel Ts) (INT i:{i. i < length Ts}. OG_Hoare.post (Ts ! i)); !!f q. P1.1 {s. f s ∈ q} (Basic f) q; !!c1 c2 p q r. [| \<parallel>- p c1 r; P1.1 p c1 r; \<parallel>- r c2 q; P1.1 r c2 q |] ==> P1.1 p (Seq c1 c2) q; !!b c1 c2 p q. [| \<parallel>- (p ∩ b) c1 q; P1.1 (p ∩ b) c1 q; \<parallel>- (p ∩ - b) c2 q; P1.1 (p ∩ - b) c2 q |] ==> P1.1 p (Cond b c1 c2) q; !!b c i p. [| \<parallel>- (p ∩ b) c p; P1.1 (p ∩ b) c p |] ==> P1.1 p (While b i c) (p ∩ - b); !!c p p' q q'. [| p' ⊆ p; \<parallel>- p c q; P1.1 p c q; q ⊆ q' |] ==> P1.1 p' c q' |] ==> \<parallel>- xe1 xd1 xc1 --> P1.1 xe1 xd1 xc1
lemmas AnnBasic:
r ⊆ {s. f s ∈ q} ==> \<turnstile> AnnBasic r f q
lemmas AnnBasic:
r ⊆ {s. f s ∈ q} ==> \<turnstile> AnnBasic r f q
lemmas AnnSeq:
[| \<turnstile> c0.0 pre c1.0; \<turnstile> c1.0 q |] ==> \<turnstile> AnnSeq c0.0 c1.0 q
lemmas AnnSeq:
[| \<turnstile> c0.0 pre c1.0; \<turnstile> c1.0 q |] ==> \<turnstile> AnnSeq c0.0 c1.0 q
lemmas AnnCond1:
[| r ∩ b ⊆ pre c1.0; \<turnstile> c1.0 q; r ∩ - b ⊆ pre c2.0; \<turnstile> c2.0 q |] ==> \<turnstile> AnnCond1 r b c1.0 c2.0 q
lemmas AnnCond1:
[| r ∩ b ⊆ pre c1.0; \<turnstile> c1.0 q; r ∩ - b ⊆ pre c2.0; \<turnstile> c2.0 q |] ==> \<turnstile> AnnCond1 r b c1.0 c2.0 q
lemmas AnnCond2:
[| r ∩ b ⊆ pre c; \<turnstile> c q; r ∩ - b ⊆ q |] ==> \<turnstile> AnnCond2 r b c q
lemmas AnnCond2:
[| r ∩ b ⊆ pre c; \<turnstile> c q; r ∩ - b ⊆ q |] ==> \<turnstile> AnnCond2 r b c q
lemmas AnnWhile:
[| r ⊆ i; i ∩ b ⊆ pre c; \<turnstile> c i; i ∩ - b ⊆ q |] ==> \<turnstile> AnnWhile r b i c q
lemmas AnnWhile:
[| r ⊆ i; i ∩ b ⊆ pre c; \<turnstile> c i; i ∩ - b ⊆ q |] ==> \<turnstile> AnnWhile r b i c q
lemmas AnnAwait:
[| atom_com c; \<parallel>- (r ∩ b) c q |] ==> \<turnstile> AnnAwait r b c q
lemmas AnnAwait:
[| atom_com c; \<parallel>- (r ∩ b) c q |] ==> \<turnstile> AnnAwait r b c q
lemmas AnnConseq:
[| \<turnstile> c q; q ⊆ q' |] ==> \<turnstile> c q'
lemmas AnnConseq:
[| \<turnstile> c q; q ⊆ q' |] ==> \<turnstile> c q'
lemmas Parallel:
[| ∀i<length Ts. ∃c q. Ts ! i = (Some c, q) ∧ \<turnstile> c q; interfree Ts |] ==> \<parallel>- (INT i:{i. i < length Ts}. pre (the (OG_Hoare.com (Ts ! i)))) Parallel Ts (INT i:{i. i < length Ts}. OG_Hoare.post (Ts ! i))
lemmas Parallel:
[| ∀i<length Ts. ∃c q. Ts ! i = (Some c, q) ∧ \<turnstile> c q; interfree Ts |] ==> \<parallel>- (INT i:{i. i < length Ts}. pre (the (OG_Hoare.com (Ts ! i)))) Parallel Ts (INT i:{i. i < length Ts}. OG_Hoare.post (Ts ! i))
lemmas Basic:
\<parallel>- {s. f s ∈ q}
Basic f
q
lemmas Basic:
\<parallel>- {s. f s ∈ q}
Basic f
q
lemmas Seq:
[| \<parallel>- p c1.0 r; \<parallel>- r c2.0 q |] ==> \<parallel>- p Seq c1.0 c2.0 q
lemmas Seq:
[| \<parallel>- p c1.0 r; \<parallel>- r c2.0 q |] ==> \<parallel>- p Seq c1.0 c2.0 q
lemmas Cond:
[| \<parallel>- (p ∩ b) c1.0 q; \<parallel>- (p ∩ - b) c2.0 q |] ==> \<parallel>- p Cond b c1.0 c2.0 q
lemmas Cond:
[| \<parallel>- (p ∩ b) c1.0 q; \<parallel>- (p ∩ - b) c2.0 q |] ==> \<parallel>- p Cond b c1.0 c2.0 q
lemmas While:
\<parallel>- (p ∩ b) c p ==> \<parallel>- p While b i c (p ∩ - b)
lemmas While:
\<parallel>- (p ∩ b) c p ==> \<parallel>- p While b i c (p ∩ - b)
lemmas Conseq:
[| p' ⊆ p; \<parallel>- p c q; q ⊆ q' |] ==> \<parallel>- p' c q'
lemmas Conseq:
[| p' ⊆ p; \<parallel>- p c q; q ⊆ q' |] ==> \<parallel>- p' c q'
lemma Basic_ntran:
[| ((Basic f, s), Parallel Ts, t) ∈ transition ^ n; All_None Ts |] ==> t = f s
lemma SEM_fwhile:
SEM S (p ∩ b) ⊆ p ==> SEM (fwhile b S k) p ⊆ p ∩ - b
lemma atom_hoare_sound:
[| \<parallel>- p c q; atom_com c |] ==> \<parallel>= p c q
lemmas ann_transition_cases:
(None, s) -1-> t ==> P
[| (Some (AnnBasic r f), s) -1-> t; t = (None, f s) ==> P |] ==> P
[| (Some (AnnSeq c1.0 c2.0), s) -1-> t; !!t. [| ((Some c1.0, s), None, t) ∈ ann_transition; t = (Some c2.0, t) |] ==> P; !!c2 t. [| ((Some c1.0, s), Some c2, t) ∈ ann_transition; t = (Some (AnnSeq c2 c2.0), t) |] ==> P |] ==> P
[| (Some (AnnCond1 r b c1.0 c2.0), s) -1-> t; [| s ∈ b; t = (Some c1.0, s) |] ==> P; [| s ∉ b; t = (Some c2.0, s) |] ==> P |] ==> P
[| (Some (AnnCond2 r b c), s) -1-> t; [| s ∈ b; t = (Some c, s) |] ==> P; [| s ∉ b; t = (None, s) |] ==> P |] ==> P
[| (Some (AnnWhile r b I c), s) -1-> t; [| s ∉ b; t = (None, s) |] ==> P; [| s ∈ b; t = (Some (AnnSeq c (AnnWhile I b I c)), s) |] ==> P |] ==> P
[| (Some (AnnAwait r b c), s) -1-> t; !!t. [| s ∈ b; atom_com c; ((c, s), Parallel [], t) ∈ transition*; t = (None, t) |] ==> P |] ==> P
lemma ann_hoare_case_analysis:
\<turnstile> C q' ==> (∀r f. C = AnnBasic r f --> (∃q. r ⊆ {s. f s ∈ q} ∧ q ⊆ q')) ∧ (∀c0 c1. C = AnnSeq c0 c1 --> (∃q⊆q'. \<turnstile> c0 pre c1 ∧ \<turnstile> c1 q)) ∧ (∀r b c1 c2. C = AnnCond1 r b c1 c2 --> (∃q⊆q'. r ∩ b ⊆ pre c1 ∧ \<turnstile> c1 q ∧ r ∩ - b ⊆ pre c2 ∧ \<turnstile> c2 q)) ∧ (∀r b c. C = AnnCond2 r b c --> (∃q⊆q'. r ∩ b ⊆ pre c ∧ \<turnstile> c q ∧ r ∩ - b ⊆ q)) ∧ (∀r i b c. C = AnnWhile r b i c --> (∃q⊆q'. r ⊆ i ∧ i ∩ b ⊆ pre c ∧ \<turnstile> c i ∧ i ∩ - b ⊆ q)) ∧ (∀r b c. C = AnnAwait r b c --> (∃q⊆q'. \<parallel>- (r ∩ b) c q))
lemma Help:
transition ∩ {(v, v, u). True} = transition
lemma Strong_Soundness_aux_aux:
[| ((co, s), co', t) ∈ ann_transition; co = Some c; s ∈ pre c; \<turnstile> c q |] ==> if co' = None then t ∈ q else t ∈ pre (the co') ∧ \<turnstile> the co' q
lemma Strong_Soundness_aux:
[| ((Some c, s), co, t) ∈ ann_transition*; s ∈ pre c; \<turnstile> c q |] ==> if co = None then t ∈ q else t ∈ pre (the co) ∧ \<turnstile> the co q
lemma Strong_Soundness:
[| ((Some c, s), co, t) ∈ ann_transition*; s ∈ pre c; \<turnstile> c q |] ==> if co = None then t ∈ q else t ∈ pre (the co)
lemma ann_hoare_sound:
\<turnstile> c q ==> \<Turnstile> c q
lemma Parallel_length_post_P1:
((Parallel Ts, s), R', t) ∈ transition ==> ∃Rs. R' = Parallel Rs ∧ length Rs = length Ts ∧ (∀i<length Ts. OG_Hoare.post (Rs ! i) = OG_Hoare.post (Ts ! i))
lemma Parallel_length_post_PStar:
((Parallel Ts, s), R', t) ∈ transition* ==> ∃Rs. R' = Parallel Rs ∧ length Rs = length Ts ∧ (∀i<length Ts. OG_Hoare.post (Ts ! i) = OG_Hoare.post (Rs ! i))
lemma assertions_lemma:
pre c ∈ assertions c
lemma interfree_aux1:
[| ((c, s), r, t) ∈ ann_transition; interfree_aux (c1.0, q1.0, c) |] ==> interfree_aux (c1.0, q1.0, r)
lemma interfree_aux2:
[| ((c, s), r, t) ∈ ann_transition; interfree_aux (c, q, a) |] ==> interfree_aux (r, q, a)
lemma interfree_lemma:
[| ((Some c, s), r, t) ∈ ann_transition; interfree Ts; i < length Ts; Ts ! i = (Some c, q) |] ==> interfree (Ts[i := (r, q)])
lemma Parallel_Strong_Soundness_Seq_aux:
[| interfree Ts; i < length Ts; OG_Hoare.com (Ts ! i) = Some (AnnSeq c0.0 c1.0) |] ==> interfree (Ts[i := (Some c0.0, pre c1.0)])
lemma Parallel_Strong_Soundness_Seq:
[| ∀i<length Ts. if OG_Hoare.com (Ts ! i) = None then b ∈ OG_Hoare.post (Ts ! i) else b ∈ pre (the (OG_Hoare.com (Ts ! i))) ∧ \<turnstile> the (OG_Hoare.com (Ts ! i)) OG_Hoare.post (Ts ! i); OG_Hoare.com (Ts ! i) = Some (AnnSeq c0.0 c1.0); i < length Ts; interfree Ts |] ==> (∀ia<length Ts. if OG_Hoare.com (Ts[i := (Some c0.0, pre c1.0)] ! ia) = None then b ∈ OG_Hoare.post (Ts[i := (Some c0.0, pre c1.0)] ! ia) else b ∈ pre (the (OG_Hoare.com (Ts[i := (Some c0.0, pre c1.0)] ! ia))) ∧ \<turnstile> the (OG_Hoare.com (Ts[i := (Some c0.0, pre c1.0)] ! ia)) OG_Hoare.post (Ts[i := (Some c0.0, pre c1.0)] ! ia)) ∧ interfree (Ts[i := (Some c0.0, pre c1.0)])
lemma Parallel_Strong_Soundness_aux_aux:
[| ((Some c, b), co, t) ∈ ann_transition; i < length Ts; OG_Hoare.com (Ts ! i) = Some c; !!i. i < length Ts ==> if OG_Hoare.com (Ts ! i) = None then b ∈ OG_Hoare.post (Ts ! i) else b ∈ pre (the (OG_Hoare.com (Ts ! i))) ∧ \<turnstile> the (OG_Hoare.com (Ts ! i)) OG_Hoare.post (Ts ! i); interfree Ts; j < length Ts ∧ i ≠ j |] ==> if OG_Hoare.com (Ts ! j) = None then t ∈ OG_Hoare.post (Ts ! j) else t ∈ pre (the (OG_Hoare.com (Ts ! j))) ∧ \<turnstile> the (OG_Hoare.com (Ts ! j)) OG_Hoare.post (Ts ! j)
lemma Parallel_Strong_Soundness_aux:
[| ((Ts', s), Rs', t) ∈ transition*; Ts' = Parallel Ts; interfree Ts; !!i. i < length Ts ==> ∃c q. Ts ! i = (Some c, q) ∧ s ∈ pre c ∧ \<turnstile> c q; Rs' = Parallel Rs |] ==> (∀j<length Rs. if OG_Hoare.com (Rs ! j) = None then t ∈ OG_Hoare.post (Ts ! j) else t ∈ pre (the (OG_Hoare.com (Rs ! j))) ∧ \<turnstile> the (OG_Hoare.com (Rs ! j)) OG_Hoare.post (Ts ! j)) ∧ interfree Rs
lemma Parallel_Strong_Soundness:
[| ((Parallel Ts, s), Parallel Rs, t) ∈ transition*; interfree Ts; j < length Rs; ∀i<length Ts. ∃c q. Ts ! i = (Some c, q) ∧ s ∈ pre c ∧ \<turnstile> c q |] ==> if OG_Hoare.com (Rs ! j) = None then t ∈ OG_Hoare.post (Ts ! j) else t ∈ pre (the (OG_Hoare.com (Rs ! j)))
lemma oghoare_sound:
\<parallel>- p c q ==> \<parallel>= p c q