theory LocalRG_HoareDef
imports "SmallStepCon" "EmbSimpl/HoarePartialProps" "HOL-Library.Countable"
begin
section ‹Validity of Correctness Formulas›
subsection ‹Aux›
abbreviation (input)
set_fun :: "'a set ⇒ 'a ⇒ bool" ("_⇩f") where
"set_fun s ≡ λv. v∈s"
abbreviation (input)
fun_set :: "('a ⇒ bool) ⇒ 'a set" ("_⇩s") where
"fun_set f ≡ {σ. f σ}"
lemma tl_pair:"Suc (Suc j) < length l ⟹
l1 = tl l ⟹
P (l!(Suc j)) (l!(Suc (Suc j)))=
P (l1!j) (l1!(Suc j))"
by (simp add: tl_zero_eq)
lemma for_all_k_sublist:
assumes a0:"Suc (Suc j)<length l" and
a1:"(∀k < j. P ((tl l)!k) ((tl l)!(Suc k)))" and
a2:"P (l!0) (l!(Suc 0))"
shows "(∀k < Suc j. P (l!k) (l!(Suc k)))"
proof -
{fix k
assume aa0:"k < Suc j"
have "P (l!k) (l!(Suc k))"
proof (cases k)
case 0 thus ?thesis using a2 by auto
next
case (Suc k1) thus ?thesis using aa0 a0 a1 a2
by (metis SmallStepCon.nth_tl Suc_less_SucD dual_order.strict_trans length_greater_0_conv nth_Cons_Suc zero_less_Suc)
qed
} thus ?thesis by auto
qed
subsection ‹Validity for Component Programs.›
type_synonym ('s,'f) tran = "('s,'f) xstate × ('s,'f) xstate"
type_synonym ('s,'p,'f,'e) rgformula =
"(('s,'p,'f,'e) com ×
('s set) ×
(('s,'f) tran) set ×
(('s,'f) tran) set ×
('s set) ×
('s set))"
type_synonym ('s,'p,'f,'e) sextuple =
"('p ×
('s set) ×
(('s,'f) tran) set ×
(('s,'f) tran) set ×
('s set) ×
('s set))"
definition Sta :: "'s set ⇒ (('s,'f) tran) set ⇒ bool" where
"Sta ≡ λf g. (∀x y x'. x' ∈ f ∧ x=Normal x' ⟶ (x,y)∈ g ⟶ (∃y'. y=Normal y' ∧ y' ∈ f))"
lemma Sta_intro:"Sta a R ⟹ Sta b R ⟹ Sta (a ∩ b) R"
unfolding Sta_def by fastforce
lemma Sta_assoc:"Sta (a ∩ (b ∩ c)) R = Sta ((a∩b)∩c) R"
unfolding Sta_def by fastforce
lemma Sta_comm:"Sta (a ∩ b) R = Sta (b ∩ a) R"
unfolding Sta_def by fastforce
lemma Sta_add:"Sta (a ∩ b) R ⟹ Sta (a ∩ c) R ⟹
Sta (a ∩ b ∩ c) R"
unfolding Sta_def by fastforce
lemma Sta_tran:"Sta a R ⟹ a = b ⟹ Sta b R"
by auto
definition Norm:: "(('s,'f) tran) set ⇒ bool" where
"Norm ≡ λg. (∀x y. (x, y) ∈ g ⟶ (∃x' y'. x=Normal x' ∧ y=Normal y'))"
definition env_tran::
"('p ⇒ ('s, 'p, 'f,'e) LanguageCon.com option)
⇒ ('s set)
⇒ (('s, 'p, 'f,'e) LanguageCon.com × ('s, 'f) xstate) list
⇒ ('s,'f) tran set ⇒ bool"
where
"env_tran Γ q l rely ≡ snd(l!0) ∈ Normal ` q ∧ (∀i. Suc i<length l ⟶
Γ⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ rely)
"
definition env_tran_right::
"('p ⇒ ('s, 'p, 'f,'e) LanguageCon.com option)
⇒ (('s, 'p, 'f,'e) LanguageCon.com × ('s, 'f) xstate) list
⇒ ('s,'f) tran set ⇒ bool"
where
"env_tran_right Γ l rely ≡
(∀i. Suc i<length l ⟶
Γ⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ rely)
"
lemma env_tran_tail:"env_tran_right Γ (x#l) R ⟹ env_tran_right Γ l R"
unfolding env_tran_right_def
by fastforce
lemma env_tran_subr:
assumes a0:"env_tran_right Γ (l1@l2) R"
shows "env_tran_right Γ l1 R"
unfolding env_tran_right_def
proof -
{fix i
assume a1:"Suc i< length l1"
assume a2:"Γ⊢⇩c l1 ! i →⇩e l1 ! Suc i"
then have "Suc i < length (l1@l2)" using a1 by fastforce
also then have "Γ⊢⇩c (l1@l2) ! i →⇩e (l1@l2) ! Suc i"
proof -
show ?thesis
by (simp add: Suc_lessD a1 a2 nth_append)
qed
ultimately have f1:"(snd ((l1@l2)! i), snd ((l1@l2) ! Suc i)) ∈ R"
using a0 unfolding env_tran_right_def by auto
then have "(snd (l1! i), snd (l1 ! Suc i)) ∈ R"
using a1
proof -
have "∀ps psa n. if n < length ps then (ps @ psa) ! n = (ps ! n::('b, 'a, 'c,'d) LanguageCon.com × ('b, 'c) xstate)
else (ps @ psa) ! n = psa ! (n - length ps)"
by (meson nth_append)
then show ?thesis
using f1 ‹Suc i < length l1› by force
qed
} then show
"∀i. Suc i < length l1 ⟶
Γ⊢⇩c l1 ! i →⇩e l1 ! Suc i ⟶
(snd (l1 ! i), snd (l1 ! Suc i)) ∈ R"
by blast
qed
definition Satis where "Satis ≡ True"
definition sep_conj where "sep_conj ≡ True"
lemma env_tran_subl:"env_tran_right Γ (l1@l2) R ⟹ env_tran_right Γ l2 R"
proof (induct "l1")
case Nil thus ?case by auto
next
case (Cons a l1) thus ?case by (fastforce intro:append_Cons env_tran_tail )
qed
lemma env_tran_R_R':"env_tran_right Γ l R ⟹
(R ⊆ R') ⟹
env_tran_right Γ l R'"
unfolding env_tran_right_def Satis_def sep_conj_def
apply clarify
apply (erule allE)
apply auto
done
lemma env_tran_normal:
assumes a0:"env_tran_right Γ l rely ∧ Sta q rely ∧ snd(l!i) = Normal s1 ∧ s1∈q" and
a1:"Suc i < length l ∧ Γ⊢⇩c(l!i) →⇩e (l!(Suc i))"
shows "∃s1 s2. snd(l!i) = Normal s1 ∧ snd(l!(Suc i)) = Normal s2 ∧ s2∈q"
using a0 a1 unfolding env_tran_right_def Sta_def by fastforce
lemma no_env_tran_not_normal:
assumes a0:"env_tran_right Γ l rely ∧ Sta q rely ∧ snd(l!i) = Normal s1 ∧ s1∈q" and
a1:"Suc i < length l ∧ Γ⊢⇩c(l!i) →⇩e (l!(Suc i))" and
a2:"(∀s1. ¬ (snd(l!i) = Normal s1)) ∨ (∀s2. ¬ (snd (l!Suc i) = Normal s2))"
shows "P"
using a0 a1 a2 unfolding env_tran_right_def Sta_def by fastforce
definition assum ::
"('s set × ('s,'f) tran set) ⇒ (('s,'p,'f,'e) confs) set" where
"assum ≡ λ(pre, rely).
{c. snd((snd c)!0) ∈ Normal ` pre ∧
(∀i. Suc i<length (snd c) ⟶
(fst c)⊢⇩c((snd c)!i) →⇩e ((snd c)!(Suc i)) ⟶
(snd((snd c)!i), snd((snd c)!(Suc i))) ∈ rely)}"
definition assum1 ::
"('s set × ('s,'f) tran set) ⇒
'f set ⇒
(('s,'p,'f,'e) confs) set" where
"assum1 ≡ λ(pre, rely) F.
{(Γ,comp). snd(comp!0) ∈ Normal ` pre ∧
(∀i. Suc i<length comp ⟶
Γ⊢⇩c(comp!i) →⇩e (comp!(Suc i)) ⟶
(snd(comp!i), snd(comp!(Suc i))) ∈ rely)}"
lemma assum_R_R':
"(Γ, l) ∈ assum(p, R) ⟹
snd(l!0) ∈ Normal ` p' ⟹
R ⊆ R' ⟹
(Γ, l) ∈ assum(p', R')"
proof -
assume a0:"(Γ, l) ∈ assum(p, R)" and
a1:"snd(l!0) ∈ Normal ` p'" and
a2: " R ⊆ R'"
then have "env_tran_right Γ l R"
unfolding assum_def using env_tran_right_def
by force
then have "env_tran_right Γ l R'"
using a2 env_tran_R_R' by blast
thus ?thesis using a1 unfolding assum_def unfolding env_tran_right_def
by fastforce
qed
lemma same_prog_p:
"(Γ,(P,s)#(P,t)#l)∈cptn ⟹
(Γ,(P,s)#(P,t)#l) ∈ assum (p,R) ⟹
Sta p R ⟹
∃t1. t=Normal t1 ∧ t1 ∈ p
"
proof -
assume a0: "(Γ,(P,s)#(P,t)#l)∈cptn" and
a1: "(Γ,(P,s)#(P,t)#l) ∈ assum (p,R)" and
a2: "Sta p R"
then have "Suc 0 < length ((P,s)#(P,t)#l)"
by fastforce
then have "Γ⊢⇩c(((P,s)#(P,t)#l)!0) →⇩c⇩e (((P,s)#(P,t)#l)!(Suc 0))"
using a0 cptn_stepc_rtran by fastforce
then have step_ce:"Γ⊢⇩c(((P,s)#(P,t)#l)!0) →⇩e (((P,s)#(P,t)#l)!(Suc 0)) ∨
Γ⊢⇩c(((P,s)#(P,t)#l)!0) → (((P,s)#(P,t)#l)!(Suc 0))"
using step_ce_elim_cases by blast
then obtain s1 where s:"s=Normal s1 ∧ s1 ∈ p"
using a1 unfolding assum_def
by fastforce
have "∃t1. t=Normal t1 ∧ t1 ∈ p "
using step_ce
proof
{assume step_e:"Γ⊢⇩c ((P, s) # (P, t) # l) ! 0 →⇩e
((P, s) # (P, t) # l) ! Suc 0"
have ?thesis
using a2 a1 s unfolding Sta_def assum_def
proof -
have "(Suc 0 < length ((P, s) # (P, t) # l))"
by fastforce
then have assm:"(s, t) ∈ R"
using s a1 step_e
unfolding assum_def by fastforce
then obtain t1 s2 where s_t:"s= Normal s2 ∧ t = Normal t1"
using a2 s unfolding Sta_def by fastforce
then have R:"(s,t)∈R"
using assm unfolding Satis_def by fastforce
then have "s2=s1" using s s_t by fastforce
then have "t1∈p"
using a2 s s_t R unfolding Sta_def Norm_def by blast
thus ?thesis using s_t by blast
qed thus ?thesis by auto
}
next
{
assume step:"Γ⊢⇩c ((P, s) # (P, t) # l) ! 0 →
((P, s) # (P, t) # l) ! Suc 0"
then have "P≠P ∨ s=t"
proof -
have "Γ⊢⇩c (P, s) → (P, t)"
using local.step by force
then show ?thesis
using step_change_p_or_eq_s by blast
qed
then show ?thesis using s by fastforce
}
qed thus ?thesis by auto
qed
lemma tl_of_assum_in_assum:
"(Γ,(P,s)#(P,t)#l)∈cptn ⟹
(Γ,(P,s)#(P,t)#l) ∈ assum (p,R) ⟹
Sta p R ⟹
(Γ,(P,t)#l) ∈ assum (p,R)
"
proof -
assume a0: "(Γ,(P,s)#(P,t)#l)∈cptn" and
a1: "(Γ,(P,s)#(P,t)#l) ∈ assum (p,R)" and
a2: "Sta p R "
then obtain t1 where t1:"t=Normal t1 ∧ t1 ∈p"
using same_prog_p by blast
then have "env_tran_right Γ ((P,s)#(P,t)#l) R"
using env_tran_right_def a1 unfolding assum_def
by force
then have "env_tran_right Γ ((P,t)#l) R"
using env_tran_tail by auto
thus ?thesis using t1 unfolding assum_def env_tran_right_def by auto
qed
lemma tl_of_assum_in_assum1:
"(Γ,(P,s)#(Q,t)#l)∈cptn ⟹
(Γ,(P,s)#(Q,t)#l) ∈ assum (p,R) ⟹
t ∈ Normal ` q ⟹
(Γ,(Q,t)#l) ∈ assum (q,R)
"
proof -
assume a0: "(Γ,(P,s)#(Q,t)#l)∈cptn" and
a1: "(Γ,(P,s)#(Q,t)#l) ∈ assum (p,R)" and
a2: "t ∈ Normal ` q"
then have "env_tran_right Γ ((P,s)#(Q,t)#l) R"
using env_tran_right_def a1 unfolding assum_def
by force
then have "env_tran_right Γ ((Q,t)#l) R"
using env_tran_tail by auto
thus ?thesis using a2 unfolding assum_def env_tran_right_def by auto
qed
lemma sub_assum:
assumes a0: "(Γ,(x#l0)@l1) ∈ assum (p,R)"
shows "(Γ,x#l0) ∈ assum (p,R)"
proof -
{have p0:"snd x ∈ Normal ` p"
using a0 unfolding assum_def by force
then have "env_tran_right Γ ((x#l0)@l1) R"
using a0 unfolding assum_def
by (auto simp add: env_tran_right_def)
then have env:"env_tran_right Γ (x#l0) R"
using env_tran_subr by blast
also have "snd ((x#l0)!0) ∈ Normal ` p"
using p0 by fastforce
ultimately have "snd ((x#l0)!0) ∈ Normal ` p ∧
(∀i. Suc i<length (x#l0) ⟶
Γ⊢⇩c((x#l0)!i) →⇩e ((x#l0)!(Suc i)) ⟶
(snd((x#l0)!i), snd((x#l0)!(Suc i))) ∈ R)"
unfolding env_tran_right_def by auto
}
then show ?thesis unfolding assum_def by auto
qed
lemma sub_assum_r:
assumes a0: "(Γ,l0@x1#l1) ∈ assum (p,R)" and
a1: "(snd x1) ∈ Normal ` q"
shows "(Γ,x1#l1) ∈ assum (q,R)"
proof -
have "env_tran_right Γ (l0@x1#l1) R"
using a0 unfolding assum_def env_tran_right_def
by fastforce
then have "env_tran_right Γ (x1#l1) R"
using env_tran_subl by auto
thus ?thesis using a1 unfolding assum_def env_tran_right_def by fastforce
qed
definition comm ::
"(('s,'f) tran) set ×
('s set × 's set) ⇒
'f set ⇒
(('s,'p,'f,'e) confs) set" where
"comm ≡ λ(guar, (q,a)) F.
{c. snd (last (snd c)) ∉ Fault ` F ⟶
(∀i.
Suc i<length (snd c) ⟶
(fst c)⊢⇩c((snd c)!i) → ((snd c)!(Suc i)) ⟶
(snd((snd c)!i), snd((snd c)!(Suc i))) ∈ guar) ∧
(final (last (snd c)) ⟶
((fst (last (snd c)) = Skip ∧
snd (last (snd c)) ∈ Normal ` q)) ∨
(fst (last (snd c)) = Throw ∧
snd (last (snd c)) ∈ Normal ` a))}"
definition comm1 ::
"(('s,'f) tran) set ×
('s set × 's set) ⇒
'f set ⇒
(('s,'p,'f,'e) confs) set" where
"comm1 ≡ λ(guar, (q,a)) F.
{(Γ,comp). snd (last comp) ∉ Fault ` F ⟶
(∀i.
Suc i<length comp ⟶
Γ⊢⇩c(comp!i) → (comp!(Suc i)) ⟶
(snd(comp!i), snd(comp!(Suc i))) ∈ guar) ∧
(final (last comp) ⟶
((fst (last comp) = Skip ∧
snd (last comp) ∈ Normal ` q)) ∨
(fst (last comp) = Throw ∧
snd (last comp) ∈ Normal ` a))}"
lemma comm_dest:
"(Γ, l)∈ comm (G,(q,a)) F ⟹
snd (last l) ∉ Fault ` F ⟹
(∀i. Suc i<length l ⟶
Γ⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
unfolding comm_def
apply clarify
apply (drule mp)
apply fastforce
apply (erule conjE)
apply (erule allE)
by auto
lemma comm_dest1:
"(Γ, l)∈ comm (G,(q,a)) F ⟹
snd (last l) ∉ Fault ` F ⟹
Suc i<length l ⟹
Γ⊢⇩c(l!i) → (l!(Suc i)) ⟹
(snd(l!i), snd(l!(Suc i))) ∈ G"
unfolding comm_def
apply clarify
apply (drule mp)
apply fastforce
apply (erule conjE)
apply (erule allE)
by auto
lemma comm_dest2:
assumes a0: "(Γ, l)∈ comm (G,(q,a)) F" and
a1: "final (last l)" and
a2: "snd (last l) ∉ Fault ` F"
shows " ((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` a)"
proof -
show ?thesis using a0 a1 a2 unfolding comm_def by auto
qed
lemma comm_des3:
assumes a0: "(Γ, l)∈ comm (G,(q,a)) F" and
a1: "snd (last l) ∉ Fault ` F"
shows "final (last l) ⟶ ((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` a)"
using a0 a1 unfolding comm_def by auto
lemma commI:
assumes a0:"snd (last l) ∉ Fault ` F ⟹
(∀i.
Suc i<length l ⟶
Γ⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G) ∧
(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` a))"
shows "(Γ,l)∈comm (G, (q,a)) F"
using a0 unfolding comm_def
apply clarify
by simp
lemma comm_conseq:
"(Γ,l) ∈ comm(G', (q',a')) F ⟹
G' ⊆ G ∧
q' ⊆ q ∧
a' ⊆ a ⟹
(Γ,l) ∈ comm (G,(q,a)) F"
proof -
assume a0:"(Γ,l) ∈ comm(G', (q',a')) F" and
a1:" G' ⊆ G ∧
q' ⊆ q ∧
a' ⊆ a"
{
assume a:"snd (last l) ∉ Fault ` F "
have l:"(∀i.
Suc i<length l ⟶
Γ⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof -
{fix i ns ns'
assume a00:"Suc i<length l" and
a11:"Γ⊢⇩c(l!i) → (l!(Suc i))"
have "(snd(l!i), snd(l!(Suc i))) ∈ G"
proof -
have "(snd(l!i), snd(l!(Suc i))) ∈ G'"
using comm_dest1 [OF a0 a a00 a11] by auto
thus ?thesis using a1 unfolding Satis_def sep_conj_def by fastforce
qed
} thus ?thesis by auto
qed
have "(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` a))"
proof -
{assume a33:"final (last l)"
then have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q')) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` a')"
using comm_dest2[OF a0 a33 a] by auto
then have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` a)"
using a1 by fastforce
} thus ?thesis by auto
qed
note res1 = conjI[OF l this]
} thus ?thesis unfolding comm_def by simp
qed
lemma no_comp_tran_no_final_comm:
assumes a0:"∀i<length l. ¬ final (l!i)" and
a1:"∀i<length l. fst (l!i) = C" and a2:"length l>0"
shows "(Γ,l)∈comm(G, (q,a)) F"
proof-
have n_comp:"∀i. Suc i < length l ⟶ ¬ (Γ⊢⇩c (l!i) → (l!Suc i))" using a1
by (metis Suc_lessD mod_env_not_component prod.collapse)
{assume a00:"snd (last l) ∉ Fault ` F"
{fix i
assume "Suc i< length(l)" and
"Γ⊢⇩c(l!i) → (l!(Suc i))"
then have False using n_comp by auto
}
moreover {
assume "final (last l)"
then have False using a0 a2
using last_conv_nth by force
}
ultimately have ?thesis unfolding comm_def using a00 by auto
} thus ?thesis unfolding comm_def by auto
qed
definition com_validity ::
"('s,'p,'f,'e) body ⇒ 'f set ⇒ ('s,'p,'f,'e) com ⇒
's set ⇒ (('s,'f) tran) set ⇒ (('s,'f) tran) set ⇒
's set ⇒ 's set ⇒ bool"
("_ ⊨⇘'/_⇙/ _ sat [_,_, _, _,_]" [61,60,0,0,0,0,0,0] 45) where
"Γ ⊨⇘/F⇙ Pr sat [p, R, G, q,a] ≡
∀s. cp Γ Pr s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
definition com_cvalidity::
"('s,'p,'f,'e) body ⇒
('s,'p,'f,'e) sextuple set ⇒
'f set ⇒
('s,'p,'f,'e) com ⇒
's set ⇒
(('s,'f) tran) set ⇒
(('s,'f) tran) set ⇒
's set ⇒
's set ⇒
bool"
("_,_ ⊨⇘'/_⇙/ _ sat [_,_, _, _,_]" [61,60,0,0,0,0,0,0] 45) where
"Γ,Θ ⊨⇘/F⇙ Pr sat [p, R, G, q,a] ≡
(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a]) ⟶
Γ ⊨⇘/F⇙ Pr sat [p, R, G, q,a]"
definition com_validityn ::
"('s,'p,'f,'e) body ⇒ nat ⇒ 'f set ⇒ ('s,'p,'f,'e) com ⇒
's set ⇒ (('s,'f) tran) set ⇒ (('s,'f) tran) set ⇒
's set ⇒ 's set ⇒ bool"
("_ ⊨_⇘'/_⇙/ _ sat [_,_, _, _,_]" [61,0,60,0,0,0,0,0,0] 45) where
"Γ ⊨n⇘/F⇙ Pr sat [p, R, G, q,a] ≡
∀s. cpn n Γ Pr s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
definition com_cvalidityn::
"('s,'p,'f,'e) body ⇒
('s,'p,'f,'e) sextuple set ⇒ nat ⇒
'f set ⇒
('s,'p,'f,'e) com ⇒
's set ⇒
(('s,'f) tran) set ⇒
(('s,'f) tran) set ⇒
's set ⇒
's set ⇒
bool"
("_,_ ⊨_⇘'/_⇙/ _ sat [_,_, _, _,_]" [61,60,0,0,0,0,0,0,0] 45) where
"Γ,Θ ⊨n⇘/F⇙ Pr sat [p, R, G, q,a] ≡
(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]) ⟶
Γ ⊨n⇘/F⇙ Pr sat [p, R, G, q,a]"
lemma com_valid_iff_nvalid:"(Γ ⊨⇘/F⇙ Pr sat [p, R, G, q,a]) = (∀n. Γ ⊨n⇘/F⇙ Pr sat [p, R, G, q,a])"
apply (simp only: com_validity_def com_validityn_def cp_def cpn_def cptn_eq_cptn_mod_nest)
by fast
lemma com_cnvalid_to_cvalid: "(∀n. (Γ,Θ⊨n⇘/F⇙ Pr sat [p, R, G, q,a])) ⟹ (Γ,Θ⊨⇘/F⇙ Pr sat [p, R, G, q,a])"
apply (unfold com_cvalidityn_def com_cvalidity_def com_valid_iff_nvalid [THEN eq_reflection])
by fast
lemma etran_in_comm:
"(Γ,(P, t) # xs) ∈ comm(G, (q,a)) F ⟹
¬ (Γ⊢⇩c((P,s)) → ((P,t))) ⟹
(Γ,(P, s) # (P, t) # xs) ∈ cptn ⟹
(Γ,(P, s) # (P, t) # xs) ∈ comm(G, (q,a)) F"
proof -
assume a1:"(Γ,(P, t) # xs) ∈ comm(G, (q,a)) F" and
a2:"¬ Γ⊢⇩c((P,s)) → ((P,t))" and
a3:"(Γ,(P, s) # (P, t) # xs) ∈ cptn"
show ?thesis using comm_def a1 a2 a3
proof -
{
let ?l1 = "(P, t) # xs"
let ?l = "(P, s) # ?l1"
assume a00:"snd (last ?l) ∉ Fault ` F"
have concl:"(∀i ns ns'. Suc i<length ?l ⟶
Γ⊢⇩c(?l!i) → (?l!(Suc i)) ⟶
(snd(?l!i), snd(?l!(Suc i))) ∈ G)"
proof -
{fix i ns ns'
assume a11:"Suc i < length ?l" and
a12:"Γ⊢⇩c (?l ! i) → ( ?l ! Suc i)"
have p1:"(∀i ns ns'. Suc i<length ?l1 ⟶
Γ⊢⇩c(?l1!i) → (?l1!(Suc i)) ⟶
(snd(?l1!i), snd(?l1!(Suc i))) ∈ G)"
using a1 a3 a00 unfolding comm_def by auto
have "(snd (?l ! i), snd (?l ! Suc i)) ∈ G"
proof (cases i)
case 0
have "Γ⊢⇩c (P, s) → (P, t)" using a12 0 by auto
thus ?thesis using a2 by auto
next
case (Suc n) thus ?thesis
proof -
have f1: "Γ⊢⇩c ((P, t) # xs) ! n → ((P, t) # xs) ! Suc n"
using Suc a12 by fastforce
have f2: "Suc n < length ((P, t) # xs)"
using Suc a11 by fastforce
have "snd (last ((P, t) # xs)) ∉ Fault ` F"
by (metis (no_types) a00 last.simps list.distinct(1))
hence "(snd (((P, t) # xs) ! n), snd (((P, t) # xs) ! Suc n)) ∈ G"
using f2 f1 a1 comm_dest1 by blast
thus ?thesis
by (simp add: Suc)
qed
qed
} thus ?thesis by auto
qed
have concr:"(final (last ?l) ⟶
((fst (last ?l) = Skip ∧
snd (last ?l) ∈ Normal ` q)) ∨
(fst (last ?l) = Throw ∧
snd (last ?l) ∈ Normal ` a))"
using a1 a00 unfolding comm_def by auto
note res1=conjI[OF concl concr] }
thus ?thesis unfolding comm_def by auto qed
qed
lemma ctran_in_comm:
" (Normal s,Normal s) ∈ G ⟹
(Γ,(Q, Normal s) # xs) ∈ comm(G, (q,a)) F ⟹
(Γ,(P, Normal s) # (Q, Normal s) # xs) ∈ comm(G, (q,a)) F"
proof -
assume a1:"(Normal s,Normal s) ∈ G" and
a2:"(Γ,(Q, Normal s) # xs) ∈ comm(G, (q,a)) F"
show ?thesis using comm_def a1 a2
proof -
{
let ?l1 = "(Q, Normal s) # xs"
let ?l = "(P, Normal s) # ?l1"
assume a00:"snd (last ?l) ∉ Fault ` F"
have concl:"(∀i. Suc i<length ?l ⟶
Γ⊢⇩c(?l!i) → (?l!(Suc i)) ⟶
(snd(?l!i), snd(?l!(Suc i))) ∈ G)"
proof -
{fix i ns ns'
assume a11:"Suc i < length ?l" and
a12:"Γ⊢⇩c (?l ! i) → ( ?l ! Suc i)"
have p1:"(∀i. Suc i<length ?l1 ⟶
Γ⊢⇩c(?l1!i) → (?l1!(Suc i)) ⟶
(snd(?l1!i), snd(?l1!(Suc i))) ∈ G)"
using a2 a00 unfolding comm_def by auto
have "(snd (?l ! i), snd (?l ! Suc i)) ∈ G"
proof (cases i)
case 0
then have "snd (((P, Normal s) # (Q, Normal s) # xs) ! i) = Normal s ∧
snd (((P, Normal s) # (Q, Normal s) # xs) ! (Suc i)) = Normal s"
by fastforce
also have "(Normal s, Normal s) ∈ G"
using Satis_def a1 by blast
ultimately show ?thesis using a1 Satis_def by auto
next
case (Suc n) thus ?thesis using p1 a2 a11 a12
proof -
have f1: "Γ⊢⇩c ((Q, Normal s) # xs) ! n → ((Q, Normal s) # xs) ! Suc n"
using Suc a12 by fastforce
have f2: "Suc n < length ((Q, Normal s) # xs)"
using Suc a11 by fastforce
thus ?thesis using Suc f1 nth_Cons_Suc p1 by auto
qed
qed
} thus ?thesis by auto
qed
have concr:"(final (last ?l) ⟶
snd (last ?l) ∉ Fault ` F ⟶
((fst (last ?l) = Skip ∧
snd (last ?l) ∈ Normal ` q)) ∨
(fst (last ?l) = Throw ∧
snd (last ?l) ∈ Normal ` a))"
using a2 unfolding comm_def by auto
note res=conjI[OF concl concr]}
thus ?thesis unfolding comm_def by auto qed
qed
lemma not_final_in_comm:
"(Γ,(Q, Normal s) # xs) ∈ comm(G, (q,a)) F ⟹
¬ final (last ((Q, Normal s) # xs)) ⟹
(Γ,(Q, Normal s) # xs) ∈ comm(G, (q',a')) F"
unfolding comm_def by force
lemma comm_union:
assumes
a0: "(Γ,xs) ∈ comm(G, (q,a)) F" and
a1: "(Γ,ys) ∈ comm(G, (q',a')) F" and
a2: "xs≠[] ∧ ys≠[]" and
a3: "( snd (last xs),snd (ys!0)) ∈ G" and
a4: "(Γ,xs@ys) ∈ cptn"
shows "(Γ,xs@ys) ∈ comm(G, (q',a')) F"
proof -
{
let ?l="xs@ys"
assume a00:"snd (last (xs@ys)) ∉ Fault ` F"
have last_ys:"last (xs@ys) = last ys" using a2 by fastforce
have concl:"(∀i. Suc i<length ?l ⟶
Γ⊢⇩c(?l!i) → (?l!(Suc i)) ⟶
(snd(?l!i), snd(?l!(Suc i))) ∈ G)"
proof -
{fix i ns ns'
assume a11:"Suc i < length ?l" and
a12:"Γ⊢⇩c (?l ! i) → ( ?l ! Suc i)"
have all_ys:"∀i≥length xs. (xs@ys)!i = ys!(i-(length xs))"
by (simp add: nth_append)
have all_xs:"∀i<length xs. (xs@ys)!i = xs!i"
by (simp add: nth_append)
have "(snd(?l!i), snd(?l!(Suc i))) ∈ G"
proof (cases "Suc i>length xs")
case True
have "Suc (i - (length xs)) < length ys" using a11 True by fastforce
moreover have "Γ⊢⇩c (ys ! (i-(length xs))) → ( ys ! ((Suc i)-(length xs)))"
using a12 all_ys True by fastforce
moreover have "snd (last ys) ∉ Fault ` F" using last_ys a00 by fastforce
ultimately have "(snd(ys!(i-(length xs))), snd(ys!Suc (i-(length xs)))) ∈ G"
using a1 comm_dest1[of Γ ys G q' a' F "i-length xs"] True Suc_diff_le by fastforce
thus ?thesis using True all_ys Suc_diff_le by fastforce
next
case False note F1=this thus ?thesis
proof (cases "Suc i < length xs")
case True
then have "snd ((xs@ys)!(length xs -1)) ∉ Fault ` F"
using a00 a2 a4
by (simp add: last_not_F )
then have "snd (last xs) ∉ Fault ` F" using all_xs a2 by (simp add: last_conv_nth )
moreover have "Γ⊢⇩c (xs ! i) → ( xs ! Suc i)"
using True all_xs a12 by fastforce
ultimately have"(snd(xs!i), snd(xs!(Suc i))) ∈ G"
using a0 comm_dest1[of Γ xs G q a F i] True by fastforce
thus ?thesis using True all_xs by fastforce
next
case False
then have suc_i:"Suc i = length xs" using F1 by fastforce
then have i:"i=length xs -1" using a2 by fastforce
then show ?thesis using a3
by (simp add: a2 all_xs all_ys last_conv_nth )
qed
qed
} thus ?thesis by auto
qed
have concr:"(final (last ?l) ⟶
((fst (last ?l) = Skip ∧
snd (last ?l) ∈ Normal ` q')) ∨
(fst (last ?l) = Throw ∧
snd (last ?l) ∈ Normal ` a'))"
using a1 last_ys a00 a2 comm_des3 by fastforce
note res=conjI[OF concl concr]}
thus ?thesis unfolding comm_def by auto
qed
lemma cpn_rule1:"(∀s. cpn n Γ P s ∩ assum(p, R) ⊆ comm(G, (q,a)) F) ⟹
(∀s l. (Γ,l)∈cpn n Γ P s ∧ (Γ,l)∈ assum (p, R) ⟶ (Γ,l) ∈ comm(G, (q,a)) F)"
proof-
assume a0:"∀s. cpn n Γ P s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
{fix s l
assume a00:"(Γ,l)∈cpn n Γ P s ∧ (Γ,l)∈ assum (p, R)"
then have "cpn n Γ P s ∩ assum(p, R) ⊆ comm(G, (q,a)) F" using a0 by auto
then have "(Γ,l) ∈ comm(G, (q,a)) F" using a00 unfolding cpn_def assum_def comm_def
by blast
} then show ?thesis by auto
qed
lemma cpn_rule2:"(∀s l. (Γ,l)∈cpn n Γ P s ∧ (Γ,l)∈ assum (p, R) ⟶ (Γ,l) ∈ comm(G, (q,a)) F) ⟹
(∀s. cpn n Γ P s ∩ assum(p, R) ⊆ comm(G, (q,a)) F)"
proof-
assume a0:"∀s l. (Γ,l)∈cpn n Γ P s ∧ (Γ,l)∈ assum (p, R) ⟶ (Γ,l) ∈ comm(G, (q,a)) F"
{fix s l
assume a00:"(Γ,l)∈ cpn n Γ P s ∧ (Γ,l)∈assum(p, R)"
then have "(Γ,l) ∈ comm(G, (q,a)) F" using a0 unfolding cpn_def assum_def comm_def
by blast
} then show ?thesis unfolding cpn_def by fastforce
qed
lemma cpn_rule:"(∀s l. (Γ,l)∈cpn n Γ P s ∧ (Γ,l)∈ assum (p, R) ⟶ (Γ,l) ∈ comm(G, (q,a)) F) =
(∀s. cpn n Γ P s ∩ assum(p, R) ⊆ comm(G, (q,a)) F)"
using cpn_rule1 cpn_rule2
by metis
lemma split_list_i:"i<length l ⟹
∃l1 l2. l = l1@(l!i#l2)"
proof(induct l arbitrary: i)
case Nil
then show ?case by auto
next
case (Cons a l)
then show ?case
using id_take_nth_drop by blast
qed
lemma sub_assum1:
assumes a0: "(Γ,l0@l1) ∈ assum (p,R)" and a1:"l0≠[]"
shows "(Γ,l0) ∈ assum (p,R)"
by (metis a0 a1 append_self_conv2 id_take_nth_drop length_greater_0_conv sub_assum take_0)
subsection ‹Validity for Parallel Programs.›
definition All_End :: "('s,'p,'f,'e) par_config ⇒ bool" where
"All_End xs ≡ fst xs ≠[] ∧ (∀i<length (fst xs). final ((fst xs)!i,snd xs))"
definition par_assum ::
"('s set ×
(('s,'f) tran) set) ⇒
(('s,'p,'f,'e) par_confs) set" where
"par_assum ≡
λ(pre, rely). {c.
snd((snd c)!0) ∈ Normal ` pre ∧ (∀i. Suc i<length (snd c) ⟶
(fst c)⊢⇩p((snd c)!i) →⇩e ((snd c)!(Suc i)) ⟶
(snd((snd c)!i), snd((snd c)!(Suc i))) ∈ rely)}"
definition par_comm ::
"((('s,'f) tran) set ×
('s set × 's set)) ⇒
'f set ⇒
(('s,'p,'f,'e) par_confs) set" where
"par_comm ≡
λ(guar, (q,a)) F.
{c. snd (last (snd c)) ∉ Fault ` F ⟶
(∀i.
Suc i<length (snd c) ⟶
(fst c)⊢⇩p((snd c)!i) → ((snd c)!(Suc i)) ⟶
(snd((snd c)!i), snd((snd c)!(Suc i))) ∈ guar) ∧
(All_End (last (snd c)) ⟶
(∃j<length (fst (last (snd c))). fst (last (snd c))!j=Throw ∧
snd (last (snd c)) ∈ Normal ` a) ∨
(∀j<length (fst (last (snd c))). fst (last (snd c))!j=Skip ∧
snd (last (snd c)) ∈ Normal ` q))}"
definition par_com_validity ::
"('s,'p,'f,'e) body ⇒
'f set ⇒
('s,'p,'f,'e) par_com ⇒
('s set) ⇒
((('s,'f) tran) set) ⇒
((('s,'f) tran) set) ⇒
('s set) ⇒
('s set) ⇒
bool"
("_ ⊨⇘'/_⇙/ _ SAT [_, _, _, _,_]" [61,60,0,0,0,0,0,0] 45) where
"Γ ⊨⇘/F⇙ Ps SAT [pre, R, G, q,a] ≡
∀s. par_cp Γ Ps s ∩ par_assum(pre, R) ⊆ par_comm(G, (q,a)) F"
definition par_com_cvalidity ::
"('s,'p,'f,'e) body ⇒
('s,'p,'f,'e) sextuple set ⇒
'f set ⇒
('s,'p,'f,'e) par_com ⇒
('s set) ⇒
((('s,'f) tran) set) ⇒
((('s,'f) tran) set) ⇒
('s set) ⇒
('s set) ⇒
bool"
("_,_ ⊨⇘'/_⇙/ _ SAT [_, _, _, _,_]" [61,60,0,0,0,0,0,0] 45) where
"Γ,Θ ⊨⇘/F⇙ Ps SAT [p, R, G, q,a] ≡
(∀(c,p,R,G,q,a)∈ Θ. (Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a])) ⟶
Γ ⊨⇘/F⇙ Ps SAT [p, R, G, q,a]"
declare Un_subset_iff [simp del] sup.bounded_iff [simp del]
inductive
lrghoare :: "[('s,'p,'f,'e) body,
('s,'p,'f,'e) sextuple set,
'f set,
('s,'p,'f,'e) com,
('s set),
(('s,'f) tran) set, (('s,'f) tran) set,
's set,
's set] ⇒ bool"
("_,_ ⊢⇘'/_⇙ _ sat [_, _, _, _,_]" [61,61,60,60,0,0,0,0] 45)
where
Skip: "⟦ Sta q R; (∀s. (Normal s, Normal s) ∈ G) ⟧ ⟹
Γ,Θ ⊢⇘/F⇙ Skip sat [q, R, G, q,a] "
|Spec: "⟦Sta p R;Sta q R;
(∀s t. s∈p ∧ (s,t)∈r ⟶ (Normal s,Normal t) ∈ G);
p ⊆ {s. (∀t. (s,t)∈r ⟶ t ∈ q) ∧ (∃t. (s,t) ∈ r)} ⟧ ⟹
Γ,Θ ⊢⇘/F⇙ (Spec r e) sat [p, R, G, q,a]"
| Basic: "⟦ Sta p R;Sta q R;
(∀s t. s∈p ∧ (t=f s) ⟶ (Normal s,Normal t) ∈ G);
p ⊆ {s. f s ∈ q} ⟧ ⟹
Γ,Θ ⊢⇘/F⇙ (Basic f e) sat [p, R, G, q,a]"
| If: "⟦Sta p R; (∀s. (Normal s, Normal s) ∈ G);
Γ,Θ ⊢⇘/F⇙ c1 sat [p ∩ b, R, G, q,a];
Γ,Θ ⊢⇘/F⇙ c2 sat [p ∩ (-b), R, G, q,a]⟧ ⟹
Γ,Θ ⊢⇘/F⇙ (Cond b c1 c2) sat [p, R, G, q,a]"
| While: "⟦ Sta p R; Sta (p ∩ (-b)) R; Sta a R; (∀s. (Normal s,Normal s) ∈ G);
Γ,Θ ⊢⇘/F⇙ c sat [p ∩ b, R, G, p,a]⟧ ⟹
Γ,Θ ⊢⇘/F⇙ (While b c) sat [p, R, G, p ∩ (-b),a]"
| Seq: "⟦Sta a R; Sta p R; (∀s. (Normal s,Normal s) ∈ G);
Γ,Θ ⊢⇘/F⇙ c1 sat [p, R, G, q,a]; Γ,Θ ⊢⇘/F⇙ c2 sat [q, R, G, r,a]⟧ ⟹
Γ,Θ ⊢⇘/F⇙ (Seq c1 c2) sat [p, R, G, r,a]"
| Await: "⟦ Sta p R; Sta q R; Sta a R;
∀V. Γ⇩¬⇩a,{}⊢⇘/F⇙
(p ∩ b ∩ {V}) c
({s. (Normal V, Normal s) ∈ G} ∩ q),
({s. (Normal V, Normal s) ∈ G} ∩ a)⟧ ⟹
Γ,Θ ⊢⇘/F⇙ (Await b c e) sat [p, R, G, q,a]"
| Guard: "⟦Sta (p ∩ g) R; (∀s. (Normal s, Normal s) ∈ G);
Γ,Θ ⊢⇘/F⇙ c sat [p ∩ g, R, G, q,a]⟧ ⟹
Γ,Θ ⊢⇘/F⇙ (Guard f g c) sat [p ∩ g, R, G, q,a]"
| Guarantee: "⟦ Sta p R; (∀s. (Normal s, Normal s) ∈ G); f∈F;
Γ,Θ ⊢⇘/F⇙ c sat [p ∩ g, R, G, q,a] ⟧ ⟹
Γ,Θ ⊢⇘/F⇙ (Guard f g c) sat [p, R, G, q,a]"
| CallRec: "⟦(c,p,R,G,q,a) ∈ Specs;
∀(c,p,R,G,q,a)∈ Specs. c ∈ dom Γ ∧
Sta p R ∧ (∀s. (Normal s,Normal s) ∈ G)∧
Γ,Θ ∪ Specs ⊢⇘/F⇙ (the (Γ c)) sat [p, R, G, q,a];
Sta p R; (∀s. (Normal s, Normal s) ∈ G)⟧ ⟹
Γ,Θ⊢⇘/F⇙ (Call c) sat [p, R, G, q,a]"
| Asm: "⟦(c,p,R,G,q,a) ∈ Θ⟧ ⟹
Γ,Θ ⊢⇘/F⇙ (Call c) sat [p, R, G, q,a]"
| Call: "⟦
Sta p R; (∀s. (Normal s, Normal s) ∈ G);c ∈ dom Γ;
Γ,Θ ⊢⇘/F⇙ (the (Γ c)) sat [p, R, G, q,a]⟧ ⟹
Γ,Θ ⊢⇘/F⇙ (Call c) sat [p, R, G, q,a]"
| DynCom: "⟦(Sta p R) ∧ (Sta q R) ∧ (Sta a R) ∧
(∀s. (Normal s,Normal s) ∈ G);
(∀s ∈ p. (Γ,Θ⊢⇘/F⇙ (c s) sat [p, R, G, q,a]))⟧ ⟹
Γ,Θ⊢⇘/F⇙ (DynCom c) sat [p, R, G, q,a]"
| Throw: "⟦Sta a R; (∀s. (Normal s, Normal s) ∈ G) ⟧ ⟹
Γ,Θ ⊢⇘/F⇙ Throw sat [a, R, G, q,a] "
| Catch: "⟦Sta q R; (∀s. (Normal s, Normal s) ∈ G);
Γ,Θ ⊢⇘/F⇙ c1 sat [p, R, G, q,r];
Γ,Θ ⊢⇘/F⇙ c2 sat [r, R, G, q,a]⟧ ⟹
Γ,Θ ⊢⇘/F⇙ (Catch c1 c2) sat [p, R, G, q,a]"
| Conseq: "∀s ∈ p.
(∃p' R' G' q' a' Θ'.
(s∈ p') ∧
R ⊆ R' ∧
G' ⊆ G ∧
q' ⊆ q ∧
a' ⊆ a ∧ Θ' ⊆ Θ ∧
(Γ,Θ'⊢⇘/F⇙ P sat [p', R', G', q',a']) )
⟹ Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q,a]"
| Conj_post: " Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q,a] ⟹
Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q',a']
⟹ Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q ∩ q',a ∩ a']"
| Conj_Inter: "sa≠({}::nat set) ⟹
∀i∈sa. Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q i,a] ⟹
Γ,Θ⊢⇘/F⇙ P sat [p, R, G,⋂i∈sa. q i,a]"
inductive_cases hoare_elim_cases [cases set]:
"Γ,Θ ⊢⇘/F⇙ Skip sat [p, R, G, q,a]"
thm hoare_elim_cases
definition Pre :: " ('s,'p,'f,'e)rgformula ⇒ ('s set)" where
"Pre x ≡ fst(snd x)"
definition Post :: "('s,'p,'f,'e) rgformula ⇒ ('s set)" where
"Post x ≡ fst(snd(snd(snd(snd x))))"
definition Abr :: "('s,'p,'f,'e) rgformula ⇒ ('s set)" where
"Abr x ≡ snd(snd(snd(snd(snd x))))"
definition Rely :: "('s,'p,'f,'e) rgformula ⇒ (('s,'f) tran) set" where
"Rely x ≡ fst(snd(snd x))"
definition Guar :: "('s,'p,'f,'e) rgformula ⇒ (('s,'f) tran) set" where
"Guar x ≡ fst(snd(snd(snd x)))"
definition Com :: "('s,'p,'f,'e) rgformula ⇒ ('s ,'p,'f,'e) com" where
"Com x ≡ fst x"
inductive
par_rghoare :: "[('s,'p,'f,'e) body,
('s,'p,'f,'e) sextuple set,
'f set,
( ('s,'p,'f,'e) rgformula) list,
's set,
(('s,'f) tran) set, (('s,'f) tran) set,
's set,
's set] ⇒ bool"
("_,_ ⊢⇘'/_⇙ _ SAT [_, _, _, _,_]" [61,60,60,0,0,0,0] 45)
where
Parallel:
"⟦ ∀i<length xs. R ∪ (⋃j∈{j. j<length xs ∧ j≠i}. (Guar(xs!j))) ⊆ (Rely(xs!i));
(⋃j<length xs. (Guar(xs!j))) ⊆ G;
p ⊆ (⋂i<length xs. (Pre(xs!i)));
(⋂i<length xs. (Post(xs!i))) ⊆ q;
(⋃i<length xs. (Abr(xs!i))) ⊆ a;
∀i<length xs. Γ,Θ ⊢⇘/F⇙ Com(xs!i) sat [Pre(xs!i),Rely(xs!i),Guar(xs!i),Post(xs!i),Abr(xs!i)] ⟧
⟹ Γ,Θ ⊢⇘/F⇙ xs SAT [p, R, G, q,a]"
section {* Soundness *}
lemma skip_suc_i:
assumes a1:"(Γ, l) ∈ cptn ∧ fst (l!i) = Skip"
assumes a2:"i+1 < length l"
shows "fst (l!(i+1)) = Skip"
proof -
from a2 a1 obtain l1 ls where "l=l1#ls"
by (metis list.exhaust list.size(3) not_less0)
then have "Γ⊢⇩c (l!i) →⇩c⇩e (l!(Suc i))" using cptn_stepc_rtran a1 a2
by fastforce
thus ?thesis using a1 a2 step_ce_elim_cases
by (metis (no_types) Suc_eq_plus1 not_eq_not_env prod.collapse stepc_elim_cases(1))
qed
lemma throw_suc_i:
assumes a1:"(Γ, l) ∈ cptn ∧ (fst(l!i) = Throw ∧ snd(l!i) = Normal s1)"
assumes a2:"Suc i < length l"
assumes a3:"env_tran_right Γ l rely ∧ Sta q rely ∧ s1 ∈ q"
shows "fst (l!(Suc i)) = Throw ∧ (∃s2. snd(l!(Suc i)) = Normal s2 ∧ s2 ∈q)"
proof -
have fin:"final (l!i)" using a1 unfolding final_def by auto
from a2 a1 obtain l1 ls where "l=l1#ls"
by (metis list.exhaust list.size(3) not_less0)
then have "Γ⊢⇩c (l!i) →⇩c⇩e (l!(Suc i))" using cptn_stepc_rtran a1 a2
by fastforce then have "Γ⊢⇩c (l!i) → (l!(Suc i)) ∨ Γ⊢⇩c (l!i) →⇩e (l!(Suc i))"
using step_ce_elim_cases by blast
thus ?thesis proof
assume "Γ⊢⇩c (l!i) → (l!(Suc i))" thus ?thesis using fin no_step_final' by blast
next
assume "Γ⊢⇩c (l!i) →⇩e (l!(Suc i))" thus ?thesis
using a1 a3 a2 env_tran_normal by (metis (no_types, lifting) env_c_c' prod.collapse)
qed
qed
lemma i_skip_all_skip:assumes a1:"(Γ, l) ∈ cptn ∧ fst (l!i) = Skip"
assumes a2: "i≤j ∧ j < (length l)"
assumes a3:"n=j-i"
shows "fst (l!j) = Skip"
using a1 a2 a3
proof (induct n arbitrary: i j)
case 0
then have "Suc i = Suc j" by simp
thus ?case using "0.prems" skip_suc_i by fastforce
next
case (Suc n)
then have "length l > Suc i" by auto
then have "i<j" using Suc by fastforce
moreover then have "j-1< length l" using Suc by fastforce
moreover then have "j - i = Suc n" using Suc by fastforce
ultimately have "fst (l ! (j)) = LanguageCon.com.Skip" using Suc skip_suc_i
by (metis (no_types, lifting) Suc_diff_Suc Suc_eq_plus1 Suc_leI `Suc i < length l` diff_Suc_1)
also have "j=j" using Cons using Suc.prems(2) by linarith
ultimately show ?case using Suc by (metis (no_types))
qed
lemma i_throw_all_throw:assumes a1:"(Γ, l) ∈ cptn ∧ (fst (l!i) = Throw ∧ snd (l!i) = Normal s1)"
assumes a2: "i≤j ∧ j < (length l)"
assumes a3:"n=j-i"
assumes a4:"env_tran_right Γ l rely ∧ Sta q rely ∧ s1∈q"
shows "fst (l!j) = Throw ∧ (∃s2. snd(l!j) = Normal s2 ∧ s2∈q)"
using a1 a2 a3 a4
proof (induct n arbitrary: i j s1)
case 0
then have "Suc i = Suc j" by simp
thus ?case using "0.prems" skip_suc_i by fastforce
next
case (Suc n)
then have l_suc:"length l > Suc i" by linarith
then have "i<j" using Suc.prems(3) by linarith
moreover then have "j-1< length l" by (simp add: Suc.prems(2) less_imp_diff_less)
moreover then have "j - Suc i = n" by (metis Suc_diff_Suc Suc_inject ‹i < j› Suc(4))
ultimately obtain s2 where "fst (l ! (j-1)) = LanguageCon.com.Throw ∧ snd (l ! (j-1)) = Normal s2 ∧ s2∈q"
using Suc(1)[of i s1 "j-1"] Suc(2) Suc(5)
by (metis (no_types, lifting) Suc_diff_Suc diff_Suc_eq_diff_pred diff_zero less_imp_Suc_add not_le not_less_eq_eq zero_less_Suc)
also have "Suc (j - 1) < length l" using Suc by arith
ultimately have "fst (l ! (j)) = LanguageCon.com.Throw ∧ (∃s2. snd(l!j) = Normal s2 ∧ s2∈q)"
using Suc(2-5) throw_suc_i[of Γ l "j-1" s2 rely q] a4
by fastforce
also have "j=j" using Cons using Suc.prems(2) by linarith
ultimately show ?case using Suc by (metis (no_types))
qed
lemma only_one_component_tran_j:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!i) = Skip ∨ fst (l!i) = Throw " and
a1': "snd (l!i) = Normal x ∧ x ∈ q" and
a2: "i≤j ∧ Suc j < length l" and
a3: "(Γ⊢⇩c(l!j) → (l!(Suc j)))" and
a4: "env_tran_right Γ l rely ∧ Sta q rely"
shows "P"
proof -
have "fst (l!j) = Skip ∨ (fst (l!i) = Throw ∧ snd(l!i) = Normal x)"
using a0 a1 a1' a2 a3 a4 i_skip_all_skip by fastforce
also have "(Γ⊢⇩c(l!j) → (l!(Suc j)))" using a3 by fastforce
ultimately show ?thesis
by (meson SmallStepCon.final_def SmallStepCon.no_step_final' Suc_lessD a0 a2 a4 i_throw_all_throw a1')
qed
lemma only_one_component_tran_all_j:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!i) = Skip ∨ (fst (l!i) = Throw ∧ snd(l!i) = Normal s1)" and
a1': "snd (l!i) = Normal x ∧ x ∈ q" and
a2: "Suc i<length l" and
a3: "∀j. i≤j ∧ Suc j < length l ⟶ (Γ⊢⇩c(l!j) → (l!(Suc j)))" and
a4: "env_tran_right Γ l rely ∧ Sta q rely"
shows "P"
using a0 a1 a2 a3 a4 a1' only_one_component_tran_j
by (metis lessI less_Suc_eq_le)
lemma zero_skip_all_skip:
assumes a1:"(Γ, l) ∈ cptn ∧ fst (l!0) = Skip ∧ i < length l"
shows "fst (l!i) = Skip"
using a1 i_skip_all_skip by blast
lemma all_skip:
assumes
a0:"(Γ,x)∈cptn" and
a1:"x!0 = (Skip,s)"
shows "(∀i<length x. fst(x!i) = Skip)"
using a0 a1 zero_skip_all_skip by fastforce
lemma zero_throw_all_throw:
assumes a1:"(Γ, l) ∈ cptn ∧ fst (l!0) = Throw ∧
snd(l!0) = Normal s1 ∧ i < length l ∧ s1∈q"
assumes a2: "env_tran_right Γ l rely ∧ Sta q rely"
shows "fst (l!i) = Throw ∧ (∃s2. snd (l!i) = Normal s2)"
using a1 a2 i_throw_all_throw by (metis le0)
lemma only_one_component_tran_0:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "(fst (l!0) = Skip) ∨ (fst (l!0) = Throw)" and
a1': "snd (l!0) = Normal x ∧ x ∈ q" and
a2: "Suc j < length l" and
a3: "(Γ⊢⇩c(l!j) → (l!(Suc j)))" and
a4: "env_tran_right Γ l rely ∧ Sta q rely"
shows "P"
proof-
have a2':"0≤j ∧ Suc j<length l" using a2 by arith
show ?thesis
using only_one_component_tran_j[OF a0 a1 a1' a2' a3 a4] by auto
qed
lemma not_step_comp_step_env:
assumes a0: "(Γ, l) ∈ cptn" and
a1: "(Suc j<length l)" and
a2: "(∀k < j. ¬((Γ⊢⇩c(l!k) → (l!(Suc k)))))"
shows "(∀k < j. ((Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))))"
proof -
{fix k
assume asm: "k<j"
also then have "Suc k<length l" using a1 a2 by auto
ultimately have "(Γ⊢⇩c(l!k) →⇩c⇩e (l!(Suc k)))" using a0 cptn_stepc_rtran
proof -
obtain nn :: "nat ⇒ nat ⇒ nat" where
f1: "∀x0 x1. (∃v2>x1. x0 = Suc v2) = (x1 < nn x0 x1 ∧ x0 = Suc (nn x0 x1))"
by moura
obtain pp :: "nat ⇒ (('b, 'a, 'c,'d) LanguageCon.com × ('b, 'c) xstate) list ⇒
('b, 'a, 'c,'d) LanguageCon.com × ('b, 'c) xstate" and
pps :: "nat ⇒ (('b, 'a, 'c,'d) LanguageCon.com × ('b, 'c) xstate) list ⇒
(('b, 'a, 'c,'d) LanguageCon.com × ('b, 'c) xstate) list" where
"∀x0 x1. (∃v2 v3. x1 = v2 # v3 ∧ length v3 = x0) = (x1 = pp x0 x1 # pps x0 x1 ∧ length (pps x0 x1) = x0)"
by moura
then have f2: "l = pp (nn (length l) k) l # pps (nn (length l) k) l ∧ length (pps (nn (length l) k) l) = nn (length l) k"
using f1 by (meson Suc_lessE ‹Suc k < length l› length_Suc_conv)
then have f3: "Suc k < length (pp (nn (length l) k) l # pps (nn (length l) k) l)"
by (metis ‹Suc k < length l›)
have "(Γ, pp (nn (length l) k) l # pps (nn (length l) k) l) ∈ cptn"
using f2 a0 by presburger
then have "Γ⊢⇩c (pp (nn (length l) k) l # pps (nn (length l) k) l) ! k →⇩c⇩e (pp (nn (length l) k) l # pps (nn (length l) k) l) ! Suc k"
using f3 by (meson cptn_stepc_rtran)
then show ?thesis
using f2 by auto
qed
also have "¬((Γ⊢⇩c(l!k) → (l!(Suc k))))" using a2 asm by auto
ultimately have "((Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))" using step_ce_elim_cases by blast
} thus ?thesis by auto
qed
lemma cptn_i_env_same_prog:
assumes a0: "(Γ, l) ∈ cptn" and
a1: "∀k < j. k≥i ⟶ (Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))" and
a2: "i≤j ∧ j < length l"
shows "fst (l!j) = fst (l!i)"
using a0 a1 a2
proof (induct "j-i" arbitrary: l j i)
case 0 thus ?case by auto
next
case (Suc n)
then have lenl:"length l>Suc 0" by fastforce
have "j>0" using Suc by linarith
then obtain j1 where prev:"j=Suc j1"
using not0_implies_Suc by blast
then obtain a0 a1 l1 where l:"l=a0#l1@[a1]"
using Suc lenl by (metis add.commute add.left_neutral length_Cons list.exhaust list.size(3) not_add_less1 rev_exhaust)
then have al1_cptn:"(Γ,a0#l1)∈ cptn"
using Suc.prems(1) Suc.prems(3) tl_in_cptn cptn_dest_2
by blast
have i_j:"i≤j1" using Suc prev by auto
have "∀k < j1. k≥i ⟶ (Γ⊢⇩c((a0#l1)!k) →⇩e ((a0#l1)!(Suc k)))"
proof -
{fix k
assume a0:"k<j1 ∧ k≥i"
then have "(Γ⊢⇩c((a0#l1)!k) →⇩e ((a0#l1)!(Suc k)))"
using l Suc(4) prev lenl Suc(5)
proof -
have suc_k_j:"Suc k < j" using a0 prev by blast
have j1_l_l1:"j1 < Suc (length l1)"
using Suc.prems(3) l prev by auto
have "k < Suc j1"
using `k < j1 ∧ i ≤ k` less_Suc_eq by blast
hence f3: "k < j"
using prev by blast
hence ksuc:"k < Suc (Suc j1)"
using less_Suc_eq prev by blast
hence f4: "k < Suc (length l1)"
using prev Suc.prems(3) l a0 j1_l_l1 less_trans
by blast
have f6: "Γ⊢⇩c l ! k →⇩e (l ! Suc k)"
using f3 Suc(4) a0 by blast
have k_l1:"k < length l1"
using f3 Suc.prems(3) i_j l suc_k_j by auto
thus ?thesis
proof (cases k)
case 0 thus ?thesis using f6 l k_l1
by (simp add: nth_append)
next
case (Suc k1) thus ?thesis
using f6 f4 l k_l1
by (simp add: nth_append)
qed
qed
}thus ?thesis by auto
qed
then have fst:"fst ((a0#l1)!i)=fst ((a0#l1)!j1)"
using Suc(1)[of j1 i "a0#l1"]
Suc(2) Suc(3) Suc(4) Suc(5) prev al1_cptn i_j
by (metis (mono_tags, lifting) Suc_diff_le Suc_less_eq diff_Suc_1 l length_Cons length_append_singleton)
have len_l:"length l = Suc (length (a0#l1))" using l by auto
then have f1:"i<length (a0#l1)" using Suc.prems(3) i_j prev by linarith
then have f2:"j1<length (a0#l1)" using Suc.prems(3) len_l prev by auto
have i_l:"fst (l!i) = fst ((a0#l1)!i)"
using l prev f1 f2 fst
by (metis (no_types) append_Cons nth_append)
also have j1_l:"fst (l!j1) = fst ((a0#l1)!j1)"
using l prev f1 f2 fst
by (metis (no_types) append_Cons nth_append)
then have "fst (l!i) = fst (l!j1)" using
i_l j1_l fst by auto
thus ?case using Suc prev by (metis env_c_c' i_j lessI prod.collapse)
qed
lemma cptn_tran_ce_i:
assumes a1:"(Γ, l) ∈ cptn ∧ i + 1 < length l"
shows "Γ⊢⇩c(l!i) →⇩c⇩e (l!(Suc i))"
proof -
from a1
obtain a1 l1 where "l=a1#l1" using cptn.simps by blast
thus ?thesis using a1 cptn_stepc_rtran by fastforce
qed
lemma zero_final_always_env_0:
assumes a1:"(Γ, l) ∈ cptn" and
a2: "fst (l!0) = Skip ∨ fst (l!0) = Throw" and
a2': "snd(l!0) = Normal s1 ∧ s1∈q" and
a3: "Suc i < length l" and
a4: "env_tran_right Γ l rely ∧ Sta q rely"
shows "Γ⊢⇩c(l!i) →⇩e (l!(Suc i))"
proof -
have "Γ⊢⇩c(l!i) →⇩c⇩e (l!(Suc i))" using a1 a2 a3 cptn_tran_ce_i by auto
also have "¬ (Γ⊢⇩c(l!i) → (l!(Suc i)))" using a1 a2 a3 a4 a2'
using only_one_component_tran_0 by metis
ultimately show ?thesis by (simp add: step_ce.simps)
qed
lemma final_always_env_i:
assumes a1:"(Γ, l) ∈ cptn" and
a2: "fst (l!0) = Skip ∨ fst (l!0) = Throw" and
a2': "snd(l!0) = Normal s1 ∧ s1∈q" and
a3: "j≥i ∧ Suc j<length l" and
a4: "env_tran_right Γ l rely ∧ Sta q rely"
shows "Γ⊢⇩c(l!j) →⇩e (l!(Suc j))"
proof -
have ce_tran:"Γ⊢⇩c(l!j) →⇩c⇩e (l!(Suc j))" using a1 a2 a3 a4 cptn_tran_ce_i by auto
then have "Γ⊢⇩c(l!j) →⇩e (l!(Suc j)) ∨ Γ⊢⇩c(l!j) → (l!(Suc j))"
using step_ce_elim_cases by blast
thus ?thesis
proof
assume "Γ⊢⇩c(l!j) →⇩e (l!(Suc j))" then show ?thesis by auto
next
assume a01:"Γ⊢⇩c(l!j) → (l!(Suc j))"
then have "¬ (Γ⊢⇩c(l!j) → (l!(Suc j)))"
using a1 a2 a3 a4 a2' only_one_component_tran_j [OF a1]
by blast
then show ?thesis using a01 ce_tran by (simp add: step_ce.simps)
qed
qed
subsection {*Skip Sound*}
lemma stable_q_r_q:
assumes a0:"Sta q R" and
a1: "snd(l!i) ∈ Normal ` q" and
a2:"(snd(l!i), snd(l!(Suc i))) ∈ R"
shows "snd(l!(Suc i)) ∈ Normal ` q"
using a0 a1 a2
unfolding Sta_def by fastforce
lemma stability:
assumes a0:"Sta q R" and
a1: "snd(l!j) ∈ Normal ` q" and
a2: "j≤k ∧ k < (length l)" and
a3: "n=k-j" and
a4: "∀i. j≤i ∧ i < k ⟶ Γ⊢⇩c(l!i) →⇩e (l!(Suc i))" and
a5:"env_tran_right Γ l R "
shows "snd (l!k) ∈ Normal ` q ∧ fst (l!j) = fst (l!k)"
using a0 a1 a2 a3 a4 a5
proof (induct n arbitrary: j k)
case 0
thus ?case by auto
next
case (Suc n)
then have "length l > j + 1" by arith
moreover then have "k-1< length l" using Suc by fastforce
moreover then have "(k - 1) - j = n" using Suc by fastforce
moreover then have "j≤k-1" using Suc by arith
moreover have "∀i. j ≤ i ∧ i < k-1 ⟶ Γ⊢⇩c (l ! i) →⇩e (l ! Suc i)"
using Suc by fastforce
ultimately have induct:"snd (l! (k-1)) ∈ Normal ` q ∧ fst (l!j) = fst (l!(k-1))" using Suc
by blast
also have j_1:"k-1+1=k" using Cons Suc.prems(4) by auto
have f1:"∀i. j≤i ∧ i < k ⟶ (snd((snd (Γ,l))!i), snd((snd (Γ,l))!(Suc i))) ∈ R"
using Suc unfolding env_tran_right_def by fastforce
have k1:"k - 1 < k"
by (metis (no_types) Suc_eq_plus1 j_1 lessI)
then have "(snd((snd (Γ,l))!(k-1)), snd((snd (Γ,l))!(Suc (k-1)))) ∈ R"
using `j ≤ k - 1` f1 by blast
ultimately have "snd (l!k) ∈ Normal ` q" using stable_q_r_q Suc(2) Suc(5) by fastforce
also have "fst (l!j) = fst (l!k)"
proof -
have "Γ⊢⇩c (l ! (k-1)) →⇩e (l ! k)" using Suc(6) k1 `j≤k-1` by fastforce
thus ?thesis using k1 prod.collapse env_c_c' induct by metis
qed
ultimately show ?case by meson
qed
lemma stable_only_env_i_j:
assumes a0:"Sta q R" and
a1: "snd(l!i) ∈ Normal ` q" and
a2: "i<j ∧ j < (length l)" and
a3: "n=j-i-1" and
a4: "∀k≥i. k < j ⟶ Γ⊢⇩c(l!k) →⇩e (l!(Suc k))" and
a5: "env_tran_right Γ l R"
shows "snd (l!j) ∈ Normal ` q"
using a0 a1 a2 a3 a4 a5 by (meson less_imp_le_nat stability)
lemma stable_only_env_1:
assumes a0:"Sta q R" and
a1: "snd(l!i) ∈ Normal ` q" and
a2: "i<j ∧ j < (length l)" and
a3: "n=j-i-1" and
a4: "∀i. Suc i < length l ⟶ Γ⊢⇩c(l!i) →⇩e (l!(Suc i))" and
a5: "env_tran_right Γ l R"
shows "snd (l!j) ∈ Normal ` q"
using a0 a1 a2 a3 a4 a5
by (meson stable_only_env_i_j less_trans_Suc)
lemma stable_only_env_q:
assumes a0:"Sta q R" and
a1: "∀i. Suc i < length l ⟶ Γ⊢⇩c(l!i) →⇩e (l!(Suc i))" and
a2: "env_tran Γ q l R"
shows "∀i. i < length l ⟶ snd (l!i) ∈ Normal ` q"
proof (cases "0 < length l")
case False thus ?thesis using a2 unfolding env_tran_def by fastforce
next
case True
thus ?thesis
proof - {
fix i
assume aa1:"i < length l"
have post_0:"snd (l ! 0) ∈ Normal ` q "
using a2 unfolding env_tran_def by auto
then have "snd (l ! i) ∈ Normal ` q"
proof (cases i)
case 0 thus ?thesis using post_0 by auto
next
case (Suc n)
have "env_tran_right Γ l R"
using a2 env_tran_right_def unfolding env_tran_def by auto
also have "0<i" using Suc by auto
ultimately show ?thesis
using post_0 stable_only_env_1 a0 a1 a2 aa1 by blast
qed
} then show ?thesis by auto qed
qed
lemma Skip_sound1:
assumes a0:"Sta q R" and
a1:"(∀s. (Normal s, Normal s) ∈ G)" and
a10:"c ∈ cp Γ Skip s" and
a11:"c ∈ assum(q, R)"
shows "c ∈ comm (G, (q,a)) F"
proof -
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
{
assume "snd (last l) ∉ Fault ` F"
have cp:"l!0=(Skip,s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1" using a10 cp_def c_prod by fastforce
have assum:"snd(l!0) ∈ Normal ` q ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ R)"
using a11 c_prod unfolding assum_def by simp
have concl:"(∀i. Suc i<length l ⟶
Γ1⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof -
{ fix i
assume asuc:"Suc i<length l"
then have "¬ (Γ1⊢⇩c(l!i) → (l!(Suc i)))"
by (metis Suc_lessD cp prod.collapse prod.sel(1) stepc_elim_cases(1) zero_skip_all_skip)
} thus ?thesis by auto qed
have concr:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof-
{
assume valid:"final (last l)"
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a l1 where l:"l=a#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a l1] l by fastforce
then have fst_last_skip:"fst (last l) = Skip"
by (metis `0 < length l` cp diff_less fst_conv zero_less_one zero_skip_all_skip)
have last_q: "snd (last l) ∈ Normal ` q"
proof -
have env: "env_tran Γ q l R" using env_tran_def assum cp by blast
have env_right:"env_tran_right Γ l R " using a0 env_tran_right_def assum cp by metis
also obtain s1 where "snd(l!0) = Normal s1 ∧ s1∈q"
using assum by auto
ultimately have all_tran_env: "∀i. Suc i < length l ⟶ Γ⊢⇩c(l!i) →⇩e (l!(Suc i))"
using final_always_env_i cp zero_final_always_env_0 a0
by fastforce
then have "∀i. i < length l ⟶ snd (l!i) ∈ Normal ` q"
using stable_only_env_q a0 env by fastforce
thus ?thesis using last_l using len_l by fastforce
qed
note res = conjI [OF fst_last_skip last_q]
} thus ?thesis by auto
qed
note res = conjI [OF concl concr]
} thus ?thesis using c_prod unfolding comm_def by auto
qed
lemma Skip_sound:
"Sta q R ⟹
(∀s. (Normal s, Normal s) ∈ G) ⟹
Γ,Θ ⊨n⇘/F⇙ Skip sat [q,R, G, q,a]"
proof-
assume
a0:"Sta q R" and
a1:"(∀s. (Normal s, Normal s) ∈ G)"
{
fix s
have ass:"cpn n Γ Skip s ∩ assum(q, R) ⊆ comm(G, (q,a)) F"
proof-
{ fix c
assume a10:"c ∈ cpn n Γ Skip s" and a11:"c ∈ assum(q, R)"
then have a10:"c∈cp Γ Skip s"
using cp_def cpn_def cptn_if_cptn_mod cptn_mod_nest_cptn_mod by blast
have "c∈comm(G, (q,a)) F" using Skip_sound1[OF a0 a1 a10 a11] by auto
} thus ?thesis by auto
qed
}
thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma Throw_sound1:
assumes a1:"Sta a R" and
a2:"(∀s. (Normal s, Normal s) ∈ G)" and
a10:"c ∈ cp Γ Throw s" and
a11:"c ∈ assum(a, R)"
shows "c ∈ comm (G, (q,a)) F"
proof -
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
{
assume "snd (last l) ∉ Fault ` F"
have cp:"l!0=(Throw,s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1" using a10 cp_def c_prod by fastforce
have assum:"snd(l!0) ∈ Normal ` (a) ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ (R))"
using a11 c_prod unfolding assum_def by simp
then have env_tran:"env_tran_right Γ l R" using cp env_tran_right_def by auto
obtain a1 where a_normal:"snd(l!0) = Normal a1 ∧ a1 ∈ a"
using assum by auto
have concl:"(∀i ns ns'. Suc i<length l ⟶
Γ1⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ (G))"
proof -
{ fix i
assume asuc:"Suc i<length l"
then have asuci:"i<length l" by fastforce
then have "fst (l ! 0) = LanguageCon.com.Throw" using cp by auto
moreover obtain s1 where "snd (l ! 0) = Normal s1" using assum by auto
ultimately have "fst (l ! i) = Throw ∧ (∃s2. snd (l ! i) = Normal s2)"
using cp a1 assum a_normal env_tran asuci zero_throw_all_throw
by fastforce
then have "¬ (Γ1⊢⇩c(l!i) → (l!(Suc i)))"
by (meson SmallStepCon.final_def SmallStepCon.no_step_final')
} thus ?thesis by auto qed
have concr:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof-
{
assume valid:"final (last l)"
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a1 l1 where l:"l=a1#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a1 l1] l by fastforce
then have fst_last_skip:"fst (last l) = Throw"
by (metis a1 a_normal cp diff_less env_tran fst_conv len_l zero_less_one zero_throw_all_throw)
have last_q: "snd (last l) ∈ Normal ` (a)"
proof -
have env: "env_tran Γ a l R" using env_tran_def assum cp by blast
have env_right:"env_tran_right Γ l R" using env_tran_right_def assum cp by metis
then have all_tran_env: "∀i. Suc i < length l ⟶ Γ⊢⇩c(l!i) →⇩e (l!(Suc i))"
using final_always_env_i a1 assum cp zero_final_always_env_0 by fastforce
then have "∀i. i < length l ⟶ snd (l!i) ∈ Normal ` (a)"
using stable_only_env_q a1 env by fastforce
thus ?thesis using last_l using len_l by fastforce
qed
note res = conjI [OF fst_last_skip last_q]
} thus ?thesis by auto qed
note res = conjI [OF concl concr]
}
thus ?thesis using c_prod unfolding comm_def by auto
qed
lemma Throw_sound:
"Sta a R ⟹
(∀s. (Normal s, Normal s) ∈ G) ⟹
Γ,Θ ⊨n⇘/F⇙ Throw sat [a, R, G, q,a]"
proof -
assume
a1:"Sta a R" and
a2:" (∀s. (Normal s, Normal s) ∈ G)" {
fix s
have ass:"cpn n Γ Throw s ∩ assum(a, R) ⊆ comm(G, (q,a)) F"
proof-
{ fix c
assume a10:"c ∈ cpn n Γ Throw s" and a11:"c ∈ assum(a, R)"
then have a10:"c∈cp Γ Throw s"
using cp_def cpn_def cptn_if_cptn_mod cptn_mod_nest_cptn_mod by blast
have "c∈comm(G, (q,a)) F" using Throw_sound1[OF a1 a2 a10 a11] by auto
} thus ?thesis by auto
qed
}
thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma no_comp_tran_before_i_0_g:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!0) = c " and
a2: "Suc i<length l ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))" and
a3: "j < i ∧ (Γ⊢⇩c(l!j) → (l!(Suc j)))" and
a4: "∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))" and
a5: "∀s1 s2 c1. Γ⊢⇩c(c, s1) → ((c1,s2)) ⟶
(c1=Skip) ∨ (c1=Throw ∧ (∃s21. s2 = Normal s21))" and
a6: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc j) ∈ Normal ` q"
shows "P"
proof -
have "Suc j < length l" using a0 a1 a2 a3 a4 by fastforce
then have "fst (l!j) = c"
using a0 a1 a2 a3 a4 cptn_env_same_prog[of Γ l j] by fastforce
then obtain s s1 c1 where l_0: "l!j = (c, s) ∧ l!(Suc j) = (c1,s1)"
by (metis (no_types) prod.collapse)
moreover have "snd (l!j) ∈ Normal ` p" using a4 stability[of p rely l 0 j j] a6 a3 a2
proof -
have "∀B r ps n na nb f. ¬ Sta B r ∨ snd (ps ! n) ∉ Normal ` B ∨ ¬ n ≤ na ∨ ¬ na < length ps ∨ na - n ≠ nb ∨ (∃nb≥n. nb < na ∧ ¬ f⊢⇩c ps ! nb →⇩e ps ! Suc nb) ∨ ¬ env_tran_right f ps r ∨ snd (ps ! na) ∈ Normal ` B ∧ (fst (ps ! n)::('b, 'a, 'c,'d) LanguageCon.com) = fst (ps ! na)"
using stability by blast
then show ?thesis
using Suc_lessD ‹Suc j < length l› a4 a6 by blast
qed
then have suc_0_skip: "(fst (l!Suc j) = Skip ∨ fst (l!Suc j) = Throw) ∧
(∃s2. snd(l!Suc j) = Normal s2 ) "
using a5 a6 a3 SmallStepCon.step_Stuck_prop using fst_conv imageE l_0 snd_conv by auto
thus ?thesis using only_one_component_tran_j
proof -
have "∀n na. ¬ n < na ∨ Suc n ≤ na"
using Suc_leI by satx
thus ?thesis using only_one_component_tran_j[OF a0] suc_0_skip a6 a0 a2 a3
using imageE by blast
qed
qed
lemma no_comp_tran_before_i:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = c" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))" and
a3: "k≤j ∧ j < i ∧ (Γ⊢⇩c(l!j) → (l!(Suc j)))" and
a4: "∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))" and
a5: "∀s1 s2 c1. Γ⊢⇩c(c, s1) → ((c1,s2)) ⟶
(c1=Skip) ∨ (c1=Throw ∧ (∃s21. s2 = Normal s21))" and
a6: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc j) ∈ Normal ` q"
shows "P"
using a0 a1 a2 a3 a4 a5 a6
proof (induct k arbitrary: l i j)
case 0 thus ?thesis using no_comp_tran_before_i_0_g by blast
next
case (Suc n)
then obtain a1 l1 where l: "l=a1#l1"
by (metis less_nat_zero_code list.exhaust list.size(3))
then have l1notempty:"l1≠[]" using Suc by force
then obtain i' where i': "i=Suc i'" using Suc
using less_imp_Suc_add by blast
then obtain j' where j': "j=Suc j'" using Suc
using Suc_le_D by blast
have "(Γ,l1)∈cptn" using Suc l
using tl_in_cptn l1notempty by blast
moreover have "fst (l1 ! n) = c"
using Suc l l1notempty by force
moreover have "Suc i' < length l1 ∧ n ≤ i' ∧ Γ⊢⇩c l1 ! i' → (l1 ! Suc i')"
using Suc l l1notempty i' by auto
moreover have "n ≤ j' ∧ j' < i' ∧ Γ⊢⇩c l1 ! j' → (l1 ! Suc j')"
using Suc l l1notempty i' j' by auto
moreover have "∀k<j'. Γ⊢⇩c l1 ! k →⇩e (l1 ! Suc k)"
using Suc l l1notempty j' by auto
moreover have "env_tran_right Γ l1 rely ∧ Sta q rely ∧ Sta p rely ∧ snd (l1!0) ∈ Normal ` p ∧
Sta q rely ∧ snd (l1!Suc j') ∈ Normal ` q"
proof -
have suc0:"Suc 0 < length l" using Suc by auto
have "j>0" using j' by auto
then have "Γ⊢⇩c(l!0) →⇩e (l!(Suc 0))" using Suc(6) by blast
then have "(snd(l!Suc 0) ∈ Normal ` p)"
using Suc(8) suc0 unfolding Sta_def env_tran_right_def by blast
also have "snd (l!Suc j) ∈ Normal ` q" using Suc(8) by auto
ultimately show ?thesis using Suc(8) l by (metis env_tran_tail j' nth_Cons_Suc)
qed
ultimately show ?case using Suc(1)[of l1 i' j' ] Suc(7) Suc(8) j' l by auto
qed
lemma exists_first_occ: "P (n::nat) ⟹ ∃m. P m ∧ (∀i<m. ¬ P i)"
proof (induct n)
case 0 thus ?case by auto
next
case (Suc n) thus ?case
by (metis ex_least_nat_le not_less0)
qed
lemma exist_first_comp_tran':
assumes a1: "Suc i<length l ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))"
shows "∃j. (Suc j<length l ∧ (Γ⊢⇩c(l!j) → (l!(Suc j)))) ∧ (∀k < j. ¬Γ⊢⇩c(l!k) → (l!(Suc k)))"
proof -
let ?P = "(λn. Suc n<length l ∧ (Γ⊢⇩c(l!n) → (l!(Suc n))))"
show ?thesis using exists_first_occ[of ?P i] a1 by auto
qed
lemma exist_first_comp_tran:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "Suc i<length l ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))"
shows "∃j. j≤i ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ (∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
proof -
obtain j where pj:"(Suc j<length l ∧ (Γ⊢⇩c(l!j) → (l!(Suc j)))) ∧
(∀k < j. ¬(Suc k<length l ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))))"
using a1 exist_first_comp_tran' by blast
then have "j≤i" using a1 pj by (cases "j≤i", auto)
moreover have "Γ⊢⇩c(l!j) → (l!(Suc j))" using pj by auto
moreover have "(∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
proof -
{fix k
assume kj:"k<j"
then have "Suc k ≥ length l ∨ ¬ ( (Γ⊢⇩c(l!k) → (l!(Suc k)))) " using pj by auto
then have "(Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))"
proof
{assume "length l ≤ Suc k"
thus ?thesis using kj pj by auto
}
{assume "¬ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
also have "k + 1 < length l" using kj pj by auto
ultimately show ?thesis
using a0 cptn_tran_ce_i step_ce_elim_cases by blast
}
qed
} thus ?thesis by auto
qed
ultimately show ?thesis by auto
qed
lemma skip_com_all_skip:
assumes a0:"(Γ, l) ∈ cptn" and
a1:"fst (l!i) = Skip" and
a2:"i<length l"
shows "∀j. j≥i ∧ j <length l ⟶ fst (l!j) = Skip"
using a0 a1 a2
proof (induct "length l - (i + 1)" arbitrary: i)
case 0 thus ?case by (metis Suc_eq_plus1 Suc_leI diff_is_0_eq nat_less_le zero_less_diff)
next
case (Suc n)
then have l:"Suc i < length l" by arith
have n:"n = (length l) - (Suc i + 1)" using Suc by arith
then have "Γ⊢⇩c l ! i →⇩c⇩e l ! Suc i" using cptn_tran_ce_i Suc
by (metis (no_types) Suc.hyps(2) a0 cptn_tran_ce_i zero_less_Suc zero_less_diff)
then have "Γ⊢⇩c l ! i → l ! Suc i ∨ Γ⊢⇩c l ! i →⇩e l ! Suc i"
using step_ce_elim_cases by blast
then have or:"fst(l!Suc i) = Skip"
proof
{assume "Γ⊢⇩c l ! i →⇩e l ! Suc i"
thus ?thesis using Suc(4) by (metis env_c_c' prod.collapse)
}
next
{assume step:"Γ⊢⇩c l ! i → l ! Suc i"
{assume "fst(l!i) = Skip"
then have ?thesis using step
using SmallStepCon.final_def SmallStepCon.no_step_final' by blast
}note left = this
{assume "fst(l!i) = Throw"
then have ?thesis using step stepc_elim_cases
proof -
have "∃x. l ! Suc i = (LanguageCon.com.Skip, x)"
by (metis (no_types) ‹fst (l ! i) = LanguageCon.com.Throw› local.step stepc_elim_cases(11) surjective_pairing)
then show ?thesis
by fastforce
qed
} then show ?thesis using Suc(4) left by auto
}
qed
show ?case using Suc(1)[OF n a0 or l] Suc(4) Suc(5) by (metis le_less_Suc_eq not_le)
qed
lemma terminal_com_all_term:
assumes a0:"(Γ, l) ∈ cptn" and
a1:"fst (l!i) = Skip ∨ fst (l!i) = Throw" and
a2:"i<length l"
shows "∀j. j≥i ∧ j <length l ⟶ fst (l!j) = Skip ∨ fst (l!j) = Throw"
using a0 a1 a2
proof (induct "length l - (i + 1)" arbitrary: i)
case 0 thus ?case by (metis Suc_eq_plus1 Suc_leI diff_is_0_eq nat_less_le zero_less_diff)
next
case (Suc n)
then have l:"Suc i < length l" by arith
have n:"n = (length l) - (Suc i + 1)" using Suc by arith
then have "Γ⊢⇩c l ! i →⇩c⇩e l ! Suc i" using cptn_tran_ce_i Suc
by (metis (no_types) Suc.hyps(2) a0 cptn_tran_ce_i zero_less_Suc zero_less_diff)
then have "Γ⊢⇩c l ! i → l ! Suc i ∨ Γ⊢⇩c l ! i →⇩e l ! Suc i"
using step_ce_elim_cases by blast
then have or:"fst(l!Suc i) = Skip ∨ fst(l!Suc i) = Throw"
proof
{assume "Γ⊢⇩c l ! i →⇩e l ! Suc i"
thus ?thesis using Suc(4) by (metis env_c_c' prod.collapse)
}
next
{assume step:"Γ⊢⇩c l ! i → l ! Suc i"
{assume "fst(l!i) = Skip"
then have ?thesis using step
using SmallStepCon.final_def SmallStepCon.no_step_final' by blast
}note left = this
{assume "fst(l!i) = Throw"
then have ?thesis using step stepc_elim_cases
proof -
have "∃x. l ! Suc i = (LanguageCon.com.Skip, x)"
by (metis (no_types) ‹fst (l ! i) = LanguageCon.com.Throw› local.step stepc_elim_cases(11) surjective_pairing)
then show ?thesis
by fastforce
qed
} then show ?thesis using Suc(4) left by auto
}
qed
show ?case using Suc(1)[OF n a0 or l] Suc(4) Suc(5) by (metis le_less_Suc_eq not_le)
qed
lemma only_one_c_comp_tran:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!0) = c" and
a2: "Suc i<length l ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))" and
a3: "i < j ∧ Suc j < length l ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ fst (l!j) = c" and
a4: "∀s1 s2 c1. Γ⊢⇩c(c, s1) → ((c1,s2)) ⟶
((c1=Skip) ∨ (c1=Throw)) " and
a5: "(∀k < i. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
shows "P"
proof -
have fst:"fst (l!i) = c" using a0 a1 a5
by (simp add: a2 cptn_env_same_prog)
then have suci:"fst (l!Suc i) = Skip ∨ fst (l!Suc i) = Throw"
using a4 by (metis a2 surjective_pairing)
then have "fst (l!j) = Skip ∨ fst (l!j) = Throw"
proof -
have "Suc i ≤ j"
using Suc_leI a3 by presburger
then show ?thesis
using Suc_lessD terminal_com_all_term[OF a0 suci] a2 a3 by blast
qed
thus ?thesis
proof
{assume "fst (l ! j) = Skip"
then show ?thesis using a3 SmallStepCon.final_def SmallStepCon.no_step_final' by blast
}
next
{assume asm:"fst (l ! j) = Throw"
then show ?thesis
proof (cases "snd (l!i)")
case Normal
thus ?thesis using a3 a2 fst asm
by (metis SmallStepCon.final_def SmallStepCon.no_step_final')
next
case Abrupt thus ?thesis using a3 a2 fst asm skip_com_all_skip
suci by (metis Suc_leI Suc_lessD a0 mod_env_not_component prod.collapse)
next
case Fault thus ?thesis using a3 a2 fst asm skip_com_all_skip
suci by (metis Suc_leI Suc_lessD a0 mod_env_not_component prod.collapse)
next
case Stuck thus ?thesis using a3 a2 fst asm skip_com_all_skip
suci by (metis Suc_leI Suc_lessD a0 mod_env_not_component prod.collapse)
qed
}
qed
qed
lemma only_one_component_tran1:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!0) = c" and
a2: "Suc i<length l ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))" and
a3: "j ≠ i ∧ Suc j < length l ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ fst (l!j) = c" and
a4: "∀s1 s2 c1. Γ⊢⇩c(c, s1) → ((c1,s2)) ⟶
((c1=Skip) ∨ (c1=Throw)) " and
a5: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc j) ∈ Normal ` q "
shows "P"
proof (cases "j=i")
case True thus ?thesis using a3 by auto
next
case False note j_neq_i=this
thus ?thesis
proof (cases "j<i")
case True
thus ?thesis
proof -
obtain bb :: "'b set ⇒ ('b ⇒ ('b, 'c) xstate) ⇒ ('b, 'c) xstate ⇒ 'b" where
"∀x0 x1 x2. (∃v3. x2 = x1 v3 ∧ v3 ∈ x0) = (x2 = x1 (bb x0 x1 x2) ∧ bb x0 x1 x2 ∈ x0)"
by moura
then have f1: "∀x f B. x ∉ f ` B ∨ x = f (bb B f x) ∧ bb B f x ∈ B"
by (meson imageE)
then have "Γ⊢⇩c (c, snd (l ! j)) → (fst (l ! Suc j), Normal (bb q Normal (snd (l ! Suc j))))"
by (metis (no_types) a3 a5 surjective_pairing)
then show ?thesis
using f1 by (meson Suc_leI a0 a2 a4 a5 True only_one_component_tran_j)
qed
next
case False
obtain j1
where all_ev:"j1≤i ∧
(Γ⊢⇩c(l!j1) → (l!(Suc j1))) ∧
(∀k < j1. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
using a0 a2 a3 exist_first_comp_tran by blast
then have fst:"fst (l!j1) = c"
using a0 a1 a2 cptn_env_same_prog le_imp_less_Suc less_trans_Suc by blast
have suc:"Suc j1 < length l ∧ Γ⊢⇩c l ! j1 → l ! Suc j1" using all_ev a2
using Suc_lessD le_eq_less_or_eq less_trans_Suc by linarith
have evs:"(∀k < j1. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))" using all_ev by auto
have j:"j1 < j ∧ Suc j < length l ∧ Γ⊢⇩c l ! j → l ! Suc j ∧ fst (l ! j) = c"
using a3 all_ev False by auto
then show ?thesis
using only_one_c_comp_tran[OF a0 a1 suc j a4 evs] by auto
qed
qed
lemma only_one_component_tran_i:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = c" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))" and
a3: "k≤j ∧ j ≠ i ∧ Suc j < length l ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ fst (l!j) = c" and
a4: "∀s1 s2 c1. Γ⊢⇩c(c, s1) → ((c1,s2)) ⟶
((c1=Skip) ∨ (c1=Throw)) " and
a5: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc j) ∈ Normal ` q "
shows "P"
using a0 a1 a2 a3 a4 a5
proof (induct k arbitrary: l i j p q)
case 0 show ?thesis using only_one_component_tran1[OF 0(1) 0(2) ] 0 by blast
next
case (Suc n)
then obtain a1 l1 where l: "l=a1#l1"
by (metis less_nat_zero_code list.exhaust list.size(3))
then have l1notempty:"l1≠[]" using Suc by force
then obtain i' where i': "i=Suc i'" using Suc
using less_imp_Suc_add using Suc_le_D by meson
then obtain j' where j': "j=Suc j'" using Suc
using Suc_le_D by meson
have a0:"(Γ,l1)∈cptn" using Suc l
using tl_in_cptn l1notempty by meson
moreover have a1:"fst (l1 ! n) = c"
using Suc l l1notempty by force
moreover have a2:"Suc i' < length l1 ∧ n ≤ i' ∧ Γ⊢⇩c l1 ! i' → (l1 ! Suc i')"
using Suc l l1notempty i' by auto
moreover have a3:"n ≤ j' ∧ j' ≠ i' ∧ Suc j' < length l1 ∧ Γ⊢⇩c l1 ! j' → (l1 ! Suc j') ∧ fst (l1!j') = c"
using Suc l l1notempty i' j' by auto
moreover have a4:"env_tran_right Γ l1 rely ∧
Sta p rely ∧ snd (l1!n) ∈ Normal ` p ∧
Sta q rely ∧ snd (l1 ! Suc j') ∈ Normal ` q "
using Suc(7) l j' unfolding env_tran_right_def by fastforce
show ?case using Suc(1)[OF a0 a1 a2 a3 Suc(6) a4] by auto
qed
lemma only_one_component_tran:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = c" and
a2: "k≤i ∧ i ≠ j ∧ Suc i<length l ∧ (Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = c" and
a3: "k≤j ∧ Suc j < length l" and
a4: "∀s1 s2 c1. Γ⊢⇩c(c,s1) → ((c1,s2)) ⟶
((c1=Skip) ∨ (c1=Throw))" and
a5: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q "
shows "(Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using a0 a1 a2 a3 a4 a5 only_one_component_tran_i
proof -
{assume "(Γ⊢⇩c(l!j) → (l!(Suc j))) ∨ (¬ Γ⊢⇩c(l!j) → (l!(Suc j)))"
then have "(Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
proof
assume "Γ⊢⇩c l ! j → (l ! Suc j)"
then have j:"Suc j<length l ∧ k≤j ∧ (Γ⊢⇩c(l!j) → (l!(Suc j)))" using a3 by auto
show ?thesis using only_one_component_tran_i[OF a0 a1 j a2 a4 a5]
by blast
next
assume "¬ Γ⊢⇩c l ! j → (l ! Suc j)"
thus ?thesis
by (metis Suc_eq_plus1 a0 a3 cptn_tran_ce_i step_ce_elim_cases)
qed
} thus ?thesis by auto
qed
lemma only_one_component_tran_all_env:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = c" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = c" and
a3: "∀s1 s2 c1. Γ⊢⇩c(c,s1) → ((c1,s2)) ⟶
((c1=Skip) ∨ (c1=Throw))" and
a4: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q "
shows "∀j. k≤j ∧ j≠i ∧ Suc j < (length l) ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
proof -
{fix j
assume ass:"k≤j ∧ j≠i ∧ Suc j < (length l)"
then have a2:"k ≤ i ∧ i ≠ j ∧ Suc i < length l ∧ Γ⊢⇩c l ! i → l ! Suc i ∧ fst (l ! i) = c"
using a2 by auto
then have "(Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using only_one_component_tran[OF a0 a1 ] a2 a3 ass a4 by blast
} thus ?thesis by auto
qed
lemma only_one_component_tran_all_not_comp:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = c" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = c" and
a3: "∀s1 s2 c1. Γ⊢⇩c(c, s1) → ((c1,s2)) ⟶
((c1=Skip) ∨ (c1=Throw))" and
a4: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q "
shows "∀j. k≤j ∧ j≠i ∧ Suc j < (length l) ⟶ ¬(Γ⊢⇩c(l!j) → (l!(Suc j)))"
proof -
{fix j
assume ass:"k≤j ∧ j≠i ∧ Suc j < (length l)"
then have "¬(Γ⊢⇩c(l!j) → (l!(Suc j)))"
using a0 a1 a2 a3 a4 only_one_component_tran_i ass by blast
} thus ?thesis by auto
qed
lemma final_exist_component_tran1:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!i) = c" and
a2: "env_tran Γ q l R ∧ Sta q R" and
a3: "i≤j ∧ j < length l ∧ final (l!j)" and
a5: "c≠Skip ∧ c≠Throw"
shows "∃k. k≥i ∧ k<j ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
proof -
{assume "∀k. k≥i ∧ k<j ⟶ ¬(Γ⊢⇩c(l!k) → (l!(Suc k)))"
then have "∀k. k≥i ∧ k<j ⟶ (Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))"
by (metis (no_types, lifting) Suc_eq_plus1 a0 a3 cptn_tran_ce_i less_trans_Suc step_ce_elim_cases)
then have "fst (l!j) = fst (l!i)" using cptn_i_env_same_prog a0 a3 by blast
then have False using a3 a1 a5 unfolding final_def by auto
}
thus ?thesis by auto
qed
lemma final_exist_component_tran:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!i) = c" and
a2: "i≤j ∧ j < length l ∧ final (l!j)" and
a3: "c≠Skip ∧ c≠Throw"
shows "∃k. k≥i ∧ k<j ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
proof -
{assume "∀k. k≥i ∧ k<j ⟶ ¬(Γ⊢⇩c(l!k) → (l!(Suc k)))"
then have "∀k. k≥i ∧ k<j ⟶ (Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))"
by (metis (no_types, lifting) Suc_eq_plus1 a0 a2 cptn_tran_ce_i less_trans_Suc step_ce_elim_cases)
then have "fst (l!j) = fst (l!i)" using cptn_i_env_same_prog a0 a2 by blast
then have False using a2 a1 a3 unfolding final_def by auto
}
thus ?thesis by auto
qed
lemma suc_not_final_final_c_tran:
assumes a0: "(Γ, l) ∈ cptn " and
a1: "Suc j< length l ∧ ¬final (l!j) ∧ final (l!Suc j)"
shows "(Γ⊢⇩c(l!j) → (l!(Suc j)))"
proof -
obtain x xs where l:"l = x#xs" using a0 cptn.simps by blast
obtain c1 s1 c2 s2 where l1:"l!j = (c1,s1) ∧ l!(Suc j) = (c2,s2)" using a1 by fastforce
have "¬ Γ⊢⇩c(l!j) →⇩e (l!(Suc j))"
proof -
{ assume a:"Γ⊢⇩c(l!j) →⇩e (l!(Suc j))"
then have eq_fst:"fst (l!j) = fst (l!Suc j)" by (metis env_c_c' prod.collapse)
{ assume "fst (l!Suc j) = Skip"
then have "False" using a1 eq_fst unfolding final_def by fastforce
}note p1=this
{ assume "fst (l!Suc j) = Throw ∧ (∃s. snd (l!Suc j) = Normal s)"
then have "False" using a1 eq_fst unfolding final_def
by (metis a eenv_normal_s'_normal_s local.l1 snd_conv)
}
then have False using a1 p1 unfolding final_def by fastforce
} thus ?thesis by auto
qed
also have "Γ⊢⇩c(l!j) →⇩c⇩e (l!(Suc j))" using l cptn_stepc_rtran a0 a1 by fastforce
ultimately show ?thesis using step_ce_not_step_e_step_c local.l1 by fastforce
qed
lemma final_exist_component_tran_final:
assumes a0:"(Γ, l) ∈ cptn" and
a2: "i≤j ∧ j < length l ∧ final (l!j)" and
a3: "¬final(l!i)"
shows "∃k. k≥i ∧ k<j ∧ (Γ⊢⇩c(l!k) → (l!(Suc k))) ∧ final(l!(Suc k))"
proof -
let ?P = "λj. i≤j ∧ j < length l ∧ final (l!j)"
obtain k where k:"?P k ∧ (∀i<k. ¬ ?P i)" using a2 exists_first_occ[of ?P j] by auto
then have i_k_not_final:"∀i'<k. i'≥i ⟶ ¬final (l!i')" using a2 by fastforce
have i_eq_j:"i<j" using a2 a3 using le_imp_less_or_eq by auto
then obtain pre_k where pre_k:"Suc pre_k = k" using a2 k
by (metis a3 eq_iff le0 lessE neq0_conv)
then have "Γ⊢⇩c(l!pre_k) → (l!k)"
proof -
have "pre_k ≥i" using pre_k i_eq_j using a3 k le_Suc_eq by blast
then have "¬(final (l!pre_k))" using i_k_not_final pre_k by auto
thus ?thesis using suc_not_final_final_c_tran a0 a2 pre_k k by fastforce
qed
thus ?thesis using pre_k by (metis a2 a3 i_k_not_final k le_Suc_eq not_less_eq)
qed
subsection {* Basic Sound *}
lemma basic_skip:
"∀s1 s2 c1. Γ⊢⇩c(Basic f e, s1) → ((c1,s2)) ⟶ c1=Skip"
proof -
{fix s1 s2 c1
assume "Γ⊢⇩c(Basic f e,s1) → ((c1,s2))"
then have "c1=Skip" using stepc_elim_cases(3) by blast
} thus ?thesis by auto
qed
lemma no_comp_tran_before_i_basic:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Basic f e" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))" and
a3: "k≤j ∧ j < i ∧ (Γ⊢⇩c(l!j) → (l!(Suc j)))" and
a4: "∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))" and
a5: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc j) ∈ Normal ` q"
shows "P"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Basic f e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using basic_skip by fastforce
thus ?thesis using a0 a1 a2 a3 a4 a5 no_comp_tran_before_i by blast
qed
lemma only_one_component_tran_i_basic:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Basic f e" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))" and
a3: "k≤j ∧ j ≠ i ∧ Suc j < length l ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ fst (l!j) = Basic f e" and
a4: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc j) ∈ Normal ` q"
shows "P"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Basic f e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using basic_skip by blast
thus ?thesis using a0 a1 a2 a3 a4 only_one_component_tran_i[OF a0 a1 a2 ] by blast
qed
lemma only_one_component_tran_basic:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Basic f e" and
a2: " k≤i ∧ i ≠ j ∧ Suc i<length l ∧ (Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = Basic f e" and
a3: "k≤j ∧ Suc j < length l" and
a4: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q"
shows "(Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Basic f e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using basic_skip by blast
thus ?thesis using a0 a1 a2 a3 a4 only_one_component_tran by blast
qed
lemma only_one_component_tran_all_env_basic:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Basic f e" and
a2: "k≤i ∧ Suc i<length l ∧ (Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = Basic f e" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q"
shows "∀j. k≤j ∧ j≠i ∧ Suc j < (length l) ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
proof -
have b: "∀s1 s2 c1. Γ⊢⇩c(Basic f e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using basic_skip by blast
show ?thesis
by (metis (no_types) a0 a1 a2 a3 only_one_component_tran_basic)
qed
lemma only_one_component_tran_all_not_comp_basic:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Basic f e" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = Basic f e" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q"
shows "∀j. k≤j ∧ j≠i ∧ Suc j < (length l) ⟶ ¬(Γ⊢⇩c(l!j) → (l!(Suc j)))"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Basic f e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using basic_skip by blast
thus ?thesis using a0 a1 a2 a3 only_one_component_tran_all_not_comp by blast
qed
lemma one_component_tran_basic:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!0) = Basic f e" and
a2: "Suc k<length l ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely" and
a4:"p ⊆ {s. f s ∈ q}"
shows "∀j. 0≤j ∧ j≠k ∧ Suc j < (length l) ⟶ ¬(Γ⊢⇩c(l!j) → (l!(Suc j)))"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Basic f e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using basic_skip by blast
also obtain j where first:"(Suc j<length l ∧ (Γ⊢⇩c(l!j) → (l!(Suc j)))) ∧
(∀k < j. ¬((Γ⊢⇩c(l!k) → (l!(Suc k)))))"
by (metis (no_types) a2 exist_first_comp_tran')
moreover then have prg_j:"fst (l!j) = Basic f e" using a1 a0
by (metis cptn_env_same_prog not_step_comp_step_env)
moreover have sta_j:"snd (l!j) ∈ Normal ` p"
proof -
have a0':"0≤j ∧ j<(length l)" using first by auto
have a1':"(∀k. 0≤k ∧ k < j ⟶ ((Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))))"
using first not_step_comp_step_env a0 by fastforce
thus ?thesis using stability first a3 a1' a0' by blast
qed
then have "snd (l!Suc j) ∈ Normal ` q" using a4 first prg_j
proof -
obtain s where "snd (l!j) = Normal s ∧ s∈p" using sta_j by fastforce
moreover then have "fst(l!Suc j) = Skip ∧ snd(l!Suc j) = Normal (f s)" using first
by (metis fst_conv prg_j snd_conv stepc_Normal_elim_cases(3) surjective_pairing)
ultimately show ?thesis using a4 by fastforce
qed
then have "∀i. 0≤i ∧ i≠j ∧ Suc i < (length l) ⟶ ¬(Γ⊢⇩c(l!i) → (l!(Suc i)))"
using only_one_component_tran_all_not_comp_basic[OF a0 a1] first a3
a0 a1 calculation(1) only_one_component_tran1 prg_j by blast
moreover then have "k=j" using a2 by fastforce
ultimately show ?thesis by auto
qed
lemma one_component_tran_basic_env:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!0) = Basic f e" and
a2: "Suc k<length l ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely" and
a4:"p ⊆ {s. f s ∈ q}"
shows "∀j. 0≤j ∧ j≠k ∧ Suc j < (length l) ⟶ Γ⊢⇩c(l!j) →⇩e (l!(Suc j))"
proof -
have "∀j. 0≤j ∧ j≠k ∧ Suc j < (length l) ⟶ ¬ (Γ⊢⇩c(l!j) → (l!(Suc j)))"
using one_component_tran_basic[OF a0 a1 a2 a3 a4] by auto
thus ?thesis using a0
by (metis Suc_eq_plus1 cptn_tran_ce_i step_ce_elim_cases)
qed
lemma final_exist_component_tran_basic:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!i) = Basic f e" and
a2: "env_tran Γ q l R" and
a3: "i≤j ∧ j < length l ∧ final (l!j)"
shows "∃k. k≥i ∧ k<j ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
proof -
show ?thesis using a0 a1 a2 a3 final_exist_component_tran by blast
qed
lemma Basic_sound1:
assumes a0:"p ⊆ {s. f s ∈ q}" and
a1:"(∀s t. s∈p ∧ (t=f s) ⟶ (Normal s,Normal t) ∈ G)" and
a2:"Sta p R" and
a3:"Sta q R" and
a10:"c ∈ cp Γ (Basic f e) s" and a11:"c ∈ assum(p, R)"
shows "c∈comm(G, (q,a)) F"
proof -
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
have cp:"l!0=(Basic f e,s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1" using a10 cp_def c_prod by fastforce
have assum:"snd(l!0) ∈ Normal ` (p) ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ R)"
using a11 c_prod unfolding assum_def by simp
have concl:"(∀i ns ns'. Suc i<length l ⟶
Γ1⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof -
{ fix k
assume a00:"Suc k<length l" and
a11:"Γ1⊢⇩c(l!k) → (l!(Suc k))"
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a l1 where l:"l=a#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a l1] l by fastforce
have env_tran:"env_tran Γ p l R" using assum env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def a2 unfolding env_tran_def by auto
then have all_event:"∀j. 0≤j ∧ j ≠ k ∧ Suc j < length l ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using one_component_tran_basic_env[of Γ l f e k R] a0 a00 a11 a2 a3 assum cp
env_tran_right fst_conv
by metis
then have before_k_all_evn:"∀j. 0≤j ∧ j < k ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using a00 a11 by fastforce
then have k_basic:"fst(l!k) = Basic f e ∧ snd (l!k) ∈ Normal ` (p)"
using cp env_tran_right a2 assum a00 a11 stability[of p R l 0 k k Γ]
by force
have suc_k_skip_q:"fst(l!Suc k) = Skip ∧ snd (l!(Suc k)) ∈ Normal ` q"
proof
show suc_skip: "fst(l!Suc k) = Skip"
using a0 a00 a11 k_basic by (metis basic_skip surjective_pairing)
next
obtain s' where k_s: "snd (l!k)=Normal s' ∧ s' ∈ (p)"
using a00 a11 k_basic by auto
then have "snd (l!(Suc k)) = Normal (f s')"
using a00 a11 k_basic stepc_Normal_elim_cases(3)
by (metis prod.inject surjective_pairing)
then show "snd (l!(Suc k)) ∈ Normal ` q" using a0 k_s by blast
qed
obtain s' s'' where
ss:"snd (l!k) = Normal s' ∧ s' ∈ (p) ∧
snd (l!(Suc k)) = Normal s'' ∧ s'' ∈ q"
using suc_k_skip_q k_basic by fastforce
then have "(snd(l!k), snd(l!(Suc k))) ∈ G"
using a0 a1 a2
by (metis Pair_inject a11 k_basic prod.exhaust_sel stepc_Normal_elim_cases(3))
} thus ?thesis by auto qed
have concr:"(final (last l) ⟶
snd (last l) ∉ Fault ` F ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof-
{
assume valid:"final (last l)"
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a l1 where l:"l=a#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a l1] l by fastforce
have env_tran:"env_tran Γ p l R" using assum env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def a2 unfolding env_tran_def by auto
have "∃k. k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
proof -
have "0≤ (length l-1)" using len_l last_l by auto
moreover have "(length l-1) < length l" using len_l by auto
moreover have "final (l!(length l-1))" using valid last_l by auto
moreover have "fst (l!0) = Basic f e" using cp by auto
ultimately show ?thesis
using cp final_exist_component_tran_basic env_tran a2 by blast
qed
then obtain k where k_comp_tran: "k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
by auto
moreover then have "Suc k < length l" by auto
ultimately have all_event:"∀j. 0≤j ∧ j ≠ k ∧ Suc j < length l ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using one_component_tran_basic_env[of Γ l f e k R] a0 a11 a2 a3 assum cp
env_tran_right fst_conv by metis
then have before_k_all_evn:"∀j. 0≤j ∧ j < k ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using k_comp_tran by fastforce
then have k_basic:"fst(l!k) = Basic f e ∧ snd (l!k) ∈ Normal ` (p)"
using cp env_tran_right a2 assum k_comp_tran stability[of p R l 0 k k Γ]
by force
have suc_k_skip_q:"fst(l!Suc k) = Skip ∧ snd (l!(Suc k)) ∈ Normal ` q"
proof
show suc_skip: "fst(l!Suc k) = Skip"
using a0 k_comp_tran k_basic by (metis basic_skip surjective_pairing)
next
obtain s' where k_s: "snd (l!k)=Normal s' ∧ s' ∈ (p)"
using k_comp_tran k_basic by auto
then have "snd (l!(Suc k)) = Normal (f s')"
using k_comp_tran k_basic stepc_Normal_elim_cases(3)
by (metis prod.inject surjective_pairing)
then show "snd (l!(Suc k)) ∈ Normal ` q" using a0 using k_s by blast
qed
have after_k_all_evn:"∀j. (Suc k)≤j ∧ Suc j < (length l) ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using all_event k_comp_tran by fastforce
then have fst_last_skip:"fst (last l) = Skip ∧
snd ((last l)) ∈ Normal ` q"
using a2 last_l len_l cp env_tran_right a3 suc_k_skip_q assum k_comp_tran
stability [of q R l "Suc k" "((length l) - 1)" _ Γ]
by fastforce
} thus ?thesis by auto qed
note res = conjI [OF concl concr]
thus ?thesis using c_prod unfolding comm_def by auto
qed
lemma Basic_sound:
"p ⊆ {s. f s ∈ q} ⟹
(∀s t. s∈p ∧ (t=f s) ⟶ (Normal s,Normal t) ∈ G) ⟹
Sta p R ⟹
Sta q R ⟹
Γ,Θ ⊨n⇘/F⇙ (Basic f e) sat [p, R, G, q,a]"
proof -
assume
a0:"p ⊆ {s. f s ∈ q}" and
a1:"(∀s t. s∈p ∧ (t=f s) ⟶ (Normal s,Normal t) ∈ G)" and
a2:"Sta p R " and
a3:"Sta q R"
{
fix s
have "cpn n Γ (Basic f e) s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ (Basic f e) s" and a11:"c ∈ assum(p, R)"
then have a10:"c∈cp Γ (Basic f e) s"
using cp_def cpn_def cptn_if_cptn_mod cptn_mod_nest_cptn_mod by blast
have "c∈comm(G, (q,a)) F" using Basic_sound1[OF a0 a1 a2 a3 a10 a11] by auto
} thus ?thesis by auto
qed
}
thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
subsection {* Spec Sound *}
lemma spec_skip:
"∀s1 s2 c1. Γ⊢⇩c(Spec r e,s1) → ((c1,s2)) ⟶ c1=Skip"
proof -
{fix s1 s2 c1
assume "Γ⊢⇩c(Spec r e,s1) → ((c1,s2))"
then have "c1=Skip" using stepc_elim_cases(4) by force
} thus ?thesis by auto
qed
lemma no_comp_tran_before_i_spec:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Spec r e" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))" and
a3: "k≤j ∧ j < i ∧ (Γ⊢⇩c(l!j) → (l!(Suc j)))" and
a4: "∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))" and
a5: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc j) ∈ Normal ` q"
shows "P"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Spec r e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using spec_skip by blast
thus ?thesis using a0 a1 a2 a3 a4 a5 no_comp_tran_before_i by blast
qed
lemma only_one_component_tran_i_spec:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Spec r e" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))" and
a3: "k≤j ∧ j ≠ i ∧ Suc j < length l ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ fst (l!j) = Spec r e" and
a4: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc j) ∈ Normal ` q"
shows "P"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Spec r e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using spec_skip by blast
thus ?thesis using a0 a1 a2 a3 a4 only_one_component_tran_i[OF a0 a1 a2 ] by blast
qed
lemma only_one_component_tran_spec:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Spec r e" and
a2: "k≤i ∧ i ≠ j ∧ Suc i<length l ∧ (Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = Spec r e" and
a3: "k≤j ∧ Suc j < length l" and
a4: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q"
shows "(Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Spec r e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using spec_skip by blast
thus ?thesis using a0 a1 a2 a3 a4 only_one_component_tran by blast
qed
lemma only_one_component_tran_all_env_spec:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Spec r e" and
a2: "k≤i ∧ Suc i<length l ∧(Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = Spec r e" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q"
shows "∀j. k≤j ∧ j≠i ∧ Suc j < (length l) ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Spec r e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using spec_skip by blast
thus ?thesis by (metis (no_types) a0 a1 a2 a3 only_one_component_tran_spec)
qed
lemma only_one_component_tran_all_not_comp_spec:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Spec r e" and
a2: "k≤i ∧ Suc i<length l ∧(Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = Spec r e" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q"
shows "∀j. k≤j ∧ j≠i ∧ Suc j < (length l) ⟶ ¬(Γ⊢⇩c(l!j) → (l!(Suc j)))"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Spec r e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using spec_skip by blast
thus ?thesis using a0 a1 a2 a3 only_one_component_tran_all_not_comp by blast
qed
lemma one_component_tran_spec:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!0) = Spec r e" and
a2: "Suc k<length l ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely" and
a4:"p ⊆ {s. (∀t. (s,t)∈r ⟶ t ∈ q) ∧ (∃t. (s,t) ∈ r)}"
shows "∀j. 0≤j ∧ j≠k ∧ Suc j < (length l) ⟶ ¬(Γ⊢⇩c(l!j) → (l!(Suc j)))"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Spec r e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using spec_skip by blast
also obtain j where first:"(Suc j<length l ∧ (Γ⊢⇩c(l!j) → (l!Suc j))) ∧
(∀k < j. ¬((Γ⊢⇩c(l!k) → (l!(Suc k)))))"
by (metis (no_types) a2 exist_first_comp_tran')
moreover then have prg_j:"fst (l!j) = Spec r e" using a1 a0
by (metis cptn_env_same_prog not_step_comp_step_env)
moreover have sta_j:"snd (l!j) ∈ Normal ` p"
proof -
have a0':"0≤j ∧ j<(length l)" using first by auto
have a1':"(∀k. 0≤k ∧ k < j ⟶ ((Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))))"
using first not_step_comp_step_env a0 by fastforce
thus ?thesis using stability first a3 a1' a0' by blast
qed
then have "snd (l!Suc j) ∈ Normal ` q" using a4 first prg_j
proof -
obtain s where s:"snd (l!j) = Normal s ∧ s∈p" using sta_j by fastforce
then have suc_skip: "fst (l!Suc j) = Skip"
using spec_skip first prg_j a4 by (metis (no_types, lifting) prod.collapse)
moreover obtain s' where "snd (l!Suc j) = Normal s' ∧ (s,s')∈r"
proof -
{ have f1:"(Γ⊢⇩c(fst(l!j),snd(l!j)) → (fst(l!Suc j),snd(l!Suc j)))" using first by auto
obtain t where "snd (l!Suc j) = Normal t"
using step_spec_skip_normal_normal[of Γ "fst(l!j)" "snd(l!j)" "fst(l!Suc j)" "snd(l!Suc j)" r]
suc_skip prg_j s a4 f1 by blast
moreover then have "(s,t)∈r" using a4 s prg_j f1 suc_skip stepc_Normal_elim_cases(4)
by (metis (no_types, lifting) stepc_Normal_elim_cases(4) prod.inject xstate.distinct(5) xstate.inject(1))
ultimately have "∃t. snd (l!Suc j)= Normal t ∧ (s,t)∈r" by auto
}
then show "(⋀s'. snd (l ! Suc j) = Normal s' ∧ (s, s') ∈ r ⟹ thesis) ⟹ thesis" ..
qed
then show ?thesis using a4 sta_j s by auto
qed
then have "∀i. 0≤i ∧ i≠j ∧ Suc i < (length l) ⟶ ¬(Γ⊢⇩c(l!i) → (l!(Suc i)))"
using only_one_component_tran_all_not_comp_spec[OF a0 a1] first a3
a0 a1 calculation(1) only_one_component_tran1 prg_j by blast
moreover then have "k=j" using a2 by fastforce
ultimately show ?thesis by auto
qed
lemma one_component_tran_spec_env:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!0) = Spec r e" and
a2: "Suc k<length l ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely" and
a4:"p ⊆ {s. (∀t. (s,t)∈r ⟶ t ∈ q) ∧ (∃t. (s,t) ∈ r)}"
shows "∀j. 0≤j ∧ j≠k ∧ Suc j < (length l) ⟶ Γ⊢⇩c(l!j) →⇩e (l!(Suc j))"
proof -
have "∀j. 0≤j ∧ j≠k ∧ Suc j < (length l) ⟶ ¬ (Γ⊢⇩c(l!j) → (l!(Suc j)))"
using one_component_tran_spec[OF a0 a1 a2 a3 a4] by auto
thus ?thesis using a0
by (metis Suc_eq_plus1 cptn_tran_ce_i step_ce_elim_cases)
qed
lemma final_exist_component_tran_spec:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!i) = Spec r e" and
a2: "env_tran Γ q l R " and
a3: "i≤j ∧ j < length l ∧ final (l!j)"
shows "∃k. k≥i ∧ k<j ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Spec r e,s1) → ((c1,s2)) ⟶ (c1=Skip)"
using spec_skip by blast
thus ?thesis using a0 a1 a2 a3 final_exist_component_tran by blast
qed
lemma Spec_sound1:
"p ⊆ {s. (∀t. (s,t)∈r ⟶ t ∈ q) ∧ (∃t. (s,t) ∈ r)} ⟹
(∀s t. s∈p ∧ (s,t)∈r ⟶ (Normal s, Normal t) ∈ G) ⟹
Sta p R ⟹
Sta q R ⟹
c ∈ cp Γ (Spec r e) s ⟹
c ∈ assum(p, R) ⟹
c ∈ comm (G, (q,a)) F"
proof -
assume
a0:"p ⊆ {s. (∀t. (s,t)∈r ⟶ t ∈ q) ∧ (∃t. (s,t) ∈ r)}" and
a1:"(∀s t. s∈p ∧ (s,t)∈r ⟶ (Normal s,Normal t) ∈ G)" and
a2:"Sta p R" and
a3:"Sta q R" and
a10:"c ∈ cp Γ (Spec r e) s" and
a11:"c ∈ assum(p, R)"
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
have cp:"l!0=(Spec r e,s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1" using a10 cp_def c_prod by fastforce
have assum:"snd(l!0) ∈ Normal ` (p) ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ R)"
using a11 c_prod unfolding assum_def by simp
have concl:"(∀i ns ns'. Suc i<length l ⟶
Γ1⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof -
{ fix k
assume a00:"Suc k<length l" and
a11:"Γ1⊢⇩c(l!k) → (l!(Suc k))"
obtain ck sk csk ssk where tran_pair:
"Γ1⊢⇩c (ck,sk) → (csk, ssk) ∧ (ck = fst (l!k)) ∧ (sk = snd (l!k)) ∧ (csk = fst (l!(Suc k))) ∧ (ssk = snd (l!(Suc k)))"
using a11 by fastforce
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a l1 where l:"l=a#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a l1] l by fastforce
have env_tran:"env_tran Γ p l R" using assum env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def unfolding env_tran_def by auto
then have all_event:"∀j. 0≤j ∧ j ≠ k ∧ Suc j < length l ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using a00 a11 one_component_tran_spec_env[of Γ l r e k R]
env_tran_right fst_conv a0 a2 a3 cp len_l assum
by fastforce
then have before_k_all_evn:"∀j. 0≤j ∧ j < k ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using a00 a11 by fastforce
then have k_basic:"ck = Spec r e ∧ sk ∈ Normal ` (p)"
using cp env_tran_right a2 assum a00 a11 stability[of p R l 0 k k Γ] tran_pair
by force
have suc_skip: "csk = Skip"
using a0 a00 k_basic tran_pair spec_skip by blast
obtain s' where ss:"sk = Normal s' ∧ s' ∈ (p)"
using k_basic by fastforce
obtain s'' where suc_k_skip_q:"ssk = Normal s'' ∧ (s',s'')∈r "
proof -
{from ss obtain t where "ssk = Normal t"
using step_spec_skip_normal_normal[of Γ1 ck sk csk ssk r e s']
k_basic tran_pair a0 suc_skip
by blast
moreover then have "(s',t)∈r" using a0 k_basic ss a11 suc_skip
by (metis (no_types, lifting) stepc_Normal_elim_cases(4) tran_pair prod.inject xstate.distinct(5) xstate.inject(1))
ultimately have "∃t. ssk= Normal t ∧ (s',t)∈r" by auto
}
then show "(⋀s''. ssk = Normal s'' ∧ (s',s'')∈r⟹ thesis) ⟹ thesis" ..
qed
then have "(snd(l!k), snd(l!(Suc k))) ∈ G"
using ss a1 tran_pair by force
} thus ?thesis by auto qed
have concr:"(final (last l) ⟶ ((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof-
{
assume valid:"final (last l)"
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a l1 where l:"l=a#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a l1] l by fastforce
have env_tran:"env_tran Γ p l R" using assum env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def unfolding env_tran_def by auto
have "∃k. k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
proof -
have "0≤ (length l-1)" using len_l last_l by auto
moreover have "(length l-1) < length l" using len_l by auto
moreover have "final (l!(length l-1))" using valid last_l by auto
moreover have "fst (l!0) = Spec r e" using cp by auto
ultimately show ?thesis
using cp final_exist_component_tran_spec env_tran by blast
qed
then obtain k where k_comp_tran: "k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
by auto
then obtain ck sk csk ssk where tran_pair:
"Γ1⊢⇩c (ck,sk) → (csk, ssk) ∧ (ck = fst (l!k)) ∧ (sk = snd (l!k)) ∧ (csk = fst (l!(Suc k))) ∧ (ssk = snd (l!(Suc k)))"
using cp by fastforce
moreover then have "Suc k < length l" using k_comp_tran by auto
ultimately have all_event:"∀j. 0≤j ∧ j ≠ k ∧ Suc j < length l ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using one_component_tran_spec_env[of Γ l r e k R] a0 a11 a2 a3 assum cp
env_tran_right fst_conv
by fastforce
then have before_k_all_evn:"∀j. 0≤j ∧ j < k ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using k_comp_tran by fastforce
then have k_basic:"ck = Spec r e ∧ sk ∈ Normal ` (p)"
using cp env_tran_right a2 assum tran_pair k_comp_tran stability[of p R l 0 k k Γ] tran_pair
by force
have suc_skip: "csk = Skip"
using a0 k_basic tran_pair spec_skip by blast
have suc_k_skip_q:"ssk ∈ Normal ` q"
proof -
obtain s' where k_s: "sk =Normal s' ∧ s' ∈ (p)"
using k_basic by auto
then obtain t where "ssk = Normal t"
using step_spec_skip_normal_normal[of Γ1 ck sk csk ssk r] k_basic tran_pair a0 suc_skip
by blast
then obtain t where "ssk= Normal t" by fastforce
then have "(s',t)∈r" using k_basic k_s a11 suc_skip
by (metis (no_types, lifting) stepc_Normal_elim_cases(4) tran_pair prod.inject xstate.distinct(5) xstate.inject(1))
thus "ssk ∈ Normal ` q" using a0 k_s `ssk = Normal t` by blast
qed
have after_k_all_evn:"∀j. (Suc k)≤j ∧ Suc j < (length l) ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using all_event k_comp_tran by fastforce
then have fst_last_skip:"fst (last l) = Skip ∧
snd ((last l)) ∈ Normal ` q"
using l tran_pair suc_skip last_l len_l cp
env_tran_right a3 suc_k_skip_q
assum k_comp_tran stability [of q R l "Suc k" "((length l) - 1)" _ Γ]
by (metis One_nat_def Suc_eq_plus1 Suc_leI Suc_mono diff_Suc_1 lessI list.size(4))
} thus ?thesis by auto qed
note res = conjI [OF concl concr]
thus ?thesis using c_prod unfolding comm_def by auto
qed
lemma Spec_sound:
"p ⊆ {s. (∀t. (s,t)∈r ⟶ t ∈ q) ∧ (∃t. (s,t) ∈ r)} ⟹
(∀s t. s∈p ∧ (s,t)∈r ⟶ (Normal s, Normal t) ∈ G) ⟹
Sta p R ⟹
Sta q R ⟹
Γ,Θ ⊨n⇘/F⇙ (Spec r e) sat [p, R, G, q,a]"
proof -
assume
a0:"p ⊆ {s. (∀t. (s,t)∈r ⟶ t ∈ q) ∧ (∃t. (s,t) ∈ r)}" and
a1:"(∀s t. s∈p ∧ (s,t)∈r ⟶ (Normal s,Normal t) ∈ G)" and
a2:"Sta p R" and
a3:"Sta q R"
{
fix s
have "cpn n Γ (Spec r e) s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ (Spec r e) s" and a11:"c ∈ assum(p, R)"
then have a10:"c∈cp Γ (Spec r e) s"
using cp_def cpn_def cptn_if_cptn_mod cptn_mod_nest_cptn_mod by blast
have "c∈comm(G, (q,a)) F" using Spec_sound1[OF a0 a1 a2 a3 a10 a11] by auto
} thus ?thesis by auto
qed
}
thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
subsection {* Await Sound *}
lemma await_skip:
"∀s1 s2 c1. Γ⊢⇩c(Await b c e,s1) → ((c1,s2)) ⟶ c1=Skip ∨ (c1 = Throw ∧ (∃s21. s2 =Normal s21 ))"
proof -
{fix s1 s2 c1
assume "Γ⊢⇩c(Await b c e,s1) → ((c1,s2))"
then have "c1=Skip ∨ (c1 = Throw ∧ (∃s21. s2 =Normal s21 ))" using stepc_elim_cases(8) by blast
} thus ?thesis by auto
qed
lemma no_comp_tran_before_i_await:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Await b c e" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))" and
a3: "k≤j ∧ j < i ∧ (Γ⊢⇩c(l!j) → (l!(Suc j)))" and
a4: "∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))" and
a5: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc j) ∈ Normal ` q"
shows "P"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Await b c e,s1) → ((c1,s2)) ⟶ c1=Skip ∨ (c1 = Throw ∧ (∃s21. s2 =Normal s21 ))"
using await_skip by blast
thus ?thesis using a0 a1 a2 a3 a4 a5 no_comp_tran_before_i by blast
qed
lemma only_one_component_tran_i_await:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Await b c e" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i)))" and
a3: "k≤j ∧ j ≠ i ∧ Suc j < length l ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ fst (l!j) = Await b c e" and
a4: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc j) ∈ Normal ` q"
shows "P"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Await b c e,s1) → ((c1,s2)) ⟶ (c1=Skip)∨(c1 = Throw ∧ (∃s21. s2 =Normal s21 ))"
using await_skip by blast
thus ?thesis using a0 a1 a2 a3 a4 only_one_component_tran_i by blast
qed
lemma only_one_component_tran_await:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Await b c e" and
a2: " k≤i ∧ i ≠ j ∧ Suc i<length l ∧ (Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = Await b c e" and
a3: "k≤j ∧ Suc j < length l" and
a4: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q"
shows "(Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Await b c e,s1) → ((c1,s2)) ⟶ (c1=Skip)∨(c1 = Throw ∧ (∃s21. s2 =Normal s21 ))"
using await_skip by blast
thus ?thesis using a0 a1 a2 a3 a4 only_one_component_tran by blast
qed
lemma only_one_component_tran_all_env_await:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Await b c e" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = Await b c e" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q"
shows "∀j. k≤j ∧ j≠i ∧ Suc j < (length l) ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
proof -
have a:"∀s1 s2 c1. Γ⊢⇩c(Await b c e,s1) → ((c1,s2)) ⟶ (c1=Skip)∨(c1 = Throw)"
using await_skip by blast
thus ?thesis by (metis (no_types) a0 a1 a2 a3 only_one_component_tran_await)
qed
lemma only_one_component_tran_all_not_comp_await:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!k) = Await b c e" and
a2: "Suc i<length l ∧ k≤i ∧ (Γ⊢⇩c(l!i) → (l!(Suc i))) ∧ fst (l!i) = Await b c e" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!k) ∈ Normal ` p ∧
Sta q rely ∧ snd (l!Suc i) ∈ Normal ` q"
shows "∀j. k≤j ∧ j≠i ∧ Suc j < (length l) ⟶ ¬(Γ⊢⇩c(l!j) → (l!(Suc j)))"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Await b c e,s1) → ((c1,s2)) ⟶ (c1=Skip)∨(c1 = Throw ∧ (∃s21. s2 =Normal s21 ))"
using await_skip by blast
thus ?thesis using a0 a1 a2 a3 only_one_component_tran_all_not_comp by blast
qed
lemma one_component_tran_await:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!0) = Await b c e" and
a2: "Suc k<length l ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely ∧
Sta a rely" and
a4:"∀V. Γ⇩¬⇩a,{}⊢⇘/F⇙
(p ∩ b ∩ {V}) c
({s. (Normal V, Normal s) ∈ G} ∩ q),
({s. (Normal V, Normal s) ∈ G} ∩ a)" and
a5:"snd (last l) ∉ Fault ` F"
shows "(∀j. 0≤j ∧ j≠k ∧ Suc j < (length l) ⟶ ¬(Γ⊢⇩c(l!j) → (l!(Suc j)))) ∧
(∃s s'. fst (l!k) = Await b c e ∧ snd (l!k) ∈ Normal ` (p) ∧ snd (l!k) = Normal s ∧ snd (l!Suc k) = Normal s' ∧
(snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ q) ∨
snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ a)))"
proof -
have suc_skip:"∀s1 s2 c1. Γ⊢⇩c(Await b c e,s1) → ((c1,s2)) ⟶ (c1=Skip)∨ (c1=Throw ∧ (∃s21. s2 =Normal s21 ))"
using await_skip by blast
also obtain j where first:"(Suc j<length l ∧ (Γ⊢⇩c(l!j) → (l!(Suc j)))) ∧
(∀k < j. ¬((Γ⊢⇩c(l!k) → (l!(Suc k)))))"
by (metis (no_types) a2 exist_first_comp_tran')
moreover then have prg_j:"fst (l!j) = Await b c e" using a1 a0
by (metis cptn_env_same_prog not_step_comp_step_env)
moreover have sta_j:"snd (l!j) ∈ Normal ` p"
proof -
have a0':"0≤j ∧ j<(length l)" using first by auto
have a1':"(∀k. 0≤k ∧ k < j ⟶ ((Γ⊢⇩c(l!k) →⇩e (l!(Suc k)))))"
using first not_step_comp_step_env a0 by fastforce
thus ?thesis using stability first a3 a1' a0' by blast
qed
from sta_j obtain s where
k_basic:"fst (l!j) = Await b c e ∧ snd (l!j) = Normal s ∧ s ∈ p ∧ snd(l!j) ∈ Normal ` p"
using sta_j prg_j by fastforce
then have conc:"snd (l!Suc j) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ q) ∨
snd (l!Suc j) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ a)"
proof -
have "Γ⇩¬⇩a,{}⊨⇘/F⇙
(p ∩ b ∩ {s}) c
({s'. (Normal s, Normal s') ∈ G} ∩ q),
({s'. (Normal s, Normal s') ∈ G} ∩ a)"
using a4 hoare_sound by fastforce
then have e_auto:"Γ⇩¬⇩a⊨⇘/F⇙ (p ∩ b ∩ {s}) c
({s'. (Normal s, Normal s') ∈ G} ∩ q),
({s'. (Normal s, Normal s') ∈ G} ∩ a)"
unfolding cvalid_def by auto
have f': "Γ⊢⇩c(fst (l!j), snd(l!j)) → (fst(l!(Suc j)), snd(l!(Suc j)))"
using first by auto
have step_await:"Suc j<length l ∧ Γ⊢⇩c (Await b c e,snd(l!j)) → (fst(l!(Suc j)), snd(l!(Suc j)))"
using f' k_basic first by fastforce
then have s'_in_bp:"s∈ b ∧ s ∈ p" using k_basic stepc_Normal_elim_cases(8) by metis
then have "s ∈ (p ∩ b)" by fastforce
moreover have test:
"∃t. Γ⇩¬⇩a⊢ ⟨c,Normal s⟩ ⇒ t ∧
((∃t'. t =Abrupt t' ∧ snd(l!Suc j) = Normal t') ∨
(∀t'. t ≠ Abrupt t' ∧ snd(l!Suc j)=t))"
proof -
fix t
{ assume "fst(l!Suc j) = Skip"
then have step:"Γ⊢⇩c (Await b c e,Normal s) → (Skip, snd(l!Suc j))"
using step_await k_basic by fastforce
have s'_b:"s∈b" using s'_in_bp by fastforce
note step = stepc_elim_cases_Await_skip[OF step]
have h:"(s ∈ b ⟹ Γ⇩¬⇩a⊢ ⟨c,Normal s⟩ ⇒ snd(l!Suc j) ⟹ ∀t'. snd(l!Suc j) ≠ Abrupt t' ⟹
Γ⇩¬⇩a⊢ ⟨c,Normal s⟩ ⇒ snd(l!Suc j) ∧ (∀t'. snd(l!Suc j) ≠ Abrupt t'))" by auto
have ?thesis
using step[OF h] by fastforce
} note left = this
{ assume "fst(l!Suc j)= Throw ∧ (∃s1. snd(l!Suc j) = Normal s1)"
then obtain s1 where step:"fst(l!Suc j)= Throw ∧ snd(l!Suc j) = Normal s1"
by fastforce
then have step: "Γ⊢⇩c (Await b c e,Normal s) → (Throw, snd(l!Suc j))"
using step_await k_basic by fastforce
have s'_b:"s∈b" using s'_in_bp by fastforce
note step = stepc_elim_cases_Await_throw[OF step]
have h:"(⋀t'. snd(l!Suc j) = Normal t' ⟹ s ∈ b ⟹ Γ⇩¬⇩a⊢ ⟨c,Normal s⟩ ⇒ Abrupt t' ⟹
Γ⇩¬⇩a⊢ ⟨c,Normal s⟩ ⇒ Abrupt t' ∧ snd(l!Suc j) = Normal t')"
by auto
have ?thesis using step[OF h] by blast
} thus ?thesis using suc_skip left step_await suc_skip by blast
qed
then obtain t where e_step:"Γ⇩¬⇩a⊢ ⟨c,Normal s⟩ ⇒ t ∧
((∃t'. t =Abrupt t' ∧ snd(l!Suc j) = Normal t') ∨
(∀t'. t ≠ Abrupt t' ∧ snd(l!Suc j)=t)) " by fastforce
moreover have "t ∉ Fault ` F"
proof -
{assume a10:"t ∈ Fault ` F"
then obtain tf where "t=Fault tf ∧ tf∈F" by fastforce
then have "snd(l!Suc j) = Fault tf ∧ tf∈F" using e_step by fastforce
also have "snd(l!Suc j) ∉ Fault ` F"
using last_not_F[of Γ l F] a5 a1 step_await a0 by blast
ultimately have False by auto
} thus ?thesis by auto
qed
ultimately have t_q_a:"t ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩q) ∪
Abrupt ` ({s'. (Normal s, Normal s') ∈ G} ∩ a)"
using e_auto unfolding valid_def by fastforce
thus ?thesis using e_step t_q_a by blast
qed
then have "∀i. 0≤i ∧ i≠j ∧ Suc i < (length l) ⟶ ¬(Γ⊢⇩c(l!i) → (l!(Suc i)))"
using only_one_component_tran_all_not_comp_await[OF a0 a1] first a3
a0 a1 calculation(1) only_one_component_tran1 prg_j by blast
moreover then have k:"k=j" using a2 by fastforce
ultimately have "(∀j. 0≤j ∧ j≠k ∧ Suc j < (length l) ⟶ ¬(Γ⊢⇩c(l!j) → (l!(Suc j))))" by auto
also from conc k k_basic have
"(∃s s'. fst (l!k) = Await b c e ∧ snd (l!k) ∈ Normal ` (p) ∧ snd (l!k) = Normal s ∧ snd (l!Suc k) = Normal s' ∧
(snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ q) ∨
snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ a)))"
by fastforce
ultimately show ?thesis by auto
qed
lemma one_component_tran_await_env:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!0) = Await b c e" and
a2: "Suc k<length l ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))" and
a3: "env_tran_right Γ l rely ∧ Sta p rely ∧ snd (l!0) ∈ Normal ` p ∧
Sta q rely ∧
Sta a rely" and
a4:"∀V. Γ⇩¬⇩a,{}⊢⇘/F⇙
(p ∩ b ∩ {V}) c
({s. (Normal V, Normal s) ∈ G} ∩ q),
({s. (Normal V, Normal s) ∈ G} ∩ a)" and
a5:"snd (last l) ∉ Fault ` F"
shows "(∀j. 0≤j ∧ j≠k ∧ Suc j < (length l) ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))) ∧
(∃s s'. fst (l!k) = Await b c e ∧ snd (l!k) ∈ Normal ` (p) ∧
snd (l!k) = Normal s ∧ snd (l!Suc k) = Normal s' ∧
(snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ q) ∨
snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ a)))"
proof -
have "(∀j. 0≤j ∧ j≠k ∧ Suc j < (length l) ⟶ ¬ (Γ⊢⇩c(l!j) → (l!(Suc j)))) ∧
(∃s s'. fst (l!k) = Await b c e ∧ snd (l!k) ∈ Normal ` (p) ∧
snd (l!k) = Normal s ∧ snd (l!Suc k) = Normal s' ∧
(snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ q) ∨
snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ a)))"
using one_component_tran_await[OF a0 a1 a2 a3 a4 a5] by auto
thus ?thesis using a0
by (metis Suc_eq_plus1 cptn_tran_ce_i step_ce_elim_cases)
qed
lemma final_exist_component_tran_await:
assumes a0:"(Γ, l) ∈ cptn" and
a1: "fst (l!i) = Await b c e" and
a2: "env_tran Γ q l R " and
a3: "i≤j ∧ j < length l ∧ final (l!j)"
shows "∃k. k≥i ∧ k<j ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
proof -
have "∀s1 s2 c1. Γ⊢⇩c(Await b c e,s1) → ((c1,s2)) ⟶ (c1=Skip)∨(c1 = Throw ∧ (∃s21. s2 =Normal s21 ))"
using await_skip by blast
thus ?thesis using a0 a1 a2 a3 final_exist_component_tran by blast
qed
inductive_cases stepc_elim_cases_Await_Fault:
"Γ⊢⇩c (Await b c e, Normal s) → (u,Fault f)"
lemma Await_sound1:
"∀V. Γ⇩¬⇩a,{}⊢⇘/F⇙
(p ∩ b ∩ {V}) e
({s. (Normal V, Normal s) ∈ G} ∩ q),
({s. (Normal V, Normal s) ∈ G} ∩ a) ⟹
Sta p R ⟹ Sta q R ⟹ Sta a R ⟹
c ∈ cp Γ (Await b e e1) s ⟹
c ∈ assum(p, R) ⟹
c ∈ comm (G, (q,a)) F"
proof -
assume
a0: "∀V. Γ⇩¬⇩a,{}⊢⇘/F⇙
(p ∩ b ∩ {V}) e
({s. (Normal V, Normal s) ∈ G} ∩ q),
({s. (Normal V, Normal s) ∈ G} ∩ a)" and
a2:"Sta p R" and
a3:"Sta q R" and
a4:"Sta a R" and
a10:"c ∈ cp Γ (Await b e e1) s" and
a11:"c ∈ assum(p, R)"
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
{assume last_fault:"snd (last l) ∉ Fault ` F"
have cp:"l!0=(Await b e e1,s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1" using a10 cp_def c_prod by fastforce
have assum:"snd(l!0) ∈ Normal ` (p) ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ R)"
using a11 c_prod unfolding assum_def by simp
have concl:"(∀i ns ns'. Suc i<length l ⟶
Γ1⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof -
{ fix k ns ns'
assume a00:"Suc k<length l" and
a11:"Γ1⊢⇩c(l!k) → (l!(Suc k))"
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a1 l1 where l:"l=a1#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have env_tran:"env_tran Γ p l R" using assum env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def unfolding env_tran_def by auto
then have all_event:
"(∃s s'. fst (l!k) = Await b e e1 ∧ snd (l!k) ∈ Normal ` (p) ∧ snd (l!k) =
Normal s ∧ snd (l!Suc k) = Normal s' ∧
(snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ q) ∨
snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ a)))"
using a00 a11 one_component_tran_await_env[of Γ l b e e1 k R p q a F G] env_tran_right cp len_l
using a0 a2 a3 a4 assum fst_conv last_fault by auto
then obtain s' s'' where ss:
"snd (l!k) = Normal s' ∧ s' ∈ (p) ∧ snd (l!Suc k) = Normal s''
∧ (s'' ∈ (({s. (Normal s', Normal s) ∈ G} ∩ q)) ∨
s''∈ (({s. (Normal s', Normal s) ∈ G} ∩ a)))"
by fastforce
then have "(snd(l!k), snd(l!(Suc k))) ∈ G"
using a2 by force
} thus ?thesis using c_prod by auto qed
have concr:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof-
{
assume valid:"final (last l)"
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a1 l1 where l:"l=a1#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a1 l1] l by fastforce
have env_tran:"env_tran Γ p l R" using assum env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def unfolding env_tran_def by auto
have "∃k. k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
proof -
have "0≤ (length l-1)" using len_l last_l by auto
moreover have "(length l-1) < length l" using len_l by auto
moreover have "final (l!(length l-1))" using valid last_l by auto
moreover have "fst (l!0) = Await b e e1" using cp by auto
ultimately show ?thesis
using cp final_exist_component_tran_await env_tran by blast
qed
then obtain k where k_comp_tran: "k≥0 ∧ Suc k < length l ∧ (Γ⊢⇩c(l!k) → (l!(Suc k)))"
by fastforce
then obtain ck sk csk ssk where tran_pair:
"Γ1⊢⇩c (ck,sk) → (csk, ssk) ∧ (ck = fst (l!k)) ∧ (sk = snd (l!k)) ∧ (csk = fst (l!(Suc k))) ∧ (ssk = snd (l!(Suc k)))"
using cp by fastforce
have all_event:
"(∀j. 0≤j ∧ j≠k ∧ Suc j < (length l) ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))) ∧
(∃s s'. fst (l!k) = Await b e e1 ∧ snd (l!k) ∈ Normal ` (p) ∧ snd (l!k) =
Normal s ∧ snd (l!Suc k) = Normal s' ∧
(snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ q) ∨
snd (l!Suc k) ∈ Normal ` ({s'. (Normal s, Normal s') ∈ G} ∩ a)))"
using one_component_tran_await_env[of Γ l b e e1 k R p q a F G] a0 a11 a2 a3 a4 assum cp
env_tran_right len_l fst_conv last_fault k_comp_tran by fastforce
then have before_k_all_evn:"∀j. 0≤j ∧ j < k ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using k_comp_tran by fastforce
then obtain s' where k_basic:"ck = Await b e e1 ∧ sk ∈ Normal ` (p) ∧ sk = Normal s'"
using cp env_tran_right a2 assum tran_pair k_comp_tran stability[of p R l 0 k k Γ] tran_pair
by force
have "Γ⇩¬⇩a,{}⊨⇘/F⇙
(p ∩ b ∩ {s'}) e
({s. (Normal s', Normal s) ∈ G} ∩ q),
({s. (Normal s', Normal s) ∈ G} ∩ a)"
using a0 hoare_sound k_basic
by fastforce
then have e_auto:"Γ⇩¬⇩a⊨⇘/F⇙ (p ∩ b ∩ {s'}) e
({s. (Normal s', Normal s) ∈ G} ∩ q),
({s. (Normal s', Normal s) ∈ G} ∩ a)"
unfolding cvalid_def by auto
have after_k_all_evn:"∀j. (Suc k)≤j ∧ Suc j < (length l) ⟶ (Γ⊢⇩c(l!j) →⇩e (l!(Suc j)))"
using all_event k_comp_tran by fastforce
have suc_skip: "csk = Skip ∨ (csk = Throw ∧ (∃s1. ssk = Normal s1))"
using a0 k_basic tran_pair await_skip by blast
moreover {
assume at:"csk = Skip"
then have atom_tran:"Γ⇩¬⇩a⊢⟨e,sk⟩ ⇒ ssk"
using k_basic tran_pair k_basic cp stepc_elim_cases_Await_skip
by metis
have sk_in_normal_pb:"sk ∈ Normal ` (p ∩ b)"
using k_basic tran_pair at cp stepc_elim_cases_Await_skip
by (metis (no_types, lifting) IntI image_iff)
then have "fst (last l) = Skip ∧
snd ((last l)) ∈ Normal ` q"
proof (cases ssk)
case (Normal t)
then have "ssk ∈ Normal ` q"
using sk_in_normal_pb k_basic e_auto Normal atom_tran unfolding valid_def
by blast
thus ?thesis
using at l tran_pair last_l len_l cp
env_tran_right a3 after_k_all_evn
assum k_comp_tran stability [of q R l "Suc k" "((length l) - 1)" _ Γ]
by (metis (no_types, hide_lams) Suc_leI diff_Suc_eq_diff_pred diff_less less_one zero_less_diff)
next
case (Abrupt t)
thus ?thesis
using at k_basic tran_pair k_basic cp stepc_elim_cases_Await_skip
by metis
next
case (Fault f1)
then have "ssk ∈ Normal ` q ∨ ssk ∈ Fault ` F"
using k_basic sk_in_normal_pb e_auto Fault atom_tran unfolding valid_def by auto
thus ?thesis
proof
assume "ssk ∈ Normal ` q" thus ?thesis using Fault by auto
next
assume suck_fault:"ssk ∈ Fault ` F"
have "∀i<length l. snd (l ! i) ∉ Fault ` F"
using last_not_F[of Γ l F] last_fault cp by auto
thus ?thesis
using cp tran_pair a11 k_comp_tran suck_fault
by (meson diff_less len_l less_imp_Suc_add less_one less_trans_Suc)
qed
next
case (Stuck)
then have "ssk ∈ Normal ` q"
using k_basic sk_in_normal_pb e_auto Stuck atom_tran unfolding valid_def
by blast
thus ?thesis using Stuck by auto
qed
}
moreover {
assume at:"(csk = Throw ∧ (∃t. ssk = Normal t))"
then obtain t where ssk_normal:"ssk=Normal t" by auto
then have atom_tran:"Γ⇩¬⇩a⊢⟨e,sk⟩ ⇒ Abrupt t"
using at k_basic tran_pair k_basic ssk_normal cp stepc_elim_cases_Await_throw xstate.inject(1)
by metis
also have "sk ∈ Normal ` (p ∩ b)"
using k_basic tran_pair k_basic ssk_normal at cp stepc_elim_cases_Await_throw
by (metis (no_types, lifting) IntI imageE image_eqI stepc_elim_cases_Await_throw)
then have "ssk ∈ Normal ` a"
using e_auto k_basic ssk_normal atom_tran unfolding valid_def
by blast
then have "(fst (last l) = Throw ∧ snd (last l) ∈ Normal ` (a))"
using at l tran_pair last_l len_l cp
env_tran_right a4 after_k_all_evn
assum k_comp_tran stability [of a R l "Suc k" "((length l) - 1)" _ Γ]
by (metis (no_types, hide_lams) Suc_leI diff_Suc_eq_diff_pred diff_less less_one zero_less_diff)
}
ultimately have "fst (last l) = Skip ∧
snd ((last l)) ∈ Normal ` q ∨
(fst (last l) = Throw ∧ snd (last l) ∈ Normal ` (a))"
by blast
} thus ?thesis by auto qed
note res = conjI [OF concl concr]
}
thus ?thesis using c_prod unfolding comm_def by auto
qed
subsection {* If sound *}
lemma cptn_assum_induct:
assumes
a0: "(Γ,l) ∈ (cp Γ c s) ∧ ((Γ,l) ∈ assum(p, R))" and
a1: "k < length l ∧ l!k=(c1,Normal s') ∧ s' ∈ p1"
shows "(Γ,drop k l)∈ ((cp Γ c1 (Normal s')) ∩ assum(p1, R) )"
proof -
have drop_k_s:"(drop k l)!0 = (c1,Normal s')" using a1 by fastforce
have p1:"s' ∈ p1" using a1 by auto
have k_l:"k < length l" using a1 by auto
show ?thesis
proof
show "(Γ, drop k l) ∈ cp Γ c1 (Normal s')"
unfolding cp_def
using dropcptn_is_cptn a0 a1 drop_k_s cp_def
by fastforce
next
let ?c= "(Γ,drop k l)"
have l:"snd((snd ?c!0)) ∈ Normal ` p1"
using p1 drop_k_s by auto
{fix i
assume a00:"Suc i<length (snd ?c)"
assume a11:"(fst ?c)⊢⇩c((snd ?c)!i) →⇩e ((snd ?c)!(Suc i))"
have "(snd((snd ?c)!i), snd((snd ?c)!(Suc i))) ∈ R "
using a0 unfolding assum_def using a00 a11 by auto
} thus "(Γ, drop k l) ∈ assum (p1, R)"
using l unfolding assum_def by fastforce
qed
qed
lemma Await_sound:
"∀V. Γ⇩¬⇩a,{}⊢⇘/F⇙
(p ∩ b ∩ {V}) e
({s. (Normal V, Normal s) ∈ G} ∩ q),
({s. (Normal V, Normal s) ∈ G} ∩ a) ⟹
Sta p R ⟹ Sta q R ⟹ Sta a R ⟹
Γ,Θ ⊨n⇘/F⇙ (Await b e e1) sat [p, R, G, q,a]"
proof -
assume
a0: "∀V. Γ⇩¬⇩a,{}⊢⇘/F⇙
(p ∩ b ∩ {V}) e
({s. (Normal V, Normal s) ∈ G} ∩ q),
({s. (Normal V, Normal s) ∈ G} ∩ a)" and
a2:"Sta p R" and
a3:"Sta q R" and
a4:"Sta a R"
{
fix s
have "cpn n Γ (Await b e e1) s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ (Await b e e1) s" and a11:"c ∈ assum(p, R)"
then have a10:"c∈cp Γ (Await b e e1) s"
using cp_def cpn_def cptn_if_cptn_mod cptn_mod_nest_cptn_mod by blast
have "c∈comm(G, (q,a)) F" using Await_sound1[OF a0 a2 a3 a4 a10 a11] by auto
} thus ?thesis by auto
qed
}
thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma cptn_comm_induct:
assumes
a0: "(Γ,l) ∈ (cp Γ c s)" and
a1: "l1 = drop j l ∧ (Γ, l1)∈ comm(G, (q,a)) F" and
a2: "k ≥ j ∧ j < length l"
shows "snd (last (l)) ∉ Fault ` F ⟶ ((Suc k < length l ⟶
Γ⊢⇩c(l!k) → (l!(Suc k)) ⟶
(snd(l!k), snd(l!(Suc k))) ∈ G)
∧ (final (last (l)) ⟶
((fst (last (l)) = Skip ∧
snd (last (l)) ∈ Normal ` q)) ∨
(fst (last (l)) = Throw ∧
snd (last (l)) ∈ Normal ` (a))))"
proof -
have pair_Γl:"fst (Γ,l1) = Γ ∧ snd (Γ,l1) = l1" by fastforce
have a03:"snd (last (l1)) ∉ Fault ` F ⟶(∀i.
Suc i<length (snd (Γ, l1)) ⟶
fst (Γ, l1)⊢⇩c((snd (Γ, l1))!i) → ((snd (Γ, l1))!(Suc i)) ⟶
(snd((snd (Γ, l1))!i), snd((snd (Γ, l1))!(Suc i))) ∈ G) ∧
(final (last (snd (Γ, l1))) ⟶
snd (last (snd (Γ, l1))) ∉ Fault ` F ⟶
((fst (last (snd (Γ, l1))) = Skip ∧
snd (last (snd (Γ, l1))) ∈ Normal ` q)) ∨
(fst (last (snd (Γ, l1))) = Throw ∧
snd (last (snd (Γ, l1))) ∈ Normal ` (a)))"
using a1 unfolding comm_def by fastforce
have last_l:"last l1 = last l" using a1 a2 by fastforce
show ?thesis
proof -
{
assume " snd (last l) ∉ Fault ` F"
then have l1_f:"snd (last l1) ∉ Fault ` F"
using a03 a1 a2 by force
{ assume "Suc k < length l"
then have a2: "k ≥ j ∧ Suc k < length l" using a2 by auto
have "k ≤ length l" using a2 by fastforce
then have l1_l:"(l!k = l1! (k - j) ) ∧ (l!Suc k = l1!Suc (k - j))"
using a1 a2 by fastforce
have a00:"Suc (k - j) < length l1" using a1 a2 by fastforce
have "Γ⊢⇩c(l1!(k-j)) → (l1!(Suc (k-j))) ⟶
(snd((snd (Γ, l1))!(k-j)), snd((snd (Γ, l1))!(Suc (k-j)))) ∈ G"
using pair_Γl a00 l1_f a03 by presburger
then have " Γ⊢⇩c(l!k) → (l!(Suc k)) ⟶
(snd (l ! k), snd (l ! Suc k)) ∈ G "
using l1_l last_l by auto
} then have l_side:"Suc k < length l ⟶
Γ⊢⇩c l ! k → l ! Suc k ⟶
(snd (l ! k), snd (l ! Suc k)) ∈ G" by auto
{
assume a10:"final (last (l))"
then have final_eq: "final (last (l1))"
using a10 a1 a2 by fastforce
also have "snd (last (l1)) ∉ Fault ` F"
using last_l l1_f by fastforce
ultimately have "((fst (last (snd (Γ, l1))) = Skip ∧
snd (last (snd (Γ, l1))) ∈ Normal ` q)) ∨
(fst (last (snd (Γ, l1))) = Throw ∧
snd (last (snd (Γ, l1))) ∈ Normal ` (a))"
using pair_Γl a03 by presburger
then have "((fst (last (snd (Γ, l))) = Skip ∧
snd (last (snd (Γ, l))) ∈ Normal ` q)) ∨
(fst (last (snd (Γ, l))) = Throw ∧
snd (last (snd (Γ, l))) ∈ Normal ` (a))"
using final_eq a1 a2 by auto
} then have
r_side:
"SmallStepCon.final (last l) ⟶
fst (last l) = LanguageCon.com.Skip ∧ snd (last l) ∈ Normal ` q ∨
fst (last l) = LanguageCon.com.Throw ∧ snd (last l) ∈ Normal ` a"
by fastforce
note res=conjI[OF l_side r_side]
} thus ?thesis by auto
qed
qed
lemma cpn_assum_induct:
assumes
a0: "(Γ,l) ∈ (cpn n Γ c s) ∧ ((Γ,l) ∈ assum(p, R))" and
a1: "k < length l ∧ l!k=(c1,Normal s') ∧ s' ∈ p1"
shows "(Γ,drop k l)∈ ((cpn n Γ c1 (Normal s')) ∩ assum(p1, R) )"
proof -
have drop_k_s:"(drop k l)!0 = (c1,Normal s')" using a1 by fastforce
have p1:"s' ∈ p1" using a1 by auto
have k_l:"k < length l" using a1 by auto
show ?thesis
proof
show "(Γ, drop k l) ∈ cpn n Γ c1 (Normal s')"
unfolding cp_def
using a0 a1
by (simp add: cpn_def dropcptn_is_cptn1)
next
let ?c= "(Γ,drop k l)"
have l:"snd((snd ?c!0)) ∈ Normal ` p1"
using p1 drop_k_s by auto
{fix i
assume a00:"Suc i<length (snd ?c)"
assume a11:"(fst ?c)⊢⇩c((snd ?c)!i) →⇩e ((snd ?c)!(Suc i))"
have "(snd((snd ?c)!i), snd((snd ?c)!(Suc i))) ∈ R "
using a0 unfolding assum_def using a00 a11 by auto
} thus "(Γ, drop k l) ∈ assum (p1, R)"
using l unfolding assum_def by fastforce
qed
qed
lemma cpn_comm_induct:
assumes
a1: "l1 = drop j l ∧ (Γ, l1)∈ comm(G, (q,a)) F" and
a2: "k ≥ j ∧ j < length l"
shows "snd (last (l)) ∉ Fault ` F ⟶ ((Suc k < length l ⟶
Γ⊢⇩c(l!k) → (l!(Suc k)) ⟶
(snd(l!k), snd(l!(Suc k))) ∈ G)
∧ (final (last (l)) ⟶
((fst (last (l)) = Skip ∧
snd (last (l)) ∈ Normal ` q)) ∨
(fst (last (l)) = Throw ∧
snd (last (l)) ∈ Normal ` (a))))"
proof -
have pair_Γl:"fst (Γ,l1) = Γ ∧ snd (Γ,l1) = l1" by fastforce
have a03:"snd (last (l1)) ∉ Fault ` F ⟶(∀i.
Suc i<length (snd (Γ, l1)) ⟶
fst (Γ, l1)⊢⇩c((snd (Γ, l1))!i) → ((snd (Γ, l1))!(Suc i)) ⟶
(snd((snd (Γ, l1))!i), snd((snd (Γ, l1))!(Suc i))) ∈ G) ∧
(final (last (snd (Γ, l1))) ⟶
snd (last (snd (Γ, l1))) ∉ Fault ` F ⟶
((fst (last (snd (Γ, l1))) = Skip ∧
snd (last (snd (Γ, l1))) ∈ Normal ` q)) ∨
(fst (last (snd (Γ, l1))) = Throw ∧
snd (last (snd (Γ, l1))) ∈ Normal ` (a)))"
using a1 unfolding comm_def by fastforce
have last_l:"last l1 = last l" using a1 a2 by fastforce
show ?thesis
proof -
{
assume " snd (last l) ∉ Fault ` F"
then have l1_f:"snd (last l1) ∉ Fault ` F"
using a03 a1 a2 by force
{ assume "Suc k < length l"
then have a2: "k ≥ j ∧ Suc k < length l" using a2 by auto
have "k ≤ length l" using a2 by fastforce
then have l1_l:"(l!k = l1! (k - j) ) ∧ (l!Suc k = l1!Suc (k - j))"
using a1 a2 by fastforce
have a00:"Suc (k - j) < length l1" using a1 a2 by fastforce
have "Γ⊢⇩c(l1!(k-j)) → (l1!(Suc (k-j))) ⟶
(snd((snd (Γ, l1))!(k-j)), snd((snd (Γ, l1))!(Suc (k-j)))) ∈ G"
using pair_Γl a00 l1_f a03 by presburger
then have " Γ⊢⇩c(l!k) → (l!(Suc k)) ⟶
(snd (l ! k), snd (l ! Suc k)) ∈ G "
using l1_l last_l by auto
} then have l_side:"Suc k < length l ⟶
Γ⊢⇩c l ! k → l ! Suc k ⟶
(snd (l ! k), snd (l ! Suc k)) ∈ G" by auto
{
assume a10:"final (last (l))"
then have final_eq: "final (last (l1))"
using a10 a1 a2 by fastforce
also have "snd (last (l1)) ∉ Fault ` F"
using last_l l1_f by fastforce
ultimately have "((fst (last (snd (Γ, l1))) = Skip ∧
snd (last (snd (Γ, l1))) ∈ Normal ` q)) ∨
(fst (last (snd (Γ, l1))) = Throw ∧
snd (last (snd (Γ, l1))) ∈ Normal ` (a))"
using pair_Γl a03 by presburger
then have "((fst (last (snd (Γ, l))) = Skip ∧
snd (last (snd (Γ, l))) ∈ Normal ` q)) ∨
(fst (last (snd (Γ, l))) = Throw ∧
snd (last (snd (Γ, l))) ∈ Normal ` (a))"
using final_eq a1 a2 by auto
} then have
r_side:
"SmallStepCon.final (last l) ⟶
fst (last l) = LanguageCon.com.Skip ∧ snd (last l) ∈ Normal ` q ∨
fst (last l) = LanguageCon.com.Throw ∧ snd (last l) ∈ Normal ` a"
by fastforce
note res=conjI[OF l_side r_side]
} thus ?thesis by auto
qed
qed
lemma If_sound:
"Γ,Θ ⊢⇘/F⇙ c1 sat [p ∩ b, R, G, q,a] ⟹
(∀n. Γ,Θ ⊨n⇘/F⇙ c1 sat [p ∩ b, R, G, q,a]) ⟹
Γ,Θ ⊢⇘/F⇙ c2 sat [p ∩ (-b), R, G, q,a] ⟹
(∀n. Γ,Θ ⊨n⇘/F⇙ c2 sat [p ∩ (-b), R, G, q,a]) ⟹
Sta p R ⟹ (∀s. (Normal s, Normal s) ∈ G) ⟹
Γ,Θ ⊨n⇘/F⇙ (Cond b c1 c2) sat [p, R, G, q,a]"
proof -
assume
a0:"Γ,Θ ⊢⇘/F⇙ c1 sat [p ∩ b, R, G, q,a]" and
a1:"Γ,Θ ⊢⇘/F⇙ c2 sat [p ∩ (-b), R, G, q,a]" and
a2: "∀n. Γ,Θ ⊨n⇘/F⇙ c1 sat [p ∩ b, R, G, q,a]" and
a3: "∀n. Γ,Θ ⊨n⇘/F⇙ c2 sat [p ∩ (-b), R, G, q,a]" and
a4: "Sta p R" and
a5: "(∀s. (Normal s, Normal s) ∈ G)"
{
fix s
assume all_call:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
then have a3:"Γ ⊨n⇘/F⇙ c2 sat [p ∩ (-b), R, G, q,a]"
using a3 com_cvalidityn_def by fastforce
have a2:"Γ ⊨n⇘/F⇙ c1 sat [p ∩ b, R, G, q,a]"
using a2 all_call com_cvalidityn_def by fastforce
have "cpn n Γ (Cond b c1 c2) s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ (Cond b c1 c2) s" and a11:"c ∈ assum(p, R)"
then have a10':"c ∈ cp Γ (Cond b c1 c2) s" unfolding cp_def cpn_def
using cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod by fastforce
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
have "c ∈ comm(G, (q,a)) F"
proof -
{assume l_f:"snd (last l) ∉ Fault ` F"
have cp:"l!0=((Cond b c1 c2),s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1" using a10' cp_def c_prod by fastforce
have Γ1:"(Γ, l) = c" using c_prod cp by blast
have assum:"snd(l!0) ∈ Normal ` (p) ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ R)"
using a11 c_prod unfolding assum_def by simp
then have env_tran:"env_tran Γ p l R" using env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def unfolding env_tran_def by auto
have concl:"(∀i. Suc i<length l ⟶
Γ1⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof -
{ fix k ns ns'
assume a00:"Suc k<length l" and
a21:"Γ⊢⇩c(l!k) → (l!(Suc k))"
obtain j where before_k_all_evnt:"j≤k ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ (∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
using a00 a21 exist_first_comp_tran cp by blast
then obtain cj sj csj ssj where pair_j:"(Γ⊢⇩c(cj,sj) → (csj,ssj)) ∧ cj = fst (l!j) ∧ sj = snd (l!j) ∧ csj = fst (l!(Suc j)) ∧ ssj = snd(l!(Suc j))"
by fastforce
have k_basic:"cj = (Cond b c1 c2) ∧ sj ∈ Normal ` (p)"
using pair_j before_k_all_evnt cp env_tran_right a4 assum a00 stability[of p R l 0 j j Γ]
by force
then obtain s' where ss:"sj = Normal s' ∧ s'∈ (p)" by auto
then have ssj_normal_s:"ssj = Normal s'" using before_k_all_evnt k_basic pair_j
by (metis prod.collapse snd_conv stepc_Normal_elim_cases(6))
have "(snd(l!k), snd(l!(Suc k))) ∈ G"
using ss a2 unfolding Satis_def
proof (cases "k=j")
case True
have "(Normal s', Normal s') ∈ G"
using a5 by blast
thus "(snd (l ! k), snd (l ! Suc k)) ∈ G"
using pair_j k_basic True ss ssj_normal_s by auto
next
case False
have j_length:"Suc j < length l" using a00 before_k_all_evnt by fastforce
have l_suc:"l!(Suc j) = (csj, Normal s')"
using before_k_all_evnt pair_j ssj_normal_s
by fastforce
have l_k:"j<k" using before_k_all_evnt False by fastforce
have "s'∈b ∨ s'∉b" by auto
thus "(snd (l ! k), snd (l ! Suc k)) ∈ G"
proof
assume a000:"s'∈b"
then have cj:"csj=c1" using k_basic pair_j ss
by (metis (no_types) fst_conv stepc_Normal_elim_cases(6))
moreover have p1:"s' ∈ (p ∩ b)" using a000 ss by blast
moreover have "cpn n Γ csj ssj ∩ assum((p ∩ b), R) ⊆ comm(G, (q,a)) F"
using calculation a2 com_validityn_def cj by blast
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using l_suc j_length a10 a11 Γ1 ssj_normal_s
cpn_assum_induct[of Γ l n "(LanguageCon.com.Cond b c1 c2)" s p R "Suc j" c1 s' "(p ∩ b)"]
by blast
show ?thesis
using l_k drop_comm a00 a21 a10 Γ1 l_f
cpn_comm_induct
by fastforce
next
assume a000:"s'∉b"
then have cj:"csj=c2" using k_basic pair_j ss
by (metis (no_types) fst_conv stepc_Normal_elim_cases(6))
moreover have p1:"s' ∈ (p ∩ (-b))" using a000 ss by fastforce
moreover then have "cpn n Γ csj ssj ∩ assum((p ∩ (-b)), R) ⊆ comm(G, (q,a)) F"
using a3 com_validityn_def cj by blast
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using l_suc j_length a10 a11 Γ1 ssj_normal_s
cpn_assum_induct[of Γ l n "(LanguageCon.com.Cond b c1 c2)" s p R "Suc j" c2 s' "(p ∩ (-b))"]
by fastforce
show ?thesis
using l_k drop_comm a00 a21 a10 Γ1 l_f
cpn_comm_induct
unfolding Satis_def by fastforce
qed
qed
} thus ?thesis by (simp add: c_prod cp) qed
have concr:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof-
{
assume valid:"final (last l)"
assume not_fault: "snd (last l) ∉ Fault ` F"
have "∃k. k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k))) ∧ final (l!(Suc k))"
proof -
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a1 l1 where l:"l=a1#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a1 l1] l by fastforce
have final_0:"¬final(l!0)" using cp unfolding final_def by auto
have "0≤ (length l-1)" using len_l last_l by auto
moreover have "(length l-1) < length l" using len_l by auto
moreover have "final (l!(length l-1))" using valid last_l by auto
moreover have "fst (l!0) = LanguageCon.com.Cond b c1 c2" using cp by auto
ultimately show ?thesis
using cp final_exist_component_tran_final env_tran_right final_0
by blast
qed
then obtain k where a21: "k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k))) ∧ final (l!(Suc k))"
by auto
then have a00:"Suc k<length l" by fastforce
then obtain j where before_k_all_evnt:"j≤k ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ (∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
using a00 a21 exist_first_comp_tran cp by blast
then obtain cj sj csj ssj where pair_j:"(Γ⊢⇩c(cj,sj) → (csj,ssj)) ∧ cj = fst (l!j) ∧ sj = snd (l!j) ∧ csj = fst (l!(Suc j)) ∧ ssj = snd(l!(Suc j))"
by fastforce
have j_length:"Suc j < length l" using a00 before_k_all_evnt by fastforce
then have k_basic:"cj = (Cond b c1 c2) ∧ sj ∈ Normal ` (p)"
using pair_j before_k_all_evnt cp env_tran_right a4 assum a00 stability[of p R l 0 j j Γ]
by fastforce
then obtain s' where ss:"sj = Normal s' ∧ s'∈ (p)" by auto
then have ssj_normal_s:"ssj = Normal s'" using before_k_all_evnt k_basic pair_j
by (metis prod.collapse snd_conv stepc_Normal_elim_cases(6))
have l_suc:"l!(Suc j) = (csj, Normal s')"
using before_k_all_evnt pair_j ssj_normal_s
by fastforce
have "s'∈b ∨ s'∉b" by auto
then have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a))"
proof
assume a000:"s'∈b"
then have cj:"csj=c1" using k_basic pair_j ss
by (metis (no_types) fst_conv stepc_Normal_elim_cases(6))
moreover have p1:"s' ∈ (p ∩ b)" using a000 ss by blast
moreover then have "cpn n Γ csj ssj ∩ assum((p ∩ b), R) ⊆ comm(G, (q,a)) F"
using a2 com_validityn_def cj by blast
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using l_suc j_length a10 a11 Γ1 ssj_normal_s
cpn_assum_induct[of Γ l n "(LanguageCon.com.Cond b c1 c2)" s p R "Suc j" c1 s' "(p ∩ b)"]
by blast
thus ?thesis
using j_length drop_comm a10 Γ1 cpn_comm_induct valid not_fault
by blast
next
assume a000:"s'∉b"
then have cj:"csj=c2" using k_basic pair_j ss
by (metis (no_types) fst_conv stepc_Normal_elim_cases(6))
moreover have p1:"s'∈(p ∩ (-b))" using a000 ss by blast
moreover then have "cpn n Γ csj ssj ∩ assum((p ∩ (-b)), R) ⊆ comm(G, (q,a)) F"
using a3 com_validityn_def cj by blast
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using l_suc j_length a10 a11 Γ1 ssj_normal_s
cpn_assum_induct[of Γ l n "(LanguageCon.com.Cond b c1 c2)" s p R "Suc j" c2 s' "(p ∩ (-b))"]
by blast
thus ?thesis
using j_length drop_comm a10 Γ1 cpn_comm_induct valid not_fault
by blast
qed
} thus ?thesis using l_f by fastforce qed
note res = conjI [OF concl concr]
}
thus ?thesis using c_prod unfolding comm_def by auto qed
} thus ?thesis by auto qed
} thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma Asm_sound:
"(c, p, R, G, q, a) ∈ Θ ⟹
Γ,Θ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]
"
proof -
assume
a0:"(c, p, R, G, q, a) ∈ Θ"
{ fix s
assume all_call:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
then have "Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]" using a0 by auto
} thus ?thesis unfolding com_cvalidityn_def by auto
qed
lemma events_p:
assumes a0:"(Γ,cfg#l)∈cptn" and
a1:"(Γ,cfg#l) ∈ assum (p,R)" and
a2:"i<length (cfg#l)" and
a3:"∀k≤i. fst ((cfg#l)!k) = fst cfg" and
a4:"Sta p R "
shows "∃t1. snd((cfg#l)!i)=Normal t1 ∧ t1 ∈ p"
using a2 a3
proof(induct i)
case 0
then show ?case using a1 a2 unfolding assum_def by auto
next
case (Suc n)
then have "∃t1. snd ((cfg # l) ! n) = Normal t1 ∧ t1 ∈ p" by auto
moreover have "Γ⊢⇩c ((cfg#l)!n) →⇩e ((cfg#l)!(Suc n))" using Suc a0
by (metis Env calculation less_Suc_eq_le less_not_refl nat_le_linear prod.collapse)
then have "(snd ((cfg#l)!n),snd ((cfg#l)!(Suc n)))∈R" using a1 Suc(2)
unfolding assum_def by auto
ultimately show ?case using a4 unfolding Sta_def by auto
qed
lemma not_val_zero:"c ∈ dom Γ ⟹ Sta p R ⟹ Γ ⊨0⇘/F⇙ Call c sat [p,R, G, q,a]"
proof-
assume a0:" c ∈ dom Γ"
assume a1:"Sta p R"
{fix l s
assume a01:"(Γ,l) ∈ cpn 0 Γ (Call c) s ∧ (Γ,l)∈ assum(p, R)"
then have "length l ≥ 1" unfolding cpn_def using CptnEmpty
by (metis (no_types, lifting) One_nat_def Product_Type.Collect_case_prodD Suc_leI length_greater_0_conv snd_conv)
moreover {assume a02:"length l = 1"
then have "l = [(Call c,s)]"
proof -
have "l ! 0 = (LanguageCon.com.Call c, s)" using a01 unfolding cpn_def
by fastforce
then show ?thesis using a02
by (metis One_nat_def Suc_leI impossible_Cons length_greater_0_conv list.size(3) neq_Nil_conv nth_Cons_0 zero_neq_one)
qed
then have "(Γ,l) ∈ comm(G, (q,a)) F" unfolding comm_def final_def by auto
}
moreover {assume a02:"length l > 1"
then obtain a1 ls where l:"l =(Call c, s)#a1#ls" using a01 unfolding cpn_def
apply auto
by (metis (no_types, hide_lams) One_nat_def Suc_eq_plus1 less_not_refl list.exhaust list.size(3) list.size(4) not_less_zero nth_Cons_0 prod.collapse)
have l_cptn:"(Γ,l)∈cptn" using a01 unfolding cpn_def
using cptn_eq_cptn_mod_nest by blast
then obtain m where min_call:"min_call m Γ l"
using cptn_eq_cptn_mod_set cptn_mod_cptn_mod_nest minimum_nest_call by blast
{ assume a03:"∀i<length l. fst (l!i) = Call c"
then have "(Γ,l) ∈ comm(G, (q,a)) F"
using no_comp_tran_no_final_comm[OF _ a03] a02 unfolding final_def
by fastforce
}
moreover { assume a03:"¬ (∀i<length l. fst (l!i) = Call c)"
then obtain i where i:"(i<length l ∧ fst (l!i) ≠ Call c)"
by auto
then obtain j where cfg_j:"fst (l!j) ≠ Call c ∧ (∀k<j. fst (l!k) = Call c)"
by (fast dest: exists_first_occ[of "λi. fst (l!i) ≠ Call c" i])
moreover have j:"j>0 ∧ j<length l" using l i calculation
by (metis gr0I fstI leI le_less_trans nth_Cons')
ultimately have step:"(Γ⊢⇩c (l!(j-1)) → (l!j))"
using l l_cptn
by (metis One_nat_def Suc_pred cptn_stepc_rtran diff_less not_eq_not_env prod.collapse
step_ce_elim_cases zero_less_one)
moreover obtain s' where j_1_cfg:"snd (l!(j-1)) = Normal s' ∧ s' ∈ p"
using cfg_j l a01[simplified l] j[simplified l] i a1 events_p[OF l_cptn[simplified l] _ _ _ a1, of "j-1"]
by force
then have j_cfg:"l!j = (the (Γ c), Normal s')" using cfg_j a0
stepc_Normal_elim_cases(9) calculation
by (metis diff_less domIff j option.sel prod.collapse zero_less_one)
ultimately have False
proof-
have "(0,Γ, drop (j-1) l) ∈ cptn_mod_nest_call"
using a01 unfolding cpn_def
by (simp add: dropcptn_is_cptn1 j less_imp_diff_less)
then show ?thesis
using redex_call_cptn_mod_min_nest_call_gr_zero j j_cfg j_1_cfg cfg_j step a0
by (metis Cons_nth_drop_Suc One_nat_def SmallStepCon.redex.simps(7) Suc_pred
diff_less domIff less_imp_diff_less min_call_def not_less_zero prod.collapse
stepc_Normal_elim_cases(9) zero_less_one)
qed
}
ultimately have "(Γ,l) ∈ comm(G, (q,a)) F" by auto
}
ultimately have "(Γ,l)∈comm(G, (q,a)) F" by fastforce
} then show ?thesis unfolding com_validityn_def cpn_def by auto
qed
lemma Call_sound:
"f ∈ dom Γ ⟹
∀n. Γ,Θ ⊨n⇘/F⇙ (the (Γ f)) sat [p, R, G, q,a] ⟹
Sta p R ⟹ (∀s. (Normal s,Normal s) ∈ G) ⟹
Γ,Θ ⊨n⇘/F⇙ (Call f) sat [p, R, G, q,a]"
proof -
assume
a0:"f ∈ dom Γ" and
a2:"∀n. Γ,Θ ⊨n⇘/F⇙ (the (Γ f)) sat [p, R, G, q,a]" and
a3: "Sta p R" and
a4: "(∀s. (Normal s, Normal s) ∈ G)"
obtain bdy where a0:"Γ f = Some bdy" using a0 by auto
{
fix s
assume all_call:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
then have a2:"Γ ⊨n⇘/F⇙ bdy sat [p, R, G, q,a]"
using a0 a2 com_cvalidityn_def by fastforce
have "cpn n Γ (Call f) s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ (Call f) s" and a11:"c ∈ assum(p, R)"
then have a10':"c ∈ cp Γ (Call f) s"
unfolding cpn_def cp_def using cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod by fastforce
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
have "c ∈ comm(G, (q,a)) F"
proof -
{assume l_f:"snd (last l) ∉ Fault ` F"
have cp:"l!0=((Call f),s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1" using a10' cp_def c_prod by fastforce
have Γ1:"(Γ, l) = c" using c_prod cp by blast
have assum:"snd(l!0) ∈ Normal ` (p) ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ R)"
using a11 c_prod unfolding assum_def by simp
then have env_tran:"env_tran Γ p l R" using env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def unfolding env_tran_def by auto
have concl:"(∀i. Suc i<length l ⟶
Γ1⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof -
{ fix k ns ns'
assume a00:"Suc k<length l" and
a21:"Γ⊢⇩c(l!k) → (l!(Suc k))"
obtain j where before_k_all_evnt:"j≤k ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ (∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
using a00 a21 exist_first_comp_tran cp by blast
then obtain cj sj csj ssj where pair_j:"(Γ⊢⇩c(cj,sj) → (csj,ssj)) ∧ cj = fst (l!j) ∧ sj = snd (l!j) ∧ csj = fst (l!(Suc j)) ∧ ssj = snd(l!(Suc j))"
by fastforce
have k_basic:"cj = (Call f) ∧ sj ∈ Normal ` (p)"
using pair_j before_k_all_evnt cp env_tran_right a3 assum a00 stability[of p R l 0 j j Γ]
by force
then obtain s' where ss:"sj = Normal s' ∧ s'∈ (p)" by auto
then have ssj_normal_s:"ssj = Normal s'"
using before_k_all_evnt k_basic pair_j a0
by (metis not_None_eq snd_conv stepc_Normal_elim_cases(9))
have "(snd(l!k), snd(l!(Suc k))) ∈ G"
using ss a2
proof (cases "k=j")
case True
have "(Normal s', Normal s') ∈ G"
using a4 by fastforce
thus "(snd (l ! k), snd (l ! Suc k)) ∈ G"
using pair_j k_basic True ss ssj_normal_s by auto
next
case False
have j_k:"j<k" using before_k_all_evnt False by fastforce
thus "(snd (l ! k), snd (l ! Suc k)) ∈ G"
proof -
have j_length:"Suc j < length l" using a00 before_k_all_evnt by fastforce
have cj:"csj=bdy" using k_basic pair_j ss a0
by (metis fst_conv option.distinct(1) option.sel stepc_Normal_elim_cases(9))
moreover have p1:"s'∈p" using ss by blast
moreover then have "cpn n Γ csj ssj ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
using a2 com_validityn_def cj by blast
moreover then have "l!(Suc j) = (csj, Normal s')"
using before_k_all_evnt pair_j cj ssj_normal_s
by fastforce
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using j_length a10 a11 Γ1 ssj_normal_s
by (meson contra_subsetD cpn_assum_induct)
then show ?thesis
using a00 a21 Γ1 j_k j_length l_f
cptn_comm_induct[of Γ l "Call f" s _ "Suc j" G q a F k ]
Suc_leI a10' by blast
qed
qed
} thus ?thesis by (simp add: c_prod cp) qed
have concr:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof-
{
assume valid:"final (last l)"
have "∃k. k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k))) ∧ final (l!(Suc k))"
proof -
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a1 l1 where l:"l=a1#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a1 l1] l by fastforce
have final_0:"¬final(l!0)" using cp unfolding final_def by auto
have "0≤ (length l-1)" using len_l last_l by auto
moreover have "(length l-1) < length l" using len_l by auto
moreover have "final (l!(length l-1))" using valid last_l by auto
moreover have "fst (l!0) = Call f" using cp by auto
ultimately show ?thesis
using cp final_exist_component_tran_final env_tran_right final_0
by blast
qed
then obtain k where a21: "k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k))) ∧ final (l!(Suc k))"
by auto
then have a00:"Suc k<length l" by fastforce
then obtain j where before_k_all_evnt:"j≤k ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ (∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
using a00 a21 exist_first_comp_tran cp by blast
then obtain cj sj csj ssj where pair_j:"(Γ⊢⇩c(cj,sj) → (csj,ssj)) ∧ cj = fst (l!j) ∧ sj = snd (l!j) ∧ csj = fst (l!(Suc j)) ∧ ssj = snd(l!(Suc j))"
by fastforce
have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a))"
proof -
have j_length:"Suc j < length l" using a00 before_k_all_evnt by fastforce
then have k_basic:"cj = (Call f) ∧ sj ∈ Normal ` (p)"
using pair_j before_k_all_evnt cp env_tran_right a3 assum a00 stability[of p R l 0 j j Γ]
by force
then obtain s' where ss:"sj = Normal s' ∧ s'∈ (p)" by auto
then have ssj_normal_s:"ssj = Normal s'"
using before_k_all_evnt k_basic pair_j a0
by (metis not_None_eq snd_conv stepc_Normal_elim_cases(9))
have cj:"csj=bdy" using k_basic pair_j ss a0
by (metis fst_conv option.distinct(1) option.sel stepc_Normal_elim_cases(9))
moreover have p1:"s'∈p" using ss by blast
moreover then have "cpn n Γ csj ssj ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
using a2 com_validityn_def cj by blast
moreover then have "l!(Suc j) = (csj, Normal s')"
using before_k_all_evnt pair_j cj ssj_normal_s
by fastforce
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using j_length a10 a11 Γ1 ssj_normal_s
by (meson contra_subsetD cpn_assum_induct)
thus ?thesis
using j_length l_f drop_comm a10' Γ1 cptn_comm_induct[of Γ l "Call f" s _ "Suc j" G q a F "Suc j"] valid
by blast
qed
} thus ?thesis by auto
qed
note res = conjI [OF concl concr]}
thus ?thesis using c_prod unfolding comm_def by force qed
} thus ?thesis by auto qed
} thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma CallRec_sound:
"(c, p, R, G, q, a) ∈ Specs ⟹
∀(c, p, R, G, q, a)∈Specs.
c ∈ dom Γ ∧
Sta p R ∧ (∀s. (Normal s,Normal s) ∈ G) ∧
Γ,Θ ∪ Specs ⊢⇘/F⇙ the (Γ c) sat [p, R, G, q,a] ∧
(∀x. Γ,Θ ∪ Specs ⊨x ⇘/F⇙ the (Γ c) sat [p,R, G, q,a]) ⟹
Sta p R ⟹ (∀s. (Normal s,Normal s) ∈ G) ⟹
Γ,Θ ⊨n⇘/F⇙ Call c sat [p,R, G, q,a]"
proof -
assume a0: "(c, p, R, G, q, a) ∈ Specs" and
a1:
"∀(c, p, R, G, q, a)∈Specs.
c ∈ dom Γ ∧ Sta p R ∧ (∀s. (Normal s,Normal s) ∈ G) ∧
Γ,Θ ∪ Specs ⊢⇘/F⇙ the (Γ c) sat [p, R, G, q,a] ∧
(∀x. Γ,Θ ∪ Specs ⊨x ⇘/F⇙ the (Γ c) sat [p,R, G, q,a])"
then have a1': "c ∈ dom Γ" and
a1'': "Γ,Θ ∪ Specs ⊨n ⇘/F⇙ the (Γ c) sat [p,R, G, q,a]" using a0 by auto
from a1 have
valid_body:
"∀(c, p, R, G, q, a)∈Specs.
c ∈ dom Γ ∧ Sta p R ∧ (∀s. (Normal s,Normal s) ∈ G) ∧
(∀x. Γ,Θ ∪ Specs ⊨x ⇘/F⇙ the (Γ c) sat [p,R, G, q,a])" by fastforce
assume a5: "Sta p R" and
a6: "(∀s. (Normal s,Normal s) ∈ G)"
obtain bdy where Γbdy:"Γ c = Some bdy" using a1' by auto
have theta_specs:
"∀(c, p, R, G, q, a)∈Θ. Γ ⊨n⇘/F⇙ Call c sat [p,R, G, q,a] ⟹
∀(c, p, R, G, q, a)∈Specs. Γ ⊨n⇘/F⇙ Call c sat [p,R, G, q,a]"
proof(induct n)
case 0
show "∀(c, p, R, G, a, d)∈Specs. Γ ⊨0⇘/F⇙ LanguageCon.com.Call c sat [p,R, G, a,d]"
proof-
{fix c p R G a d
assume a00:"(c, p, R, G, a, d) ∈ Specs"
then have "c∈dom Γ ∧ Sta p R" using a1 by auto
then have " Γ ⊨0⇘/F⇙ (LanguageCon.com.Call c) sat [p,R, G, a,d]"
using not_val_zero by fastforce
} then show ?thesis by auto
qed
next
case (Suc n)
have hyp:"∀(c, p, R, G, q, a)∈Θ. Γ ⊨n⇘/F⇙ Call c sat [p,R, G, q,a] ⟹
∀(c, p, R, G, q, a)∈Specs. Γ ⊨n⇘/F⇙ Call c sat [p,R, G, q,a]" by fact
have body:"∀(c, p, R, G, q, a)∈Θ. Γ ⊨Suc n⇘/F⇙ Call c sat [p,R, G, q,a]" by fact
then show ?case
proof-
{ fix c p R G q a
assume a000:"(c, p, R, G, q, a) ∈ Specs"
have ctxt_m:"∀(c, p, R, G, q, a)∈Θ. Γ ⊨n⇘/F⇙ Call c sat [p,R, G, q,a]"
using body cptn_mod_nest_mono unfolding com_validityn_def cpn_def
by (fastforce simp add: cpn_rule)
then have valid_Proc:"∀(c, p, R, G, q, a)∈Specs. Γ ⊨n⇘/F⇙ Call c sat [p,R, G, q,a]"
using hyp by auto
have Sta:"Sta p R" using a1 a000 by auto
have c_dom:" c ∈ dom Γ" using a1 a000 by auto
have guar:"∀s. (Normal s,Normal s)∈G" using a1 a000 by auto
let ?Θ'= "Θ ∪ Specs"
from valid_Proc ctxt_m
have "∀(c, p, R, G, q, a)∈?Θ'. Γ ⊨n⇘/F⇙ Call c sat [p,R, G, q,a]"
by fastforce
with valid_body
have valid_body_m:
"∀(c, p, R, G, q, a)∈Specs. Γ ⊨n⇘/F⇙ (the (Γ c)) sat [p,R, G, q,a]"
by (fastforce simp:com_cvalidityn_def)
then have valid_body:"Γ ⊨n⇘/F⇙ (the (Γ c)) sat [p,R, G, q,a]" using a000 by auto
then have "Γ ⊨Suc n⇘/F⇙ Call c sat [p,R, G, q,a]"
proof-
{ fix l s
assume a01:"(Γ,l)∈cpn (Suc n) Γ (Call c) s ∧ (Γ,l)∈ assum(p, R)"
then have "length l ≥ 1" unfolding cpn_def using CptnEmpty
by (metis (no_types, lifting) One_nat_def Product_Type.Collect_case_prodD Suc_leI length_greater_0_conv snd_conv)
moreover {
assume a02:"length l = 1"
then have "l = [(Call c,s)]"
proof -
have "l ! 0 = (LanguageCon.com.Call c, s)" using a01 unfolding cpn_def
by fastforce
then show ?thesis using a02
by (metis One_nat_def Suc_leI impossible_Cons
length_greater_0_conv list.size(3) neq_Nil_conv nth_Cons_0 zero_neq_one)
qed
then have "(Γ,l) ∈ comm(G, (q,a)) F" unfolding comm_def final_def by auto
}
moreover {assume a02:"length l > 1"
then obtain a1 ls where l:"l =(Call c, s)#a1#ls" using a01 unfolding cpn_def
apply auto
by (metis (no_types, hide_lams) One_nat_def Suc_eq_plus1 less_not_refl list.exhaust list.size(3) list.size(4) not_less_zero nth_Cons_0 prod.collapse)
have l_cptn:"(Γ,l)∈cptn" using a01 unfolding cpn_def
using cptn_eq_cptn_mod_nest by blast
then obtain m where min_call:"min_call m Γ l"
using cptn_eq_cptn_mod_set cptn_mod_cptn_mod_nest minimum_nest_call by blast
{ assume a03:"∀i<length l. fst (l!i) = Call c"
then have "(Γ,l) ∈ comm(G, (q,a)) F"
using no_comp_tran_no_final_comm[OF _ a03] a02 unfolding final_def
by fastforce
}
moreover{
assume a03:"¬(∀i<length l. fst (l!i) = Call c)"
then obtain i where i:"(i<length l ∧ fst (l!i) ≠ Call c)"
by auto
then obtain j where cfg_j:"fst (l!j) ≠ Call c ∧ (∀k<j. fst (l!k) = Call c)"
by (fast dest: exists_first_occ[of "λi. fst (l!i) ≠ Call c" i])
moreover have j:"j>0 ∧ j<length l" using l i calculation
by (metis gr0I fstI leI le_less_trans nth_Cons')
ultimately have step:"(Γ⊢⇩c (l!(j-1)) → (l!j))"
using l l_cptn
by (metis One_nat_def Suc_pred cptn_stepc_rtran diff_less not_eq_not_env prod.collapse
step_ce_elim_cases zero_less_one)
then obtain s' where j_1_cfg:"snd (l!(j-1)) = Normal s' ∧ s' ∈ p"
using cfg_j l a01[simplified l] j[simplified l] i Sta events_p[OF l_cptn[simplified l] _ _ _ Sta, of "j-1"]
by force
then have j_cfg:"l!j = (the (Γ c), Normal s')"
using cfg_j c_dom stepc_Normal_elim_cases(9) step
by (metis diff_less domIff j option.sel prod.collapse zero_less_one)
then have suc_n_call:"(Suc n,Γ, drop (j-1) l) ∈ cptn_mod_nest_call"
using a01 unfolding cpn_def
by (simp add: dropcptn_is_cptn1 j less_imp_diff_less)
have "(n,Γ, drop j l) ∈ cptn_mod_nest_call"
proof-
have "¬ (Γ⊢⇩c (l!(j-1)) →⇩e (l!j))" using step
by (metis etranE mod_env_not_component)
then have "(Suc n,Γ, (Call c, Normal s')#(the (Γ c),Normal s')#(drop (j+1) l)) ∈ cptn_mod_nest_call"
using a01 j step cfg_j j_cfg j_1_cfg suc_n_call
by (metis (no_types, lifting) Cons_nth_drop_Suc One_nat_def Suc_eq_plus1
Suc_less_eq Suc_pred diff_less less_SucI prod.collapse zero_less_one)
then have "(n,Γ, (the (Γ c),Normal s')#(drop (j+1) l)) ∈ cptn_mod_nest_call"
using cfg_j j_cfg elim_cptn_mod_nest_call_n_dec[OF _ ] c_dom by fastforce
then show ?thesis
by (metis Cons_nth_drop_Suc Suc_eq_plus1 j j_cfg)
qed
moreover have "(Γ, drop j l) ∈ assum(p,R)"
proof-
have "(Γ, take j l @ l ! j # drop (Suc j) l) ∈ assum (p, R)"
using conjunct2[OF a01] id_take_nth_drop[OF conjunct2[OF j]] by auto
then show ?thesis
using sub_assum_r[OF ] j_1_cfg l j_cfg j
by (metis Cons_nth_drop_Suc image_eqI snd_conv)
qed
ultimately have comm_drop:"(Γ, drop j l)∈ comm(G, (q,a)) F"
using valid_body j_cfg j unfolding com_validityn_def cpn_def
by fastforce
have "(Γ,l) ∈ comm(G, (q,a)) F"
proof-
have h:"∀j<length (take j l). fst ((take j l)!j) = (Call c)" using j cfg_j by fastforce
then have comm_take:"(Γ,take j l) ∈ comm(G, (q,a)) F"
using no_comp_tran_no_final_comm[of "take j l" "Call c"] j_1_cfg l j_cfg j cfg_j
unfolding final_def by auto
moreover have "(snd (last (take j l)), snd (drop j l ! 0)) ∈ G"
proof-
have "length (take j l) = j"using l j_1_cfg j j_cfg by auto
moreover have "(take j l)!(j-1) = l!(j-1)"
using l j_1_cfg j j_cfg by auto
ultimately have "last (take j l) = l!(j-1)"
using j by (metis last_conv_nth less_numeral_extra(3) list.size(3))
then show ?thesis using l j_1_cfg j j_cfg guar by auto
qed
ultimately show ?thesis using j_1_cfg j_cfg j cfg_j j l_cptn
comm_union[OF comm_take comm_drop] by fastforce
qed
} ultimately have "(Γ,l)∈comm(G, (q,a)) F" by auto
}ultimately have "(Γ,l)∈comm(G, (q,a)) F" by fastforce
} thus ?thesis unfolding com_validityn_def using cpn_rule2 by blast
qed
} thus ?case by fastforce
qed
qed
then show ?thesis using a0 unfolding com_cvalidityn_def by auto
qed
lemma Seq_env_P:assumes a0:"Γ⊢⇩c(Seq P Q,s) →⇩e (Seq P Q,t)"
shows "Γ⊢⇩c(P,s) →⇩e (P,t)"
using a0
by (metis env_not_normal_s snormal_enviroment)
lemma map_eq_state:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Seq c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) " and
a2:"l1=map (lift c2) l2"
shows
"∀i<length l1. snd (l1!i) = snd (l2!i)"
using a0 a1 a2 unfolding cp_def
by (simp add: snd_lift)
lemma map_eq_seq_c:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Seq c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) " and
a2:"l1=map (lift c2) l2"
shows
"∀i<length l1. fst (l1!i) = Seq (fst (l2!i)) c2"
proof -
{fix i
assume a3:"i<length l1"
have "fst (l1!i) = Seq (fst (l2!i)) c2"
using a0 a1 a2 a3 unfolding lift_def
by (simp add: case_prod_unfold)
}thus ?thesis by auto
qed
lemma same_env_seq_c:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Seq c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) " and
a2:"l1=map (lift c2) l2"
shows
"∀i. Suc i<length l2 ⟶ Γ⊢⇩c(l2!i) →⇩e (l2!(Suc i)) =
Γ⊢⇩c(l1!i) →⇩e (l1!(Suc i))"
proof -
have a0a:"(Γ,l1) ∈cptn ∧ l1!0 = ((Seq c1 c2),s)"
using a0 unfolding cp_def by blast
have a1a: "(Γ,l2) ∈cptn ∧ l2!0 = (c1,s)"
using a1 unfolding cp_def by blast
{
fix i
assume a3:"Suc i< length l2"
have "Γ⊢⇩c(l2!i) →⇩e (l2!(Suc i)) =
Γ⊢⇩c(l1!i) →⇩e (l1!(Suc i))"
proof
{
assume a4:"Γ⊢⇩c l2 ! i →⇩e l2 ! Suc i"
obtain c1i s1i c1si s1si where l1prod:"l1 ! i=(c1i,s1i) ∧ l1!Suc i = (c1si,s1si)"
by fastforce
obtain c2i s2i c2si s2si where l2prod:"l2 ! i=(c2i,s2i) ∧ l2!Suc i = (c2si,s2si)"
by fastforce
then have "c1i = (Seq c2i c2) ∧ c1si = (Seq c2si c2)"
using a0 a1 a2 a3 a4 map_eq_seq_c l1prod
by (metis Suc_lessD fst_conv length_map)
also have "s2i=s1i ∧ s2si=s1si"
using a0 a1 a4 a2 a3 l2prod map_eq_state l1prod
by (metis Suc_lessD nth_map snd_conv snd_lift)
ultimately show "Γ⊢⇩c l1 ! i →⇩e (l1 ! Suc i)"
using a4 l1prod l2prod
by (metis Env_n env_c_c' env_not_normal_s step_e.Env)
}
{
assume a4:"Γ⊢⇩c l1 ! i →⇩e l1 ! Suc i"
obtain c1i s1i c1si s1si where l1prod:"l1 ! i=(c1i,s1i) ∧ l1!Suc i = (c1si,s1si)"
by fastforce
obtain c2i s2i c2si s2si where l2prod:"l2 ! i=(c2i,s2i) ∧ l2!Suc i = (c2si,s2si)"
by fastforce
then have "c1i = (Seq c2i c2) ∧ c1si = (Seq c2si c2)"
using a0 a1 a2 a3 a4 map_eq_seq_c l1prod
by (metis Suc_lessD fst_conv length_map)
also have "s2i=s1i ∧ s2si=s1si"
using a0 a1 a4 a2 a3 l2prod map_eq_state l1prod
by (metis Suc_lessD nth_map snd_conv snd_lift)
ultimately show "Γ⊢⇩c l2 ! i →⇩e (l2 ! Suc i)"
using a4 l1prod l2prod
by (metis Env_n LanguageCon.com.inject(3) env_c_c' env_not_normal_s step_e.Env)
}
qed
}
thus ?thesis by auto
qed
lemma same_comp_seq_c:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Seq c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) " and
a2:"l1=map (lift c2) l2"
shows
"∀i. Suc i<length l2 ⟶ Γ⊢⇩c(l2!i) → (l2!(Suc i)) =
Γ⊢⇩c(l1!i) → (l1!(Suc i))"
proof -
have a0a:"(Γ,l1) ∈cptn ∧ l1!0 = ((Seq c1 c2),s)"
using a0 unfolding cp_def by blast
have a1a: "(Γ,l2) ∈cptn ∧ l2!0 = (c1,s)"
using a1 unfolding cp_def by blast
{
fix i
assume a3:"Suc i< length l2"
have "Γ⊢⇩c(l2!i) → (l2!(Suc i)) =
Γ⊢⇩c(l1!i) → (l1!(Suc i))"
proof
{
assume a4:"Γ⊢⇩c l2 ! i → l2 ! Suc i"
obtain c1i s1i c1si s1si where l1prod:"l1 ! i=(c1i,s1i) ∧ l1!Suc i = (c1si,s1si)"
by fastforce
obtain c2i s2i c2si s2si where l2prod:"l2 ! i=(c2i,s2i) ∧ l2!Suc i = (c2si,s2si)"
by fastforce
then have "c1i = (Seq c2i c2) ∧ c1si = (Seq c2si c2)"
using a0 a1 a2 a3 a4 map_eq_seq_c l1prod
by (metis Suc_lessD fst_conv length_map)
also have "s2i=s1i ∧ s2si=s1si"
using a0 a1 a4 a2 a3 l2prod map_eq_state l1prod
by (metis Suc_lessD nth_map snd_conv snd_lift)
ultimately show "Γ⊢⇩c l1 ! i → (l1 ! Suc i)"
using a4 l1prod l2prod
by (simp add: Seqc)
}
{
assume a4:"Γ⊢⇩c l1 ! i → l1 ! Suc i"
obtain c1i s1i c1si s1si where l1prod:"l1 ! i=(c1i,s1i) ∧ l1!Suc i = (c1si,s1si)"
by fastforce
obtain c2i s2i c2si s2si where l2prod:"l2 ! i=(c2i,s2i) ∧ l2!Suc i = (c2si,s2si)"
by fastforce
then have "c1i = (Seq c2i c2) ∧ c1si = (Seq c2si c2)"
using a0 a1 a2 a3 a4 map_eq_seq_c l1prod
by (metis Suc_lessD fst_conv length_map)
also have "s2i=s1i ∧ s2si=s1si"
using a0 a1 a4 a2 a3 l2prod map_eq_state l1prod
by (metis Suc_lessD nth_map snd_conv snd_lift)
ultimately show "Γ⊢⇩c l2 ! i → (l2 ! Suc i)"
using a4 l1prod l2prod stepc_elim_cases_Seq_Seq
by auto
}
qed
}
thus ?thesis by auto
qed
lemma assum_map:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Seq c1 c2) s) ∧ ((Γ,l1) ∈ assum(p, R))" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) " and
a2:"l1=map (lift c2) l2"
shows
"((Γ,l2) ∈ assum(p, R))"
proof -
have a3: "∀i. Suc i<length l2 ⟶ Γ⊢⇩c(l2!i) →⇩e (l2!(Suc i)) =
Γ⊢⇩c(l1!i) →⇩e (l1!(Suc i))"
using a0 a1 a2 same_env_seq_c by fastforce
have pair_Γl1:"fst (Γ,l1) = Γ ∧ snd (Γ,l1) = l1" by fastforce
have pair_Γl2:"fst (Γ,l2) = Γ ∧ snd (Γ,l2) = l2" by fastforce
have drop_k_s:"l2!0 = (c1,s)" using a1 cp_def by blast
have eq_length:"length l1 = length l2" using a2 by auto
obtain s' where normal_s:"s = Normal s'"
using a0 unfolding cp_def assum_def by fastforce
then have p1:"s'∈p" using a0 unfolding cp_def assum_def by fastforce
show ?thesis
proof -
let ?c= "(Γ,l2)"
have l:"snd((snd ?c!0)) ∈ Normal ` (p)"
using p1 drop_k_s a1 normal_s unfolding cp_def by auto
{fix i
assume a00:"Suc i<length (snd ?c)"
assume a11:"(fst ?c)⊢⇩c((snd ?c)!i) →⇩e ((snd ?c)!(Suc i))"
have "(snd((snd ?c)!i), snd((snd ?c)!(Suc i))) ∈ R"
using a0 a1 a2 a3 map_eq_state unfolding assum_def
using a00 a11 eq_length by fastforce
} thus "(Γ, l2) ∈ assum (p, R)"
using l unfolding assum_def by fastforce
qed
qed
lemma comm_map':
assumes
a0:"(Γ,l1) ∈ (cp Γ (Seq c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) ∧ (Γ, l2)∈ comm(G, (q,a)) F" and
a2:"l1=map (lift c2) l2"
shows
"snd (last l1) ∉ Fault ` F ⟶(Suc k < length l1 ⟶
Γ⊢⇩c(l1!k) → (l1!(Suc k)) ⟶
(snd(l1!k), snd(l1!(Suc k))) ∈ G) ∧
(fst (last l1) = (Seq c c2) ∧ final (c, snd (last l1)) ⟶
(fst (last l1) = (Seq Skip c2) ∧
(snd (last l1) ∈ Normal ` q) ∨
(fst (last l1) = (Seq Throw c2) ∧
snd (last l1) ∈ Normal ` (a))))
"
proof -
have a3:"∀i. Suc i<length l2 ⟶ Γ⊢⇩c(l2!i) → (l2!(Suc i)) =
Γ⊢⇩c(l1!i) → (l1!(Suc i))"
using a0 a1 a2 same_comp_seq_c
by fastforce
have pair_Γl1:"fst (Γ,l1) = Γ ∧ snd (Γ,l1) = l1" by fastforce
have pair_Γl2:"fst (Γ,l2) = Γ ∧ snd (Γ,l2) = l2" by fastforce
have drop_k_s:"l2!0 = (c1,s)" using a1 cp_def by blast
have eq_length:"length l1 = length l2" using a2 by auto
then have len0:"length l1>0" using a0 unfolding cp_def
using Collect_case_prodD drop_k_s eq_length by auto
then have l1_not_empty:"l1≠[]" by auto
then have l2_not_empty:"l2 ≠ []" using a2 by blast
have last_lenl1:"last l1 = l1!((length l1) -1)"
using last_conv_nth l1_not_empty by auto
have last_lenl2:"last l2 = l2!((length l2) -1)"
using last_conv_nth l2_not_empty by auto
have a03:"snd (last l2) ∉ Fault ` F ⟶(∀i ns ns'.
Suc i<length (snd (Γ, l2)) ⟶
fst (Γ, l2)⊢⇩c((snd (Γ, l2))!i) → ((snd (Γ, l2))!(Suc i)) ⟶
(snd((snd (Γ, l2))!i), snd((snd (Γ, l2))!(Suc i))) ∈ G) ∧
(final (last (snd (Γ, l2))) ⟶
((fst (last (snd (Γ, l2))) = Skip ∧
snd (last (snd (Γ, l2))) ∈ Normal ` q)) ∨
(fst (last (snd (Γ, l2))) = Throw ∧
snd (last (snd (Γ, l2))) ∈ Normal ` (a)))"
using a1 unfolding comm_def by fastforce
show ?thesis unfolding comm_def
proof -
{ fix k ns ns'
assume a00a:"snd (last l1) ∉ Fault ` F"
assume a00:"Suc k < length l1"
then have "k ≤ length l1" using a2 by fastforce
have a00:"Suc k < length l2" using eq_length a00 by fastforce
then have a00a:"snd (last l2) ∉ Fault ` F"
proof-
have "snd (l1!((length l1) -1)) = snd (l2!((length l2) -1))"
using a2 a1 a0 map_eq_state eq_length l2_not_empty last_snd
by fastforce
then have "snd(last l2) = snd (last l1)"
using last_lenl1 last_lenl2 by auto
thus ?thesis using a00a by auto
qed
then have "snd (last l1) ∉ Fault ` F ⟶Γ⊢⇩c(l1!k) → (l1!(Suc k)) ⟶
(snd((snd (Γ, l1))!k), snd((snd (Γ, l1))!(Suc k))) ∈ G"
using pair_Γl1 pair_Γl2 a00 a03 a3 eq_length a00a
by (metis Suc_lessD a0 a1 a2 map_eq_state)
} note l=this
{
assume a00: "fst (last l1) = (Seq c c2) ∧ final (c, snd (last l1))" and
a01:"snd (last (l1)) ∉ Fault ` F"
then have c:"c=Skip ∨ c = Throw"
unfolding final_def by auto
then have fst_last_l2:"fst (last l2) = c"
using last_lenl1 a00 l1_not_empty eq_length len0 a2 last_conv_nth last_lift
by fastforce
also have last_eq:"snd (last l2) = snd (last l1)"
using l2_not_empty a2 last_conv_nth last_lenl1 last_snd
by fastforce
ultimately have "final (fst (last l2),snd (last l2))"
using a00 by auto
then have "final (last l2)" by auto
also have "snd (last (l2)) ∉ Fault ` F"
using last_eq a01 by auto
ultimately have "(fst (last l2)) = Skip ∧
snd (last l2) ∈ Normal ` q ∨
(fst (last l2) = Throw ∧
snd (last l2) ∈ Normal ` (a))"
using a03 by auto
then have "(fst (last l1) = (Seq Skip c2) ∧
snd (last l1) ∈ Normal ` q) ∨
(fst (last l1) = (Seq Throw c2) ∧
snd (last l1) ∈ Normal ` (a))"
using last_eq fst_last_l2 a00 by force
}
thus ?thesis using l by auto qed
qed
lemma comm_map'':
assumes
a0:"(Γ,l1) ∈ (cp Γ (Seq c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) ∧ (Γ, l2)∈ comm(G, (q,a)) F" and
a2:"l1=map (lift c2) l2"
shows
"snd (last l1) ∉ Fault ` F ⟶ ((Suc k < length l1 ⟶
Γ⊢⇩c(l1!k) → (l1!(Suc k)) ⟶
(snd(l1!k), snd(l1!(Suc k))) ∈ G) ∧
(final (last l1) ⟶
(fst (last l1) = Skip ∧
(snd (last l1) ∈ Normal ` r) ∨
(fst (last l1) = Throw ∧
snd (last l1) ∈ Normal ` (a)))))
"
proof -
have a3:"∀i. Suc i<length l2 ⟶ Γ⊢⇩c(l2!i) → (l2!(Suc i)) =
Γ⊢⇩c(l1!i) → (l1!(Suc i))"
using a0 a1 a2 same_comp_seq_c
by fastforce
have pair_Γl1:"fst (Γ,l1) = Γ ∧ snd (Γ,l1) = l1" by fastforce
have pair_Γl2:"fst (Γ,l2) = Γ ∧ snd (Γ,l2) = l2" by fastforce
have drop_k_s:"l2!0 = (c1,s)" using a1 cp_def by blast
have eq_length:"length l1 = length l2" using a2 by auto
then have len0:"length l1>0" using a0 unfolding cp_def
using Collect_case_prodD drop_k_s eq_length by auto
then have l1_not_empty:"l1≠[]" by auto
then have l2_not_empty:"l2 ≠ []" using a2 by blast
have last_lenl1:"last l1 = l1!((length l1) -1)"
using last_conv_nth l1_not_empty by auto
have last_lenl2:"last l2 = l2!((length l2) -1)"
using last_conv_nth l2_not_empty by auto
have a03:"snd (last l2) ∉ Fault ` F ⟶(∀i ns ns'.
Suc i<length (snd (Γ, l2)) ⟶
fst (Γ, l2)⊢⇩c((snd (Γ, l2))!i) → ((snd (Γ, l2))!(Suc i)) ⟶
(snd((snd (Γ, l2))!i), snd((snd (Γ, l2))!(Suc i))) ∈ G) ∧
(final (last (snd (Γ, l2))) ⟶
((fst (last (snd (Γ, l2))) = Skip ∧
snd (last (snd (Γ, l2))) ∈ Normal ` q)) ∨
(fst (last (snd (Γ, l2))) = Throw ∧
snd (last (snd (Γ, l2))) ∈ Normal ` (a)))"
using a1 unfolding comm_def by fastforce
show ?thesis unfolding comm_def
proof -
{ fix k ns ns'
assume a00a:"snd (last l1) ∉ Fault ` F"
assume a00:"Suc k < length l1"
then have "k ≤ length l1" using a2 by fastforce
have a00:"Suc k < length l2" using eq_length a00 by fastforce
then have a00a:"snd (last l2) ∉ Fault ` F"
proof-
have "snd (l1!((length l1) -1)) = snd (l2!((length l2) -1))"
using a2 a1 a0 map_eq_state eq_length l2_not_empty last_snd
by fastforce
then have "snd(last l2) = snd (last l1)"
using last_lenl1 last_lenl2 by auto
thus ?thesis using a00a by auto
qed
then have "Γ⊢⇩c(l1!k) → (l1!(Suc k)) ⟶
(snd((snd (Γ, l1))!k), snd((snd (Γ, l1))!(Suc k))) ∈ G"
using pair_Γl1 pair_Γl2 a00 a03 a3 eq_length a00a
by (metis (no_types,lifting) a2 Suc_lessD nth_map snd_lift)
} note l= this
{
assume a00: "final (last l1)"
then have c:"fst (last l1)=Skip ∨ fst (last l1) = Throw"
unfolding final_def by auto
moreover have "fst (last l1) = Seq (fst (last l2)) c2"
using a2 last_lenl1 eq_length
proof -
have "last l2 = l2 ! (length l2 - 1)"
using l2_not_empty last_conv_nth by blast
then show ?thesis
by (metis One_nat_def a2 l2_not_empty last_lenl1 last_lift)
qed
ultimately have False by simp
} thus ?thesis using l by auto qed
qed
lemma comm_map:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Seq c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) ∧ (Γ, l2)∈ comm(G, (q,a)) F" and
a2:"l1=map (lift c2) l2"
shows
"(Γ, l1)∈ comm(G, (r,a)) F"
proof -
{fix i
have "snd (last l1) ∉ Fault ` F ⟶(Suc i < length (l1) ⟶
Γ⊢⇩c (l1 ! i) → (l1 ! (Suc i)) ⟶
(snd (l1 ! i), snd (l1 ! Suc i)) ∈ G) ∧
(SmallStepCon.final (last l1) ⟶
fst (last l1) = LanguageCon.com.Skip ∧
snd (last l1) ∈ Normal ` r ∨
fst (last l1) = LanguageCon.com.Throw ∧
snd (last l1) ∈ Normal ` a) "
using comm_map''[of Γ l1 c1 c2 s l2 G q a F i r] a0 a1 a2
by fastforce
} then show ?thesis using comm_def unfolding comm_def by force
qed
lemma Seq_sound1:
assumes
a0:"(n,Γ,x)∈cptn_mod_nest_call" and
a1:"x!0 = ((Seq P Q),s)" and
a2:"∀i<length x. fst (x!i)≠ Q" and
a3:"¬ final (last x)" and
a4:"env_tran_right Γ x rely" and
a5:"snd (x!0)∈ Normal ` p ∧ Sta p rely ∧ Sta a rely " and
a6: "Γ ⊨n⇘/F⇙ P sat [p, rely, G, q,a]"
shows
"∃xs. (Γ,xs) ∈ cpn n Γ P s ∧ x = map (lift Q) xs"
using a0 a1 a2 a3 a4 a5 a6
proof (induct arbitrary: P s p)
case (CptnModNestOne n Γ C s1)
then have "(Γ, [(P,s)]) ∈ cpn n Γ P s ∧ [(C, s1)] = map (lift Q) [(P,s)]"
unfolding cpn_def lift_def
by (simp add: cptn_mod_nest_call.CptnModNestOne)
thus ?case by fastforce
next
case (CptnModNestEnv Γ C s1 t1 n xsa)
then have C:"C=Seq P Q" unfolding lift_def by fastforce
have "∃xs. (Γ, xs) ∈ cpn n Γ P t1 ∧ (C, t1) # xsa = map (lift Q) xs"
proof -
have "((C, t1) # xsa) ! 0 = (LanguageCon.com.Seq P Q, t1)" using C by auto
moreover have "∀i<length ((C, t1) # xsa). fst (((C, t1) # xsa) ! i) ≠ Q"
using CptnModNestEnv(5) by fastforce
moreover have "¬ SmallStepCon.final (last ((C, t1) # xsa))" using CptnModNestEnv(6)
by fastforce
moreover have "snd (((C, t1) # xsa) ! 0) ∈ Normal ` p"
using CptnModNestEnv(8) CptnModNestEnv(1) CptnModNestEnv(7)
unfolding env_tran_right_def Sta_def by fastforce
ultimately show ?thesis
using CptnModNestEnv(3) CptnModNestEnv(7) CptnModNestEnv(8) CptnModNestEnv(9) env_tran_tail by blast
qed
then obtain xs where hi:"(Γ, xs) ∈ cpn n Γ P t1 ∧ (C, t1) # xsa = map (lift Q) xs"
by fastforce
have s1_s:"s1=s" using CptnModNestEnv unfolding cpn_def by auto
obtain xsa' where xs:"xs=((P,t1)#xsa') ∧ (n, Γ,((P,t1)#xsa'))∈cptn_mod_nest_call ∧ (C, t1) # xsa = map (lift Q) ((P,t1)#xsa')"
using hi unfolding cpn_def by fastforce
have env_tran:"Γ⊢⇩c(P,s1) →⇩e (P,t1)" using CptnModNestEnv Seq_env_P by (metis fst_conv nth_Cons_0)
then have "(n, Γ,(P,s1)#(P,t1)#xsa')∈cptn_mod_nest_call"
using xs env_tran cptn_mod_nest_call.CptnModNestEnv by fastforce
then have "(Γ,(P,s1)#(P,t1)#xsa') ∈ cpn n Γ P s"
using cpn_def s1_s by fastforce
moreover have "(C,s1)#(C, t1) # xsa = map (lift Q) ((P,s1)#(P,t1)#xsa')"
using xs C unfolding lift_def by fastforce
ultimately show ?case by auto
next
case (CptnModNestSkip)
thus ?case by (metis SmallStepCon.redex_not_Seq fst_conv nth_Cons_0)
next
case (CptnModNestThrow)
thus ?case by (metis SmallStepCon.redex_not_Seq fst_conv nth_Cons_0)
next
case (CptnModNestSeq1 n Γ P0 sa xsa zs P1)
then have a1:"LanguageCon.com.Seq P Q = LanguageCon.com.Seq P0 P1"
by fastforce
have f1: "sa = s"
using CptnModNestSeq1.prems(1) by force
have f2: "P = P0 ∧ Q = P1" using a1 by auto
hence "(Γ, (P0, sa) # xsa) ∈ cpn n Γ P s"
using f2 f1 CptnModNestSeq1.hyps(1) by (simp add: cpn_def)
thus ?case
using Cons_lift CptnModNestSeq1.hyps(3) a1 by fastforce
next
case (CptnModNestSeq2 n Γ P0 sa xsa P1 ys zs)
then have "P0 = P ∧ P1 = Q" by auto
then obtain i where zs:"fst (zs!i) = Q ∧ (i< (length zs))" using CptnModNestSeq2
by (metis (no_types, lifting) add_diff_cancel_left' fst_conv length_Cons length_append nth_append_length zero_less_Suc zero_less_diff)
then have "Suc i< length ((Seq P0 P1,sa)#zs)" by fastforce
then have "fst (((Seq P0 P1, sa) # zs)!Suc i) = Q" using zs by fastforce
thus ?case using CptnModNestSeq2(8) zs by auto
next
case (CptnModNestSeq3 n Γ P1 sa xsa s' ys zs Q1 )
have s'_a:"s' ∈ a"
proof -
have cpP1:"(Γ, (P1, Normal sa) # xsa) ∈ cpn n Γ P1 (Normal sa)"
using CptnModNestSeq3.hyps(1) unfolding cpn_def by fastforce
then have cpP1':"(Γ, (P1, Normal sa) # xsa) ∈ cp Γ P1 (Normal sa)"
using CptnModNestSeq3.hyps(1) cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod
unfolding cp_def by fastforce
have map:"((Seq P1 Q1), Normal sa)#(map (lift Q1) xsa) = map (lift Q1) ((P1, Normal sa) # xsa) "
using CptnModSeq3 by (simp add: Cons_lift)
then
have "(Γ,((LanguageCon.com.Seq P1 Q1, Normal sa) # (map (lift Q1) xsa))) ∈ assum (p,rely)"
proof -
have "env_tran_right Γ ((LanguageCon.com.Seq P1 Q1, Normal sa) # (map (lift Q1) xsa)) rely"
using CptnModNestSeq3(11) CptnModNestSeq3(7) map
by (metis (no_types) Cons_lift_append CptnModNestSeq3.hyps(7) CptnModNestSeq3.prems(4) env_tran_subr)
thus ?thesis using CptnModNestSeq3(12)
unfolding assum_def env_tran_right_def by fastforce
qed
moreover have "(Γ,((Seq P1 Q1), Normal sa)#(map (lift Q1) xsa)) ∈ cpn n Γ (Seq P1 Q1) (Normal sa)"
using CptnModNestSeq3(7) CptnModNestSeq3.hyps(1) cptn_mod_nest_call.CptnModNestSeq1
unfolding cpn_def by fastforce
then have "(Γ,((Seq P1 Q1), Normal sa)#(map (lift Q1) xsa)) ∈ cp Γ (Seq P1 Q1) (Normal sa)"
using CptnModNestSeq3.hyps(1) cptn_eq_cptn_mod_set cptn_mod.CptnModSeq1 cptn_mod_nest_cptn_mod
unfolding cp_def by fastforce
ultimately have "(Γ, (P1, Normal sa) # xsa) ∈ assum (p,rely)"
using assum_map map cpP1' by fastforce
then have "(Γ, (P1, Normal sa) # xsa) ∈ comm (G,(q,a)) F"
using cpP1 CptnModNestSeq3(13) CptnModNestSeq3.prems(1) unfolding com_validityn_def by auto
thus ?thesis
using CptnModNestSeq3(3) CptnModNestSeq3(4)
unfolding comm_def final_def by fastforce
qed
have "final (last ((LanguageCon.com.Throw, Normal s')# ys))"
proof -
have cptn_mod:"(n, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call"
using CptnModNestSeq3(5) by (simp add: cptn_eq_cptn_mod_set)
then have cptn:"(Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn"
using cptn_eq_cptn_mod_nest by auto
moreover have throw_0:"((LanguageCon.com.Throw, Normal s') # ys)!0 = (Throw, Normal s') ∧ 0 < length((LanguageCon.com.Throw, Normal s') # ys)"
by force
moreover have last:"last ((LanguageCon.com.Throw, Normal s') # ys) = ((LanguageCon.com.Throw, Normal s') # ys)!((length ((LanguageCon.com.Throw, Normal s') # ys)) - 1)"
using last_conv_nth by auto
moreover have env_tran:"env_tran_right Γ ((LanguageCon.com.Throw, Normal s') # ys) rely"
using CptnModNestSeq3(11) CptnModNestSeq3(7) env_tran_subl env_tran_tail by blast
ultimately obtain st' where "fst (last ((LanguageCon.com.Throw, Normal s') # ys)) = Throw ∧
snd (last ((LanguageCon.com.Throw, Normal s') # ys)) = Normal st'"
using zero_throw_all_throw[of Γ "((Throw, Normal s') # ys)" "s'" "(length ((Throw, Normal s') # ys))-1" a rely]
s'_a CptnModNestSeq3(11) CptnModNestSeq3(12) by fastforce
thus ?thesis using CptnModNestSeq3(10) final_def by blast
qed
thus ?case using CptnModNestSeq3(10) CptnModNestSeq3(7)
by force
qed (auto)
lemma Seq_sound2:
assumes
a0:"(Γ,x)∈cptn_mod" and
a1:"x!0 = ((Seq P Q),s)" and
a2:"∀i<length x. fst (x!i)≠ Q" and
a3:"fst (last x) = Throw ∧ snd (last x) = Normal s'" and
a4:"env_tran_right Γ x rely"
shows
"∃xs s' ys. (Γ,xs) ∈ cp Γ P s ∧ x = ((map (lift Q) xs)@((Throw, Normal s')#ys))"
using a0 a1 a2 a3 a4
proof (induct arbitrary: P s s')
case (CptnModOne Γ C s1)
then have "(Γ, [(P,s)]) ∈ cp Γ P s ∧ [(C, s1)] = map (lift Q) [(P,s)]@[(Throw, Normal s')]"
unfolding cp_def lift_def by (simp add: cptn.CptnOne)
thus ?case by fastforce
next
case (CptnModEnv Γ C s1 t1 xsa)
then have C:"C=Seq P Q" unfolding lift_def by fastforce
have "∃xs s' ys. (Γ, xs) ∈ cp Γ P t1 ∧ (C, t1) # xsa = map (lift Q) xs@((Throw, Normal s')#ys)"
proof -
have "((C, t1) # xsa) ! 0 = (LanguageCon.com.Seq P Q, t1)" using C by auto
moreover have "∀i<length ((C, t1) # xsa). fst (((C, t1) # xsa) ! i) ≠ Q"
using CptnModEnv(5) by fastforce
moreover have "fst (last ((C, t1) # xsa)) = Throw ∧ snd (last ((C, t1) # xsa)) = Normal s'" using CptnModEnv(6)
by fastforce
ultimately show ?thesis
using CptnModEnv(3) CptnModEnv(7) env_tran_tail by blast
qed
then obtain xs s'' ys where hi:"(Γ, xs) ∈ cp Γ P t1 ∧ (C, t1) # xsa = map (lift Q) xs@((Throw, Normal s'')#ys)"
by fastforce
have s1_s:"s1=s" using CptnModEnv unfolding cp_def by auto
have "∃xsa' s'' ys. xs=((P,t1)#xsa') ∧ (Γ,((P,t1)#xsa'))∈cptn ∧ (C, t1) # xsa = map (lift Q) ((P,t1)#xsa')@((Throw, Normal s'')#ys)"
using hi unfolding cp_def
proof -
have "(Γ,xs)∈cptn ∧ xs!0 = (P,t1)" using hi unfolding cp_def by fastforce
moreover then have "xs≠[]" using cptn.simps by fastforce
ultimately obtain xsa' where "xs=((P,t1)#xsa')" using SmallStepCon.nth_tl by fastforce
thus ?thesis
using hi using ‹(Γ, xs) ∈ cptn ∧ xs ! 0 = (P, t1)› by auto
qed
then obtain xsa' s'' ys where xs:"xs=((P,t1)#xsa') ∧ (Γ,((P,t1)#xsa'))∈cptn ∧ (C, t1) # xsa = map (lift Q) ((P,t1)#xsa')@((Throw, Normal s'')#ys)"
by fastforce
have env_tran:"Γ⊢⇩c(P,s1) →⇩e (P,t1)" using CptnModEnv Seq_env_P by (metis fst_conv nth_Cons_0)
then have "(Γ,(P,s1)#(P,t1)#xsa')∈cptn" using xs env_tran CptnEnv by fastforce
then have "(Γ,(P,s1)#(P,t1)#xsa') ∈ cp Γ P s"
using cp_def s1_s by fastforce
moreover have "(C,s1)#(C, t1) # xsa = map (lift Q) ((P,s1)#(P,t1)#xsa')@((Throw, Normal s'')#ys)"
using xs C unfolding lift_def by fastforce
ultimately show ?case by auto
next
case (CptnModSkip)
thus ?case by (metis SmallStepCon.redex_not_Seq fst_conv nth_Cons_0)
next
case (CptnModThrow)
thus ?case by (metis SmallStepCon.redex_not_Seq fst_conv nth_Cons_0)
next
case (CptnModSeq1 Γ P0 sa xsa zs P1)
thus ?case
proof -
have a1:"∀c p. fst (case p of (ca::('s, 'a, 'd,'e) LanguageCon.com, x::('s, 'd) xstate) ⇒
(LanguageCon.com.Seq ca c, x)) = LanguageCon.com.Seq (fst p) c"
by simp
then have "[] = xsa"
proof -
have "[] ≠ zs"
using CptnModSeq1 by force
then show ?thesis
by (metis (no_types) LanguageCon.com.distinct(71) One_nat_def CptnModSeq1(3,6)
last.simps last_conv_nth last_lift)
qed
then have "∀c. Throw = c ∨ [] = zs"
using CptnModSeq1(3) by fastforce
then show ?thesis
using CptnModSeq1.prems(3) by force
qed
next
case (CptnModSeq2 Γ P0 sa xsa P1 ys zs)
then have "P0 = P ∧ P1 = Q" by auto
then obtain i where zs:"fst (zs!i) = Q ∧ (i< (length zs))" using CptnModSeq2
by (metis (no_types, lifting) add_diff_cancel_left' fst_conv length_Cons length_append nth_append_length zero_less_Suc zero_less_diff)
then have "Suc i< length ((Seq P0 P1,sa)#zs)" by fastforce
then have "fst (((Seq P0 P1, sa) # zs)!Suc i) = Q" using zs by fastforce
thus ?case using CptnModSeq2(8) zs by auto
next
case (CptnModSeq3 Γ P0 sa xsa s'' ys zs P1)
then have "P0 = P ∧ P1 = Q ∧ s=Normal sa" by auto
moreover then have "(Γ, (P0, Normal sa) # xsa)∈ cp Γ P s"
using CptnModSeq3(1)
by (simp add: cp_def cptn_eq_cptn_mod_set)
moreover have "last zs=(Throw, Normal s')" using CptnModSeq3(10) CptnModSeq3.hyps(7)
by (simp add: prod_eqI)
ultimately show ?case using CptnModSeq3(7)
using Cons_lift_append by blast
qed (auto)
lemma Seq_sound2':
assumes
a0:"(n,Γ,x)∈cptn_mod_nest_call" and
a1:"x!0 = ((Seq P Q),s)" and
a2:"∀i<length x. fst (x!i)≠ Q" and
a3:"fst (last x) = Throw ∧ snd (last x) = Normal s'" and
a4:"env_tran_right Γ x rely"
shows
"∃xs s' ys. (Γ,xs) ∈ cpn n Γ P s ∧ x = ((map (lift Q) xs)@((Throw, Normal s')#ys))"
using a0 a1 a2 a3 a4
proof (induct arbitrary: P s s')
case (CptnModNestOne n Γ C s1)
then have "(Γ, [(P,s)]) ∈ cpn n Γ P s ∧ [(C, s1)] = map (lift Q) [(P,s)]@[(Throw, Normal s')]"
unfolding cp_def lift_def by (simp add: cptn.CptnOne)
thus ?case by fastforce
next
case (CptnModNestEnv Γ C s1 t1 n xsa)
then have C:"C=Seq P Q" unfolding lift_def by fastforce
have "∃xs s' ys. (Γ, xs) ∈ cpn n Γ P t1 ∧ (C, t1) # xsa = map (lift Q) xs@((Throw, Normal s')#ys)"
proof -
have "((C, t1) # xsa) ! 0 = (LanguageCon.com.Seq P Q, t1)" using C by auto
moreover have "∀i<length ((C, t1) # xsa). fst (((C, t1) # xsa) ! i) ≠ Q"
using CptnModNestEnv(5) by fastforce
moreover have "fst (last ((C, t1) # xsa)) = Throw ∧ snd (last ((C, t1) # xsa)) = Normal s'"
using CptnModNestEnv(6)
by fastforce
ultimately show ?thesis
using CptnModNestEnv(3) CptnModNestEnv(7) env_tran_tail by blast
qed
then obtain xs s'' ys where hi:"(Γ, xs) ∈ cpn n Γ P t1 ∧ (C, t1) # xsa = map (lift Q) xs@((Throw, Normal s'')#ys)"
by fastforce
have s1_s:"s1=s" using CptnModNestEnv unfolding cp_def by auto
have "∃xsa' s'' ys. xs=((P,t1)#xsa') ∧ (n, Γ,((P,t1)#xsa'))∈cptn_mod_nest_call ∧ (C, t1) # xsa = map (lift Q) ((P,t1)#xsa')@((Throw, Normal s'')#ys)"
using hi unfolding cp_def
proof -
have "(n, Γ,xs)∈cptn_mod_nest_call ∧ xs!0 = (P,t1)" using hi unfolding cpn_def by fastforce
moreover then have "xs≠[]" using cptn_mod_nest_call.simps by fastforce
ultimately obtain xsa' where "xs=((P,t1)#xsa')" using SmallStepCon.nth_tl by fastforce
thus ?thesis
using hi using ‹(n, Γ, xs) ∈ cptn_mod_nest_call ∧ xs ! 0 = (P, t1)› by auto
qed
then obtain xsa' s'' ys where xs:"xs=((P,t1)#xsa') ∧ (n, Γ,((P,t1)#xsa'))∈cptn_mod_nest_call ∧
(C, t1) # xsa = map (lift Q) ((P,t1)#xsa')@((Throw, Normal s'')#ys)"
by fastforce
have env_tran:"Γ⊢⇩c(P,s1) →⇩e (P,t1)" using CptnModNestEnv Seq_env_P by (metis fst_conv nth_Cons_0)
then have "(n, Γ,(P,s1)#(P,t1)#xsa')∈cptn_mod_nest_call" using xs env_tran cptn_mod_nest_call.CptnModNestEnv by blast
then have "(Γ,(P,s1)#(P,t1)#xsa') ∈ cpn n Γ P s"
using cpn_def s1_s by fastforce
moreover have "(C,s1)#(C, t1) # xsa = map (lift Q) ((P,s1)#(P,t1)#xsa')@((Throw, Normal s'')#ys)"
using xs C unfolding lift_def by fastforce
ultimately show ?case by auto
next
case (CptnModNestSkip)
thus ?case by (metis SmallStepCon.redex_not_Seq fst_conv nth_Cons_0)
next
case (CptnModNestThrow)
thus ?case by (metis SmallStepCon.redex_not_Seq fst_conv nth_Cons_0)
next
case (CptnModNestSeq1 n Γ P0 sa xsa zs P1)
thus ?case
proof -
have a1:"∀c p. fst (case p of (ca::('s, 'a, 'd,'e) LanguageCon.com, x::('s, 'd) xstate) ⇒
(LanguageCon.com.Seq ca c, x)) = LanguageCon.com.Seq (fst p) c"
by simp
then have "[] = xsa"
proof -
have "[] ≠ zs"
using CptnModNestSeq1 by force
then show ?thesis
by (metis (no_types) LanguageCon.com.distinct(71) One_nat_def CptnModNestSeq1(3,6)
last.simps last_conv_nth last_lift)
qed
then have "∀c. Throw = c ∨ [] = zs"
using CptnModNestSeq1(3) by fastforce
then show ?thesis
using CptnModNestSeq1.prems(3) by force
qed
next
case (CptnModNestSeq2 n Γ P0 sa xsa P1 ys zs)
then have "P0 = P ∧ P1 = Q" by auto
then obtain i where zs:"fst (zs!i) = Q ∧ (i< (length zs))" using CptnModNestSeq2
by (metis (no_types, lifting) add_diff_cancel_left' fst_conv length_Cons length_append nth_append_length zero_less_Suc zero_less_diff)
then have "Suc i< length ((Seq P0 P1,sa)#zs)" by fastforce
then have "fst (((Seq P0 P1, sa) # zs)!Suc i) = Q" using zs by fastforce
thus ?case using CptnModNestSeq2(8) zs by auto
next
case (CptnModNestSeq3 n Γ P0 sa xsa s'' ys zs P1)
then have "P0 = P ∧ P1 = Q ∧ s=Normal sa" by auto
moreover then have "(Γ, (P0, Normal sa) # xsa)∈ cpn n Γ P s"
using CptnModNestSeq3(1)
by (simp add: cpn_def)
moreover have "last zs=(Throw, Normal s')" using CptnModNestSeq3(10) CptnModNestSeq3.hyps(7)
by (simp add: prod_eqI)
ultimately show ?case using CptnModNestSeq3(7)
using Cons_lift_append by blast
qed (auto)
lemma Last_Skip_Exist_Final:
assumes
a0:"(Γ,x)∈cptn" and
a1:"x!0 = ((Seq P Q),s)" and
a2:"∀i<length x. fst (x!i)≠ Q" and
a3:"fst(last x) = Skip"
shows
"∃c s' i. i<length x ∧ x!i = (Seq c Q,s') ∧ final (c,s')"
using a0 a1 a2 a3
proof (induct arbitrary: P s)
case (CptnOne Γ c s1) thus ?case by fastforce
next
case (CptnEnv Γ C st t xsa)
thus ?case
proof -
have "LanguageCon.com.Seq P Q = C"
using CptnEnv.prems(1) by auto
then show ?thesis
using CptnEnv.hyps(3) CptnEnv.prems(2) CptnEnv.prems(3) by fastforce
qed
next
case (CptnComp Γ C st C' st' xsa)
then have c_seq:"C = (Seq P Q) ∧ st = s" by force
from CptnComp show ?case proof(cases)
case (Seqc P1 P1' P2)
then have "∃c s' i. i < length ((C', st') # xsa) ∧
((C', st') # xsa) ! i = (LanguageCon.com.Seq c Q, s') ∧
SmallStepCon.final (c, s')"
using CptnComp last.simps by fastforce
thus ?thesis by fastforce
next
case (SeqThrowc C2 s')
thus ?thesis
proof -
have "LanguageCon.com.Seq LanguageCon.com.Throw Q = C"
using ‹C = LanguageCon.com.Seq LanguageCon.com.Throw C2› c_seq by blast
then show ?thesis
using ‹st = Normal s'› unfolding final_def by force
qed
next
case (FaultPropc) thus ?thesis
using c_seq redex_not_Seq by blast
next
case (StuckPropc) thus ?thesis
using c_seq redex_not_Seq by blast
next
case (AbruptPropc) thus ?thesis
using c_seq redex_not_Seq by blast
qed (auto)
qed
lemma Seq_sound3:
assumes
a0:"(n,Γ,x)∈cptn_mod_nest_call" and
a1:"x!0 = ((Seq P Q),s)" and
a2:"∀i<length x. fst (x!i)≠ Q" and
a3:"fst(last x) = Skip" and
a4:"env_tran_right Γ x rely " and
a5:"snd (x!0)∈ Normal ` p ∧ Sta p rely ∧ Sta a rely" and
a6: "Γ ⊨n⇘/F⇙ P sat [p, rely, G, q,a]"
shows
"False"
using a0 a1 a2 a3 a4 a5 a6
proof (induct arbitrary: P s p)
case (CptnModNestOne n Γ C s1)
thus ?case by fastforce
next
case (CptnModNestEnv Γ C s1 t1 n xsa)
then have C:"C=Seq P Q" unfolding lift_def by fastforce
thus ?case
proof -
have "((C, t1) # xsa) ! 0 = (LanguageCon.com.Seq P Q, t1)" using C by auto
moreover have "∀i<length ((C, t1) # xsa). fst (((C, t1) # xsa) ! i) ≠ Q"
using CptnModNestEnv(5) by fastforce
moreover have "fst (last ((C, t1) # xsa)) = LanguageCon.com.Skip" using CptnModNestEnv(6)
by (simp add: SmallStepCon.final_def)
moreover have "snd (((C, t1) # xsa) ! 0) ∈ Normal ` p"
using CptnModNestEnv(8) CptnModNestEnv(1) CptnModNestEnv(7)
unfolding env_tran_right_def Sta_def by fastforce
ultimately show ?thesis
using CptnModNestEnv(3) CptnModNestEnv(7) CptnModNestEnv(8) CptnModNestEnv(9) env_tran_tail
by blast
qed
next
case (CptnModNestSkip)
thus ?case by (metis SmallStepCon.redex_not_Seq fst_conv nth_Cons_0)
next
case (CptnModNestThrow)
thus ?case by (metis SmallStepCon.redex_not_Seq fst_conv nth_Cons_0)
next
case (CptnModNestSeq1 n Γ P0 sa xsa zs P1)
obtain cl where "fst (last ((LanguageCon.com.Seq P0 P1, sa) # zs)) = Seq cl P1"
using CptnModNestSeq1(3) by (metis One_nat_def fst_conv last.simps last_conv_nth last_lift map_is_Nil_conv)
thus ?case using CptnModNestSeq1(6) by auto
next
case (CptnModNestSeq2 n Γ P0 sa xsa P1 ys zs)
then have "P0 = P ∧ P1 = Q" by auto
then obtain i where zs:"fst (zs!i) = Q ∧ (i< (length zs))" using CptnModNestSeq2
by (metis (no_types, lifting) add_diff_cancel_left' fst_conv length_Cons length_append nth_append_length zero_less_Suc zero_less_diff)
thus ?case using CptnModNestSeq2(8) zs by auto
next
case (CptnModNestSeq3 n Γ P1 sa xsa s' ys zs Q1 )
have s'_a:"s' ∈ a"
proof -
have cpnP1:"(Γ, (P1, Normal sa) # xsa) ∈ cpn n Γ P1 (Normal sa)"
using CptnModNestSeq3.hyps(1) unfolding cpn_def
by fastforce
then have cpP1:"(Γ, (P1, Normal sa) # xsa) ∈ cp Γ P1 (Normal sa)"
using CptnModNestSeq3.hyps(1) cptn_mod_nest_cptn_mod cptn_if_cptn_mod unfolding cp_def cpn_def
by fastforce
have map:"((Seq P1 Q1), Normal sa)#(map (lift Q1) xsa) = map (lift Q1) ((P1, Normal sa) # xsa) "
using CptnModNestSeq3 by (simp add: Cons_lift)
then
have "(Γ,((LanguageCon.com.Seq P1 Q1, Normal sa) # (map (lift Q1) xsa))) ∈ assum (p,rely)"
proof -
have "env_tran_right Γ ((LanguageCon.com.Seq P1 Q1, Normal sa) # (map (lift Q1) xsa)) rely"
using CptnModNestSeq3(11) CptnModNestSeq3(7) map
by (metis (no_types) Cons_lift_append CptnModNestSeq3.hyps(7) CptnModNestSeq3.prems(4) env_tran_subr)
thus ?thesis using CptnModNestSeq3(12)
unfolding assum_def env_tran_right_def by fastforce
qed
moreover have "(Γ,((Seq P1 Q1), Normal sa)#(map (lift Q1) xsa)) ∈ cpn n Γ (Seq P1 Q1) (Normal sa)"
using CptnModNestSeq3.hyps(1)
CptnModNestSeq1
unfolding cpn_def by fastforce
then have "(Γ,((Seq P1 Q1), Normal sa)#(map (lift Q1) xsa)) ∈ cp Γ (Seq P1 Q1) (Normal sa)"
using CptnModNestSeq3.hyps(1) cp_def cptn_eq_cptn_mod_set
cptn_mod.CptnModSeq1 cptn_mod_nest_cptn_mod by fastforce
ultimately have "(Γ, (P1, Normal sa) # xsa) ∈ assum (p,rely)"
using assum_map map cpP1 by fastforce
then have "(Γ, (P1, Normal sa) # xsa) ∈ comm (G,(q,a)) F"
using cpnP1 CptnModNestSeq3(13) CptnModNestSeq3.prems(1) unfolding com_validityn_def by auto
thus ?thesis
using CptnModNestSeq3(3) CptnModNestSeq3(4)
unfolding comm_def final_def by fastforce
qed
have "fst (last ((LanguageCon.com.Throw, Normal s') # ys)) = Throw"
proof -
have cptn:"(Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn"
using CptnModNestSeq3(5)
using cptn_eq_cptn_mod_nest by blast
moreover have throw_0:"((LanguageCon.com.Throw, Normal s') # ys)!0 = (Throw, Normal s') ∧ 0 < length((LanguageCon.com.Throw, Normal s') # ys)"
by force
moreover have last:"last ((LanguageCon.com.Throw, Normal s') # ys) = ((LanguageCon.com.Throw, Normal s') # ys)!((length ((LanguageCon.com.Throw, Normal s') # ys)) - 1)"
using last_conv_nth by auto
moreover have env_tran:"env_tran_right Γ ((LanguageCon.com.Throw, Normal s') # ys) rely"
using CptnModNestSeq3(11) CptnModNestSeq3(7) env_tran_subl env_tran_tail by blast
ultimately obtain st' where "fst (last ((LanguageCon.com.Throw, Normal s') # ys)) = Throw ∧
snd (last ((LanguageCon.com.Throw, Normal s') # ys)) = Normal st'"
using zero_throw_all_throw[of Γ "((Throw, Normal s') # ys)" "s'" "(length ((Throw, Normal s') # ys))-1" a rely]
s'_a CptnModNestSeq3(11) CptnModNestSeq3(12) by fastforce
thus ?thesis using CptnModNestSeq3(10) final_def by blast
qed
thus ?case using CptnModNestSeq3(10) CptnModNestSeq3(7)
by force
qed(auto)
lemma map_xs_ys:
assumes
a0:"(Γ, (P0, sa) # xsa) ∈ cptn_mod" and
a1:"fst (last ((P0, sa) # xsa)) = C" and
a2:"(Γ, (P1, snd (last ((P0, sa) # xsa))) # ys) ∈ cptn_mod" and
a3:"zs = map (lift P1) xsa @ (P1, snd (last ((P0, sa) # xsa))) # ys" and
a4:"((LanguageCon.com.Seq P0 P1, sa) # zs) ! 0 = (LanguageCon.com.Seq P Q, s)" and
a5:"i < length ((LanguageCon.com.Seq P0 P1, sa) # zs) ∧ ((LanguageCon.com.Seq P0 P1, sa) # zs) ! i = (Q, sj)" and
a6:"∀j<i. fst (((LanguageCon.com.Seq P0 P1, sa) # zs) ! j) ≠ Q"
shows
"∃xs ys. (Γ, xs) ∈ cp Γ P s ∧
(Γ, ys) ∈ cp Γ Q (snd (xs ! (i - 1))) ∧ (LanguageCon.com.Seq P0 P1, sa) # zs = map (lift Q) xs @ ys"
proof -
let ?P0 = "(P0, sa) # xsa"
have P_Q:"P=P0 ∧ s=sa ∧ Q = P1" using a4 by force
have i:"i=(length ((P0, sa) # xsa))"
proof (cases "i=(length ((P0, sa) # xsa))")
case True thus ?thesis by auto
next
case False
then have i:"i<(length ((P0, sa) # xsa)) ∨ i > (length ((P0, sa) # xsa))" by auto
{
assume i:"i<(length ((P0, sa) # xsa))"
then have eq_map:"((LanguageCon.com.Seq P0 P1, sa) # zs) ! i = map (lift P1) ((P0, sa) # xsa) ! i"
using a3 Cons_lift_append by (metis (no_types, lifting) length_map nth_append)
then have "∃ci si. map (lift P1) ((P0, sa) # xsa) ! i = (Seq ci P1,si)"
using i unfolding lift_def
proof -
have "map (λ(c, y). (LanguageCon.com.Seq c P1, y)) ((P0, sa) # xsa) ! i = (case ((P0, sa) # xsa) ! i of (c, x) ⇒ (LanguageCon.com.Seq c P1, x))"
by (meson ‹i < length ((P0, sa) # xsa)› nth_map)
then show "∃c x. map (λ(c, x). (LanguageCon.com.Seq c P1, x)) ((P0, sa) # xsa) ! i = (LanguageCon.com.Seq c P1, x)"
by (simp add: case_prod_beta)
qed
then have "((LanguageCon.com.Seq P0 P1, sa) # zs) ! i ≠ (Q, sj)"
using P_Q eq_map by fastforce
then have ?thesis using a5 by auto
}note l=this
{
assume i:"i>(length ((P0, sa) # xsa))"
have "fst (((LanguageCon.com.Seq P0 P1, sa) # zs) ! (length ?P0)) = Q"
using a3 P_Q Cons_lift_append by (metis fstI length_map nth_append_length)
then have ?thesis using a6 i by auto
}
thus ?thesis using l i by auto
qed
then have "(Γ, (P0, sa) # xsa) ∈ cp Γ P s"
using a0 cptn_eq_cptn_mod P_Q unfolding cp_def by fastforce
also have "(Γ, (P1, snd (last ((P0, sa) # xsa))) # ys) ∈ cp Γ Q (snd (?P0 ! ((length ?P0) -1)))"
using a3 cptn_eq_cptn_mod P_Q unfolding cp_def
proof -
have "(Γ, (Q, snd (last ((P0, sa) # xsa))) # ys) ∈ cptn_mod"
using a2 P_Q by blast
then have "(Γ, (Q, snd (last ((P0, sa) # xsa))) # ys) ∈ {(f, ps). ps ! 0 = (Q, snd (((P0, sa) # xsa) ! (Suc (length xsa) - 1))) ∧ (Γ, ps) ∈ cptn ∧ f = Γ}"
by (simp add: cptn_eq_cptn_mod last_length)
then show "(Γ, (P1, snd (last ((P0, sa) # xsa))) # ys) ∈ {(f, ps). ps ! 0 = (Q, snd (((P0, sa) # xsa) ! (length ((P0, sa) # xsa) - 1))) ∧ (Γ, ps) ∈ cptn ∧ f = Γ}"
using P_Q by force
qed
ultimately show ?thesis using a3 P_Q i using Cons_lift_append by blast
qed
lemma map_xs_ys':
assumes
a0:"(n, Γ, (P0, sa) # xsa) ∈ cptn_mod_nest_call" and
a1:"fst (last ((P0, sa) # xsa)) = C" and
a2:"(n,Γ, (P1, snd (last ((P0, sa) # xsa))) # ys) ∈ cptn_mod_nest_call" and
a3:"zs = map (lift P1) xsa @ (P1, snd (last ((P0, sa) # xsa))) # ys" and
a4:"((LanguageCon.com.Seq P0 P1, sa) # zs) ! 0 = (LanguageCon.com.Seq P Q, s)" and
a5:"i < length ((LanguageCon.com.Seq P0 P1, sa) # zs) ∧ ((LanguageCon.com.Seq P0 P1, sa) # zs) ! i = (Q, sj)" and
a6:"∀j<i. fst (((LanguageCon.com.Seq P0 P1, sa) # zs) ! j) ≠ Q"
shows
"∃xs ys. (Γ, xs) ∈ cpn n Γ P s ∧
(Γ, ys) ∈ cpn n Γ Q (snd (xs ! (i - 1))) ∧ (LanguageCon.com.Seq P0 P1, sa) # zs = map (lift Q) xs @ ys"
proof -
let ?P0 = "(P0, sa) # xsa"
have P_Q:"P=P0 ∧ s=sa ∧ Q = P1" using a4 by force
have i:"i=(length ((P0, sa) # xsa))"
proof (cases "i=(length ((P0, sa) # xsa))")
case True thus ?thesis by auto
next
case False
then have i:"i<(length ((P0, sa) # xsa)) ∨ i > (length ((P0, sa) # xsa))" by auto
{
assume i:"i<(length ((P0, sa) # xsa))"
then have eq_map:"((LanguageCon.com.Seq P0 P1, sa) # zs) ! i = map (lift P1) ((P0, sa) # xsa) ! i"
using a3 Cons_lift_append by (metis (no_types, lifting) length_map nth_append)
then have "∃ci si. map (lift P1) ((P0, sa) # xsa) ! i = (Seq ci P1,si)"
using i unfolding lift_def
proof -
have "map (λ(c, y). (LanguageCon.com.Seq c P1, y)) ((P0, sa) # xsa) ! i = (case ((P0, sa) # xsa) ! i of (c, x) ⇒ (LanguageCon.com.Seq c P1, x))"
by (meson ‹i < length ((P0, sa) # xsa)› nth_map)
then show "∃c x. map (λ(c, x). (LanguageCon.com.Seq c P1, x)) ((P0, sa) # xsa) ! i = (LanguageCon.com.Seq c P1, x)"
by (simp add: case_prod_beta)
qed
then have "((LanguageCon.com.Seq P0 P1, sa) # zs) ! i ≠ (Q, sj)"
using P_Q eq_map by fastforce
then have ?thesis using a5 by auto
}note l=this
{
assume i:"i>(length ((P0, sa) # xsa))"
have "fst (((LanguageCon.com.Seq P0 P1, sa) # zs) ! (length ?P0)) = Q"
using a3 P_Q Cons_lift_append by (metis fstI length_map nth_append_length)
then have ?thesis using a6 i by auto
}
thus ?thesis using l i by auto
qed
then have "(Γ, (P0, sa) # xsa) ∈ cpn n Γ P s"
using a0 P_Q unfolding cpn_def by fastforce
also have "(Γ, (P1, snd (last ((P0, sa) # xsa))) # ys) ∈ cpn n Γ Q (snd (?P0 ! ((length ?P0) -1)))"
using a3 cptn_eq_cptn_mod P_Q unfolding cpn_def
proof -
have "(n, Γ, (Q, snd (last ((P0, sa) # xsa))) # ys) ∈ cptn_mod_nest_call"
using a2 P_Q by blast
then have "(Γ, (Q, snd (last ((P0, sa) # xsa))) # ys) ∈ {(f, ps). ps ! 0 = (Q, snd (((P0, sa) # xsa) ! (Suc (length xsa) - 1))) ∧
(n, Γ, ps) ∈ cptn_mod_nest_call ∧ f = Γ}"
by (simp add: cptn_eq_cptn_mod last_length)
then show "(Γ, (P1, snd (last ((P0, sa) # xsa))) # ys) ∈ {(f, ps). ps ! 0 = (Q, snd (((P0, sa) # xsa) ! (length ((P0, sa) # xsa) - 1))) ∧ (n,Γ, ps) ∈ cptn_mod_nest_call ∧ f = Γ}"
using P_Q by force
qed
ultimately show ?thesis using a3 P_Q i using Cons_lift_append by blast
qed
lemma Seq_sound4:
assumes
a0:"(n,Γ,x)∈cptn_mod_nest_call" and
a1:"x!0 = ((Seq P Q),s)" and
a2:"i<length x ∧ x!i=(Q,sj)" and
a3:"∀j<i. fst(x!j)≠Q" and
a4:"env_tran_right Γ x rely" and
a5:"snd (x!0)∈ Normal ` p ∧ Sta p rely ∧ Sta a rely" and
a6: "Γ ⊨n⇘/F⇙ P sat [p, rely, G, q,a]"
shows
"∃xs ys. (Γ,xs) ∈ (cpn n Γ P s) ∧ (Γ,ys) ∈ (cpn n Γ Q (snd (xs!(i-1)))) ∧ x = (map (lift Q) xs)@ys"
using a0 a1 a2 a3 a4 a5 a6
proof (induct arbitrary: i sj P s p)
case (CptnModNestOne Γ C s1)
thus ?case by fastforce
next
case (CptnModNestEnv Γ C st t n xsa)
have a1:"Seq P Q ≠ Q" by simp
then have C_seq:"C=(Seq P Q)" using CptnModNestEnv by fastforce
then have "fst(((C, st) # (C, t) # xsa)!0) ≠Q" using CptnEnv a1 by auto
moreover have n_q:"fst(((C, st) # (C, t) # xsa)!1) ≠Q" using CptnModNestEnv a1 by auto
moreover have "fst(((C, st) # (C, t) # xsa)!i) =Q" using CptnModNestEnv by auto
ultimately have i_suc: "i> (Suc 0)"
by (metis Suc_eq_plus1 Suc_lessI add.left_neutral neq0_conv)
then obtain i' where i':"i=Suc i'" by (meson lessE)
then have i_minus:"i'=i-1" by auto
have c_init:"((C, t) # xsa) ! 0 = ((Seq P Q), t)"
using CptnModNestEnv by auto
moreover have "i'< length ((C,t)#xsa) ∧ ((C,t)#xsa)!i' = (Q,sj)"
using i' CptnModNestEnv(5) by force
moreover have "∀j<i'. fst (((C, t) # xsa) ! j) ≠ Q"
using i' CptnModNestEnv(6) by force
moreover have "snd (((C, t) # xsa) ! 0) ∈ Normal ` p"
using CptnModNestEnv(8) CptnModNestEnv(1) CptnModNestEnv(7)
unfolding env_tran_right_def Sta_def by fastforce
ultimately have hyp:"∃xs ys.
(Γ, xs) ∈ cpn n Γ P t ∧
(Γ, ys) ∈ cpn n Γ Q (snd (xs ! (i'-1))) ∧ (C, t) # xsa = map (lift Q) xs @ ys"
using CptnModNestEnv(3) env_tran_tail CptnModNestEnv(8) CptnModNestEnv(9)
CptnModNestEnv.prems(4) by blast
then obtain xs ys where xs_cp:"(Γ, xs) ∈ cpn n Γ P t ∧
(Γ, ys) ∈ cpn n Γ Q (snd (xs ! (i'-1))) ∧ (C, t) # xsa = map (lift Q) xs @ ys"
by fast
have "(Γ, (P,s)#xs) ∈ cpn n Γ P s"
proof -
have "xs!0 = (P,t)"
using xs_cp unfolding cpn_def by blast
moreover have "xs≠[]"
using xs_cp n_q c_init unfolding cpn_def by auto
ultimately obtain xs' where xs':"(n, Γ, (P,t)#xs') ∈ cptn_mod_nest_call ∧ xs=(P,t)#xs'"
using SmallStepCon.nth_tl xs_cp unfolding cpn_def by force
thus ?thesis
proof -
have "(LanguageCon.com.Seq P Q, s) = (C, st)"
using CptnModNestEnv.prems(1) by auto
then have "Γ⊢⇩c (P, s) →⇩e (P, t)"
using Seq_env_P CptnModNestEnv(1) by blast
then show ?thesis
by (simp add:xs' cpn_def cptn_mod_nest_call.CptnModNestEnv)
qed
qed
thus ?case
using i_suc Cons_lift_append CptnModNestEnv.prems(1) i' i_minus xs_cp
by fastforce
next
case (CptnModNestSkip)
thus ?case by (metis SmallStepCon.redex_not_Seq fst_conv nth_Cons_0)
next
case (CptnModNestThrow)
thus ?case by (metis SmallStepCon.redex_not_Seq fst_conv nth_Cons_0)
next
case (CptnModNestSeq1 n Γ P0 sa xsa zs P1)
then have P1_Q:"P1 = Q" by auto
let ?x = "(LanguageCon.com.Seq P0 P1, sa) # zs"
have "∀j<length ?x. ∃c s. ?x!j = (Seq c P1,s)" using CptnModNestSeq1(3)
proof (induct xsa arbitrary: zs P0 P1 sa)
case Nil thus ?case by auto
next
case (Cons a xsa)
then obtain ac as where "a=(ac,as)" by fastforce
then have zs:"zs = (Seq ac P1,as)#(map (lift P1) xsa)"
using Cons(2)
unfolding lift_def by auto
have zs_eq:"(map (lift P1) xsa)=(map (lift P1) xsa)" by auto
note hyp=Cons(1)[OF zs_eq]
note hyp[of ac as]
thus ?case using zs Cons(2) by (metis One_nat_def diff_Suc_Suc diff_zero length_Cons less_Suc_eq_0_disj nth_Cons')
qed
thus ?case using P1_Q CptnModNestSeq1(5) using fstI seq_not_eq2 by auto
next
case (CptnModNestSeq2 n Γ P0 sa xsa P1 ys zs)
then show ?case using map_xs_ys'[OF CptnModNestSeq2(1) CptnModNestSeq2(3) CptnModNestSeq2(4) CptnModNestSeq2(6)
CptnModNestSeq2(7) CptnModNestSeq2(8) CptnModNestSeq2(9)] by blast
next
case (CptnModNestSeq3 n Γ P1 sa xsa s' ys zs Q1 )
then have P_Q:"P=P1 ∧ Q = Q1" by force
thus ?case
proof (cases "Q1 = Throw")
case True thus ?thesis using map_xs_ys'[of n Γ P1 "Normal sa" xsa Throw Throw ys zs]
CptnModNestSeq3 by fastforce
next
case False note q_not_throw=this
have "∀x. x< length ((LanguageCon.com.Seq P1 Q1, Normal sa) # zs) ⟶
((LanguageCon.com.Seq P1 Q1, Normal sa) # zs) ! x ≠ (Q, sj)"
proof -
{
fix x
assume x_less:"x< length ((LanguageCon.com.Seq P1 Q1, Normal sa) # zs)"
have "((LanguageCon.com.Seq P1 Q1, Normal sa) # zs) ! x ≠ (Q, sj)"
proof (cases "x < length ((LanguageCon.com.Seq P1 Q1, Normal sa)#map (lift Q1) xsa)")
case True
then have eq_map:"((LanguageCon.com.Seq P1 Q1, Normal sa) # zs) ! x = map (lift Q1) ((P1, Normal sa) # xsa) ! x"
by (metis (no_types) Cons_lift Cons_lift_append CptnModNestSeq3.hyps(7) True nth_append)
then have "∃ci si. map (lift Q1) ((P1, Normal sa) # xsa) ! x = (Seq ci Q1,si)"
using True unfolding lift_def
proof -
have "x < length ((P1, Normal sa) # xsa)"
using True by auto
then have "map (λ(c, y). (LanguageCon.com.Seq c Q1, y)) ((P1, Normal sa) # xsa) ! x = (case ((P1, Normal sa) # xsa) ! x of (c, x) ⇒ (LanguageCon.com.Seq c Q1, x))"
using nth_map by blast
then show "∃c x1. map (λ(c, x1). (LanguageCon.com.Seq c Q1, x1)) ((P1, Normal sa) # xsa) ! x = (LanguageCon.com.Seq c Q1, x1)"
by (simp add: case_prod_beta')
qed
then have "((LanguageCon.com.Seq P1 Q1, Normal sa) # zs) ! x ≠ (Q, sj)"
using P_Q eq_map by fastforce
thus ?thesis using CptnModNestSeq3(10) by auto
next
case False
have s'_a:"s' ∈ a"
proof -
have cpP1:"(Γ, (P1, Normal sa) # xsa) ∈ cpn n Γ P1 (Normal sa)"
using CptnModNestSeq3.hyps(1) cptn_eq_cptn_mod_set unfolding cpn_def by fastforce
then have cpP1':"(Γ, (P1, Normal sa) # xsa) ∈ cp Γ P1 (Normal sa)"
unfolding cpn_def cp_def
using cptn_if_cptn_mod cptn_mod_nest_cptn_mod by fastforce
have map:"((Seq P1 Q1), Normal sa)#(map (lift Q1) xsa) = map (lift Q1) ((P1, Normal sa) # xsa) "
using CptnModSeq3 by (simp add: Cons_lift)
then
have "(Γ,((LanguageCon.com.Seq P1 Q1, Normal sa) # (map (lift Q1) xsa))) ∈ assum (p,rely)"
proof -
have "env_tran_right Γ ((LanguageCon.com.Seq P1 Q1, Normal sa) # (map (lift Q1) xsa)) rely"
using CptnModNestSeq3(11) CptnModNestSeq3(7) map
by (metis (no_types) Cons_lift_append CptnModNestSeq3.hyps(7) CptnModNestSeq3.prems(4) env_tran_subr)
thus ?thesis using CptnModNestSeq3(12)
unfolding assum_def env_tran_right_def by fastforce
qed
moreover have "(Γ,((Seq P1 Q1), Normal sa)#(map (lift Q1) xsa)) ∈ cpn n Γ (Seq P1 Q1) (Normal sa)"
using CptnModNestSeq3(7) CptnModNestSeq3.hyps(1) cptn_eq_cptn_mod_set cptn_mod_nest_call.CptnModNestSeq1
unfolding cpn_def by fastforce
then have "(Γ,((Seq P1 Q1), Normal sa)#(map (lift Q1) xsa)) ∈ cp Γ (Seq P1 Q1) (Normal sa)"
unfolding cpn_def cp_def
by (simp add: cptn_if_cptn_mod cptn_mod_nest_cptn_mod)
ultimately have "(Γ, (P1, Normal sa) # xsa) ∈ assum (p,rely)"
using assum_map map cpP1' by fastforce
then have "(Γ, (P1, Normal sa) # xsa) ∈ comm (G,(q,a)) F"
using cpP1 CptnModNestSeq3(13) CptnModNestSeq3.prems(1) unfolding com_validityn_def by auto
thus ?thesis
using CptnModNestSeq3(3) CptnModNestSeq3(4)
unfolding comm_def final_def by fastforce
qed
have all_throw:"∀i<length ((LanguageCon.com.Throw, Normal s')# ys).
fst (((LanguageCon.com.Throw, Normal s')# ys)!i) = Throw"
proof -
{fix i
assume i:"i< length ((LanguageCon.com.Throw, Normal s')# ys)"
have cptn:"(n, Γ, (LanguageCon.com.Throw, Normal s') # ys) ∈ cptn_mod_nest_call"
using CptnModNestSeq3(5) by auto
moreover have throw_0:"((LanguageCon.com.Throw, Normal s') # ys)!0 = (Throw, Normal s') ∧ 0 < length((LanguageCon.com.Throw, Normal s') # ys)"
by force
moreover have last:"last ((LanguageCon.com.Throw, Normal s') # ys) = ((LanguageCon.com.Throw, Normal s') # ys)!((length ((LanguageCon.com.Throw, Normal s') # ys)) - 1)"
using last_conv_nth by auto
moreover have env_tran:"env_tran_right Γ ((LanguageCon.com.Throw, Normal s') # ys) rely"
using CptnModNestSeq3(11) CptnModNestSeq3(7) env_tran_subl env_tran_tail by blast
ultimately have
"fst (((LanguageCon.com.Throw, Normal s')# ys)!i) = Throw"
using zero_throw_all_throw[of Γ "((Throw, Normal s') # ys)" "s'" "i" a rely]
s'_a CptnModNestSeq3(12) i
using cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod by auto
}
thus ?thesis using CptnModNestSeq3(10) final_def by blast
qed
then have
"∀x≥ length ((LanguageCon.com.Seq P1 Q1, Normal sa) # map (lift Q1) xsa).
x<length (((LanguageCon.com.Seq P1 Q1, Normal sa) # zs)) ⟶
fst (((LanguageCon.com.Seq P1 Q1, Normal sa) # zs) ! x) = Throw"
proof-
{
fix x
assume a1:"x≥ length ((LanguageCon.com.Seq P1 Q1, Normal sa) # map (lift Q1) xsa)" and
a2:"x<length (((LanguageCon.com.Seq P1 Q1, Normal sa) # zs))"
then have "((LanguageCon.com.Seq P1 Q1, Normal sa) # zs) ! x =
((LanguageCon.com.Throw, Normal s')# ys) !(x - (length ((LanguageCon.com.Seq P1 Q1, Normal sa) # map (lift Q1) xsa)))"
using CptnModNestSeq3(7) by (metis Cons_lift Cons_lift_append not_le nth_append)
then have"fst (((LanguageCon.com.Seq P1 Q1, Normal sa) # zs) ! x) = Throw"
using all_throw a1 a2 CptnModNestSeq3.hyps(7) by auto
} thus ?thesis by auto
qed
thus ?thesis using False CptnModNestSeq3(7) q_not_throw P_Q x_less
by (metis fst_conv not_le)
qed
} thus ?thesis by auto
qed
thus ?thesis using CptnModNestSeq3(9) by fastforce
qed
qed(auto)
inductive_cases stepc_elim_cases_Seq_throw:
"Γ⊢⇩c(Seq c1 c2,s) → (Throw, Normal s1)"
inductive_cases stepc_elim_cases_Seq_skip_c2:
"Γ⊢⇩c(Seq c1 c2,s) → (c2,s)"
lemma seq_skip_throw:
"Γ⊢⇩c(Seq c1 c2,s) → (c2,s) ⟹ c1= Skip ∨ (c1=Throw ∧ (∃s2'. s=Normal s2'))"
apply (rule stepc_elim_cases_Seq_skip_c2)
apply fastforce
apply (auto)+
apply (fastforce intro:redex_not_Seq)+
done
lemma Seq_sound:
"Γ,Θ ⊢⇘/F⇙ c1 sat [p, R, G, q,a] ⟹
∀n. Γ,Θ ⊨n⇘/F⇙ c1 sat [p, R, G, q,a] ⟹
Γ,Θ ⊢⇘/F⇙ c2 sat [q, R, G, r,a] ⟹
∀n. Γ,Θ ⊨n⇘/F⇙ c2 sat [q, R, G, r,a] ⟹
Sta a R ∧ Sta p R ⟹ (∀s. (Normal s,Normal s) ∈ G) ⟹
Γ,Θ ⊨n⇘/F⇙ (Seq c1 c2) sat [p, R, G, r,a]"
proof -
assume
a0:"Γ,Θ ⊢⇘/F⇙ c1 sat [p, R, G, q,a]" and
a1:"∀n. Γ,Θ ⊨n⇘/F⇙ c1 sat [p, R, G, q,a]" and
a2:"Γ,Θ ⊢⇘/F⇙ c2 sat [q, R, G, r,a]" and
a3: "∀n. Γ,Θ ⊨n⇘/F⇙ c2 sat [q, R, G, r,a]" and
a4: "Sta a R ∧ Sta p R" and
a5: "(∀s. (Normal s, Normal s) ∈ G)"
{
fix s
assume all_call:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
then have a1:"Γ ⊨n⇘/F⇙ c1 sat [p, R, G, q,a]"
using a1 com_cvalidityn_def by fastforce
then have a3: "Γ ⊨n⇘/F⇙ c2 sat [q, R, G, r,a]"
using a3 com_cvalidityn_def all_call by fastforce
have "cpn n Γ (Seq c1 c2) s ∩ assum(p, R) ⊆ comm(G, (r,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ (Seq c1 c2) s" and a11:"c ∈ assum(p, R)"
then have a10':"c ∈ cp Γ (Seq c1 c2) s" unfolding cpn_def cp_def
using cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod by fastforce
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
have cp:"l!0=((Seq c1 c2),s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1" using a10' cp_def c_prod by fastforce
have cptn_nest:"l!0=((Seq c1 c2),s) ∧ (n,Γ,l) ∈ cptn_mod_nest_call ∧ Γ=Γ1" using a10 cpn_def c_prod by fastforce
have Γ1:"(Γ, l) = c" using c_prod cp by blast
have "c ∈ comm(G, (r,a)) F"
proof -
{
assume l_f:"snd (last l) ∉ Fault ` F"
have assum:"snd(l!0) ∈ Normal ` (p) ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ R)"
using a11 c_prod unfolding assum_def by simp
then have env_tran:"env_tran Γ p l R" using env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def unfolding env_tran_def by auto
have "(∀i. Suc i<length l ⟶
Γ⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)∧
(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` r)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` a))"
proof (cases "∀i<length l. fst (l!i)≠ c2")
case True
then have no_c2:"∀i<length l. fst (l!i)≠ c2" by assumption
show ?thesis
proof (cases "final (last l)")
case True
then obtain s' where "fst (last l) = Skip ∨ (fst (last l) = Throw ∧ snd (last l) = Normal s')"
using final_def by fast
thus ?thesis
proof
assume "fst (last l) = LanguageCon.com.Skip"
then have "False"
using no_c2 env_tran_right cptn_nest cptn_eq_cptn_mod_set Seq_sound3 a4 a1 assum
by blast
thus ?thesis by auto
next
assume asm0:"fst (last l) = LanguageCon.com.Throw ∧ snd (last l) = Normal s'"
then obtain lc1 s1' ys where cpn_lc1:"(Γ,lc1) ∈ cpn n Γ c1 s ∧ l = ((map (lift c2) lc1)@((Throw, Normal s1')#ys))"
using Seq_sound2'[of n Γ l c1 c2 s s'] cptn_nest cptn_eq_cptn_mod_set env_tran_right no_c2 by blast
then have cp_lc1:"(Γ,lc1) ∈ cp Γ c1 s"
using cptn_if_cptn_mod cptn_mod_nest_cptn_mod split_conv
unfolding cp_def cpn_def by blast
let ?m_lc1 = "map (lift c2) lc1"
let ?lm_lc1 = "(length ?m_lc1)"
let ?last_m_lc1 = "?m_lc1!(?lm_lc1-1)"
have lc1_not_empty:"lc1 ≠ []"
using Γ1 a10 cpn_lc1 cp by auto
then have map_cpn:"(Γ,?m_lc1) ∈ cpn n Γ (Seq c1 c2) s"
proof -
have f1: "lc1 ! 0 = (c1, s) ∧ (n,Γ, lc1) ∈ cptn_mod_nest_call ∧ Γ = Γ"
using cpn_lc1 cpn_def by blast
then have f2: "(n, Γ, ?m_lc1) ∈ cptn_mod_nest_call"
by (metis (no_types) Cons_lift cptn_mod_nest_call.CptnModNestSeq1 f1 lc1_not_empty list.exhaust nth_Cons_0)
then show ?thesis
using f2 f1 lc1_not_empty by (simp add: cpn_def lift_def)
qed
then have map_cp:"(Γ,?m_lc1) ∈ cp Γ (Seq c1 c2) s"
by (metis (no_types, lifting) cp_def cp_lc1 cpn_def lift_is_cptn mem_Collect_eq split_conv)
also have map_assum:"(Γ,?m_lc1) ∈ assum (p,R)"
using sub_assum a10 a11 Γ1 cpn_lc1 lc1_not_empty
by (metis SmallStepCon.nth_tl map_is_Nil_conv)
ultimately have "((Γ,lc1) ∈ assum(p, R))"
using Γ1 assum_map cp_lc1 by blast
then have lc1_comm:"(Γ,lc1) ∈ comm(G, (q,a)) F"
using a1 cpn_lc1 unfolding com_validityn_def by blast
then have m_lc1_comm:"(Γ,?m_lc1) ∈ comm(G, (q,a)) F"
using map_cp map_assum comm_map cp_lc1 by fastforce
then have last_m_lc1:"last (?m_lc1) = (Seq (fst (last lc1)) c2,snd (last lc1))"
proof -
have a000:"∀p c. (LanguageCon.com.Seq (fst p) c, snd p) = lift c p"
using Cons_lift by force
then show ?thesis
by (simp add: last_map a000 lc1_not_empty)
qed
then have last_length:"last (?m_lc1) = ?last_m_lc1"
using lc1_not_empty last_conv_nth list.map_disc_iff by blast
then have l_map:"l!(?lm_lc1-1)= ?last_m_lc1"
using cpn_lc1
by (simp add:lc1_not_empty nth_append)
then have lm_lc1:"l!(?lm_lc1) = (Throw, Normal s1')"
using cpn_lc1 by (meson nth_append_length)
then have step:"Γ⊢⇩c(l!(?lm_lc1-1)) → (l!(?lm_lc1))"
proof -
have "Γ⊢⇩c(l!(?lm_lc1-1)) →⇩c⇩e (l!(?lm_lc1))"
proof -
have f1: "∀n na. ¬ n < na ∨ Suc (na - Suc n) = na - n"
by (meson Suc_diff_Suc)
have "map (lift c2) lc1 ≠ []"
by (metis lc1_not_empty map_is_Nil_conv)
then have f2: "0 < length (map (lift c2) lc1)"
by (meson length_greater_0_conv)
then have "length (map (lift c2) lc1) - 1 + 1 < length (map (lift c2) lc1 @ (LanguageCon.com.Throw, Normal s1') # ys)"
by simp
then show ?thesis
using f2 f1
by (metis Suc_pred' cp cpn_lc1 cptn_tran_ce_i)
qed
moreover have "¬Γ⊢⇩c(l!(?lm_lc1-1)) →⇩e (l!(?lm_lc1))"
using last_m_lc1 last_length l_map
proof -
have "(LanguageCon.com.Seq (fst (last lc1)) c2, snd (last lc1)) = l ! (length (map (lift c2) lc1) - 1)"
using l_map last_m_lc1 local.last_length by presburger
then show ?thesis
by (metis (no_types) LanguageCon.com.distinct(71) ‹l ! length (map (lift c2) lc1) = (LanguageCon.com.Throw, Normal s1')› env_c_c')
qed
ultimately show ?thesis using step_ce_elim_cases by blast
qed
then have last_lc1_suc:"snd (l!(?lm_lc1-1)) = snd (l!?lm_lc1) ∧ fst (l!(?lm_lc1-1)) = Seq Throw c2"
using lm_lc1 stepc_elim_cases_Seq_throw
by (metis One_nat_def asm0 append_is_Nil_conv cpn_lc1 diff_Suc_less fst_conv l_map last_conv_nth last_m_lc1 length_greater_0_conv list.simps(3) local.last_length no_c2 snd_conv)
then have a_normal:"snd (l!?lm_lc1) ∈ Normal ` (a)"
proof
have last_lc1:"fst (last lc1) = Throw ∧ snd (last lc1) = Normal s1'"
using last_length l_map lm_lc1 last_m_lc1 last_lc1_suc
by (metis LanguageCon.com.inject(3) fst_conv snd_conv)
have "final (last lc1)" using last_lc1 final_def
by blast
moreover have "snd (last lc1)∉ Fault ` F"
using last_lc1 by fastforce
ultimately have "(fst (last lc1) = Throw ∧
snd (last lc1) ∈ Normal ` (a))"
using lc1_comm last_lc1 unfolding comm_def by force
thus ?thesis using l_map last_lc1_suc last_m_lc1 last_length by auto
qed
have concl:"(∀i. Suc i<length l ⟶
Γ⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof-
{ fix k ns ns'
assume a00:"Suc k<length l" and
a21:"Γ⊢⇩c(l!k) → (l!(Suc k))"
then have i_m_l:"∀i <?lm_lc1 . l!i = ?m_lc1!i"
using cp_lc1
proof -
have "map (lift c2) lc1 ≠ []"
by (meson lc1_not_empty list.map_disc_iff)
then show ?thesis
by (metis (no_types) cpn_lc1 nth_append)
qed
have last_not_F:"snd (last ?m_lc1) ∉ Fault ` F"
using l_map last_lc1_suc lm_lc1 last_length by auto
have "(snd(l!k), snd(l!(Suc k))) ∈ G"
proof (cases "Suc k< ?lm_lc1")
case True
then have a11': "Γ⊢⇩c(?m_lc1!k) → (?m_lc1!(Suc k))"
using a11 i_m_l True
proof -
have "∀n na. ¬ 0 < n - Suc na ∨ na < n "
using diff_Suc_eq_diff_pred zero_less_diff by presburger
then show ?thesis
by (metis (no_types) True a21 i_m_l zero_less_diff)
qed
then have "(snd(?m_lc1!k), snd(?m_lc1!(Suc k))) ∈ G"
using a11' m_lc1_comm True comm_dest1 l_f last_not_F by fastforce
thus ?thesis using i_m_l using True by fastforce
next
case False
then have "(Suc k=?lm_lc1) ∨ (Suc k>?lm_lc1)" by auto
thus ?thesis
proof
{assume suck:"(Suc k=?lm_lc1)"
then have k:"k=?lm_lc1-1" by auto
have G_s1':"(Normal s1', Normal s1')∈G"
using a5 by auto
then show "(snd (l!k), snd (l!Suc k)) ∈ G"
proof -
have "snd (l!Suc k) = Normal s1'"
using lm_lc1 suck by fastforce
then show ?thesis using suck k G_s1' last_lc1_suc by fastforce
qed
}
next
{
assume a001:"Suc k>?lm_lc1"
have "∀i. i≥(length lc1) ∧ (Suc i < length l) ⟶
¬(Γ⊢⇩c(l!i) → (l!(Suc i)))"
using lm_lc1 lc1_not_empty
proof -
have "env_tran_right Γ l R"
by (metis env_tran_right)
then show ?thesis
using a_normal cp fst_conv length_map
lm_lc1 only_one_component_tran_j[of Γ l ?lm_lc1 s1' a k R] snd_conv a21 a001 a00
a4 by auto
qed
then have "¬(Γ⊢⇩c(l!k) → (l!(Suc k)))"
using a00 a001 by auto
then show ?thesis using a21 by fastforce
}
qed
qed
} thus ?thesis by auto
qed
have concr:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` r)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` a))"
proof -
have l_t:"fst (last l) = Throw"
using lm_lc1 by (simp add: asm0)
have "?lm_lc1 ≤ length l -1" using cpn_lc1 by fastforce
then have "snd (l ! (length l - 1)) ∈ Normal ` a"
using cp a_normal a4 fst_conv lm_lc1 snd_conv
env_tran_right i_throw_all_throw[of Γ l ?lm_lc1 s1' "(length l -1)" _ R a ]
by (metis (no_types, lifting) One_nat_def diff_is_0_eq diff_less diff_less_Suc diff_zero image_iff length_greater_0_conv lessI less_antisym list.size(3) xstate.inject(1))
thus ?thesis using l_t
by (simp add: cpn_lc1 last_conv_nth)
qed
note res = conjI [OF concl concr]
then show ?thesis using Γ1 c_prod unfolding comm_def by auto
qed
next
case False
then obtain lc1 where cpn_lc1:"(Γ,lc1) ∈ cpn n Γ c1 s ∧ l = map (lift c2) lc1"
using Seq_sound1 assum False no_c2 env_tran_right cptn_nest cptn_eq_cptn_mod_set a4 a1
by blast
then have cp_lc1:"(Γ,lc1) ∈ cp Γ c1 s "
using cp_def cpn_def cptn_if_cptn_mod cptn_mod_nest_cptn_mod by fastforce
then have "((Γ,lc1) ∈ assum(p, R))"
using Γ1 cpn_lc1 a10' a11 assum_map by blast
then have "(Γ, lc1)∈ comm(G, (q,a)) F" using cpn_lc1 a1
by (meson IntI com_validityn_def contra_subsetD)
then have "(Γ, l)∈ comm(G, (r,a)) F"
using comm_map a10' Γ1 cp_lc1 cpn_lc1 by blast
then show ?thesis using l_f
unfolding comm_def by auto
qed
next
case False
then obtain k where k_len:"k<length l ∧ fst (l ! k) = c2"
by blast
then have "∃m. (m < length l ∧ fst (l ! m) = c2) ∧
(∀i<m. ¬ (i < length l ∧ fst (l ! i) = c2))"
using a0 exists_first_occ[of "(λi. i<length l ∧ fst (l ! i) = c2)" k]
by blast
then obtain i where a0:"i<length l ∧ fst (l !i) = c2 ∧
(∀j<i. (fst (l ! j) ≠ c2))"
by fastforce
then obtain s2 where li:"l!i =(c2,s2)" by (meson eq_fst_iff)
then obtain lc1 lc2 where cp_lc1:"(Γ,lc1) ∈ (cpn n Γ c1 s) ∧
(Γ,lc2) ∈ (cpn n Γ c2 (snd (lc1!(i-1)))) ∧
l = (map (lift c2) lc1)@lc2"
using Seq_sound4[of n Γ l c1 c2 s] a0 env_tran_right a4 a1 cptn_nest assum by blast
then have cp_lc1':"(Γ,lc1) ∈ (cp Γ c1 s) ∧
(Γ,lc2) ∈ (cp Γ c2 (snd (lc1!(i-1))))"
unfolding cp_def cpn_def cptn_eq_cptn_mod_nest by fastforce
have "∀i < length l. snd (l!i) ∉ Fault ` F"
using cp l_f last_not_F[of Γ l F] by blast
then have i_not_fault:"snd (l!i) ∉ Fault ` F" using a0 by blast
have length_c1_map:"length lc1 = length (map (lift c2) lc1)"
by fastforce
then have i_map:"i=length lc1"
using cp_lc1 li a0 unfolding lift_def
proof -
assume a1: "(Γ, lc1) ∈ cpn n Γ c1 s ∧ (Γ, lc2) ∈ cpn n Γ c2 (snd (lc1 ! (i - 1))) ∧ l = map (λ(P, s). (LanguageCon.com.Seq P c2, s)) lc1 @ lc2"
have f2: "i < length l ∧ fst (l ! i) = c2 ∧ (∀n. ¬ n < i ∨ fst (l ! n) ≠ c2)"
using a0 by blast
have f3: "(LanguageCon.com.Seq (fst (lc1 ! i)) c2, snd (lc1 ! i)) = lift c2 (lc1 ! i)"
by (simp add: case_prod_unfold lift_def)
then have "fst (l ! length lc1) = c2"
using a1 by (simp add: cpn_def nth_append)
thus ?thesis
using f3 f2 by (metis (no_types) nth_append cp_lc1
fst_conv length_map lift_nth linorder_neqE_nat seq_and_if_not_eq(4))
qed
have lc2_l:"∀j<length lc2. lc2!j=l!(i+j)"
using cp_lc1 length_c1_map i_map a0
by (metis nth_append_length_plus)
have lc1_not_empty:"lc1 ≠ []"
using cp cp_lc1 unfolding cpn_def by fastforce
have lc2_not_empty:"lc2 ≠ []"
using a0 cp_lc1 i_map by auto
have l_is:"s2= snd (last lc1)"
using cp_lc1 li a0 lc1_not_empty i_map unfolding cpn_def
by (auto simp add: last_conv_nth lc2_l)
let ?m_lc1 = "map (lift c2) lc1"
have last_m_lc1:"l!(i-1) = (Seq (fst (last lc1)) c2,s2)"
proof -
have a000:"∀p c. (LanguageCon.com.Seq (fst p) c, snd p) = lift c p"
using Cons_lift by force
have "length (map (lift c2) lc1) = i"
using i_map by fastforce
then show ?thesis
by (metis (no_types) One_nat_def l_is a000 cp_lc1 diff_less last_conv_nth last_map
lc1_not_empty length_c1_map length_greater_0_conv less_Suc0 nth_append)
qed
have last_mcl1_not_F:"snd (last ?m_lc1) ∉ Fault ` F"
proof -
have "map (lift c2) lc1 ≠ []"
by (metis lc1_not_empty list.map_disc_iff)
then show ?thesis
by (metis (full_types) One_nat_def i_not_fault l_is last_conv_nth last_snd lc1_not_empty li snd_conv)
qed
have map_cp:"(Γ,?m_lc1) ∈ cpn n Γ (Seq c1 c2) s"
proof -
have f1: "lc1 ! 0 = (c1, s) ∧ (n,Γ, lc1) ∈ cptn_mod_nest_call ∧ Γ = Γ"
using cp_lc1 cpn_def by blast
then have f2: "(n,Γ, ?m_lc1) ∈ cptn_mod_nest_call" using lc1_not_empty
by (metis Cons_lift SmallStepCon.nth_tl cptn_mod_nest_call.CptnModNestSeq1)
then show ?thesis
using f2 f1 lc1_not_empty by (simp add: cpn_def lift_def)
qed
then have map_cp':"(Γ,?m_lc1) ∈ cp Γ (Seq c1 c2) s"
unfolding cpn_def cp_def
using cptn_eq_cptn_mod_nest by fastforce
also have map_assum:"(Γ,?m_lc1) ∈ assum (p,R)"
using sub_assum a10 a11 Γ1 cp_lc1 lc1_not_empty
by (metis SmallStepCon.nth_tl map_is_Nil_conv)
ultimately have "((Γ,lc1) ∈ assum(p, R))"
using Γ1 assum_map using assum_map cp_lc1' by blast
then have lc1_comm:"(Γ,lc1) ∈ comm(G, (q,a)) F"
using a1 cp_lc1 by (meson IntI com_validityn_def contra_subsetD)
then have m_lc1_comm:"(Γ,?m_lc1) ∈ comm(G, (q,a)) F"
using map_cp' map_assum comm_map cp_lc1' by fastforce
then have i_step:"Γ⊢⇩c(l!(i-1)) → (l!i)"
proof -
have "Γ⊢⇩c(l!(i-1)) →⇩c⇩e (l!(i))"
proof -
have f1: "∀n na. ¬ n < na ∨ Suc (na - Suc n) = na - n"
by (meson Suc_diff_Suc)
have "map (lift c2) lc1 ≠ []"
by (metis lc1_not_empty map_is_Nil_conv)
then have f2: "0 < length (map (lift c2) lc1)"
by (meson length_greater_0_conv)
then have "length (map (lift c2) lc1) - 1 + 1 < length (map (lift c2) lc1 @ lc2)"
using f2 lc2_not_empty by simp
then show ?thesis
using f2 f1
proof -
have "0 < i"
using f2 i_map by blast
then show ?thesis
by (metis (no_types) One_nat_def Suc_diff_1 a0 add.right_neutral add_Suc_right cp cptn_tran_ce_i)
qed
qed
moreover have "¬Γ⊢⇩c(l!(i-1)) →⇩e (l!i)"
using li last_m_lc1
by (metis (no_types, lifting) env_c_c' seq_and_if_not_eq(4))
ultimately show ?thesis using step_ce_elim_cases by blast
qed
then have step:"Γ⊢⇩c(Seq (fst (last lc1)) c2,s2) → (c2, s2)"
using last_m_lc1 li by fastforce
then obtain s2' where
last_lc1:"fst (last lc1) = Skip ∨
fst (last lc1) = Throw ∧ (s2 = Normal s2')"
using seq_skip_throw by blast
have final:"final (last lc1)"
using last_lc1 l_is unfolding final_def by auto
have normal_last:"fst (last lc1) = Skip ∧ snd (last lc1) ∈ Normal ` q ∨
fst (last lc1) = Throw ∧ snd (last lc1) ∈ Normal ` (a)"
proof -
have "snd (last lc1) ∉ Fault ` F"
using i_not_fault l_is li by auto
then show ?thesis
using final comm_dest2 lc1_comm by blast
qed
obtain s2' where lastlc1_normal:"snd (last lc1) = Normal s2'"
using normal_last by blast
then have Normals2:"s2 = Normal s2'" by (simp add: l_is )
have Gs2':"(Normal s2', Normal s2')∈G" using a5 by auto
have concl:
"(∀i. Suc i<length l ⟶
Γ⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof-
{ fix k
assume a00:"Suc k<length l" and
a21:"Γ⊢⇩c(l!k) → (l!(Suc k))"
have i_m_l:"∀j <i . l!j = ?m_lc1!j"
proof -
have "map (lift c2) lc1 ≠ []"
by (meson lc1_not_empty list.map_disc_iff)
then show ?thesis
using cp_lc1 i_map length_c1_map by (fastforce simp:nth_append)
qed
have "(snd(l!k), snd(l!(Suc k))) ∈ G"
proof (cases "Suc k< i")
case True
then have a11': "Γ⊢⇩c(?m_lc1!k) → (?m_lc1!(Suc k))"
using a11 i_m_l True
proof -
have "∀n na. ¬ 0 < n - Suc na ∨ na < n "
using diff_Suc_eq_diff_pred zero_less_diff by presburger
then show ?thesis using True a21 i_m_l by force
qed
have "Suc k < length ?m_lc1" using True i_map length_c1_map by metis
then have "(snd(?m_lc1!k), snd(?m_lc1!(Suc k))) ∈ G"
using a11' last_mcl1_not_F m_lc1_comm True i_map length_c1_map
comm_dest1[of Γ]
by blast
thus ?thesis using i_m_l using True by fastforce
next
case False
have "(Suc k=i) ∨ (Suc k>i)" using False by auto
thus ?thesis
proof
{ assume suck:"(Suc k=i)"
then have k:"k=i-1" by auto
then show "(snd (l!k), snd (l!Suc k)) ∈ G"
proof -
have "snd (l!Suc k) = Normal s2'"
using Normals2 suck li by auto
moreover have "snd (l ! k) = Normal s2'"
using Normals2 k last_m_lc1 by fastforce
moreover have "∃p. p ∈ G "
by (meson case_prod_conv mem_Collect_eq Gs2')
ultimately show ?thesis using suck k Normals2
using Gs2' by force
qed
}
next
{
assume a001:"Suc k>i"
then have k:"k≥i" by fastforce
then obtain k' where k':"k=i+k'"
using add.commute le_Suc_ex by blast
{assume throw:"c2=Throw ∧ fst (last lc1) = Throw"
then have s2_in:"s2' ∈ a"
using Normals2 i_map normal_last li lastlc1_normal
using image_iff snd_conv xstate.inject(1) by auto
then have "∀k. k≥i ∧ (Suc k < length l) ⟶
¬(Γ⊢⇩c(l!k) → (l!(Suc k)))"
using Normals2 li lastlc1_normal a21 a001 a00 a4
a0 throw env_tran_right only_one_component_tran_j snd_conv
by (metis cp env_tran_right)
then have ?thesis using a21 a001 k a00 by blast
} note left=this
{assume "¬(c2=Throw ∧ fst (last lc1) = Throw)"
then have "fst (last lc1) = Skip"
using last_m_lc1 last_lc1
by (metis step a0 l_is li prod.collapse stepc_Normal_elim_cases(11) stepc_Normal_elim_cases(5))
then have s2_normal:"s2 ∈ Normal ` q"
using normal_last lastlc1_normal Normals2
by fastforce
have length_lc2:"length l=i+length lc2"
using i_map cp_lc1 by fastforce
have "(Γ,lc2) ∈ assum (q,R)"
proof -
have left:"snd (lc2!0) ∈ Normal ` q"
using li lc2_l s2_normal lc2_not_empty by fastforce
{
fix j
assume j_len:"Suc j<length lc2" and
j_step:"Γ⊢⇩c(lc2!j) →⇩e (lc2!(Suc j))"
then have suc_len:"Suc (i + j)<length l" using j_len length_lc2
by fastforce
also then have "Γ⊢⇩c(l!(i+j)) →⇩e (l! (Suc (i+ j)))"
using lc2_l j_step j_len by fastforce
ultimately have "(snd(lc2!j), snd(lc2!(Suc j))) ∈ R"
using assum suc_len lc2_l j_len cp by fastforce
}
then show ?thesis using left
unfolding assum_def by fastforce
qed
also have "(Γ,lc2) ∈ cpn n Γ c2 s2"
using cp_lc1 i_map l_is last_conv_nth lc1_not_empty by fastforce
ultimately have comm_lc2:"(Γ,lc2) ∈ comm (G, (r,a)) F"
using a3 unfolding com_validityn_def by auto
have lc2_last_f:"snd (last lc2)∉ Fault ` F"
using lc2_l lc2_not_empty l_f cp_lc1 by fastforce
have suck':"Suc k' < length lc2"
using k' a00 length_lc2 by arith
moreover then have "Γ⊢⇩c(lc2!k') → (lc2!(Suc k'))"
using k' lc2_l a21 by fastforce
ultimately have "(snd (lc2! k'), snd (lc2 ! Suc k')) ∈ G"
using comm_lc2 lc2_last_f comm_dest1[of Γ lc2 G r a F k']
by blast
then have ?thesis using suck' lc2_l k' by fastforce
}
then show ?thesis using left by auto
}
qed
qed
} thus ?thesis by auto
qed note left=this
have right:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` r)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` a))"
proof -
{ assume final_l:"final (last l)"
have eq_last_lc2_l:"last l=last lc2" by (simp add: cp_lc1 lc2_not_empty)
then have final_lc2:"final (last lc2)" using final_l by auto
{
assume lst_lc1_throw:"fst (last lc1) = Throw"
then have c2_throw:"c2 = Throw"
using lst_lc1_throw step lastlc1_normal stepc_elim_cases_Seq_skip_c2
by fastforce
have s2_a:"s2 ∈ Normal ` (a)"
using normal_last
by (simp add: lst_lc1_throw l_is)
have all_ev:"∀k<length l - 1. k≥i ∧ (Suc k < length l) ⟶
Γ⊢⇩c(l!k) →⇩e (l!(Suc k))"
proof -
have s2_in:"s2' ∈ a"
using Normals2 i_map normal_last li lastlc1_normal
using image_iff snd_conv xstate.inject(1) lst_lc1_throw by auto
then have "∀k. k≥i ∧ (Suc k < length l) ⟶
¬(Γ⊢⇩c(l!k) → (l!(Suc k)))"
using Normals2 li lastlc1_normal a4
a0 c2_throw env_tran_right only_one_component_tran_j snd_conv
by (metis cp env_tran_right)
thus ?thesis by (metis Suc_eq_plus1 cp cptn_tran_ce_i step_ce_elim_cases)
qed
then have Throw:"fst (l!(length l - 1)) = Throw"
using cp c2_throw a0 cptn_i_env_same_prog[of Γ l "((length l)-1)" i]
by fastforce
then have "snd (l!(length l - 1)) ∈ Normal ` (a) ∧ fst (l!(length l - 1)) = Throw"
using all_ev a0 s2_a li a4 env_tran_right stability[of a R l i "(length l) -1" _ Γ] Throw
by (metis One_nat_def Suc_pred length_greater_0_conv
lessI linorder_not_less list.size(3)
not_less0 not_less_eq_eq snd_conv)
then have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` r)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a))"
using a0 by (metis last_conv_nth list.size(3) not_less0)
} note left = this
{ assume "fst (last lc1) = Skip"
then have s2_normal:"s2 ∈ Normal ` q"
using normal_last lastlc1_normal Normals2
by fastforce
have length_lc2:"length l=i+length lc2"
using i_map cp_lc1 by fastforce
have "(Γ,lc2) ∈ assum (q,R)"
proof -
have left:"snd (lc2!0) ∈ Normal ` q"
using li lc2_l s2_normal lc2_not_empty by fastforce
{
fix j
assume j_len:"Suc j<length lc2" and
j_step:"Γ⊢⇩c(lc2!j) →⇩e (lc2!(Suc j))"
then have suc_len:"Suc (i + j)<length l" using j_len length_lc2
by fastforce
also then have "Γ⊢⇩c(l!(i+j)) →⇩e (l! (Suc (i+ j)))"
using lc2_l j_step j_len by fastforce
ultimately have "(snd(lc2!j), snd(lc2!(Suc j))) ∈ R"
using assum suc_len lc2_l j_len cp by fastforce
}
then show ?thesis using left
unfolding assum_def by fastforce
qed
also have "(Γ,lc2) ∈ cpn n Γ c2 s2"
using cp_lc1 i_map l_is last_conv_nth lc1_not_empty by fastforce
ultimately have comm_lc2:"(Γ,lc2) ∈ comm (G, (r,a)) F"
using a3 unfolding com_validityn_def by auto
have lc2_last_f:"snd (last lc2)∉ Fault ` F"
using lc2_l lc2_not_empty l_f cp_lc1 by fastforce
then have "((fst (last lc2) = Skip ∧
snd (last lc2) ∈ Normal ` r)) ∨
(fst (last lc2) = Throw ∧
snd (last lc2) ∈ Normal ` a)"
using final_lc2 comm_lc2 unfolding comm_def by auto
then have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` r)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` a)"
using eq_last_lc2_l by auto
}
then have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` r)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` a)"
using left using last_lc1 by auto
} thus ?thesis by auto qed
thus ?thesis using left l_f Γ1 unfolding comm_def by force
qed
} thus ?thesis using Γ1 unfolding comm_def by auto qed
} thus ?thesis by auto qed
} thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma Catch_env_P:assumes a0:"Γ⊢⇩c(Catch P Q,s) →⇩e (Catch P Q,t)"
shows "Γ⊢⇩c(P,s) →⇩e (P,t)"
using a0
by (metis env_not_normal_s snormal_enviroment)
lemma map_catch_eq_state:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Catch c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) " and
a2:"l1=map (lift_catch c2) l2"
shows
"∀i<length l1. snd (l1!i) = snd (l2!i)"
using a0 a1 a2 unfolding cp_def
by (simp add: snd_lift_catch)
lemma map_eq_catch_c:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Catch c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) " and
a2:"l1=map (lift_catch c2) l2"
shows
"∀i<length l1. fst (l1!i) = Catch (fst (l2!i)) c2"
proof -
{fix i
assume a3:"i<length l1"
have "fst (l1!i) = Catch (fst (l2!i)) c2"
using a0 a1 a2 a3 unfolding lift_catch_def
by (simp add: case_prod_unfold)
}thus ?thesis by auto
qed
lemma same_env_catch_c:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Catch c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) " and
a2:"l1=map (lift_catch c2) l2"
shows
"∀i. Suc i<length l2 ⟶ Γ⊢⇩c(l2!i) →⇩e (l2!(Suc i)) =
Γ⊢⇩c(l1!i) →⇩e (l1!(Suc i))"
proof -
have a0a:"(Γ,l1) ∈cptn ∧ l1!0 = ((Catch c1 c2),s)"
using a0 unfolding cp_def by blast
have a1a: "(Γ,l2) ∈cptn ∧ l2!0 = (c1,s)"
using a1 unfolding cp_def by blast
{
fix i
assume a3:"Suc i< length l2"
have "Γ⊢⇩c(l2!i) →⇩e (l2!(Suc i)) =
Γ⊢⇩c(l1!i) →⇩e (l1!(Suc i))"
proof
{
assume a4:"Γ⊢⇩c l2 ! i →⇩e l2 ! Suc i"
obtain c1i s1i c1si s1si where l1prod:"l1 ! i=(c1i,s1i) ∧ l1!Suc i = (c1si,s1si)"
by fastforce
obtain c2i s2i c2si s2si where l2prod:"l2 ! i=(c2i,s2i) ∧ l2!Suc i = (c2si,s2si)"
by fastforce
then have "c1i = (Catch c2i c2) ∧ c1si = (Catch c2si c2)"
using a0 a1 a2 a3 a4 l1prod
by (simp add: lift_catch_def)
also have "s2i=s1i ∧ s2si=s1si"
using a0 a1 a4 a2 a3 l2prod l1prod
by (simp add: lift_catch_def)
ultimately show "Γ⊢⇩c l1 ! i →⇩e (l1 ! Suc i)"
using a4 l1prod l2prod
by (metis Env_n env_c_c' env_not_normal_s step_e.Env)
}
{
assume a4:"Γ⊢⇩c l1 ! i →⇩e l1 ! Suc i"
obtain c1i s1i c1si s1si where l1prod:"l1 ! i=(c1i,s1i) ∧ l1!Suc i = (c1si,s1si)"
by fastforce
obtain c2i s2i c2si s2si where l2prod:"l2 ! i=(c2i,s2i) ∧ l2!Suc i = (c2si,s2si)"
by fastforce
then have "c1i = (Catch c2i c2) ∧ c1si = (Catch c2si c2)"
using a0 a1 a2 a3 a4 l1prod
by (simp add: lift_catch_def)
also have "s2i=s1i ∧ s2si=s1si"
using a0 a1 a4 a2 a3 l2prod l1prod
by (simp add: lift_catch_def)
ultimately show "Γ⊢⇩c l2 ! i →⇩e (l2 ! Suc i)"
using a4 l1prod l2prod
by (metis Env_n LanguageCon.com.inject(9) env_c_c' env_not_normal_s step_e.Env)
}
qed
}
thus ?thesis by auto
qed
lemma same_comp_catch_c:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Catch c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) " and
a2:"l1=map (lift_catch c2) l2"
shows
"∀i. Suc i<length l2 ⟶ Γ⊢⇩c(l2!i) → (l2!(Suc i)) =
Γ⊢⇩c(l1!i) → (l1!(Suc i))"
proof -
have a0a:"(Γ,l1) ∈cptn ∧ l1!0 = ((Catch c1 c2),s)"
using a0 unfolding cp_def by blast
have a1a: "(Γ,l2) ∈cptn ∧ l2!0 = (c1,s)"
using a1 unfolding cp_def by blast
{
fix i
assume a3:"Suc i< length l2"
have "Γ⊢⇩c(l2!i) → (l2!(Suc i)) =
Γ⊢⇩c(l1!i) → (l1!(Suc i))"
proof
{
assume a4:"Γ⊢⇩c l2 ! i → l2 ! Suc i"
obtain c1i s1i c1si s1si where l1prod:"l1 ! i=(c1i,s1i) ∧ l1!Suc i = (c1si,s1si)"
by fastforce
obtain c2i s2i c2si s2si where l2prod:"l2 ! i=(c2i,s2i) ∧ l2!Suc i = (c2si,s2si)"
by fastforce
then have "c1i = (Catch c2i c2) ∧ c1si = (Catch c2si c2)"
using a0 a1 a2 a3 a4 map_eq_catch_c l1prod
by (simp add: lift_catch_def)
also have "s2i=s1i ∧ s2si=s1si"
using a0 a1 a4 a2 a3 l2prod map_eq_state l1prod
by (simp add: lift_catch_def)
ultimately show "Γ⊢⇩c l1 ! i → (l1 ! Suc i)"
using a4 l1prod l2prod
by (simp add: Catchc)
}
{
assume a4:"Γ⊢⇩c l1 ! i → l1 ! Suc i"
obtain c1i s1i c1si s1si where l1prod:"l1 ! i=(c1i,s1i) ∧ l1!Suc i = (c1si,s1si)"
by fastforce
obtain c2i s2i c2si s2si where l2prod:"l2 ! i=(c2i,s2i) ∧ l2!Suc i = (c2si,s2si)"
by fastforce
then have "c1i = (Catch c2i c2) ∧ c1si = (Catch c2si c2)"
using a0 a1 a2 a3 a4 l1prod
by (simp add: lift_catch_def)
also have "s2i=s1i ∧ s2si=s1si"
using a0 a1 a4 a2 a3 l2prod l1prod
by (simp add: lift_catch_def)
ultimately show "Γ⊢⇩c l2 ! i → (l2 ! Suc i)"
using a4 l1prod l2prod stepc_elim_cases_Catch_Catch Catch_not_c
by (metis (no_types))
}
qed
}
thus ?thesis by auto
qed
lemma assum_map_catch:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Catch c1 c2) s) ∧ ((Γ,l1) ∈ assum(p, R))" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) " and
a2:"l1=map (lift_catch c2) l2"
shows
"((Γ,l2) ∈ assum(p, R))"
proof -
have a3: "∀i. Suc i<length l2 ⟶ Γ⊢⇩c(l2!i) →⇩e (l2!(Suc i)) =
Γ⊢⇩c(l1!i) →⇩e (l1!(Suc i))"
using a0 a1 a2 same_env_catch_c by fastforce
have pair_Γl1:"fst (Γ,l1) = Γ ∧ snd (Γ,l1) = l1" by fastforce
have pair_Γl2:"fst (Γ,l2) = Γ ∧ snd (Γ,l2) = l2" by fastforce
have drop_k_s:"l2!0 = (c1,s)" using a1 cp_def by blast
have eq_length:"length l1 = length l2" using a2 by auto
obtain s' where normal_s:"s = Normal s'"
using a0 unfolding cp_def assum_def by fastforce
then have p1:"s'∈p" using a0 unfolding cp_def assum_def by fastforce
show ?thesis
proof -
let ?c= "(Γ,l2)"
have l:"snd((snd ?c!0)) ∈ Normal ` (p)"
using p1 drop_k_s a1 normal_s unfolding cp_def by auto
{fix i
assume a00:"Suc i<length (snd ?c)"
assume a11:"(fst ?c)⊢⇩c((snd ?c)!i) →⇩e ((snd ?c)!(Suc i))"
have "(snd((snd ?c)!i), snd((snd ?c)!(Suc i))) ∈ R"
using a0 a1 a2 a3 map_catch_eq_state unfolding assum_def
using a00 a11 eq_length by fastforce
} thus "(Γ, l2) ∈ assum (p, R)"
using l unfolding assum_def by fastforce
qed
qed
lemma comm_map'_catch:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Catch c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) ∧ (Γ, l2)∈ comm(G, (q,a)) F" and
a2:"l1=map (lift_catch c2) l2"
shows
"snd (last l1) ∉ Fault ` F ⟶(Suc k < length l1 ⟶
Γ⊢⇩c(l1!k) → (l1!(Suc k)) ⟶
(snd(l1!k), snd(l1!(Suc k))) ∈ G) ∧
(fst (last l1) = (Catch c c2) ∧ final (c, snd (last l1)) ⟶
(fst (last l1) = (Catch Skip c2) ∧
(snd (last l1) ∈ Normal ` q) ∨
(fst (last l1) = (Catch Throw c2) ∧
snd (last l1) ∈ Normal ` (a))))
"
proof -
have a3:"∀i. Suc i<length l2 ⟶ Γ⊢⇩c(l2!i) → (l2!(Suc i)) =
Γ⊢⇩c(l1!i) → (l1!(Suc i))"
using a0 a1 a2 same_comp_catch_c
by fastforce
have pair_Γl1:"fst (Γ,l1) = Γ ∧ snd (Γ,l1) = l1" by fastforce
have pair_Γl2:"fst (Γ,l2) = Γ ∧ snd (Γ,l2) = l2" by fastforce
have drop_k_s:"l2!0 = (c1,s)" using a1 cp_def by blast
have eq_length:"length l1 = length l2" using a2 by auto
have len0:"length l2>0" using a1 unfolding cp_def
using cptn.simps by fastforce
then have len0:"length l1>0" using eq_length by auto
then have l1_not_empty:"l1≠[]" by auto
then have l2_not_empty:"l2 ≠ []" using a2 by blast
have last_lenl1:"last l1 = l1!((length l1) -1)"
using last_conv_nth l1_not_empty by auto
have last_lenl2:"last l2 = l2!((length l2) -1)"
using last_conv_nth l2_not_empty by auto
have a03:"snd (last l2) ∉ Fault ` F ⟶(∀i ns ns'.
Suc i<length (snd (Γ, l2)) ⟶
fst (Γ, l2)⊢⇩c((snd (Γ, l2))!i) → ((snd (Γ, l2))!(Suc i)) ⟶
(snd((snd (Γ, l2))!i), snd((snd (Γ, l2))!(Suc i))) ∈ G) ∧
(final (last (snd (Γ, l2))) ⟶
((fst (last (snd (Γ, l2))) = Skip ∧
snd (last (snd (Γ, l2))) ∈ Normal ` q)) ∨
(fst (last (snd (Γ, l2))) = Throw ∧
snd (last (snd (Γ, l2))) ∈ Normal ` (a)))"
using a1 unfolding comm_def by fastforce
show ?thesis unfolding comm_def
proof -
{ fix k ns ns'
assume a00a:"snd (last l1) ∉ Fault ` F"
assume a00:"Suc k < length l1"
then have "k ≤ length l1" using a2 by fastforce
have a00:"Suc k < length l2" using eq_length a00 by fastforce
then have a00a:"snd (last l2) ∉ Fault ` F"
proof-
have "snd (l1!((length l1) -1)) = snd (l2!((length l2) -1))"
using a2 a1 a0 map_catch_eq_state eq_length l2_not_empty last_snd
by fastforce
then have "snd(last l2) = snd (last l1)"
using last_lenl1 last_lenl2 by auto
thus ?thesis using a00a by auto
qed
then have "snd (last l1) ∉ Fault ` F ⟶Γ⊢⇩c(l1!k) → (l1!(Suc k)) ⟶
(snd((snd (Γ, l1))!k), snd((snd (Γ, l1))!(Suc k))) ∈ G"
using pair_Γl1 pair_Γl2 a00 a03 a3 eq_length a00a
by (metis Suc_lessD a0 a1 a2 map_catch_eq_state)
} note l=this
{
assume a00: "fst (last l1) = (Catch c c2) ∧ final (c, snd (last l1))" and
a01:"snd (last (l1)) ∉ Fault ` F"
then have c:"c=Skip ∨ c = Throw"
unfolding final_def by auto
then have fst_last_l2:"fst (last l2) = c"
using last_lenl1 a00 l1_not_empty eq_length len0 a2 last_conv_nth last_lift_catch
by fastforce
also have last_eq:"snd (last l2) = snd (last l1)"
using l2_not_empty a2 last_conv_nth last_lenl1 last_snd_catch
by fastforce
ultimately have "final (fst (last l2),snd (last l2))"
using a00 by auto
then have "final (last l2)" by auto
also have "snd (last (l2)) ∉ Fault ` F"
using last_eq a01 by auto
ultimately have "(fst (last l2)) = Skip ∧
snd (last l2) ∈ Normal ` q ∨
(fst (last l2) = Throw ∧
snd (last l2) ∈ Normal ` (a))"
using a03 by auto
then have "(fst (last l1) = (Catch Skip c2) ∧
snd (last l1) ∈ Normal ` q) ∨
(fst (last l1) = (Catch Throw c2) ∧
snd (last l1) ∈ Normal ` (a))"
using last_eq fst_last_l2 a00 by force
}
thus ?thesis using l by auto qed
qed
lemma comm_map''_catch:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Catch c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) ∧ (Γ, l2)∈ comm(G, (q,a)) F" and
a2:"l1=map (lift_catch c2) l2"
shows
"snd (last l1) ∉ Fault ` F ⟶ ((Suc k < length l1 ⟶
Γ⊢⇩c(l1!k) → (l1!(Suc k)) ⟶
(snd(l1!k), snd(l1!(Suc k))) ∈ G) ∧
(final (last l1) ⟶
(fst (last l1) = Skip ∧
(snd (last l1) ∈ Normal ` r) ∨
(fst (last l1) = Throw ∧
snd (last l1) ∈ Normal ` a))))
"
proof -
have a3:"∀i. Suc i<length l2 ⟶ Γ⊢⇩c(l2!i) → (l2!(Suc i)) =
Γ⊢⇩c(l1!i) → (l1!(Suc i))"
using a0 a1 a2 same_comp_catch_c
by fastforce
have pair_Γl1:"fst (Γ,l1) = Γ ∧ snd (Γ,l1) = l1" by fastforce
have pair_Γl2:"fst (Γ,l2) = Γ ∧ snd (Γ,l2) = l2" by fastforce
have drop_k_s:"l2!0 = (c1,s)" using a1 cp_def by blast
have eq_length:"length l1 = length l2" using a2 by auto
have len0:"length l2>0" using a1 unfolding cp_def
using cptn.simps by fastforce
then have len0:"length l1>0" using eq_length by auto
then have l1_not_empty:"l1≠[]" by auto
then have l2_not_empty:"l2 ≠ []" using a2 by blast
have last_lenl1:"last l1 = l1!((length l1) -1)"
using last_conv_nth l1_not_empty by auto
have last_lenl2:"last l2 = l2!((length l2) -1)"
using last_conv_nth l2_not_empty by auto
have a03:"snd (last l2) ∉ Fault ` F ⟶(∀i ns ns'.
Suc i<length (snd (Γ, l2)) ⟶
fst (Γ, l2)⊢⇩c((snd (Γ, l2))!i) → ((snd (Γ, l2))!(Suc i)) ⟶
(snd((snd (Γ, l2))!i), snd((snd (Γ, l2))!(Suc i))) ∈ G) ∧
(final (last (snd (Γ, l2))) ⟶
((fst (last (snd (Γ, l2))) = Skip ∧
snd (last (snd (Γ, l2))) ∈ Normal ` q)) ∨
(fst (last (snd (Γ, l2))) = Throw ∧
snd (last (snd (Γ, l2))) ∈ Normal ` (a)))"
using a1 unfolding comm_def by fastforce
show ?thesis unfolding comm_def
proof -
{ fix k ns ns'
assume a00a:"snd (last l1) ∉ Fault ` F"
assume a00:"Suc k < length l1"
then have "k ≤ length l1" using a2 by fastforce
have a00:"Suc k < length l2" using eq_length a00 by fastforce
then have a00a:"snd (last l2) ∉ Fault ` F"
proof-
have "snd (l1!((length l1) -1)) = snd (l2!((length l2) -1))"
using a2 a1 a0 map_catch_eq_state eq_length l2_not_empty last_snd
by fastforce
then have "snd(last l2) = snd (last l1)"
using last_lenl1 last_lenl2 by auto
thus ?thesis using a00a by auto
qed
then have "Γ⊢⇩c(l1!k) → (l1!(Suc k)) ⟶
(snd((snd (Γ, l1))!k), snd((snd (Γ, l1))!(Suc k))) ∈ G"
using pair_Γl1 pair_Γl2 a00 a03 a3 eq_length a00a
by (metis (no_types,lifting) a2 Suc_lessD nth_map snd_lift_catch)
} note l= this
{
assume a00: "final (last l1)"
then have c:"fst (last l1)=Skip ∨ fst (last l1) = Throw"
unfolding final_def by auto
moreover have "fst (last l1) = Catch (fst (last l2)) c2"
using a2 last_lenl1 eq_length
proof -
have "last l2 = l2 ! (length l2 - 1)"
using l2_not_empty last_conv_nth by blast
then show ?thesis
by (metis One_nat_def a2 l2_not_empty last_lenl1 last_lift_catch)
qed
ultimately have False by simp
} thus ?thesis using l by auto qed
qed
lemma comm_map_catch:
assumes
a0:"(Γ,l1) ∈ (cp Γ (Catch c1 c2) s)" and
a1:"(Γ,l2) ∈ (cp Γ c1 s) ∧ (Γ, l2)∈ comm(G, (q,a)) F" and
a2:"l1=map (lift_catch c2) l2"
shows
"(Γ, l1)∈ comm(G, (r,a)) F"
proof -
{fix i ns ns'
have "snd (last l1) ∉ Fault ` F ⟶(Suc i < length (l1) ⟶
Γ⊢⇩c (l1 ! i) → (l1 ! (Suc i)) ⟶
(snd (l1 ! i), snd (l1 ! Suc i)) ∈ G) ∧
(SmallStepCon.final (last l1) ⟶
fst (last l1) = LanguageCon.com.Skip ∧
snd (last l1) ∈ Normal ` r ∨
fst (last l1) = LanguageCon.com.Throw ∧
snd (last l1) ∈ Normal ` a) "
using comm_map''_catch[of Γ l1 c1 c2 s l2 G q a F i r] a0 a1 a2
by fastforce
} then show ?thesis using comm_def unfolding comm_def by force
qed
lemma Catch_sound1:
assumes
a0:"(n,Γ,x)∈cptn_mod_nest_call" and
a1:"x!0 = ((Catch P Q),s)" and
a2:"∀i<length x. fst (x!i)≠ Q" and
a3:"¬ final (last x)" and
a4:"env_tran_right Γ x rely"
shows
"∃xs. (Γ,xs) ∈ cpn n Γ P s ∧ x = map (lift_catch Q) xs"
using a0 a1 a2 a3 a4
proof (induct arbitrary: P s)
case (CptnModNestOne n Γ C s1)
then have "(Γ, [(P,s)]) ∈ cpn n Γ P s ∧ [(C, s1)] = map (lift_catch Q) [(P,s)]"
unfolding cpn_def lift_catch_def
by (simp add: cptn_mod_nest_call.CptnModNestOne)
thus ?case by fastforce
next
case (CptnModNestEnv Γ C s1 t1 n xsa)
then have C:"C=Catch P Q" unfolding lift_catch_def by fastforce
have "∃xs. (Γ, xs) ∈ cpn n Γ P t1 ∧ (C, t1) # xsa = map (lift_catch Q) xs"
proof -
have "((C, t1) # xsa) ! 0 = (Catch P Q, t1)" using C by auto
moreover have "∀i<length ((C, t1) # xsa). fst (((C, t1) # xsa) ! i) ≠ Q"
using CptnModNestEnv(5) by fastforce
moreover have "¬ SmallStepCon.final (last ((C, t1) # xsa))" using CptnModNestEnv(6)
by fastforce
ultimately show ?thesis
using CptnModNestEnv(3) CptnModNestEnv(7) env_tran_tail by blast
qed
then obtain xs where hi:"(Γ, xs) ∈ cpn n Γ P t1 ∧ (C, t1) # xsa = map (lift_catch Q) xs"
by fastforce
have s1_s:"s1=s" using CptnModNestEnv unfolding cpn_def by auto
obtain xsa' where xs:"xs=((P,t1)#xsa') ∧ (n,Γ,((P,t1)#xsa'))∈cptn_mod_nest_call ∧
(C, t1) # xsa = map (lift_catch Q) ((P,t1)#xsa')"
using hi unfolding cpn_def by fastforce
have env_tran:"Γ⊢⇩c(P,s1) →⇩e (P,t1)" using CptnModNestEnv Catch_env_P by (metis fst_conv nth_Cons_0)
then have "(n,Γ,(P,s1)#(P,t1)#xsa')∈cptn_mod_nest_call" using xs env_tran
using cptn_mod_nest_call.CptnModNestEnv by blast
then have "(Γ,(P,s1)#(P,t1)#xsa') ∈ cpn n Γ P s"
using cpn_def s1_s by fastforce
moreover have "(C,s1)#(C, t1) # xsa = map (lift_catch Q) ((P,s1)#(P,t1)#xsa')"
using xs C unfolding lift_catch_def by fastforce
ultimately show ?case by auto
next
case (CptnModNestSkip)
thus ?case by (metis SmallStepCon.redex_not_Catch fst_conv nth_Cons_0)
next
case (CptnModNestThrow)
thus ?case by (metis SmallStepCon.redex_not_Catch fst_conv nth_Cons_0)
next
case (CptnModNestCatch1 n Γ P0 sa xsa zs P1)
then have a1:"LanguageCon.com.Catch P Q = LanguageCon.com.Catch P0 P1"
by fastforce
have f1: "sa = s"
using CptnModNestCatch1.prems(1) by force
have f2: "P = P0 ∧ Q = P1" using a1 by auto
have "(n,Γ, (P0, sa) # xsa) ∈ cptn_mod_nest_call"
by (metis CptnModNestCatch1.hyps(1))
hence "(Γ, (P0, sa) # xsa) ∈ cpn n Γ P s"
using f2 f1 by (simp add: cpn_def)
thus ?case
using Cons_lift_catch CptnModNestCatch1.hyps(3) a1 by blast
next
case (CptnModNestCatch2 n Γ P1 sa xsa ys zs Q1)
have "final (last ((Skip, sa)# ys))"
proof -
have cptn:"(n, Γ, (Skip,snd (last ((P1, sa) # xsa))) # ys) ∈ cptn_mod_nest_call"
using CptnModNestCatch2(4) by (simp add: cptn_eq_cptn_mod_set)
then have cptn':"(Γ, (Skip,snd (last ((P1, sa) # xsa))) # ys) ∈ cptn"
using cptn_eq_cptn_mod_nest by blast
moreover have throw_0:"((Skip,snd (last ((P1, sa) # xsa))) # ys)!0 = (Skip, snd (last ((P1, sa) # xsa))) ∧ 0 < length((Skip, snd (last ((P1, sa) # xsa))) # ys)"
by force
moreover have last:"last ((Skip,snd (last ((P1, sa) # xsa))) # ys) =
((Skip,snd (last ((P1, sa) # xsa))) # ys)!((length ((Skip,snd (last ((P1, sa) # xsa))) # ys)) - 1)"
using last_conv_nth by auto
moreover have env_tran:"env_tran_right Γ ((Skip,snd (last ((P1, sa) # xsa))) # ys) rely"
using CptnModNestCatch2.hyps(6) CptnModNestCatch2.prems(4) env_tran_subl env_tran_tail by blast
ultimately obtain st' where "fst (last ((Skip,snd (last ((P1, sa) # xsa))) # ys)) = Skip ∧
snd (last ((Skip,snd (last ((P1, sa) # xsa))) # ys)) = Normal st'"
using CptnModNestCatch2 zero_skip_all_skip[of Γ "((Skip,snd (last ((P1, sa) # xsa))) # ys)" "(length ((Skip,snd (last ((P1, sa) # xsa))) # ys))-1"]
by (metis (no_types, hide_lams) SmallStepCon.final_def append_Cons diff_less fst_conv
last_appendR list.simps(3) zero_less_one)
thus ?thesis using final_def by (metis fst_conv last.simps)
qed
thus ?case
by (metis (no_types, lifting) CptnModNestCatch2.hyps(3) CptnModNestCatch2.hyps(6)
CptnModNestCatch2.prems(3) SmallStepCon.final_def append_is_Nil_conv last.simps
last_appendR list.simps(3) prod.collapse)
next
case (CptnModNestCatch3 n Γ P0 sa xsa sa' P1 ys zs)
then have "P0 = P ∧ P1 = Q" by auto
then obtain i where zs:"fst (zs!i) = Q ∧ (i< (length zs))"
using CptnModNestCatch3
by (metis (no_types, lifting) add_diff_cancel_left' fst_conv length_Cons length_append nth_append_length zero_less_Suc zero_less_diff)
then have "Suc i< length ((Catch P0 P1,Normal sa)#zs)" by fastforce
then have "fst (((Catch P0 P1, Normal sa) # zs)!Suc i) = Q" using zs by fastforce
thus ?case using CptnModNestCatch3(9) zs by auto
qed (auto)
lemma Catch_sound2:
assumes
a0:"(n,Γ,x)∈cptn_mod_nest_call" and
a1:"x!0 = ((Catch P Q),s)" and
a2:"∀i<length x. fst (x!i)≠ Q" and
a3:"fst (last x) = Skip" and
a4:"env_tran_right Γ x rely"
shows
"∃xs ys. (Γ,xs) ∈ cpn n Γ P s ∧ x = ((map (lift_catch Q) xs)@((Skip,snd(last xs))#ys))"
using a0 a1 a2 a3 a4
proof (induct arbitrary: P s)
case (CptnModNestOne n Γ C s1)
thus ?case by fastforce
next
case (CptnModNestEnv Γ C s1 t1 n xsa)
then have C:"C=Catch P Q" unfolding lift_catch_def by fastforce
have "∃xs ys. (Γ, xs) ∈ cpn n Γ P t1 ∧ (C, t1) # xsa =
map (lift_catch Q) xs@((Skip, snd(last xs))#ys)"
proof -
have "((C, t1) # xsa) ! 0 = (LanguageCon.com.Catch P Q, t1)" using C by auto
moreover have "∀i<length ((C, t1) # xsa). fst (((C, t1) # xsa) ! i) ≠ Q"
using CptnModNestEnv(5) by fastforce
moreover have "fst (last ((C, t1) # xsa)) = Skip" using CptnModNestEnv(6)
by fastforce
ultimately show ?thesis
using CptnModNestEnv(3) CptnModNestEnv(7) env_tran_tail by blast
qed
then obtain xs ys where hi:"(Γ, xs) ∈ cpn n Γ P t1 ∧ (C, t1) # xsa = map (lift_catch Q) xs@((Skip,snd(last ((P, t1)#xs)))#ys)"
by fastforce
have s1_s:"s1=s" using CptnModNestEnv unfolding cp_def by auto
have "∃xsa' ys. xs=((P,t1)#xsa') ∧ (n,Γ,((P,t1)#xsa'))∈cptn_mod_nest_call ∧ (C, t1) # xsa = map (lift_catch Q) ((P,t1)#xsa')@((Skip, snd(last xs))#ys)"
using hi unfolding cp_def
proof -
have "(n,Γ,xs)∈cptn_mod_nest_call ∧ xs!0 = (P,t1)" using hi unfolding cpn_def by fastforce
moreover then have "xs≠[]" using CptnEmpty calculation by blast
moreover obtain xsa' where "xs=((P,t1)#xsa')" using SmallStepCon.nth_tl calculation by fastforce
ultimately show ?thesis
using hi by auto
qed
then obtain xsa' ys where xs:"xs=((P,t1)#xsa') ∧ (n,Γ,((P,t1)#xsa'))∈cptn_mod_nest_call ∧ (C, t1) # xsa =
map (lift_catch Q) ((P,t1)#xsa')@((Skip,snd(last ((P,s1)#(P,t1)#xsa')))#ys)"
by fastforce
have env_tran:"Γ⊢⇩c(P,s1) →⇩e (P,t1)" using CptnModNestEnv Catch_env_P by (metis fst_conv nth_Cons_0)
then have "(n,Γ,(P,s1)#(P,t1)#xsa')∈cptn_mod_nest_call" using xs env_tran CptnEnv
by (simp add: cptn_mod_nest_call.CptnModNestEnv)
then have "(Γ,(P,s1)#(P,t1)#xsa') ∈ cpn n Γ P s"
using cpn_def s1_s by fastforce
moreover have "(C,s1)#(C, t1) # xsa = map (lift_catch Q) ((P,s1)#(P,t1)#xsa')@((Skip,snd(last ((P,s1)#(P,t1)#xsa')))#ys)"
using xs C unfolding lift_catch_def
by auto
ultimately show ?case by fastforce
next
case (CptnModNestSkip)
thus ?case by (metis SmallStepCon.redex_not_Catch fst_conv nth_Cons_0)
next
case (CptnModNestThrow)
thus ?case by (metis SmallStepCon.redex_not_Catch fst_conv nth_Cons_0)
next
case (CptnModNestCatch1 n Γ P0 sa xsa zs P1)
thus ?case
proof -
have "∀c x. (LanguageCon.com.Catch c P1, x) # zs = map (lift_catch P1) ((c, x) # xsa)"
using Cons_lift_catch CptnModNestCatch1.hyps(3) by blast
then have "(P0, sa) # xsa = []"
by (metis (no_types) CptnModNestCatch1.prems(3) LanguageCon.com.distinct(19) One_nat_def last_conv_nth last_lift_catch map_is_Nil_conv)
then show ?thesis
by force
qed
next
case (CptnModNestCatch2 n Γ P1 sa xsa ys zs Q1)
then have "P1 = P ∧ Q1 = Q ∧ sa = s" by auto
moreover then have "(Γ, (P1,sa) # xsa)∈ cpn n Γ P s"
using CptnModNestCatch2(1)
by (simp add: cpn_def cptn_eq_cptn_mod_set)
moreover obtain s' where "last zs=(Skip, s')"
proof -
assume a1: "⋀s'. last zs = (LanguageCon.com.Skip, s') ⟹ thesis"
have "∃x. last zs = (LanguageCon.com.Skip, x)"
by (metis (no_types) CptnModNestCatch2.hyps(6) CptnModNestCatch2.prems(3) append_is_Nil_conv last_ConsR list.simps(3) prod.exhaust_sel)
then show ?thesis
using a1 by metis
qed
ultimately show ?case
using Cons_lift_catch_append CptnModNestCatch2.hyps(6) by fastforce
next
case (CptnModNestCatch3 n Γ P0 sa xsa sa' P1 ys zs)
then have "P0 = P ∧ P1 = Q ∧ s=Normal sa" by auto
then obtain i where zs:"fst (zs!i) = Q ∧ (i< (length zs))"
using CptnModNestCatch3
by (metis (no_types, lifting) add_diff_cancel_left' fst_conv length_Cons length_append nth_append_length zero_less_Suc zero_less_diff)
then have si:"Suc i< length ((Catch P0 P1,Normal sa)#zs)" by fastforce
then have "fst (((Seq P0 P1, Normal sa) # zs)!Suc i) = Q" using zs by fastforce
thus ?case using CptnModNestCatch3(9) zs
by (metis si nth_Cons_Suc)
qed (auto)
lemma Catch_sound3:
assumes
a0:"(Γ,x)∈cptn" and
a1:"x!0 = ((Catch P Q),s)" and
a2:"∀i<length x. fst (x!i)≠ Q" and
a3:"fst(last x) = Throw" and
a4:"env_tran_right Γ x rely "
shows
"False"
using a0 a1 a2 a3 a4
proof (induct arbitrary: P s)
case (CptnOne Γ C s1) thus ?case by auto
next
case (CptnEnv Γ C st t xsa)
thus ?case
proof -
have f1: "env_tran_right Γ ((C, t) # xsa) rely"
using CptnEnv.prems(4) env_tran_tail by blast
have "LanguageCon.com.Catch P Q = C"
using CptnEnv.prems(1) by auto
then show ?thesis
using f1 CptnEnv.hyps(3) CptnEnv.prems(2) CptnEnv.prems(3) by moura
qed
next
case (CptnComp Γ C st C' st' xsa)
then have c_catch:"C = (Catch P Q) ∧ st = s" by force
from CptnComp show ?case proof(cases)
case (Catchc P1 P1' P2) thus ?thesis
proof -
have f1: "env_tran_right Γ ((C', st') # xsa) rely"
using CptnComp.prems(4) env_tran_tail by blast
have "Q = P2"
using c_catch Catchc(1) by blast
then show ?thesis
using f1 CptnComp.hyps(3) CptnComp.prems(2) CptnComp.prems(3) Catchc(2) by moura
qed
next
case (CatchSkipc) thus ?thesis
proof -
have "fst (((C', st') # xsa) ! 0) = LanguageCon.com.Skip"
by (simp add: local.CatchSkipc(2))
then show ?thesis
by (metis (no_types) CptnComp.hyps(2) CptnComp.prems(3) LanguageCon.com.distinct(17)
last_ConsR last_length length_Cons lessI list.simps(3) zero_skip_all_skip)
qed
next
case (SeqThrowc C2 s') thus ?thesis
by (simp add: c_catch)
next
case (FaultPropc) thus ?thesis
using c_catch redex_not_Catch by blast
next
case (StuckPropc) thus ?thesis
using c_catch redex_not_Catch by blast
next
case (AbruptPropc) thus ?thesis
using c_catch redex_not_Catch by blast
qed (auto)
qed
lemma catch_map_xs_ys':
assumes
a0:"(n, Γ, (P0, sa) # xsa) ∈ cptn_mod_nest_call" and
a1:"fst (last ((P0, sa) # xsa)) = C" and
a2:"(n,Γ, (P1, snd (last ((P0, sa) # xsa))) # ys) ∈ cptn_mod_nest_call" and
a3:"zs = map (lift_catch P1) xsa @ (P1, snd (last ((P0, sa) # xsa))) # ys" and
a4:"((LanguageCon.com.Catch P0 P1, sa) # zs) ! 0 = (LanguageCon.com.Catch P Q, s)" and
a5:"i < length ((LanguageCon.com.Catch P0 P1, sa) # zs) ∧ ((LanguageCon.com.Catch P0 P1, sa) # zs) ! i = (Q, sj)" and
a6:"∀j<i. fst (((LanguageCon.com.Catch P0 P1, sa) # zs) ! j) ≠ Q"
shows
"∃xs ys. (Γ, xs) ∈ cpn n Γ P s ∧
(Γ, ys) ∈ cpn n Γ Q (snd (xs ! (i - 1))) ∧ (LanguageCon.com.Catch P0 P1, sa) # zs = map (lift_catch Q) xs @ ys"
proof -
let ?P0 = "(P0, sa) # xsa"
have P_Q:"P=P0 ∧ s=sa ∧ Q = P1" using a4 by force
have i:"i=(length ((P0, sa) # xsa))"
proof (cases "i=(length ((P0, sa) # xsa))")
case True thus ?thesis by auto
next
case False
then have i:"i<(length ((P0, sa) # xsa)) ∨ i > (length ((P0, sa) # xsa))" by auto
{
assume i:"i<(length ((P0, sa) # xsa))"
then have eq_map:"((LanguageCon.com.Catch P0 P1, sa) # zs) ! i = map (lift_catch P1) ((P0, sa) # xsa) ! i"
using a3 Cons_lift_catch_append
by (metis (no_types, lifting) length_map nth_append)
then have "∃ci si. map (lift_catch P1) ((P0, sa) # xsa) ! i = (Catch ci P1,si)"
using i unfolding lift_catch_def
by (metis a5 eq_map fst_conv length_map map_lift_catch_all_catch)
then have "((LanguageCon.com.Catch P0 P1, sa) # zs) ! i ≠ (Q, sj)"
using P_Q eq_map by fastforce
then have ?thesis using a5 by auto
}note l=this
{
assume i:"i>(length ((P0, sa) # xsa))"
have "fst (((LanguageCon.com.Catch P0 P1, sa) # zs) ! (length ?P0)) = Q"
using a3 P_Q Cons_lift_catch_append by (metis fstI length_map nth_append_length)
then have ?thesis using a6 i by auto
}
thus ?thesis using l i by auto
qed
then have "(Γ, (P0, sa) # xsa) ∈ cpn n Γ P s"
using a0 P_Q unfolding cpn_def by fastforce
also have "(Γ, (P1, snd (last ((P0, sa) # xsa))) # ys) ∈ cpn n Γ Q (snd (?P0 ! ((length ?P0) -1)))"
using a3 cptn_eq_cptn_mod P_Q unfolding cpn_def
proof -
have "(n, Γ, (Q, snd (last ((P0, sa) # xsa))) # ys) ∈ cptn_mod_nest_call"
using a2 P_Q by blast
then have "(Γ, (Q, snd (last ((P0, sa) # xsa))) # ys) ∈ {(f, ps). ps ! 0 = (Q, snd (((P0, sa) # xsa) ! (Suc (length xsa) - 1))) ∧
(n, Γ, ps) ∈ cptn_mod_nest_call ∧ f = Γ}"
by (simp add: cptn_eq_cptn_mod last_length)
then show "(Γ, (P1, snd (last ((P0, sa) # xsa))) # ys) ∈ {(f, ps). ps ! 0 = (Q, snd (((P0, sa) # xsa) ! (length ((P0, sa) # xsa) - 1))) ∧ (n,Γ, ps) ∈ cptn_mod_nest_call ∧ f = Γ}"
using P_Q by force
qed
ultimately show ?thesis using a3 P_Q i using Cons_lift_catch_append by blast
qed
lemma Catch_sound4:
assumes
a0:"(n,Γ,x)∈cptn_mod_nest_call" and
a1:"x!0 = ((Catch P Q),s)" and
a2:"i<length x ∧ x!i=(Q,sj)" and
a3:"∀j<i. fst(x!j)≠Q" and
a4:"env_tran_right Γ x rely "
shows
"∃xs ys. (Γ,xs) ∈ (cpn n Γ P s) ∧ (Γ,ys) ∈ (cpn n Γ Q (snd (xs!(i-1)))) ∧ x = (map (lift_catch Q) xs)@ys"
using a0 a1 a2 a3 a4
proof (induct arbitrary: i sj P s)
case (CptnModNestEnv Γ C st t n xsa)
have a1:"Catch P Q ≠ Q" by simp
then have C_catch:"C=(Catch P Q)" using CptnModNestEnv by fastforce
then have "fst(((C, st) # (C, t) # xsa)!0) ≠Q" using CptnEnv a1 by auto
moreover have "fst(((C, st) # (C, t) # xsa)!1) ≠Q" using CptnModNestEnv a1 by auto
moreover have "fst(((C, st) # (C, t) # xsa)!i) =Q" using CptnModNestEnv by auto
ultimately have i_suc: "i> (Suc 0)" using CptnModNestEnv
by (metis Suc_eq_plus1 Suc_lessI add.left_neutral neq0_conv)
then obtain i' where i':"i=Suc i'" by (meson lessE)
then have i_minus:"i'=i-1" by auto
have "((C, t) # xsa) ! 0 = ((Catch P Q), t)"
using CptnModNestEnv by auto
moreover have "i'< length ((C,t)#xsa) ∧ ((C,t)#xsa)!i' = (Q,sj)"
using i' CptnModNestEnv(5) by force
moreover have "∀j<i'. fst (((C, t) # xsa) ! j) ≠ Q"
using i' CptnModNestEnv(6) by force
ultimately have hyp:"∃xs ys.
(Γ, xs) ∈ cpn n Γ P t ∧
(Γ, ys) ∈ cpn n Γ Q (snd (xs ! (i'-1))) ∧ (C, t) # xsa = map (lift_catch Q) xs @ ys"
using CptnModNestEnv(3) env_tran_tail CptnModNestEnv.prems(4) by blast
then obtain xs ys where xs_cp:"(Γ, xs) ∈ cpn n Γ P t ∧
(Γ, ys) ∈ cpn n Γ Q (snd (xs ! (i'-1))) ∧ (C, t) # xsa = map (lift_catch Q) xs @ ys"
by fast
have "(Γ, (P, s)#xs) ∈ cpn n Γ P s"
proof -
have "xs!0 = (P,t)"
using xs_cp unfolding cpn_def by blast
moreover have "xs≠[]"
using cpn_def CptnEmpty xs_cp by blast
ultimately obtain xs' where xs':"(n,Γ, (P,t)#xs') ∈ cptn_mod_nest_call ∧ xs=(P,t)#xs'"
using SmallStepCon.nth_tl xs_cp unfolding cpn_def by force
thus ?thesis using cpn_def
proof -
have "(Catch P Q, s) = (C, st)"
using CptnModNestEnv.prems(1) by auto
then have "Γ⊢⇩c (P, s) →⇩e (P, t)"
using Catch_env_P CptnModNestEnv(1) by blast
then show ?thesis
by (simp add: cpn_def cptn_mod_nest_call.CptnModNestEnv xs')
qed
qed
thus ?case
using i_suc Cons_lift_catch_append CptnModNestEnv.prems(1) i' i_minus xs_cp
by fastforce
next
case (CptnModNestSkip Γ P s t n xs)
then show ?case
by (metis (no_types) CptnModNestSkip.hyps(2) CptnModNestSkip.prems(1)
fst_conv nth_Cons_0 redex_not_Catch)
next
case (CptnModNestThrow Γ P s t n xs)
then show ?case
by (metis (no_types) CptnModNestThrow.hyps(2) CptnModNestThrow.prems(1)
fst_conv nth_Cons_0 redex_not_Catch)
next
case (CptnModNestCatch1 n Γ P0 s xs zs P1)
then show ?case
by (metis Catch_not_c Cons_lift_catch LanguageCon.com.inject(9)
fst_conv map_lift_catch_all_catch nth_Cons_0)
next
case (CptnModNestCatch2 n Γ P1 sa xsa ys zs Q1)
then have P_Q:"P=P1 ∧ Q = Q1" by force
thus ?case
proof (cases "Q1 = Skip")
case True thus ?thesis using catch_map_xs_ys'
CptnModNestCatch2 by blast
next
case False note q_not_throw=this
have "∀x. x< length ((LanguageCon.com.Catch P1 Q1, sa) # zs) ⟶
((LanguageCon.com.Catch P1 Q1, sa) # zs) ! x ≠ (Q, sj)" using CptnModNestCatch2
proof -
{
fix x
assume x_less:"x< length ((LanguageCon.com.Catch P1 Q1, sa) # zs)"
have "((LanguageCon.com.Catch P1 Q1, sa) # zs) ! x ≠ (Q, sj)"
proof (cases "x < length ((LanguageCon.com.Catch P1 Q1, sa)#map (lift_catch Q1) xsa)")
case True
then have eq_map:"((LanguageCon.com.Catch P1 Q1, sa) # zs) ! x = map (lift_catch Q1) ((P1, sa) # xsa) ! x"
by (metis (no_types) Cons_lift_catch Cons_lift_catch_append CptnModNestCatch2(6) True nth_append)
then have "∃ci si. map (lift_catch Q1) ((P1, sa) # xsa) ! x = (Catch ci Q1,si)"
using True unfolding lift_catch_def
by (metis Cons_lift_catch True eq_map map_lift_catch_all_catch surjective_pairing)
then have "((LanguageCon.com.Catch P1 Q1, sa) # zs) ! x ≠ (Q, sj)"
using P_Q eq_map by fastforce
thus ?thesis using CptnModNestCatch2(10) by auto
next
case False
let ?s' = "snd (last ((P1, sa) # xsa))"
have all_throw:"∀i<length ((LanguageCon.com.Skip, ?s')# ys).
fst (((Skip, ?s')# ys)!i) = Skip" using CptnModNestCatch2
by (metis cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod skip_all_skip)
then have
"∀x≥ length ((LanguageCon.com.Catch P1 Q1, sa) # map (lift_catch Q1) xsa).
x<length (((LanguageCon.com.Catch P1 Q1, sa) # zs)) ⟶
fst (((LanguageCon.com.Catch P1 Q1, sa) # zs) ! x) = Skip"
using Cons_lift_catch Cons_lift_catch_append CptnModNestCatch2.hyps(1) CptnModNestCatch2.hyps(3)
CptnModNestCatch2.hyps(4) CptnModNestCatch2.hyps(6) cptn_eq_cptn_mod_set
cptn_mod_nest_call.CptnModNestCatch2 cptn_mod_nest_cptn_mod skip_com_all_skip
proof-
{
fix x
assume a1:"x≥ length ((Catch P1 Q1, sa) # map (lift_catch Q1) xsa)" and
a2:"x<length (((Catch P1 Q1, sa) # zs))"
then have "((Catch P1 Q1, sa) # zs) ! x =
((Skip, ?s')# ys) !(x - (length ((Catch P1 Q1, sa) # map (lift_catch Q1) xsa)))"
using CptnModNestCatch2(6) by (metis Cons_lift_catch Cons_lift_catch_append not_le nth_append)
then have"fst (((Catch P1 Q1, sa) # zs) ! x) = Skip"
using all_throw a1 a2 CptnModNestCatch2.hyps(6) by auto
} thus ?thesis by auto
qed
thus ?thesis using False q_not_throw P_Q x_less
by (metis fst_conv not_le)
qed
} thus ?thesis by auto
qed
thus ?thesis using CptnModNestCatch2(8) by fastforce
qed
next
case (CptnModNestCatch3 n Γ P1 sa xsa ys zs Q1)
then show ?case using catch_map_xs_ys'[OF CptnModNestCatch3(1) CptnModNestCatch3(3) CptnModNestCatch3(5)
CptnModNestCatch3(7) CptnModNestCatch3(8) CptnModNestCatch3(9)]
by blast
qed(auto)
inductive_cases stepc_elim_cases_Catch_throw:
"Γ⊢⇩c(Catch c1 c2,s) → (Throw, Normal s1)"
inductive_cases stepc_elim_cases_Catch_skip_c2:
"Γ⊢⇩c(Catch c1 c2,s) → (c2,s)"
inductive_cases stepc_elim_cases_Catch_skip_2:
"Γ⊢⇩c(Catch c1 c2,s) → (Skip, s)"
lemma catch_skip_throw:
"Γ⊢⇩c(Catch c1 c2,s) → (c2,s) ⟹ (c2= Skip ∧ c1=Skip) ∨ (c1=Throw ∧ (∃s2'. s=Normal s2'))"
apply (rule stepc_elim_cases_Catch_skip_c2)
apply fastforce
apply (auto)+
using redex_not_Catch apply auto
done
lemma catch_skip_throw1:
"Γ⊢⇩c(Catch c1 c2,s) → (Skip,s) ⟹ (c1=Skip) ∨ (c1=Throw ∧ (∃s2'. s=Normal s2') ∧ c2 = Skip)"
apply (rule stepc_elim_cases_Catch_skip_2)
using redex_not_Catch apply auto
using redex_not_Catch by auto
lemma Catch_sound:
"Γ,Θ ⊢⇘/F⇙ c1 sat [p, R, G, q,r] ⟹
∀n. Γ,Θ ⊨n⇘/F⇙ c1 sat [p, R, G, q,r] ⟹
Γ,Θ ⊢⇘/F⇙ c2 sat [r, R, G, q,a] ⟹
∀n. Γ,Θ ⊨n⇘/F⇙ c2 sat [r, R, G, q,a] ⟹
Sta q R ⟹ (∀s. (Normal s,Normal s) ∈ G) ⟹
Γ,Θ ⊨n⇘/F⇙ (Catch c1 c2) sat [p, R, G, q,a]"
proof -
assume
a0:"Γ,Θ ⊢⇘/F⇙ c1 sat [p, R, G, q,r]" and
a1:"∀n. Γ,Θ ⊨n⇘/F⇙ c1 sat [p, R, G, q,r]" and
a2:"Γ,Θ ⊢⇘/F⇙ c2 sat [r, R, G, q,a]" and
a3: "∀n. Γ,Θ ⊨n⇘/F⇙ c2 sat [r, R, G, q,a]" and
a4: "Sta q R " and
a5: "(∀s. (Normal s, Normal s) ∈ G)"
{
fix s
assume all_call:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
then have a1:"Γ ⊨n⇘/F⇙ c1 sat [p, R, G, q,r]"
using a1 com_cvalidityn_def by fastforce
then have a3: "Γ ⊨n⇘/F⇙ c2 sat [r, R, G, q,a]"
using a3 com_cvalidityn_def all_call by fastforce
have "cpn n Γ (Catch c1 c2) s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ (Catch c1 c2) s" and a11:"c ∈ assum(p, R)"
then have a10':"c ∈ cp Γ (Catch c1 c2) s"
unfolding cpn_def cp_def
using cptn_if_cptn_mod cptn_mod_nest_cptn_mod by blast
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
have cp:"l!0=((Catch c1 c2),s) ∧ (n,Γ,l) ∈ cptn_mod_nest_call ∧ Γ=Γ1"
using a10 cpn_def c_prod by fastforce
then have cp':"l!0=((Catch c1 c2),s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1"
using cptn_eq_cptn_mod_nest by auto
have Γ1:"(Γ, l) = c" using c_prod cp by blast
have "c ∈ comm(G, (q,a)) F"
proof -
{
assume l_f:"snd (last l) ∉ Fault ` F"
have assum:"snd(l!0) ∈ Normal ` (p) ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ R)"
using a11 c_prod unfolding assum_def by simp
then have env_tran:"env_tran Γ p l R" using env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def unfolding env_tran_def by auto
have "(∀i. Suc i<length l ⟶
Γ⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)∧
(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof (cases "∀i<length l. fst (l!i)≠ c2")
case True
then have no_c2:"∀i<length l. fst (l!i)≠ c2" by assumption
show ?thesis
proof (cases "final (last l)")
case True
then obtain s' where "fst (last l) = Skip ∨ (fst (last l) = Throw ∧ snd (last l) = Normal s')"
using final_def by fast
thus ?thesis
proof
assume "fst (last l)= LanguageCon.com.Throw ∧ snd (last l) = Normal s'"
then have "False" using no_c2 env_tran_right cp' cptn_eq_cptn_mod_set Catch_sound3
by blast
thus ?thesis by auto
next
assume asm0:"fst (last l) = Skip"
then obtain lc1 ys where cp_lc1:"(Γ,lc1) ∈ cpn n Γ c1 s ∧ l = ((map (lift_catch c2) lc1)@((Skip,snd(last lc1))#ys))"
using Catch_sound2 cp cptn_eq_cptn_mod_set env_tran_right no_c2 by blast
then have cp_lc1':"(Γ,lc1) ∈ cp Γ c1 s"
unfolding cpn_def cp_def
using cptn_if_cptn_mod cptn_mod_nest_cptn_mod by fastforce
let ?m_lc1 = "map (lift_catch c2) lc1"
let ?lm_lc1 = "(length ?m_lc1)"
let ?last_m_lc1 = "?m_lc1!(?lm_lc1-1)"
have lc1_not_empty:"lc1 ≠ []"
using Γ1 a10' cpn_def cp_lc1 cp by auto
then have map_cp:"(Γ,?m_lc1) ∈ cpn n Γ (Catch c1 c2) s"
proof -
have f1: "lc1 ! 0 = (c1, s) ∧ (n, Γ, lc1) ∈ cptn_mod_nest_call ∧ Γ = Γ"
using cp_lc1 unfolding cpn_def by blast
then have f2: "(n,Γ, ?m_lc1) ∈ cptn_mod_nest_call" using lc1_not_empty
by (metis Cons_lift_catch SmallStepCon.nth_tl cptn_mod_nest_call.CptnModNestCatch1)
then show ?thesis
using f2 f1 lc1_not_empty by (simp add: cpn_def lift_catch_def)
qed
then have map_cp':"(Γ,?m_lc1) ∈ cp Γ (Catch c1 c2) s"
unfolding cp_def cpn_def
using cp_def cp_lc1' lift_catch_is_cptn by fastforce
also have map_assum:"(Γ,?m_lc1) ∈ assum (p,R)"
using sub_assum a10 a11 Γ1 cp_lc1 lc1_not_empty
by (metis SmallStepCon.nth_tl map_is_Nil_conv)
ultimately have "((Γ,lc1) ∈ assum(p, R))"
using Γ1 assum_map_catch cp_lc1' by blast
then have lc1_comm:"(Γ,lc1) ∈ comm(G, (q,r)) F"
using a1 cp_lc1 by (meson IntI com_validityn_def contra_subsetD)
then have m_lc1_comm:"(Γ,?m_lc1) ∈ comm(G, (q,r)) F"
using map_cp map_assum comm_map_catch cp_lc1' map_cp' by blast
then have last_m_lc1:"last (?m_lc1) = (Catch (fst (last lc1)) c2,snd (last lc1))"
proof -
have a000:"∀p c. (LanguageCon.com.Catch (fst p) c, snd p) = lift_catch c p"
using Cons_lift_catch by force
then show ?thesis
by (simp add: last_map a000 lc1_not_empty)
qed
then have last_length:"last (?m_lc1) = ?last_m_lc1"
using lc1_not_empty last_conv_nth list.map_disc_iff by blast
then have l_map:"l!(?lm_lc1-1)= ?last_m_lc1"
using cp_lc1
by (simp add:lc1_not_empty nth_append)
then have lm_lc1:"l!(?lm_lc1) = (Skip, snd (last lc1))"
using cp_lc1 by (meson nth_append_length)
then have step:"Γ⊢⇩c(l!(?lm_lc1-1)) → (l!(?lm_lc1))"
proof -
have "Γ⊢⇩c(l!(?lm_lc1-1)) →⇩c⇩e (l!(?lm_lc1))"
proof -
have f1: "0 ≤ length ys"
by blast
moreover have "Suc (length (map (lift_catch c2) lc1) + length ys) =
length (map (lift_catch c2) lc1 @ (LanguageCon.com.Skip, snd (last lc1)) # ys)"
by force
ultimately show ?thesis
by (metis (no_types) Suc_diff_1 Suc_eq_plus1 cp' cp_lc1 cptn_tran_ce_i
lc1_not_empty le_add_same_cancel1 length_greater_0_conv
less_Suc_eq_le list.map_disc_iff)
qed
moreover have "¬Γ⊢⇩c(l!(?lm_lc1-1)) →⇩e (l!(?lm_lc1))"
using last_m_lc1 last_length l_map
proof -
have "(LanguageCon.com.Catch (fst (last lc1)) c2, snd (last lc1)) = l ! (length (map (lift_catch c2) lc1) - 1)"
using l_map last_m_lc1 local.last_length by presburger
then show ?thesis
by (metis LanguageCon.com.simps(30) env_c_c' lm_lc1)
qed
ultimately show ?thesis using step_ce_elim_cases by blast
qed
have last_lc1_suc:"snd (l!(?lm_lc1-1)) = snd (l!?lm_lc1)"
using l_map last_m_lc1 lm_lc1 local.last_length by force
then have step_catch:"Γ⊢⇩c(Catch (fst (last lc1)) c2,snd (last lc1)) → (Skip, snd (last lc1))"
using l_map last_m_lc1 lm_lc1 local.last_length local.step
by presburger
then obtain s2' where
last_lc1:"fst (last lc1) = Skip ∨
fst (last lc1) = Throw ∧ (snd (last lc1) = Normal s2') ∧ c2 = Skip"
using catch_skip_throw1 by fastforce
then have last_lc1_skip:"fst (last lc1) = Skip"
proof
assume "fst (last lc1) = LanguageCon.com.Throw ∧
snd (last lc1) = Normal s2' ∧ c2 = LanguageCon.com.Skip"
thus ?thesis using no_c2 asm0
by (simp add: cp_lc1 last_conv_nth )
qed auto
have last_not_F:"snd (last ?m_lc1) ∉ Fault ` F"
proof -
have "snd ?last_m_lc1 = snd (l!(?lm_lc1-1))"
using l_map by auto
have "(?lm_lc1-1)< length l"using cp_lc1 by fastforce
also then have "snd (l!(?lm_lc1-1))∉ Fault ` F"
using cp' cp_lc1 l_f last_not_F[of Γ l F]
by fastforce
ultimately show ?thesis using l_map last_length by fastforce
qed
then have q_normal:"snd (l!?lm_lc1) ∈ Normal ` q "
proof -
have last_lc1:"fst (last lc1) = Skip"
using last_lc1_skip by fastforce
have "final (last lc1)" using last_lc1 final_def
by blast
then show ?thesis
using lc1_comm last_lc1 last_not_F
unfolding comm_def
using last_lc1_suc comm_dest2 l_map lm_lc1 local.last_length
by force
qed
then obtain s1' where normal_lm_lc1:"snd (l!?lm_lc1) = Normal s1' ∧ s1' ∈q"
by auto
have concl:"(∀i ns ns'. Suc i<length l ⟶
Γ⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof-
{ fix k ns ns'
assume a00:"Suc k<length l" and
a21:"Γ⊢⇩c(l!k) → (l!(Suc k))"
then have i_m_l:"∀i <?lm_lc1 . l!i = ?m_lc1!i"
using cp_lc1
proof -
have "map (lift c2) lc1 ≠ []"
by (meson lc1_not_empty list.map_disc_iff)
then show ?thesis
by (metis (no_types) cp_lc1 nth_append)
qed
have "(snd(l!k), snd(l!(Suc k))) ∈ G"
proof (cases "Suc k< ?lm_lc1")
case True
then have a11': "Γ⊢⇩c(?m_lc1!k) → (?m_lc1!(Suc k))"
using a11 i_m_l True
proof -
have "∀n na. ¬ 0 < n - Suc na ∨ na < n "
using diff_Suc_eq_diff_pred zero_less_diff by presburger
then show ?thesis
by (metis (no_types) True a21 i_m_l zero_less_diff)
qed
then have "(snd(?m_lc1!k), snd(?m_lc1!(Suc k))) ∈ G"
using a11' m_lc1_comm True comm_dest1 l_f last_not_F by fastforce
thus ?thesis using i_m_l True by auto
next
case False
then have "(Suc k=?lm_lc1) ∨ (Suc k>?lm_lc1)" by auto
thus ?thesis
proof
{assume suck:"(Suc k=?lm_lc1)"
then have k:"k=?lm_lc1-1" by auto
then obtain s1' where s1'_normal:"snd(l!?lm_lc1) = Normal s1'"
using q_normal by fastforce
have G_s1':"(Normal s1', Normal s1')∈ G" using a5 by auto
then show "(snd (l!k), snd (l!Suc k)) ∈ G"
proof -
have "snd (l ! k) = Normal s1'"
using k last_lc1_suc s1'_normal by presburger
then show ?thesis
using G_s1' s1'_normal suck by force
qed
}
next
{
assume a001:"Suc k>?lm_lc1"
have "∀i. i≥(length lc1) ∧ (Suc i < length l) ⟶
¬(Γ⊢⇩c(l!i) → (l!(Suc i)))"
using lm_lc1 lc1_not_empty
proof -
have "env_tran_right Γ1 l R"
by (metis cp env_tran_right)
then show ?thesis
using cp' fst_conv length_map lm_lc1 a001 a21 a00 a4
normal_lm_lc1
by (metis (no_types) only_one_component_tran_j)
qed
then have "¬(Γ⊢⇩c(l!k) → (l!(Suc k)))"
using a00 a001 by auto
then show ?thesis using a21 by fastforce
}
qed
qed
} thus ?thesis by auto
qed
have concr:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof -
have l_t:"fst (last l) = Skip"
using lm_lc1 by (simp add: asm0)
have "?lm_lc1 ≤ length l -1" using cp_lc1 by fastforce
also have "∀i. ?lm_lc1 ≤i ∧ i<(length l -1) ⟶ Γ⊢⇩c(l!i) →⇩e (l!(Suc i))"
using cp' fst_conv length_map lm_lc1 a4
normal_lm_lc1 only_one_component_tran_j[of Γ l ?lm_lc1 s1' q ]
by (metis Suc_eq_plus1 cptn_tran_ce_i env_tran_right less_diff_conv step_ce_elim_cases)
ultimately have "snd (l ! (length l - 1)) ∈ Normal ` q"
using cp_lc1 q_normal a4 env_tran_right stability[of q R l ?lm_lc1 "(length l) -1" _ Γ]
by fastforce
thus ?thesis using l_t
by (simp add: cp_lc1 last_conv_nth)
qed
note res = conjI [OF concl concr]
then show ?thesis using Γ1 c_prod unfolding comm_def by auto
qed
next
case False
then obtain lc1 where cp_lc1:"(Γ,lc1) ∈ cpn n Γ c1 s ∧ l = map (lift_catch c2) lc1"
using Catch_sound1 False no_c2 env_tran_right cp cptn_eq_cptn_mod_set
by blast
then have cp_lc1':"(Γ,lc1) ∈ cp Γ c1 s"
unfolding cpn_def cp_def
using cptn_eq_cptn_mod_nest by fastforce
then have "((Γ,lc1) ∈ assum(p, R))"
using Γ1 a10' a11 assum_map_catch cp_lc1
by blast
then have "(Γ, lc1)∈ comm(G, (q,r)) F" using cp_lc1 a1
by (meson IntI com_validityn_def contra_subsetD)
then have "(Γ, l)∈ comm(G, (q,r)) F"
using comm_map_catch a10' Γ1 cp_lc1 cp_lc1' by fastforce
then show ?thesis using l_f False
unfolding comm_def by fastforce
qed
next
case False
then obtain k where k_len:"k<length l ∧ fst (l ! k) = c2"
by blast
then have "∃m. (m < length l ∧ fst (l ! m) = c2) ∧
(∀i<m. ¬ (i < length l ∧ fst (l ! i) = c2))"
using a0 exists_first_occ[of "(λi. i<length l ∧ fst (l ! i) = c2)" k]
by blast
then obtain i where a0:"i<length l ∧ fst (l !i) = c2 ∧
(∀j<i. (fst (l ! j) ≠ c2))"
by fastforce
then obtain s2 where li:"l!i =(c2,s2)" by (meson eq_fst_iff)
then obtain lc1 lc2 where cp_lc1:"(Γ,lc1) ∈ (cpn n Γ c1 s) ∧
(Γ,lc2) ∈ (cpn n Γ c2 (snd (lc1!(i-1)))) ∧
l = (map (lift_catch c2) lc1)@lc2"
using Catch_sound4 a0 cp env_tran_right by blast
then have cp_lc1':"(Γ,lc1) ∈ (cp Γ c1 s) ∧
(Γ,lc2) ∈ (cp Γ c2 (snd (lc1!(i-1))))"
unfolding cp_def cpn_def using cptn_eq_cptn_mod_nest by fastforce
have i_not_fault:"snd (l!i) ∉ Fault ` F" using a0 cp' l_f last_not_F[of Γ l F] by blast
have length_c1_map:"length lc1 = length (map (lift_catch c2) lc1)"
by fastforce
then have i_map:"i=length lc1"
using cp_lc1 li a0 unfolding lift_catch_def
proof -
assume a1: "(Γ, lc1) ∈ cpn n Γ c1 s ∧ (Γ, lc2) ∈ cpn n Γ c2 (snd (lc1 ! (i - 1))) ∧ l = map (λ(P, s). (Catch P c2, s)) lc1 @ lc2"
have f2: "i < length l ∧ fst (l ! i) = c2 ∧ (∀n. ¬ n < i ∨ fst (l ! n) ≠ c2)"
using a0 by blast
have f3: "(Catch (fst (lc1 ! i)) c2, snd (lc1 ! i)) = lift_catch c2 (lc1 ! i)"
by (simp add: case_prod_unfold lift_catch_def)
then have "fst (l ! length lc1) = c2"
using a1 by (simp add: cpn_def nth_append)
thus ?thesis
using f3 f2
by (metis (no_types, lifting) Pair_inject a0 cp_lc1 f3
length_c1_map li linorder_neqE_nat nth_append nth_map seq_and_if_not_eq(12))
qed
have lc2_l:"∀j<length lc2. lc2!j=l!(i+j)"
using cp_lc1 length_c1_map i_map a0
by (metis nth_append_length_plus)
have lc1_not_empty:"lc1 ≠ []"
using cp cp_lc1 unfolding cpn_def by fastforce
have lc2_not_empty:"lc2 ≠ []"
using cpn_def cp_lc1 a0 i_map by force
have l_is:"s2= snd (last lc1)"
using cp_lc1 li a0 lc1_not_empty unfolding cpn_def
by (auto simp add: i_map last_conv_nth lc2_l)
let ?m_lc1 = "map (lift_catch c2) lc1"
have last_m_lc1:"l!(i-1) = (Catch (fst (last lc1)) c2,s2)"
using i_map cp_lc1 l_is last_lift_catch last_snd_catch lc1_not_empty length_c1_map
by (metis (no_types, lifting) One_nat_def diff_Suc_less last_conv_nth length_greater_0_conv nth_append prod.collapse)
have last_mcl1_not_F:"snd (last ?m_lc1) ∉ Fault ` F"
by (metis One_nat_def i_not_fault l_is last_conv_nth last_snd_catch li list.map_disc_iff snd_conv)
have map_cp:"(Γ,?m_lc1) ∈ cpn n Γ (Catch c1 c2) s"
proof -
have f1: "lc1 ! 0 = (c1, s) ∧ (n,Γ, lc1) ∈ cptn_mod_nest_call ∧ Γ = Γ"
using cp_lc1 cpn_def by blast
then have f2: "(n, Γ, ?m_lc1) ∈ cptn_mod_nest_call" using lc1_not_empty
by (metis Cons_lift_catch SmallStepCon.nth_tl cptn_mod_nest_call.CptnModNestCatch1)
then show ?thesis
using f2 f1 lc1_not_empty by (simp add: cpn_def lift_catch_def)
qed
then have map_cp':"(Γ,?m_lc1) ∈ cp Γ (Catch c1 c2) s"
unfolding cp_def cpn_def using cptn_eq_cptn_mod_nest by fastforce
also have map_assum:"(Γ,?m_lc1) ∈ assum (p,R)"
using sub_assum a10 a11 Γ1 cp_lc1 lc1_not_empty
by (metis SmallStepCon.nth_tl map_is_Nil_conv)
ultimately have "((Γ,lc1) ∈ assum(p, R))"
using Γ1 assum_map_catch using assum_map cp_lc1 cp_lc1' by blast
then have lc1_comm:"(Γ,lc1) ∈ comm(G, (q,r)) F"
using a1 cp_lc1 by (meson IntI com_validityn_def contra_subsetD)
then have m_lc1_comm:"(Γ,?m_lc1) ∈ comm(G, (q,r)) F"
using map_cp' map_assum comm_map_catch cp_lc1 cp_lc1' by blast
then have "Γ⊢⇩c(l!(i-1)) → (l!i)"
proof -
have "Γ⊢⇩c(l!(i-1)) →⇩c⇩e (l!(i))"
by (metis Suc_eq_plus1 Suc_pred' a0 cp' cptn_tran_ce_i i_map
lc1_not_empty length_greater_0_conv)
moreover have "¬Γ⊢⇩c(l!(i-1)) →⇩e (l!i)"
using li last_m_lc1
by (metis (no_types, lifting) env_c_c' seq_and_if_not_eq(12))
ultimately show ?thesis using step_ce_elim_cases by blast
qed
then have step:"Γ⊢⇩c(Catch (fst (last lc1)) c2,s2) → (c2, s2)"
using last_m_lc1 li by fastforce
then obtain s2' where
last_lc1:"(fst (last lc1) = Skip ∧ c2 = Skip) ∨
fst (last lc1) = Throw ∧ (s2 = Normal s2')"
using catch_skip_throw by blast
have final:"final (last lc1)"
using last_lc1 l_is unfolding final_def by auto
have normal_last:"fst (last lc1) = Skip ∧ snd (last lc1) ∈ Normal ` q ∨
fst (last lc1) = Throw ∧ snd (last lc1) ∈ Normal ` r"
proof -
have "snd (last lc1) ∉ Fault ` F"
using i_not_fault l_is li by auto
then show ?thesis
using final comm_dest2 lc1_comm by blast
qed
obtain s2' where lastlc1_normal:"snd (last lc1) = Normal s2'"
using normal_last by blast
then have Normals2:"s2 = Normal s2'" by (simp add: l_is)
have Gs2':" (Normal s2', Normal s2')∈G" using a5 by auto
have concl:
"(∀i. Suc i<length l ⟶
Γ⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof-
{ fix k ns ns'
assume a00:"Suc k<length l" and
a21:"Γ⊢⇩c(l!k) → (l!(Suc k))"
have i_m_l:"∀j <i . l!j = ?m_lc1!j"
proof -
have "map (lift c2) lc1 ≠ []"
by (meson lc1_not_empty list.map_disc_iff)
then show ?thesis
using cp_lc1 i_map length_c1_map by (fastforce simp:nth_append)
qed
have "(snd(l!k), snd(l!(Suc k))) ∈ G"
proof (cases "Suc k< i")
case True
then have a11': "Γ⊢⇩c(?m_lc1!k) → (?m_lc1!(Suc k))"
using a11 i_m_l True
proof -
have "∀n na. ¬ 0 < n - Suc na ∨ na < n "
using diff_Suc_eq_diff_pred zero_less_diff by presburger
then show ?thesis using True a21 i_m_l by force
qed
have "Suc k < length ?m_lc1" using True i_map length_c1_map by metis
then have "(snd(?m_lc1!k), snd(?m_lc1!(Suc k))) ∈ G"
using a11' last_mcl1_not_F m_lc1_comm True i_map length_c1_map comm_dest1[of Γ]
by blast
thus ?thesis using i_m_l True by auto
next
case False
have "(Suc k=i) ∨ (Suc k>i)" using False by auto
thus ?thesis
proof
{ assume suck:"(Suc k=i)"
then have k:"k=i-1" by auto
then show "(snd (l!k), snd (l!Suc k)) ∈ G"
using Gs2' Normals2 last_m_lc1 li suck by auto
}
next
{
assume a001:"Suc k>i"
then have k:"k≥i" by fastforce
then obtain k' where k':"k=i+k'"
using add.commute le_Suc_ex by blast
{assume skip:"c2=Skip"
then have "∀k. k≥i ∧ (Suc k < length l) ⟶
¬(Γ⊢⇩c(l!k) → (l!(Suc k)))"
using Normals2 li lastlc1_normal a21 a001 a00 a4
a0 skip env_tran_right cp'
by (metis SmallStepCon.final_def SmallStepCon.no_step_final' Suc_lessD skip_com_all_skip)
then have ?thesis using a21 a001 k a00 by blast
} note left=this
{assume "c2≠Skip"
then have "fst (last lc1) = Throw"
using last_m_lc1 last_lc1 by simp
then have s2_normal:"s2 ∈ Normal ` r"
using normal_last lastlc1_normal Normals2
by fastforce
have length_lc2:"length l=i+length lc2"
using i_map cp_lc1 by fastforce
have "(Γ,lc2) ∈ assum (r,R)"
proof -
have left:"snd (lc2!0) ∈ Normal ` r"
using li lc2_l s2_normal lc2_not_empty by fastforce
{
fix j
assume j_len:"Suc j<length lc2" and
j_step:"Γ⊢⇩c(lc2!j) →⇩e (lc2!(Suc j))"
then have suc_len:"Suc (i + j)<length l" using j_len length_lc2
by fastforce
also then have "Γ⊢⇩c(l!(i+j)) →⇩e (l! (Suc (i+ j)))"
using lc2_l j_step j_len by fastforce
ultimately have "(snd(lc2!j), snd(lc2!(Suc j))) ∈ R"
using assum suc_len lc2_l j_len cp by fastforce
}
then show ?thesis using left
unfolding assum_def by fastforce
qed
also have "(Γ,lc2) ∈ cpn n Γ c2 s2"
using cp_lc1 i_map l_is last_conv_nth lc1_not_empty by fastforce
ultimately have comm_lc2:"(Γ,lc2) ∈ comm (G, (q,a)) F"
using a3 unfolding com_validityn_def by auto
have lc2_last_f:"snd (last lc2)∉ Fault ` F"
using lc2_l lc2_not_empty l_f cp_lc1 by fastforce
have suck':"Suc k' < length lc2"
using k' a00 length_lc2 by arith
moreover then have "Γ⊢⇩c(lc2!k') → (lc2!(Suc k'))"
using k' lc2_l a21 by fastforce
ultimately have "(snd (lc2! k'), snd (lc2 ! Suc k')) ∈ G"
using comm_lc2 lc2_last_f comm_dest1[of Γ lc2 G q a F k']
by blast
then have ?thesis using suck' lc2_l k' by fastforce
}
then show ?thesis using left by auto
}
qed
qed
} thus ?thesis by auto
qed note left=this
have right:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof -
{ assume final_l:"final (last l)"
have eq_last_lc2_l:"last l=last lc2" by (simp add: cp_lc1 lc2_not_empty)
then have final_lc2:"final (last lc2)" using final_l by auto
{
assume lst_lc1_skip:"fst (last lc1) = Skip"
then have c2_skip:"c2 = Skip"
using step lastlc1_normal LanguageCon.com.distinct(17) last_lc1
by auto
have Skip:"fst (l!(length l - 1)) = Skip"
using li Normals2 env_tran_right cp' c2_skip a0
i_skip_all_skip[of Γ l i "(length l) - 1" _]
by fastforce
have s2_a:"s2 ∈ Normal ` q"
using normal_last
by (simp add: lst_lc1_skip l_is)
then have "∀ia. i ≤ ia ∧ ia < length l - 1 ⟶ Γ⊢⇩c l ! ia →⇩e l ! Suc ia"
using c2_skip li Normals2 a0 cp' env_tran_right final_def
by (metis (no_types, hide_lams) One_nat_def SmallStepCon.no_step_final'
Suc_lessD add.right_neutral add_Suc_right
cptn_tran_ce_i i_skip_all_skip less_diff_conv step_ce_elim_cases)
then have "snd (l!(length l - 1)) ∈ Normal ` q ∧ fst (l!(length l - 1)) = Skip"
using a0 s2_a li a4 env_tran_right stability[of q R l i "(length l) -1" _ Γ] Skip
by (metis One_nat_def Suc_pred length_greater_0_conv lessI linorder_not_less list.size(3)
not_less0 not_less_eq_eq snd_conv)
then have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a))"
using a0 by (metis last_conv_nth list.size(3) not_less0)
} note left = this
{ assume "fst (last lc1) = Throw"
then have s2_normal:"s2 ∈ Normal ` r"
using normal_last lastlc1_normal Normals2
by fastforce
have length_lc2:"length l=i+length lc2"
using i_map cp_lc1 by fastforce
have "(Γ,lc2) ∈ assum (r,R)"
proof -
have left:"snd (lc2!0) ∈ Normal ` r"
using li lc2_l s2_normal lc2_not_empty by fastforce
{
fix j
assume j_len:"Suc j<length lc2" and
j_step:"Γ⊢⇩c(lc2!j) →⇩e (lc2!(Suc j))"
then have suc_len:"Suc (i + j)<length l" using j_len length_lc2
by fastforce
also then have "Γ⊢⇩c(l!(i+j)) →⇩e (l! (Suc (i+ j)))"
using lc2_l j_step j_len by fastforce
ultimately have "(snd(lc2!j), snd(lc2!(Suc j))) ∈ R"
using assum suc_len lc2_l j_len cp by fastforce
}
then show ?thesis using left
unfolding assum_def by fastforce
qed
also have "(Γ,lc2) ∈ cpn n Γ c2 s2"
using cp_lc1 i_map l_is last_conv_nth lc1_not_empty by fastforce
ultimately have comm_lc2:"(Γ,lc2) ∈ comm (G, (q,a)) F"
using a3 unfolding com_validityn_def by auto
have lc2_last_f:"snd (last lc2)∉ Fault ` F"
using lc2_l lc2_not_empty l_f cp_lc1 by fastforce
then have "((fst (last lc2) = Skip ∧
snd (last lc2) ∈ Normal ` q)) ∨
(fst (last lc2) = Throw ∧
snd (last lc2) ∈ Normal ` (a))"
using final_lc2 comm_lc2 unfolding comm_def by auto
then have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a))"
using eq_last_lc2_l by auto
}
then have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a))"
using left using last_lc1 by auto
} thus ?thesis by auto qed
thus ?thesis using left l_f Γ1 unfolding comm_def by force
qed
} thus ?thesis using Γ1 unfolding comm_def by auto qed
} thus ?thesis by auto qed
} thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma DynCom_sound:
"(∀s ∈ p. ((Γ,Θ⊢⇘/F⇙ (c1 s) sat [p, R, G, q,a]) ∧
(∀n. (Γ,Θ ⊨n⇘/F⇙ (c1 s) sat [p,R, G, q,a])))) ⟹
(∀s. (Normal s, Normal s) ∈ G) ⟹
(Sta p R) ∧ (Sta q R) ∧ (Sta a R) ⟹
Γ,Θ ⊨n⇘/F⇙ (DynCom c1) sat [p, R, G, q,a]"
proof -
assume
a0:"(∀s ∈ p. ((Γ,Θ⊢⇘/F⇙ (c1 s) sat [p, R, G, q,a]) ∧
(∀n. (Γ,Θ ⊨n⇘/F⇙ (c1 s) sat [p, R, G, q,a]))))" and
a1:"∀s. (Normal s, Normal s) ∈ G" and
a2: "(Sta p R) ∧ (Sta q R) ∧ (Sta a R)"
{
fix s
assume all_DynCom:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
then have a0:"(∀s ∈ p. (Γ ⊨n⇘/F⇙ (c1 s) sat [p, R, G, q,a]))"
using a0 unfolding com_cvalidityn_def by fastforce
have "cpn n Γ(DynCom c1) s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ (DynCom c1) s" and a11:"c ∈ assum(p, R)"
then have a10':"c ∈ cp Γ (DynCom c1) s"
unfolding cp_def cpn_def
using cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod by fastforce
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
have "c ∈ comm(G, (q,a)) F"
proof -
{assume l_f:"snd (last l) ∉ Fault ` F"
have cp:"l!0=(DynCom c1,s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1"
using a10' cp_def c_prod by fastforce
have Γ1:"(Γ, l) = c" using c_prod cp by blast
have assum:"snd(l!0) ∈ Normal ` (p) ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ R)"
using a11 c_prod unfolding assum_def by simp
then have env_tran:"env_tran Γ p l R" using env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def unfolding env_tran_def by auto
obtain ns where s_normal:"s=Normal ns ∧ ns∈ p"
using cp assum by fastforce
have concl:"(∀i. Suc i<length l ⟶
Γ1⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof -
{ fix k ns ns'
assume a00:"Suc k<length l" and
a21:"Γ⊢⇩c(l!k) → (l!(Suc k))"
obtain j where before_k_all_evnt:"j≤k ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ (∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
using a00 a21 exist_first_comp_tran cp by blast
then obtain cj sj csj ssj where pair_j:"(Γ⊢⇩c(cj,sj) → (csj,ssj)) ∧ cj = fst (l!j) ∧ sj = snd (l!j) ∧ csj = fst (l!(Suc j)) ∧ ssj = snd(l!(Suc j))"
by fastforce
have k_basic:"cj = (DynCom c1) ∧ sj ∈ Normal ` (p)"
using pair_j before_k_all_evnt a2 cp env_tran_right assum a00 stability[of p R l 0 j j Γ]
by force
then obtain s' where ss:"sj = Normal s' ∧ s'∈ (p)" by auto
then have ssj_normal_s:"ssj = Normal s'"
using before_k_all_evnt k_basic pair_j a0
by (metis snd_conv stepc_Normal_elim_cases(10))
have "(snd(l!k), snd(l!(Suc k))) ∈ G"
using ss a2 unfolding Satis_def
proof (cases "k=j")
case True
have "(Normal s', Normal s')∈G" using a1 by fastforce
thus "(snd (l ! k), snd (l ! Suc k)) ∈ G"
using pair_j k_basic True ss ssj_normal_s by auto
next
case False
have j_k:"j<k" using before_k_all_evnt False by fastforce
thus "(snd (l ! k), snd (l ! Suc k)) ∈ G"
proof -
have j_length:"Suc j < length l" using a00 before_k_all_evnt by fastforce
have p1:"s'∈p ∧ ssj=Normal s'" using ss ssj_normal_s by fastforce
then have c1_valid:"(Γ ⊨n⇘/F⇙ (c1 s') sat [p, R, G, q,a])"
using a0 by fastforce
have cj:"csj= (c1 s')" using k_basic pair_j ss a0 s_normal
proof -
have "Γ⊢⇩c (LanguageCon.com.DynCom c1, Normal s') → (csj, ssj)"
using k_basic pair_j ss by force
then have "(csj, ssj) = (c1 s', Normal s')"
by (meson stepc_Normal_elim_cases(10))
then show ?thesis
by blast
qed
moreover then have "cpn n Γ csj ssj ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
using a2 com_validityn_def cj p1 c1_valid by blast
moreover then have "l!(Suc j) = (csj, Normal s')"
using before_k_all_evnt pair_j cj ssj_normal_s
by fastforce
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using p1 j_length a11 Γ1 ssj_normal_s
using a10 cpn_assum_induct by fastforce
then show ?thesis
using a00 a21 a10' Γ1 j_k j_length l_f
cptn_comm_induct[of Γ l "DynCom c1" s _ "Suc j" G q a F k ]
unfolding Satis_def by fastforce
qed
qed
} thus ?thesis by (simp add: c_prod cp) qed
have concr:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof-
{
assume valid:"final (last l)"
have "∃k. k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k))) ∧ final (l!(Suc k))"
proof -
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a1 l1 where l:"l=a1#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a1 l1] l by fastforce
have final_0:"¬final(l!0)" using cp unfolding final_def by auto
have "0≤ (length l-1)" using len_l last_l by auto
moreover have "(length l-1) < length l" using len_l by auto
moreover have "final (l!(length l-1))" using valid last_l by auto
moreover have "fst (l!0) = DynCom c1" using cp by auto
ultimately show ?thesis
using a2 cp final_exist_component_tran_final env_tran_right final_0
by blast
qed
then obtain k where a21: "k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k))) ∧ final (l!(Suc k))"
by auto
then have a00:"Suc k<length l" by fastforce
then obtain j where before_k_all_evnt:"j≤k ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ (∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
using a00 a21 exist_first_comp_tran cp by blast
then obtain cj sj csj ssj where pair_j:"(Γ⊢⇩c(cj,sj) → (csj,ssj)) ∧ cj = fst (l!j) ∧ sj = snd (l!j) ∧ csj = fst (l!(Suc j)) ∧ ssj = snd(l!(Suc j))"
by fastforce
have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a))"
proof -
have j_length:"Suc j < length l" using a00 before_k_all_evnt by fastforce
then have k_basic:"cj = (DynCom c1) ∧ sj ∈ Normal ` (p)"
using a2 pair_j before_k_all_evnt cp env_tran_right assum stability[of p R l 0 j j Γ]
by force
then obtain s' where ss:"sj = Normal s' ∧ s'∈ (p)" by auto
then have ssj_normal_s:"ssj = Normal s'"
using before_k_all_evnt k_basic pair_j a0
by (metis snd_conv stepc_Normal_elim_cases(10))
have cj:"csj=c1 s'" using k_basic pair_j ss a0
by (metis fst_conv stepc_Normal_elim_cases(10))
moreover have p1:"s'∈p" using ss by blast
moreover then have "cpn n Γ csj ssj ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
using a0 com_validityn_def cj by blast
moreover then have "l!(Suc j) = (csj, Normal s')"
using before_k_all_evnt pair_j cj ssj_normal_s
by fastforce
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using j_length a10 a11 Γ1 ssj_normal_s
by (meson contra_subsetD cpn_assum_induct)
thus ?thesis
using j_length l_f drop_comm a10' Γ1 cptn_comm_induct[of Γ l "DynCom c1" s _ "Suc j" G q a F "Suc j"] valid
by blast
qed
} thus ?thesis by auto
qed
note res = conjI [OF concl concr]}
thus ?thesis using c_prod unfolding comm_def by force qed
} thus ?thesis by auto qed
} thus ?thesis by (auto simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma Guard_sound:
"Γ,Θ ⊢⇘/F⇙ c1 sat [p ∩ g, R, G, q,a] ⟹
(∀n. Γ,Θ ⊨n⇘/F⇙ c1 sat [p ∩ g, R, G, q,a]) ⟹
Sta (p ∩ g) R ⟹ (∀s. (Normal s, Normal s) ∈ G) ⟹
Γ,Θ ⊨n⇘/F⇙ (Guard f g c1) sat [p ∩ g, R, G, q,a]"
proof -
assume
a0:"Γ,Θ ⊢⇘/F⇙ c1 sat [(p ∩ g) , R, G, q,a]" and
a1:"(∀n. Γ,Θ ⊨n⇘/F⇙ c1 sat [p ∩ g, R, G, q,a])" and
a2: "Sta (p ∩ g) R" and
a3: "∀s. (Normal s, Normal s) ∈ G"
{
fix s
assume all_call:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
then have a1:"Γ ⊨n⇘/F⇙ c1 sat [p ∩ g, R, G, q,a]"
using a1 com_cvalidityn_def by fastforce
have "cpn n Γ (Guard f g c1) s ∩ assum(p ∩ g, R) ⊆ comm(G, (q,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ (Guard f g c1) s" and a11:"c ∈ assum(p ∩ g, R)"
then have a10':"c ∈ cp Γ (Guard f g c1) s"
unfolding cpn_def cp_def using cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod by fastforce
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
have "c ∈ comm(G, (q,a)) F"
proof -
{assume l_f:"snd (last l) ∉ Fault ` F"
have cp:"l!0=((Guard f g c1),s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1" using a10' cp_def c_prod by fastforce
have Γ1:"(Γ, l) = c" using c_prod cp by blast
have assum:"snd(l!0) ∈ Normal ` (p ∩ g) ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ R)"
using a11 c_prod unfolding assum_def by simp
then have env_tran:"env_tran Γ (p ∩ g) l R" using env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def unfolding env_tran_def by auto
have concl:"(∀i. Suc i<length l ⟶
Γ1⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof -
{ fix k ns ns'
assume a00:"Suc k<length l" and
a21:"Γ⊢⇩c(l!k) → (l!(Suc k))"
obtain j where before_k_all_evnt:"j≤k ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ (∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
using a00 a21 exist_first_comp_tran cp by blast
then obtain cj sj csj ssj where pair_j:"(Γ⊢⇩c(cj,sj) → (csj,ssj)) ∧ cj = fst (l!j) ∧ sj = snd (l!j) ∧ csj = fst (l!(Suc j)) ∧ ssj = snd(l!(Suc j))"
by fastforce
have k_basic:"cj =(Guard f g c1) ∧ sj ∈ Normal ` (p ∩ g)"
using pair_j before_k_all_evnt cp env_tran_right a2 assum a00 stability[of "p ∩ g" R l 0 j j Γ]
by force
then obtain s' where ss:"sj = Normal s' ∧ s'∈ (p ∩ g)" by auto
then have ssj_normal_s:"ssj = Normal s'"
using before_k_all_evnt k_basic pair_j a0 stepc_Normal_elim_cases(2)
by (metis (no_types, lifting) IntD2 prod.inject)
have "(snd(l!k), snd(l!(Suc k))) ∈ G"
using ss a2 unfolding Satis_def
proof (cases "k=j")
case True
have " (Normal s', Normal s')∈G" using a3 by auto
thus "(snd (l ! k), snd (l ! Suc k)) ∈ G"
using pair_j k_basic True ss ssj_normal_s by auto
next
case False
have j_k:"j<k" using before_k_all_evnt False by fastforce
thus "(snd (l ! k), snd (l ! Suc k)) ∈ G"
proof -
have j_length:"Suc j < length l" using a00 before_k_all_evnt by fastforce
have cj:"csj=c1" using k_basic pair_j ss a0
by (metis (no_types, lifting) IntD2 fst_conv stepc_Normal_elim_cases(2))
moreover have p1:"s' ∈ (p ∩ g)" using ss by blast
moreover then have "cpn n Γ csj ssj ∩ assum(p ∩ g, R) ⊆ comm(G, (q,a)) F"
using a1 com_validityn_def cj by blast
moreover then have "l!(Suc j) = (csj, Normal s')"
using before_k_all_evnt pair_j cj ssj_normal_s
by fastforce
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using j_length a10 a11 Γ1 ssj_normal_s
cpn_assum_induct
by fastforce
then show ?thesis
using a00 a21 a10' Γ1 j_k j_length l_f
cptn_comm_induct[of Γ l "(Guard f g c1)" s _ "Suc j" G q a F k ]
unfolding Satis_def by fastforce
qed
qed
} thus ?thesis by (simp add: c_prod cp) qed
have concr:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof-
{
assume valid:"final (last l)"
have "∃k. k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k))) ∧ final (l!(Suc k))"
proof -
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a1 l1 where l:"l=a1#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a1 l1] l by fastforce
have final_0:"¬final(l!0)" using cp unfolding final_def by auto
have "0≤ (length l-1)" using len_l last_l by auto
moreover have "(length l-1) < length l" using len_l by auto
moreover have "final (l!(length l-1))" using valid last_l by auto
moreover have "fst (l!0) = (Guard f g c1)" using cp by auto
ultimately show ?thesis
using cp final_exist_component_tran_final env_tran_right final_0
by blast
qed
then obtain k where a21: "k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k))) ∧ final (l!(Suc k))"
by auto
then have a00:"Suc k<length l" by fastforce
then obtain j where before_k_all_evnt:"j≤k ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ (∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
using a00 a21 exist_first_comp_tran cp by blast
then obtain cj sj csj ssj where pair_j:"(Γ⊢⇩c(cj,sj) → (csj,ssj)) ∧ cj = fst (l!j) ∧ sj = snd (l!j) ∧ csj = fst (l!(Suc j)) ∧ ssj = snd(l!(Suc j))"
by fastforce
have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a))"
proof -
have j_length:"Suc j < length l" using a00 before_k_all_evnt by fastforce
then have k_basic:"cj = (Guard f g c1) ∧ sj ∈ Normal ` (p ∩ g)"
using pair_j before_k_all_evnt cp env_tran_right a2 assum a00 stability[of "p ∩ g" R l 0 j j Γ]
by force
then obtain s' where ss:"sj = Normal s' ∧ s'∈ (p ∩ g)" by auto
then have ssj_normal_s:"ssj = Normal s'"
using before_k_all_evnt k_basic pair_j a1
by (metis (no_types, lifting) IntD2 Pair_inject stepc_Normal_elim_cases(2))
have cj:"csj=c1" using k_basic pair_j ss a0
by (metis (no_types, lifting) fst_conv IntD2 stepc_Normal_elim_cases(2))
moreover have p1:"s' ∈ (p ∩ g)" using ss by blast
moreover then have "cpn n Γ csj ssj ∩ assum((p ∩ g), R) ⊆ comm(G, (q,a)) F"
using a1 com_validityn_def cj by blast
moreover then have "l!(Suc j) = (csj, Normal s')"
using before_k_all_evnt pair_j cj ssj_normal_s
by fastforce
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using j_length a10 a11 Γ1 ssj_normal_s cpn_assum_induct
by fastforce
thus ?thesis
using j_length l_f drop_comm a10' Γ1 cptn_comm_induct[of Γ l "(Guard f g c1)" s _ "Suc j" G q a F "Suc j"] valid
by blast
qed
} thus ?thesis by auto
qed
note res = conjI [OF concl concr]}
thus ?thesis using c_prod unfolding comm_def by force qed
} thus ?thesis by auto qed
} thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma Guarantee_sound:
"Γ,Θ ⊢⇘/F⇙ c1 sat [(p ∩ g), R, G, q,a] ⟹
∀n. Γ,Θ ⊨n⇘/F⇙ c1 sat [p ∩ g, R, G, q,a] ⟹
Sta p R ⟹
f∈F ⟹
(∀s. (Normal s, Normal s) ∈ G) ⟹
Γ,Θ ⊨n⇘/F⇙ (Guard f g c1) sat [p, R, G, q,a]"
proof -
assume
a0:"Γ,Θ ⊢⇘/F⇙ c1 sat [p ∩ g, R, G, q,a]" and
a1:"∀n. Γ,Θ ⊨n⇘/F⇙ c1 sat [p ∩ g, R, G, q,a]" and
a2: "Sta p R" and
a3: "(∀s. (Normal s, Normal s) ∈ G)" and
a4: "f∈F"
{
fix s
assume all_call:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
then have a1:"Γ ⊨n⇘/F⇙ c1 sat [p ∩ g, R, G, q,a]"
using a1 com_cvalidityn_def by fastforce
have "cpn n Γ (Guard f g c1) s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ (Guard f g c1) s" and a11:"c ∈ assum(p, R)"
then have a10':"c ∈ cp Γ (Guard f g c1) s"
unfolding cp_def cpn_def using cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod by fast
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
have "c ∈ comm(G, (q,a)) F"
proof -
{assume l_f:"snd (last l) ∉ Fault ` F"
have cp:"l!0=((Guard f g c1),s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1" using a10' cp_def c_prod by fastforce
have Γ1:"(Γ, l) = c" using c_prod cp by blast
have assum:"snd(l!0) ∈ Normal ` (p) ∧ (∀i. Suc i<length l ⟶
(Γ1)⊢⇩c(l!i) →⇩e (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ R)"
using a11 c_prod unfolding assum_def by simp
then have env_tran:"env_tran Γ p l R" using env_tran_def cp by blast
then have env_tran_right: "env_tran_right Γ l R"
using env_tran env_tran_right_def unfolding env_tran_def by auto
have concl:"(∀i ns ns'. Suc i<length l ⟶
Γ1⊢⇩c(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G)"
proof -
{ fix k ns ns'
assume a00:"Suc k<length l" and
a21:"Γ⊢⇩c(l!k) → (l!(Suc k))"
obtain j where before_k_all_evnt:"j≤k ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ (∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
using a00 a21 exist_first_comp_tran cp by blast
then obtain cj sj csj ssj where pair_j:"(Γ⊢⇩c(cj,sj) → (csj,ssj)) ∧ cj = fst (l!j) ∧ sj = snd (l!j) ∧ csj = fst (l!(Suc j)) ∧ ssj = snd(l!(Suc j))"
by fastforce
have k_basic:"cj =(Guard f g c1) ∧ sj ∈ Normal ` (p)"
using pair_j before_k_all_evnt cp env_tran_right a2 assum a00 stability[of p R l 0 j j Γ]
by force
then obtain s' where ss:"sj = Normal s' ∧ s'∈ (p)" by auto
have or:"s'∈ (g ∪ (-g))" by fastforce
{assume "s' ∈ g"
then have k_basic:"cj =(Guard f g c1) ∧ sj ∈ Normal ` (p ∩ g)"
using ss k_basic by fastforce
then have ss: "sj = Normal s' ∧ s'∈ (p ∩ g)"
using ss by fastforce
have ssj_normal_s:"ssj = Normal s'"
using ss before_k_all_evnt k_basic pair_j a0 stepc_Normal_elim_cases(2)
by (metis (no_types, lifting) IntD2 prod.inject)
have "(snd(l!k), snd(l!(Suc k))) ∈ G"
using ss a2 unfolding Satis_def
proof (cases "k=j")
case True
have "(Normal s', Normal s') ∈ G" using a3 by auto
thus "(snd (l ! k), snd (l ! Suc k)) ∈ G"
using pair_j k_basic True ss ssj_normal_s by auto
next
case False
have j_k:"j<k" using before_k_all_evnt False by fastforce
thus "(snd (l ! k), snd (l ! Suc k)) ∈ G"
proof -
have j_length:"Suc j < length l" using a00 before_k_all_evnt by fastforce
have cj:"csj=c1" using k_basic pair_j ss a0
by (metis (no_types, lifting) fst_conv IntD2 stepc_Normal_elim_cases(2))
moreover have p1:"s' ∈ (p ∩ g)" using ss by blast
moreover then have "cpn n Γ csj ssj ∩ assum((p ∩ g), R) ⊆ comm(G, (q,a)) F"
using a1 com_validityn_def cj by blast
moreover then have "l!(Suc j) = (csj, Normal s')"
using before_k_all_evnt pair_j cj ssj_normal_s
by fastforce
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using j_length a10 a11 Γ1 ssj_normal_s cpn_assum_induct
by fastforce
then show ?thesis
using a3 a00 a21 a10' Γ1 j_k j_length l_f
cptn_comm_induct[of Γ l "(Guard f g c1)" s _ "Suc j" G q a F k]
unfolding Satis_def by fastforce
qed
qed
} note p1=this
{
assume "s'∉g"
then have csj_skip:"csj= Skip ∧ ssj=Fault f" using k_basic ss pair_j
by (meson Pair_inject stepc_Normal_elim_cases(2))
then have "snd (last l) = Fault f" using pair_j
proof -
have "j = k"
proof -
have f1: "k < length l"
using a00 by linarith
have "¬ SmallStepCon.final (l ! k)"
by (metis SmallStepCon.no_step_final' a21)
then have "¬ Suc j ≤ k"
using f1 SmallStepCon.final_def cp csj_skip i_skip_all_skip pair_j by blast
then show ?thesis
by (metis Suc_leI before_k_all_evnt le_eq_less_or_eq)
qed
then have False
using pair_j csj_skip by (metis a00 a4 cp image_eqI l_f last_not_F)
then show ?thesis
by metis
qed
then have False using a4 l_f by auto
}
then have "(snd(l!k), snd(l!(Suc k))) ∈ G"
using p1 or by fastforce
} thus ?thesis by (simp add: c_prod cp) qed
have concr:"(final (last l) ⟶
((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a)))"
proof-
{
assume valid:"final (last l)"
have "∃k. k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k))) ∧ final (l!(Suc k))"
proof -
have len_l:"length l > 0" using cp using cptn.simps by blast
then obtain a1 l1 where l:"l=a1#l1" by (metis SmallStepCon.nth_tl length_greater_0_conv)
have last_l:"last l = l!(length l-1)"
using last_length [of a1 l1] l by fastforce
have final_0:"¬final(l!0)" using cp unfolding final_def by auto
have "0≤ (length l-1)" using len_l last_l by auto
moreover have "(length l-1) < length l" using len_l by auto
moreover have "final (l!(length l-1))" using valid last_l by auto
moreover have "fst (l!0) = (Guard f g c1)" using cp by auto
ultimately show ?thesis
using cp final_exist_component_tran_final env_tran_right final_0
by blast
qed
then obtain k where a21: "k≥0 ∧ k<((length l) - 1) ∧ (Γ⊢⇩c(l!k) → (l!(Suc k))) ∧ final (l!(Suc k))"
by auto
then have a00:"Suc k<length l" by fastforce
then obtain j where before_k_all_evnt:"j≤k ∧ (Γ⊢⇩c(l!j) → (l!(Suc j))) ∧ (∀k < j. (Γ⊢⇩c(l!k) →⇩e (l!(Suc k))))"
using a00 a21 exist_first_comp_tran cp by blast
then obtain cj sj csj ssj where pair_j:"(Γ⊢⇩c(cj,sj) → (csj,ssj)) ∧ cj = fst (l!j) ∧ sj = snd (l!j) ∧ csj = fst (l!(Suc j)) ∧ ssj = snd(l!(Suc j))"
by fastforce
have "((fst (last l) = Skip ∧
snd (last l) ∈ Normal ` q)) ∨
(fst (last l) = Throw ∧
snd (last l) ∈ Normal ` (a))"
proof -
have j_length:"Suc j < length l" using a00 before_k_all_evnt by fastforce
have k_basic:"cj =(Guard f g c1) ∧ sj ∈ Normal ` (p)"
using pair_j before_k_all_evnt cp env_tran_right a2 assum a00 stability[of p R l 0 j j Γ]
by force
then obtain s' where ss:"sj = Normal s' ∧ s'∈ (p)" by auto
have or:"s'∈ (g ∪ (-g))" by fastforce
{assume "s' ∈ g"
then have k_basic:"cj =(Guard f g c1) ∧ sj ∈ Normal ` (p ∩ g)"
using ss k_basic by fastforce
then have ss: "sj = Normal s' ∧ s'∈ (p ∩ g)"
using ss by fastforce
then have ssj_normal_s:"ssj = Normal s'"
using before_k_all_evnt k_basic pair_j a1
by (metis (no_types, lifting) Pair_inject IntD2 stepc_Normal_elim_cases(2))
have cj:"csj=c1" using k_basic pair_j ss a0
by (metis (no_types, lifting) fst_conv IntD2 stepc_Normal_elim_cases(2))
moreover have p1:"s'∈(p ∩ g)" using ss by blast
moreover then have "cpn n Γ csj ssj ∩ assum((p ∩ g), R) ⊆ comm(G, (q,a)) F"
using a1 com_validityn_def cj by blast
moreover then have "l!(Suc j) = (csj, Normal s')"
using before_k_all_evnt pair_j cj ssj_normal_s
by fastforce
ultimately have drop_comm:"((Γ, drop (Suc j) l))∈ comm(G, (q,a)) F"
using j_length a10 a11 Γ1 ssj_normal_s cpn_assum_induct
by fastforce
then have ?thesis
using j_length l_f drop_comm a10' Γ1
cptn_comm_induct[of Γ l "(Guard f g c1)" s _ "Suc j" G q a F "Suc j"] valid
by blast
}note left=this
{
assume "s'∉g"
then have "csj= Skip ∧ ssj=Fault f" using k_basic ss pair_j
by (meson Pair_inject stepc_Normal_elim_cases(2))
then have "snd (last l) = Fault f" using pair_j
by (metis a4 cp imageI j_length l_f last_not_F)
then have False using a4 l_f by auto
}
thus ?thesis using or left by auto qed
} thus ?thesis by auto
qed
note res = conjI [OF concl concr]}
thus ?thesis using c_prod unfolding comm_def by force qed
} thus ?thesis by auto qed
} thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma WhileNone:
"Γ⊢⇩c (While b c1, s1) → (LanguageCon.com.Skip, t1) ⟹
(n,Γ, (Skip, t1) # xsa) ∈ cptn_mod_nest_call ⟹
Γ ⊨n⇘/F⇙ c1 sat [p ∩ b,R, G, p,a] ⟹
Sta p R ⟹
Sta (p ∩ (-b)) R ⟹
Sta a R ⟹
(∀s. (Normal s, Normal s) ∈ G) ⟹
(Γ, (While b c1, s1) # (LanguageCon.com.Skip, t1) # xsa) ∈ assum (p, R) ⟹
(Γ, (While b c1, s1) # (LanguageCon.com.Skip, t1) # xsa) ∈ comm (G,(p ∩ (-b)),a) F"
proof -
assume a0:"Γ⊢⇩c (While b c1, s1) → (LanguageCon.com.Skip, t1)" and
a1:"(n,Γ, (Skip, t1) # xsa) ∈ cptn_mod_nest_call" and
a2:" Γ ⊨n⇘/F⇙ c1 sat [p ∩ b,R, G, p,a]" and
a3:"Sta p R" and
a4:"Sta (p ∩ (-b)) R" and
a5:"Sta a R" and
a6:"∀s. (Normal s, Normal s) ∈ G" and
a7:"(Γ, (While b c1, s1) # (LanguageCon.com.Skip, t1) # xsa) ∈ assum (p, R)"
obtain s1' where s1N:"s1=Normal s1' ∧ s1'∈p" using a7 unfolding assum_def by fastforce
then have s1_t1:"s1'∉ b ∧ t1=s1" using a0
using LanguageCon.com.distinct(5) prod.inject
by (fastforce elim:stepc_Normal_elim_cases(7))
then have t1_Normal_post:"t1∈ Normal ` (p ∩ (-b))"
using s1N by fastforce
also have "(Γ, (While b c1, s1) # (LanguageCon.com.Skip, t1) # xsa)∈cptn"
using a1 a0 cptn.simps
using cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod by fastforce
ultimately have assum_skip:
"(Γ,(LanguageCon.com.Skip, t1) # xsa) ∈ assum (( p ∩ (-b)), R)"
using a1 a7 tl_of_assum_in_assum1 t1_Normal_post by fastforce
have skip_comm:"(Γ,(LanguageCon.com.Skip, t1) # xsa) ∈
comm (G,(( p ∩ (-b)),a)) F"
proof-
obtain Θ where "(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p , R, G, q,a])" by auto
moreover have "Γ,Θ ⊨n⇘/F⇙ Skip sat [( p ∩ (-b)), R, G, ( p ∩ (-b)),a]"
using Skip_sound[of "(p ∩ - b)"] a4 a6 by blast
ultimately show ?thesis
using assum_skip a1 unfolding com_cvalidityn_def com_validityn_def cpn_def
by fastforce
qed
have G_ref:"(Normal s1', Normal s1')∈G" using a6 by fastforce
thus ?thesis using skip_comm ctran_in_comm[of s1'] s1N s1_t1 by blast
qed
lemma while1:
"(n,Γ, ((c, Normal s1) # xs1)) ∈ cptn_mod_nest_call ⟹
s1 ∈ b ⟹
xsa = map (lift (While b c)) xs1 ⟹
Γ ⊨n⇘/F⇙ c sat [p ∩ b,R, G, p,a] ⟹
(Γ, (While b c, Normal s1) #
(Seq c (LanguageCon.com.While b c), Normal s1) # xsa)
∈ assum (p, R) ⟹
∀s. (Normal s, Normal s) ∈ G ⟹
(Γ, (LanguageCon.com.While b c, Normal s1) #
(LanguageCon.com.Seq c (LanguageCon.com.While b c), Normal s1) # xsa)
∈ comm (G, p∩(-b), a) F"
proof -
assume
a0:"(n,Γ, ((c, Normal s1) # xs1)) ∈ cptn_mod_nest_call" and
a1:"s1 ∈ b" and
a2:"xsa = map (lift (While b c)) xs1" and
a3:"Γ ⊨n⇘/F⇙ c sat [p ∩ b,R, G, p,a]" and
a4:"(Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) # xsa)
∈ assum (p, R) " and
a5:"∀s. (Normal s,Normal s) ∈ G"
have seq_map:"(Seq c (While b c), Normal s1) # xsa=
map (lift (While b c)) ((c,Normal s1)#xs1)"
using a2 unfolding lift_def by fastforce
have step:"Γ⊢⇩c(While b c,Normal s1) → (Seq c (While b c),Normal s1)" using a1
WhileTruec by fastforce
have s1_normal:"s1 ∈ p ∧ s1 ∈ b " using a4 a1 unfolding assum_def by fastforce
then have G_ref:"(Normal s1, Normal s1) ∈ G" using a5 by fastforce
have s1_collect_p: "Normal s1∈ Normal ` (p ∩ b)" using s1_normal by fastforce
have "(Γ, map (lift (While b c)) ((c,Normal s1)#xs1))∈cptn"
using a2 cptn_eq_cptn_mod_nest lift_is_cptn a0 by blast
then have cptn_seq:"(Γ,(Seq c (While b c), Normal s1) # xsa) ∈cptn"
using seq_map by auto
then have "(Γ, (While b c, Normal s1) # (Seq c (While b c), Normal s1) # xsa) ∈ cptn"
using step by (simp add: cptn.CptnComp)
then have assum_seq:"(Γ,(Seq c (While b c), Normal s1) # xsa)∈assum (p, R)"
using a4 tl_of_assum_in_assum1 s1_collect_p by fastforce
have cp_c:"(Γ, ((c, Normal s1) # xs1)) ∈ (cpn n Γ c (Normal s1))"
using a0 unfolding cpn_def by fastforce
then have cp_c':"(Γ, ((c, Normal s1) # xs1)) ∈ (cp Γ c (Normal s1))"
unfolding cp_def cpn_def using cptn_eq_cptn_mod_nest by fastforce
also have cp_seq:"(Γ, (Seq c (While b c), Normal s1) # xsa) ∈ (cp Γ (Seq c (While b c)) (Normal s1))"
using cptn_seq unfolding cp_def by fastforce
ultimately have "(Γ, ((c, Normal s1) # xs1)) ∈ assum(p,R)"
using assum_map assum_seq seq_map by fastforce
then have "(Γ, ((c, Normal s1) # xs1)) ∈ assum((p ∩ b),R)"
unfolding assum_def using s1_collect_p by fastforce
then have "(Γ, ((c, Normal s1) # xs1)) ∈ comm(G,(p,a)) F"
using a3 cp_c unfolding com_validityn_def by fastforce
then have "(Γ, (Seq c (While b c), Normal s1) # xsa) ∈ comm(G,(p,a)) F"
using cp_seq cp_c' comm_map seq_map by fastforce
then have "(Γ, (While b c, Normal s1) # (Seq c (While b c), Normal s1) # xsa) ∈ comm(G,(p,a)) F"
using G_ref ctran_in_comm by fastforce
also have "¬ final (last ((While b c, Normal s1) # (Seq c (While b c), Normal s1) # xsa))"
using seq_map unfolding final_def lift_def by (simp add: case_prod_beta' last_map)
ultimately show ?thesis using not_final_in_comm[of Γ] by blast
qed
lemma while2:
" (n,Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) # xsa) ∈cptn_mod_nest_call ⟹
(n, Γ, (c, Normal s1) # xs1) ∈ cptn_mod_nest_call ⟹
fst (last ((c, Normal s1) # xs1)) = LanguageCon.com.Skip ⟹
s1 ∈ b ⟹
xsa = map (lift (While b c)) xs1 @
(While b c, snd (last ((c, Normal s1) # xs1))) # ys ⟹
(n, Γ, (While b c, snd (last ((c, Normal s1) # xs1))) # ys)
∈ cptn_mod_nest_call ⟹
(Γ ⊨n⇘/F⇙ c sat [p ∩ b, R, G, p,a] ⟹
(Γ, (While b c, snd (last ((c, Normal s1) # xs1))) # ys)
∈ assum (p, R) ⟹
(Γ, (While b c, snd (last ((c, Normal s1) # xs1))) # ys)
∈ comm (G, p ∩ (-b), a) F) ⟹
Γ ⊨n⇘/F⇙ c sat [ p ∩ b, R, G, p,a] ⟹
(Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) # xsa)
∈ assum (p, R) ⟹
∀s. (Normal s,Normal s) ∈ G ⟹
(Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) # xsa)
∈ comm (G,( p ∩ (-b), a)) F"
proof -
assume a00:"(n, Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) # xsa) ∈cptn_mod_nest_call" and
a0:"(n,Γ, (c, Normal s1) # xs1) ∈ cptn_mod_nest_call" and
a1:" fst (last ((c, Normal s1) # xs1)) = LanguageCon.com.Skip" and
a2:"s1 ∈ b" and
a3:"xsa = map (lift (While b c)) xs1 @
(While b c, snd (last ((c, Normal s1) # xs1))) # ys" and
a4:"(n,Γ, (While b c, snd (last ((c, Normal s1) # xs1))) # ys)
∈ cptn_mod_nest_call" and
a5:"Γ ⊨n⇘/F⇙ c sat [p ∩ b, R, G, p,a]" and
a6:"(Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) # xsa)
∈ assum (p, R)" and
a7:"(Γ ⊨n⇘/F⇙ c sat [p ∩ b, R, G, p,a] ⟹
(Γ, (While b c, snd (last ((c, Normal s1) # xs1))) # ys)
∈ assum (p, R) ⟹
(Γ, (While b c, snd (last ((c, Normal s1) # xs1))) # ys)
∈ comm (G,p ∩ (-b), a) F)" and
a8:"∀s. (Normal s, Normal s) ∈ G"
let ?l= "(While b c, Normal s1) #
(Seq c (While b c), Normal s1) # xsa"
let ?sub_l="((While b c, Normal s1) #
(Seq c (While b c), Normal s1) #
map (lift (While b c)) xs1)"
{
assume final_not_fault:"snd (last ?l) ∉ Fault ` F"
have a0':"(Γ, (c, Normal s1) # xs1) ∈ cptn"
using cptn_eq_cptn_mod_nest using a0 by auto
have a4:"(Γ, (While b c, snd (last ((c, Normal s1) # xs1))) # ys) ∈ cptn"
using cptn_eq_cptn_mod_nest using a4 by blast
have seq_map:"(Seq c (While b c), Normal s1) # map (lift (While b c)) xs1=
map (lift (While b c)) ((c,Normal s1)#xs1)"
using a2 unfolding lift_def by fastforce
have step:"Γ⊢⇩c(While b c,Normal s1) → (Seq c (While b c),Normal s1)" using a2
WhileTruec by fastforce
have s1_normal:"s1∈p ∧ s1 ∈ b " using a6 a2 unfolding assum_def by fastforce
have G_ref:"(Normal s1, Normal s1)∈G "
using a8 by blast
have s1_collect_p: "Normal s1∈ Normal ` (p ∩ b)" using s1_normal by fastforce
have "(Γ, map (lift (While b c)) ((c,Normal s1)#xs1))∈cptn"
using a2 cptn_eq_cptn_mod lift_is_cptn a0' by fastforce
then have cptn_seq:"(Γ,(Seq c (While b c), Normal s1) # map (lift (While b c)) xs1) ∈cptn"
using seq_map by auto
then have "(Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) #
map (lift (While b c)) xs1) ∈ cptn"
using step by (simp add: cptn.CptnComp)
also have "(Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) #
map (lift (While b c)) xs1)
∈ assum (p, R)"
using a6 a3 sub_assum by force
ultimately have assum_seq:"(Γ,(Seq c (While b c), Normal s1) #
map (lift (While b c)) xs1) ∈ assum (p, R)"
using a6 tl_of_assum_in_assum1 s1_collect_p
tl_of_assum_in_assum by fastforce
have cpn_c:"(Γ, ((c, Normal s1) # xs1)) ∈ (cpn n Γ c (Normal s1))"
using a0 unfolding cpn_def by fastforce
have cp_c:"(Γ, ((c, Normal s1) # xs1)) ∈ (cp Γ c (Normal s1))"
using a0' unfolding cp_def by fastforce
also have cp_seq:"(Γ, (Seq c (While b c), Normal s1) # map (lift (While b c)) xs1) ∈ (cp Γ (Seq c (While b c)) (Normal s1))"
using cptn_seq unfolding cp_def by fastforce
ultimately have "(Γ, ((c, Normal s1) # xs1)) ∈ assum(p,R) "
using assum_map assum_seq seq_map by fastforce
then have "(Γ, ((c, Normal s1) # xs1)) ∈ assum((p ∩ b),R) "
unfolding assum_def using s1_collect_p by fastforce
then have c_comm:"(Γ, ((c, Normal s1) # xs1)) ∈ comm(G,(p,a)) F"
using a5 cpn_c unfolding com_validityn_def by fastforce
then have "(Γ, (Seq c (While b c), Normal s1) # map (lift (While b c)) xs1) ∈ comm(G,(p,a)) F"
using cp_seq cp_c comm_map seq_map by fastforce
then have comm_while:"(Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) #
map (lift (While b c)) xs1) ∈ comm(G,(p,a)) F"
using G_ref ctran_in_comm by fastforce
have final_last_c:"final (last ((c,Normal s1)#xs1))"
using a1 a3 unfolding final_def by fastforce
have last_while1:"snd (last (map (lift (While b c)) ((c,Normal s1)#xs1))) = snd (last ((c, Normal s1) # xs1))"
unfolding lift_def by (simp add: case_prod_beta' last_map)
have last_while2:"(last (map (lift (While b c)) ((c,Normal s1)#xs1))) =
last ((While b c, Normal s1) # (Seq c (While b c), Normal s1) # map (lift (While b c)) xs1)"
using seq_map by fastforce
have not_fault_final_last_c:
"snd (last ( (c,Normal s1)#xs1)) ∉ Fault ` F"
proof -
have "(length ?sub_l) - 1 < length ?l"
using a3 by fastforce
then have "snd (?l!((length ?sub_l) - 1))∉ Fault ` F"
using final_not_fault a3 a00 last_not_F[of Γ ?l F]
cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod by blast
thus ?thesis using last_while2 last_while1 seq_map
by (metis (no_types) Cons_lift_append a3 diff_Suc_1 last_length length_Cons lessI nth_Cons_Suc nth_append)
qed
then have last_c_normal:"snd (last ( (c,Normal s1)#xs1)) ∈ Normal ` (p)"
using c_comm a1 unfolding comm_def final_def by fastforce
then obtain sl where sl:"snd (last ( (c,Normal s1)#xs1)) = Normal sl" by fastforce
have while_comm:"(Γ, (While b c, snd (last ((c, Normal s1) # xs1))) # ys) ∈ comm(G,(p∩(-b),a)) F"
proof -
have assum_while: "(Γ, (While b c, snd (last ((c, Normal s1) # xs1))) # ys)
∈ assum (p, R)"
using last_c_normal a3 a6 sub_assum_r[of Γ ?sub_l "(While b c, snd (last ((c, Normal s1) # xs1)))" ys p R p]
by fastforce
thus ?thesis using a5 a7 by fastforce
qed
have "sl∈p" using last_c_normal sl by fastforce
then have G1_ref:"(Normal sl, Normal sl)∈G" using a8 by auto
also have "snd (last ?sub_l) = Normal sl"
using last_while1 last_while2 sl by fastforce
ultimately have ?thesis
using cptn_eq_cptn_mod_nest a00 a3 sl while_comm comm_union[OF comm_while]
by fastforce
} note p1 =this
{
assume final_not_fault:"¬ (snd (last ?l) ∉ Fault ` F)"
then have ?thesis unfolding comm_def by fastforce
} thus ?thesis using p1 by fastforce
qed
lemma while3:
"(n, Γ, (c, Normal s1) # xs1) ∈ cptn_mod_nest_call ⟹
fst (last ((c, Normal s1) # xs1)) = Throw ⟹
s1 ∈ b ⟹
snd (last ((c, Normal s1) # xs1)) = Normal sl ⟹
(n,Γ, (Throw, Normal sl) # ys) ∈ cptn_mod_nest_call ⟹
Γ ⊨n⇘/F⇙ c sat [p ∩ b,R, G, p,a] ⟹
(Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) #
(map (lift (While b c)) xs1 @
(Throw, Normal sl) # ys))
∈ assum (p, R) ⟹
Sta p R ⟹
Sta a R ⟹ ∀s. (Normal s, Normal s) ∈ G ⟹
(Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) #
((map (lift (While b c)) xs1 @
(Throw, Normal sl) # ys))) ∈ comm (G, p∩ (-b), a) F
"
proof -
assume a0:"(n,Γ, (c, Normal s1) # xs1) ∈ cptn_mod_nest_call" and
a1:"fst (last ((c, Normal s1) # xs1)) = Throw" and
a2:"s1 ∈ b" and
a3:"snd (last ((c, Normal s1) # xs1)) = Normal sl" and
a4:"(n,Γ, (Throw, Normal sl) # ys) ∈ cptn_mod_nest_call" and
a5:"Γ ⊨n⇘/F⇙ c sat [p ∩ b, R, G, p,a]" and
a6:"(Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) #
(map (lift (While b c)) xs1 @
(Throw, Normal sl) # ys))
∈ assum (p, R)" and
a7: "Sta p R" and
a8: "Sta a R" and
a10:"∀s. (Normal s,Normal s) ∈ G"
have seq_map:"(Seq c (While b c), Normal s1) # map (lift (While b c)) xs1=
map (lift (While b c)) ((c,Normal s1)#xs1)"
using a2 unfolding lift_def by fastforce
have step:"Γ⊢⇩c(While b c,Normal s1) → (Seq c (While b c),Normal s1)" using a2
WhileTruec by fastforce
have s1_normal:"s1∈p ∧ s1 ∈ b " using a6 a2 unfolding assum_def by fastforce
then have G_ref:"(Normal s1, Normal s1)∈G" using a10 by auto
have s1_collect_p: "Normal s1∈ Normal ` (p ∩ b)" using s1_normal by fastforce
have "(n, Γ, map (lift (While b c)) ((c,Normal s1)#xs1))∈cptn_mod_nest_call"
using a2 lift_is_cptn a0
by (metis cptn_mod_nest_call.CptnModNestSeq1 seq_map)
then have cptn_seq:"(n,Γ,(Seq c (While b c), Normal s1) # map (lift (While b c)) xs1) ∈cptn_mod_nest_call"
using seq_map by auto
then have cptn:"(n,Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) #
map (lift (While b c)) xs1) ∈ cptn_mod_nest_call"
by (meson a0 a2 cptn_mod_nest_call.CptnModNestWhile1)
also have "(Γ, (LanguageCon.com.While b c, Normal s1) #
(LanguageCon.com.Seq c (LanguageCon.com.While b c), Normal s1) #
map (lift (LanguageCon.com.While b c)) xs1)
∈ assum (p, R) "
using a6 sub_assum by force
ultimately have assum_seq:"(Γ,(Seq c (While b c), Normal s1) #
map (lift (While b c)) xs1) ∈ assum (p, R)"
using a6 tl_of_assum_in_assum1 s1_collect_p
tl_of_assum_in_assum cptn_eq_cptn_mod_nest by fast
have cpn_c:"(Γ, ((c, Normal s1) # xs1)) ∈ (cpn n Γ c (Normal s1))"
using a0 unfolding cpn_def by fastforce
then have cp_c:"(Γ, ((c, Normal s1) # xs1)) ∈ (cp Γ c (Normal s1))"
unfolding cp_def cpn_def using cptn_eq_cptn_mod_nest by auto
moreover have cp_seq:"(Γ, (Seq c (While b c), Normal s1) # map (lift (While b c)) xs1) ∈ (cpn n Γ (Seq c (While b c)) (Normal s1))"
using cptn_seq unfolding cpn_def by fastforce
then have cp_seq':"(Γ, (Seq c (While b c), Normal s1) # map (lift (While b c)) xs1) ∈ (cp Γ (Seq c (While b c)) (Normal s1))"
unfolding cp_def cpn_def using cptn_eq_cptn_mod_nest by auto
ultimately have "(Γ, ((c, Normal s1) # xs1)) ∈ assum(p,R)"
using assum_map assum_seq seq_map by fastforce
then have "(Γ, ((c, Normal s1) # xs1)) ∈ assum((p ∩ b),R)"
unfolding assum_def using s1_collect_p by fastforce
then have c_comm:"(Γ, ((c, Normal s1) # xs1)) ∈ comm(G,(p,a)) F"
using a5 cpn_c unfolding com_validityn_def by fastforce
then have "(Γ, (Seq c (While b c), Normal s1) # map (lift (While b c)) xs1) ∈ comm(G,(p,a)) F"
using cp_seq' cp_c comm_map seq_map by fastforce
then have comm_while:"(Γ, (While b c, Normal s1) # (Seq c (While b c), Normal s1) # map (lift (While b c)) xs1) ∈ comm(G,(p,a)) F"
using G_ref ctran_in_comm by fastforce
have final_last_c:"final (last ((c,Normal s1)#xs1))"
using a1 a3 unfolding final_def by fastforce
have not_fault_final_last_c:
"snd (last ( (c,Normal s1)#xs1)) ∉ Fault ` F"
using a3 by fastforce
then have sl_a:"Normal sl ∈ Normal ` (a)"
using final_last_c a1 c_comm unfolding comm_def
using a3 comm_dest2
by auto
have last_while1:"snd (last (map (lift (While b c)) ((c,Normal s1)#xs1))) = snd (last ((c, Normal s1) # xs1))"
unfolding lift_def by (simp add: case_prod_beta' last_map)
have last_while2:"(last (map (lift (While b c)) ((c,Normal s1)#xs1))) =
last ((While b c, Normal s1) # (Seq c (While b c), Normal s1) # map (lift (While b c)) xs1)"
using seq_map by fastforce
have throw_comm:"(Γ, (Throw, Normal sl) # ys) ∈ comm(G,(p∩(-b),a)) F"
proof -
have assum_throw: "(Γ, (Throw, Normal sl) # ys) ∈ assum (a,R)"
using sl_a a6 sub_assum_r[of _ "(LanguageCon.com.While b c, Normal s1) #
(LanguageCon.com.Seq c (LanguageCon.com.While b c), Normal s1) #
map (lift (LanguageCon.com.While b c)) xs1" "(Throw, Normal sl)" ]
by fastforce
also have "(Γ,(Throw, Normal sl) # ys) ∈ cpn n Γ Throw (Normal sl)"
unfolding cpn_def using a4 by fastforce
ultimately show ?thesis using Throw_sound[of a R G Γ] a10 a8
unfolding com_cvalidityn_def com_validityn_def by fast
qed
have p1:"(LanguageCon.com.While b c, Normal s1) #
(LanguageCon.com.Seq c (LanguageCon.com.While b c), Normal s1) #
map (lift (LanguageCon.com.While b c)) xs1 ≠
[] ∧
(LanguageCon.com.Throw, Normal sl) # ys ≠ []" by auto
have "sl ∈ a" using sl_a by fastforce
then have G1_ref:"(Normal sl, Normal sl) ∈ G" using a10 by auto
moreover have "snd (last ((While b c, Normal s1) #
(Seq c (While b c), Normal s1) #
map (lift (While b c)) xs1)) = Normal sl"
using last_while1 last_while2 a3 by fastforce
moreover have "snd (((LanguageCon.com.Throw, Normal sl) # ys) ! 0) = Normal sl"
by (metis nth_Cons_0 snd_conv)
ultimately have G:"(snd (last ((While b c, Normal s1) #
(Seq c (While b c), Normal s1) #
map (lift (While b c)) xs1)),
snd (((LanguageCon.com.Throw, Normal sl) # ys) ! 0)) ∈ G" by auto
have cptn:"(Γ, ((LanguageCon.com.While b c, Normal s1) #
(LanguageCon.com.Seq c (LanguageCon.com.While b c), Normal s1) #
map (lift (LanguageCon.com.While b c)) xs1) @
(LanguageCon.com.Throw, Normal sl) # ys)
∈ cptn" using cptn a4 a0 a1 a3 a4 cptn_eq_cptn_mod_set cptn_mod.CptnModWhile3 s1_normal
cptn_eq_cptn_mod_nest by (metis append_Cons)
show ?thesis using a0 comm_union[OF comm_while throw_comm p1 G cptn] by auto
qed
inductive_cases stepc_elim_cases_while_throw [cases set]:
"Γ⊢⇩c(While b c, s) → (Throw, t)"
lemma WhileSound_aux:
"Γ ⊨n⇘/F⇙ c1 sat [p ∩ b, R, G, p,a] ⟹
Sta p R ⟹
Sta (p ∩ (-b)) R ⟹
Sta a R ⟹
(n, Γ,x)∈ cptn_mod_nest_call ⟹
∀s. (Normal s, Normal s) ∈ G ⟹
∀s xs. x = ((While b c1),s)#xs ⟶
(Γ,x)∈assum(p,R) ⟶
(Γ,x) ∈ comm (G,(( p ∩ (-b)),a)) F"
proof -
assume a0: "Γ ⊨n⇘/F⇙ c1 sat [p ∩ b, R, G, p,a] " and
a1: "Sta p R" and
a2: "Sta (p ∩ (-b)) R" and
a3: "Sta a R" and
a4: "(n,Γ,x)∈ cptn_mod_nest_call" and
a5: "∀s. (Normal s, Normal s) ∈ G"
{fix xs s
assume while_xs:"x=((While b c1),s)#xs" and
x_assum:"(Γ,x)∈assum(p,R)"
have "(Γ,x) ∈ comm (G,(( p ∩ (-b)),a)) F"
using a4 a0 while_xs x_assum
proof (induct arbitrary: xs s c1 rule:cptn_mod_nest_call.induct)
case (CptnModNestOne Γ C s1) thus ?case
using CptnModOne unfolding comm_def final_def
by auto
next
case (CptnModNestEnv Γ C s1 t1 n xsa)
then have c_while:"C = While b c1" by fastforce
have "(Γ, (C, t1) # xsa) ∈ assum (p, R) ⟶
(Γ, (C, t1) # xsa) ∈ comm (G, p ∩ (-b), a) F"
using CptnModNestEnv by fastforce
moreover have"(n,Γ,(C, s1)#(C, t1) # xsa) ∈ cptn_mod_nest_call"
using CptnModNestEnv(1,2) CptnModNestEnv.hyps(1) CptnModNestEnv.hyps(2)
using cptn_mod_nest_call.CptnModNestEnv by blast
then have cptn_mod:"(Γ,(C, s1)#(C, t1) # xsa) ∈ cptn"
using cptn_eq_cptn_mod_nest by blast
then have "(Γ, (C, t1) # xsa) ∈ assum (p, R)"
using tl_of_assum_in_assum CptnModNestEnv(6) a1 a2 a3 a4 a5
by blast
ultimately have "(Γ, (C, t1) # xsa) ∈ comm (G, p ∩ (-b), a) F"
by auto
also have " ¬ (Γ⊢⇩c((C,s1)) → ((C,t1)))"
by (simp add: mod_env_not_component)
ultimately show ?case
using cptn_mod etran_in_comm by blast
next
case (CptnModNestSkip Γ C s1 t1 n xsa)
then have "C=While b c1" by auto
also have "(n,Γ, (LanguageCon.com.Skip, t1) # xsa) ∈ cptn_mod_nest_call"
using cptn_eq_cptn_mod_set CptnModNestSkip(4) by fastforce
thus ?case using WhileNone CptnModNestSkip a1 a2 a3 a4 a5 by blast
next
case (CptnModNestThrow Γ C s1 t1 n xsa)
then have "C = While b c1" by auto
thus ?case using stepc_elim_cases_while_throw CptnModNestThrow(1)
by blast
next
case (CptnModNestWhile1 n Γ c s1 xs1 b1 xsa zs)
then have "b=b1 ∧ c=c1 ∧ s=Normal s1" by auto
thus ?case
using a4 a5 CptnModNestWhile1 while1[of n Γ] by blast
next
case (CptnModNestWhile2 n Γ c s1 xs1 b1 xsa ys zs)
then have a00: "(n,Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) # xsa)∈cptn_mod_nest_call"
using cptn_mod_nest_call.CptnModNestWhile2 by fast
then have eqs:"b=b1 ∧ c=c1 ∧ s=Normal s1"using CptnModNestWhile2 by auto
thus ?case using a00 a4 a5 CptnModNestWhile2 while2[of n Γ b c s1 xsa xs1 ys F p R G a]
by blast
next
case (CptnModNestWhile3 n Γ c s1 xs1 b1 sl ys zs)
then have eqs:"b=b1 ∧ c=c1 ∧ s=Normal s1" by auto
then have "(Γ, (While b c, Normal s1) #
(Seq c (While b c), Normal s1) #
((map (lift (While b c)) xs1 @
(Throw, Normal sl) # ys))) ∈ comm (G, p∩(-b), a) F"
using a1 a3 a4 a5 CptnModNestWhile3 while3[of n Γ c s1 xs1 b sl ys F p R G a]
by fastforce
thus ?case using eqs CptnModNestWhile3 by auto
qed (auto)
}
then show ?thesis by auto
qed
lemma While_sound:
"Γ,Θ ⊢⇘/F⇙ c1 sat [p ∩ b, R, G, p,a] ⟹
(∀n. Γ,Θ ⊨n⇘/F⇙ c1 sat [p ∩ b, R, G, p,a]) ⟹
Sta p R ⟹
Sta (p ∩ (-b)) R ⟹ Sta a R ⟹ ∀s. (Normal s, Normal s) ∈ G ⟹
Γ,Θ ⊨n⇘/F⇙ (While b c1) sat [p, R, G, p ∩ (-b),a]"
proof -
assume
a0:"Γ,Θ ⊢⇘/F⇙ c1 sat [p ∩ b, R, G, p,a]" and
a1:"∀n. Γ,Θ ⊨n ⇘/F⇙ c1 sat [p ∩ b, R, G, p,a]" and
a2: "Sta p R" and
a3: "Sta (p ∩ (-b)) R" and
a4: "Sta a R" and
a5: "∀s. (Normal s, Normal s) ∈ G"
{
fix s
assume all_call:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
then have a1:"(Γ ⊨n⇘/F⇙ c1 sat [p ∩ b, R, G, p,a])"
using a1 com_cvalidityn_def by fastforce
have "cpn n Γ (While b c1) s ∩ assum(p, R) ⊆ comm(G, (p ∩ (-b),a)) F"
proof-
{fix c
assume a10:"c ∈ cpn n Γ (While b c1) s" and a11:"c ∈ assum(p, R)"
then have a10': "c ∈ cp Γ (While b c1) s"
unfolding cp_def cpn_def using cptn_eq_cptn_mod_set cptn_mod_nest_cptn_mod by fastforce
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
have cp:"l!0=((While b c1),s) ∧ (Γ,l) ∈ cptn ∧ Γ=Γ1" using a10' cp_def c_prod by fastforce
have Γ1:"(Γ, l) = c" using c_prod cp by blast
obtain xs where "l=((While b c1),s)#xs" using cp
proof -
assume a1: "⋀xs. l = (LanguageCon.com.While b c1, s) # xs ⟹ thesis"
have "[] ≠ l"
using cp cptn.simps by auto
then show ?thesis
using a1 by (metis (full_types) SmallStepCon.nth_tl cp)
qed
moreover have "(n,Γ,l)∈cptn_mod_nest_call" using a10
using Γ1 cpn_def by fastforce
ultimately have "c ∈ comm(G, (p ∩ (-b),a)) F"
using a1 a2 a3 a4 WhileSound_aux a11 Γ1 a5
by blast
} thus ?thesis by auto qed
}
thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma Conseq_sound:
"(∀s∈ p.
∃p' R' G' q' a' I' Θ'.
s ∈ p' ∧
R ⊆ R' ∧
G' ⊆ G ∧
q' ⊆ q ∧
a' ⊆ a ∧ Θ' ⊆ Θ ∧
Γ,Θ' ⊢⇘/F⇙ P sat [p',R', G', q',a'] ∧
(∀n. Γ,Θ' ⊨n⇘/F⇙ P sat [p', R', G', q',a'])) ⟹
Γ,Θ ⊨n⇘/F⇙ P sat [p,R, G, q,a]"
proof -
assume
a0: "(∀s∈ p.
∃p' R' G' q' a' I' Θ'.
s ∈ p' ∧
R ⊆ R' ∧
G' ⊆ G ∧
q' ⊆ q ∧
a' ⊆ a ∧ Θ' ⊆ Θ ∧
Γ,Θ' ⊢⇘/F⇙ P sat [p',R', G', q',a'] ∧
(∀n. Γ,Θ' ⊨n⇘/F⇙ P sat [p', R', G', q',a']))"
{
fix s
assume all_call:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
have "cpn n Γ P s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ P s" and a11:"c ∈ assum(p, R)"
then have a10':"c∈cp Γ P s" unfolding cp_def cpn_def cptn_eq_cptn_mod_nest by auto
obtain Γ1 l where c_prod:"c=(Γ1,l)" by fastforce
have cp:"l!0=(P,s) ∧ (n,Γ,l) ∈ cptn_mod_nest_call ∧ Γ=Γ1" using a10 cpn_def c_prod by fastforce
have Γ1:"(Γ, l) = c" using c_prod cp by blast
obtain xs where "l=(P,s)#xs" using cp
proof -
assume a1: "⋀xs. l = (P, s) # xs ⟹ thesis"
have "[] ≠ l"
using cp cptn.simps
using CptnEmpty by force
then show ?thesis
using a1 by (metis (full_types) SmallStepCon.nth_tl cp)
qed
obtain ns where s:"(s = Normal ns)" using a10 a11 unfolding assum_def cpn_def by fastforce
then have "ns ∈ p" using a10 a11 unfolding assum_def cpn_def by fastforce
then have ns:"ns∈p" by auto
then have
"∀s. s ∈ p ⟶ (∃p' R' G' q' a' Θ'. (s∈p') ∧
R ⊆ R' ∧
G' ⊆ G ∧
q' ⊆ q ∧
a' ⊆ a ∧ Θ' ⊆ Θ ∧
(Γ,Θ' ⊢⇘/F⇙ P sat [p',R', G', q',a']) ∧
(∀n. Γ,Θ' ⊨n⇘/F⇙ P sat [p', R', G', q',a']))" using a0 by auto
then have
"ns ∈ p ⟶ (∃p' R' G' q' a' Θ'. (ns ∈ p' ) ∧
R ⊆ R' ∧
G' ⊆ G ∧
q' ⊆ q ∧
a' ⊆ a ∧ Θ' ⊆ Θ ∧
(Γ,Θ' ⊢⇘/F⇙ P sat [p',R', G', q',a']) ∧
(∀n. Γ,Θ' ⊨n⇘/F⇙ P sat [p', R', G', q',a']))" apply (rule allE) by auto
then obtain p' R' G' q' a' Θ' where
rels:
"ns ∈ p' ∧
R ⊆ R' ∧
G' ⊆ G ∧
q' ⊆ q ∧
a' ⊆ a ∧ Θ' ⊆ Θ ∧
(∀n. Γ,Θ' ⊨n⇘/F⇙ P sat [p', R', G', q',a'])" using ns by auto
then have "s ∈ Normal ` p'" using s by fastforce
then have "(Γ,l) ∈ assum(p', R')"
using a11 rels cp a11 c_prod assum_R_R'[of Γ l p R p' R']
by fastforce
then have "(Γ,l) ∈ comm(G',(q',a')) F"
using rels all_call a10 c_prod cp unfolding com_cvalidityn_def com_validityn_def
by blast
then have "(Γ,l) ∈ comm(G, (q,a)) F"
using c_prod cp comm_conseq[of Γ l G' q' a' F G q a] rels by fastforce
then have "c ∈ comm(G, (q,a)) F" using c_prod cp by fastforce
}
thus ?thesis unfolding comm_def by force qed
} thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma Conj_post_sound:
"Γ,Θ ⊢⇘/F⇙ P sat [p,R, G, q,a] ∧
(∀n. Γ,Θ ⊨n⇘/F⇙ P sat [p, R, G, q,a]) ⟹
Γ,Θ ⊢⇘/F⇙ P sat [p,R, G, q',a'] ∧
(∀n. Γ,Θ ⊨n ⇘/F⇙ P sat [p, R, G, q',a']) ⟹
Γ,Θ ⊨n⇘/F⇙ P sat [p,R, G, q ∩ q' ,a ∩ a']"
proof -
assume a0: "Γ,Θ ⊢⇘/F⇙ P sat [p,R, G, q,a] ∧
(∀n. Γ,Θ ⊨n⇘/F⇙ P sat [p, R, G, q,a])" and
a1: " Γ,Θ ⊢⇘/F⇙ P sat [p,R, G, q',a'] ∧
(∀n. Γ,Θ ⊨n ⇘/F⇙ P sat [p, R, G, q',a'])"
{
fix s
assume all_call:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
with a0 have a0:"cpn n Γ P s ∩ assum(p, R) ⊆ comm(G, (q,a)) F"
unfolding com_cvalidityn_def com_validityn_def by auto
with a1 all_call have a1:"cpn n Γ P s ∩ assum(p, R) ⊆ comm(G, (q',a')) F"
unfolding com_cvalidityn_def com_validityn_def by auto
have "cpn n Γ P s ∩ assum(p, R) ⊆ comm(G, (q∩q',a∩a')) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ P s" and a11:"c ∈ assum(p, R)"
then have "c ∈ comm(G,(q,a)) F ∧ c ∈ comm(G,(q',a')) F"
using a0 a1 by auto
then have "c∈comm(G, (q∩q',a∩a')) F"
unfolding comm_def by fastforce
}
thus ?thesis unfolding comm_def by force qed
} thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma x91:"sa≠{} ⟹ c∈comm(G, (⋂i∈sa. q i,a)) F = (∀i∈sa. c∈comm(G, q i,a) F)"
unfolding comm_def apply (auto simp add: Ball_def)
apply (frule spec , force)
by (frule spec, force)
lemma conj_inter_sound:
"sa ≠ {} ⟹
∀i∈sa. Γ,Θ ⊢⇘/F⇙ P sat [p, R, G, q i,a] ∧ (∀n. Γ,Θ ⊨n⇘/F⇙ P sat [p,R, G, q i,a]) ⟹
Γ,Θ ⊨n⇘/F⇙ P sat [p,R, G, ⋂i∈sa. q i,a]"
proof -
assume a0':"sa≠{}" and
a0: "∀i∈sa. Γ,Θ ⊢⇘/F⇙ P sat [p, R, G, q i,a] ∧
(∀n. Γ,Θ ⊨n⇘/F⇙ P sat [p,R, G, q i,a])"
{
fix s
assume all_call:"∀(c,p,R,G,q,a)∈ Θ. Γ ⊨n⇘/F⇙ (Call c) sat [p, R, G, q,a]"
with a0 have a0:"∀i∈sa. cpn n Γ P s ∩ assum(p, R) ⊆ comm(G, (q i,a)) F"
unfolding com_cvalidityn_def com_validityn_def by auto
have "cpn n Γ P s ∩ assum(p, R) ⊆ comm(G, (⋂i∈sa. q i,a)) F"
proof -
{
fix c
assume a10:"c ∈ cpn n Γ P s" and a11:"c ∈ assum(p, R)"
then have "(∀i∈sa. c∈comm(G, q i,a) F)"
using a0 by fastforce
then have "c∈comm(G, (⋂i∈sa. q i,a)) F" using x91[OF a0'] by blast
}
thus ?thesis unfolding comm_def by force qed
} thus ?thesis by (simp add: com_validityn_def[of Γ] com_cvalidityn_def)
qed
lemma localRG_sound: "Γ,Θ ⊢⇘/F⇙ c sat [p, R, G, q,a] ⟹ (⋀n. Γ,Θ ⊨n ⇘/F⇙ c sat [p, R, G, q,a])"
proof (induct rule:lrghoare.induct)
case Skip
thus ?case by (simp add: Skip_sound)
next
case Spec
thus ?case by (simp add: Spec_sound)
next
case Basic
thus ?case by (simp add: Basic_sound)
next
case Await
thus ?case by (simp add: Await_sound)
next
case Throw thus ?case by (simp add: Throw_sound)
next
case If thus ?case using If_sound by (simp add: If_sound)
next
case Asm thus ?case by (simp add: Asm_sound)
next
case CallRec thus ?case by (simp add: CallRec_sound)
next
case Call thus ?case using Call_sound by (simp add: Call_sound)
next
case Seq thus ?case by (simp add: Seq_sound)
next
case Catch thus ?case by (simp add: Catch_sound)
next
case DynCom thus ?case by (simp add: DynCom_sound)
next
case Guard thus ?case by (simp add: Guard_sound)
next
case Guarantee thus ?case by (simp add: Guarantee_sound)
next
case While thus ?case by (simp add: While_sound)
next
case (Conseq p R G q a Γ Θ F P) thus ?case
using Conseq_sound by simp
next
case (Conj_post Γ Θ F P p' R' G' q a q' a') thus ?case
using Conj_post_sound[of Γ Θ] by simp
next
case (Conj_Inter sa Γ Θ F P p' R' G' q a )
thus ?case using conj_inter_sound[of sa Γ Θ] by simp
qed
definition ParallelCom :: "('s,'p,'f,'e) rgformula list ⇒ ('s,'p,'f,'e) par_com"
where
"ParallelCom Ps ≡ map fst Ps"
lemma ParallelCom_Com:"i<length xs ⟹ (ParallelCom xs)!i = Com (xs!i)"
unfolding ParallelCom_def Com_def by fastforce
lemma etran_ctran_eq_p_normal_s: "Γ⊢⇩c s1 → s1' ⟹
Γ⊢⇩c s1 →⇩e s1' ⟹
fst s1 = fst s1' ∧ snd s1 = snd s1' ∧ (∃ns1. snd s1 = Normal ns1)"
proof -
assume a0: "Γ⊢⇩c s1 → s1'" and
a1: "Γ⊢⇩c s1 →⇩e s1'"
then obtain ps1 ss1 ps1' ss1' where prod:"s1 = (ps1,ss1) ∧ s1' = (ps1', ss1')"
by fastforce
then have "ps1=ps1'" using a1 etranE by fastforce
thus ?thesis using prod a0 by (simp add: mod_env_not_component)
qed
lemma step_e_step_c_eq:"⟦
(Γ,l) ∝ clist;
Suc m < length l;
i < length clist;
(fst (clist!i))⊢⇩c((snd (clist!i))!m) →⇩e ((snd (clist!i))!Suc m);
(fst (clist!i))⊢⇩c((snd (clist!i))!m) → ((snd (clist!i))!Suc m);
(∀l<length clist.
l≠i ⟶ (fst (clist!l))⊢⇩c (snd (clist!l))!m →⇩e ((snd (clist!l))!(Suc m)))
⟧ ⟹
l!m = l!(Suc m) ∧ (∃ns. snd (l!m) = Normal ns )"
proof -
assume a0:"(Γ,l) ∝ clist" and
a1:"Suc m < length l" and
a2:"i < length clist" and
a3:"(fst (clist!i))⊢⇩c((snd (clist!i))!m) →⇩e ((snd (clist!i))!Suc m)" and
a4:"(fst (clist!i))⊢⇩c((snd (clist!i))!m) → ((snd (clist!i))!Suc m)" and
a5:"(∀l<length clist.
l≠i ⟶ (fst (clist!l))⊢⇩c (snd (clist!l))!m →⇩e ((snd (clist!l))!(Suc m)))"
obtain fp fs sp ss
where prod_step: "
Γ⊢⇩c (fp, fs) → (sp,ss) ∧
fp = fst (((snd (clist!i))!m)) ∧ fs = snd (((snd (clist!i))!m)) ∧
sp = fst ((snd (clist!i))!(Suc m)) ∧ ss = snd((snd (clist!i))!(Suc m)) ∧
Γ = fst (clist!i)"
using a0 a2 a1 a4 unfolding conjoin_def same_functions_def by fastforce
have snd_lj:"(snd (l!m)) = snd ((snd (clist!i))!m)"
using a0 a1 a2 unfolding conjoin_def same_state_def
by fastforce
have fst_clist_Γ:"∀i<length clist. fst(clist!i) = Γ"
using a0 unfolding conjoin_def same_functions_def by fastforce
have all_env: "∀l<length clist.
(fst (clist!l))⊢⇩c (snd (clist!l))!m →⇩e ((snd (clist!l))!(Suc m))"
using a3 a5 a2 fst_clist_Γ by fastforce
then have allP:"∀l< length clist. fst ((snd (clist!l))!m) = fst ((snd (clist!l))!(Suc m))"
by (fastforce elim:etranE)
then have "fst (l!m) = (fst (l!(Suc m)))"
using a0 conjoin_same_program_i_j [of "(Γ,l)"] a1 by fastforce
also have snd_l_normal:"snd (l!m) = snd (l!(Suc m)) ∧ (∃ns. snd (l!m) = Normal ns )"
proof -
have "(snd (l!Suc m)) = snd ((snd (clist!i))!(Suc m))"
using a0 a1 a2 unfolding conjoin_def same_state_def
by fastforce
also have "fs = ss ∧
(∃ns. (snd ((snd (clist!i))!m) = Normal ns ))"
using a1 a2 all_env prod_step allP
by (metis step_change_p_or_eq_s)
ultimately show ?thesis using snd_lj prod_step a1 by fastforce
qed
ultimately show ?thesis using prod_eq_iff by blast
qed
lemma two':
"⟦ ∀i<length xs. R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i));
p ⊆ (⋂i<length xs. (Pre (xs ! i)));
∀i<length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs ! i) sat [Pre (xs!i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)];
length xs=length clist; (Γ,l) ∈ par_cp Γ (ParallelCom xs) s; (Γ,l)∈par_assum (p, R) ;
∀i<length clist. clist!i∈cp Γ (Com(xs!i)) s; (Γ,l) ∝ clist;(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a]);
snd (last l) ∉ Fault ` F⟧
⟹ ∀j i ns ns'. i<length clist ∧ Suc j<length l ⟶
Γ⊢⇩c((snd (clist!i))!j) →⇩e ((snd (clist!i))!Suc j) ⟶
(snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈ Rely(xs!i)"
proof -
assume a0:"∀i<length xs. R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i))" and
a1:"p ⊆ (⋂i<length xs. (Pre (xs ! i)))" and
a2:"∀i<length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs ! i) sat [Pre (xs!i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)]" and
a3: "length xs=length clist" and
a4: "(Γ,l) ∈ par_cp Γ (ParallelCom xs) s" and
a5: "(Γ,l)∈par_assum (p, R)" and
a6: "∀i<length clist. clist!i∈cp Γ (Com(xs!i)) s" and
a7: "(Γ,l) ∝ clist" and
a8: "(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a])" and
a9: "snd (last l) ∉ Fault ` F"
{
assume a10:"∃i j ns ns'.
i<length clist ∧ Suc j<length l ∧
Γ⊢⇩c((snd (clist!i))!j) →⇩e ((snd (clist!i))!Suc j) ∧
¬(snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈ Rely(xs!i)"
then obtain j where
a10:"∃i ns ns'.
i<length clist ∧ Suc j<length l ∧
Γ⊢⇩c((snd (clist!i))!j) →⇩e ((snd (clist!i))!Suc j) ∧
¬(snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈ Rely(xs!i)" by fastforce
let ?P = "λj. ∃i. i<length clist ∧ Suc j<length l ∧
Γ⊢⇩c((snd (clist!i))!j) →⇩e ((snd (clist!i))!Suc j) ∧
(¬ (snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈ Rely(xs!i))"
obtain m where fist_occ:"(?P m) ∧ (∀i<m. ¬ ?P i)" using exists_first_occ[of ?P j] a10 by blast
then have "?P m" by fastforce
then obtain i where
fst_occ:"i<length clist ∧ Suc m<length l ∧
Γ⊢⇩c((snd (clist!i))!m) →⇩e ((snd (clist!i))!Suc m) ∧
(¬ (snd((snd (clist!i))!m), snd((snd (clist!i))!Suc m)) ∈ Rely(xs!i))"
by fastforce
have notP:"(∀i<m. ¬ ?P i)" using fist_occ by blast
have fst_clist_Γ:"∀i<length clist. fst(clist!i) = Γ"
using a7 unfolding conjoin_def same_functions_def by fastforce
have compat:"(Γ⊢⇩p(l!m) → (l!(Suc m))) ∧
(∃i<length clist.
((fst (clist!i))⊢⇩c ((snd (clist!i))!m) → ((snd (clist!i))!(Suc m))) ∧
(∀l<length clist.
l≠i ⟶ (fst (clist!l))⊢⇩c (snd (clist!l))!m →⇩e ((snd (clist!l))!(Suc m)))) ∨
(Γ⊢⇩p(l!m) →⇩e (l!(Suc m)) ∧
(∀i<length clist. (fst (clist!i))⊢⇩c (snd (clist!i))!m →⇩e ((snd (clist!i))!(Suc m))))"
using a7 fst_occ unfolding conjoin_def compat_label_def by simp
{
assume a20: "(Γ⊢⇩p(l!m) →⇩e (l!(Suc m)) ∧
(∀i<length clist. (fst (clist!i))⊢⇩c (snd (clist!i))!m →⇩e ((snd (clist!i))!(Suc m))))"
then have "(snd (l!m),snd (l!(Suc m))) ∈ R"
using fst_occ a5 unfolding par_assum_def by fastforce
then have "(snd(l!m), snd(l!(Suc m))) ∈ Rely(xs!i)"
using fst_occ a3 a0 by fastforce
then have "(snd ((snd (clist!i))!m), snd ((snd (clist!i))!(Suc m)) ) ∈ Rely(xs!i)"
using a7 fst_occ unfolding conjoin_def same_state_def by fastforce
then have False using fst_occ by auto
}note l = this
{
assume a20:"(Γ⊢⇩p(l!m) → (l!(Suc m))) ∧
(∃i<length clist.
((fst (clist!i))⊢⇩c ((snd (clist!i))!m) → ((snd (clist!i))!(Suc m))) ∧
(∀l<length clist.
l≠i ⟶ (fst (clist!l))⊢⇩c (snd (clist!l))!m →⇩e ((snd (clist!l))!(Suc m))))"
then obtain i'
where i':"i'<length clist ∧
((fst (clist!i'))⊢⇩c ((snd (clist!i'))!m) → ((snd (clist!i'))!(Suc m))) ∧
(∀l<length clist.
l≠i' ⟶ (fst (clist!l))⊢⇩c (snd (clist!l))!m →⇩e ((snd (clist!l))!(Suc m)))"
by fastforce
then have eq_Γ:"Γ = fst (clist!i')" using a7 unfolding conjoin_def same_functions_def by fastforce
obtain fp fs sp ss
where prod_step: "
Γ⊢⇩c (fp, fs) → (sp,ss) ∧
fp = fst (((snd (clist!i'))!m)) ∧ fs = snd (((snd (clist!i'))!m)) ∧
sp = fst ((snd (clist!i'))!(Suc m)) ∧ ss = snd((snd (clist!i'))!(Suc m)) ∧
Γ = fst (clist!i') "
using a7 i' unfolding conjoin_def same_functions_def by fastforce
then have False
proof (cases "i = i'")
case True
then have "l!m = l!(Suc m) ∧ (∃ns. snd (l!m) = Normal ns )"
using step_e_step_c_eq[OF a7] i' fst_occ eq_Γ by blast
then have "Γ⊢⇩p(l!m) →⇩e (l!(Suc m))"
using step_pe.ParEnv by (metis prod.collapse)
then have "(snd (l ! m), snd (l ! Suc m)) ∈ R "
using fst_occ a5 unfolding par_assum_def by fastforce
then have "(snd (l ! m), snd (l ! Suc m)) ∈ Rely (xs ! i)"
using a0 a3 fst_occ by fastforce
then show ?thesis using fst_occ a7
unfolding conjoin_def same_state_def
by fastforce
next
case False note not_eq = this
thus ?thesis
proof (cases "fp = sp")
case True
then have "fs = ss ∧ (∃ns. fs=Normal ns)"
using prod_step prod_step
using step_change_p_or_eq_s by blast
then have "Γ⊢⇩c (fp, fs) →⇩e (sp, ss)" using True step_e.Env
by fastforce
then have "l!m = l!(Suc m) ∧ (∃ns. snd (l!m) = Normal ns )"
using step_e_step_c_eq[OF a7] prod_step i' fst_occ prod.collapse by auto
then have "Γ⊢⇩p(l!m) →⇩e (l!(Suc m))"
using step_pe.ParEnv by (metis prod.collapse)
then have "(snd (l ! m), snd (l ! Suc m)) ∈ R "
using fst_occ a5 unfolding par_assum_def by fastforce
then have "(snd (l ! m), snd (l ! Suc m)) ∈ Rely (xs ! i)"
using a0 a3 fst_occ by fastforce
then show ?thesis using fst_occ a7
unfolding conjoin_def same_state_def
by fastforce
next
case False
let ?l1 = "take (Suc (Suc m)) (snd(clist!i'))"
have clist_cptn:"(Γ,snd(clist!i')) ∈ cptn" using a6 i' unfolding cp_def by fastforce
have sucm_len:"Suc m < length (snd (clist!i'))"
using i' fst_occ a7 unfolding conjoin_def same_length_def by fastforce
then have summ_lentake:"Suc m < length ?l1" by fastforce
have len_l: "0<length l" using fst_occ by fastforce
also then have "snd (clist!i')≠[]"
using i' a7 unfolding conjoin_def same_length_def by fastforce
ultimately have "snd (last (snd (clist ! i'))) = snd (last l)"
using a7 i' conjoin_last_same_state by fastforce
then have last_i_notF:"snd (last (snd(clist!i'))) ∉ Fault ` F"
using a9 by auto
have "∀i<length (snd(clist!i')). snd (snd(clist!i') ! i) ∉ Fault ` F "
using last_not_F[OF clist_cptn last_i_notF] by auto
also have suc_m_i':"Suc m < length (snd (clist !i'))"
using fst_occ i' a7 unfolding conjoin_def same_length_def by fastforce
ultimately have last_take_not_f:"snd (last (take (Suc (Suc m)) (snd(clist!i')))) ∉ Fault ` F"
by (simp add: take_Suc_conv_app_nth)
have not_env_step:"¬ Γ⊢⇩c snd (clist ! i') ! m →⇩e snd (clist ! i') ! Suc m"
using False etran_ctran_eq_p_normal_s i' prod_step by blast
then have "snd ((snd(clist!i'))!0)∈ Normal ` p"
using len_l a7 i' a5 unfolding conjoin_def same_state_def par_assum_def by fastforce
then have "snd ((snd(clist!i'))!0)∈ Normal ` (Pre (xs ! i'))"
using a1 i' a3 by fastforce
then have "snd ((take (Suc (Suc m)) (snd(clist!i')))!0)∈ Normal `(Pre (xs ! i'))"
by fastforce
moreover have
"∀j. Suc j < Suc (Suc m) ⟶
Γ⊢⇩c snd (clist ! i') ! j →⇩e snd (clist ! i') ! Suc j ⟶
(snd (snd (clist ! i') ! j), snd (snd (clist ! i') ! Suc j)) ∈ Rely (xs ! i')"
using not_env_step fst_occ Suc_less_eq fist_occ i' less_SucE less_trans_Suc by auto
then have "∀j. Suc j < length (take (Suc (Suc m)) (snd(clist!i'))) ⟶
Γ⊢⇩c snd (clist ! i') ! j →⇩e snd (clist ! i') ! Suc j ⟶
(snd (snd (clist ! i') ! j), snd (snd (clist ! i') ! Suc j)) ∈ Rely (xs ! i')"
by fastforce
ultimately have "(Γ, (take (Suc (Suc m)) (snd(clist!i')))) ∈
assum ((Pre (xs ! i')),Rely (xs ! i'))"
unfolding assum_def by fastforce
moreover have "(Γ,snd(clist!i')) ∈ cptn" using a6 i' unfolding cp_def by fastforce
then have "(Γ,take (Suc (Suc m)) (snd(clist!i'))) ∈ cptn"
by (simp add: takecptn_is_cptn)
then have "(Γ,take (Suc (Suc m)) (snd(clist!i'))) ∈ cp Γ (Com(xs!i')) s"
using i' a3 a6 unfolding cp_def by fastforce
ultimately have t:"(Γ,take (Suc (Suc m)) (snd(clist!i'))) ∈
comm (Guar (xs ! i'), (Post (xs ! i'),Abr (xs ! i'))) F "
using a8 a2 a3 i' unfolding com_cvalidity_def com_validity_def by fastforce
have "(snd(take (Suc (Suc m)) (snd(clist!i'))!m),
snd(take (Suc (Suc m)) (snd(clist!i'))!(Suc m))) ∈ Guar (xs ! i')"
using eq_Γ i' comm_dest1[OF t last_take_not_f summ_lentake] by fastforce
then have "(snd( (snd(clist!i'))!m),
snd((snd(clist!i'))!(Suc m))) ∈ Guar (xs ! i')"
by fastforce
then have "(snd( (snd(clist!i))!m),
snd((snd(clist!i))!(Suc m))) ∈ Guar (xs ! i')"
using a7 fst_occ unfolding conjoin_def same_state_def by (metis Suc_lessD i' snd_conv)
then have "(snd( (snd(clist!i))!m),
snd((snd(clist!i))!(Suc m))) ∈ Rely (xs ! i)"
using not_eq a0 i' a3 fst_occ by auto
then have "False" using fst_occ by auto
then show ?thesis by auto
qed
qed
}
then have False using compat l by auto
} thus ?thesis by auto
qed
lemma two:
"⟦ ∀i<length xs. R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i));
p ⊆ (⋂i<length xs. (Pre (xs ! i)));
∀i<length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs ! i) sat [Pre (xs!i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)];
length xs=length clist; (Γ,l) ∈ par_cp Γ (ParallelCom xs) s; (Γ,l)∈par_assum (p, R);
∀i<length clist. clist!i∈cp Γ (Com(xs!i)) s; (Γ,l) ∝ clist;(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a]);
snd (last l) ∉ Fault ` F⟧
⟹ ∀j i ns ns'. i<length clist ∧ Suc j<length l ⟶
Γ⊢⇩c((snd (clist!i))!j) → ((snd (clist!i))!Suc j) ⟶
(snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈ Guar(xs!i) "
proof -
assume a0:"∀i<length xs. R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i))" and
a1:"p ⊆ (⋂i<length xs. (Pre (xs ! i)))" and
a2:"∀i<length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs ! i) sat [Pre (xs!i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)]" and
a3: "length xs=length clist" and
a4: "(Γ,l) ∈ par_cp Γ (ParallelCom xs) s" and
a5: "(Γ,l)∈par_assum (p, R)" and
a6: "∀i<length clist. clist!i∈cp Γ (Com(xs!i)) s" and
a7: "(Γ,l) ∝ clist" and
a8: "(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a])" and
a9: "snd (last l) ∉ Fault ` F"
{
assume a10:"(∃i j. i<length clist ∧ Suc j<length l ∧
Γ⊢⇩c((snd (clist!i))!j) → ((snd (clist!i))!Suc j) ∧
¬ (snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈ Guar(xs!i)) "
then obtain j where a10: "∃i. i<length clist ∧ Suc j<length l ∧
Γ⊢⇩c((snd (clist!i))!j) → ((snd (clist!i))!Suc j) ∧
¬ (snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈ Guar(xs!i)"
by fastforce
let ?P = "λj. ∃i. i<length clist ∧ Suc j<length l ∧
Γ⊢⇩c((snd (clist!i))!j) → ((snd (clist!i))!Suc j) ∧
¬ (snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈ Guar(xs!i)"
obtain m where fist_occ:"?P m ∧ (∀i<m. ¬ ?P i)" using exists_first_occ[of ?P j] a10 by blast
then have P:"?P m" by fastforce
then have notP:"(∀i<m. ¬ ?P i)" using fist_occ by blast
obtain i ns ns' where fst_occ:"i<length clist ∧ Suc m<length l ∧
Γ⊢⇩c((snd (clist!i))!m) → ((snd (clist!i))!Suc m) ∧
(¬ (snd((snd (clist!i))!m), snd((snd (clist!i))!Suc m)) ∈ Guar(xs!i))"
using P by fastforce
have fst_clist_i: "fst (clist!i) = Γ"
using a7 fst_occ unfolding conjoin_def same_functions_def
by fastforce
have "clist!i∈cp Γ (Com(xs!i)) s" using a6 fst_occ by fastforce
then have clistcp:"(Γ, snd (clist!i))∈cp Γ (Com(xs!i)) s"
using fst_occ a7 unfolding conjoin_def same_functions_def by fastforce
let ?li="take (Suc (Suc m)) (snd (clist!i))"
have "Γ ⊨⇘/F⇙ Com (xs ! i) sat [Pre (xs!i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)]"
using a8 a2 a3 fst_occ unfolding com_cvalidity_def by fastforce
moreover have take_in_ass:"(Γ, take (Suc (Suc m)) (snd (clist!i))) ∈ assum (Pre(xs!i), Rely(xs!i))"
proof -
have length_take_length_l:"length (take (Suc (Suc m)) (snd (clist!i))) ≤ length l"
using a7 fst_occ unfolding conjoin_def same_length_def by auto
have "snd((?li!0)) ∈ Normal ` Pre(xs!i)"
proof -
have "(take (Suc (Suc m)) (snd (clist!i)))!0 = (snd (clist!i))!0" by fastforce
moreover have "snd (snd(clist!i)!0) = snd (l!0)"
using a7 fst_occ unfolding conjoin_def same_state_def by fastforce
moreover have "snd (l!0) ∈ Normal ` p"
using a5 unfolding par_assum_def by fastforce
ultimately show ?thesis using a1 a3 fst_occ by fastforce
qed note left=this
thus ?thesis
using two'[OF a0 a1 a2 a3 a4 a5 a6 a7 a8 a9] fst_occ unfolding assum_def by fastforce
qed
moreover have "(Γ,take (Suc (Suc m)) (snd (clist!i))) ∈ cp Γ (Com(xs!i)) s"
using takecptn_is_cptn clistcp unfolding cp_def by fastforce
ultimately have comm:"(Γ, take (Suc (Suc m)) (snd (clist!i)))∈comm(Guar(xs!i),(Post (xs ! i),Abr (xs ! i))) F"
unfolding com_validity_def by fastforce
also have not_fault:"snd (last (take (Suc (Suc m)) (snd (clist!i)))) ∉ Fault ` F"
proof -
have cptn:"(Γ, snd (clist!i)) ∈ cptn"
using fst_clist_i a6 fst_occ unfolding cp_def by fastforce
then have "(snd (clist!i))≠[]"
using cptn.simps list.simps(3)
by fastforce
then have "snd (last (snd (clist!i))) = snd (last l)"
using conjoin_last_same_state fst_occ a7 by fastforce
then have "snd (last (snd (clist!i))) ∉ Fault ` F" using a9
by simp
also have sucm:"Suc m < length (snd (clist!i))"
using fst_occ a7 unfolding conjoin_def same_length_def by fastforce
ultimately have sucm_not_fault:"snd ((snd (clist!i))!(Suc m)) ∉ Fault ` F"
using last_not_F cptn by blast
have "length (take (Suc (Suc m)) (snd (clist!i))) = Suc (Suc m)"
using sucm by fastforce
then have "last (take (Suc (Suc m)) (snd (clist!i))) = (take (Suc (Suc m)) (snd (clist!i)))!(Suc m)"
by (metis Suc_diff_1 Suc_inject last_conv_nth list.size(3) old.nat.distinct(2) zero_less_Suc)
moreover have "(take (Suc (Suc m)) (snd (clist!i)))!(Suc m) = (snd (clist!i))!(Suc m)"
by fastforce
ultimately show ?thesis using sucm_not_fault by fastforce
qed
then have " (Suc m < length (snd (clist ! i)) ) ⟶
(Γ⊢⇩c (snd (clist ! i)) ! m → (snd (clist ! i)) ! Suc m) ⟶
(snd ((snd (clist ! i)) ! m), snd ((snd (clist ! i)) ! Suc m)) ∈ Guar(xs!i)"
using comm_dest [OF comm not_fault] by auto
then have "False" using fst_occ using a7 unfolding conjoin_def same_length_def by fastforce
} thus ?thesis by fastforce
qed
lemma par_cptn_env_comp:
"(Γ,l) ∈ par_cptn ∧ Suc i<length l ⟹
Γ⊢⇩p l!i →⇩e (l!(Suc i)) ∨ Γ ⊢⇩p l!i → (l!(Suc i))"
proof -
assume a0:"(Γ,l) ∈ par_cptn ∧ Suc i<length l"
then obtain c1 s1 c2 s2 where li:"l!i=(c1,s1) ∧ l!(Suc i) = (c2,s2)" by fastforce
obtain xs ys where l:"l= xs@((l!i)#(l!(Suc i))#ys)" using a0
by (metis Cons_nth_drop_Suc Suc_less_SucD id_take_nth_drop less_SucI)
moreover then have "(drop (length xs) l) = ((l!i)#(l!(Suc i))#ys)"
by (metis append_eq_conv_conj)
moreover then have "length xs < length l" using leI by fastforce
ultimately have "(Γ,((l!i)#(l!(Suc i))#ys))∈par_cptn"
using a0 droppar_cptn_is_par_cptn by fastforce
also then have "(Γ,(l!(Suc i))#ys)∈par_cptn" using par_cptn_dest li by fastforce
ultimately show ?thesis using li par_cptn_elim_cases(2)
by metis
qed
lemma three:
"⟦xs≠[]; ∀i<length xs. R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i));
p ⊆ (⋂i<length xs. (Pre (xs ! i)));
∀i<length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs ! i) sat [Pre (xs!i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)];
length xs=length clist; (Γ,l) ∈ par_cp Γ (ParallelCom xs) s; (Γ,l) ∈ par_assum(p, R);
∀i<length clist. clist!i∈cp Γ (Com(xs!i)) s; (Γ,l) ∝ clist; (∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a]);
snd (last l) ∉ Fault ` F⟧
⟹ ∀j i. i<length clist ∧ Suc j<length l ⟶ Γ⊢⇩c((snd (clist!i))!j) →⇩e ((snd (clist!i))!Suc j) ⟶
(snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈
(R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j))))"
proof -
assume a0:"xs≠[]" and
a1:"∀i<length xs. R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i))" and
a2: "p ⊆ (⋂i<length xs. (Pre (xs ! i)))" and
a3: "∀i<length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs ! i) sat [Pre (xs!i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)]" and
a4: "length xs=length clist" and
a5: "(Γ,l) ∈ par_cp Γ (ParallelCom xs) s" and
a6: "(Γ,l) ∈ par_assum(p, R)" and
a7: "∀i<length clist. clist!i∈cp Γ (Com(xs!i)) s" and
a8: "(Γ,l) ∝ clist" and
a9: "(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a])" and
10: "snd (last l) ∉ Fault ` F"
{
fix j i ns ns'
assume a00:"i<length clist ∧ Suc j<length l" and
a11: "Γ⊢⇩c((snd (clist!i))!j) →⇩e ((snd (clist!i))!Suc j)"
then have two:"∀j i ns ns'. i<length clist ∧ Suc j<length l ⟶
Γ⊢⇩c((snd (clist!i))!j) → ((snd (clist!i))!Suc j) ⟶
(snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈ (Guar(xs!i))"
using two[OF a1 a2 a3 a4 a5 a6 a7 a8 a9 10] by auto
then have j_lenl:"Suc j<length l" using a00 by fastforce
have i_lj:"i<length (fst (l!j)) ∧ i<length (fst (l!(Suc j)))"
using conjoin_same_length a00 a8 by fastforce
have fst_clist_Γ:"∀i<length clist. fst(clist!i) = Γ" using a8 unfolding conjoin_def same_functions_def by fastforce
have "(Γ⊢⇩p(l!j) → (l!(Suc j))) ∧
(∃i<length clist.
((fst (clist!i))⊢⇩c ((snd (clist!i))!j) → ((snd (clist!i))!(Suc j))) ∧
(∀l<length clist.
l≠i ⟶ (fst (clist!l))⊢⇩c (snd (clist!l))!j →⇩e ((snd (clist!l))!(Suc j)))) ∨
(Γ⊢⇩p(l!j) →⇩e (l!(Suc j)) ∧
(∀i<length clist. (fst (clist!i))⊢⇩c (snd (clist!i))!j →⇩e ((snd (clist!i))!(Suc j))))"
using a8 a00 unfolding conjoin_def compat_label_def by simp
then have compat_label:"(Γ⊢⇩p(l!j) → (l!(Suc j))) ∧
(∃i<length clist.
(Γ⊢⇩c ((snd (clist!i))!j) → ((snd (clist!i))!(Suc j))) ∧
(∀l<length clist.
l≠i ⟶ Γ⊢⇩c (snd (clist!l))!j →⇩e ((snd (clist!l))!(Suc j)))) ∨
(Γ⊢⇩p(l!j) →⇩e (l!(Suc j)) ∧
(∀i<length clist. Γ⊢⇩c (snd (clist!i))!j →⇩e ((snd (clist!i))!(Suc j))))"
using fst_clist_Γ by blast
then have "(snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈
(R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. Guar (xs ! j)))"
proof
assume a10:"(Γ⊢⇩p(l!j) → (l!(Suc j))) ∧
(∃i<length clist.
(Γ⊢⇩c ((snd (clist!i))!j) → ((snd (clist!i))!(Suc j))) ∧
(∀l<length clist.
l≠i ⟶ Γ⊢⇩c (snd (clist!l))!j →⇩e ((snd (clist!l))!(Suc j))))"
then obtain i' where
a20:"i'<length clist ∧
(Γ⊢⇩c ((snd (clist!i'))!j) → ((snd (clist!i'))!(Suc j))) ∧
(∀l<length clist.
l≠i' ⟶ Γ⊢⇩c (snd (clist!l))!j →⇩e ((snd (clist!l))!(Suc j)))" by blast
thus ?thesis
proof (cases "i'=i")
case True note eq_i = this
then obtain P S1 S2 where P:"(snd (clist!i'))!j=(P,S1) ∧ ((snd (clist!i'))!(Suc j)) = (P,S2)"
using a11 by (fastforce elim:etranE)
thus ?thesis
proof (cases "S1 = S2")
case True
have snd_lj:"(snd (l!j)) = snd ((snd (clist!i'))!j)"
using a8 a20 a00 unfolding conjoin_def same_state_def
by fastforce
have all_e:"(∀l<length clist. Γ⊢⇩c (snd (clist!l))!j →⇩e ((snd (clist!l))!(Suc j)))"
using a11 a20 eq_i by fastforce
then have allP:"∀l< length clist. fst ((snd (clist!l))!j) = fst ((snd (clist!l))!(Suc j))"
by (fastforce elim:etranE)
then have "fst (l!j) = (fst (l!(Suc j)))"
using a8 conjoin_same_program_i_j [of "(Γ,l)"] a00 by fastforce
also have "snd (l!j) = snd (l!(Suc j))"
proof -
have "(snd (l!Suc j)) = snd ((snd (clist!i'))!(Suc j))"
using a8 a20 a00 unfolding conjoin_def same_state_def
by fastforce
then show ?thesis using snd_lj P True by auto
qed
ultimately have "l!j = l!(Suc j)" by (simp add: prod_eq_iff)
moreover have ns1:"∃ns1. S1=Normal ns1"
using P a20 step_change_p_or_eq_s by fastforce
ultimately have "Γ⊢⇩p(l!j) →⇩e (l!(Suc j))"
using P step_pe.ParEnv snd_lj by (metis prod.collapse snd_conv)
then have "(snd (l ! j), snd (l ! Suc j)) ∈ R "
using a00 a6 unfolding par_assum_def by fastforce
then show ?thesis using a8 a00
unfolding conjoin_def same_state_def
by fastforce
next
case False thus ?thesis
using a20 P a11 step_change_p_or_eq_s by fastforce
qed
next
case False
have i'_clist:"i' < length clist" using a20 by fastforce
then have clist_i'_Guardxs:"(snd((snd (clist!i'))!j), snd((snd (clist!i'))!Suc j)) ∈ Guar(xs!i')"
using two a00 False a8 unfolding conjoin_def same_state_def
by (metis a20)
have "snd((snd (clist!i))!j) = snd (l!j) ∧ snd((snd (clist!i))!Suc j) = snd (l!Suc j)"
using a00 a20 a8 unfolding conjoin_def same_state_def by fastforce
also have "snd((snd (clist!i'))!j) = snd (l!j) ∧ snd((snd (clist!i'))!Suc j) = snd (l!Suc j)"
using j_lenl a20 a8 unfolding conjoin_def same_state_def by fastforce
ultimately have "snd((snd (clist!i))!j) = snd((snd (clist!i'))!j) ∧
snd((snd (clist!i))!Suc j) = snd((snd (clist!i'))!Suc j)"
by fastforce
then have clist_i_Guardxs:
"(snd((snd (clist!i))!j), snd((snd (clist!i))!Suc j)) ∈
Guar(xs!i')"
using clist_i'_Guardxs by fastforce
thus ?thesis
using False a20 a4 by fastforce
qed
next
assume a10:"(Γ⊢⇩p(l!j) →⇩e (l!(Suc j)) ∧
(∀i<length clist. Γ⊢⇩c (snd (clist!i))!j →⇩e ((snd (clist!i))!(Suc j))))"
then have "(snd (l ! j), snd (l ! Suc j)) ∈ R"
using a00 a10 a6 unfolding par_assum_def by fastforce
then show ?thesis using a8 a00
unfolding conjoin_def same_state_def
by fastforce
qed
} thus ?thesis by blast
qed
definition tran_True where "tran_True ≡ True"
definition after where "after ≡ True"
lemma four:
"⟦xs≠[]; ∀i<length xs. R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i));
(⋃j<length xs. (Guar (xs ! j))) ⊆ (G);
p ⊆ (⋂i<length xs. (Pre (xs ! i)));
∀i<length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs ! i) sat [Pre (xs!i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)];
(Γ,l) ∈ par_cp Γ (ParallelCom xs) s; (Γ,l) ∈ par_assum(p, R); Suc i < length l;
Γ⊢⇩p (l!i) → (l!(Suc i));
(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a]);
snd (last l) ∉ Fault ` F⟧
⟹ (snd (l ! i), snd (l ! Suc i)) ∈ G"
proof -
assume a0:"xs≠[]" and
a1:"∀i<length xs. R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i))" and
a2:"(⋃j<length xs. (Guar (xs ! j))) ⊆ (G)" and
a3:"p ⊆ (⋂i<length xs. (Pre (xs ! i)))" and
a4:"∀i<length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs ! i) sat [Pre (xs!i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)]" and
a5:"(Γ,l) ∈ par_cp Γ (ParallelCom xs) s" and
a6:"(Γ,l) ∈ par_assum(p, R)" and
a7: "Suc i < length l" and
a8:"Γ⊢⇩p (l!i) → (l!(Suc i))" and
a10:"(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a])" and
a11:"snd (last l) ∉ Fault ` F"
have length_par_xs:"length (ParallelCom xs) = length xs" unfolding ParallelCom_def by fastforce
then have "(ParallelCom xs)≠[]" using a0 by fastforce
then have "(Γ,l) ∈{(Γ1,c). ∃clist. (length clist)=(length (ParallelCom xs)) ∧
(∀i<length clist. (clist!i) ∈ cp Γ ((ParallelCom xs)!i) s) ∧ (Γ,c) ∝ clist ∧ Γ1=Γ}"
using one a5 by fastforce
then obtain clist where "(length clist)=(length xs) ∧
(∀i<length clist. (clist!i) ∈ cp Γ ((ParallelCom xs)!i) s) ∧ (Γ,l) ∝ clist"
using length_par_xs by auto
then have conjoin:"(length clist)=(length xs) ∧
(∀i<length clist. (clist!i) ∈ cp Γ (Com (xs ! i)) s) ∧ (Γ,l) ∝ clist"
using ParallelCom_Com by fastforce
then have length_xs_clist:"length xs = length clist" by auto
have clist_cp:"∀i<length clist. (clist!i) ∈ cp Γ (Com (xs ! i)) s" using conjoin by auto
have conjoin:"(Γ,l) ∝ clist" using conjoin by auto
have l_not_empty:"l≠[]" using a5 par_cptn.simps unfolding par_cp_def by fastforce
then have l_g0:"0<length l" by fastforce
then have last_l:"last l = l!((length l) - 1)" by (simp add: last_conv_nth)
have "∀i< length l. fst (l!i) = map (λx. fst ((snd x)!i)) clist"
using conjoin unfolding conjoin_def same_program_def by fastforce
obtain Ps si Ps' ssi where li:"l!i = (Ps,si) ∧ l!(Suc i) = (Ps', ssi)" by fastforce
then have "∃j r. j<length Ps ∧ Ps' = Ps[j:=r] ∧ (Γ⊢⇩c((Ps!j),si) → (r, ssi))"
using a8 par_ctranE by fastforce
then obtain j r where step_c:"j<length Ps ∧ Ps' = Ps[j:=r] ∧ (Γ⊢⇩c((Ps!j),si) → (r, ssi))"
by auto
have length_Ps_clist:
"length Ps = length clist ∧ length Ps = length Ps'"
using conjoin a7 conjoin_same_length li step_c by fastforce
have from_step:"(snd (clist!j))!i = ((Ps!j),si) ∧ (snd (clist!j))!(Suc i) = (Ps'!j,ssi)"
proof -
have f2: "Ps = fst (snd (Γ, l) ! i)" and f2':"Ps' = fst (snd (Γ, l) ! (Suc i))"
using li by auto
have f3:"si = snd (snd (Γ, l) ! i) ∧ ssi = snd (snd (Γ, l) ! (Suc i))"
by (simp add: li)
then have "(snd (clist!j))!i = ((Ps!j),si)"
using f2 conjoin a7 step_c unfolding conjoin_def same_program_def same_state_def by force
moreover have "(snd (clist!j))!(Suc i) = (Ps'!j,ssi)"
using f2' f3 conjoin a7 step_c length_Ps_clist
unfolding conjoin_def same_program_def same_state_def
by auto
ultimately show ?thesis by auto
qed
then have step_clist:"Γ⊢⇩c(snd (clist!j))!i → (snd (clist!j))!(Suc i)"
using from_step step_c by fastforce
have j_xs:"j<length xs" using step_c length_Ps_clist length_xs_clist by auto
have "j<length clist" using j_xs length_xs_clist by auto
also have
"∀i j ns ns'. j < length clist ∧ Suc i < length l ⟶
Γ⊢⇩c snd (clist ! j) ! i → snd (clist ! j) ! Suc i ⟶
(snd (snd (clist ! j) ! i), snd (snd (clist ! j) ! Suc i)) ∈ Guar (xs ! j)"
using two[OF a1 a3 a4 length_xs_clist a5 a6 clist_cp conjoin a10 a11] by auto
ultimately have "(snd (snd (clist ! j) ! i), snd (snd (clist ! j) ! Suc i)) ∈ Guar (xs ! j)"
using a7 step_c length_Ps_clist step_clist by metis
then have "(snd (l!i), snd (l!(Suc i)))∈ Guar (xs ! j)"
using from_step a2 length_xs_clist step_c li by fastforce
then show ?thesis using a2 j_xs
unfolding sep_conj_def tran_True_def after_def Satis_def by fastforce
qed
lemma same_program_last:"l≠[] ⟹ (Γ,l) ∝ clist ⟹ i<length clist ⟹fst (last (snd (clist!i))) = fst (last l) ! i"
proof -
assume l_not_empty:"l≠[]" and
conjoin: "(Γ,l) ∝ clist" and
i_clist: "i<length clist"
have last_clist_eq_l:"∀i<length clist. last (snd (clist!i)) = (snd (clist!i))!((length l) - 1)"
using conjoin last_conv_nth l_not_empty
unfolding conjoin_def same_length_def
by (metis length_0_conv snd_eqD)
then have last_l:"last l = l!((length l)-1)" using l_not_empty by (simp add: last_conv_nth)
have "fst (last l) = map (λx. fst (snd x ! ((length l)-1))) clist"
using l_not_empty last_l conjoin unfolding conjoin_def same_program_def by auto
also have "(map (λx. fst (snd x ! ((length l)-1))) clist)!i =
fst ((snd (clist!i))! ((length l)-1))" using i_clist by fastforce
also have "fst ((snd (clist!i))! ((length l)-1)) =
fst ((snd (clist!i))! ((length (snd (clist!i)))-1))"
using conjoin i_clist unfolding conjoin_def same_length_def by fastforce
also then have "fst ((snd (clist!i))! ((length (snd (clist!i)))-1)) = fst (last (snd (clist!i)))"
using i_clist l_not_empty conjoin last_clist_eq_l last_conv_nth unfolding conjoin_def same_length_def
by presburger
finally show ?thesis by auto
qed
lemma five:
"⟦xs≠[]; ∀i<length xs. R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i));
p ⊆ (⋂i<length xs. (Pre (xs ! i)));
(⋂i<length xs. (Post (xs ! i))) ⊆ q;
(⋃i<length xs. (Abr (xs ! i))) ⊆ a ;
∀i < length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs ! i) sat [Pre (xs!i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)];
(Γ,l) ∈ par_cp Γ (ParallelCom xs) s; (Γ,l) ∈ par_assum(p, R);
All_End (last l); snd (last l) ∉ Fault ` F;(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a]) ⟧ ⟹
(∃j<length (fst (last l)). fst (last l)!j=Throw ∧
snd (last l) ∈ Normal ` (a)) ∨
(∀j<length (fst (last l)). fst (last l)!j=Skip ∧
snd (last l) ∈ Normal ` q)"
proof-
assume a0:"xs≠[]" and
a1:"∀i<length xs. R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i))" and
a2:"p ⊆ (⋂i<length xs. (Pre (xs ! i)))" and
a3:"(⋂i<length xs. (Post (xs ! i))) ⊆ q" and
a4:"(⋃i<length xs. (Abr (xs ! i))) ⊆ a" and
a5:"∀i < length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs ! i) sat [Pre (xs!i),
Rely (xs ! i), Guar (xs ! i),
Post (xs ! i),Abr (xs ! i)]" and
a6:"(Γ,l) ∈ par_cp Γ (ParallelCom xs) s" and
a7:"(Γ,l) ∈ par_assum(p, R)"and
a8:"All_End (last l)" and
a9:"snd (last l) ∉ Fault ` F" and
a10:"(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a])"
have length_par_xs:"length (ParallelCom xs) = length xs" unfolding ParallelCom_def by fastforce
then have "(ParallelCom xs)≠[]" using a0 by fastforce
then have "(Γ,l) ∈{(Γ1,c). ∃clist. (length clist)=(length (ParallelCom xs)) ∧
(∀i<length clist. (clist!i) ∈ cp Γ ((ParallelCom xs)!i) s) ∧ (Γ,c) ∝ clist ∧ Γ1=Γ}"
using one a6 by fastforce
then obtain clist where "(length clist)=(length xs) ∧
(∀i<length clist. (clist!i) ∈ cp Γ ((ParallelCom xs)!i) s) ∧ (Γ,l) ∝ clist"
using length_par_xs by auto
then have conjoin:"(length clist)=(length xs) ∧
(∀i<length clist. (clist!i) ∈ cp Γ (Com (xs ! i)) s) ∧ (Γ,l) ∝ clist"
using ParallelCom_Com by fastforce
then have length_xs_clist:"length xs = length clist" by auto
have clist_cp:"∀i<length clist. (clist!i) ∈ cp Γ (Com (xs ! i)) s" using conjoin by auto
have conjoin:"(Γ,l) ∝ clist" using conjoin by auto
have l_not_empty:"l≠[]" using a6 par_cptn.simps unfolding par_cp_def by fastforce
then have l_g0:"0<length l" by fastforce
then have last_l:"last l = l!((length l) - 1)" by (simp add: last_conv_nth)
have clist_assum:"∀i<length clist. (clist!i) ∈ assum (Pre (xs!i),Rely (xs!i))"
proof -
{ fix i
assume i_length:"i<length clist"
obtain Γ1 li where clist:"clist!i=(Γ1,li)" by fastforce
then have Γeq:"Γ1=Γ"
using conjoin i_length unfolding conjoin_def same_functions_def by fastforce
have "(Γ1,li) ∈ assum (Pre (xs!i),Rely (xs!i))"
proof-
have l:"snd (li!0) ∈ Normal ` ( (Pre (xs!i)))"
proof -
have snd_l:"snd (Γ,l) = l" by fastforce
have "snd (l!0) ∈ Normal ` (p)"
using a7 unfolding par_assum_def by fastforce
also have "snd (l!0) = snd (li!0)"
using i_length conjoin l_g0 clist
unfolding conjoin_def same_state_def by fastforce
finally show ?thesis using a2 i_length length_xs_clist
by auto
qed
have r:"(∀j. Suc j < length li ⟶
Γ⊢⇩c(li!j) →⇩e (li!(Suc j)) ⟶
(snd(li!j), snd(li!(Suc j))) ∈ Rely (xs!i))"
using three[OF a0 a1 a2 a5 length_xs_clist a6 a7 clist_cp conjoin a10 a9]
i_length conjoin a1 length_xs_clist clist
unfolding assum_def conjoin_def same_length_def by fastforce
show ?thesis using l r Γeq unfolding assum_def by fastforce
qed
then have "clist!i ∈ assum (Pre (xs!i),Rely (xs!i))" using clist by auto
} thus ?thesis by auto
qed
then have clist_com:"∀i<length clist. (clist!i) ∈ comm (Guar (xs!i),(Post(xs!i),Abr (xs!i))) F"
using a5 unfolding com_cvalidity_def
using a10 unfolding com_validity_def using clist_cp length_xs_clist
by force
have last_clist_eq_l:"∀i<length clist. last (snd (clist!i)) = (snd (clist!i))!((length l) - 1)"
using conjoin last_conv_nth l_not_empty
unfolding conjoin_def same_length_def
by (metis length_0_conv snd_eqD)
then have last_clist_l:"∀i<length clist. snd (last (snd (clist!i))) = snd (last l)"
using last_l conjoin l_not_empty unfolding conjoin_def same_state_def same_length_def
by simp
show ?thesis
proof(cases "∀i<length (fst (last l)). fst (last l)!i = Skip")
assume ac1:"∀i<length (fst (last l)). fst (last l)!i = Skip"
have "(∀j<length (fst (last l)). fst (last l) ! j = LanguageCon.com.Skip ∧ snd (last l) ∈ Normal ` q)"
proof -
{fix j
assume aj:"j<length (fst (last l))"
have "∀i<length clist. snd (last (snd (clist!i))) ∈ Normal ` Post(xs!i)"
proof-
{fix i
assume a20:"i<length clist"
then have snd_last:"snd (last (snd (clist!i))) = snd (last l)"
using last_clist_l by fastforce
have last_clist_not_F:"snd (last (snd (clist!i)))∉ Fault ` F"
using a9 last_clist_l a20 by fastforce
have "fst (last l) ! i = Skip"
using a20 ac1 conjoin_same_length[OF conjoin]
by (simp add: l_not_empty last_l )
also have "fst (last l) ! i=fst (last (snd (clist!i)))"
using same_program_last[OF l_not_empty conjoin a20] by auto
finally have "fst (last (snd (clist!i))) = Skip" .
then have "snd (last (snd (clist!i))) ∈ Normal ` Post(xs!i)"
using clist_com last_clist_not_F a20
unfolding comm_def final_def by fastforce
} thus ?thesis by auto
qed
then have "∀i<length xs. snd (last l) ∈ Normal ` Post(xs!i)"
using last_clist_l length_xs_clist by fastforce
then have "∀i<length xs. ∃x∈( Post(xs!i)). snd (last l) = Normal x"
by fastforce
moreover have "∀t. (∀i<length xs. t∈ Post (xs ! i))⟶ t∈ q" using a3
by fastforce
ultimately have "(∃x∈ q. snd (last l) = Normal x)" using a0
by (metis (mono_tags, lifting) length_greater_0_conv xstate.inject(1))
then have "snd (last l) ∈ Normal ` q" by fastforce
then have "fst (last l) ! j = LanguageCon.com.Skip ∧ snd (last l) ∈ Normal ` q"
using aj ac1 by fastforce
} thus ?thesis by auto
qed
thus ?thesis by auto
next
assume "¬ (∀i<length (fst (last l)). fst (last l)!i = Skip)"
then obtain i where a20:"i< length (fst (last l)) ∧ fst (last l)!i ≠ Skip"
by fastforce
then have last_i_throw:"fst (last l)!i = Throw ∧ (∃n. snd (last l) = Normal n)"
using a8 unfolding All_End_def final_def by fastforce
have "length (fst (last l)) = length clist"
using conjoin_same_length[OF conjoin] l_not_empty last_l
by simp
then have i_length:"i<length clist" using a20 by fastforce
then have snd_last:"snd (last (snd (clist!i))) = snd (last l)"
using last_clist_l by fastforce
have last_clist_not_F:"snd (last (snd (clist!i)))∉ Fault ` F"
using a9 last_clist_l i_length by fastforce
then have "fst (last (snd (clist!i))) = fst (last l) ! i"
using i_length same_program_last [OF l_not_empty conjoin] by fastforce
then have "fst (last (snd (clist!i))) = Throw"
using last_i_throw by fastforce
then have "snd (last (snd (clist!i))) ∈ Normal ` Abr(xs!i)"
using clist_com last_clist_not_F i_length last_i_throw snd_last
unfolding comm_def final_def by fastforce
then have "snd (last l)∈ Normal ` Abr(xs!i)" using last_clist_l i_length
by fastforce
then have "snd (last l)∈ Normal ` (a)" using a4 a0 i_length length_xs_clist by fastforce
then have "∃j<length (fst (last l)).
fst (last l) ! j = LanguageCon.com.Throw ∧ snd (last l) ∈ Normal ` a"
using last_i_throw a20 by fastforce
thus ?thesis by auto
qed
qed
lemma ParallelEmpty [rule_format]:
"∀i s. (Γ,l) ∈ par_cp Γ (ParallelCom []) s ⟶
Suc i < length l ⟶ ¬ (Γ ⊢⇩p (l!i) → (l!Suc i))"
apply(induct_tac l)
apply simp
apply clarify
apply(case_tac list,simp,simp)
apply(case_tac i)
apply(simp add:par_cp_def ParallelCom_def)
apply(erule par_ctranE,simp)
apply(simp add:par_cp_def ParallelCom_def)
apply clarify
apply(erule par_cptn.cases,simp)
apply simp
by (metis list.inject list.size(3) not_less0 step_p_pair_elim_cases)
lemma ParallelEmpty2:
assumes a0:"(Γ,l) ∈ par_cp Γ (ParallelCom []) s" and
a1: "i < length l"
shows "fst (l!i) = []"
proof -
have paremp:"ParallelCom [] = []" unfolding ParallelCom_def by auto
then have l0:"l!0 =([],s)" using a0 unfolding par_cp_def by auto
then have "(Γ,l) ∈ par_cptn" using a0 unfolding par_cp_def by fastforce
thus ?thesis using l0 a1
proof (induct arbitrary: i s)
case ParCptnOne thus ?case by auto
next
case (ParCptnEnv Γ P s1 t xs i s)
thus ?case
proof -
have f1: "i < Suc (Suc (length xs))"
using ParCptnEnv.prems(2) by auto
have "(P, s1) = ([], s)"
using ParCptnEnv.prems(1) by auto
then show ?thesis
using f1 by (metis (no_types) ParCptnEnv.hyps(3) diff_Suc_1 fst_conv length_Cons less_Suc_eq_0_disj nth_Cons')
qed
next
case (ParCptnComp Γ P s1 Q t xs)
have "(Γ, (P,s1)#(Q, t) # xs) ∈ par_cp Γ (ParallelCom []) s1"
using ParCptnComp(4) ParCptnComp(1) step_p_elim_cases by fastforce
then have "¬ Γ⊢⇩p (P, s1) → (Q, t)" using ParallelEmpty ParCptnComp by fastforce
thus ?case using ParCptnComp by auto
qed
qed
lemma parallel_sound:
"∀i<length xs.
R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i)) ⟹
(⋃j<length xs. (Guar (xs ! j))) ⊆ G ⟹
p ⊆ (⋂i<length xs. (Pre (xs ! i))) ⟹
(⋂i<length xs. (Post (xs ! i))) ⊆ q ⟹
(⋃i<length xs. (Abr (xs ! i))) ⊆ a ⟹
∀i<length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs !i) sat [Pre (xs !i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)] ⟹
Γ,Θ ⊨⇘/F⇙ ParallelCom xs SAT [p, R, G, q,a]
"
proof -
assume
a0:"∀i<length xs.
R ∪ (⋃j∈{j. j < length xs ∧ j ≠ i}. (Guar (xs ! j)))
⊆ (Rely (xs ! i))" and
a1:"(⋃j<length xs. (Guar (xs ! j))) ⊆ G" and
a2:"p ⊆ (⋂i<length xs. (Pre (xs ! i)))" and
a3:"(⋂i<length xs. (Post (xs ! i))) ⊆ q" and
a4:"(⋃i<length xs. (Abr (xs ! i))) ⊆ a" and
a5:"∀i<length xs.
Γ,Θ ⊨⇘/F⇙ Com (xs !i) sat [Pre (xs !i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i),Abr (xs ! i)]"
{
assume a00:"(∀(c,p,R,G,q,a)∈ Θ. Γ ⊨⇘/F⇙ (Call c) sat [p, R, G, q,a])"
{ fix s l
assume a10: "(Γ,l) ∈ par_cp Γ (ParallelCom xs) s ∧ (Γ,l) ∈ par_assum(p, R)"
then have c_par_cp:"(Γ,l) ∈ par_cp Γ (ParallelCom xs) s" by auto
have c_par_assum: "(Γ,l) ∈ par_assum(p, R)" using a10 by auto
{ fix i ns ns'
assume a20:"snd (last l) ∉ Fault ` F"
{
assume a30:"Suc i<length l" and
a31: "Γ⊢⇩p(l!i) → (l!(Suc i))"
have xs_not_empty:"xs≠[]"
proof -
{
assume "xs = []"
then have "¬ (Γ ⊢⇩p (l!i) → (l!Suc i))"
using a30 a10 ParallelEmpty by fastforce
then have False using a31 by auto
} thus ?thesis by auto
qed
then have "(snd(l!i), snd(l!(Suc i))) ∈ G"
using four[OF xs_not_empty a0 a1 a2 a5 c_par_cp c_par_assum a30 a31 a00 a20] by blast
} then have "Suc i<length l ⟶
Γ⊢⇩p(l!i) → (l!(Suc i)) ⟶
(snd(l!i), snd(l!(Suc i))) ∈ G " by auto
note l = this
{ assume a30:"All_End (last l)"
then have xs_not_empty:"xs≠[]"
proof -
{ assume xs_emp:"xs=[]"
have lenl:"0<length l" using a10 unfolding par_cp_def using par_cptn.simps by fastforce
then have "(length l) - 1 < length l" by fastforce
then have "fst(l!((length l) - 1)) = []" using ParallelEmpty2 a10 xs_emp by fastforce
then have False using a30 lenl unfolding All_End_def
by (simp add: last_conv_nth )
} thus ?thesis by auto
qed
then have "(∃j<length (fst (last l)). fst (last l)!j=Throw ∧
snd (last l) ∈ Normal ` (a)) ∨
(∀j<length (fst (last l)). fst (last l)!j=Skip ∧
snd (last l) ∈ Normal ` q)"
using five[OF xs_not_empty a0 a2 a3 a4 a5 c_par_cp c_par_assum a30 a20 a00] by blast
} then have "All_End (last l) ⟶
(∃j<length (fst (last l)). fst (last l)!j=Throw ∧
snd (last l) ∈ Normal ` (a)) ∨
(∀j<length (fst (last l)). fst (last l)!j=Skip ∧
snd (last l) ∈ Normal ` q)" by auto
note res1 = conjI[OF l this]
}
then have "(Γ,l) ∈ par_comm(G, (q,a)) F" unfolding par_comm_def by auto
}
then have "Γ ⊨⇘/F⇙ (ParallelCom xs) SAT [p, R, G, q,a]"
unfolding par_com_validity_def par_cp_def by fastforce
} thus ?thesis using par_com_cvalidity_def by fastforce
qed
theorem
par_rgsound:"Γ,Θ ⊢⇘/F⇙ Ps SAT [p, R, G, q,a] ⟹
Γ,Θ ⊨⇘/F⇙ (ParallelCom Ps) SAT [p, R, G, q,a]"
proof (induction rule:par_rghoare.induct)
case (Parallel xs R G p q a Γ Θ F)
thus ?case using localRG_sound com_cnvalid_to_cvalid parallel_sound[of xs R G p q a Γ Θ F]
by fast
qed
lemma Conseq':"∀s. s∈p ⟶
(∃p' q' a' R' G'.
(∀Z. Γ,Θ⊢⇘/F⇙ P sat [(p' Z), (R' Z), (G' Z), (q' Z),(a' Z)]) ∧
(∃ Z. s∈p' Z ∧ (q' Z ⊆ q) ∧ (a' Z ⊆ a) ∧ (G' Z ⊆ G) ∧ (R ⊆ R' Z)))
⟹
Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q,a]"
apply (rule Conseq)
by (meson order_refl)
lemma conseq:"⟦∀Z. Γ,Θ'⊢⇘/F⇙ P sat [(p' Z), (R' Z), (G' Z), (q' Z),(a' Z)]; Θ' ⊆ Θ ;
∀s. s ∈ p ⟶ (∃ Z. s∈p' Z ∧ (q' Z ⊆ q) ∧ (a' Z ⊆ a) ∧ (G' Z ⊆ G) ∧ (R ⊆ R' Z))⟧
⟹
Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q,a]"
by (rule Conseq) (meson order_refl)
lemma conseqPrePost[trans]:
"Γ,Θ'⊢⇘/F⇙ P sat [p', R', G', q',a'] ⟹ Θ' ⊆ Θ ⟹
p⊆p' ⟹ q' ⊆ q ⟹ a' ⊆ a ⟹ G' ⊆ G ⟹ R ⊆ R' ⟹
Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q,a]"
by (rule conseq) auto
lemma conseqPre[trans]:
"Γ,Θ⊢⇘/F⇙ P sat [p', R, G, q,a] ⟹
p⊆p' ⟹
Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q,a]"
by (rule conseq) auto
lemma conseqPost[trans]:
"Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q',a'] ⟹
q'⊆q ⟹ a'⊆a ⟹
Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q,a]"
by (rule conseq) auto
lemma shows x:"∃(sa'::nat set). (∀x. (x∈ sa) = ((to_nat x) ∈ sa'))"
by (metis (mono_tags, hide_lams) from_nat_to_nat imageE image_eqI)
lemma not_empty_set_countable:
assumes a0:"sa≠({}::('a::countable) set)"
shows "{i. ((λi. i∈ sa) o from_nat) i}≠{}"
by (metis (full_types) Collect_empty_eq_bot assms comp_apply empty_def equals0I from_nat_to_nat)
lemma eq_set_countable:"(⋂i∈{i. ((λi. i∈ sa) o from_nat) i}. (q o from_nat) i) = ((⋂i∈sa. q i))"
apply auto
by (metis (no_types) from_nat_to_nat)
lemma conj_inter_countable[trans]:
assumes a0:"sa≠({}::('a::countable) set)" and
a1:"∀i∈sa. Γ,Θ⊢⇘/F⇙ P sat [p, R, G, q i,a]"
shows"Γ,Θ⊢⇘/F⇙ P sat [p, R, G,(⋂i∈sa. q i),a]"
proof-
have "∀i∈{i. ((λi. i∈ sa) o from_nat) i}. Γ,Θ⊢⇘/F⇙ P sat [p, R, G,(q o from_nat) i,a]"
using a1 by auto
then have "Γ,Θ ⊢⇘/F⇙ P sat [p, R, G,⋂i∈{i. ((λi. i∈ sa) o from_nat) i}. (q o from_nat) i,a]"
using Conj_Inter[OF not_empty_set_countable[OF a0]] by auto
thus ?thesis using eq_set_countable
by metis
qed
lemma all_Post[trans]:
assumes a0:"∀p_n::('a::countable). Γ,Θ⊢⇘/F⇙ C sat [P, R, G, Q p_n, Qa]"
shows"Γ,Θ⊢⇘/F⇙ C sat [P, R, G,{s. ∀p_n. s∈Q p_n},Qa]"
proof-
have "Γ,Θ⊢⇘/F⇙ C sat [P, R, G,(⋂p_n. Q p_n),Qa]"
using a0 conj_inter_countable[of UNIV] by auto
moreover have s1:"∀P. {s. ∀p_n. s∈P p_n} = (⋂p_n. P p_n)"
by auto
ultimately show ?thesis
by (simp add: s1)
qed
lemma all_Pre[trans]:
assumes a0:"∀p_n. Γ,Θ⊢⇘/F⇙ C sat [P p_n, R, G, Q, Qa]"
shows"Γ,Θ⊢⇘/F⇙ C sat [{s. ∀p_n. s∈P p_n}, R, G,Q,Qa]"
proof-
{fix p_n
have "Γ,Θ⊢⇘/F⇙ C sat [{s. ∀p_n. s∈P p_n}, R, G,Q,Qa]"
proof-
have "{v. ∀n. v ∈ P n} ⊆ P p_n" by force
then show ?thesis by (meson a0 LocalRG_HoareDef.conseqPrePost subset_eq)
qed
} thus ?thesis by auto
qed
lemma Pre_Post_all:
assumes a0:"∀p_n::('a::countable). Γ,Θ⊢⇘/F⇙ C sat [P p_n, R, G, Q p_n, Qa]"
shows"Γ,Θ⊢⇘/F⇙ C sat [{s. ∀p_n. s∈P p_n}, R, G,{s. ∀p_n. s∈Q p_n},Qa]"
proof-
{fix p_n
have "Γ,Θ⊢⇘/F⇙ C sat [{s. ∀p_n. s∈P p_n}, R, G,Q p_n,Qa]"
proof-
have "{v. ∀n. v ∈ P n} ⊆ P p_n" by force
then show ?thesis by (meson a0 LocalRG_HoareDef.conseqPrePost subset_eq)
qed
}
then have f3:"∀p_n. Γ,Θ⊢⇘/F⇙ C sat [{s. ∀p_n. s∈P p_n}, R, G,Q p_n,Qa]"
by auto
then have "∀p_n. Γ,Θ⊢⇘/F⇙ C sat [{s. ∀p_n. s∈P p_n}, R, G,{s. ∀p_n. s∈Q p_n},Qa]"
using all_Post by auto
moreover have s1:"∀P. {s. ∀p_n. s∈P p_n} = (⋂p_n. P p_n)"
by auto
ultimately show ?thesis
by (simp add: s1)
qed
inductive_cases hoare_elim_skip_cases [cases set]:
"Γ,Θ⊢⇘/F⇙ Skip sat [p, R, G, q,a]"
end