Theory ReduceStoreBufferSimulation
theory ReduceStoreBufferSimulation
imports ReduceStoreBuffer
begin
locale initial⇩s⇩b = simple_ownership_distinct + read_only_unowned + unowned_shared +
constrains ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes empty_sb: "⟦i < length ts; ts!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟧ ⟹ sb=[]"
assumes empty_is: "⟦i < length ts; ts!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟧ ⟹ is=[]"
assumes empty_rels: "⟦i < length ts; ts!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟧ ⟹ ℛ=Map.empty"
sublocale initial⇩s⇩b ⊆ outstanding_non_volatile_refs_owned_or_read_only
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "non_volatile_owned_or_read_only False 𝒮 𝒪 sb"
using empty_sb [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ outstanding_volatile_writes_unowned_by_others
proof
fix i j p⇩i is⇩i 𝒪⇩i ℛ⇩i 𝒟⇩i θ⇩i sb⇩i p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j
assume i_bound: "i < length ts" and
j_bound: "j < length ts" and
neq_i_j: "i ≠ j" and
ts_i: "ts ! i = (p⇩i, is⇩i, θ⇩i, sb⇩i, 𝒟⇩i, 𝒪⇩i, ℛ⇩i)" and
ts_j: "ts ! j = (p⇩j, is⇩j, θ⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j, ℛ⇩j)"
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩i = {}"
using empty_sb [OF i_bound ts_i] empty_sb [OF j_bound ts_j] by auto
qed
sublocale initial⇩s⇩b ⊆ read_only_reads_unowned
proof
fix i j p⇩i is⇩i 𝒪⇩i ℛ⇩i 𝒟⇩i θ⇩i sb⇩i p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j
assume i_bound: "i < length ts" and
j_bound: "j < length ts" and
neq_i_j: "i ≠ j" and
ts_i: "ts ! i = (p⇩i, is⇩i, θ⇩i, sb⇩i, 𝒟⇩i, 𝒪⇩i, ℛ⇩i)" and
ts_j: "ts ! j = (p⇩j, is⇩j, θ⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j, ℛ⇩j)"
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩
read_only_reads (acquired True
(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩i) 𝒪⇩i)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩i) = {}"
using empty_sb [OF i_bound ts_i] empty_sb [OF j_bound ts_j] by auto
qed
sublocale initial⇩s⇩b ⊆ ownership_distinct
proof
fix i j p⇩i is⇩i 𝒪⇩i ℛ⇩i 𝒟⇩i θ⇩i sb⇩i p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j
assume i_bound: "i < length ts" and
j_bound: "j < length ts" and
neq_i_j: "i ≠ j" and
ts_i: "ts ! i = (p⇩i, is⇩i, θ⇩i, sb⇩i, 𝒟⇩i, 𝒪⇩i, ℛ⇩i)" and
ts_j: "ts ! j = (p⇩j, is⇩j, θ⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j, ℛ⇩j)"
show "(𝒪⇩i ∪ all_acquired sb⇩i) ∩ (𝒪⇩j ∪ all_acquired sb⇩j) = {}"
using simple_ownership_distinct [OF i_bound j_bound neq_i_j ts_i ts_j] empty_sb [OF i_bound ts_i] empty_sb [OF j_bound ts_j]
by auto
qed
sublocale initial⇩s⇩b ⊆ valid_ownership ..
sublocale initial⇩s⇩b ⊆ outstanding_non_volatile_writes_unshared
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "non_volatile_writes_unshared 𝒮 sb"
using empty_sb [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ sharing_consis
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "sharing_consistent 𝒮 𝒪 sb"
using empty_sb [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ no_outstanding_write_to_read_only_memory
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "no_write_to_read_only_memory 𝒮 sb"
using empty_sb [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ valid_sharing ..
sublocale initial⇩s⇩b ⊆ valid_ownership_and_sharing ..
sublocale initial⇩s⇩b ⊆ load_tmps_distinct
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "distinct_load_tmps is"
using empty_is [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ read_tmps_distinct
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "distinct_read_tmps sb"
using empty_sb [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ load_tmps_read_tmps_distinct
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "load_tmps is ∩ read_tmps sb = {}"
using empty_sb [OF i_bound ts_i] empty_is [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ load_tmps_read_tmps_distinct ..
sublocale initial⇩s⇩b ⊆ valid_write_sops
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "∀sop ∈ write_sops sb. valid_sop sop"
using empty_sb [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ valid_store_sops
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "∀sop ∈ store_sops is. valid_sop sop"
using empty_is [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ valid_sops ..
sublocale initial⇩s⇩b ⊆ valid_reads
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "reads_consistent False 𝒪 m sb"
using empty_sb [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ valid_history
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "program.history_consistent program_step θ (hd_prog p sb) sb"
using empty_sb [OF i_bound ts_i] by (auto simp add: program.history_consistent.simps)
qed
sublocale initial⇩s⇩b ⊆ valid_data_dependency
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "data_dependency_consistent_instrs (dom θ) is"
using empty_is [OF i_bound ts_i] by auto
next
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "load_tmps is ∩ ⋃(fst ` write_sops sb) = {}"
using empty_is [OF i_bound ts_i] empty_sb [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ load_tmps_fresh
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "load_tmps is ∩ dom θ = {}"
using empty_is [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ enough_flushs
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "outstanding_refs is_volatile_Write⇩s⇩b sb = {}"
using empty_sb [OF i_bound ts_i] by auto
qed
sublocale initial⇩s⇩b ⊆ valid_program_history
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p sb⇩1 sb⇩2
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
assume sb: "sb=sb⇩1@sb⇩2"
show "∃isa. instrs sb⇩2 @ is = isa @ prog_instrs sb⇩2"
using empty_sb [OF i_bound ts_i] empty_is [OF i_bound ts_i] sb by auto
next
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "last_prog p sb = p"
using empty_sb [OF i_bound ts_i] by auto
qed
inductive
sim_config:: "('p,'p store_buffer,bool,owns,rels) thread_config list × memory × shared ⇒
('p, unit,bool,owns,rels) thread_config list × memory × shared ⇒ bool"
(‹_ ∼ _› [60,60] 100)
where
"⟦m = flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b;
𝒮 = share_all_until_volatile_write ts⇩s⇩b 𝒮⇩s⇩b;
length ts⇩s⇩b = length ts;
∀i < length ts⇩s⇩b.
let (p, is⇩s⇩b, θ, sb, 𝒟⇩s⇩b, 𝒪, ℛ) = ts⇩s⇩b!i;
suspends = dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb
in ∃is 𝒟. instrs suspends @ is⇩s⇩b = is @ prog_instrs suspends ∧
𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b sb ≠ {}) ∧
ts!i = (hd_prog p suspends,
is,
θ |` (dom θ - read_tmps suspends),(),
𝒟,
acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪,
release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) (dom 𝒮⇩s⇩b) ℛ )
⟧
⟹
(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m,𝒮)"
text ‹The machine without
history only stores writes in the store-buffer.›
inductive sim_history_config::
"('p,'p store_buffer,'dirty,'owns,'rels) thread_config list ⇒ ('p,'p store_buffer,bool,owns,rels) thread_config list ⇒ bool"
(‹_ ∼⇩h _ › [60,60] 100)
where
"⟦length ts = length ts⇩h;
∀i < length ts.
(∃𝒪' 𝒟' ℛ'.
let (p,is, θ, sb,𝒟, 𝒪,ℛ) = ts⇩h!i in
ts!i=(p,is, θ, filter is_Write⇩s⇩b sb,𝒟',𝒪',ℛ') ∧
(filter is_Write⇩s⇩b sb = [] ⟶ sb=[]))
⟧
⟹
ts ∼⇩h ts⇩h"
lemma (in initial⇩s⇩b) history_refl:"ts ∼⇩h ts"
apply -
apply (rule sim_history_config.intros)
apply simp
apply clarsimp
subgoal for i
apply (case_tac "ts!i")
apply (drule_tac i=i in empty_sb)
apply assumption
apply auto
done
done
lemma share_all_empty: "∀i p is xs sb 𝒟 𝒪 ℛ. i < length ts ⟶ ts!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟶ sb=[]
⟹ share_all_until_volatile_write ts 𝒮 = 𝒮"
apply (induct ts)
apply clarsimp
apply clarsimp
apply (frule_tac x=0 in spec)
apply clarsimp
apply force
done
lemma flush_all_empty: "∀i p is xs sb 𝒟 𝒪 ℛ. i < length ts ⟶ ts!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟶ sb=[]
⟹ flush_all_until_volatile_write ts m = m"
apply (induct ts)
apply clarsimp
apply clarsimp
apply (frule_tac x=0 in spec)
apply clarsimp
apply force
done
lemma sim_config_emptyE:
assumes empty:
"∀i p is xs sb 𝒟 𝒪 ℛ. i < length ts⇩s⇩b ⟶ ts⇩s⇩b!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟶ sb=[]"
assumes sim: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m,𝒮)"
shows "𝒮 = 𝒮⇩s⇩b ∧ m = m⇩s⇩b ∧ length ts = length ts⇩s⇩b ∧
(∀i < length ts⇩s⇩b.
let (p, is, θ, sb, 𝒟, 𝒪, ℛ) = ts⇩s⇩b!i
in ts!i = (p, is, θ, (), 𝒟, 𝒪, ℛ))"
proof -
from sim
show ?thesis
apply cases
apply (clarsimp simp add: flush_all_empty [OF empty] share_all_empty [OF empty])
subgoal for i
apply (drule_tac x=i in spec)
apply (cut_tac i=i in empty [rule_format])
apply clarsimp
apply assumption
apply (auto simp add: Let_def)
done
done
qed
lemma sim_config_emptyI:
assumes empty:
"∀i p is xs sb 𝒟 𝒪 ℛ. i < length ts⇩s⇩b ⟶ ts⇩s⇩b!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟶ sb=[]"
assumes leq: "length ts = length ts⇩s⇩b"
assumes ts: "(∀i < length ts⇩s⇩b.
let (p, is, θ, sb, 𝒟, 𝒪, ℛ) = ts⇩s⇩b!i
in ts!i = (p, is, θ, (), 𝒟, 𝒪, ℛ))"
shows "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m⇩s⇩b,𝒮⇩s⇩b)"
apply (rule sim_config.intros)
apply (simp add: flush_all_empty [OF empty])
apply (simp add: share_all_empty [OF empty])
apply (simp add: leq)
apply (clarsimp)
apply (frule (1) empty [rule_format])
using ts
apply (auto simp add: Let_def)
done
lemma mem_eq_un_eq: "⟦length ts'=length ts; ∀i< length ts'. P (ts'!i) = Q (ts!i) ⟧ ⟹ (⋃x∈set ts'. P x) = (⋃x∈set ts. Q x)"
apply (auto simp add: in_set_conv_nth )
apply (force dest!: nth_mem)
apply (frule nth_mem)
subgoal for x i
apply (drule_tac x=i in spec)
apply auto
done
done
lemma (in program) trace_to_steps:
assumes trace: "trace c 0 k"
shows steps: "c 0 ⇒⇩d⇧* c k"
using trace
proof (induct k)
case 0
show "c 0 ⇒⇩d⇧* c 0"
by auto
next
case (Suc k)
have prem: "trace c 0 (Suc k)" by fact
hence "trace c 0 k"
by (auto simp add: program_trace_def)
from Suc.hyps [OF this]
have "c 0 ⇒⇩d⇧* c k" .
also
term program_trace
from prem interpret program_trace program_step c 0 "Suc k" .
from step [of k] have "c (k) ⇒⇩d c (Suc k)"
by auto
finally show ?case .
qed
lemma (in program) safe_reach_to_safe_reach_upto:
assumes safe_reach: "safe_reach_direct safe c⇩0"
shows "safe_reach_upto n safe c⇩0"
proof
fix k c l
assume k_n: "k ≤ n"
assume trace: "trace c 0 k"
assume c_0: "c 0 = c⇩0"
assume l_k: "l ≤ k"
show "safe (c l)"
proof -
from trace k_n l_k have trace': "trace c 0 l"
by (auto simp add: program_trace_def)
from trace_to_steps [OF trace']
have "c 0 ⇒⇩d⇧* c l".
with safe_reach c_0 show "safe (c l)"
by (cases "c l") (auto simp add: safe_reach_def)
qed
qed
lemma (in program_progress) safe_free_flowing_implies_safe_delayed':
assumes init: "initial⇩s⇩b ts⇩s⇩b 𝒮⇩s⇩b"
assumes sim: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m,𝒮)"
assumes safe_reach_ff: "safe_reach_direct safe_free_flowing (ts,m,𝒮)"
shows "safe_reach_direct safe_delayed (ts,m,𝒮)"
proof -
from init
interpret ini: initial⇩s⇩b ts⇩s⇩b 𝒮⇩s⇩b .
from sim obtain
m: "m = flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b" and
𝒮: "𝒮 = share_all_until_volatile_write ts⇩s⇩b 𝒮⇩s⇩b" and
leq: "length ts⇩s⇩b = length ts" and
t_sim: "∀i < length ts⇩s⇩b.
let (p, is⇩s⇩b, θ, sb, 𝒟⇩s⇩b, 𝒪, ℛ) = ts⇩s⇩b!i;
suspends = dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb
in ∃is 𝒟. instrs suspends @ is⇩s⇩b = is @ prog_instrs suspends ∧
𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b sb ≠ {}) ∧
ts!i = (hd_prog p suspends,
is,
θ |` (dom θ - read_tmps suspends),(),
𝒟,
acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪,
release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) (dom 𝒮⇩s⇩b) ℛ )"
by cases auto
from ini.empty_sb
have shared_eq: "𝒮 = 𝒮⇩s⇩b"
apply (simp only: 𝒮)
apply (rule share_all_empty)
apply force
done
have sd: "simple_ownership_distinct ts"
proof
fix i j p⇩i is⇩i 𝒪⇩i ℛ⇩i 𝒟⇩i θ⇩i sb⇩i p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j
assume i_bound: "i < length ts" and
j_bound: "j < length ts" and
neq_i_j: "i ≠ j" and
ts_i: "ts ! i = (p⇩i, is⇩i, θ⇩i, sb⇩i, 𝒟⇩i, 𝒪⇩i, ℛ⇩i)" and
ts_j: "ts ! j = (p⇩j, is⇩j, θ⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j, ℛ⇩j)"
show "(𝒪⇩i) ∩ (𝒪⇩j ) = {}"
proof -
from t_sim [simplified leq, rule_format, OF i_bound] ini.empty_sb [simplified leq, OF i_bound]
have ts_i: "ts⇩s⇩b!i = (p⇩i,is⇩i,θ⇩i,[],𝒟⇩i,𝒪⇩i,ℛ⇩i)"
using ts_i
by (force simp add: Let_def)
from t_sim [simplified leq, rule_format, OF j_bound] ini.empty_sb [simplified leq, OF j_bound]
have ts_j: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,[],𝒟⇩j,𝒪⇩j,ℛ⇩j)"
using ts_j
by (force simp add: Let_def)
from ini.simple_ownership_distinct [simplified leq, OF i_bound j_bound neq_i_j ts_i ts_j]
show ?thesis .
qed
qed
have ro: "read_only_unowned 𝒮 ts"
proof
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
show "𝒪 ∩ read_only 𝒮 = {}"
proof -
from t_sim [simplified leq, rule_format, OF i_bound] ini.empty_sb [simplified leq, OF i_bound]
have ts_i: "ts⇩s⇩b!i = (p,is,θ,[],𝒟,𝒪,ℛ)"
using ts_i
by (force simp add: Let_def)
from ini.read_only_unowned [simplified leq, OF i_bound ts_i] shared_eq
show ?thesis by simp
qed
qed
have us: "unowned_shared 𝒮 ts"
proof
show "- (⋃((λ(_, _, _, _, _, 𝒪, _). 𝒪) ` set ts)) ⊆ dom 𝒮"
proof -
have "(⋃((λ(_, _, _, _, _, 𝒪, _). 𝒪) ` set ts⇩s⇩b)) = (⋃((λ(_, _, _, _, _, 𝒪, _). 𝒪) ` set ts))"
apply clarsimp
apply (rule mem_eq_un_eq)
apply (simp add: leq)
apply clarsimp
apply (frule t_sim [rule_format])
apply (clarsimp simp add: Let_def)
apply (drule (1) ini.empty_sb)
apply auto
done
with ini.unowned_shared show ?thesis by (simp only: shared_eq)
qed
qed
{
fix i "is" 𝒪 ℛ 𝒟 θ sb p
assume i_bound: "i < length ts"
assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
have "ℛ = Map.empty"
proof -
from t_sim [simplified leq, rule_format, OF i_bound] ini.empty_sb [simplified leq, OF i_bound]
have ts_i: "ts⇩s⇩b!i = (p,is,θ,[],𝒟,𝒪,ℛ)"
using ts_i
by (force simp add: Let_def)
from ini.empty_rels [simplified leq, OF i_bound ts_i]
show ?thesis .
qed
}
with us have initial: "initial (ts, m, 𝒮)"
by (fastforce simp add: initial_def)
{
fix ts' 𝒮' m'
assume steps: "(ts,m,𝒮) ⇒⇩d⇧* (ts',m',𝒮')"
have "safe_delayed (ts',m',𝒮')"
proof -
from steps_to_trace [OF steps] obtain c k
where trace: "trace c 0 k" and c_0: "c 0 = (ts,m,𝒮)" and c_k: "c k = (ts',m',𝒮')"
by auto
from safe_reach_to_safe_reach_upto [OF safe_reach_ff]
have safe_upto_k: "safe_reach_upto k safe_free_flowing (ts, m, 𝒮)".
from safe_free_flowing_implies_safe_delayed [OF _ _ _ _ safe_upto_k, simplified, OF initial sd ro us]
have "safe_reach_upto k safe_delayed (ts, m, 𝒮)".
then interpret program_safe_reach_upto program_step k safe_delayed "(ts,m,𝒮)" .
from safe_config [where c=c and k=k and l=k, OF _ trace c_0] c_k show ?thesis by simp
qed
}
then show ?thesis
by (clarsimp simp add: safe_reach_def)
qed
lemma map_onws_sb_owned:"⋀j. j < length ts ⟹ map 𝒪_sb ts ! j = (𝒪⇩j,sb⇩j) ⟹ map owned ts ! j = 𝒪⇩j"
apply (induct ts)
apply simp
subgoal for t ts j
apply (case_tac j)
apply (case_tac t)
apply auto
done
done
lemma map_onws_sb_owned':"⋀j. j < length ts ⟹ 𝒪_sb (ts ! j) = (𝒪⇩j,sb⇩j) ⟹ owned (ts ! j) = 𝒪⇩j"
apply (induct ts)
apply simp
subgoal for t ts j
apply (case_tac j)
apply (case_tac t)
apply auto
done
done
lemma read_only_read_acquired_unforwarded_acquire_witness:
"⋀𝒮 𝒪 X.⟦non_volatile_owned_or_read_only True 𝒮 𝒪 sb;
sharing_consistent 𝒮 𝒪 sb; a ∉ read_only 𝒮; a ∉ 𝒪;
a ∈ unforwarded_non_volatile_reads sb X⟧
⟹(∃sop a' v ys zs A L R W.
sb = ys @ Write⇩s⇩b True a' sop v A L R W # zs ∧
a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a' ≠ a) ∨
(∃A L R W ys zs. sb = ys @ Ghost⇩s⇩b A L R W# zs ∧ a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys)"
proof (induct sb)
case Nil thus ?case by simp
next
case (Cons x sb)
show ?case
proof (cases x)
case (Write⇩s⇩b volatile a' sop v A L R W)
show ?thesis
proof (cases volatile)
case True
note volatile=this
from Cons.prems obtain
nvo': "non_volatile_owned_or_read_only True (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) (𝒪 ∪ A - R) sb" and
a_nro: "a ∉ read_only 𝒮" and
a_unowned: "a ∉ 𝒪" and
A_shared_owns: "A ⊆ dom 𝒮 ∪ 𝒪" and L_A: "L ⊆ A" and A_R: "A ∩ R = {}" and
R_owns: "R ⊆ 𝒪" and
consis': "sharing_consistent (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) (𝒪 ∪ A - R) sb" and
a_unforw: "a ∈ unforwarded_non_volatile_reads sb (insert a' X)"
by (clarsimp simp add: Write⇩s⇩b True)
from unforwarded_not_written [OF a_unforw]
have a_notin: "a ∉ insert a' X".
hence a'_a: "a' ≠ a"
by simp
from R_owns a_unowned
have a_R: "a ∉ R"
by auto
show ?thesis
proof (cases "a ∈ A")
case True
then show ?thesis
apply -
apply (rule disjI1)
apply (rule_tac x=sop in exI)
apply (rule_tac x=a' in exI)
apply (rule_tac x=v in exI)
apply (rule_tac x="[]" in exI)
apply (rule_tac x=sb in exI)
apply (simp add: Write⇩s⇩b volatile True a'_a)
done
next
case False
with a_unowned R_owns a_nro L_A A_R
obtain a_nro': "a ∉ read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)" and a_unowned': "a ∉ 𝒪 ∪ A - R"
by (force simp add: in_read_only_convs)
from Cons.hyps [OF nvo' consis' a_nro' a_unowned' a_unforw]
have "(∃sop a' v ys zs A L R W.
sb = ys @ Write⇩s⇩b True a' sop v A L R W # zs ∧
a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a' ≠ a) ∨
(∃A L R W ys zs. sb = ys @ Ghost⇩s⇩b A L R W# zs ∧ a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys)"
(is "?write ∨ ?ghst")
by simp
then show ?thesis
proof
assume ?write
then obtain sop' a'' v' ys zs A' L' R' W' where
sb: "sb = ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs" and
props: "a ∈ A'" "a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a'' ≠ a"
by auto
show ?thesis
using props False a_notin sb
apply -
apply (rule disjI1)
apply (rule_tac x=sop' in exI)
apply (rule_tac x=a'' in exI)
apply (rule_tac x=v' in exI)
apply (rule_tac x="(x#ys)" in exI)
apply (rule_tac x=zs in exI)
apply (simp add: Write⇩s⇩b volatile False a'_a)
done
next
assume ?ghst
then obtain ys zs A' L' R' W' where
sb: "sb = ys @ Ghost⇩s⇩b A' L' R' W'# zs" and
props: "a ∈ A'" "a ∉ outstanding_refs is_Write⇩s⇩b ys"
by auto
show ?thesis
using props False a_notin sb
apply -
apply (rule disjI2)
apply (rule_tac x=A' in exI)
apply (rule_tac x=L' in exI)
apply (rule_tac x=R' in exI)
apply (rule_tac x=W' in exI)
apply (rule_tac x="(x#ys)" in exI)
apply (rule_tac x=zs in exI)
apply (simp add: Write⇩s⇩b volatile False a'_a)
done
qed
qed
next
case False
from Cons.prems obtain
consis': "sharing_consistent 𝒮 𝒪 sb" and
a_nro': "a ∉ read_only 𝒮" and
a_unowned: "a ∉ 𝒪" and
a_ro': "a' ∈ 𝒪" and
nvo': "non_volatile_owned_or_read_only True 𝒮 𝒪 sb" and
a_unforw': "a ∈ unforwarded_non_volatile_reads sb (insert a' X)"
by (auto simp add: Write⇩s⇩b False split: if_split_asm)
from unforwarded_not_written [OF a_unforw']
have a_notin: "a ∉ insert a' X".
from Cons.hyps [OF nvo' consis' a_nro' a_unowned a_unforw']
have "(∃sop a' v ys zs A L R W.
sb = ys @ Write⇩s⇩b True a' sop v A L R W # zs ∧
a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a' ≠ a) ∨
(∃A L R W ys zs. sb = ys @ Ghost⇩s⇩b A L R W# zs ∧ a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys)"
(is "?write ∨ ?ghst")
by simp
then show ?thesis
proof
assume ?write
then obtain sop' a'' v' ys zs A' L' R' W' where
sb: "sb = ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs" and
props: "a ∈ A'" "a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a'' ≠ a"
by auto
show ?thesis
using props False a_notin sb
apply -
apply (rule disjI1)
apply (rule_tac x=sop' in exI)
apply (rule_tac x=a'' in exI)
apply (rule_tac x=v' in exI)
apply (rule_tac x="(x#ys)" in exI)
apply (rule_tac x=zs in exI)
apply (simp add: Write⇩s⇩b False )
done
next
assume ?ghst
then obtain ys zs A' L' R' W' where
sb: "sb = ys @ Ghost⇩s⇩b A' L' R' W' # zs" and
props: "a ∈ A'" "a ∉ outstanding_refs is_Write⇩s⇩b ys"
by auto
show ?thesis
using props False a_notin sb
apply -
apply (rule disjI2)
apply (rule_tac x=A' in exI)
apply (rule_tac x=L' in exI)
apply (rule_tac x=R' in exI)
apply (rule_tac x=W' in exI)
apply (rule_tac x="(x#ys)" in exI)
apply (rule_tac x=zs in exI)
apply (simp add: Write⇩s⇩b False )
done
qed
qed
next
case (Read⇩s⇩b volatile a' t v)
from Cons.prems
obtain
consis': "sharing_consistent 𝒮 𝒪 sb" and
a_nro': "a ∉ read_only 𝒮" and
a_unowned: "a ∉ 𝒪" and
nvo': "non_volatile_owned_or_read_only True 𝒮 𝒪 sb" and
a_unforw: "a ∈ unforwarded_non_volatile_reads sb X"
by (auto simp add: Read⇩s⇩b split: if_split_asm)
from Cons.hyps [OF nvo' consis' a_nro' a_unowned a_unforw]
have "(∃sop a' v ys zs A L R W.
sb = ys @ Write⇩s⇩b True a' sop v A L R W # zs ∧
a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a' ≠ a) ∨
(∃A L R W ys zs. sb = ys @ Ghost⇩s⇩b A L R W# zs ∧ a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys)"
(is "?write ∨ ?ghst")
by simp
then show ?thesis
proof
assume ?write
then obtain sop' a'' v' ys zs A' L' R' W' where
sb: "sb = ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs" and
props: "a ∈ A'" "a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a'' ≠ a"
by auto
show ?thesis
using props sb
apply -
apply (rule disjI1)
apply (rule_tac x=sop' in exI)
apply (rule_tac x=a'' in exI)
apply (rule_tac x=v' in exI)
apply (rule_tac x="(x#ys)" in exI)
apply (rule_tac x=zs in exI)
apply (simp add: Read⇩s⇩b)
done
next
assume ?ghst
then obtain ys zs A' L' R' W' where
sb: "sb = ys @ Ghost⇩s⇩b A' L' R' W'# zs" and
props: "a ∈ A'" "a ∉ outstanding_refs is_Write⇩s⇩b ys"
by auto
show ?thesis
using props sb
apply -
apply (rule disjI2)
apply (rule_tac x=A' in exI)
apply (rule_tac x=L' in exI)
apply (rule_tac x=R' in exI)
apply (rule_tac x=W' in exI)
apply (rule_tac x="(x#ys)" in exI)
apply (rule_tac x=zs in exI)
apply (simp add: Read⇩s⇩b )
done
qed
next
case Prog⇩s⇩b
from Cons.prems
obtain
consis': "sharing_consistent 𝒮 𝒪 sb" and
a_nro': "a ∉ read_only 𝒮" and
a_unowned: "a ∉ 𝒪" and
nvo': "non_volatile_owned_or_read_only True 𝒮 𝒪 sb" and
a_unforw: "a ∈ unforwarded_non_volatile_reads sb X"
by (auto simp add: Prog⇩s⇩b)
from Cons.hyps [OF nvo' consis' a_nro' a_unowned a_unforw]
have "(∃sop a' v ys zs A L R W.
sb = ys @ Write⇩s⇩b True a' sop v A L R W # zs ∧
a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a' ≠ a) ∨
(∃A L R W ys zs. sb = ys @ Ghost⇩s⇩b A L R W# zs ∧ a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys)"
(is "?write ∨ ?ghst")
by simp
then show ?thesis
proof
assume ?write
then obtain sop' a'' v' ys zs A' L' R' W' where
sb: "sb = ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs" and
props: "a ∈ A'" "a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a'' ≠ a"
by auto
show ?thesis
using props sb
apply -
apply (rule disjI1)
apply (rule_tac x=sop' in exI)
apply (rule_tac x=a'' in exI)
apply (rule_tac x=v' in exI)
apply (rule_tac x="(x#ys)" in exI)
apply (rule_tac x=zs in exI)
apply (simp add: Prog⇩s⇩b)
done
next
assume ?ghst
then obtain ys zs A' L' R' W' where
sb: "sb = ys @ Ghost⇩s⇩b A' L' R' W'# zs" and
props: "a ∈ A'" "a ∉ outstanding_refs is_Write⇩s⇩b ys"
by auto
show ?thesis
using props sb
apply -
apply (rule disjI2)
apply (rule_tac x=A' in exI)
apply (rule_tac x=L' in exI)
apply (rule_tac x=R' in exI)
apply (rule_tac x=W' in exI)
apply (rule_tac x="(x#ys)" in exI)
apply (rule_tac x=zs in exI)
apply (simp add: Prog⇩s⇩b )
done
qed
next
case (Ghost⇩s⇩b A L R W)
from Cons.prems obtain
nvo': "non_volatile_owned_or_read_only True (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) (𝒪 ∪ A - R) sb" and
a_nro: "a ∉ read_only 𝒮" and
a_unowned: "a ∉ 𝒪" and
A_shared_owns: "A ⊆ dom 𝒮 ∪ 𝒪" and L_A: "L ⊆ A" and A_R: "A ∩ R = {}" and
R_owns: "R ⊆ 𝒪" and
consis': "sharing_consistent (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) (𝒪 ∪ A - R) sb" and
a_unforw: "a ∈ unforwarded_non_volatile_reads sb X"
by (clarsimp simp add: Ghost⇩s⇩b)
show ?thesis
proof (cases "a ∈ A")
case True
then show ?thesis
apply -
apply (rule disjI2)
apply (rule_tac x=A in exI)
apply (rule_tac x=L in exI)
apply (rule_tac x=R in exI)
apply (rule_tac x=W in exI)
apply (rule_tac x="[]" in exI)
apply (rule_tac x=sb in exI)
apply (simp add: Ghost⇩s⇩b True)
done
next
case False
with a_unowned a_nro L_A R_owns a_nro L_A A_R
obtain a_nro': "a ∉ read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)" and a_unowned': "a ∉ 𝒪 ∪ A - R"
by (force simp add: in_read_only_convs)
from Cons.hyps [OF nvo' consis' a_nro' a_unowned' a_unforw]
have "(∃sop a' v ys zs A L R W.
sb = ys @ Write⇩s⇩b True a' sop v A L R W # zs ∧
a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a' ≠ a) ∨
(∃A L R W ys zs. sb = ys @ Ghost⇩s⇩b A L R W# zs ∧ a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys)"
(is "?write ∨ ?ghst")
by simp
then show ?thesis
proof
assume ?write
then obtain sop' a'' v' ys zs A' L' R' W' where
sb: "sb = ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs" and
props: "a ∈ A'" "a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a'' ≠ a"
by auto
show ?thesis
using props sb
apply -
apply (rule disjI1)
apply (rule_tac x=sop' in exI)
apply (rule_tac x=a'' in exI)
apply (rule_tac x=v' in exI)
apply (rule_tac x="(x#ys)" in exI)
apply (rule_tac x=zs in exI)
apply (simp add: Ghost⇩s⇩b False )
done
next
assume ?ghst
then obtain ys zs A' L' R' W' where
sb: "sb = ys @ Ghost⇩s⇩b A' L' R' W'# zs" and
props: "a ∈ A'" "a ∉ outstanding_refs is_Write⇩s⇩b ys"
by auto
show ?thesis
using props sb
apply -
apply (rule disjI2)
apply (rule_tac x=A' in exI)
apply (rule_tac x=L' in exI)
apply (rule_tac x=R' in exI)
apply (rule_tac x=W' in exI)
apply (rule_tac x="(x#ys)" in exI)
apply (rule_tac x=zs in exI)
apply (simp add: Ghost⇩s⇩b False )
done
qed
qed
qed
qed
lemma release_shared_exchange_weak:
assumes shared_eq: "∀a ∈ 𝒪 ∪ all_acquired sb. (𝒮'::shared) a = 𝒮 a"
assumes consis: "weak_sharing_consistent 𝒪 sb"
shows "release sb (dom 𝒮') ℛ = release sb (dom 𝒮) ℛ"
using shared_eq consis
proof (induct sb arbitrary: 𝒮 𝒮' 𝒪 ℛ)
case Nil thus ?case by auto
next
case (Cons x sb)
show ?case
proof (cases x)
case (Write⇩s⇩b volatile a sop v A L R W)
show ?thesis
proof (cases volatile)
case True
from Cons.prems obtain
L_A: "L ⊆ A" and A_R: "A ∩ R = {}" and R_owns: "R ⊆ 𝒪" and
consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb" and
shared_eq: "∀a ∈ 𝒪 ∪ A ∪ all_acquired sb. 𝒮' a = 𝒮 a"
by (clarsimp simp add: Write⇩s⇩b True )
from shared_eq
have shared_eq': "∀a∈𝒪 ∪ A - R ∪ all_acquired sb. (𝒮' ⊕⇘W⇙ R ⊖⇘A⇙ L) a = (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) a"
by (auto simp add: augment_shared_def restrict_shared_def)
from Cons.hyps [OF shared_eq' consis']
have "release sb (dom (𝒮' ⊕⇘W⇙ R ⊖⇘A⇙ L)) Map.empty = release sb (dom (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)) Map.empty" .
then show ?thesis
by (auto simp add: Write⇩s⇩b True domIff)
next
case False with Cons show ?thesis
by (auto simp add: Write⇩s⇩b)
qed
next
case Read⇩s⇩b with Cons show ?thesis
by auto
next
case Prog⇩s⇩b with Cons show ?thesis
by auto
next
case (Ghost⇩s⇩b A L R W)
from Cons.prems obtain
L_A: "L ⊆ A" and A_R: "A ∩ R = {}" and R_owns: "R ⊆ 𝒪" and
consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb" and
shared_eq: "∀a ∈ 𝒪 ∪ A ∪ all_acquired sb. 𝒮' a = 𝒮 a"
by (clarsimp simp add: Ghost⇩s⇩b )
from shared_eq
have shared_eq': "∀a∈𝒪 ∪ A - R ∪ all_acquired sb. (𝒮' ⊕⇘W⇙ R ⊖⇘A⇙ L) a = (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) a"
by (auto simp add: augment_shared_def restrict_shared_def)
from shared_eq R_owns have "∀a∈R. (a ∈ dom 𝒮) = (a ∈ dom 𝒮')"
by (auto simp add: domIff)
from augment_rels_shared_exchange [OF this]
have "(augment_rels (dom 𝒮') R ℛ) = (augment_rels (dom 𝒮) R ℛ)".
with Cons.hyps [OF shared_eq' consis']
have "release sb (dom (𝒮' ⊕⇘W⇙ R ⊖⇘A⇙ L)) (augment_rels (dom 𝒮') R ℛ) =
release sb (dom (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)) (augment_rels (dom 𝒮) R ℛ)" by simp
then show ?thesis
by (clarsimp simp add: Ghost⇩s⇩b domIff)
qed
qed
lemma read_only_share_all_shared: "⋀𝒮. ⟦ a ∈ read_only (share sb 𝒮)⟧
⟹ a ∈ read_only 𝒮 ∪ all_shared sb"
proof (induct sb)
case Nil thus ?case by simp
next
case (Cons x sb)
show ?case
proof (cases x)
case (Write⇩s⇩b volatile a sop v A L R W)
show ?thesis
proof (cases volatile)
case True
with Write⇩s⇩b Cons.hyps [of "(𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"] Cons.prems
show ?thesis
by (auto simp add: read_only_def augment_shared_def restrict_shared_def
split: if_split_asm option.splits)
next
case False with Write⇩s⇩b Cons show ?thesis by auto
qed
next
case Read⇩s⇩b with Cons show ?thesis by auto
next
case Prog⇩s⇩b with Cons show ?thesis by auto
next
case (Ghost⇩s⇩b A L R W)
with Cons.hyps [of "(𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"] Cons.prems
show ?thesis
by (auto simp add: read_only_def augment_shared_def restrict_shared_def
split: if_split_asm option.splits)
qed
qed
lemma read_only_shared_all_until_volatile_write_subset':
"⋀𝒮.
read_only (share_all_until_volatile_write ts 𝒮) ⊆
read_only 𝒮 ∪ (⋃((λ(_, _, _, sb, _, _ ,_). all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)) ` set ts))"
proof (induct ts)
case Nil thus ?case by simp
next
case (Cons t ts)
obtain p "is" 𝒪 ℛ 𝒟 θ sb where
t: "t = (p,is,θ,sb,𝒟,𝒪,ℛ)"
by (cases t)
have aargh: "(Not ∘ is_volatile_Write⇩s⇩b) = (λa. ¬ is_volatile_Write⇩s⇩b a)"
by (rule ext) auto
let ?take_sb = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
let ?drop_sb = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
{
fix a
assume a_in: "a ∈ read_only
(share_all_until_volatile_write ts
(share ?take_sb 𝒮))" and
a_notin_shared: "a ∉ read_only 𝒮" and
a_notin_rest: "a ∉ (⋃((λ(_, _, _, sb, _, _ ,_). all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)) ` set ts))"
have "a ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
proof -
from Cons.hyps [of "(share ?take_sb 𝒮)"] a_in a_notin_rest
have "a ∈ read_only (share ?take_sb 𝒮)"
by (auto simp add: aargh)
from read_only_share_all_shared [OF this] a_notin_shared
show ?thesis by auto
qed
}
then show ?case
by (auto simp add: t aargh)
qed
lemma read_only_share_acquired_all_shared:
"⋀𝒪 𝒮. weak_sharing_consistent 𝒪 sb ⟹ 𝒪 ∩ read_only 𝒮 = {} ⟹
a ∈ read_only (share sb 𝒮) ⟹ a ∈ 𝒪 ∪ all_acquired sb ⟹ a ∈ all_shared sb"
proof (induct sb)
case Nil thus ?case by auto
next
case (Cons x sb)
show ?case
proof (cases x)
case (Write⇩s⇩b volatile a' sop v A L R W)
show ?thesis
proof (cases volatile)
case True
note volatile=this
from Cons.prems obtain
owns_ro: "𝒪 ∩ read_only 𝒮 = {}" and L_A: " L ⊆ A" and A_R: "A ∩ R = {}" and
R_owns: "R ⊆ 𝒪" and consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb" and
a_share: "a ∈ read_only (share sb (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))" and
a_A_all: "a ∈ 𝒪 ∪ A ∪ all_acquired sb"
by (clarsimp simp add: Write⇩s⇩b True)
from owns_ro A_R R_owns have owns_ro': "(𝒪 ∪ A - R) ∩ read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by (auto simp add: in_read_only_convs)
from Cons.hyps [OF consis' owns_ro' a_share]
show ?thesis
using L_A A_R R_owns owns_ro a_A_all
by (auto simp add: Write⇩s⇩b volatile augment_shared_def restrict_shared_def read_only_def domIff
split: if_split_asm)
next
case False
with Cons Write⇩s⇩b show ?thesis by (auto)
qed
next
case Read⇩s⇩b with Cons show ?thesis by auto
next
case Prog⇩s⇩b with Cons show ?thesis by auto
next
case (Ghost⇩s⇩b A L R W)
from Cons.prems obtain
owns_ro: "𝒪 ∩ read_only 𝒮 = {}" and L_A: " L ⊆ A" and A_R: "A ∩ R = {}" and
R_owns: "R ⊆ 𝒪" and consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb" and
a_share: "a ∈ read_only (share sb (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))" and
a_A_all: "a ∈ 𝒪 ∪ A ∪ all_acquired sb"
by (clarsimp simp add: Ghost⇩s⇩b)
from owns_ro A_R R_owns have owns_ro': "(𝒪 ∪ A - R) ∩ read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by (auto simp add: in_read_only_convs)
from Cons.hyps [OF consis' owns_ro' a_share]
show ?thesis
using L_A A_R R_owns owns_ro a_A_all
by (auto simp add: Ghost⇩s⇩b augment_shared_def restrict_shared_def read_only_def domIff
split: if_split_asm)
qed
qed
lemma read_only_share_unowned': "⋀𝒪 𝒮.
⟦weak_sharing_consistent 𝒪 sb; 𝒪 ∩ read_only 𝒮 = {}; a ∉ 𝒪 ∪ all_acquired sb; a ∈ read_only 𝒮⟧
⟹ a ∈ read_only (share sb 𝒮)"
proof (induct sb)
case Nil thus ?case by simp
next
case (Cons x sb)
show ?case
proof (cases x)
case (Write⇩s⇩b volatile a' sop v A L R W)
show ?thesis
proof (cases volatile)
case False
with Cons Write⇩s⇩b show ?thesis by auto
next
case True
from Cons.prems obtain
owns_ro: "𝒪 ∩ read_only 𝒮 = {}" and L_A: " L ⊆ A" and A_R: "A ∩ R = {}" and
R_owns: "R ⊆ 𝒪" and consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb" and
a_share: "a ∈ read_only 𝒮" and
a_notin: "a ∉ 𝒪" "a ∉ A" "a ∉ all_acquired sb"
by (clarsimp simp add: Write⇩s⇩b True)
from owns_ro A_R R_owns have owns_ro': "(𝒪 ∪ A - R) ∩ read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by (auto simp add: in_read_only_convs)
from a_notin have a_notin': "a ∉ 𝒪 ∪ A - R ∪ all_acquired sb"
by auto
from a_share a_notin L_A A_R R_owns have a_ro': "a ∈ read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: read_only_def restrict_shared_def augment_shared_def)
from Cons.hyps [OF consis' owns_ro' a_notin' a_ro']
have "a ∈ read_only (share sb (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))"
by auto
then show ?thesis
by (auto simp add: Write⇩s⇩b True)
qed
next
case Read⇩s⇩b with Cons show ?thesis by auto
next
case Prog⇩s⇩b with Cons show ?thesis by auto
next
case (Ghost⇩s⇩b A L R W)
from Cons.prems obtain
owns_ro: "𝒪 ∩ read_only 𝒮 = {}" and L_A: " L ⊆ A" and A_R: "A ∩ R = {}" and
R_owns: "R ⊆ 𝒪" and consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb" and
a_share: "a ∈ read_only 𝒮" and
a_notin: "a ∉ 𝒪" "a ∉ A" "a ∉ all_acquired sb"
by (clarsimp simp add: Ghost⇩s⇩b)
from owns_ro A_R R_owns have owns_ro': "(𝒪 ∪ A - R) ∩ read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by (auto simp add: in_read_only_convs)
from a_notin have a_notin': "a ∉ 𝒪 ∪ A - R ∪ all_acquired sb"
by auto
from a_share a_notin L_A A_R R_owns have a_ro': "a ∈ read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: read_only_def restrict_shared_def augment_shared_def)
from Cons.hyps [OF consis' owns_ro' a_notin' a_ro']
have "a ∈ read_only (share sb (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))"
by auto
then show ?thesis
by (auto simp add: Ghost⇩s⇩b)
qed
qed
lemma release_False_mono:
"⋀S ℛ. ℛ a = Some False ⟹ outstanding_refs is_volatile_Write⇩s⇩b sb = {} ⟹
release sb S ℛ a = Some False "
proof (induct sb)
case Nil thus ?case by simp
next
case (Cons x sb)
show ?case
proof (cases x)
case (Ghost⇩s⇩b A L R W)
have rels_a: "ℛ a = Some False" by fact
then have "(augment_rels S R ℛ) a = Some False"
by (auto simp add: augment_rels_def)
from Cons.hyps [where ℛ = "(augment_rels S R ℛ)", OF this] Cons.prems
show ?thesis
by (clarsimp simp add: Ghost⇩s⇩b)
next
case Write⇩s⇩b with Cons show ?thesis by auto
next
case Read⇩s⇩b with Cons show ?thesis by auto
next
case Prog⇩s⇩b with Cons show ?thesis by auto
qed
qed
lemma release_False_mono_take:
"⋀S ℛ. ℛ a = Some False ⟹ release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) S ℛ a = Some False "
proof (induct sb)
case Nil thus ?case by simp
next
case (Cons x sb)
show ?case
proof (cases x)
case (Ghost⇩s⇩b A L R W)
have rels_a: "ℛ a = Some False" by fact
then have "(augment_rels S R ℛ) a = Some False"
by (auto simp add: augment_rels_def)
from Cons.hyps [where ℛ = "(augment_rels S R ℛ)", OF this]
show ?thesis
by (clarsimp simp add: Ghost⇩s⇩b)
next
case Write⇩s⇩b with Cons show ?thesis by auto
next
case Read⇩s⇩b with Cons show ?thesis by auto
next
case Prog⇩s⇩b with Cons show ?thesis by auto
qed
qed
lemma shared_switch:
"⋀𝒮 𝒪. ⟦weak_sharing_consistent 𝒪 sb; read_only 𝒮 ∩ 𝒪 = {};
𝒮 a ≠ Some False; share sb 𝒮 a = Some False⟧
⟹ a ∈ 𝒪 ∪ all_acquired sb "
proof (induct sb)
case Nil thus ?case by (auto simp add: read_only_def)
next
case (Cons x sb)
have aargh: "(Not ∘ is_volatile_Write⇩s⇩b) = (λa. ¬ is_volatile_Write⇩s⇩b a)"
by (rule ext) auto
show ?case
proof (cases x)
case (Ghost⇩s⇩b A L R W)
from Cons.prems obtain
dist: "read_only 𝒮 ∩ 𝒪 = {}" and
share: "𝒮 a ≠ Some False" and
share': "share sb (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) a = Some False" and
L_A: "L ⊆ A" and A_R: "A ∩ R = {}" and R_owns: "R ⊆ 𝒪" and
consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb" by (clarsimp simp add: Ghost⇩s⇩b aargh)
from dist L_A A_R R_owns have dist': "read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) ∩ (𝒪 ∪ A - R)= {}"
by (auto simp add: in_read_only_convs)
show ?thesis
proof (cases "(𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) a = Some False")
case False
from Cons.hyps [OF consis' dist' this share']
show ?thesis by (auto simp add: Ghost⇩s⇩b)
next
case True
with share L_A A_R R_owns dist
have "a ∈ 𝒪 ∪ A"
by (cases "𝒮 a")
(auto simp add: augment_shared_def restrict_shared_def read_only_def split: if_split_asm )
thus ?thesis by (auto simp add: Ghost⇩s⇩b)
qed
next
case (Write⇩s⇩b volatile a' sop v A L R W)
show ?thesis
proof (cases volatile)
case True
note volatile=this
from Cons.prems obtain
dist: "read_only 𝒮 ∩ 𝒪 = {}" and
share: "𝒮 a ≠ Some False" and
share': "share sb (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) a = Some False" and
L_A: "L ⊆ A" and A_R: "A ∩ R = {}" and R_owns: "R ⊆ 𝒪" and
consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb" by (clarsimp simp add: Write⇩s⇩b True aargh)
from dist L_A A_R R_owns have dist': "read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) ∩ (𝒪 ∪ A - R)= {}"
by (auto simp add: in_read_only_convs)
show ?thesis
proof (cases "(𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) a = Some False")
case False
from Cons.hyps [OF consis' dist' this share']
show ?thesis by (auto simp add: Write⇩s⇩b True)
next
case True
with share L_A A_R R_owns dist
have "a ∈ 𝒪 ∪ A"
by (cases "𝒮 a")
(auto simp add: augment_shared_def restrict_shared_def read_only_def split: if_split_asm )
thus ?thesis by (auto simp add: Write⇩s⇩b volatile)
qed
next
case False
with Cons show ?thesis by (auto simp add: Write⇩s⇩b)
qed
next
case Read⇩s⇩b with Cons show ?thesis by (auto)
next
case Prog⇩s⇩b with Cons show ?thesis by (auto)
qed
qed
lemma shared_switch_release_False:
"⋀𝒮 ℛ. ⟦
outstanding_refs is_volatile_Write⇩s⇩b sb = {};
a ∉ dom 𝒮;
a ∈ dom (share sb 𝒮)⟧
⟹
release sb (dom 𝒮) ℛ a = Some False"
proof (induct sb)
case Nil thus ?case by (auto simp add: read_only_def)
next
case (Cons x sb)
have aargh: "(Not ∘ is_volatile_Write⇩s⇩b) = (λa. ¬ is_volatile_Write⇩s⇩b a)"
by (rule ext) auto
show ?case
proof (cases x)
case (Ghost⇩s⇩b A L R W)
from Cons.prems obtain
a_notin: "a ∉ dom 𝒮" and
share: "a ∈ dom (share sb (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))" and
out': "outstanding_refs is_volatile_Write⇩s⇩b sb = {}"
by (clarsimp simp add: Ghost⇩s⇩b aargh)
show ?thesis
proof (cases "a ∈ R")
case False
with a_notin have "a ∉ dom (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by auto
from Cons.hyps [OF out' this share]
show ?thesis
by (auto simp add: Ghost⇩s⇩b)
next
case True
with a_notin have "augment_rels (dom 𝒮) R ℛ a = Some False"
by (auto simp add: augment_rels_def split: option.splits)
from release_False_mono [of "augment_rels (dom 𝒮) R ℛ", OF this out']
show ?thesis
by (auto simp add: Ghost⇩s⇩b)
qed
next
case Write⇩s⇩b with Cons show ?thesis by (clarsimp split: if_split_asm)
next
case Read⇩s⇩b with Cons show ?thesis by auto
next
case Prog⇩s⇩b with Cons show ?thesis by auto
qed
qed
lemma release_not_unshared_no_write:
"⋀𝒮 ℛ. ⟦
outstanding_refs is_volatile_Write⇩s⇩b sb = {};
non_volatile_writes_unshared 𝒮 sb;
release sb (dom 𝒮) ℛ a ≠ Some False;
a ∈ dom (share sb 𝒮)⟧
⟹
a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b sb"
proof (induct sb)
case Nil thus ?case by (auto simp add: read_only_def)
next
case (Cons x sb)
have aargh: "(Not ∘ is_volatile_Write⇩s⇩b) = (λa. ¬ is_volatile_Write⇩s⇩b a)"
by (rule ext) auto
show ?case
proof (cases x)
case (Ghost⇩s⇩b A L R W)
from Cons.prems obtain
share: "a ∈ dom (share sb (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))" and
rel: "release sb
(dom (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)) (augment_rels (dom 𝒮) R ℛ) a ≠ Some False" and
out': "outstanding_refs is_volatile_Write⇩s⇩b sb = {}" and
nvu: "non_volatile_writes_unshared (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) sb"
by (clarsimp simp add: Ghost⇩s⇩b )
from Cons.hyps [OF out' nvu rel share]
show ?thesis by (auto simp add: Ghost⇩s⇩b)
next
case (Write⇩s⇩b volatile a' sop v A L R W)
show ?thesis
proof (cases volatile)
case True with Write⇩s⇩b Cons.prems have False by auto
thus ?thesis ..
next
case False
note not_vol = this
from Cons.prems obtain
rel: "release sb (dom 𝒮) ℛ a ≠ Some False" and
out': "outstanding_refs is_volatile_Write⇩s⇩b sb = {}" and
nvo: "non_volatile_writes_unshared 𝒮 sb" and
a'_not_dom: "a' ∉ dom 𝒮" and
a_dom: "a ∈ dom (share sb 𝒮)"
by (auto simp add: Write⇩s⇩b False)
from Cons.hyps [OF out' nvo rel a_dom]
have a_notin_rest: "a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b sb".
show ?thesis
proof (cases "a'=a")
case False with a_notin_rest
show ?thesis by (clarsimp simp add: Write⇩s⇩b not_vol )
next
case True
from shared_switch_release_False [OF out' a'_not_dom [simplified True] a_dom]
have "release sb (dom 𝒮) ℛ a = Some False".
with rel have False by simp
thus ?thesis ..
qed
qed
next
case Read⇩s⇩b with Cons show ?thesis by auto
next
case Prog⇩s⇩b with Cons show ?thesis by auto
qed
qed
corollary release_not_unshared_no_write_take:
assumes nvw: "non_volatile_writes_unshared 𝒮 (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
assumes rel: "release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) (dom 𝒮) ℛ a ≠ Some False"
assumes a_in: "a ∈ dom (share (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒮)"
shows
"a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
using release_not_unshared_no_write[OF takeWhile_not_vol_write_outstanding_refs [of sb] nvw rel a_in]
by simp
lemma read_only_unacquired_share':
"⋀S 𝒪. ⟦𝒪 ∩ read_only S = {}; weak_sharing_consistent 𝒪 sb; a ∈ read_only S;
a ∉ all_shared sb; a ∉ acquired True sb 𝒪 ⟧
⟹ a ∈ read_only (share sb S)"
proof (induct sb)
case Nil thus ?case by simp
next
case (Cons x sb)
show ?case
proof (cases x)
case (Write⇩s⇩b volatile a' sop v A L R W)
show ?thesis
proof (cases volatile)
case True
note volatile=this
from Cons.prems
obtain a_ro: "a ∈ read_only S" and a_R: "a ∉ R" and a_unsh: "a ∉ all_shared sb" and
owns_ro: "𝒪 ∩ read_only S = {}" and
L_A: "L ⊆ A" and A_R: "A ∩ R = {}" and R_owns: "R ⊆ 𝒪" and
consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb" and
a_notin: "a ∉ acquired True sb (𝒪 ∪ A - R)"
by (clarsimp simp add: Write⇩s⇩b True)
show ?thesis
proof (cases "a ∈ A")
case True
with a_R have "a ∈ 𝒪 ∪ A - R" by auto
from all_shared_acquired_in [OF this a_unsh]
have "a ∈ acquired True sb (𝒪 ∪ A - R)" by auto
with a_notin have False by auto
thus ?thesis ..
next
case False
from owns_ro A_R R_owns have owns_ro': "(𝒪 ∪ A - R) ∩ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by (auto simp add: in_read_only_convs)
from a_ro False owns_ro R_owns L_A have a_ro': "a ∈ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs)
from Cons.hyps [OF owns_ro' consis' a_ro' a_unsh a_notin]
show ?thesis
by (clarsimp simp add: Write⇩s⇩b True)
qed
next
case False
with Cons show ?thesis
by (clarsimp simp add: Write⇩s⇩b False)
qed
next
case Read⇩s⇩b with Cons show ?thesis by (clarsimp)
next
case Prog⇩s⇩b with Cons show ?thesis by (clarsimp)
next
case (Ghost⇩s⇩b A L R W)
from Cons.prems
obtain a_ro: "a ∈ read_only S" and a_R: "a ∉ R" and a_unsh: "a ∉ all_shared sb" and
owns_ro: "𝒪 ∩ read_only S = {}" and
L_A: "L ⊆ A" and A_R: "A ∩ R = {}" and R_owns: "R ⊆ 𝒪" and
consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb" and
a_notin: "a ∉ acquired True sb (𝒪 ∪ A - R)"
by (clarsimp simp add: Ghost⇩s⇩b)
show ?thesis
proof (cases "a ∈ A")
case True
with a_R have "a ∈ 𝒪 ∪ A - R" by auto
from all_shared_acquired_in [OF this a_unsh]
have "a ∈ acquired True sb (𝒪 ∪ A - R)" by auto
with a_notin have False by auto
thus ?thesis ..
next
case False
from owns_ro A_R R_owns have owns_ro': "(𝒪 ∪ A - R) ∩ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by (auto simp add: in_read_only_convs)
from a_ro False owns_ro R_owns L_A have a_ro': "a ∈ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs)
from Cons.hyps [OF owns_ro' consis' a_ro' a_unsh a_notin]
show ?thesis
by (clarsimp simp add: Ghost⇩s⇩b)
qed
qed
qed
lemma read_only_share_all_until_volatile_write_unacquired':
"⋀𝒮. ⟦ownership_distinct ts; read_only_unowned 𝒮 ts; weak_sharing_consis ts;
∀i < length ts. (let (_,_,_,sb,_,𝒪,ℛ) = ts!i in
a ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪 ∧
a ∉ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb
));
a ∈ read_only 𝒮⟧
⟹ a ∈ read_only (share_all_until_volatile_write ts 𝒮)"
proof (induct ts)
case Nil thus ?case by simp
next
case (Cons t ts)
obtain p "is" 𝒪 ℛ 𝒟 θ sb where
t: "t = (p,is,θ,sb,𝒟,𝒪,ℛ)"
by (cases t)
have dist: "ownership_distinct (t#ts)" by fact
then interpret ownership_distinct "t#ts" .
from ownership_distinct_tl [OF dist]
have dist': "ownership_distinct ts".
have aargh: "(Not ∘ is_volatile_Write⇩s⇩b) = (λa. ¬ is_volatile_Write⇩s⇩b a)"
by (rule ext) auto
have a_ro: "a ∈ read_only 𝒮" by fact
have ro_unowned: "read_only_unowned 𝒮 (t#ts)" by fact
then interpret read_only_unowned 𝒮 "t#ts" .
have consis: "weak_sharing_consis (t#ts)" by fact
then interpret weak_sharing_consis "t#ts" .
note consis' = weak_sharing_consis_tl [OF consis]
let ?take_sb = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
let ?drop_sb = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
from weak_sharing_consis [of 0] t
have consis_sb: "weak_sharing_consistent 𝒪 sb"
by force
with weak_sharing_consistent_append [of 𝒪 ?take_sb ?drop_sb]
have consis_take: "weak_sharing_consistent 𝒪 ?take_sb"
by auto
have ro_unowned': "read_only_unowned (share ?take_sb 𝒮) ts"
proof
fix j
fix p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j
assume j_bound: "j < length ts"
assume jth: "ts!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "𝒪⇩j ∩ read_only (share ?take_sb 𝒮) = {}"
proof -
{
fix a
assume a_owns: "a ∈ 𝒪⇩j"
assume a_ro: "a ∈ read_only (share ?take_sb 𝒮)"
have False
proof -
from ownership_distinct [of 0 "Suc j"] j_bound jth t
have dist: "(𝒪 ∪ all_acquired sb) ∩ (𝒪⇩j ∪ all_acquired sb⇩j) = {}"
by fastforce
from read_only_unowned [of "Suc j"] j_bound jth
have dist_ro: "𝒪⇩j ∩ read_only 𝒮 = {}" by force
show ?thesis
proof (cases "a ∈ (𝒪 ∪ all_acquired sb)")
case True
with dist a_owns show False by auto
next
case False
hence "a ∉ (𝒪 ∪ all_acquired ?take_sb)"
using all_acquired_append [of ?take_sb ?drop_sb]
by auto
from read_only_share_unowned [OF consis_take this a_ro]
have "a ∈ read_only 𝒮".
with dist_ro a_owns show False by auto
qed
qed
}
thus ?thesis by auto
qed
qed
from Cons.prems
obtain unacq_ts: "∀i < length ts. (let (_,_,_,sb,_,𝒪,_) = ts!i in
a ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪 ∧
a ∉ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)) " and
unacq_sb: "a ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪" and
unsh_sb: "a ∉ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) "
apply clarsimp
apply (rule that)
apply (auto simp add: t aargh)
done
from read_only_unowned [of 0] t
have owns_ro: "𝒪 ∩ read_only 𝒮 = {}"
by force
from read_only_unacquired_share' [OF owns_ro consis_take a_ro unsh_sb unacq_sb]
have "a ∈ read_only (share (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒮)".
from Cons.hyps [OF dist' ro_unowned' consis' unacq_ts this]
show ?case
by (simp add: t)
qed
lemma not_shared_not_acquired_switch:
"⋀X Y. ⟦a ∉ all_shared sb; a ∉ X; a ∉ acquired True sb X; a ∉ Y⟧ ⟹ a ∉ acquired True sb Y"
proof (induct sb)
case Nil thus ?case by simp
next
case (Cons x sb)
show ?case
proof (cases x)
case (Write⇩s⇩b volatile a' sop v A L R W)
show ?thesis
proof (cases volatile)
case True
from Cons.prems obtain
a_X: "a ∉ X" and a_acq: "a ∉ acquired True sb (X ∪ A - R)" and
a_Y: "a ∉ Y" and a_R: "a ∉ R" and
a_shared: "a ∉ all_shared sb"
by (clarsimp simp add: Write⇩s⇩b True)
show ?thesis
proof (cases "a ∈ A")
case True
with a_X a_R
have "a ∈ X ∪ A - R" by auto
from all_shared_acquired_in [OF this a_shared]
have "a ∈ acquired True sb (X ∪ A - R)".
with a_acq have False by simp
thus ?thesis ..
next
case False
with a_X a_Y obtain a_X': "a ∉ X ∪ A - R" and a_Y': "a ∉ Y ∪ A - R"
by auto
from Cons.hyps [OF a_shared a_X' a_acq a_Y']
show ?thesis
by (auto simp add: Write⇩s⇩b True)
qed
next
case False with Cons.hyps [of X Y] Cons.prems show ?thesis by (auto simp add: Write⇩s⇩b)
qed
next
case Read⇩s⇩b with Cons.hyps [of X Y] Cons.prems show ?thesis by (auto)
next
case Prog⇩s⇩b with Cons.hyps [of X Y] Cons.prems show ?thesis by (auto)
next
case (Ghost⇩s⇩b A L R W)
from Cons.prems obtain
a_X: "a ∉ X" and a_acq: "a ∉ acquired True sb (X ∪ A - R)" and
a_Y: "a ∉ Y" and a_R: "a ∉ R" and
a_shared: "a ∉ all_shared sb"
by (clarsimp simp add: Ghost⇩s⇩b)
show ?thesis
proof (cases "a ∈ A")
case True
with a_X a_R
have "a ∈ X ∪ A - R" by auto
from all_shared_acquired_in [OF this a_shared]
have "a ∈ acquired True sb (X ∪ A - R)".
with a_acq have False by simp
thus ?thesis ..
next
case False
with a_X a_Y obtain a_X': "a ∉ X ∪ A - R" and a_Y': "a ∉ Y ∪ A - R"
by auto
from Cons.hyps [OF a_shared a_X' a_acq a_Y']
show ?thesis
by (auto simp add: Ghost⇩s⇩b)
qed
qed
qed
lemma read_only_share_all_acquired_in':
"⋀S 𝒪. ⟦𝒪 ∩ read_only S = {}; weak_sharing_consistent 𝒪 sb; a ∈ read_only (share sb S)⟧
⟹ a ∈ read_only (share sb Map.empty) ∨ (a ∈ read_only S ∧ a ∉ acquired True sb 𝒪 ∧ a ∉ all_shared sb )"
proof (induct sb)
case Nil thus ?case by auto
next
case (Cons x sb)
show ?case
proof (cases x)
case (Write⇩s⇩b volatile a' sop v A L R W)
show ?thesis
proof (cases volatile)
case True
note volatile=this
from Cons.prems
obtain a_in: "a ∈ read_only (share sb (S ⊕⇘W⇙ R ⊖⇘A⇙ L))" and
owns_ro: "𝒪 ∩ read_only S = {}" and
L_A: "L ⊆ A" and A_R: "A ∩ R = {}" and R_owns: "R ⊆ 𝒪" and
consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb"
by (clarsimp simp add: Write⇩s⇩b True)
from owns_ro A_R R_owns have owns_ro': "(𝒪 ∪ A - R) ∩ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by (auto simp add: in_read_only_convs)
from Cons.hyps [OF owns_ro' consis' a_in]
have hyp: "a ∈ read_only (share sb Map.empty) ∨
(a ∈ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L) ∧ a ∉ acquired True sb (𝒪 ∪ A - R) ∧ a ∉ all_shared sb)".
have "a ∈ read_only (share sb (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L)) ∨
(a ∈ read_only S ∧ a ∉ R ∧ a ∉ acquired True sb (𝒪 ∪ A - R) ∧ a ∉ all_shared sb)"
proof -
{
assume a_emp: "a ∈ read_only (share sb Map.empty)"
have "read_only Map.empty ⊆ read_only (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs)
from share_read_only_mono_in [OF a_emp this]
have "a ∈ read_only (share sb (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L))".
}
moreover
{
assume a_ro: "a ∈ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L)" and
a_not_acq: "a ∉ acquired True sb (𝒪 ∪ A - R)" and
a_unsh: "a ∉ all_shared sb"
have ?thesis
proof (cases "a ∈ read_only S")
case True
with a_ro obtain a_A: "a ∉ A"
by (auto simp add: in_read_only_convs)
with True a_not_acq a_unsh R_owns owns_ro
show ?thesis
by auto
next
case False
with a_ro have a_ro_empty: "a ∈ read_only (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs split: if_split_asm)
have "read_only (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L) ⊆ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs)
with owns_ro'
have owns_ro_empty: "(𝒪 ∪ A - R) ∩ read_only (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by blast
from read_only_unacquired_share' [OF owns_ro_empty consis' a_ro_empty a_unsh a_not_acq]
have "a ∈ read_only (share sb (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L))".
thus ?thesis
by simp
qed
}
moreover note hyp
ultimately show ?thesis by blast
qed
then show ?thesis
by (clarsimp simp add: Write⇩s⇩b True)
next
case False with Cons show ?thesis
by (auto simp add: Write⇩s⇩b)
qed
next
case Read⇩s⇩b with Cons show ?thesis by auto
next
case Prog⇩s⇩b with Cons show ?thesis by auto
next
case (Ghost⇩s⇩b A L R W)
from Cons.prems
obtain a_in: "a ∈ read_only (share sb (S ⊕⇘W⇙ R ⊖⇘A⇙ L))" and
owns_ro: "𝒪 ∩ read_only S = {}" and
L_A: "L ⊆ A" and A_R: "A ∩ R = {}" and R_owns: "R ⊆ 𝒪" and
consis': "weak_sharing_consistent (𝒪 ∪ A - R) sb"
by (clarsimp simp add: Ghost⇩s⇩b)
from owns_ro A_R R_owns have owns_ro': "(𝒪 ∪ A - R) ∩ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by (auto simp add: in_read_only_convs)
from Cons.hyps [OF owns_ro' consis' a_in]
have hyp: "a ∈ read_only (share sb Map.empty) ∨
(a ∈ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L) ∧ a ∉ acquired True sb (𝒪 ∪ A - R) ∧ a ∉ all_shared sb)".
have "a ∈ read_only (share sb (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L)) ∨
(a ∈ read_only S ∧ a ∉ R ∧ a ∉ acquired True sb (𝒪 ∪ A - R) ∧ a ∉ all_shared sb)"
proof -
{
assume a_emp: "a ∈ read_only (share sb Map.empty)"
have "read_only Map.empty ⊆ read_only (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs)
from share_read_only_mono_in [OF a_emp this]
have "a ∈ read_only (share sb (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L))".
}
moreover
{
assume a_ro: "a ∈ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L)" and
a_not_acq: "a ∉ acquired True sb (𝒪 ∪ A - R)" and
a_unsh: "a ∉ all_shared sb"
have ?thesis
proof (cases "a ∈ read_only S")
case True
with a_ro obtain a_A: "a ∉ A"
by (auto simp add: in_read_only_convs)
with True a_not_acq a_unsh R_owns owns_ro
show ?thesis
by auto
next
case False
with a_ro have a_ro_empty: "a ∈ read_only (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs split: if_split_asm)
have "read_only (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L) ⊆ read_only (S ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs)
with owns_ro'
have owns_ro_empty: "(𝒪 ∪ A - R) ∩ read_only (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by blast
from read_only_unacquired_share' [OF owns_ro_empty consis' a_ro_empty a_unsh a_not_acq]
have "a ∈ read_only (share sb (Map.empty ⊕⇘W⇙ R ⊖⇘A⇙ L))".
thus ?thesis
by simp
qed
}
moreover note hyp
ultimately show ?thesis by blast
qed
then show ?thesis
by (clarsimp simp add: Ghost⇩s⇩b)
qed
qed
lemma in_read_only_share_all_until_volatile_write':
assumes dist: "ownership_distinct ts"
assumes consis: "sharing_consis 𝒮 ts"
assumes ro_unowned: "read_only_unowned 𝒮 ts"
assumes i_bound: "i < length ts"
assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
assumes a_unacquired_others: "∀j < length ts. i≠j ⟶
(let (_,_,_,sb⇩j,_,𝒪,_) = ts!j in
a ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪 ∧
a ∉ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j ))"
assumes a_ro_share: "a ∈ read_only (share sb 𝒮)"
shows "a ∈ read_only (share (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)
(share_all_until_volatile_write ts 𝒮))"
proof -
from consis
interpret sharing_consis 𝒮 ts .
interpret read_only_unowned 𝒮 ts by fact
from sharing_consis [OF i_bound ts_i]
have consis_sb: "sharing_consistent 𝒮 𝒪 sb".
from sharing_consistent_weak_sharing_consistent [OF this]
have weak_consis: "weak_sharing_consistent 𝒪 sb".
from read_only_unowned [OF i_bound ts_i]
have owns_ro: "𝒪 ∩ read_only 𝒮 = {}".
from read_only_share_all_acquired_in' [OF owns_ro weak_consis a_ro_share]
have "a ∈ read_only (share sb Map.empty) ∨ a ∈ read_only 𝒮 ∧ a ∉ acquired True sb 𝒪 ∧ a ∉ all_shared sb".
moreover
let ?take_sb = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
let ?drop_sb = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
from weak_consis weak_sharing_consistent_append [of 𝒪 ?take_sb ?drop_sb]
obtain weak_consis': "weak_sharing_consistent (acquired True ?take_sb 𝒪) ?drop_sb" and
weak_consis_take: "weak_sharing_consistent 𝒪 ?take_sb"
by auto
{
assume "a ∈ read_only (share sb Map.empty)"
with share_append [of ?take_sb ?drop_sb]
have a_in': "a ∈ read_only (share ?drop_sb (share ?take_sb Map.empty))"
by auto
have owns_empty: "𝒪 ∩ read_only Map.empty = {}"
by auto
from weak_sharing_consistent_preserves_distinct [OF weak_consis_take owns_empty]
have "acquired True ?take_sb 𝒪 ∩ read_only (share ?take_sb Map.empty) = {}".
from read_only_share_all_acquired_in [OF this weak_consis' a_in']
have "a ∈ read_only (share ?drop_sb Map.empty) ∨ a ∈ read_only (share ?take_sb Map.empty) ∧ a ∉ all_acquired ?drop_sb".
moreover
{
assume a_ro_drop: "a ∈ read_only (share ?drop_sb Map.empty)"
have "read_only Map.empty ⊆ read_only (share_all_until_volatile_write ts 𝒮)"
by auto
from share_read_only_mono_in [OF a_ro_drop this]
have ?thesis .
}
moreover
{
assume a_ro_take: "a ∈ read_only (share ?take_sb Map.empty)"
assume a_unacq_drop: "a ∉ all_acquired ?drop_sb"
from read_only_share_unowned_in [OF weak_consis_take a_ro_take]
have "a ∈ 𝒪 ∪ all_acquired ?take_sb" by auto
hence "a ∈ 𝒪 ∪ all_acquired sb" using all_acquired_append [of ?take_sb ?drop_sb]
by auto
from share_all_until_volatile_write_thread_local' [OF dist consis i_bound ts_i this] a_ro_share
have ?thesis by (auto simp add: read_only_def)
}
ultimately have ?thesis by blast
}
moreover
{
assume a_ro: "a ∈ read_only 𝒮"
assume a_unacq: "a ∉ acquired True sb 𝒪"
assume a_unsh: "a ∉ all_shared sb"
with all_shared_append [of ?take_sb ?drop_sb]
obtain a_notin_take: "a ∉ all_shared ?take_sb" and a_notin_drop: "a ∉ all_shared ?drop_sb"
by auto
have ?thesis
proof (cases "a ∈ acquired True ?take_sb 𝒪")
case True
from all_shared_acquired_in [OF this a_notin_drop] acquired_append [of True ?take_sb ?drop_sb 𝒪] a_unacq
have False
by auto
thus ?thesis ..
next
case False
with a_unacquired_others i_bound ts_i a_notin_take
have a_unacq': "∀j < length ts.
(let (_,_,_,sb⇩j,_,𝒪,_) = ts!j in
a ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪 ∧
a ∉ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j ))"
by (auto simp add: Let_def)
from local.weak_sharing_consis_axioms have "weak_sharing_consis ts" .
from read_only_share_all_until_volatile_write_unacquired' [OF dist ro_unowned
‹weak_sharing_consis ts› a_unacq' a_ro]
have a_ro_all: "a ∈ read_only (share_all_until_volatile_write ts 𝒮)" .
from weak_consis weak_sharing_consistent_append [of 𝒪 ?take_sb ?drop_sb]
have weak_consis_drop: "weak_sharing_consistent (acquired True ?take_sb 𝒪) ?drop_sb"
by auto
from weak_sharing_consistent_preserves_distinct_share_all_until_volatile_write [OF dist
ro_unowned ‹weak_sharing_consis ts› i_bound ts_i]
have "acquired True ?take_sb 𝒪 ∩
read_only (share_all_until_volatile_write ts 𝒮) = {}".
from read_only_unacquired_share' [OF this weak_consis_drop a_ro_all a_notin_drop]
acquired_append [of True ?take_sb ?drop_sb 𝒪] a_unacq
show ?thesis by auto
qed
}
ultimately show ?thesis by blast
qed
lemma all_acquired_unshared_acquired:
"⋀𝒪. a ∈ all_acquired sb ==> a ∉ all_shared sb ==> a ∈ acquired True sb 𝒪"
apply (induct sb)
apply (auto split: memref.split intro: all_shared_acquired_in)
done
lemma safe_RMW_common:
assumes safe: "𝒪s,ℛs,i⊢ (RMW a t (D,f) cond ret A L R W# is, θ, m, 𝒟, 𝒪, 𝒮)√"
shows "(a ∈ 𝒪 ∨ a ∈ dom 𝒮) ∧ (∀j < length 𝒪s. i≠j ⟶ (ℛs!j) a ≠ Some False)"
using safe
apply (cases)
apply (auto simp add: domIff)
done
lemma acquired_reads_all_acquired': "⋀𝒪.
acquired_reads True sb 𝒪 ⊆ acquired True sb 𝒪 ∪ all_shared sb"
apply (induct sb)
apply clarsimp
apply (auto split: memref.splits dest: all_shared_acquired_in)
done
lemma release_all_shared_exchange:
"⋀ℛ S' S. ∀a ∈ all_shared sb. (a ∈ S') = (a ∈ S) ⟹ release sb S' ℛ = release sb S ℛ"
proof (induct sb)
case Nil thus ?case by auto
next
case (Cons x sb)
show ?case
proof (cases x)
case (Write⇩s⇩b volatile a' sop v A L R W)
show ?thesis
proof (cases volatile)
case True
note volatile=this
from Cons.hyps [of "(S' ∪ R - L)" "(S ∪ R - L)" Map.empty] Cons.prems
show ?thesis
by (auto simp add: Write⇩s⇩b volatile)
next
case False with Cons Write⇩s⇩b show ?thesis by auto
qed
next
case Read⇩s⇩b with Cons show ?thesis by auto
next
case Prog⇩s⇩b with Cons show ?thesis by auto
next
case (Ghost⇩s⇩b A L R W)
from augment_rels_shared_exchange [of R S S' ℛ] Cons.prems
have "augment_rels S' R ℛ = augment_rels S R ℛ"
by (auto simp add: Ghost⇩s⇩b)
with Cons.hyps [of "(S' ∪ R - L)" "(S ∪ R - L)" "augment_rels S R ℛ"] Cons.prems
show ?thesis
by (auto simp add: Ghost⇩s⇩b)
qed
qed
lemma release_append_Prog⇩s⇩b:
"⋀S ℛ. (release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb @ [Prog⇩s⇩b p⇩1 p⇩2 mis])) S ℛ) =
(release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) S ℛ) "
by (induct sb) (auto split: memref.splits)
subsection ‹Simulation of Store Buffer Machine with History by Virtual Machine with Delayed Releases›
theorem (in xvalid_program) concurrent_direct_steps_simulates_store_buffer_history_step:
assumes step_sb: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ⇒⇩s⇩b⇩h (ts⇩s⇩b',m⇩s⇩b',𝒮⇩s⇩b')"
assumes valid_own: "valid_ownership 𝒮⇩s⇩b ts⇩s⇩b"
assumes valid_sb_reads: "valid_reads m⇩s⇩b ts⇩s⇩b"
assumes valid_hist: "valid_history program_step ts⇩s⇩b"
assumes valid_sharing: "valid_sharing 𝒮⇩s⇩b ts⇩s⇩b"
assumes tmps_distinct: "tmps_distinct ts⇩s⇩b"
assumes valid_sops: "valid_sops ts⇩s⇩b"
assumes valid_dd: "valid_data_dependency ts⇩s⇩b"
assumes load_tmps_fresh: "load_tmps_fresh ts⇩s⇩b"
assumes enough_flushs: "enough_flushs ts⇩s⇩b"
assumes valid_program_history: "valid_program_history ts⇩s⇩b"
assumes valid: "valid ts⇩s⇩b"
assumes sim: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m,𝒮)"
assumes safe_reach: "safe_reach_direct safe_delayed (ts,m,𝒮)"
shows "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b' ∧ valid_reads m⇩s⇩b' ts⇩s⇩b' ∧ valid_history program_step ts⇩s⇩b' ∧
valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b' ∧ tmps_distinct ts⇩s⇩b' ∧ valid_data_dependency ts⇩s⇩b' ∧
valid_sops ts⇩s⇩b' ∧ load_tmps_fresh ts⇩s⇩b' ∧ enough_flushs ts⇩s⇩b' ∧
valid_program_history ts⇩s⇩b' ∧ valid ts⇩s⇩b' ∧
(∃ts' 𝒮' m'. (ts,m,𝒮) ⇒⇩d⇧* (ts',m',𝒮') ∧
(ts⇩s⇩b',m⇩s⇩b',𝒮⇩s⇩b') ∼ (ts',m',𝒮'))"
proof -
interpret direct_computation:
computation direct_memop_step empty_storebuffer_step program_step "λp p' is sb. sb" .
interpret sbh_computation:
computation sbh_memop_step flush_step program_step
"λp p' is sb. sb @ [Prog⇩s⇩b p p' is]" .
interpret valid_ownership 𝒮⇩s⇩b ts⇩s⇩b by fact
interpret valid_reads m⇩s⇩b ts⇩s⇩b by fact
interpret valid_history program_step ts⇩s⇩b by fact
interpret valid_sharing 𝒮⇩s⇩b ts⇩s⇩b by fact
interpret tmps_distinct ts⇩s⇩b by fact
interpret valid_sops ts⇩s⇩b by fact
interpret valid_data_dependency ts⇩s⇩b by fact
interpret load_tmps_fresh ts⇩s⇩b by fact
interpret enough_flushs ts⇩s⇩b by fact
interpret valid_program_history ts⇩s⇩b by fact
from valid_own valid_sharing
have valid_own_sharing: "valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b"
by (simp add: valid_sharing_def valid_ownership_and_sharing_def)
then
interpret valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b .
from safe_reach_safe_refl [OF safe_reach]
have safe: "safe_delayed (ts,m,𝒮)".
from step_sb
show ?thesis
proof (cases)
case (Memop i p⇩s⇩b "is⇩s⇩b" θ⇩s⇩b sb 𝒟⇩s⇩b 𝒪⇩s⇩b ℛ⇩s⇩b "is⇩s⇩b'" θ⇩s⇩b' sb' 𝒟⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b')
then obtain
ts⇩s⇩b': "ts⇩s⇩b' = ts⇩s⇩b[i := (p⇩s⇩b, is⇩s⇩b',θ⇩s⇩b', sb', 𝒟⇩s⇩b', 𝒪⇩s⇩b',ℛ⇩s⇩b')]" and
i_bound: "i < length ts⇩s⇩b" and
ts⇩s⇩b_i: "ts⇩s⇩b ! i = (p⇩s⇩b, is⇩s⇩b,θ⇩s⇩b,sb, 𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ⇩s⇩b)" and
sbh_step: "(is⇩s⇩b, θ⇩s⇩b, sb, m⇩s⇩b, 𝒟⇩s⇩b, 𝒪⇩s⇩b, ℛ⇩s⇩b,𝒮⇩s⇩b) →⇩s⇩b⇩h
(is⇩s⇩b', θ⇩s⇩b', sb', m⇩s⇩b', 𝒟⇩s⇩b', 𝒪⇩s⇩b', ℛ⇩s⇩b', 𝒮⇩s⇩b')"
by auto
from sim obtain
m: "m = flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b" and
𝒮: "𝒮 = share_all_until_volatile_write ts⇩s⇩b 𝒮⇩s⇩b" and
leq: "length ts⇩s⇩b = length ts" and
ts_sim: "∀i<length ts⇩s⇩b.
let (p, is⇩s⇩b, θ, sb, 𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ) = ts⇩s⇩b ! i;
suspends = dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb
in ∃is 𝒟. instrs suspends @ is⇩s⇩b = is @ prog_instrs suspends ∧
𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b sb ≠ {}) ∧
ts ! i =
(hd_prog p suspends,
is,
θ |` (dom θ - read_tmps suspends), (),
𝒟,
acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪⇩s⇩b,
release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) (dom 𝒮⇩s⇩b) ℛ)"
by cases blast
from i_bound leq have i_bound': "i < length ts"
by auto
have split_sb: "sb = takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb"
(is "sb = ?take_sb@?drop_sb")
by simp
from ts_sim [rule_format, OF i_bound] ts⇩s⇩b_i obtain suspends "is" 𝒟 where
suspends: "suspends = dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb" and
is_sim: "instrs suspends @ is⇩s⇩b = is @ prog_instrs suspends" and
𝒟: "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b sb ≠ {})" and
ts_i: "ts ! i =
(hd_prog p⇩s⇩b suspends, is,
θ⇩s⇩b |` (dom θ⇩s⇩b - read_tmps suspends), (), 𝒟, acquired True ?take_sb 𝒪⇩s⇩b,
release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by (auto simp add: Let_def)
from sbh_step_preserves_valid [OF i_bound ts⇩s⇩b_i sbh_step valid]
have valid': "valid ts⇩s⇩b'"
by (simp add: ts⇩s⇩b')
from 𝒟 have 𝒟⇩s⇩b: "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b ?drop_sb ≠ {})"
apply -
apply (case_tac "outstanding_refs is_volatile_Write⇩s⇩b sb = {}")
apply (fastforce simp add: outstanding_refs_conv dest: set_dropWhileD)
apply (clarsimp)
apply (drule outstanding_refs_non_empty_dropWhile)
apply blast
done
let ?ts' = "ts[i := (p⇩s⇩b, is⇩s⇩b, θ⇩s⇩b, (), 𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b,
release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)]"
have i_bound_ts': "i < length ?ts'"
using i_bound'
by auto
hence ts'_i: "?ts'!i = (p⇩s⇩b, is⇩s⇩b, θ⇩s⇩b, (),
𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
from local.sharing_consis_axioms
have sharing_consis_ts⇩s⇩b: "sharing_consis 𝒮⇩s⇩b ts⇩s⇩b" .
from sharing_consis [OF i_bound ts⇩s⇩b_i]
have sharing_consis_sb: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b sb".
from sharing_consistent_weak_sharing_consistent [OF this]
have weak_consis_sb: "weak_sharing_consistent 𝒪⇩s⇩b sb".
from this weak_sharing_consistent_append [of 𝒪⇩s⇩b ?take_sb ?drop_sb]
have weak_consis_drop:"weak_sharing_consistent (acquired True ?take_sb 𝒪⇩s⇩b) ?drop_sb"
by auto
from local.ownership_distinct_axioms
have ownership_distinct_ts⇩s⇩b: "ownership_distinct ts⇩s⇩b" .
have steps_flush_sb: "(ts,m,𝒮) ⇒⇩d⇧* (?ts', flush ?drop_sb m, share ?drop_sb 𝒮)"
proof -
from valid_reads [OF i_bound ts⇩s⇩b_i]
have reads_consis: "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b sb".
from reads_consistent_drop_volatile_writes_no_volatile_reads [OF this]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b ?drop_sb = {}".
from valid_program_history [OF i_bound ts⇩s⇩b_i]
have "causal_program_history is⇩s⇩b sb".
then have cph: "causal_program_history is⇩s⇩b ?drop_sb"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb"] )
apply (simp)
done
from valid_last_prog [OF i_bound ts⇩s⇩b_i] have last_prog: "last_prog p⇩s⇩b sb = p⇩s⇩b".
then
have lp: "last_prog p⇩s⇩b ?drop_sb = p⇩s⇩b"
apply -
apply (rule last_prog_same_append [where sb="?take_sb"])
apply simp
done
from reads_consistent_flush_all_until_volatile_write [OF valid_own_sharing i_bound
ts⇩s⇩b_i reads_consis]
have reads_consis_m: "reads_consistent True (acquired True ?take_sb 𝒪⇩s⇩b) m ?drop_sb"
by (simp add: m)
from valid_history [OF i_bound ts⇩s⇩b_i]
have h_consis: "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b (?take_sb@?drop_sb)) (?take_sb@?drop_sb)"
by (simp)
have last_prog_hd_prog: "last_prog (hd_prog p⇩s⇩b sb) ?take_sb = (hd_prog p⇩s⇩b ?drop_sb)"
proof -
from last_prog_hd_prog_append' [OF h_consis] last_prog
have "last_prog (hd_prog p⇩s⇩b ?drop_sb) ?take_sb = hd_prog p⇩s⇩b ?drop_sb"
by (simp)
moreover
have "last_prog (hd_prog p⇩s⇩b (?take_sb @ ?drop_sb)) ?take_sb =
last_prog (hd_prog p⇩s⇩b ?drop_sb) ?take_sb"
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp)
qed
from valid_write_sops [OF i_bound ts⇩s⇩b_i]
have "∀sop∈write_sops (?take_sb@?drop_sb). valid_sop sop"
by (simp)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops ?drop_sb. valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_read_tmps (?take_sb@?drop_sb)"
by (simp)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb ∩ read_tmps ?drop_sb = {}" and
distinct_read_tmps_drop: "distinct_read_tmps ?drop_sb"
by (simp only: distinct_read_tmps_append)
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop h_consis]
have hist_consis': "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b ?drop_sb) ?drop_sb"
by (simp add: last_prog_hd_prog)
have rel_eq: "release ?drop_sb (dom 𝒮) (release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b) =
release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b"
proof -
from release_append [of ?take_sb ?drop_sb]
have "release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b =
release ?drop_sb (dom (share ?take_sb 𝒮⇩s⇩b)) (release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
also
have dist: "ownership_distinct ts⇩s⇩b" by fact
have consis: "sharing_consis 𝒮⇩s⇩b ts⇩s⇩b" by fact
have "release ?drop_sb (dom (share ?take_sb 𝒮⇩s⇩b)) (release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b) =
release ?drop_sb (dom 𝒮) (release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b) "
apply (simp only: 𝒮)
apply (rule release_shared_exchange_weak [rule_format, OF _ weak_consis_drop])
apply (rule share_all_until_volatile_write_thread_local [OF dist consis i_bound ts⇩s⇩b_i, symmetric])
using acquired_all_acquired [of True ?take_sb 𝒪⇩s⇩b] all_acquired_append [of ?take_sb ?drop_sb]
by auto
finally
show ?thesis by simp
qed
from flush_store_buffer [OF i_bound' is_sim [simplified suspends]
cph ts_i [simplified suspends] refl lp reads_consis_m hist_consis'
valid_sops_drop distinct_read_tmps_drop no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], of 𝒮]
show ?thesis by (simp add: acquired_take_drop [where pending_write=True, simplified] 𝒟⇩s⇩b rel_eq)
qed
from safe_reach_safe_rtrancl [OF safe_reach steps_flush_sb]
have safe_ts': "safe_delayed (?ts', flush ?drop_sb m, share ?drop_sb 𝒮)".
from safe_delayedE [OF safe_ts' i_bound_ts' ts'_i]
have safe_memop_flush_sb: "map owned ?ts',map released ?ts',i⊢
(is⇩s⇩b, θ⇩s⇩b, flush ?drop_sb m, 𝒟⇩s⇩b,acquired True sb 𝒪⇩s⇩b,
share ?drop_sb 𝒮) √".
from acquired_takeWhile_non_volatile_Write⇩s⇩b
have acquired_take_sb: "acquired True ?take_sb 𝒪⇩s⇩b ⊆ 𝒪⇩s⇩b ∪ all_acquired ?take_sb ".
from sbh_step
show ?thesis
proof (cases)
case (SBHReadBuffered a v volatile t)
then obtain
"is⇩s⇩b": "is⇩s⇩b = Read volatile a t # is⇩s⇩b'" and
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b" and
𝒟⇩s⇩b': "𝒟⇩s⇩b'=𝒟⇩s⇩b" and
θ⇩s⇩b': "θ⇩s⇩b' = θ⇩s⇩b(t↦v)" and
sb': "sb'=sb@[Read⇩s⇩b volatile a t v]" and
m⇩s⇩b': "m⇩s⇩b' = m⇩s⇩b" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b" and
ℛ⇩s⇩b': "ℛ⇩s⇩b'=ℛ⇩s⇩b" and
buf_v: "buffered_val sb a = Some v"
by auto
from safe_memop_flush_sb [simplified is⇩s⇩b]
obtain access_cond': "a ∈ acquired True sb 𝒪⇩s⇩b ∨
a ∈ read_only (share ?drop_sb 𝒮) ∨
(volatile ∧ a ∈ dom (share ?drop_sb 𝒮))" and
volatile_clean: "volatile ⟶ ¬ 𝒟⇩s⇩b" and
rels_cond: "∀j < length ts. i≠j ⟶ released (ts!j) a ≠ Some False" and
rels_nv_cond: "¬volatile ⟶ (∀j < length ts. i≠j ⟶ a ∉ dom (released (ts!j)))"
by cases auto
from clean_no_outstanding_volatile_Write⇩s⇩b [OF i_bound ts⇩s⇩b_i] volatile_clean
have volatile_cond: "volatile ⟶ outstanding_refs is_volatile_Write⇩s⇩b sb ={}"
by auto
from buffered_val_witness [OF buf_v] obtain volatile' sop' A' L' R' W'
where
witness: "Write⇩s⇩b volatile' a sop' v A' L' R' W' ∈ set sb"
by auto
{
fix j p⇩j "is⇩s⇩b⇩j" 𝒪⇩j ℛ⇩j 𝒟⇩s⇩b⇩j θ⇩s⇩b⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume neq_i_j: "i ≠ j"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩s⇩b⇩j, θ⇩s⇩b⇩j, sb⇩j, 𝒟⇩s⇩b⇩j, 𝒪⇩j,ℛ⇩j)"
assume non_vol: "¬ volatile"
have "a ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
proof
assume a_j: "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
from ts_sim [rule_format, OF j_bound] jth
obtain suspends⇩j "is⇩j" 𝒟⇩j where
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by (auto simp add: Let_def)
from a_j ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
have a_notin_sb: "a ∉ 𝒪⇩s⇩b ∪ all_acquired sb"
by auto
with acquired_all_acquired [of True sb 𝒪⇩s⇩b]
have a_not_acq: "a ∉ acquired True sb 𝒪⇩s⇩b" by blast
with access_cond' non_vol
have a_ro: "a ∈ read_only (share ?drop_sb 𝒮)"
by auto
from read_only_share_unowned_in [OF weak_consis_drop a_ro] a_notin_sb
acquired_all_acquired [of True ?take_sb 𝒪⇩s⇩b]
all_acquired_append [of ?take_sb ?drop_sb]
have a_ro_shared: "a ∈ read_only 𝒮"
by auto
from rels_nv_cond [rule_format, OF non_vol j_bound [simplified leq] neq_i_j] ts⇩j
have "a ∉ dom (release ?take_sb⇩j (dom (𝒮⇩s⇩b)) ℛ⇩j)"
by auto
with dom_release_takeWhile [of sb⇩j "(dom (𝒮⇩s⇩b))" ℛ⇩j]
obtain
a_rels⇩j: "a ∉ dom ℛ⇩j" and
a_shared⇩j: "a ∉ all_shared ?take_sb⇩j"
by auto
have "a ∉ ⋃((λ(_, _, _, sb, _, _, _). all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)) `
set ts⇩s⇩b)"
proof -
{
fix k p⇩k "is⇩k" θ⇩k sb⇩k 𝒟⇩k 𝒪⇩k ℛ⇩k
assume k_bound: "k < length ts⇩s⇩b"
assume ts_k: "ts⇩s⇩b ! k = (p⇩k,is⇩k,θ⇩k,sb⇩k,𝒟⇩k,𝒪⇩k,ℛ⇩k)"
assume a_in: "a ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k)"
have False
proof (cases "k=j")
case True with a_shared⇩j jth ts_k a_in show False by auto
next
case False
from ownership_distinct [OF j_bound k_bound False [symmetric] jth ts_k] a_j
have "a ∉ (𝒪⇩k ∪ all_acquired sb⇩k)" by auto
with all_shared_acquired_or_owned [OF sharing_consis [OF k_bound ts_k]] a_in
show False
using all_acquired_append [of "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"
"dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"]
all_shared_append [of "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"
"dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"] by auto
qed
}
thus ?thesis by (fastforce simp add: in_set_conv_nth)
qed
with a_ro_shared
read_only_shared_all_until_volatile_write_subset' [of ts⇩s⇩b 𝒮⇩s⇩b]
have a_ro_shared⇩s⇩b: "a ∈ read_only 𝒮⇩s⇩b"
by (auto simp add: 𝒮)
with read_only_unowned [OF j_bound jth]
have a_notin_owns_j: "a ∉ 𝒪⇩j"
by auto
have own_dist: "ownership_distinct ts⇩s⇩b" by fact
have share_consis: "sharing_consis 𝒮⇩s⇩b ts⇩s⇩b" by fact
from sharing_consistent_share_all_until_volatile_write [OF own_dist share_consis i_bound ts⇩s⇩b_i]
have consis': "sharing_consistent 𝒮 (acquired True ?take_sb 𝒪⇩s⇩b) ?drop_sb"
by (simp add: 𝒮)
from share_all_until_volatile_write_thread_local [OF own_dist share_consis j_bound jth a_j] a_ro_shared
have a_ro_take: "a ∈ read_only (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto simp add: domIff 𝒮 read_only_def)
from sharing_consis [OF j_bound jth]
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
from sharing_consistent_weak_sharing_consistent [OF this] weak_sharing_consistent_append [of 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have weak_consis_drop:"weak_sharing_consistent 𝒪⇩j ?take_sb⇩j"
by auto
from read_only_share_acquired_all_shared [OF this read_only_unowned [OF j_bound jth] a_ro_take ] a_notin_owns_j a_shared⇩j
have "a ∉ all_acquired ?take_sb⇩j"
by auto
with a_j a_notin_owns_j
have a_drop: "a ∈ all_acquired ?drop_sb⇩j"
using all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
by simp
from i_bound j_bound leq have j_bound_ts': "j < length ?ts'"
by auto
note conflict_drop = a_drop [simplified suspends⇩j [symmetric]]
from split_all_acquired_in [OF conflict_drop]
show False
proof
assume "∃sop a' v ys zs A L R W.
(suspends⇩j = ys @ Write⇩s⇩b True a' sop v A L R W# zs) ∧ a ∈ A"
then
obtain a' sop' v' ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Write⇩s⇩b True a' sop' v' A' L' R' W'# zs"
(is "suspends⇩j = ?suspends") and
a_A': "a ∈ A'"
by blast
from sharing_consis [OF j_bound jth]
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
then have A'_R': "A' ∩ R' = {}"
by (simp add: sharing_consistent_append [of _ _ ?take_sb⇩j ?drop_sb⇩j, simplified]
suspends⇩j [symmetric] split_suspends⇩j sharing_consistent_append)
from valid_program_history [OF j_bound jth]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts'_j: "?ts'!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j, release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by auto
from valid_last_prog [OF j_bound jth] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j suspends⇩j = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound jth]
have reads_consis_j: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound
jth reads_consis_j]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m_j]
have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
hence reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W'])"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_write_sops [OF j_bound jth]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound jth]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound jth]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis_j]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b
(ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
have acq_simp:
"acquired True (ys @ [Write⇩s⇩b True a' sop' v' A' L' R' W'])
(acquired True ?take_sb⇩j 𝒪⇩j) =
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R'"
by (simp add: acquired_append)
from flush_store_buffer_append [where sb="ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']" and sb'="zs", simplified,
OF j_bound_ts' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts'_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="share ?drop_sb 𝒮"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs zs @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs" and
steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d⇧*
(?ts'[j:=(last_prog
(hd_prog p⇩j (Write⇩s⇩b True a' sop' v' A' L' R' W'# zs)) (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']),
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs),
(), True, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R',ℛ⇩j')],
flush (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) (flush ?drop_sb m),
share (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) (share ?drop_sb 𝒮))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append outstanding_refs_append)
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b, θ⇩s⇩b,(),
𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
note conflict_computation = rtranclp_trans [OF steps_flush_sb steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this i_bound_ys ts_ys_i, simplified is⇩s⇩b] non_vol a_not_acq
have "a ∈ read_only (share (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) (share ?drop_sb 𝒮))"
apply cases
apply (auto simp add: Let_def is⇩s⇩b)
done
with a_A'
show False
by (simp add: share_append in_read_only_convs)
next
assume "∃A L R W ys zs. suspends⇩j = ys @ Ghost⇩s⇩b A L R W # zs ∧ a ∈ A"
then
obtain A' L' R' W' ys zs where
split_suspends⇩j: "suspends⇩j = ys @ Ghost⇩s⇩b A' L' R' W'# zs"
(is "suspends⇩j = ?suspends") and
a_A': "a ∈ A'"
by blast
from valid_program_history [OF j_bound jth]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts'_j: "?ts'!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j, release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by auto
from valid_last_prog [OF j_bound jth] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j suspends⇩j = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound jth]
have reads_consis_j: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound
jth reads_consis_j]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m_j]
have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
hence reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) (ys@[Ghost⇩s⇩b A' L' R' W'])"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_write_sops [OF j_bound jth]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys@[Ghost⇩s⇩b A' L' R' W']). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound jth]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound jth]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis_j]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b
(ys@[Ghost⇩s⇩b A' L' R' W']) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
have acq_simp:
"acquired True (ys @ [Ghost⇩s⇩b A' L' R' W'])
(acquired True ?take_sb⇩j 𝒪⇩j) =
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R'"
by (simp add: acquired_append)
from flush_store_buffer_append [where sb="ys@[Ghost⇩s⇩b A' L' R' W']" and sb'="zs", simplified,
OF j_bound_ts' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts'_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="share ?drop_sb 𝒮"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs zs @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs" and
steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d⇧*
(?ts'[j:=(last_prog
(hd_prog p⇩j (Ghost⇩s⇩b A' L' R' W'# zs)) (ys@[Ghost⇩s⇩b A' L' R' W']),
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs),
(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b (ys @ [Ghost⇩s⇩b A' L' R' W']) ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R',ℛ⇩j')],
flush (ys@[Ghost⇩s⇩b A' L' R' W']) (flush ?drop_sb m),
share (ys@[Ghost⇩s⇩b A' L' R' W']) (share ?drop_sb 𝒮))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append)
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b,θ⇩s⇩b,(),
𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
note conflict_computation = rtranclp_trans [OF steps_flush_sb steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this i_bound_ys ts_ys_i, simplified is⇩s⇩b] non_vol a_not_acq
have "a ∈ read_only (share (ys@[Ghost⇩s⇩b A' L' R' W']) (share ?drop_sb 𝒮))"
apply cases
apply (auto simp add: Let_def is⇩s⇩b)
done
with a_A'
show False
by (simp add: share_append in_read_only_convs)
qed
qed
}
note non_volatile_unowned_others = this
{
assume a_in: "a ∈ read_only (share (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒮)"
assume nv: "¬ volatile"
have "a ∈ read_only (share sb 𝒮⇩s⇩b)"
proof (cases "a ∈ 𝒪⇩s⇩b ∪ all_acquired sb")
case True
from share_all_until_volatile_write_thread_local' [OF ownership_distinct_ts⇩s⇩b
sharing_consis_ts⇩s⇩b i_bound ts⇩s⇩b_i True] True a_in
show ?thesis
by (simp add: 𝒮 read_only_def)
next
case False
from read_only_share_unowned [OF weak_consis_drop _ a_in] False
acquired_all_acquired [of True ?take_sb 𝒪⇩s⇩b] all_acquired_append [of ?take_sb ?drop_sb]
have a_ro_shared: "a ∈ read_only 𝒮"
by auto
have "a ∉ ⋃((λ(_, _, _, sb, _, _, _).
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)) ` set ts⇩s⇩b)"
proof -
{
fix k p⇩k "is⇩k" θ⇩k sb⇩k 𝒟⇩k 𝒪⇩k ℛ⇩k
assume k_bound: "k < length ts⇩s⇩b"
assume ts_k: "ts⇩s⇩b ! k = (p⇩k,is⇩k,θ⇩k,sb⇩k,𝒟⇩k,𝒪⇩k,ℛ⇩k)"
assume a_in: "a ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k)"
have False
proof (cases "k=i")
case True with False ts⇩s⇩b_i ts_k a_in
all_shared_acquired_or_owned [OF sharing_consis [OF k_bound ts_k]]
all_shared_append [of "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"
"dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"] show False by auto
next
case False
from rels_nv_cond [rule_format, OF nv k_bound [simplified leq] False [symmetric] ]
ts_sim [rule_format, OF k_bound] ts_k
have "a ∉ dom (release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k) (dom (𝒮⇩s⇩b)) ℛ⇩k)"
by (auto simp add: Let_def)
with dom_release_takeWhile [of sb⇩k "(dom (𝒮⇩s⇩b))" ℛ⇩k]
obtain
a_rels⇩j: "a ∉ dom ℛ⇩k" and
a_shared⇩j: "a ∉ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k)"
by auto
with False a_in show ?thesis
by auto
qed
}
thus ?thesis by (fastforce simp add: in_set_conv_nth)
qed
with read_only_shared_all_until_volatile_write_subset' [of ts⇩s⇩b 𝒮⇩s⇩b] a_ro_shared
have "a ∈ read_only 𝒮⇩s⇩b"
by (auto simp add: 𝒮)
from read_only_share_unowned' [OF weak_consis_sb read_only_unowned [OF i_bound ts⇩s⇩b_i] False this]
show ?thesis .
qed
} note non_vol_ro_reduction = this
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (cases volatile)
case False
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b sb".
then
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b (sb@[Read⇩s⇩b False a t v])"
using access_cond' False non_vol_ro_reduction
by (auto simp add: non_volatile_owned_or_read_only_append)
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (auto simp add: False ts⇩s⇩b' sb' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
case True
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b sb".
then
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b (sb@[Read⇩s⇩b True a t v])"
using True
by (simp add: non_volatile_owned_or_read_only_append)
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (auto simp add: True ts⇩s⇩b' sb' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof -
have out: "outstanding_refs is_volatile_Write⇩s⇩b (sb @ [Read⇩s⇩b volatile a t v]) ⊆
outstanding_refs is_volatile_Write⇩s⇩b sb"
by (auto simp add: outstanding_refs_append)
have "all_acquired (sb @ [Read⇩s⇩b volatile a t v]) ⊆ all_acquired sb"
by (auto simp add: all_acquired_append)
from outstanding_volatile_writes_unowned_by_others_store_buffer
[OF i_bound ts⇩s⇩b_i out this]
show ?thesis by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b')
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof (cases volatile)
case True
have r: "read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb @ [Read⇩s⇩b volatile a t v])) 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb @ [Read⇩s⇩b volatile a t v]))
⊆ read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
apply (case_tac "outstanding_refs (is_volatile_Write⇩s⇩b) sb = {}")
apply (simp_all add: outstanding_vol_write_take_drop_appends
acquired_append read_only_reads_append True)
done
have "𝒪⇩s⇩b ∪ all_acquired (sb @ [Read⇩s⇩b volatile a t v]) ⊆ 𝒪⇩s⇩b ∪ all_acquired sb"
by (simp add: all_acquired_append)
from read_only_reads_unowned_nth_update [OF i_bound ts⇩s⇩b_i r this]
show ?thesis
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb')
next
case False
show ?thesis
proof (unfold_locales)
fix n m
fix p⇩n "is⇩n" 𝒪⇩n ℛ⇩n 𝒟⇩n θ⇩n sb⇩n p⇩m "is⇩m" 𝒪⇩m ℛ⇩m 𝒟⇩m θ⇩m sb⇩m
assume n_bound: "n < length ts⇩s⇩b'"
and m_bound: "m < length ts⇩s⇩b'"
and neq_n_m: "n≠m"
and nth: "ts⇩s⇩b'!n = (p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n, ℛ⇩n)"
and mth: "ts⇩s⇩b'!m =(p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m, ℛ⇩m)"
from n_bound have n_bound': "n < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
from m_bound have m_bound': "m < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
have acq_eq: "(𝒪⇩s⇩b' ∪ all_acquired sb') = (𝒪⇩s⇩b ∪ all_acquired sb)"
by (simp add: all_acquired_append sb' 𝒪⇩s⇩b')
show "(𝒪⇩m ∪ all_acquired sb⇩m) ∩
read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) 𝒪⇩n)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) =
{}"
proof (cases "m=i")
case True
with neq_n_m have neq_n_i: "n≠i"
by auto
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n, ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
note read_only_reads_unowned [OF n_bound' i_bound neq_n_i nth' ts⇩s⇩b_i]
moreover
note acq_eq
ultimately show ?thesis
using True ts⇩s⇩b_i nth mth n_bound' m_bound'
by (simp add: ts⇩s⇩b')
next
case False
note neq_m_i = this
with m_bound mth i_bound have mth': "ts⇩s⇩b!m = (p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
by (auto simp add: ts⇩s⇩b')
show ?thesis
proof (cases "n=i")
case True
note read_only_reads_unowned [OF i_bound m_bound' neq_m_i [symmetric] ts⇩s⇩b_i mth']
moreover
note acq_eq
moreover
note non_volatile_unowned_others [OF m_bound' neq_m_i [symmetric] mth']
ultimately show ?thesis
using True ts⇩s⇩b_i nth mth n_bound' m_bound' neq_m_i
apply (case_tac "outstanding_refs (is_volatile_Write⇩s⇩b) sb = {}")
apply (clarsimp simp add: outstanding_vol_write_take_drop_appends
acquired_append read_only_reads_append ts⇩s⇩b' sb' 𝒪⇩s⇩b')+
done
next
case False
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n, ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
from read_only_reads_unowned [OF n_bound' m_bound' neq_n_m nth' mth'] False neq_m_i
show ?thesis
by (clarsimp)
qed
qed
qed
qed
next
show "ownership_distinct ts⇩s⇩b'"
proof -
have "all_acquired (sb @ [Read⇩s⇩b volatile a t v]) ⊆ all_acquired sb"
by (auto simp add: all_acquired_append)
from ownership_distinct_instructions_read_value_store_buffer_independent
[OF i_bound ts⇩s⇩b_i this]
show ?thesis by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b')
qed
qed
have valid_hist': "valid_history program_step ts⇩s⇩b'"
proof -
from valid_history [OF i_bound ts⇩s⇩b_i]
have hcons: "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b sb) sb".
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have t_notin_reads: "t ∉ read_tmps sb"
by (auto simp add: "is⇩s⇩b")
from load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have t_notin_writes: "t ∉ ⋃(fst ` write_sops sb)"
by (auto simp add: "is⇩s⇩b")
from valid_write_sops [OF i_bound ts⇩s⇩b_i]
have valid_sops: "∀sop ∈ write_sops sb. valid_sop sop"
by auto
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have t_fresh: "t ∉ dom θ⇩s⇩b"
using "is⇩s⇩b"
by simp
have "history_consistent (θ⇩s⇩b(t↦v))
(hd_prog p⇩s⇩b (sb@ [Read⇩s⇩b volatile a t v])) (sb@ [Read⇩s⇩b volatile a t v])"
using t_notin_writes valid_sops t_fresh hcons
valid_implies_valid_prog_hd [OF i_bound ts⇩s⇩b_i valid]
apply -
apply (rule history_consistent_appendI)
apply (auto simp add: hd_prog_append_Read⇩s⇩b)
done
from valid_history_nth_update [OF i_bound this]
show ?thesis
by (auto simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b' θ⇩s⇩b')
qed
from reads_consistent_buffered_snoc [OF buf_v valid_reads [OF i_bound ts⇩s⇩b_i]
volatile_cond]
have reads_consis': "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b (sb @ [Read⇩s⇩b volatile a t v])"
by (simp split: if_split_asm)
from valid_reads_nth_update [OF i_bound this]
have valid_reads': "valid_reads m⇩s⇩b ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b')
have valid_sharing': "valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
from outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i]
have "non_volatile_writes_unshared 𝒮⇩s⇩b (sb @ [Read⇩s⇩b volatile a t v])"
by (auto simp add: non_volatile_writes_unshared_append)
from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
show "outstanding_non_volatile_writes_unshared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' 𝒮⇩s⇩b')
next
from sharing_consis [OF i_bound ts⇩s⇩b_i]
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b sb".
then
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b (sb @ [Read⇩s⇩b volatile a t v])"
by (simp add: sharing_consistent_append)
from sharing_consis_nth_update [OF i_bound this]
show "sharing_consis 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb' 𝒮⇩s⇩b')
next
note read_only_unowned [OF i_bound ts⇩s⇩b_i]
from read_only_unowned_nth_update [OF i_bound this]
show "read_only_unowned 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b')
next
from unowned_shared_nth_update [OF i_bound ts⇩s⇩b_i subset_refl]
show "unowned_shared 𝒮⇩s⇩b' ts⇩s⇩b'" by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
have "no_write_to_read_only_memory 𝒮⇩s⇩b sb".
hence "no_write_to_read_only_memory 𝒮⇩s⇩b (sb@[Read⇩s⇩b volatile a t v])"
by (simp add: no_write_to_read_only_memory_append)
from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
show "no_outstanding_write_to_read_only_memory 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒮⇩s⇩b' sb')
qed
have tmps_distinct': "tmps_distinct ts⇩s⇩b'"
proof (intro_locales)
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_load_tmps is⇩s⇩b'"
by (auto split: instr.splits simp add: is⇩s⇩b)
from load_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b')
next
from read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_read_tmps sb".
moreover
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "t ∉ read_tmps sb"
by (auto simp add: is⇩s⇩b)
ultimately have "distinct_read_tmps (sb @ [Read⇩s⇩b volatile a t v])"
by (auto simp add: distinct_read_tmps_append)
from read_tmps_distinct_nth_update [OF i_bound this]
show "read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb')
next
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ read_tmps (sb @ [Read⇩s⇩b volatile a t v]) = {}"
by (clarsimp simp add: read_tmps_append "is⇩s⇩b")
from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb')
qed
have valid_sops': "valid_sops ts⇩s⇩b'"
proof -
from valid_store_sops [OF i_bound ts⇩s⇩b_i]
have valid_store_sops': "∀sop∈store_sops is⇩s⇩b'. valid_sop sop"
by (auto simp add: "is⇩s⇩b")
from valid_write_sops [OF i_bound ts⇩s⇩b_i]
have valid_write_sops': "∀sop∈write_sops (sb@ [Read⇩s⇩b volatile a t v]). valid_sop sop"
by (auto simp add: write_sops_append)
from valid_sops_nth_update [OF i_bound valid_write_sops' valid_store_sops']
show ?thesis by (simp add: ts⇩s⇩b' sb')
qed
have valid_dd': "valid_data_dependency ts⇩s⇩b'"
proof -
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
have dd_is: "data_dependency_consistent_instrs (dom θ⇩s⇩b') is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b" θ⇩s⇩b')
from load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ ⋃(fst ` write_sops (sb@ [Read⇩s⇩b volatile a t v])) = {}"
by (auto simp add: write_sops_append "is⇩s⇩b")
from valid_data_dependency_nth_update [OF i_bound dd_is this]
show ?thesis by (simp add: ts⇩s⇩b' sb')
qed
have load_tmps_fresh': "load_tmps_fresh ts⇩s⇩b'"
proof -
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have "load_tmps (Read volatile a t # is⇩s⇩b') ∩ dom θ⇩s⇩b = {}"
by (simp add: "is⇩s⇩b")
moreover
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i] have "t ∉ load_tmps is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b")
ultimately have "load_tmps is⇩s⇩b' ∩ dom (θ⇩s⇩b(t ↦ v)) = {}"
by auto
from load_tmps_fresh_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' θ⇩s⇩b')
qed
have enough_flushs': "enough_flushs ts⇩s⇩b'"
proof -
from clean_no_outstanding_volatile_Write⇩s⇩b [OF i_bound ts⇩s⇩b_i]
have "¬ 𝒟⇩s⇩b ⟶ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Read⇩s⇩b volatile a t v]) = {}"
by (auto simp add: outstanding_refs_append )
from enough_flushs_nth_update [OF i_bound this]
show ?thesis
by (simp add: ts⇩s⇩b' sb' 𝒟⇩s⇩b')
qed
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
from valid_program_history [OF i_bound ts⇩s⇩b_i]
have "causal_program_history is⇩s⇩b sb" .
then have causal': "causal_program_history is⇩s⇩b' (sb@[Read⇩s⇩b volatile a t v])"
by (auto simp: causal_program_history_Read "is⇩s⇩b")
from valid_last_prog [OF i_bound ts⇩s⇩b_i]
have "last_prog p⇩s⇩b sb = p⇩s⇩b".
hence "last_prog p⇩s⇩b (sb @ [Read⇩s⇩b volatile a t v]) = p⇩s⇩b"
by (simp add: last_prog_append_Read⇩s⇩b)
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b' sb')
qed
show ?thesis
proof (cases "outstanding_refs is_volatile_Write⇩s⇩b sb = {}")
case True
from True have flush_all: "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = sb"
by (auto simp add: outstanding_refs_conv )
from True have suspend_nothing: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = []"
by (auto simp add: outstanding_refs_conv)
hence suspends_empty: "suspends = []"
by (simp add: suspends)
from suspends_empty is_sim have "is": "is = Read volatile a t # is⇩s⇩b'"
by (simp add: "is⇩s⇩b")
with suspends_empty ts_i
have ts_i: "ts!i = (p⇩s⇩b, Read volatile a t # is⇩s⇩b', θ⇩s⇩b,(), 𝒟, acquired True ?take_sb 𝒪⇩s⇩b, release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
from direct_memop_step.Read
have "(Read volatile a t # is⇩s⇩b', θ⇩s⇩b, (), m, 𝒟, acquired True ?take_sb 𝒪⇩s⇩b,
release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b, 𝒮) →
(is⇩s⇩b', θ⇩s⇩b(t ↦ m a), (), m, 𝒟, acquired True ?take_sb 𝒪⇩s⇩b,release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b, 𝒮)".
from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
have "(ts, m, 𝒮) ⇒⇩d (ts[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b(t ↦ m a), (),
𝒟, acquired True ?take_sb 𝒪⇩s⇩b, release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)], m, 𝒮)" .
moreover
from flush_all_until_volatile_write_Read_commute [OF i_bound ts⇩s⇩b_i [simplified "is⇩s⇩b"] ]
have flush_commute: "flush_all_until_volatile_write
(ts⇩s⇩b[i := (p⇩s⇩b,is⇩s⇩b',
θ⇩s⇩b(t↦v), sb @ [Read⇩s⇩b volatile a t v], 𝒟⇩s⇩b, 𝒪⇩s⇩b, ℛ⇩s⇩b)]) m⇩s⇩b =
flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b".
from True witness have not_volatile': "volatile' = False"
by (auto simp add: outstanding_refs_conv)
from witness not_volatile' have a_out_sb: "a ∈ outstanding_refs (Not ∘ is_volatile) sb"
apply (cases sop')
apply (fastforce simp add: outstanding_refs_conv is_volatile_def split: memref.splits)
done
with non_volatile_owned_or_read_only_outstanding_refs
[OF outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]]
have a_owned: "a ∈ 𝒪⇩s⇩b ∪ all_acquired sb ∪ read_only_reads 𝒪⇩s⇩b sb"
by auto
have "flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b a = v"
proof -
have "∀j < length ts⇩s⇩b. i ≠ j ⟶
(let (_,_,_,sb⇩j,_,_,_) = ts⇩s⇩b!j
in a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j))"
proof -
{
fix j p⇩j "is⇩j" 𝒪⇩j ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume neq_i_j: "i ≠ j"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j, xs⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j, ℛ⇩j)"
have "a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
proof
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume a_in: "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b ?take_sb⇩j"
with outstanding_refs_takeWhile [where P'= "Not ∘ is_volatile_Write⇩s⇩b"]
have a_in': "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
by auto
with non_volatile_owned_or_read_only_outstanding_non_volatile_writes
[OF outstanding_non_volatile_refs_owned_or_read_only [OF j_bound jth]]
have j_owns: "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
with ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
have a_not_owns: "a ∉ 𝒪⇩s⇩b ∪ all_acquired sb"
by blast
from non_volatile_owned_or_read_only_append [of False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
outstanding_non_volatile_refs_owned_or_read_only [OF j_bound jth]
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j"
by simp
from non_volatile_owned_or_read_only_outstanding_non_volatile_writes [OF this] a_in
have j_owns_drop: "a ∈ 𝒪⇩j ∪ all_acquired ?take_sb⇩j"
by auto
from rels_cond [rule_format, OF j_bound [simplified leq] neq_i_j] ts_sim [rule_format, OF j_bound] jth
have no_unsharing:"release ?take_sb⇩j (dom (𝒮⇩s⇩b)) ℛ⇩j a ≠ Some False"
by (auto simp add: Let_def)
{
assume "a ∈ acquired True sb 𝒪⇩s⇩b"
with acquired_all_acquired_in [OF this] ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
j_owns
have False
by auto
}
moreover
{
assume a_ro: "a ∈ read_only (share ?drop_sb 𝒮)"
from read_only_share_unowned_in [OF weak_consis_drop a_ro] a_not_owns
acquired_all_acquired [of True ?take_sb 𝒪⇩s⇩b]
all_acquired_append [of ?take_sb ?drop_sb]
have "a ∈ read_only 𝒮"
by auto
with share_all_until_volatile_write_thread_local [OF ownership_distinct_ts⇩s⇩b sharing_consis_ts⇩s⇩b j_bound jth j_owns]
have "a ∈ read_only (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto simp add: read_only_def 𝒮)
hence a_dom: "a ∈ dom (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto simp add: read_only_def domIff)
from outstanding_non_volatile_writes_unshared [OF j_bound jth]
non_volatile_writes_unshared_append [of 𝒮⇩s⇩b ?take_sb⇩j ?drop_sb⇩j]
have nvw: "non_volatile_writes_unshared 𝒮⇩s⇩b ?take_sb⇩j" by auto
from release_not_unshared_no_write_take [OF this no_unsharing a_dom] a_in
have False by auto
}
moreover
{
assume a_share: "volatile ∧ a ∈ dom (share ?drop_sb 𝒮)"
from outstanding_non_volatile_writes_unshared [OF j_bound jth]
have "non_volatile_writes_unshared 𝒮⇩s⇩b sb⇩j".
with non_volatile_writes_unshared_append [of 𝒮⇩s⇩b "?take_sb⇩j"
"?drop_sb⇩j"]
have unshared_take: "non_volatile_writes_unshared 𝒮⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
by clarsimp
from valid_own have own_dist: "ownership_distinct ts⇩s⇩b"
by (simp add: valid_ownership_def)
from valid_sharing have "sharing_consis 𝒮⇩s⇩b ts⇩s⇩b"
by (simp add: valid_sharing_def)
from sharing_consistent_share_all_until_volatile_write [OF own_dist this i_bound ts⇩s⇩b_i]
have sc: "sharing_consistent 𝒮 (acquired True ?take_sb 𝒪⇩s⇩b) ?drop_sb"
by (simp add: 𝒮)
from sharing_consistent_share_all_shared
have "dom (share ?drop_sb 𝒮) ⊆ dom 𝒮 ∪ all_shared ?drop_sb"
by auto
also from sharing_consistent_all_shared [OF sc]
have "… ⊆ dom 𝒮 ∪ acquired True ?take_sb 𝒪⇩s⇩b" by auto
also from acquired_all_acquired all_acquired_takeWhile
have "… ⊆ dom 𝒮 ∪ (𝒪⇩s⇩b ∪ all_acquired sb)" by force
finally
have a_shared: "a ∈ dom 𝒮"
using a_share a_not_owns
by auto
with share_all_until_volatile_write_thread_local [OF ownership_distinct_ts⇩s⇩b sharing_consis_ts⇩s⇩b j_bound jth j_owns]
have a_dom: "a ∈ dom (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto simp add: 𝒮 domIff)
from release_not_unshared_no_write_take [OF unshared_take no_unsharing a_dom] a_in
have False by auto
}
ultimately show False
using access_cond'
by auto
qed
}
thus ?thesis
by (fastforce simp add: Let_def)
qed
from flush_all_until_volatile_write_buffered_val_conv
[OF True i_bound ts⇩s⇩b_i this]
show ?thesis
by (simp add: buf_v)
qed
hence m_a_v: "m a = v"
by (simp add: m)
have tmps_commute: "θ⇩s⇩b(t ↦ v) = (θ⇩s⇩b |` (dom θ⇩s⇩b - {t}))(t ↦ v)"
apply (rule ext)
apply (auto simp add: restrict_map_def domIff)
done
from suspend_nothing
have suspend_nothing': "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') = []"
by (simp add: sb')
from 𝒟
have 𝒟': "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Read⇩s⇩b volatile a t v]) ≠ {})"
by (auto simp: outstanding_refs_append)
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b') ∼ (ts[i := (p⇩s⇩b,is⇩s⇩b',
θ⇩s⇩b(t↦m a),(),𝒟, acquired True ?take_sb 𝒪⇩s⇩b,
release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)],m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute ts⇩s⇩b' 𝒪⇩s⇩b' θ⇩s⇩b' sb' 𝒟⇩s⇩b' ℛ⇩s⇩b')
using share_all_until_volatile_write_Read_commute [OF i_bound ts⇩s⇩b_i [simplified is⇩s⇩b]]
apply (simp add: 𝒮 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b' θ⇩s⇩b' ℛ⇩s⇩b')
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim ts_i True 𝒟'
apply (clarsimp simp add: Let_def nth_list_update
outstanding_refs_conv m_a_v ts⇩s⇩b' 𝒪⇩s⇩b' 𝒮⇩s⇩b' θ⇩s⇩b' sb' ℛ⇩s⇩b' suspend_nothing'
𝒟⇩s⇩b' flush_all acquired_append release_append
split: if_split_asm )
apply (rule tmps_commute)
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct'
valid_sops' valid_dd' load_tmps_fresh' enough_flushs'
valid_program_history' valid'
m⇩s⇩b' 𝒮⇩s⇩b' 𝒪⇩s⇩b'
by (auto simp del: fun_upd_apply )
next
case False
then obtain r where r_in: "r ∈ set sb" and volatile_r: "is_volatile_Write⇩s⇩b r"
by (auto simp add: outstanding_refs_conv)
from takeWhile_dropWhile_real_prefix
[OF r_in, of "(Not ∘ is_volatile_Write⇩s⇩b)", simplified, OF volatile_r]
obtain a' v' sb'' sop' A' L' R' W' where
sb_split: "sb = takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
and
drop: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
apply (auto)
subgoal for y ys
apply (case_tac y)
apply auto
done
done
from drop suspends have suspends: "suspends = Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
by simp
have "(ts, m, 𝒮) ⇒⇩d⇧* (ts, m, 𝒮)" by auto
moreover
from flush_all_until_volatile_write_Read_commute [OF i_bound ts⇩s⇩b_i
[simplified "is⇩s⇩b"] ]
have flush_commute: "flush_all_until_volatile_write
(ts⇩s⇩b[i := (p⇩s⇩b,is⇩s⇩b', θ⇩s⇩b(t ↦ v), sb @ [Read⇩s⇩b volatile a t v], 𝒟⇩s⇩b, 𝒪⇩s⇩b, ℛ⇩s⇩b)]) m⇩s⇩b =
flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b".
have "Write⇩s⇩b True a' sop' v' A' L' R' W'∈ set sb"
by (subst sb_split) auto
from dropWhile_append1 [OF this, of "(Not ∘ is_volatile_Write⇩s⇩b)"]
have drop_app_comm:
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb @ [Read⇩s⇩b volatile a t v])) =
dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ [Read⇩s⇩b volatile a t v]"
by simp
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have "t ∉ dom θ⇩s⇩b"
by (auto simp add: "is⇩s⇩b")
then have tmps_commute:
"θ⇩s⇩b |` (dom θ⇩s⇩b - read_tmps sb'') =
θ⇩s⇩b |` (dom θ⇩s⇩b - insert t (read_tmps sb''))"
apply -
apply (rule ext)
apply auto
done
from 𝒟
have 𝒟': "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Read⇩s⇩b volatile a t v]) ≠ {})"
by (auto simp: outstanding_refs_append)
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute ts⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b' sb' 𝒟⇩s⇩b' )
using share_all_until_volatile_write_Read_commute [OF i_bound ts⇩s⇩b_i [simplified is⇩s⇩b]]
apply (simp add: 𝒮 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b')
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim ts_i is_sim 𝒟'
apply (clarsimp simp add: Let_def nth_list_update is_sim drop_app_comm
read_tmps_append suspends prog_instrs_append_Read⇩s⇩b instrs_append_Read⇩s⇩b
hd_prog_append_Read⇩s⇩b
drop "is⇩s⇩b" ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b' 𝒟⇩s⇩b' acquired_append takeWhile_append1 [OF r_in] volatile_r
split: if_split_asm)
apply (simp add: drop tmps_commute)+
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' valid_dd'
valid_sops' load_tmps_fresh' enough_flushs'
valid_program_history' valid' m⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply )
qed
next
case (SBHReadUnbuffered a volatile t)
then obtain
"is⇩s⇩b": "is⇩s⇩b = Read volatile a t # is⇩s⇩b'" and
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b" and
ℛ⇩s⇩b': "ℛ⇩s⇩b'=ℛ⇩s⇩b" and
θ⇩s⇩b': "θ⇩s⇩b' = θ⇩s⇩b(t↦(m⇩s⇩b a))" and
sb': "sb'=sb@[Read⇩s⇩b volatile a t (m⇩s⇩b a)]" and
m⇩s⇩b': "m⇩s⇩b' = m⇩s⇩b" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b" and
𝒟⇩s⇩b': "𝒟⇩s⇩b'=𝒟⇩s⇩b" and
buf_None: "buffered_val sb a = None"
by auto
from safe_memop_flush_sb [simplified is⇩s⇩b]
obtain access_cond': "a ∈ acquired True sb 𝒪⇩s⇩b ∨
a ∈ read_only (share ?drop_sb 𝒮) ∨ (volatile ∧ a ∈ dom (share ?drop_sb 𝒮))" and
volatile_clean: "volatile ⟶ ¬ 𝒟⇩s⇩b" and
rels_cond: "∀j < length ts. i≠j ⟶ released (ts!j) a ≠ Some False" and
rels_nv_cond: "¬volatile ⟶ (∀j < length ts. i≠j ⟶ a ∉ dom (released (ts!j)))"
by cases auto
from clean_no_outstanding_volatile_Write⇩s⇩b [OF i_bound ts⇩s⇩b_i] volatile_clean
have volatile_cond: "volatile ⟶ outstanding_refs is_volatile_Write⇩s⇩b sb ={}"
by auto
{
fix j p⇩j "is⇩s⇩b⇩j" 𝒪⇩j ℛ⇩j 𝒟⇩s⇩b⇩j θ⇩s⇩b⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume neq_i_j: "i ≠ j"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩s⇩b⇩j, θ⇩s⇩b⇩j, sb⇩j, 𝒟⇩s⇩b⇩j, 𝒪⇩j,ℛ⇩j)"
assume non_vol: "¬ volatile"
have "a ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
proof
assume a_j: "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
from ts_sim [rule_format, OF j_bound] jth
obtain suspends⇩j "is⇩j" 𝒟⇩j where
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by (auto simp add: Let_def)
from a_j ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
have a_notin_sb: "a ∉ 𝒪⇩s⇩b ∪ all_acquired sb"
by auto
with acquired_all_acquired [of True sb 𝒪⇩s⇩b]
have a_not_acq: "a ∉ acquired True sb 𝒪⇩s⇩b" by blast
with access_cond' non_vol
have a_ro: "a ∈ read_only (share ?drop_sb 𝒮)"
by auto
from read_only_share_unowned_in [OF weak_consis_drop a_ro] a_notin_sb
acquired_all_acquired [of True ?take_sb 𝒪⇩s⇩b]
all_acquired_append [of ?take_sb ?drop_sb]
have a_ro_shared: "a ∈ read_only 𝒮"
by auto
from rels_nv_cond [rule_format, OF non_vol j_bound [simplified leq] neq_i_j] ts⇩j
have "a ∉ dom (release ?take_sb⇩j (dom (𝒮⇩s⇩b)) ℛ⇩j)"
by auto
with dom_release_takeWhile [of sb⇩j "(dom (𝒮⇩s⇩b))" ℛ⇩j]
obtain
a_rels⇩j: "a ∉ dom ℛ⇩j" and
a_shared⇩j: "a ∉ all_shared ?take_sb⇩j"
by auto
have "a ∉ ⋃((λ(_, _, _, sb, _, _, _). all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)) `
set ts⇩s⇩b)"
proof -
{
fix k p⇩k "is⇩k" θ⇩k sb⇩k 𝒟⇩k 𝒪⇩k ℛ⇩k
assume k_bound: "k < length ts⇩s⇩b"
assume ts_k: "ts⇩s⇩b ! k = (p⇩k,is⇩k,θ⇩k,sb⇩k,𝒟⇩k,𝒪⇩k,ℛ⇩k)"
assume a_in: "a ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k)"
have False
proof (cases "k=j")
case True with a_shared⇩j jth ts_k a_in show False by auto
next
case False
from ownership_distinct [OF j_bound k_bound False [symmetric] jth ts_k] a_j
have "a ∉ (𝒪⇩k ∪ all_acquired sb⇩k)" by auto
with all_shared_acquired_or_owned [OF sharing_consis [OF k_bound ts_k]] a_in
show False
using all_acquired_append [of "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"
"dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"]
all_shared_append [of "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"
"dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"] by auto
qed
}
thus ?thesis by (fastforce simp add: in_set_conv_nth)
qed
with a_ro_shared
read_only_shared_all_until_volatile_write_subset' [of ts⇩s⇩b 𝒮⇩s⇩b]
have a_ro_shared⇩s⇩b: "a ∈ read_only 𝒮⇩s⇩b"
by (auto simp add: 𝒮)
with read_only_unowned [OF j_bound jth]
have a_notin_owns_j: "a ∉ 𝒪⇩j"
by auto
have own_dist: "ownership_distinct ts⇩s⇩b" by fact
have share_consis: "sharing_consis 𝒮⇩s⇩b ts⇩s⇩b" by fact
from sharing_consistent_share_all_until_volatile_write [OF own_dist share_consis i_bound ts⇩s⇩b_i]
have consis': "sharing_consistent 𝒮 (acquired True ?take_sb 𝒪⇩s⇩b) ?drop_sb"
by (simp add: 𝒮)
from share_all_until_volatile_write_thread_local [OF own_dist share_consis j_bound jth a_j] a_ro_shared
have a_ro_take: "a ∈ read_only (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto simp add: domIff 𝒮 read_only_def)
from sharing_consis [OF j_bound jth]
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
from sharing_consistent_weak_sharing_consistent [OF this] weak_sharing_consistent_append [of 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have weak_consis_drop:"weak_sharing_consistent 𝒪⇩j ?take_sb⇩j"
by auto
from read_only_share_acquired_all_shared [OF this read_only_unowned [OF j_bound jth] a_ro_take ] a_notin_owns_j a_shared⇩j
have "a ∉ all_acquired ?take_sb⇩j"
by auto
with a_j a_notin_owns_j
have a_drop: "a ∈ all_acquired ?drop_sb⇩j"
using all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
by simp
from i_bound j_bound leq have j_bound_ts': "j < length ?ts'"
by auto
note conflict_drop = a_drop [simplified suspends⇩j [symmetric]]
from split_all_acquired_in [OF conflict_drop]
show False
proof
assume "∃sop a' v ys zs A L R W.
(suspends⇩j = ys @ Write⇩s⇩b True a' sop v A L R W# zs) ∧ a ∈ A"
then
obtain a' sop' v' ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Write⇩s⇩b True a' sop' v' A' L' R' W'# zs"
(is "suspends⇩j = ?suspends") and
a_A': "a ∈ A'"
by blast
from sharing_consis [OF j_bound jth]
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
then have A'_R': "A' ∩ R' = {}"
by (simp add: sharing_consistent_append [of _ _ ?take_sb⇩j ?drop_sb⇩j, simplified]
suspends⇩j [symmetric] split_suspends⇩j sharing_consistent_append)
from valid_program_history [OF j_bound jth]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts'_j: "?ts'!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j, release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by auto
from valid_last_prog [OF j_bound jth] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j suspends⇩j = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound jth]
have reads_consis_j: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound
jth reads_consis_j]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m_j]
have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
hence reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W'])"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_write_sops [OF j_bound jth]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound jth]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound jth]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis_j]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b
(ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
have acq_simp:
"acquired True (ys @ [Write⇩s⇩b True a' sop' v' A' L' R' W'])
(acquired True ?take_sb⇩j 𝒪⇩j) =
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R'"
by (simp add: acquired_append)
from flush_store_buffer_append [where sb="ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']" and sb'="zs", simplified,
OF j_bound_ts' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts'_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="share ?drop_sb 𝒮"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs zs @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs" and
steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d⇧*
(?ts'[j:=(last_prog
(hd_prog p⇩j (Write⇩s⇩b True a' sop' v' A' L' R' W'# zs)) (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']),
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs),
(), True, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R',ℛ⇩j')],
flush (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) (flush ?drop_sb m),
share (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) (share ?drop_sb 𝒮))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append outstanding_refs_append)
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b, θ⇩s⇩b,(),
𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
note conflict_computation = rtranclp_trans [OF steps_flush_sb steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this i_bound_ys ts_ys_i, simplified is⇩s⇩b] non_vol a_not_acq
have "a ∈ read_only (share (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) (share ?drop_sb 𝒮))"
apply cases
apply (auto simp add: Let_def is⇩s⇩b)
done
with a_A'
show False
by (simp add: share_append in_read_only_convs)
next
assume "∃A L R W ys zs. suspends⇩j = ys @ Ghost⇩s⇩b A L R W # zs ∧ a ∈ A"
then
obtain A' L' R' W' ys zs where
split_suspends⇩j: "suspends⇩j = ys @ Ghost⇩s⇩b A' L' R' W'# zs"
(is "suspends⇩j = ?suspends") and
a_A': "a ∈ A'"
by blast
from valid_program_history [OF j_bound jth]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts'_j: "?ts'!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j, release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by auto
from valid_last_prog [OF j_bound jth] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j suspends⇩j = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound jth]
have reads_consis_j: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound
jth reads_consis_j]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m_j]
have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
hence reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) (ys@[Ghost⇩s⇩b A' L' R' W'])"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_write_sops [OF j_bound jth]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys@[Ghost⇩s⇩b A' L' R' W']). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound jth]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound jth]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis_j]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b
(ys@[Ghost⇩s⇩b A' L' R' W']) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
have acq_simp:
"acquired True (ys @ [Ghost⇩s⇩b A' L' R' W'])
(acquired True ?take_sb⇩j 𝒪⇩j) =
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R'"
by (simp add: acquired_append)
from flush_store_buffer_append [where sb="ys@[Ghost⇩s⇩b A' L' R' W']" and sb'="zs", simplified,
OF j_bound_ts' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts'_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="share ?drop_sb 𝒮"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs zs @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs" and
steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d⇧*
(?ts'[j:=(last_prog
(hd_prog p⇩j (Ghost⇩s⇩b A' L' R' W'# zs)) (ys@[Ghost⇩s⇩b A' L' R' W']),
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs),
(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b (ys @ [Ghost⇩s⇩b A' L' R' W']) ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R',ℛ⇩j')],
flush (ys@[Ghost⇩s⇩b A' L' R' W']) (flush ?drop_sb m),
share (ys@[Ghost⇩s⇩b A' L' R' W']) (share ?drop_sb 𝒮))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append)
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b,θ⇩s⇩b,(),
𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
note conflict_computation = rtranclp_trans [OF steps_flush_sb steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this i_bound_ys ts_ys_i, simplified is⇩s⇩b] non_vol a_not_acq
have "a ∈ read_only (share (ys@[Ghost⇩s⇩b A' L' R' W']) (share ?drop_sb 𝒮))"
apply cases
apply (auto simp add: Let_def is⇩s⇩b)
done
with a_A'
show False
by (simp add: share_append in_read_only_convs)
qed
qed
}
note non_volatile_unowned_others = this
{
assume a_in: "a ∈ read_only (share (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒮)"
assume nv: "¬ volatile"
have "a ∈ read_only (share sb 𝒮⇩s⇩b)"
proof (cases "a ∈ 𝒪⇩s⇩b ∪ all_acquired sb")
case True
from share_all_until_volatile_write_thread_local' [OF ownership_distinct_ts⇩s⇩b
sharing_consis_ts⇩s⇩b i_bound ts⇩s⇩b_i True] True a_in
show ?thesis
by (simp add: 𝒮 read_only_def)
next
case False
from read_only_share_unowned [OF weak_consis_drop _ a_in] False
acquired_all_acquired [of True ?take_sb 𝒪⇩s⇩b] all_acquired_append [of ?take_sb ?drop_sb]
have a_ro_shared: "a ∈ read_only 𝒮"
by auto
have "a ∉ ⋃((λ(_, _, _, sb, _, _, _).
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)) ` set ts⇩s⇩b)"
proof -
{
fix k p⇩k "is⇩k" θ⇩k sb⇩k 𝒟⇩k 𝒪⇩k ℛ⇩k
assume k_bound: "k < length ts⇩s⇩b"
assume ts_k: "ts⇩s⇩b ! k = (p⇩k,is⇩k,θ⇩k,sb⇩k,𝒟⇩k,𝒪⇩k,ℛ⇩k)"
assume a_in: "a ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k)"
have False
proof (cases "k=i")
case True with False ts⇩s⇩b_i ts_k a_in
all_shared_acquired_or_owned [OF sharing_consis [OF k_bound ts_k]]
all_shared_append [of "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"
"dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k"] show False by auto
next
case False
from rels_nv_cond [rule_format, OF nv k_bound [simplified leq] False [symmetric] ]
ts_sim [rule_format, OF k_bound] ts_k
have "a ∉ dom (release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k) (dom (𝒮⇩s⇩b)) ℛ⇩k)"
by (auto simp add: Let_def)
with dom_release_takeWhile [of sb⇩k "(dom (𝒮⇩s⇩b))" ℛ⇩k]
obtain
a_rels⇩j: "a ∉ dom ℛ⇩k" and
a_shared⇩j: "a ∉ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩k)"
by auto
with False a_in show ?thesis
by auto
qed
}
thus ?thesis
by (auto simp add: in_set_conv_nth)
qed
with read_only_shared_all_until_volatile_write_subset' [of ts⇩s⇩b 𝒮⇩s⇩b] a_ro_shared
have "a ∈ read_only 𝒮⇩s⇩b"
by (auto simp add: 𝒮)
from read_only_share_unowned' [OF weak_consis_sb read_only_unowned [OF i_bound ts⇩s⇩b_i] False this]
show ?thesis .
qed
} note non_vol_ro_reduction = this
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (cases volatile)
case False
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b sb".
then
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b (sb@[Read⇩s⇩b False a t (m⇩s⇩b a)])"
using access_cond' False non_vol_ro_reduction
by (auto simp add: non_volatile_owned_or_read_only_append)
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (auto simp add: False ts⇩s⇩b' sb' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
case True
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b sb".
then
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b (sb@[Read⇩s⇩b True a t (m⇩s⇩b a)])"
using True
by (simp add: non_volatile_owned_or_read_only_append)
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (auto simp add: True ts⇩s⇩b' sb' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof -
have out: "outstanding_refs is_volatile_Write⇩s⇩b (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)]) ⊆
outstanding_refs is_volatile_Write⇩s⇩b sb"
by (auto simp add: outstanding_refs_append)
have "all_acquired (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)]) ⊆ all_acquired sb"
by (auto simp add: all_acquired_append)
from outstanding_volatile_writes_unowned_by_others_store_buffer
[OF i_bound ts⇩s⇩b_i out this]
show ?thesis by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b')
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof (cases volatile)
case True
have r: "read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b)
(sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)])) 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)]))
⊆ read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
apply (case_tac "outstanding_refs (is_volatile_Write⇩s⇩b) sb = {}")
apply (simp_all add: outstanding_vol_write_take_drop_appends
acquired_append read_only_reads_append True)
done
have "𝒪⇩s⇩b ∪ all_acquired (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)]) ⊆ 𝒪⇩s⇩b ∪ all_acquired sb"
by (simp add: all_acquired_append)
from read_only_reads_unowned_nth_update [OF i_bound ts⇩s⇩b_i r this]
show ?thesis
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb')
next
case False
show ?thesis
proof (unfold_locales)
fix n m
fix p⇩n "is⇩n" 𝒪⇩n ℛ⇩n 𝒟⇩n θ⇩n sb⇩n p⇩m "is⇩m" 𝒪⇩m ℛ⇩m 𝒟⇩m θ⇩m sb⇩m
assume n_bound: "n < length ts⇩s⇩b'"
and m_bound: "m < length ts⇩s⇩b'"
and neq_n_m: "n≠m"
and nth: "ts⇩s⇩b'!n = (p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
and mth: "ts⇩s⇩b'!m =(p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
from n_bound have n_bound': "n < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
from m_bound have m_bound': "m < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
have acq_eq: "(𝒪⇩s⇩b' ∪ all_acquired sb') = (𝒪⇩s⇩b ∪ all_acquired sb)"
by (simp add: all_acquired_append sb' 𝒪⇩s⇩b')
show "(𝒪⇩m ∪ all_acquired sb⇩m) ∩
read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) 𝒪⇩n)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) =
{}"
proof (cases "m=i")
case True
with neq_n_m have neq_n_i: "n≠i"
by auto
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
note read_only_reads_unowned [OF n_bound' i_bound neq_n_i nth' ts⇩s⇩b_i]
moreover
note acq_eq
ultimately show ?thesis
using True ts⇩s⇩b_i nth mth n_bound' m_bound'
by (simp add: ts⇩s⇩b')
next
case False
note neq_m_i = this
with m_bound mth i_bound have mth': "ts⇩s⇩b!m = (p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
by (auto simp add: ts⇩s⇩b')
show ?thesis
proof (cases "n=i")
case True
note read_only_reads_unowned [OF i_bound m_bound' neq_m_i [symmetric] ts⇩s⇩b_i mth']
moreover
note acq_eq
moreover
note non_volatile_unowned_others [OF m_bound' neq_m_i [symmetric] mth']
ultimately show ?thesis
using True ts⇩s⇩b_i nth mth n_bound' m_bound' neq_m_i
apply (case_tac "outstanding_refs (is_volatile_Write⇩s⇩b) sb = {}")
apply (clarsimp simp add: outstanding_vol_write_take_drop_appends
acquired_append read_only_reads_append ts⇩s⇩b' sb' 𝒪⇩s⇩b')+
done
next
case False
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
from read_only_reads_unowned [OF n_bound' m_bound' neq_n_m nth' mth'] False neq_m_i
show ?thesis
by (clarsimp)
qed
qed
qed
qed
show "ownership_distinct ts⇩s⇩b'"
proof -
have "all_acquired (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)]) ⊆ all_acquired sb"
by (auto simp add: all_acquired_append)
from ownership_distinct_instructions_read_value_store_buffer_independent
[OF i_bound ts⇩s⇩b_i this]
show ?thesis by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b')
qed
qed
have valid_hist': "valid_history program_step ts⇩s⇩b'"
proof -
from valid_history [OF i_bound ts⇩s⇩b_i]
have hcons: "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b sb) sb".
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have t_notin_reads: "t ∉ read_tmps sb"
by (auto simp add: "is⇩s⇩b")
from load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have t_notin_writes: "t ∉ ⋃(fst ` write_sops sb )"
by (auto simp add: "is⇩s⇩b")
from valid_write_sops [OF i_bound ts⇩s⇩b_i]
have valid_sops: "∀sop ∈ write_sops sb. valid_sop sop"
by auto
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have t_fresh: "t ∉ dom θ⇩s⇩b"
using "is⇩s⇩b"
by simp
from valid_implies_valid_prog_hd [OF i_bound ts⇩s⇩b_i valid]
have "history_consistent (θ⇩s⇩b(t↦m⇩s⇩b a))
(hd_prog p⇩s⇩b (sb@ [Read⇩s⇩b volatile a t (m⇩s⇩b a)]))
(sb@ [Read⇩s⇩b volatile a t (m⇩s⇩b a)])"
using t_notin_writes valid_sops t_fresh hcons
apply -
apply (rule history_consistent_appendI)
apply (auto simp add: hd_prog_append_Read⇩s⇩b)
done
from valid_history_nth_update [OF i_bound this]
show ?thesis
by (auto simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b' θ⇩s⇩b')
qed
from
reads_consistent_unbuffered_snoc [OF buf_None refl valid_reads [OF i_bound ts⇩s⇩b_i] volatile_cond ]
have reads_consis': "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)])"
by (simp split: if_split_asm)
from valid_reads_nth_update [OF i_bound this]
have valid_reads': "valid_reads m⇩s⇩b ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b')
have valid_sharing': "valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
from outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i]
have "non_volatile_writes_unshared 𝒮⇩s⇩b (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)])"
by (auto simp add: non_volatile_writes_unshared_append)
from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
show "outstanding_non_volatile_writes_unshared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' 𝒮⇩s⇩b')
next
from sharing_consis [OF i_bound ts⇩s⇩b_i]
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b sb".
then
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)])"
by (simp add: sharing_consistent_append)
from sharing_consis_nth_update [OF i_bound this]
show "sharing_consis 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb' 𝒮⇩s⇩b')
next
note read_only_unowned [OF i_bound ts⇩s⇩b_i]
from read_only_unowned_nth_update [OF i_bound this]
show "read_only_unowned 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b')
next
from unowned_shared_nth_update [OF i_bound ts⇩s⇩b_i subset_refl]
show "unowned_shared 𝒮⇩s⇩b' ts⇩s⇩b'" by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
have "no_write_to_read_only_memory 𝒮⇩s⇩b sb".
hence "no_write_to_read_only_memory 𝒮⇩s⇩b (sb@[Read⇩s⇩b volatile a t (m⇩s⇩b a)])"
by (simp add: no_write_to_read_only_memory_append)
from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
show "no_outstanding_write_to_read_only_memory 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒮⇩s⇩b' sb')
qed
have tmps_distinct': "tmps_distinct ts⇩s⇩b'"
proof (intro_locales)
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_load_tmps is⇩s⇩b'"
by (auto split: instr.splits simp add: is⇩s⇩b)
from load_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b')
next
from read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_read_tmps sb".
moreover
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "t ∉ read_tmps sb"
by (auto simp add: is⇩s⇩b)
ultimately have "distinct_read_tmps (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)])"
by (auto simp add: distinct_read_tmps_append)
from read_tmps_distinct_nth_update [OF i_bound this]
show "read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb')
next
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ read_tmps (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)]) = {}"
by (clarsimp simp add: read_tmps_append "is⇩s⇩b")
from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb')
qed
have valid_sops': "valid_sops ts⇩s⇩b'"
proof -
from valid_store_sops [OF i_bound ts⇩s⇩b_i]
have valid_store_sops': "∀sop∈store_sops is⇩s⇩b'. valid_sop sop"
by (auto simp add: "is⇩s⇩b")
from valid_write_sops [OF i_bound ts⇩s⇩b_i]
have valid_write_sops': "∀sop∈write_sops (sb@ [Read⇩s⇩b volatile a t (m⇩s⇩b a)]).
valid_sop sop"
by (auto simp add: write_sops_append)
from valid_sops_nth_update [OF i_bound valid_write_sops' valid_store_sops']
show ?thesis by (simp add: ts⇩s⇩b' sb')
qed
have valid_dd': "valid_data_dependency ts⇩s⇩b'"
proof -
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
have dd_is: "data_dependency_consistent_instrs (dom θ⇩s⇩b') is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b" θ⇩s⇩b')
from load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ ⋃(fst ` write_sops (sb@ [Read⇩s⇩b volatile a t (m⇩s⇩b a)])) = {}"
by (auto simp add: write_sops_append "is⇩s⇩b")
from valid_data_dependency_nth_update [OF i_bound dd_is this]
show ?thesis by (simp add: ts⇩s⇩b' sb')
qed
have load_tmps_fresh': "load_tmps_fresh ts⇩s⇩b'"
proof -
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have "load_tmps (Read volatile a t # is⇩s⇩b') ∩ dom θ⇩s⇩b = {}"
by (simp add: "is⇩s⇩b")
moreover
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i] have "t ∉ load_tmps is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b")
ultimately have "load_tmps is⇩s⇩b' ∩ dom (θ⇩s⇩b(t ↦ (m⇩s⇩b a))) = {}"
by auto
from load_tmps_fresh_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' θ⇩s⇩b')
qed
have enough_flushs': "enough_flushs ts⇩s⇩b'"
proof -
from clean_no_outstanding_volatile_Write⇩s⇩b [OF i_bound ts⇩s⇩b_i]
have "¬ 𝒟⇩s⇩b ⟶ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Read⇩s⇩b volatile a t (m⇩s⇩b a)]) = {}"
by (auto simp add: outstanding_refs_append )
from enough_flushs_nth_update [OF i_bound this]
show ?thesis
by (simp add: ts⇩s⇩b' sb' 𝒟⇩s⇩b')
qed
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
from valid_program_history [OF i_bound ts⇩s⇩b_i]
have "causal_program_history is⇩s⇩b sb" .
then have causal': "causal_program_history is⇩s⇩b' (sb@[Read⇩s⇩b volatile a t (m⇩s⇩b a)])"
by (auto simp: causal_program_history_Read "is⇩s⇩b")
from valid_last_prog [OF i_bound ts⇩s⇩b_i]
have "last_prog p⇩s⇩b sb = p⇩s⇩b".
hence "last_prog p⇩s⇩b (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)]) = p⇩s⇩b"
by (simp add: last_prog_append_Read⇩s⇩b)
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b' sb')
qed
show ?thesis
proof (cases "outstanding_refs is_volatile_Write⇩s⇩b sb = {}")
case True
from True have flush_all: "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = sb"
by (auto simp add: outstanding_refs_conv )
from True have suspend_nothing: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = []"
by (auto simp add: outstanding_refs_conv)
hence suspends_empty: "suspends = []"
by (simp add: suspends)
from suspends_empty is_sim have "is": "is = Read volatile a t # is⇩s⇩b'"
by (simp add: "is⇩s⇩b")
with suspends_empty ts_i
have ts_i: "ts!i = (p⇩s⇩b, Read volatile a t # is⇩s⇩b', θ⇩s⇩b,(),
𝒟, acquired True ?take_sb 𝒪⇩s⇩b, release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
from direct_memop_step.Read
have "(Read volatile a t # is⇩s⇩b',θ⇩s⇩b, (), m,
𝒟, acquired True ?take_sb 𝒪⇩s⇩b,release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b,𝒮) →
(is⇩s⇩b', θ⇩s⇩b(t ↦ m a), (), m, 𝒟, acquired True ?take_sb 𝒪⇩s⇩b,
release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b, 𝒮)".
from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
have "(ts, m, 𝒮) ⇒⇩d (ts[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b(t ↦ m a), (),
𝒟, acquired True ?take_sb 𝒪⇩s⇩b,release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)], m, 𝒮)".
moreover
from flush_all_until_volatile_write_Read_commute [OF i_bound ts⇩s⇩b_i [simplified "is⇩s⇩b"] ]
have flush_commute: "flush_all_until_volatile_write
(ts⇩s⇩b[i := (p⇩s⇩b,is⇩s⇩b', θ⇩s⇩b(t↦m⇩s⇩b a), sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)], 𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ⇩s⇩b)])
m⇩s⇩b =
flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b".
have "flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b a = m⇩s⇩b a"
proof -
have "∀j < length ts⇩s⇩b. i ≠ j ⟶
(let (_,_,_,sb⇩j,_,_,_) = ts⇩s⇩b!j
in a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j))"
proof -
{
fix j p⇩j "is⇩j" 𝒪⇩j ℛ⇩j 𝒟⇩j acq⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume neq_i_j: "i ≠ j"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j, xs⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j, ℛ⇩j)"
have "a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
proof
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume a_in: "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b ?take_sb⇩j"
with outstanding_refs_takeWhile [where P'= "Not ∘ is_volatile_Write⇩s⇩b"]
have a_in': "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
by auto
with non_volatile_owned_or_read_only_outstanding_non_volatile_writes
[OF outstanding_non_volatile_refs_owned_or_read_only [OF j_bound jth]]
have j_owns: "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
with ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
have a_not_owns: "a ∉ 𝒪⇩s⇩b ∪ all_acquired sb"
by blast
from non_volatile_owned_or_read_only_append [of False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
outstanding_non_volatile_refs_owned_or_read_only [OF j_bound jth]
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j"
by simp
from non_volatile_owned_or_read_only_outstanding_non_volatile_writes [OF this] a_in
have j_owns_drop: "a ∈ 𝒪⇩j ∪ all_acquired ?take_sb⇩j"
by auto
from rels_cond [rule_format, OF j_bound [simplified leq] neq_i_j] ts_sim [rule_format, OF j_bound] jth
have no_unsharing:"release ?take_sb⇩j (dom (𝒮⇩s⇩b)) ℛ⇩j a ≠ Some False"
by (auto simp add: Let_def)
{
assume "a ∈ acquired True sb 𝒪⇩s⇩b"
with acquired_all_acquired_in [OF this] ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
j_owns
have False
by auto
}
moreover
{
assume a_ro: "a ∈ read_only (share ?drop_sb 𝒮)"
from read_only_share_unowned_in [OF weak_consis_drop a_ro] a_not_owns
acquired_all_acquired [of True ?take_sb 𝒪⇩s⇩b]
all_acquired_append [of ?take_sb ?drop_sb]
have "a ∈ read_only 𝒮"
by auto
with share_all_until_volatile_write_thread_local [OF ownership_distinct_ts⇩s⇩b sharing_consis_ts⇩s⇩b j_bound jth j_owns]
have "a ∈ read_only (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto simp add: read_only_def 𝒮)
hence a_dom: "a ∈ dom (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto simp add: read_only_def domIff)
from outstanding_non_volatile_writes_unshared [OF j_bound jth]
non_volatile_writes_unshared_append [of 𝒮⇩s⇩b ?take_sb⇩j ?drop_sb⇩j]
have nvw: "non_volatile_writes_unshared 𝒮⇩s⇩b ?take_sb⇩j" by auto
from release_not_unshared_no_write_take [OF this no_unsharing a_dom] a_in
have False by auto
}
moreover
{
assume a_share: "volatile ∧ a ∈ dom (share ?drop_sb 𝒮)"
from outstanding_non_volatile_writes_unshared [OF j_bound jth]
have "non_volatile_writes_unshared 𝒮⇩s⇩b sb⇩j".
with non_volatile_writes_unshared_append [of 𝒮⇩s⇩b "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
have unshared_take: "non_volatile_writes_unshared 𝒮⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
by clarsimp
from valid_own have own_dist: "ownership_distinct ts⇩s⇩b"
by (simp add: valid_ownership_def)
from valid_sharing have "sharing_consis 𝒮⇩s⇩b ts⇩s⇩b"
by (simp add: valid_sharing_def)
from sharing_consistent_share_all_until_volatile_write [OF own_dist this i_bound ts⇩s⇩b_i]
have sc: "sharing_consistent 𝒮 (acquired True ?take_sb 𝒪⇩s⇩b) ?drop_sb"
by (simp add: 𝒮)
from sharing_consistent_share_all_shared
have "dom (share ?drop_sb 𝒮) ⊆ dom 𝒮 ∪ all_shared ?drop_sb"
by auto
also from sharing_consistent_all_shared [OF sc]
have "… ⊆ dom 𝒮 ∪ acquired True ?take_sb 𝒪⇩s⇩b" by auto
also from acquired_all_acquired all_acquired_takeWhile
have "… ⊆ dom 𝒮 ∪ (𝒪⇩s⇩b ∪ all_acquired sb)" by force
finally
have a_shared: "a ∈ dom 𝒮"
using a_share a_not_owns
by auto
with share_all_until_volatile_write_thread_local [OF ownership_distinct_ts⇩s⇩b sharing_consis_ts⇩s⇩b j_bound jth j_owns]
have a_dom: "a ∈ dom (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto simp add: 𝒮 domIff)
from release_not_unshared_no_write_take [OF unshared_take no_unsharing a_dom] a_in
have False by auto
}
ultimately show False
using access_cond'
by auto
qed
}
thus ?thesis
by (fastforce simp add: Let_def)
qed
from flush_all_until_volatile_write_buffered_val_conv
[OF True i_bound ts⇩s⇩b_i this]
show ?thesis
by (simp add: buf_None)
qed
hence m_a: "m a = m⇩s⇩b a"
by (simp add: m)
have tmps_commute: "θ⇩s⇩b(t ↦ (m⇩s⇩b a)) =
(θ⇩s⇩b |` (dom θ⇩s⇩b - {t}))(t ↦ (m⇩s⇩b a))"
apply (rule ext)
apply (auto simp add: restrict_map_def domIff)
done
from suspend_nothing
have suspend_nothing': "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') = []"
by (simp add: sb')
from 𝒟
have 𝒟': "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Read⇩s⇩b volatile a t (m⇩s⇩b a)]) ≠ {})"
by (auto simp: outstanding_refs_append)
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b') ∼ (ts[i := (p⇩s⇩b,is⇩s⇩b', θ⇩s⇩b(t↦m a),(), 𝒟, acquired True ?take_sb 𝒪⇩s⇩b,release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)], m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute ts⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b' sb' 𝒟⇩s⇩b' )
using share_all_until_volatile_write_Read_commute [OF i_bound ts⇩s⇩b_i [simplified is⇩s⇩b]]
apply (simp add: 𝒮 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b')
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim ts_i True 𝒟'
apply (clarsimp simp add: Let_def nth_list_update
outstanding_refs_conv m_a ts⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b' 𝒮⇩s⇩b' θ⇩s⇩b' sb' 𝒟⇩s⇩b' suspend_nothing'
flush_all acquired_append release_append
split: if_split_asm )
apply (rule tmps_commute)
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct'
valid_sops' valid_dd' load_tmps_fresh' enough_flushs'
valid_program_history' valid'
m⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply )
next
case False
then obtain r where r_in: "r ∈ set sb" and volatile_r: "is_volatile_Write⇩s⇩b r"
by (auto simp add: outstanding_refs_conv)
from takeWhile_dropWhile_real_prefix
[OF r_in, of "(Not ∘ is_volatile_Write⇩s⇩b)", simplified, OF volatile_r]
obtain a' v' sb'' sop' A' L' R' W' where
sb_split: "sb = takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
and
drop: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
apply (auto)
subgoal for y ys
apply (case_tac y)
apply auto
done
done
from drop suspends have suspends: "suspends = Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
by simp
have "(ts, m, 𝒮) ⇒⇩d⇧* (ts, m, 𝒮)" by auto
moreover
note flush_commute = flush_all_until_volatile_write_Read_commute [OF i_bound ts⇩s⇩b_i
[simplified "is⇩s⇩b"] ]
have "Write⇩s⇩b True a' sop' v' A' L' R' W'∈ set sb"
by (subst sb_split) auto
from dropWhile_append1 [OF this, of "(Not ∘ is_volatile_Write⇩s⇩b)"]
have drop_app_comm:
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)])) =
dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ [Read⇩s⇩b volatile a t (m⇩s⇩b a)]"
by simp
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have "t ∉ dom θ⇩s⇩b"
by (auto simp add: "is⇩s⇩b")
then have tmps_commute:
"θ⇩s⇩b |` (dom θ⇩s⇩b - read_tmps sb'') =
θ⇩s⇩b |` (dom θ⇩s⇩b - insert t (read_tmps sb''))"
apply -
apply (rule ext)
apply auto
done
from 𝒟
have 𝒟': "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Read⇩s⇩b volatile a t (m⇩s⇩b a)]) ≠ {})"
by (auto simp: outstanding_refs_append)
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute ts⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b' sb')
using share_all_until_volatile_write_Read_commute [OF i_bound ts⇩s⇩b_i [simplified is⇩s⇩b]]
apply (simp add: 𝒮 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b')
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim ts_i is_sim 𝒟'
apply (clarsimp simp add: Let_def nth_list_update is_sim drop_app_comm
read_tmps_append suspends prog_instrs_append_Read⇩s⇩b instrs_append_Read⇩s⇩b
hd_prog_append_Read⇩s⇩b
drop "is⇩s⇩b" ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b' 𝒟⇩s⇩b' acquired_append takeWhile_append1 [OF r_in] volatile_r split: if_split_asm)
apply (simp add: drop tmps_commute)+
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' valid_dd'
valid_sops' load_tmps_fresh' enough_flushs'
valid_program_history' valid'
m⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply )
qed
next
case (SBHWriteNonVolatile a D f A L R W)
then obtain
"is⇩s⇩b": "is⇩s⇩b = Write False a (D, f) A L R W# is⇩s⇩b'" and
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b" and
ℛ⇩s⇩b': "ℛ⇩s⇩b'=ℛ⇩s⇩b" and
θ⇩s⇩b': "θ⇩s⇩b' = θ⇩s⇩b" and
𝒟⇩s⇩b': "𝒟⇩s⇩b'=𝒟⇩s⇩b" and
sb': "sb'=sb@[Write⇩s⇩b False a (D, f) (f θ⇩s⇩b) A L R W]" and
m⇩s⇩b': "m⇩s⇩b' = m⇩s⇩b" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b"
by auto
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
have D_tmps: "D ⊆ dom θ⇩s⇩b"
by (simp add: is⇩s⇩b)
from safe_memop_flush_sb [simplified is⇩s⇩b]
obtain a_owned': "a ∈ acquired True sb 𝒪⇩s⇩b" and a_unshared': "a ∉ dom (share ?drop_sb 𝒮)" and
rels_cond: "∀j < length ts. i≠j ⟶ a ∉ dom (released (ts!j))"
by cases auto
from a_owned' acquired_all_acquired
have a_owned'': "a ∈ 𝒪⇩s⇩b ∪ all_acquired sb"
by auto
{
fix j
fix p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume ts⇩s⇩b_j: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume neq_i_j: "i ≠ j"
have "a ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
proof -
from ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j] a_owned''
show ?thesis
by auto
qed
} note a_unowned_others = this
have a_unshared: "a ∉ dom (share sb 𝒮⇩s⇩b)"
proof
assume a_share: "a ∈ dom (share sb 𝒮⇩s⇩b)"
from valid_sharing have "sharing_consis 𝒮⇩s⇩b ts⇩s⇩b"
by (simp add: valid_sharing_def)
from in_shared_sb_share_all_until_volatile_write [OF this i_bound ts⇩s⇩b_i a_owned'' a_share]
have "a ∈ dom (share ?drop_sb 𝒮)"
by (simp add: 𝒮)
with a_unshared'
show False
by auto
qed
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof -
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b sb".
with a_owned'
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b (sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W])"
by (simp add: non_volatile_owned_or_read_only_append)
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof -
have "outstanding_refs is_volatile_Write⇩s⇩b (sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W]) ⊆
outstanding_refs is_volatile_Write⇩s⇩b sb"
by (auto simp add: outstanding_refs_append)
from outstanding_volatile_writes_unowned_by_others_store_buffer
[OF i_bound ts⇩s⇩b_i this]
show ?thesis by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' all_acquired_append)
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof -
have r: "read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b)
(sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W])) 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W]))
⊆
read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
apply (case_tac "outstanding_refs (is_volatile_Write⇩s⇩b) sb = {}")
apply (simp_all add: outstanding_vol_write_take_drop_appends
acquired_append read_only_reads_append )
done
have "𝒪⇩s⇩b ∪ all_acquired (sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W]) ⊆ 𝒪⇩s⇩b ∪ all_acquired sb"
by (simp add: all_acquired_append)
from read_only_reads_unowned_nth_update [OF i_bound ts⇩s⇩b_i r this]
show ?thesis
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb')
qed
next
show "ownership_distinct ts⇩s⇩b'"
proof -
from ownership_distinct_instructions_read_value_store_buffer_independent
[OF i_bound ts⇩s⇩b_i]
show ?thesis by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' all_acquired_append)
qed
qed
have valid_hist': "valid_history program_step ts⇩s⇩b'"
proof -
from valid_history [OF i_bound ts⇩s⇩b_i]
have "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b sb) sb".
with valid_write_sops [OF i_bound ts⇩s⇩b_i] D_tmps
valid_implies_valid_prog_hd [OF i_bound ts⇩s⇩b_i valid]
have "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b (sb@[Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W]))
(sb@ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W])"
apply -
apply (rule history_consistent_appendI)
apply (auto simp add: hd_prog_append_Write⇩s⇩b)
done
from valid_history_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' θ⇩s⇩b')
qed
have valid_reads': "valid_reads m⇩s⇩b ts⇩s⇩b'"
proof -
from valid_reads [OF i_bound ts⇩s⇩b_i]
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b sb" .
from reads_consistent_snoc_Write⇩s⇩b [OF this]
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b (sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W])".
from valid_reads_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' θ⇩s⇩b')
qed
have valid_sharing': "valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
from outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i] a_unshared
have "non_volatile_writes_unshared 𝒮⇩s⇩b
(sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W])"
by (auto simp add: non_volatile_writes_unshared_append)
from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
show "outstanding_non_volatile_writes_unshared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' θ⇩s⇩b' 𝒮⇩s⇩b')
next
from sharing_consis [OF i_bound ts⇩s⇩b_i]
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b sb".
then
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b (sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W])"
by (simp add: sharing_consistent_append)
from sharing_consis_nth_update [OF i_bound this]
show "sharing_consis 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb' 𝒮⇩s⇩b')
next
from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound ts⇩s⇩b_i] ]
show "read_only_unowned 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' 𝒪⇩s⇩b')
next
from unowned_shared_nth_update [OF i_bound ts⇩s⇩b_i subset_refl]
show "unowned_shared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' θ⇩s⇩b' 𝒮⇩s⇩b')
next
from a_unshared
have "a ∉ read_only (share sb 𝒮⇩s⇩b)"
by (auto simp add: read_only_def dom_def)
with no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
have "no_write_to_read_only_memory 𝒮⇩s⇩b (sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W])"
by (simp add: no_write_to_read_only_memory_append)
from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
show "no_outstanding_write_to_read_only_memory 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' sb')
qed
have tmps_distinct': "tmps_distinct ts⇩s⇩b'"
proof (intro_locales)
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_load_tmps is⇩s⇩b'"
by (auto split: instr.splits simp add: "is⇩s⇩b")
from load_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_distinct ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' θ⇩s⇩b')
next
from read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_read_tmps sb".
hence "distinct_read_tmps (sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W])"
by (simp add: distinct_read_tmps_append)
from read_tmps_distinct_nth_update [OF i_bound this]
show "read_tmps_distinct ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' θ⇩s⇩b')
next
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ read_tmps (sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W]) = {}"
by (clarsimp simp add: read_tmps_append "is⇩s⇩b")
from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_read_tmps_distinct ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' θ⇩s⇩b')
qed
have valid_sops': "valid_sops ts⇩s⇩b'"
proof -
from valid_store_sops [OF i_bound ts⇩s⇩b_i]
obtain valid_Df: "valid_sop (D,f)" and
valid_store_sops': "∀sop∈store_sops is⇩s⇩b'. valid_sop sop"
by (auto simp add: "is⇩s⇩b")
from valid_Df valid_write_sops [OF i_bound ts⇩s⇩b_i]
have valid_write_sops': "∀sop∈write_sops (sb@ [Write⇩s⇩b False a (D, f) (f θ⇩s⇩b) A L R W]).
valid_sop sop"
by (auto simp add: write_sops_append)
from valid_sops_nth_update [OF i_bound valid_write_sops' valid_store_sops']
show ?thesis
by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' θ⇩s⇩b')
qed
have valid_dd': "valid_data_dependency ts⇩s⇩b'"
proof -
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
obtain D_indep: "D ∩ load_tmps is⇩s⇩b' = {}" and
dd_is: "data_dependency_consistent_instrs (dom θ⇩s⇩b') is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b" θ⇩s⇩b')
from load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i] D_indep
have "load_tmps is⇩s⇩b' ∩
⋃(fst ` write_sops (sb@ [Write⇩s⇩b False a (D, f) (f θ⇩s⇩b) A L R W])) = {}"
by (auto simp add: write_sops_append "is⇩s⇩b")
from valid_data_dependency_nth_update [OF i_bound dd_is this]
show ?thesis
by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' θ⇩s⇩b')
qed
have load_tmps_fresh': "load_tmps_fresh ts⇩s⇩b'"
proof -
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ dom θ⇩s⇩b = {}"
by (auto simp add: "is⇩s⇩b")
from load_tmps_fresh_nth_update [OF i_bound this]
show ?thesis
by (simp add: ts⇩s⇩b' "is⇩s⇩b" sb' 𝒪⇩s⇩b' θ⇩s⇩b')
qed
have enough_flushs': "enough_flushs ts⇩s⇩b'"
proof -
from clean_no_outstanding_volatile_Write⇩s⇩b [OF i_bound ts⇩s⇩b_i]
have "¬ 𝒟⇩s⇩b ⟶ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W]) = {}"
by (auto simp add: outstanding_refs_append )
from enough_flushs_nth_update [OF i_bound this]
show ?thesis
by (simp add: ts⇩s⇩b' sb' 𝒟⇩s⇩b')
qed
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
from valid_program_history [OF i_bound ts⇩s⇩b_i]
have "causal_program_history is⇩s⇩b sb" .
then have causal': "causal_program_history is⇩s⇩b' (sb@[Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W])"
by (auto simp: causal_program_history_Write "is⇩s⇩b")
from valid_last_prog [OF i_bound ts⇩s⇩b_i]
have "last_prog p⇩s⇩b sb = p⇩s⇩b".
hence "last_prog p⇩s⇩b (sb @ [Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W]) = p⇩s⇩b"
by (simp add: last_prog_append_Write⇩s⇩b)
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b' sb')
qed
from valid_store_sops [OF i_bound ts⇩s⇩b_i, rule_format]
have "valid_sop (D,f)" by (auto simp add: "is⇩s⇩b")
then interpret valid_sop "(D,f)" .
show ?thesis
proof (cases "outstanding_refs is_volatile_Write⇩s⇩b sb = {}")
case True
from True have flush_all: "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = sb"
by (auto simp add: outstanding_refs_conv)
from True have suspend_nothing: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = []"
by (auto simp add: outstanding_refs_conv)
hence suspends_empty: "suspends = []"
by (simp add: suspends)
from suspends_empty is_sim have "is": "is = Write False a (D,f) A L R W# is⇩s⇩b'"
by (simp add: "is⇩s⇩b")
with suspends_empty ts_i
have ts_i: "ts!i = (p⇩s⇩b, Write False a (D,f) A L R W# is⇩s⇩b',
θ⇩s⇩b,(),
𝒟, acquired True ?take_sb 𝒪⇩s⇩b,release ?take_sb (dom (𝒮⇩s⇩b)) ℛ⇩s⇩b)"
by simp
from direct_memop_step.WriteNonVolatile [OF ]
have "(Write False a (D, f) A L R W# is⇩s⇩b',
θ⇩s⇩b, (),m,𝒟,acquired True ?take_sb 𝒪⇩s⇩b ,release ?take_sb (dom (𝒮⇩s⇩b)) ℛ⇩s⇩b, 𝒮) →
(is⇩s⇩b',
θ⇩s⇩b, (), m(a := f θ⇩s⇩b), 𝒟, acquired True ?take_sb 𝒪⇩s⇩b,
release ?take_sb (dom (𝒮⇩s⇩b)) ℛ⇩s⇩b, 𝒮)".
from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
have "(ts, m, 𝒮) ⇒⇩d
(ts[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b, (),𝒟, acquired True ?take_sb 𝒪⇩s⇩b,
release ?take_sb (dom (𝒮⇩s⇩b)) ℛ⇩s⇩b)],
m(a := f θ⇩s⇩b),𝒮)".
moreover
have "∀j<length ts⇩s⇩b. i ≠ j ⟶
(let (_,_, _, sb⇩j,_,_,_) = ts⇩s⇩b ! j
in a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j))"
proof -
{
fix j p⇩j "is⇩j" 𝒪⇩j ℛ⇩j 𝒟⇩j acq⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume neq_i_j: "i ≠ j"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j, xs⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j,ℛ⇩j)"
have "a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
proof
assume a_in: "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
hence "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
using outstanding_refs_append [of is_non_volatile_Write⇩s⇩b "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
by auto
with non_volatile_owned_or_read_only_outstanding_non_volatile_writes
[OF outstanding_non_volatile_refs_owned_or_read_only [OF j_bound jth]]
have j_owns: "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
from j_owns a_owned'' ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
show False
by auto
qed
}
thus ?thesis by (fastforce simp add: Let_def)
qed
note flush_commute = flush_all_until_volatile_write_append_non_volatile_write_commute
[OF True i_bound ts⇩s⇩b_i this]
from suspend_nothing
have suspend_nothing': "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') = []"
by (simp add: sb')
from 𝒟
have 𝒟': "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b
(sb@[Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W]) ≠ {})"
by (auto simp: outstanding_refs_append)
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b') ∼
(ts[i := (p⇩s⇩b,is⇩s⇩b', θ⇩s⇩b,(),𝒟, acquired True ?take_sb 𝒪⇩s⇩b,
release ?take_sb (dom (𝒮⇩s⇩b)) ℛ⇩s⇩b)],
m(a:=f θ⇩s⇩b),𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute ts⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b' sb' θ⇩s⇩b' 𝒟⇩s⇩b' )
using share_all_until_volatile_write_Write_commute
[OF i_bound ts⇩s⇩b_i [simplified is⇩s⇩b]]
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b')
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim ts_i True 𝒟'
apply (clarsimp simp add: Let_def nth_list_update
outstanding_refs_conv ts⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b' 𝒮⇩s⇩b' θ⇩s⇩b' sb' 𝒟⇩s⇩b' suspend_nothing' flush_all
acquired_append release_append split: if_split_asm)
done
ultimately
show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' valid_sops'
valid_dd' load_tmps_fresh' enough_flushs'
valid_program_history' valid' m⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply)
next
case False
then obtain r where r_in: "r ∈ set sb" and volatile_r: "is_volatile_Write⇩s⇩b r"
by (auto simp add: outstanding_refs_conv)
from takeWhile_dropWhile_real_prefix
[OF r_in, of "(Not ∘ is_volatile_Write⇩s⇩b)", simplified, OF volatile_r]
obtain a' v' sb'' sop' A' L' R' W' where
sb_split: "sb = takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
and
drop: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
apply (auto)
subgoal for y ys
apply (case_tac y)
apply auto
done
done
from drop suspends have suspends: "suspends = Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
by simp
have "(ts, m, 𝒮) ⇒⇩d⇧* (ts, m, 𝒮)" by auto
moreover
note flush_commute =
flush_all_until_volatile_write_append_unflushed [OF False i_bound ts⇩s⇩b_i]
have "Write⇩s⇩b True a' sop' v' A' L' R' W' ∈ set sb"
by (subst sb_split) auto
note drop_app = dropWhile_append1 [OF this, of "(Not ∘ is_volatile_Write⇩s⇩b)", simplified]
from 𝒟
have 𝒟': "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Write⇩s⇩b False a (D,f) (f θ⇩s⇩b) A L R W]) ≠ {})"
by (auto simp: outstanding_refs_append)
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b') ∼ (ts,m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute ts⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b' sb')
using share_all_until_volatile_write_Write_commute
[OF i_bound ts⇩s⇩b_i [simplified is⇩s⇩b]]
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b')
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim ts_i is_sim 𝒟'
apply (clarsimp simp add: Let_def nth_list_update is_sim drop_app
read_tmps_append suspends
prog_instrs_append_Write⇩s⇩b instrs_append_Write⇩s⇩b hd_prog_append_Write⇩s⇩b
drop "is⇩s⇩b" ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' 𝒮⇩s⇩b'
θ⇩s⇩b' 𝒟⇩s⇩b' acquired_append takeWhile_append1 [OF r_in]
volatile_r
split: if_split_asm)
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' valid_dd'
valid_sops' load_tmps_fresh' enough_flushs'
valid_program_history' valid' m⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply )
qed
next
case (SBHWriteVolatile a D f A L R W)
then obtain
"is⇩s⇩b": "is⇩s⇩b = Write True a (D, f) A L R W# is⇩s⇩b'" and
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b" and
ℛ⇩s⇩b': "ℛ⇩s⇩b'=ℛ⇩s⇩b" and
θ⇩s⇩b': "θ⇩s⇩b' = θ⇩s⇩b" and
𝒟⇩s⇩b': "𝒟⇩s⇩b'=True" and
sb': "sb'=sb@[Write⇩s⇩b True a (D, f) (f θ⇩s⇩b) A L R W]" and
m⇩s⇩b': "m⇩s⇩b' = m⇩s⇩b" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b"
by auto
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
have D_subset: "D ⊆ dom θ⇩s⇩b"
by (simp add: is⇩s⇩b)
from safe_memop_flush_sb [simplified is⇩s⇩b] obtain
a_unowned_others_ts:
"∀j<length (map owned ts). i ≠ j ⟶ (a ∉ owned (ts!j) ∪ dom (released (ts!j)))" and
L_subset: "L ⊆ A" and
A_shared_owned: "A ⊆ dom (share ?drop_sb 𝒮) ∪ acquired True sb 𝒪⇩s⇩b" and
R_acq: "R ⊆ acquired True sb 𝒪⇩s⇩b" and
A_R: "A ∩ R = {}" and
A_unowned_by_others_ts:
"∀j<length (map owned ts). i≠j ⟶ (A ∩ (owned (ts!j) ∪ dom (released (ts!j))) = {})" and
a_not_ro': "a ∉ read_only (share ?drop_sb 𝒮)"
by cases auto
from a_unowned_others_ts ts_sim leq
have a_unowned_others:
"∀j<length ts⇩s⇩b. i ≠ j ⟶
(let (_,_,_,sb⇩j,_,𝒪⇩j,_) = ts⇩s⇩b!j in
a ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j ∧
a ∉ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j))"
apply (clarsimp simp add: Let_def)
subgoal for j
apply (drule_tac x=j in spec)
apply (auto simp add: dom_release_takeWhile)
done
done
have a_not_ro: "a ∉ read_only (share sb 𝒮⇩s⇩b)"
proof
assume a: "a ∈ read_only (share sb 𝒮⇩s⇩b)"
from local.read_only_unowned_axioms have "read_only_unowned 𝒮⇩s⇩b ts⇩s⇩b".
from in_read_only_share_all_until_volatile_write' [OF ownership_distinct_ts⇩s⇩b sharing_consis_ts⇩s⇩b
‹read_only_unowned 𝒮⇩s⇩b ts⇩s⇩b› i_bound ts⇩s⇩b_i a_unowned_others a]
have "a ∈ read_only (share ?drop_sb 𝒮)"
by (simp add: 𝒮)
with a_not_ro' show False by simp
qed
from A_unowned_by_others_ts ts_sim leq
have A_unowned_by_others:
"∀j<length ts⇩s⇩b. i≠j ⟶ (let (_,_,_,sb⇩j,_,𝒪⇩j,_) = ts⇩s⇩b!j
in A ∩ (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j ∪
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)) = {})"
apply (clarsimp simp add: Let_def)
subgoal for j
apply (drule_tac x=j in spec)
apply (force simp add: dom_release_takeWhile)
done
done
have a_not_acquired_others: "∀j<length (map 𝒪_sb ts⇩s⇩b). i ≠ j ⟶
(let (𝒪⇩j,sb⇩j) = (map 𝒪_sb ts⇩s⇩b)!j in a ∉ all_acquired sb⇩j)"
proof -
{
fix j 𝒪⇩j sb⇩j
assume j_bound: "j < length (map owned ts⇩s⇩b)"
assume neq_i_j: "i≠j"
assume ts⇩s⇩b_j: "(map 𝒪_sb ts⇩s⇩b)!j = (𝒪⇩j,sb⇩j)"
assume conflict: "a ∈ all_acquired sb⇩j"
have False
proof -
from j_bound leq
have j_bound': "j < length (map owned ts)"
by auto
from j_bound have j_bound'': "j < length ts⇩s⇩b"
by auto
from j_bound' have j_bound''': "j < length ts"
by simp
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
from ts_sim [rule_format, OF j_bound''] ts⇩s⇩b_j j_bound''
obtain p⇩j suspends⇩j "is⇩s⇩b⇩j" ℛ⇩j 𝒟⇩s⇩b⇩j 𝒟⇩j θ⇩s⇩b⇩j "is⇩j" where
ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j,is⇩s⇩b⇩j, θ⇩s⇩b⇩j, sb⇩j, 𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)" and
suspends⇩j: "suspends⇩j = dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j,
acquired True ?take_sb⇩j 𝒪⇩j,
release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
apply (cases "ts⇩s⇩b!j")
apply (force simp add: Let_def)
done
from a_unowned_others [rule_format,OF _ neq_i_j] ts⇩s⇩b_j j_bound
obtain a_unacq: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j" and a_not_shared: "a ∉ all_shared ?take_sb⇩j"
by auto
have conflict_drop: "a ∈ all_acquired suspends⇩j"
proof (rule ccontr)
assume "a ∉ all_acquired suspends⇩j"
with all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j] conflict
have "a ∈ all_acquired ?take_sb⇩j"
by (auto simp add: suspends⇩j)
from all_acquired_unshared_acquired [OF this a_not_shared] a_unacq
show False by auto
qed
from j_bound''' i_bound' have j_bound_ts': "j < length ?ts'"
by simp
from split_all_acquired_in [OF conflict_drop]
show ?thesis
proof
assume "∃sop a' v ys zs A L R W.
suspends⇩j = ys @ Write⇩s⇩b True a' sop v A L R W# zs ∧ a ∈ A"
then
obtain a' sop' v' ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Write⇩s⇩b True a' sop' v' A' L' R' W'# zs"
(is "suspends⇩j = ?suspends") and
a_A': "a ∈ A'"
by blast
from sharing_consis [OF j_bound'' ts⇩s⇩b_j]
have sharing_consis_j: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
then have A'_R': "A' ∩ R' = {}"
by (simp add: sharing_consistent_append [of _ _ ?take_sb⇩j ?drop_sb⇩j, simplified]
suspends⇩j [symmetric] split_suspends⇩j sharing_consistent_append)
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts'_j: "?ts'!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by auto
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j suspends⇩j = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis_j: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b›
j_bound'' ts⇩s⇩b_j this]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m_j]
have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
hence reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W'])"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis_j]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b
(ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
have acq_simp:
"acquired True (ys @ [Write⇩s⇩b True a' sop' v' A' L' R' W'])
(acquired True ?take_sb⇩j 𝒪⇩j) =
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R'"
by (simp add: acquired_append)
from flush_store_buffer_append [where sb="ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']" and sb'="zs", simplified,
OF j_bound_ts' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts'_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="share ?drop_sb 𝒮"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs zs @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs" and
steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d⇧*
(?ts'[j:=(last_prog
(hd_prog p⇩j (Write⇩s⇩b True a' sop' v' A' L' R' W'# zs)) (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']),
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs),
(), True, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R',ℛ⇩j')],
flush (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) (flush ?drop_sb m),
share (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) (share ?drop_sb 𝒮))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append outstanding_refs_append)
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b, θ⇩s⇩b,(),
𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
note conflict_computation = rtranclp_trans [OF steps_flush_sb steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this i_bound_ys ts_ys_i, simplified is⇩s⇩b]
have a_unowned:
"∀j < length ?ts_ys. i≠j ⟶ (let (𝒪⇩j) = map owned ?ts_ys!j in a ∉ 𝒪⇩j)"
apply cases
apply (auto simp add: Let_def is⇩s⇩b)
done
from a_A' a_unowned [rule_format, of j] neq_i_j j_bound' A'_R'
show False
by (auto simp add: Let_def)
next
assume "∃A L R W ys zs. suspends⇩j = ys @ Ghost⇩s⇩b A L R W# zs ∧ a ∈ A"
then
obtain A' L' R' W' ys zs where
split_suspends⇩j: "suspends⇩j = ys @ Ghost⇩s⇩b A' L' R' W'# zs"
(is "suspends⇩j = ?suspends") and
a_A': "a ∈ A'"
by blast
from sharing_consis [OF j_bound'' ts⇩s⇩b_j]
have sharing_consis_j: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
then have A'_R': "A' ∩ R' = {}"
by (simp add: sharing_consistent_append [of _ _ ?take_sb⇩j ?drop_sb⇩j, simplified]
suspends⇩j [symmetric] split_suspends⇩j sharing_consistent_append)
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts'_j: "?ts'!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j, release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by auto
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j suspends⇩j = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis_j: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b›
j_bound'' ts⇩s⇩b_j this]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m_j]
have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
hence reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) (ys@[Ghost⇩s⇩b A' L' R' W'])"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys@[Ghost⇩s⇩b A' L' R' W']). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis_j]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b
(ys@[Ghost⇩s⇩b A' L' R' W']) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
have acq_simp:
"acquired True (ys @ [Ghost⇩s⇩b A' L' R' W'])
(acquired True ?take_sb⇩j 𝒪⇩j) =
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R'"
by (simp add: acquired_append)
from flush_store_buffer_append [where sb="ys@[Ghost⇩s⇩b A' L' R' W']" and sb'="zs", simplified,
OF j_bound_ts' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts'_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="share ?drop_sb 𝒮"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs zs @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs" and
steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d⇧*
(?ts'[j:=(last_prog
(hd_prog p⇩j (Ghost⇩s⇩b A' L' R' W'# zs)) (ys@[Ghost⇩s⇩b A' L' R' W']),
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs),
(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b (ys @ [Ghost⇩s⇩b A' L' R' W']) ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R',ℛ⇩j')],
flush (ys@[Ghost⇩s⇩b A' L' R' W']) (flush ?drop_sb m),
share (ys@[Ghost⇩s⇩b A' L' R' W']) (share ?drop_sb 𝒮))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append)
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b,θ⇩s⇩b,(),
𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
note conflict_computation = rtranclp_trans [OF steps_flush_sb steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this i_bound_ys ts_ys_i, simplified is⇩s⇩b]
have a_unowned:
"∀j < length ?ts_ys. i≠j ⟶ (let (𝒪⇩j) = map owned ?ts_ys!j in a ∉ 𝒪⇩j)"
apply cases
apply (auto simp add: Let_def is⇩s⇩b)
done
from a_A' a_unowned [rule_format, of j] neq_i_j j_bound' A'_R'
show False
by (auto simp add: Let_def)
qed
qed
}
thus ?thesis
by (auto simp add: Let_def)
qed
have A_unused_by_others:
"∀j<length (map 𝒪_sb ts⇩s⇩b). i ≠ j ⟶
(let (𝒪⇩j, sb⇩j) = map 𝒪_sb ts⇩s⇩b! j
in A ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j = {})"
proof -
{
fix j 𝒪⇩j sb⇩j
assume j_bound: "j < length (map owned ts⇩s⇩b)"
assume neq_i_j: "i≠j"
assume ts⇩s⇩b_j: "(map 𝒪_sb ts⇩s⇩b)!j = (𝒪⇩j,sb⇩j)"
assume conflict: "A ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {}"
have False
proof -
from j_bound leq
have j_bound': "j < length (map owned ts)"
by auto
from j_bound have j_bound'': "j < length ts⇩s⇩b"
by auto
from j_bound' have j_bound''': "j < length ts"
by simp
from conflict obtain a' where
a'_in: "a' ∈ A" and
a'_in_j: "a' ∈ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j"
by auto
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
from ts_sim [rule_format, OF j_bound''] ts⇩s⇩b_j j_bound''
obtain p⇩j suspends⇩j "is⇩s⇩b⇩j" 𝒟⇩s⇩b⇩j 𝒟⇩j ℛ⇩j θ⇩s⇩b⇩j "is⇩j" where
ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j,is⇩s⇩b⇩j, θ⇩s⇩b⇩j, sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)" and
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(), 𝒟⇩j,
acquired True ?take_sb⇩j 𝒪⇩j,
release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
apply (cases "ts⇩s⇩b!j")
apply (force simp add: Let_def)
done
have "a' ∈ outstanding_refs is_volatile_Write⇩s⇩b suspends⇩j"
proof -
from a'_in_j
have "a' ∈ outstanding_refs is_volatile_Write⇩s⇩b (?take_sb⇩j @ ?drop_sb⇩j)"
by simp
thus ?thesis
apply (simp only: outstanding_refs_append suspends⇩j)
apply (auto simp add: outstanding_refs_conv dest: set_takeWhileD)
done
qed
from split_volatile_Write⇩s⇩b_in_outstanding_refs [OF this]
obtain sop v ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Write⇩s⇩b True a' sop v A' L' R' W'# zs" (is "suspends⇩j = ?suspends")
by blast
from direct_memop_step.WriteVolatile [where θ=θ⇩s⇩b and m="flush ?drop_sb m"]
have "(Write True a (D, f) A L R W# is⇩s⇩b',
θ⇩s⇩b, (), flush ?drop_sb m,𝒟⇩s⇩b,acquired True sb 𝒪⇩s⇩b,
release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b,
share ?drop_sb 𝒮) →
(is⇩s⇩b', θ⇩s⇩b, (), (flush ?drop_sb m)(a := f θ⇩s⇩b), True, acquired True sb 𝒪⇩s⇩b ∪ A - R, Map.empty,
share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF
i_bound_ts' [simplified is⇩s⇩b] ts'_i [simplified is⇩s⇩b] this [simplified is⇩s⇩b]]
have store_step: "(?ts', flush ?drop_sb m,share ?drop_sb 𝒮 ) ⇒⇩d
(?ts'[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b, (),
True, acquired True sb 𝒪⇩s⇩b ∪ A - R,Map.empty)],
(flush ?drop_sb m)(a := f θ⇩s⇩b), share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L )"
(is " _ ⇒⇩d (?ts_A, ?m_A, ?share_A)")
by (simp add: is⇩s⇩b)
from i_bound' have i_bound'': "i < length ?ts_A"
by simp
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts_A_j: "?ts_A!j = (hd_prog p⇩j (ys @ Write⇩s⇩b True a' sop v A' L' R' W'# zs), is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (ys @ Write⇩s⇩b True a' sop v A' L' R' W'# zs)), (), 𝒟⇩j,
acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by (simp add: split_suspends⇩j)
from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
by simp
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound''
ts⇩s⇩b_j reads_consis]
have reads_consis_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m]
have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
from a_unowned_others [rule_format, OF _ neq_i_j] j_bound ts⇩s⇩b_j
obtain a_notin_owns_j: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j" and a_unshared: "a ∉ all_shared ?take_sb⇩j"
by auto
from a_not_acquired_others [rule_format, OF _ neq_i_j] j_bound ts⇩s⇩b_j
have a_not_acquired_j: "a ∉ all_acquired sb⇩j"
by auto
from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' ts⇩s⇩b_j]
have nvo_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
have a_no_non_vol_read: "a ∉ outstanding_refs is_non_volatile_Read⇩s⇩b ?drop_sb⇩j"
proof
assume a_in_nvr:"a ∈ outstanding_refs is_non_volatile_Read⇩s⇩b ?drop_sb⇩j"
from reads_consistent_drop [OF reads_consis]
have rc: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) (flush ?take_sb⇩j m⇩s⇩b) ?drop_sb⇩j".
from non_volatile_owned_or_read_only_drop [OF nvo_j]
have nvo_j_drop: "non_volatile_owned_or_read_only True (share ?take_sb⇩j 𝒮⇩s⇩b)
(acquired True ?take_sb⇩j 𝒪⇩j)
?drop_sb⇩j"
by simp
from outstanding_refs_non_volatile_Read⇩s⇩b_all_acquired [OF rc this a_in_nvr]
have a_owns_acq_ror:
"a ∈ 𝒪⇩j ∪ all_acquired sb⇩j ∪ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by (auto dest!: acquired_all_acquired_in all_acquired_takeWhile_dropWhile_in
simp add: acquired_takeWhile_non_volatile_Write⇩s⇩b)
have a_unowned_j: "a ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
proof (cases "a ∈ 𝒪⇩j")
case False with a_not_acquired_j show ?thesis by auto
next
case True
from all_shared_acquired_in [OF True a_unshared] a_notin_owns_j
have False by auto thus ?thesis ..
qed
with a_owns_acq_ror
have a_ror: "a ∈ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
with read_only_reads_unowned [OF j_bound'' i_bound neq_i_j [symmetric] ts⇩s⇩b_j ts⇩s⇩b_i]
have a_unowned_sb: "a ∉ 𝒪⇩s⇩b ∪ all_acquired sb"
by auto
from sharing_consis [OF j_bound'' ts⇩s⇩b_j] sharing_consistent_append [of 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have consis_j_drop: "sharing_consistent (share ?take_sb⇩j 𝒮⇩s⇩b) (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
from read_only_reads_read_only [OF nvo_j_drop consis_j_drop] a_ror a_unowned_j
all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j] acquired_takeWhile_non_volatile_Write⇩s⇩b [of sb⇩j 𝒪⇩j]
have "a ∈ read_only (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto simp add: )
from read_only_share_all_shared [OF this] a_unshared
have "a ∈ read_only 𝒮⇩s⇩b"
by fastforce
from read_only_unacquired_share [OF read_only_unowned [OF i_bound ts⇩s⇩b_i]
weak_sharing_consis [OF i_bound ts⇩s⇩b_i] this] a_unowned_sb
have "a ∈ read_only (share sb 𝒮⇩s⇩b)"
by auto
with a_not_ro show False
by simp
qed
with reads_consistent_mem_eq_on_non_volatile_reads [OF _ subset_refl reads_consis_flush_m]
have "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A suspends⇩j"
by (auto simp add: suspends⇩j)
hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A ys"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops ys. valid_sop sop"
apply (simp only: write_sops_append )
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop h_consis]
last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b ys = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from flush_store_buffer_append [
OF j_bound'''' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_A_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_m_A_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?share_A"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs (Write⇩s⇩b True a' sop v A' L' R' W'# zs) @ is⇩s⇩b⇩j =
is⇩j' @ prog_instrs (Write⇩s⇩b True a' sop v A' L' R' W'# zs)" and
steps_ys: "(?ts_A, ?m_A, ?share_A) ⇒⇩d⇧*
(?ts_A[j:= (last_prog (hd_prog p⇩j (Write⇩s⇩b True a' sop v A' L' R' W'# zs)) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Write⇩s⇩b True a' sop v A' L' R' W' # zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j') ],
flush ys ?m_A,
share ys ?share_A)"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto)
note conflict_computation = rtranclp_trans [OF rtranclp_r_rtranclp [OF steps_flush_sb, OF store_step] steps_ys]
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Write⇩s⇩b True a' sop v A' L' R' W']) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Write True a' sop A' L' R' W'#is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from j_bound'''
have j_bound_ys: "j < length ?ts_ys"
by auto
from j_bound_ys neq_i_j
have ts_ys_j: "?ts_ys!j=(last_prog (hd_prog p⇩j (Write⇩s⇩b True a' sop v A' L' R' W'# zs)) ys, is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Write⇩s⇩b True a' sop v A' L' R' W'# zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by auto
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j, simplified is⇩j']
have a_unowned:
"∀i < length ?ts_ys. j≠i ⟶ (let (𝒪⇩i) = map owned ?ts_ys!i in a' ∉ 𝒪⇩i)"
apply cases
apply (auto simp add: Let_def is⇩s⇩b)
done
from a'_in a_unowned [rule_format, of i] neq_i_j i_bound' A_R
show False
by (auto simp add: Let_def)
qed
}
thus ?thesis
by (auto simp add: Let_def)
qed
have A_unaquired_by_others:
"∀j<length (map 𝒪_sb ts⇩s⇩b). i ≠ j ⟶
(let (𝒪⇩j, sb⇩j) = map 𝒪_sb ts⇩s⇩b! j
in A ∩ all_acquired sb⇩j = {})"
proof -
{
fix j 𝒪⇩j sb⇩j
assume j_bound: "j < length (map owned ts⇩s⇩b)"
assume neq_i_j: "i≠j"
assume ts⇩s⇩b_j: "(map 𝒪_sb ts⇩s⇩b)!j = (𝒪⇩j,sb⇩j)"
assume conflict: "A ∩ all_acquired sb⇩j ≠ {}"
have False
proof -
from j_bound leq
have j_bound': "j < length (map owned ts)"
by auto
from j_bound have j_bound'': "j < length ts⇩s⇩b"
by auto
from j_bound' have j_bound''': "j < length ts"
by simp
from conflict obtain a' where
a'_in: "a' ∈ A" and
a'_in_j: "a' ∈ all_acquired sb⇩j"
by auto
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
from ts_sim [rule_format, OF j_bound''] ts⇩s⇩b_j j_bound''
obtain p⇩j suspends⇩j "is⇩s⇩b⇩j" 𝒟⇩s⇩b⇩j 𝒟⇩j ℛ⇩j θ⇩s⇩b⇩j "is⇩j" where
ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j,is⇩s⇩b⇩j, θ⇩s⇩b⇩j, sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)" and
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
apply (cases "ts⇩s⇩b!j")
apply (force simp add: Let_def)
done
from a'_in_j all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have "a' ∈ all_acquired ?take_sb⇩j ∨ a' ∈ all_acquired suspends⇩j"
by (auto simp add: suspends⇩j)
thus False
proof
assume "a' ∈ all_acquired ?take_sb⇩j"
with A_unowned_by_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound a'_in
show False
by (auto dest: all_acquired_unshared_acquired)
next
assume conflict_drop: "a' ∈ all_acquired suspends⇩j"
from split_all_acquired_in [OF conflict_drop]
show False
proof
assume "∃sop a'' v ys zs A L R W.
suspends⇩j = ys @ Write⇩s⇩b True a'' sop v A L R W# zs ∧ a' ∈ A"
then
obtain a'' sop' v' ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W'# zs"
(is "suspends⇩j = ?suspends") and
a'_A': "a' ∈ A'"
by auto
from direct_memop_step.WriteVolatile [where θ=θ⇩s⇩b and m="flush ?drop_sb m"]
have "(Write True a (D, f) A L R W # is⇩s⇩b',
θ⇩s⇩b, (), flush ?drop_sb m ,𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b,
release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b,
share ?drop_sb 𝒮) →
(is⇩s⇩b', θ⇩s⇩b, (), (flush ?drop_sb m)(a := f θ⇩s⇩b), True, acquired True sb 𝒪⇩s⇩b ∪ A - R,Map.empty,
share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF
i_bound_ts' [simplified is⇩s⇩b] ts'_i [simplified is⇩s⇩b] this [simplified is⇩s⇩b]]
have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d
(?ts'[i := (p⇩s⇩b, is⇩s⇩b',
θ⇩s⇩b, (),True, acquired True sb 𝒪⇩s⇩b ∪ A - R,Map.empty)],
(flush ?drop_sb m)(a := f θ⇩s⇩b),share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
(is " _ ⇒⇩d (?ts_A, ?m_A, ?share_A)")
by (simp add: is⇩s⇩b)
from i_bound' have i_bound'': "i < length ?ts_A"
by simp
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts_A_j: "?ts_A!j = (hd_prog p⇩j (ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W'# zs), is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W'# zs)), (), 𝒟⇩j,
acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by (simp add: split_suspends⇩j)
from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
by simp
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b›
j_bound''
ts⇩s⇩b_j reads_consis]
have reads_consis_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m]
have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
from a_unowned_others [rule_format, OF _ neq_i_j] j_bound ts⇩s⇩b_j
obtain a_notin_owns_j: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j" and a_unshared: "a ∉ all_shared ?take_sb⇩j"
by auto
from a_not_acquired_others [rule_format, OF _ neq_i_j] j_bound ts⇩s⇩b_j
have a_not_acquired_j: "a ∉ all_acquired sb⇩j"
by auto
from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' ts⇩s⇩b_j]
have nvo_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
have a_no_non_vol_read: "a ∉ outstanding_refs is_non_volatile_Read⇩s⇩b ?drop_sb⇩j"
proof
assume a_in_nvr:"a ∈ outstanding_refs is_non_volatile_Read⇩s⇩b ?drop_sb⇩j"
from reads_consistent_drop [OF reads_consis]
have rc: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) (flush ?take_sb⇩j m⇩s⇩b) ?drop_sb⇩j".
from non_volatile_owned_or_read_only_drop [OF nvo_j]
have nvo_j_drop: "non_volatile_owned_or_read_only True (share ?take_sb⇩j 𝒮⇩s⇩b)
(acquired True ?take_sb⇩j 𝒪⇩j)
?drop_sb⇩j"
by simp
from outstanding_refs_non_volatile_Read⇩s⇩b_all_acquired [OF rc this a_in_nvr]
have a_owns_acq_ror:
"a ∈ 𝒪⇩j ∪ all_acquired sb⇩j ∪ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by (auto dest!: acquired_all_acquired_in all_acquired_takeWhile_dropWhile_in
simp add: acquired_takeWhile_non_volatile_Write⇩s⇩b)
have a_unowned_j: "a ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
proof (cases "a ∈ 𝒪⇩j")
case False with a_not_acquired_j show ?thesis by auto
next
case True
from all_shared_acquired_in [OF True a_unshared] a_notin_owns_j
have False by auto thus ?thesis ..
qed
with a_owns_acq_ror
have a_ror: "a ∈ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
with read_only_reads_unowned [OF j_bound'' i_bound neq_i_j [symmetric] ts⇩s⇩b_j ts⇩s⇩b_i]
have a_unowned_sb: "a ∉ 𝒪⇩s⇩b ∪ all_acquired sb"
by auto
from sharing_consis [OF j_bound'' ts⇩s⇩b_j] sharing_consistent_append [of 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have consis_j_drop: "sharing_consistent (share ?take_sb⇩j 𝒮⇩s⇩b) (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
from read_only_reads_read_only [OF nvo_j_drop consis_j_drop] a_ror a_unowned_j
all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j] acquired_takeWhile_non_volatile_Write⇩s⇩b [of sb⇩j 𝒪⇩j]
have "a ∈ read_only (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto)
from read_only_share_all_shared [OF this] a_unshared
have "a ∈ read_only 𝒮⇩s⇩b"
by fastforce
from read_only_unacquired_share [OF read_only_unowned [OF i_bound ts⇩s⇩b_i]
weak_sharing_consis [OF i_bound ts⇩s⇩b_i] this] a_unowned_sb
have "a ∈ read_only (share sb 𝒮⇩s⇩b)"
by auto
with a_not_ro show False
by simp
qed
with reads_consistent_mem_eq_on_non_volatile_reads [OF _ subset_refl reads_consis_flush_m]
have "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A suspends⇩j"
by (auto simp add: suspends⇩j)
hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A ys"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops ys. valid_sop sop"
apply (simp only: write_sops_append )
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop h_consis]
last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b ys = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from flush_store_buffer_append [
OF j_bound'''' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_A_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_m_A_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?share_A"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs (Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs) @ is⇩s⇩b⇩j =
is⇩j' @ prog_instrs (Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs)" and
steps_ys: "(?ts_A, ?m_A, ?share_A) ⇒⇩d⇧*
(?ts_A[j:= (last_prog (hd_prog p⇩j (Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs)) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j') ],
flush ys ?m_A, share ys ?share_A)"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto)
note conflict_computation = rtranclp_trans [OF rtranclp_r_rtranclp [OF steps_flush_sb, OF store_step] steps_ys]
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Write⇩s⇩b True a'' sop' v' A' L' R' W']) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Write True a'' sop' A' L' R' W'#is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from j_bound'''
have j_bound_ys: "j < length ?ts_ys"
by auto
from j_bound_ys neq_i_j
have ts_ys_j: "?ts_ys!j=(last_prog (hd_prog p⇩j (Write⇩s⇩b True a'' sop' v' A' L' R' W'# zs)) ys, is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Write⇩s⇩b True a'' sop' v' A' L' R' W'# zs)),(),𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by auto
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j, simplified is⇩j']
have A'_unowned:
"∀i < length ?ts_ys. j≠i ⟶ (let (𝒪⇩i) = map owned ?ts_ys!i in A' ∩ 𝒪⇩i = {})"
apply cases
apply (fastforce simp add: Let_def is⇩s⇩b)+
done
from a'_in a'_A' A'_unowned [rule_format, of i] neq_i_j i_bound' A_R
show False
by (auto simp add: Let_def)
next
assume "∃A L R W ys zs.
suspends⇩j = ys @ Ghost⇩s⇩b A L R W # zs ∧ a' ∈ A"
then
obtain ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Ghost⇩s⇩b A' L' R' W'# zs" (is "suspends⇩j = ?suspends") and
a'_A': "a' ∈ A'"
by auto
from direct_memop_step.WriteVolatile [where θ=θ⇩s⇩b and m="flush ?drop_sb m"]
have "(Write True a (D, f) A L R W# is⇩s⇩b',
θ⇩s⇩b, (), flush ?drop_sb m,𝒟⇩s⇩b,acquired True sb 𝒪⇩s⇩b,
release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b,
share ?drop_sb 𝒮) →
(is⇩s⇩b', θ⇩s⇩b, (), (flush ?drop_sb m)(a := f θ⇩s⇩b), True, acquired True sb 𝒪⇩s⇩b ∪ A - R, Map.empty,
share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF
i_bound_ts' [simplified is⇩s⇩b] ts'_i [simplified is⇩s⇩b] this [simplified is⇩s⇩b]]
have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d
(?ts'[i := (p⇩s⇩b, is⇩s⇩b',
θ⇩s⇩b, (), True, acquired True sb 𝒪⇩s⇩b ∪ A - R,Map.empty)],
(flush ?drop_sb m)(a := f θ⇩s⇩b),share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
(is " _ ⇒⇩d (?ts_A, ?m_A, ?share_A)")
by (simp add: is⇩s⇩b)
from i_bound' have i_bound'': "i < length ?ts_A"
by simp
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts_A_j: "?ts_A!j = (hd_prog p⇩j (ys @ Ghost⇩s⇩b A' L' R' W'# zs), is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (ys @ Ghost⇩s⇩b A' L' R' W'# zs)), (),𝒟⇩j,
acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by (simp add: split_suspends⇩j)
from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
by simp
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b›
j_bound''
ts⇩s⇩b_j reads_consis]
have reads_consis_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m]
have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
from a_unowned_others [rule_format, OF _ neq_i_j] j_bound ts⇩s⇩b_j
obtain a_notin_owns_j: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j" and a_unshared: "a ∉ all_shared ?take_sb⇩j"
by auto
from a_not_acquired_others [rule_format, OF _ neq_i_j] j_bound ts⇩s⇩b_j
have a_not_acquired_j: "a ∉ all_acquired sb⇩j"
by auto
from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' ts⇩s⇩b_j]
have nvo_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
have a_no_non_vol_read: "a ∉ outstanding_refs is_non_volatile_Read⇩s⇩b ?drop_sb⇩j"
proof
assume a_in_nvr:"a ∈ outstanding_refs is_non_volatile_Read⇩s⇩b ?drop_sb⇩j"
from reads_consistent_drop [OF reads_consis]
have rc: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) (flush ?take_sb⇩j m⇩s⇩b) ?drop_sb⇩j".
from non_volatile_owned_or_read_only_drop [OF nvo_j]
have nvo_j_drop: "non_volatile_owned_or_read_only True (share ?take_sb⇩j 𝒮⇩s⇩b)
(acquired True ?take_sb⇩j 𝒪⇩j)
?drop_sb⇩j"
by simp
from outstanding_refs_non_volatile_Read⇩s⇩b_all_acquired [OF rc this a_in_nvr]
have a_owns_acq_ror:
"a ∈ 𝒪⇩j ∪ all_acquired sb⇩j ∪ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by (auto dest!: acquired_all_acquired_in all_acquired_takeWhile_dropWhile_in
simp add: acquired_takeWhile_non_volatile_Write⇩s⇩b)
have a_unowned_j: "a ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
proof (cases "a ∈ 𝒪⇩j")
case False with a_not_acquired_j show ?thesis by auto
next
case True
from all_shared_acquired_in [OF True a_unshared] a_notin_owns_j
have False by auto thus ?thesis ..
qed
with a_owns_acq_ror
have a_ror: "a ∈ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
with read_only_reads_unowned [OF j_bound'' i_bound neq_i_j [symmetric] ts⇩s⇩b_j ts⇩s⇩b_i]
have a_unowned_sb: "a ∉ 𝒪⇩s⇩b ∪ all_acquired sb"
by auto
from sharing_consis [OF j_bound'' ts⇩s⇩b_j] sharing_consistent_append [of 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have consis_j_drop: "sharing_consistent (share ?take_sb⇩j 𝒮⇩s⇩b) (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
from read_only_reads_read_only [OF nvo_j_drop consis_j_drop] a_ror a_unowned_j
all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j] acquired_takeWhile_non_volatile_Write⇩s⇩b [of sb⇩j 𝒪⇩j]
have "a ∈ read_only (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto)
from read_only_share_all_shared [OF this] a_unshared
have "a ∈ read_only 𝒮⇩s⇩b"
by fastforce
from read_only_unacquired_share [OF read_only_unowned [OF i_bound ts⇩s⇩b_i]
weak_sharing_consis [OF i_bound ts⇩s⇩b_i] this] a_unowned_sb
have "a ∈ read_only (share sb 𝒮⇩s⇩b)"
by auto
with a_not_ro show False
by simp
qed
with reads_consistent_mem_eq_on_non_volatile_reads [OF _ subset_refl reads_consis_flush_m]
have "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A suspends⇩j"
by (auto simp add: suspends⇩j)
hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A ys"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops ys. valid_sop sop"
apply (simp only: write_sops_append )
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop h_consis]
last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b ys = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from flush_store_buffer_append [
OF j_bound'''' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_A_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_m_A_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?share_A"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs (Ghost⇩s⇩b A' L' R' W' # zs) @ is⇩s⇩b⇩j =
is⇩j' @ prog_instrs (Ghost⇩s⇩b A' L' R' W'# zs)" and
steps_ys: "(?ts_A, ?m_A, ?share_A) ⇒⇩d⇧*
(?ts_A[j:= (last_prog (hd_prog p⇩j (Ghost⇩s⇩b A' L' R' W'# zs)) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Ghost⇩s⇩b A' L' R' W'# zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j') ],
flush ys ?m_A,
share ys ?share_A)"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto)
note conflict_computation = rtranclp_trans [OF rtranclp_r_rtranclp [OF steps_flush_sb, OF store_step] steps_ys]
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Ghost⇩s⇩b A' L' R' W']) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Ghost A' L' R' W'#is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from j_bound'''
have j_bound_ys: "j < length ?ts_ys"
by auto
from j_bound_ys neq_i_j
have ts_ys_j: "?ts_ys!j=(last_prog (hd_prog p⇩j (Ghost⇩s⇩b A' L' R' W'# zs)) ys, is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Write⇩s⇩b True a'' sop' v' A' L' R' W'# zs)),(),𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by auto
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j, simplified is⇩j']
have A'_unowned:
"∀i < length ?ts_ys. j≠i ⟶ (let (𝒪⇩i) = map owned ?ts_ys!i in A' ∩ 𝒪⇩i = {})"
apply cases
apply (fastforce simp add: Let_def is⇩s⇩b)+
done
from a'_in a'_A' A'_unowned [rule_format, of i] neq_i_j i_bound' A_R
show False
by (auto simp add: Let_def)
qed
qed
qed
}
thus ?thesis
by (auto simp add: Let_def)
qed
have A_no_read_only_reads_by_others:
"∀j<length (map 𝒪_sb ts⇩s⇩b). i ≠ j ⟶
(let (𝒪⇩j, sb⇩j) = map 𝒪_sb ts⇩s⇩b! j
in A ∩ read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) = {})"
proof -
{
fix j 𝒪⇩j sb⇩j
assume j_bound: "j < length (map 𝒪_sb ts⇩s⇩b)"
assume neq_i_j: "i≠j"
assume ts⇩s⇩b_j: "(map 𝒪_sb ts⇩s⇩b)!j = (𝒪⇩j,sb⇩j)"
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume conflict: "A ∩ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j ≠ {}"
have False
proof -
from j_bound leq
have j_bound': "j < length (map owned ts)"
by auto
from j_bound have j_bound'': "j < length ts⇩s⇩b"
by auto
from j_bound' have j_bound''': "j < length ts"
by simp
from conflict obtain a' where
a'_in: "a' ∈ A" and
a'_in_j: "a' ∈ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
from ts_sim [rule_format, OF j_bound''] ts⇩s⇩b_j j_bound''
obtain p⇩j suspends⇩j "is⇩s⇩b⇩j" 𝒟⇩s⇩b⇩j 𝒟⇩j ℛ⇩j θ⇩s⇩b⇩j "is⇩j" where
ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j,is⇩s⇩b⇩j, θ⇩s⇩b⇩j, sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)" and
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(), 𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
apply (cases "ts⇩s⇩b!j")
apply (force simp add: Let_def)
done
from split_in_read_only_reads [OF a'_in_j [simplified suspends⇩j [symmetric]]]
obtain t v ys zs where
split_suspends⇩j: "suspends⇩j = ys @ Read⇩s⇩b False a' t v# zs" (is "suspends⇩j = ?suspends") and
a'_unacq: "a' ∉ acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j)"
by blast
from direct_memop_step.WriteVolatile [where θ=θ⇩s⇩b and m="flush ?drop_sb m"]
have "(Write True a (D, f) A L R W# is⇩s⇩b',
θ⇩s⇩b, (), flush ?drop_sb m, 𝒟⇩s⇩b,acquired True sb 𝒪⇩s⇩b,
release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b, share ?drop_sb 𝒮) →
(is⇩s⇩b', θ⇩s⇩b, (), (flush ?drop_sb m)(a := f θ⇩s⇩b), True, acquired True sb 𝒪⇩s⇩b ∪ A - R,Map.empty,
share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF
i_bound_ts' [simplified is⇩s⇩b] ts'_i [simplified is⇩s⇩b] this [simplified is⇩s⇩b]]
have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d
(?ts'[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b, (),
True, acquired True sb 𝒪⇩s⇩b ∪ A - R,Map.empty)],
(flush ?drop_sb m)(a := f θ⇩s⇩b),share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
(is " _ ⇒⇩d (?ts_A, ?m_A, ?share_A)")
by (simp add: is⇩s⇩b)
from i_bound' have i_bound'': "i < length ?ts_A"
by simp
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts_A_j: "?ts_A!j = (hd_prog p⇩j (ys @ Read⇩s⇩b False a' t v# zs), is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (ys @ Read⇩s⇩b False a' t v# zs)), (), 𝒟⇩j,
acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by (simp add: split_suspends⇩j)
from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
by simp
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound''
ts⇩s⇩b_j reads_consis]
have reads_consis_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m]
have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
from a_unowned_others [rule_format, OF j_bound'' neq_i_j ] j_bound ts⇩s⇩b_j
obtain a_notin_owns_j: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j" and a_unshared: "a ∉ all_shared ?take_sb⇩j"
by auto
from a_not_acquired_others [rule_format, OF j_bound neq_i_j] j_bound ts⇩s⇩b_j
have a_not_acquired_j: "a ∉ all_acquired sb⇩j"
by auto
from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' ts⇩s⇩b_j]
have nvo_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
have a_no_non_vol_read: "a ∉ outstanding_refs is_non_volatile_Read⇩s⇩b ?drop_sb⇩j"
proof
assume a_in_nvr:"a ∈ outstanding_refs is_non_volatile_Read⇩s⇩b ?drop_sb⇩j"
from reads_consistent_drop [OF reads_consis]
have rc: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) (flush ?take_sb⇩j m⇩s⇩b) ?drop_sb⇩j".
from non_volatile_owned_or_read_only_drop [OF nvo_j]
have nvo_j_drop: "non_volatile_owned_or_read_only True (share ?take_sb⇩j 𝒮⇩s⇩b)
(acquired True ?take_sb⇩j 𝒪⇩j)
?drop_sb⇩j"
by simp
from outstanding_refs_non_volatile_Read⇩s⇩b_all_acquired [OF rc this a_in_nvr]
have a_owns_acq_ror:
"a ∈ 𝒪⇩j ∪ all_acquired sb⇩j ∪ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by (auto dest!: acquired_all_acquired_in all_acquired_takeWhile_dropWhile_in
simp add: acquired_takeWhile_non_volatile_Write⇩s⇩b)
have a_unowned_j: "a ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
proof (cases "a ∈ 𝒪⇩j")
case False with a_not_acquired_j show ?thesis by auto
next
case True
from all_shared_acquired_in [OF True a_unshared] a_notin_owns_j
have False by auto thus ?thesis ..
qed
with a_owns_acq_ror
have a_ror: "a ∈ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
with read_only_reads_unowned [OF j_bound'' i_bound neq_i_j [symmetric] ts⇩s⇩b_j ts⇩s⇩b_i]
have a_unowned_sb: "a ∉ 𝒪⇩s⇩b ∪ all_acquired sb"
by auto
from sharing_consis [OF j_bound'' ts⇩s⇩b_j] sharing_consistent_append [of 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have consis_j_drop: "sharing_consistent (share ?take_sb⇩j 𝒮⇩s⇩b) (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
from read_only_reads_read_only [OF nvo_j_drop consis_j_drop] a_ror a_unowned_j
all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j] acquired_takeWhile_non_volatile_Write⇩s⇩b [of sb⇩j 𝒪⇩j]
have "a ∈ read_only (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto)
from read_only_share_all_shared [OF this] a_unshared
have "a ∈ read_only 𝒮⇩s⇩b"
by fastforce
from read_only_unacquired_share [OF read_only_unowned [OF i_bound ts⇩s⇩b_i]
weak_sharing_consis [OF i_bound ts⇩s⇩b_i] this] a_unowned_sb
have "a ∈ read_only (share sb 𝒮⇩s⇩b)"
by auto
with a_not_ro show False
by simp
qed
with reads_consistent_mem_eq_on_non_volatile_reads [OF _ subset_refl reads_consis_flush_m]
have "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A suspends⇩j"
by (auto simp add: suspends⇩j)
hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A ys"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops ys. valid_sop sop"
apply (simp only: write_sops_append )
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop h_consis]
last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b ys = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from flush_store_buffer_append [
OF j_bound'''' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_A_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_m_A_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?share_A"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs (Read⇩s⇩b False a' t v# zs) @ is⇩s⇩b⇩j =
is⇩j' @ prog_instrs (Read⇩s⇩b False a' t v# zs)" and
steps_ys: "(?ts_A, ?m_A, ?share_A) ⇒⇩d⇧*
(?ts_A[j:= (last_prog (hd_prog p⇩j (Read⇩s⇩b False a' t v# zs)) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Read⇩s⇩b False a' t v# zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j') ],
flush ys ?m_A,
share ys ?share_A)"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto)
note conflict_computation = rtranclp_trans [OF rtranclp_r_rtranclp [OF steps_flush_sb, OF store_step] steps_ys]
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Read⇩s⇩b False a' t v]) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Read False a' t#is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from j_bound'''
have j_bound_ys: "j < length ?ts_ys"
by auto
from j_bound_ys neq_i_j
have ts_ys_j: "?ts_ys!j=(last_prog (hd_prog p⇩j (Read⇩s⇩b False a' t v# zs)) ys, is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Read⇩s⇩b False a' t v# zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by auto
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j, simplified is⇩j']
have "a' ∈ acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∨
a' ∈ read_only (share ys (share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))"
apply cases
apply (auto simp add: Let_def is⇩s⇩b)
done
with a'_unacq
have a'_ro: "a' ∈ read_only (share ys (share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))"
by auto
from a'_in
have a'_not_ro: "a' ∉ read_only (share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs)
have "a' ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
proof -
{
assume a_notin: "a' ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
from weak_sharing_consis [OF j_bound'' ts⇩s⇩b_j]
have "weak_sharing_consistent 𝒪⇩j sb⇩j".
with weak_sharing_consistent_append [of 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have "weak_sharing_consistent (acquired True ?take_sb⇩j 𝒪⇩j) suspends⇩j"
by (auto simp add: suspends⇩j)
with split_suspends⇩j
have weak_consis: "weak_sharing_consistent (acquired True ?take_sb⇩j 𝒪⇩j) ys"
by (simp add: weak_sharing_consistent_append)
from all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have "all_acquired ys ⊆ all_acquired sb⇩j"
apply (clarsimp)
apply (clarsimp simp add: suspends⇩j [symmetric] split_suspends⇩j all_acquired_append)
done
with a_notin acquired_takeWhile_non_volatile_Write⇩s⇩b [of sb⇩j 𝒪⇩j]
all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have "a' ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j ∪ all_acquired ys"
by auto
from read_only_share_unowned [OF weak_consis this a'_ro]
have "a' ∈ read_only (share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)" .
with a'_not_ro have False
by auto
}
thus ?thesis by blast
qed
moreover
from A_unaquired_by_others [rule_format, OF j_bound neq_i_j] ts⇩s⇩b_j j_bound
have "A ∩ all_acquired sb⇩j = {}"
by (auto simp add: Let_def)
moreover
from A_unowned_by_others [rule_format, OF j_bound'' neq_i_j] ts⇩s⇩b_j j_bound
have "A ∩ 𝒪⇩j = {}"
by (auto simp add: Let_def dest: all_shared_acquired_in)
moreover note a'_in
ultimately
show False
by auto
qed
}
thus ?thesis
by (auto simp add: Let_def)
qed
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof -
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b (sb @ [Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W]) "
by (auto simp add: non_volatile_owned_or_read_only_append)
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof (unfold_locales)
fix i⇩1 j p⇩1 "is⇩1" 𝒪⇩1 ℛ⇩1 𝒟⇩1 xs⇩1 sb⇩1 p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume i⇩1_bound: "i⇩1 < length ts⇩s⇩b'"
assume j_bound: "j < length ts⇩s⇩b'"
assume i⇩1_j: "i⇩1 ≠ j"
assume ts_i⇩1: "ts⇩s⇩b'!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
assume ts_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}"
proof (cases "i⇩1=i")
case True
with i⇩1_j have i_j: "i≠j"
by simp
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
hence j_bound'': "j < length (map owned ts⇩s⇩b)"
by simp
from ts_j i_j have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from a_unowned_others [rule_format, OF _ i_j] i_j ts_j j_bound
obtain a_notin_j: "a ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j" and
a_unshared: "a ∉ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
by (auto simp add: Let_def ts⇩s⇩b')
from a_not_acquired_others [rule_format, OF _ i_j] i_j ts_j j_bound
have a_notin_acq: "a ∉ all_acquired sb⇩j"
by (auto simp add: Let_def ts⇩s⇩b')
from outstanding_volatile_writes_unowned_by_others
[OF i_bound j_bound' i_j ts⇩s⇩b_i ts_j']
have "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb = {}".
with ts_i⇩1 a_notin_j a_unshared a_notin_acq True i_bound show ?thesis
by (auto simp add: ts⇩s⇩b' sb' outstanding_refs_append
acquired_takeWhile_non_volatile_Write⇩s⇩b dest: all_shared_acquired_in)
next
case False
note i⇩1_i = this
from i⇩1_bound have i⇩1_bound': "i⇩1 < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_i⇩1 False have ts_i⇩1': "ts⇩s⇩b!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
by (simp add: ts⇩s⇩b')
show ?thesis
proof (cases "j=i")
case True
from i⇩1_bound'
have i⇩1_bound'': "i⇩1 < length (map owned ts⇩s⇩b)"
by simp
from outstanding_volatile_writes_unowned_by_others
[OF i⇩1_bound' i_bound i⇩1_i ts_i⇩1' ts⇩s⇩b_i]
have "(𝒪⇩s⇩b ∪ all_acquired sb) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}".
moreover
from A_unused_by_others [rule_format, OF _ False [symmetric]] False ts_i⇩1 i⇩1_bound
have "A ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}"
by (auto simp add: Let_def ts⇩s⇩b')
ultimately
show ?thesis
using ts_j True ts⇩s⇩b'
by (auto simp add: i_bound ts⇩s⇩b' 𝒪⇩s⇩b' sb' all_acquired_append)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_j False have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from outstanding_volatile_writes_unowned_by_others
[OF i⇩1_bound' j_bound' i⇩1_j ts_i⇩1' ts_j']
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}" .
qed
qed
qed
next
show "ownership_distinct ts⇩s⇩b'"
proof -
have "∀j<length ts⇩s⇩b. i ≠ j ⟶
(let (p⇩j, is⇩j, θ⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j,ℛ⇩j) = ts⇩s⇩b ! j
in (𝒪⇩s⇩b ∪ all_acquired sb') ∩ (𝒪⇩j ∪ all_acquired sb⇩j) = {})"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j acq⇩j θ⇩j sb⇩j
assume neq_i_j: "i ≠ j"
assume j_bound: "j < length ts⇩s⇩b"
assume ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j, is⇩j, θ⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j,ℛ⇩j)"
have "(𝒪⇩s⇩b ∪ all_acquired sb') ∩ (𝒪⇩j ∪ all_acquired sb⇩j) = {}"
proof -
{
fix a'
assume a'_in_i: "a' ∈ (𝒪⇩s⇩b ∪ all_acquired sb')"
assume a'_in_j: "a' ∈ (𝒪⇩j ∪ all_acquired sb⇩j)"
have False
proof -
from a'_in_i have "a' ∈ (𝒪⇩s⇩b ∪ all_acquired sb) ∨ a' ∈ A"
by (simp add: sb' all_acquired_append)
then show False
proof
assume "a' ∈ (𝒪⇩s⇩b ∪ all_acquired sb)"
with ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j] a'_in_j
show ?thesis
by auto
next
assume "a' ∈ A"
moreover
have j_bound': "j < length (map owned ts⇩s⇩b)"
using j_bound by auto
from A_unowned_by_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
obtain "A ∩ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j = {}" and
"A ∩ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) = {}"
by (auto simp add: Let_def)
moreover
from A_unaquired_by_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
have "A ∩ all_acquired sb⇩j = {}"
by auto
ultimately
show ?thesis
using a'_in_j
by (auto dest: all_shared_acquired_in)
qed
qed
}
then show ?thesis by auto
qed
}
then show ?thesis by (fastforce simp add: Let_def)
qed
from ownership_distinct_nth_update [OF i_bound ts⇩s⇩b_i this]
show ?thesis by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb')
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof
fix n m
fix p⇩n "is⇩n" 𝒪⇩n ℛ⇩n 𝒟⇩n θ⇩n sb⇩n p⇩m "is⇩m" 𝒪⇩m ℛ⇩m 𝒟⇩m θ⇩m sb⇩m
assume n_bound: "n < length ts⇩s⇩b'"
and m_bound: "m < length ts⇩s⇩b'"
and neq_n_m: "n≠m"
and nth: "ts⇩s⇩b'!n = (p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
and mth: "ts⇩s⇩b'!m =(p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
from n_bound have n_bound': "n < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
from m_bound have m_bound': "m < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
show "(𝒪⇩m ∪ all_acquired sb⇩m) ∩
read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) 𝒪⇩n)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) =
{}"
proof (cases "m=i")
case True
with neq_n_m have neq_n_i: "n≠i"
by auto
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
note read_only_reads_unowned [OF n_bound' i_bound neq_n_i nth' ts⇩s⇩b_i]
moreover
from A_no_read_only_reads_by_others [rule_format, OF _ neq_n_i [symmetric]] n_bound' nth'
have "A ∩ read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) 𝒪⇩n)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) =
{}"
by auto
ultimately
show ?thesis
using True ts⇩s⇩b_i nth' mth n_bound' m_bound'
by (auto simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb' all_acquired_append)
next
case False
note neq_m_i = this
with m_bound mth i_bound have mth': "ts⇩s⇩b!m = (p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
by (auto simp add: ts⇩s⇩b')
show ?thesis
proof (cases "n=i")
case True
note read_only_reads_unowned [OF i_bound m_bound' neq_m_i [symmetric] ts⇩s⇩b_i mth']
then show ?thesis
using True neq_m_i ts⇩s⇩b_i nth mth n_bound' m_bound'
apply (case_tac "outstanding_refs (is_volatile_Write⇩s⇩b) sb = {}")
apply (clarsimp simp add: outstanding_vol_write_take_drop_appends
acquired_append read_only_reads_append ts⇩s⇩b' sb' 𝒪⇩s⇩b')+
done
next
case False
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
from read_only_reads_unowned [OF n_bound' m_bound' neq_n_m nth' mth'] False neq_m_i
show ?thesis
by (clarsimp)
qed
qed
qed
qed
have valid_hist': "valid_history program_step ts⇩s⇩b'"
proof -
from valid_history [OF i_bound ts⇩s⇩b_i]
have "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b sb) sb".
with valid_write_sops [OF i_bound ts⇩s⇩b_i] D_subset
valid_implies_valid_prog_hd [OF i_bound ts⇩s⇩b_i valid]
have "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b (sb@[Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W]))
(sb@ [Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W])"
apply -
apply (rule history_consistent_appendI)
apply (auto simp add: hd_prog_append_Write⇩s⇩b)
done
from valid_history_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' θ⇩s⇩b')
qed
have valid_reads': "valid_reads m⇩s⇩b ts⇩s⇩b'"
proof -
from valid_reads [OF i_bound ts⇩s⇩b_i]
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b sb" .
from reads_consistent_snoc_Write⇩s⇩b [OF this]
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b (sb @ [Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W])".
from valid_reads_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b')
qed
have valid_sharing': "valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
from outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i]
have "non_volatile_writes_unshared 𝒮⇩s⇩b (sb @ [Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W])"
by (auto simp add: non_volatile_writes_unshared_append)
from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
show "outstanding_non_volatile_writes_unshared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' 𝒮⇩s⇩b')
next
from sharing_consis [OF i_bound ts⇩s⇩b_i]
have consis': "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b sb".
from A_shared_owned
have "A ⊆ dom (share ?drop_sb 𝒮) ∪ acquired True sb 𝒪⇩s⇩b"
by (simp add: sharing_consistent_append acquired_takeWhile_non_volatile_Write⇩s⇩b)
moreover have "dom (share ?drop_sb 𝒮) ⊆ dom 𝒮 ∪ dom (share sb 𝒮⇩s⇩b)"
proof
fix a'
assume a'_in: "a' ∈ dom (share ?drop_sb 𝒮)"
from share_unshared_in [OF a'_in]
show "a' ∈ dom 𝒮 ∪ dom (share sb 𝒮⇩s⇩b)"
proof
assume "a' ∈ dom (share ?drop_sb Map.empty)"
from share_mono_in [OF this] share_append [of ?take_sb ?drop_sb]
have "a' ∈ dom (share sb 𝒮⇩s⇩b)"
by auto
thus ?thesis
by simp
next
assume "a' ∈ dom 𝒮 ∧ a' ∉ all_unshared ?drop_sb"
thus ?thesis by auto
qed
qed
ultimately
have A_subset: "A ⊆ dom 𝒮 ∪ dom (share sb 𝒮⇩s⇩b) ∪ acquired True sb 𝒪⇩s⇩b"
by auto
with A_unowned_by_others
have "A ⊆ dom (share sb 𝒮⇩s⇩b) ∪ acquired True sb 𝒪⇩s⇩b"
proof -
{
fix x
assume x_A: "x ∈ A"
have "x ∈ dom (share sb 𝒮⇩s⇩b) ∪ acquired True sb 𝒪⇩s⇩b"
proof -
{
assume "x ∈ dom 𝒮"
from share_all_until_volatile_write_share_acquired [OF ‹sharing_consis 𝒮⇩s⇩b ts⇩s⇩b›
i_bound ts⇩s⇩b_i this [simplified 𝒮]]
A_unowned_by_others x_A
have ?thesis
by (fastforce simp add: Let_def)
}
with A_subset show ?thesis using x_A by auto
qed
}
thus ?thesis by blast
qed
with consis' L_subset A_R R_acq
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b (sb @ [Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W])"
by (simp add: sharing_consistent_append acquired_takeWhile_non_volatile_Write⇩s⇩b)
from sharing_consis_nth_update [OF i_bound this]
show "sharing_consis 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb' 𝒮⇩s⇩b')
next
from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound ts⇩s⇩b_i] ]
show "read_only_unowned 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' 𝒪⇩s⇩b')
next
from unowned_shared_nth_update [OF i_bound ts⇩s⇩b_i subset_refl]
show "unowned_shared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from a_not_ro no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
have "no_write_to_read_only_memory 𝒮⇩s⇩b (sb @ [Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W])"
by (simp add: no_write_to_read_only_memory_append)
from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
show "no_outstanding_write_to_read_only_memory 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' sb')
qed
have tmps_distinct': "tmps_distinct ts⇩s⇩b'"
proof (intro_locales)
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_load_tmps is⇩s⇩b'" by (simp add: "is⇩s⇩b")
from load_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b')
next
from read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_read_tmps (sb @ [Write⇩s⇩b True a (D, f) (f θ⇩s⇩b) A L R W])"
by (auto simp add: distinct_read_tmps_append)
from read_tmps_distinct_nth_update [OF i_bound this]
show "read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb')
next
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ read_tmps (sb @ [Write⇩s⇩b True a (D, f) (f θ⇩s⇩b) A L R W]) ={}"
by (auto simp add: read_tmps_append "is⇩s⇩b")
from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb')
qed
have valid_sops': "valid_sops ts⇩s⇩b'"
proof -
from valid_store_sops [OF i_bound ts⇩s⇩b_i]
obtain valid_Df: "valid_sop (D,f)" and
valid_store_sops': "∀sop∈store_sops is⇩s⇩b'. valid_sop sop"
by (auto simp add: "is⇩s⇩b")
from valid_Df valid_write_sops [OF i_bound ts⇩s⇩b_i]
have valid_write_sops': "∀sop∈write_sops (sb@ [Write⇩s⇩b True a (D, f) (f θ⇩s⇩b) A L R W]).
valid_sop sop"
by (auto simp add: write_sops_append)
from valid_sops_nth_update [OF i_bound valid_write_sops' valid_store_sops']
show ?thesis by (simp add: ts⇩s⇩b' sb')
qed
have valid_dd': "valid_data_dependency ts⇩s⇩b'"
proof -
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
obtain D_indep: "D ∩ load_tmps is⇩s⇩b' = {}" and
dd_is: "data_dependency_consistent_instrs (dom θ⇩s⇩b') is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b" θ⇩s⇩b')
from load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i] D_indep
have "load_tmps is⇩s⇩b' ∩ ⋃(fst ` write_sops (sb@ [Write⇩s⇩b True a (D, f) (f θ⇩s⇩b) A L R W])) ={}"
by (auto simp add: write_sops_append "is⇩s⇩b")
from valid_data_dependency_nth_update [OF i_bound dd_is this]
show ?thesis by (simp add: ts⇩s⇩b' sb')
qed
have load_tmps_fresh': "load_tmps_fresh ts⇩s⇩b'"
proof -
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ dom θ⇩s⇩b = {}"
by (auto simp add: "is⇩s⇩b")
from load_tmps_fresh_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' θ⇩s⇩b')
qed
have enough_flushs': "enough_flushs ts⇩s⇩b'"
proof -
from clean_no_outstanding_volatile_Write⇩s⇩b [OF i_bound ts⇩s⇩b_i]
have "¬ True ⟶ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W]) = {}"
by (auto simp add: outstanding_refs_append )
from enough_flushs_nth_update [OF i_bound this]
show ?thesis
by (simp add: ts⇩s⇩b' sb' 𝒟⇩s⇩b')
qed
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
from valid_program_history [OF i_bound ts⇩s⇩b_i]
have "causal_program_history is⇩s⇩b sb" .
then have causal': "causal_program_history is⇩s⇩b' (sb@[Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W])"
by (auto simp: causal_program_history_Write "is⇩s⇩b")
from valid_last_prog [OF i_bound ts⇩s⇩b_i]
have "last_prog p⇩s⇩b sb = p⇩s⇩b".
hence "last_prog p⇩s⇩b (sb @ [Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W]) = p⇩s⇩b"
by (simp add: last_prog_append_Write⇩s⇩b)
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b' sb')
qed
show ?thesis
proof (cases "outstanding_refs is_volatile_Write⇩s⇩b sb = {}")
case True
from True have flush_all: "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = sb"
by (auto simp add: outstanding_refs_conv)
from True have suspend_nothing: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = []"
by (auto simp add: outstanding_refs_conv)
hence suspends_empty: "suspends = []"
by (simp add: suspends)
from suspends_empty is_sim have "is": "is = Write True a (D,f) A L R W# is⇩s⇩b'"
by (simp add: "is⇩s⇩b")
with suspends_empty ts_i
have ts_i: "ts!i = (p⇩s⇩b, Write True a (D,f) A L R W# is⇩s⇩b', θ⇩s⇩b,(),𝒟, acquired True ?take_sb 𝒪⇩s⇩b, release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
have "(ts, m, 𝒮) ⇒⇩d⇧* (ts, m, 𝒮)" by auto
moreover
note flush_commute =
flush_all_until_volatile_write_append_volatile_write_commute
[OF True i_bound ts⇩s⇩b_i]
from True
have drop_app: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b)
(sb@[Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W]) =
[Write⇩s⇩b True a (D,f) (f θ⇩s⇩b) A L R W]"
by (auto simp add: outstanding_refs_conv)
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b') ∼ (ts,m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute ts⇩s⇩b' θ⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b' sb')
using share_all_until_volatile_write_Write_commute
[OF i_bound ts⇩s⇩b_i [simplified is⇩s⇩b]]
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b')
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim ts_i
apply (clarsimp simp add: Let_def nth_list_update drop_app
ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' 𝒮⇩s⇩b' θ⇩s⇩b' 𝒟⇩s⇩b' outstanding_refs_append takeWhile_tail flush_all
split: if_split_asm )
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct'
valid_sops'
valid_dd' load_tmps_fresh' enough_flushs'
valid_program_history' valid' m⇩s⇩b' 𝒮⇩s⇩b'
by auto
next
case False
then obtain r where r_in: "r ∈ set sb" and volatile_r: "is_volatile_Write⇩s⇩b r"
by (auto simp add: outstanding_refs_conv)
from takeWhile_dropWhile_real_prefix
[OF r_in, of "(Not ∘ is_volatile_Write⇩s⇩b)", simplified, OF volatile_r]
obtain a' v' sb'' A'' L'' R'' W'' sop' where
sb_split: "sb = takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ Write⇩s⇩b True a' sop' v' A'' L'' R'' W''# sb''"
and
drop: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = Write⇩s⇩b True a' sop' v' A'' L'' R'' W''# sb''"
apply (auto)
subgoal for y ys
apply (case_tac y)
apply auto
done
done
from drop suspends have suspends: "suspends = Write⇩s⇩b True a' sop' v' A'' L'' R'' W''# sb''"
by simp
have "(ts, m, 𝒮) ⇒⇩d⇧* (ts, m, 𝒮)" by auto
moreover
note flush_commute =
flush_all_until_volatile_write_append_unflushed [OF False i_bound ts⇩s⇩b_i]
have "Write⇩s⇩b True a' sop' v' A'' L'' R'' W'' ∈ set sb"
by (subst sb_split) auto
note drop_app = dropWhile_append1
[OF this, of "(Not ∘ is_volatile_Write⇩s⇩b)", simplified]
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b') ∼ (ts,m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute ts⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b' sb')
using share_all_until_volatile_write_Write_commute
[OF i_bound ts⇩s⇩b_i [simplified is⇩s⇩b]]
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b')
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim ts_i is_sim
apply (clarsimp simp add: Let_def nth_list_update is_sim drop_app
read_tmps_append suspends
prog_instrs_append_Write⇩s⇩b instrs_append_Write⇩s⇩b hd_prog_append_Write⇩s⇩b
drop "is⇩s⇩b" ts⇩s⇩b' sb' 𝒪⇩s⇩b' 𝒮⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b' 𝒟⇩s⇩b' outstanding_refs_append takeWhile_tail release_append split: if_split_asm)
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' valid_dd'
valid_sops' load_tmps_fresh' enough_flushs'
valid_program_history' valid' m⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply )
qed
next
case SBHFence
then obtain
"is⇩s⇩b": "is⇩s⇩b = Fence # is⇩s⇩b'" and
sb: "sb=[]" and
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b" and
ℛ⇩s⇩b': "ℛ⇩s⇩b'=Map.empty" and
θ⇩s⇩b': "θ⇩s⇩b' = θ⇩s⇩b" and
𝒟⇩s⇩b': "¬ 𝒟⇩s⇩b'" and
sb': "sb'=sb" and
m⇩s⇩b': "m⇩s⇩b' = m⇩s⇩b" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b"
by auto
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof -
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b []"
by simp
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' 𝒮⇩s⇩b')
qed
next
from outstanding_volatile_writes_unowned_by_others_store_buffer
[OF i_bound ts⇩s⇩b_i subset_refl]
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
next
from read_only_reads_unowned_nth_update [OF i_bound ts⇩s⇩b_i, of "[]" 𝒪⇩s⇩b]
show "read_only_reads_unowned ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
next
from ownership_distinct_instructions_read_value_store_buffer_independent
[OF i_bound ts⇩s⇩b_i]
show "ownership_distinct ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have valid_hist': "valid_history program_step ts⇩s⇩b'"
proof -
from valid_history [OF i_bound ts⇩s⇩b_i]
have "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b []) []" by simp
from valid_history_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' θ⇩s⇩b')
qed
have valid_reads': "valid_reads m⇩s⇩b ts⇩s⇩b'"
proof -
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b []" by simp
from valid_reads_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have valid_sharing': "valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
have "non_volatile_writes_unshared 𝒮⇩s⇩b []"
by (simp add: sb)
from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
show "outstanding_non_volatile_writes_unshared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb sb' 𝒮⇩s⇩b')
next
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b []" by simp
from sharing_consis_nth_update [OF i_bound this]
show "sharing_consis 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb' sb 𝒮⇩s⇩b')
next
from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound ts⇩s⇩b_i] ]
show "read_only_unowned 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' 𝒪⇩s⇩b')
next
from unowned_shared_nth_update [OF i_bound ts⇩s⇩b_i subset_refl]
show "unowned_shared 𝒮⇩s⇩b' ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound, of "[]"]
show "no_outstanding_write_to_read_only_memory 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' sb' sb)
qed
have tmps_distinct': "tmps_distinct ts⇩s⇩b'"
proof (intro_locales)
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_load_tmps is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b" split: instr.splits)
from load_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' "is⇩s⇩b")
next
from read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_read_tmps []" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
from read_tmps_distinct_nth_update [OF i_bound this]
show "read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
next
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ read_tmps [] = {}"
by (clarsimp)
from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have valid_sops': "valid_sops ts⇩s⇩b'"
proof -
from valid_store_sops [OF i_bound ts⇩s⇩b_i]
obtain
valid_store_sops': "∀sop∈store_sops is⇩s⇩b'. valid_sop sop"
by (auto simp add: "is⇩s⇩b" ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
from valid_sops_nth_update [OF i_bound _ valid_store_sops', where sb= "[]" ]
show ?thesis by (auto simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have valid_dd': "valid_data_dependency ts⇩s⇩b'"
proof -
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
obtain
dd_is: "data_dependency_consistent_instrs (dom θ⇩s⇩b') is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b" θ⇩s⇩b')
from load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ ⋃(fst ` write_sops []) = {}"
by (auto simp add: write_sops_append)
from valid_data_dependency_nth_update [OF i_bound dd_is this]
show ?thesis by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have load_tmps_fresh': "load_tmps_fresh ts⇩s⇩b'"
proof -
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ dom θ⇩s⇩b = {}"
by (auto simp add: "is⇩s⇩b")
from load_tmps_fresh_nth_update [OF i_bound this]
show ?thesis by (simp add: "is⇩s⇩b" ts⇩s⇩b' sb' sb θ⇩s⇩b')
qed
from enough_flushs_nth_update [OF i_bound, where sb="[]" ]
have enough_flushs': "enough_flushs ts⇩s⇩b'"
by (auto simp add: ts⇩s⇩b' sb' sb)
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
have causal': "causal_program_history is⇩s⇩b' sb'"
by (simp add: "is⇩s⇩b" sb sb')
have "last_prog p⇩s⇩b sb' = p⇩s⇩b"
by (simp add: sb' sb)
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b' sb')
qed
from is_sim have "is": "is = Fence # is⇩s⇩b'"
by (simp add: suspends sb "is⇩s⇩b")
with ts_i
have ts_i: "ts!i = (p⇩s⇩b, Fence # is⇩s⇩b', θ⇩s⇩b,(), 𝒟, acquired True ?take_sb 𝒪⇩s⇩b, release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by (simp add: suspends sb)
from direct_memop_step.Fence
have "(Fence # is⇩s⇩b',
θ⇩s⇩b, (),m,𝒟, acquired True ?take_sb 𝒪⇩s⇩b, release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b, 𝒮) →
(is⇩s⇩b', θ⇩s⇩b, (), m, False, acquired True ?take_sb 𝒪⇩s⇩b, Map.empty, 𝒮)".
from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
have "(ts, m, 𝒮) ⇒⇩d
(ts[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b, (), False, acquired True ?take_sb 𝒪⇩s⇩b,Map.empty)], m, 𝒮)".
moreover
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b') ∼ (ts[i := (p⇩s⇩b,is⇩s⇩b', θ⇩s⇩b,(), False,acquired True ?take_sb 𝒪⇩s⇩b,Map.empty)],m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' 𝒮⇩s⇩b' m
flush_all_until_volatile_nth_update_unused [OF i_bound ts⇩s⇩b_i])
using share_all_until_volatile_write_Fence_commute
[OF i_bound ts⇩s⇩b_i [simplified is⇩s⇩b sb]]
apply (clarsimp simp add: 𝒮 ts⇩s⇩b' 𝒮⇩s⇩b' is⇩s⇩b 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b' sb' sb)
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim
apply (clarsimp simp add: Let_def nth_list_update
ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' ℛ⇩s⇩b' 𝒮⇩s⇩b' 𝒟⇩s⇩b' ex_not θ⇩s⇩b'
split: if_split_asm )
done
ultimately
show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' valid_sops'
valid_dd' load_tmps_fresh' enough_flushs'
valid_program_history' valid' m⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply)
next
case (SBHRMWReadOnly cond t a D f ret A L R W)
then obtain
"is⇩s⇩b": "is⇩s⇩b = RMW a t (D,f) cond ret A L R W # is⇩s⇩b'" and
cond: "¬ (cond (θ⇩s⇩b(t↦m⇩s⇩b a)))" and
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b" and
ℛ⇩s⇩b': "ℛ⇩s⇩b'=Map.empty" and
θ⇩s⇩b': "θ⇩s⇩b' = θ⇩s⇩b(t↦m⇩s⇩b a)" and
𝒟⇩s⇩b': "¬ 𝒟⇩s⇩b'" and
sb: "sb=[]" and
sb': "sb'=[]" and
m⇩s⇩b': "m⇩s⇩b' = m⇩s⇩b" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b"
by auto
from safe_RMW_common [OF safe_memop_flush_sb [simplified is⇩s⇩b]]
obtain access_cond: "a ∈ 𝒪⇩s⇩b ∨ a ∈ dom 𝒮" and
rels_cond: " ∀j < length ts. i≠j ⟶ released (ts!j) a ≠ Some False"
by (auto simp add: 𝒮 sb)
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof -
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b []"
by simp
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' 𝒮⇩s⇩b')
qed
next
from outstanding_volatile_writes_unowned_by_others_store_buffer
[OF i_bound ts⇩s⇩b_i subset_refl]
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from read_only_reads_unowned_nth_update [OF i_bound ts⇩s⇩b_i, of "[]" 𝒪⇩s⇩b]
show "read_only_reads_unowned ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
next
from ownership_distinct_instructions_read_value_store_buffer_independent
[OF i_bound ts⇩s⇩b_i]
show "ownership_distinct ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have valid_hist': "valid_history program_step ts⇩s⇩b'"
proof -
from valid_history [OF i_bound ts⇩s⇩b_i]
have "history_consistent (θ⇩s⇩b(t↦m⇩s⇩b a)) (hd_prog p⇩s⇩b []) []" by simp
from valid_history_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' θ⇩s⇩b')
qed
have valid_reads': "valid_reads m⇩s⇩b ts⇩s⇩b'"
proof -
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b []" by simp
from valid_reads_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have valid_sharing': "valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
from outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i]
have "non_volatile_writes_unshared 𝒮⇩s⇩b []"
by (simp add: sb)
from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
show "outstanding_non_volatile_writes_unshared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb sb' 𝒮⇩s⇩b')
next
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b []" by simp
from sharing_consis_nth_update [OF i_bound this]
show "sharing_consis 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb' sb 𝒮⇩s⇩b')
next
from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound ts⇩s⇩b_i] ]
show "read_only_unowned 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' 𝒪⇩s⇩b')
next
from unowned_shared_nth_update [OF i_bound ts⇩s⇩b_i subset_refl]
show "unowned_shared 𝒮⇩s⇩b' ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound, of "[]"]
show "no_outstanding_write_to_read_only_memory 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' sb' sb)
qed
have tmps_distinct': "tmps_distinct ts⇩s⇩b'"
proof (intro_locales)
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_load_tmps is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b" split: instr.splits)
from load_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' "is⇩s⇩b")
next
from read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_read_tmps []" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
from read_tmps_distinct_nth_update [OF i_bound this]
show "read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
next
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ read_tmps [] = {}"
by (clarsimp)
from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have valid_sops': "valid_sops ts⇩s⇩b'"
proof -
from valid_store_sops [OF i_bound ts⇩s⇩b_i]
obtain
valid_store_sops': "∀sop∈store_sops is⇩s⇩b'. valid_sop sop"
by (auto simp add: "is⇩s⇩b" ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
from valid_sops_nth_update [OF i_bound _ valid_store_sops', where sb= "[]" ]
show ?thesis by (auto simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have valid_dd': "valid_data_dependency ts⇩s⇩b'"
proof -
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
obtain
dd_is: "data_dependency_consistent_instrs (dom θ⇩s⇩b') is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b" θ⇩s⇩b')
from load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ ⋃(fst ` write_sops []) = {}"
by (auto simp add: write_sops_append)
from valid_data_dependency_nth_update [OF i_bound dd_is this]
show ?thesis by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have load_tmps_fresh': "load_tmps_fresh ts⇩s⇩b'"
proof -
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have "load_tmps (RMW a t (D,f) cond ret A L R W# is⇩s⇩b') ∩ dom θ⇩s⇩b = {}"
by (simp add: "is⇩s⇩b")
moreover
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i] have "t ∉ load_tmps is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b")
ultimately have "load_tmps is⇩s⇩b' ∩ dom (θ⇩s⇩b(t ↦ m⇩s⇩b a)) = {}"
by auto
from load_tmps_fresh_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' θ⇩s⇩b')
qed
from enough_flushs_nth_update [OF i_bound, where sb="[]" ]
have enough_flushs': "enough_flushs ts⇩s⇩b'"
by (auto simp add: ts⇩s⇩b' sb' sb)
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
have causal': "causal_program_history is⇩s⇩b' sb'"
by (simp add: "is⇩s⇩b" sb sb')
have "last_prog p⇩s⇩b sb' = p⇩s⇩b"
by (simp add: sb' sb)
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b' sb')
qed
from is_sim have "is": "is = RMW a t (D,f) cond ret A L R W# is⇩s⇩b'"
by (simp add: suspends sb "is⇩s⇩b")
with ts_i
have ts_i: "ts!i = (p⇩s⇩b, RMW a t (D,f) cond ret A L R W# is⇩s⇩b', θ⇩s⇩b,(),
𝒟, acquired True ?take_sb 𝒪⇩s⇩b, release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by (simp add: suspends sb)
have "flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b a = m⇩s⇩b a"
proof -
have "∀j < length ts⇩s⇩b. i ≠ j ⟶
(let (_,_,_,sb⇩j,_,_,_) = ts⇩s⇩b!j
in a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j))"
proof -
{
fix j p⇩j "is⇩j" 𝒪⇩j ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume neq_i_j: "i ≠ j"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j, xs⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j,ℛ⇩j)"
have "a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
proof
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume a_in: "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b ?take_sb⇩j"
with outstanding_refs_takeWhile [where P'= "Not ∘ is_volatile_Write⇩s⇩b"]
have a_in': "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
by auto
with non_volatile_owned_or_read_only_outstanding_non_volatile_writes
[OF outstanding_non_volatile_refs_owned_or_read_only [OF j_bound jth]]
have j_owns: "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
from rels_cond [rule_format, OF j_bound [simplified leq] neq_i_j] ts_sim [rule_format, OF j_bound] jth
have no_unsharing:"release ?take_sb⇩j (dom (𝒮⇩s⇩b)) ℛ⇩j a ≠ Some False"
by (auto simp add: Let_def)
from access_cond
show False
proof
assume "a ∈ 𝒪⇩s⇩b"
with ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
j_owns
show False
by auto
next
assume a_shared: "a ∈ dom 𝒮"
with share_all_until_volatile_write_thread_local [OF ownership_distinct_ts⇩s⇩b sharing_consis_ts⇩s⇩b j_bound jth j_owns]
have a_dom: "a ∈ dom (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto simp add: 𝒮 domIff)
from outstanding_non_volatile_writes_unshared [OF j_bound jth]
have "non_volatile_writes_unshared 𝒮⇩s⇩b sb⇩j".
with non_volatile_writes_unshared_append [of 𝒮⇩s⇩b "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
have unshared_take: "non_volatile_writes_unshared 𝒮⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
by clarsimp
from release_not_unshared_no_write_take [OF unshared_take no_unsharing a_dom] a_in
show False by auto
qed
qed
}
thus ?thesis
by (fastforce simp add: Let_def)
qed
from flush_all_until_volatile_write_buffered_val_conv
[OF _ i_bound ts⇩s⇩b_i this]
show ?thesis
by (simp add: sb)
qed
hence m_a: "m a = m⇩s⇩b a"
by (simp add: m)
from cond have cond': "¬ cond (θ⇩s⇩b(t ↦ m a))"
by (simp add: m_a)
from direct_memop_step.RMWReadOnly [where cond=cond and θ=θ⇩s⇩b and m=m, OF cond']
have "(RMW a t (D, f) cond ret A L R W # is⇩s⇩b',
θ⇩s⇩b, (),m, 𝒟, 𝒪⇩s⇩b, ℛ⇩s⇩b, 𝒮) →
(is⇩s⇩b', θ⇩s⇩b(t ↦ m a), (), m, False, 𝒪⇩s⇩b, Map.empty, 𝒮)".
from direct_computation.concurrent_step.Memop [OF i_bound' ts_i [simplified sb, simplified] this]
have "(ts, m, 𝒮) ⇒⇩d (ts[i := (p⇩s⇩b, is⇩s⇩b',
θ⇩s⇩b(t ↦ m a), (), False, 𝒪⇩s⇩b,Map.empty)], m, 𝒮)".
moreover
have tmps_commute: "θ⇩s⇩b(t ↦ (m⇩s⇩b a)) =
(θ⇩s⇩b |` (dom θ⇩s⇩b - {t}))(t ↦ (m⇩s⇩b a))"
apply (rule ext)
apply (auto simp add: restrict_map_def domIff)
done
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b') ∼ (ts[i := (p⇩s⇩b,is⇩s⇩b', θ⇩s⇩b(t ↦ m a),(), False,𝒪⇩s⇩b,Map.empty)],m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' m
flush_all_until_volatile_nth_update_unused [OF i_bound ts⇩s⇩b_i, simplified sb])
using share_all_until_volatile_write_RMW_commute [OF i_bound ts⇩s⇩b_i [simplified is⇩s⇩b sb]]
apply (clarsimp simp add: 𝒮 ts⇩s⇩b' 𝒮⇩s⇩b' is⇩s⇩b 𝒪⇩s⇩b' θ⇩s⇩b' sb' sb)
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim
apply (clarsimp simp add: Let_def nth_list_update
ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' ℛ⇩s⇩b' 𝒮⇩s⇩b' θ⇩s⇩b' 𝒟⇩s⇩b' ex_not m_a
split: if_split_asm)
apply (rule tmps_commute)
done
ultimately
show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' valid_sops'
valid_dd' load_tmps_fresh' enough_flushs'
valid_program_history' valid' m⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply)
next
case (SBHRMWWrite cond t a D f ret A L R W)
then obtain
"is⇩s⇩b": "is⇩s⇩b = RMW a t (D,f) cond ret A L R W # is⇩s⇩b'" and
cond: "(cond (θ⇩s⇩b(t↦m⇩s⇩b a)))" and
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b ∪ A - R" and
ℛ⇩s⇩b': "ℛ⇩s⇩b'=Map.empty" and
𝒟⇩s⇩b': "¬ 𝒟⇩s⇩b'" and
θ⇩s⇩b': "θ⇩s⇩b' = θ⇩s⇩b(t↦ret (m⇩s⇩b a) (f (θ⇩s⇩b(t↦m⇩s⇩b a))))" and
sb: "sb=[]" and
sb': "sb'=[]" and
m⇩s⇩b': "m⇩s⇩b' = m⇩s⇩b(a := f (θ⇩s⇩b(t↦m⇩s⇩b a)))" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L"
by auto
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
have D_subset: "D ⊆ dom θ⇩s⇩b"
by (simp add: is⇩s⇩b)
from is_sim have "is": "is = RMW a t (D,f) cond ret A L R W # is⇩s⇩b'"
by (simp add: suspends sb "is⇩s⇩b")
with ts_i
have ts_i: "ts!i = (p⇩s⇩b, RMW a t (D,f) cond ret A L R W # is⇩s⇩b', θ⇩s⇩b,(), 𝒟, 𝒪⇩s⇩b,ℛ⇩s⇩b)"
by (simp add: suspends sb)
from safe_RMW_common [OF safe_memop_flush_sb [simplified is⇩s⇩b]]
obtain access_cond: "a ∈ 𝒪⇩s⇩b ∨ a ∈ dom 𝒮" and
rels_cond: " ∀j < length ts. i≠j ⟶ released (ts!j) a ≠ Some False"
by (auto simp add: 𝒮 sb)
have a_unflushed:
"∀j < length ts⇩s⇩b. i ≠ j ⟶
(let (_,_,_,sb⇩j,_,_,_) = ts⇩s⇩b!j
in a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j))"
proof -
{
fix j p⇩j "is⇩j" 𝒪⇩j ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume neq_i_j: "i ≠ j"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j, xs⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j, ℛ⇩j)"
have "a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
proof
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume a_in: "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b ?take_sb⇩j"
with outstanding_refs_takeWhile [where P'= "Not ∘ is_volatile_Write⇩s⇩b"]
have a_in': "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
by auto
with non_volatile_owned_or_read_only_outstanding_non_volatile_writes
[OF outstanding_non_volatile_refs_owned_or_read_only [OF j_bound jth]]
have j_owns: "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
with ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
have a_not_owns: "a ∉ 𝒪⇩s⇩b ∪ all_acquired sb"
by blast
assume a_in: "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b
(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
with outstanding_refs_takeWhile [where P'= "Not ∘ is_volatile_Write⇩s⇩b"]
have a_in': "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
by auto
from rels_cond [rule_format, OF j_bound [simplified leq] neq_i_j] ts_sim [rule_format, OF j_bound] jth
have no_unsharing:"release ?take_sb⇩j (dom (𝒮⇩s⇩b)) ℛ⇩j a ≠ Some False"
by (auto simp add: Let_def)
from access_cond
show False
proof
assume "a ∈ 𝒪⇩s⇩b"
with ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
j_owns
show False
by auto
next
assume a_shared: "a ∈ dom 𝒮"
with share_all_until_volatile_write_thread_local [OF ownership_distinct_ts⇩s⇩b sharing_consis_ts⇩s⇩b j_bound jth j_owns]
have a_dom: "a ∈ dom (share ?take_sb⇩j 𝒮⇩s⇩b)"
by (auto simp add: 𝒮 domIff)
from outstanding_non_volatile_writes_unshared [OF j_bound jth]
have "non_volatile_writes_unshared 𝒮⇩s⇩b sb⇩j".
with non_volatile_writes_unshared_append [of 𝒮⇩s⇩b "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
have unshared_take: "non_volatile_writes_unshared 𝒮⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
by clarsimp
from release_not_unshared_no_write_take [OF unshared_take no_unsharing a_dom] a_in
show False by auto
qed
qed
}
thus ?thesis
by (fastforce simp add: Let_def)
qed
have "flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b a = m⇩s⇩b a"
proof -
from flush_all_until_volatile_write_buffered_val_conv
[OF _ i_bound ts⇩s⇩b_i a_unflushed]
show ?thesis
by (simp add: sb)
qed
hence m_a: "m a = m⇩s⇩b a"
by (simp add: m)
from cond have cond': "cond (θ⇩s⇩b(t ↦ m a))"
by (simp add: m_a)
from safe_memop_flush_sb [simplified is⇩s⇩b] cond'
obtain
L_subset: "L ⊆ A" and
A_shared_owned: "A ⊆ dom 𝒮 ∪ 𝒪⇩s⇩b" and
R_owned: "R ⊆ 𝒪⇩s⇩b" and
A_R: "A ∩ R = {}" and
a_unowned_others_ts:
"∀j<length ts. i ≠ j ⟶ (a ∉ owned (ts!j) ∪ dom (released (ts!j)))" and
A_unowned_by_others_ts:
"∀j<length ts. i ≠ j ⟶ (A ∩ (owned (ts!j) ∪ dom (released (ts!j))) = {})" and
a_not_ro: "a ∉ read_only 𝒮"
by cases (auto simp add: sb)
from a_unowned_others_ts ts_sim leq
have a_unowned_others:
"∀j<length ts⇩s⇩b. i ≠ j ⟶
(let (_,_,_,sb⇩j,_,𝒪⇩j,_) = ts⇩s⇩b!j in
a ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j ∧
a ∉ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j))"
apply (clarsimp simp add: Let_def)
subgoal for j
apply (drule_tac x=j in spec)
apply (auto simp add: dom_release_takeWhile)
done
done
from A_unowned_by_others_ts ts_sim leq
have A_unowned_by_others:
"∀j<length ts⇩s⇩b. i≠j ⟶ (let (_,_,_,sb⇩j,_,𝒪⇩j,_) = ts⇩s⇩b!j
in A ∩ (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j ∪
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)) = {})"
apply (clarsimp simp add: Let_def)
subgoal for j
apply (drule_tac x=j in spec)
apply (force simp add: dom_release_takeWhile)
done
done
have a_not_ro': "a ∉ read_only 𝒮⇩s⇩b"
proof
assume a: "a ∈ read_only (𝒮⇩s⇩b)"
from local.read_only_unowned_axioms have "read_only_unowned 𝒮⇩s⇩b ts⇩s⇩b".
from in_read_only_share_all_until_volatile_write' [OF ownership_distinct_ts⇩s⇩b sharing_consis_ts⇩s⇩b
‹read_only_unowned 𝒮⇩s⇩b ts⇩s⇩b› i_bound ts⇩s⇩b_i a_unowned_others, simplified sb, simplified,
OF a]
have "a ∈ read_only (𝒮)"
by (simp add: 𝒮)
with a_not_ro show False by simp
qed
{
fix j
fix p⇩j is⇩s⇩b⇩j 𝒪⇩j ℛ⇩j 𝒟⇩s⇩b⇩j θ⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume ts⇩s⇩b_j: "ts⇩s⇩b!j=(p⇩j,is⇩s⇩b⇩j,θ⇩j,sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)"
assume neq_i_j: "i≠j"
have "a ∉ unforwarded_non_volatile_reads (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) {}"
proof
let ?take_sb⇩j = "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
let ?drop_sb⇩j = "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
assume a_in: "a ∈ unforwarded_non_volatile_reads ?drop_sb⇩j {}"
from a_unowned_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
obtain a_unacq_take: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j" and a_not_shared: "a ∉ all_shared ?take_sb⇩j"
by auto
note nvo_j = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound ts⇩s⇩b_j]
from non_volatile_owned_or_read_only_drop [OF nvo_j]
have nvo_drop_j: "non_volatile_owned_or_read_only True (share ?take_sb⇩j 𝒮⇩s⇩b)
(acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j" .
note consis_j = sharing_consis [OF j_bound ts⇩s⇩b_j]
with sharing_consistent_append [of 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
obtain consis_take_j: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j" and
consis_drop_j: "sharing_consistent (share ?take_sb⇩j 𝒮⇩s⇩b)
(acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
from in_unforwarded_non_volatile_reads_non_volatile_Read⇩s⇩b [OF a_in]
have a_in': "a ∈ outstanding_refs is_non_volatile_Read⇩s⇩b ?drop_sb⇩j".
note reads_consis_j = valid_reads [OF j_bound ts⇩s⇩b_j]
from reads_consistent_drop [OF this]
have reads_consis_drop_j:
"reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) (flush ?take_sb⇩j m⇩s⇩b) ?drop_sb⇩j".
from read_only_share_all_shared [of a ?take_sb⇩j 𝒮⇩s⇩b] a_not_ro' a_not_shared
have a_not_ro_j: "a ∉ read_only (share ?take_sb⇩j 𝒮⇩s⇩b)"
by auto
from ts_sim [rule_format, OF j_bound] ts⇩s⇩b_j j_bound
obtain suspends⇩j "is⇩j" 𝒟⇩j where
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩j |` (dom θ⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by (auto simp: Let_def)
from ts⇩j neq_i_j j_bound
have ts'_j: "?ts'!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩j |` (dom θ⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by auto
from valid_last_prog [OF j_bound ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
from j_bound i_bound' leq have j_bound_ts': "j < length ?ts'"
by simp
from read_only_read_acquired_unforwarded_acquire_witness [OF nvo_drop_j consis_drop_j
a_not_ro_j a_unacq_take a_in]
have False
proof
assume "∃sop a' v ys zs A L R W.
?drop_sb⇩j= ys @ Write⇩s⇩b True a' sop v A L R W # zs ∧ a ∈ A ∧
a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a'≠a"
with suspends⇩j
obtain a' sop' v' ys zs' A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Write⇩s⇩b True a' sop' v' A' L' R' W'# zs'" (is "suspends⇩j=?suspends") and
a_A': "a ∈ A'" and
no_write: "a ∉ outstanding_refs is_Write⇩s⇩b (ys @ [Write⇩s⇩b True a' sop' v' A' L' R' W'])"
by(auto simp add: outstanding_refs_append )
from last_prog
have lp: "last_prog p⇩j suspends⇩j = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from sharing_consis [OF j_bound ts⇩s⇩b_j]
have sharing_consis_j: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
then have A'_R': "A' ∩ R' = {}"
by (simp add: sharing_consistent_append [of _ _ ?take_sb⇩j ?drop_sb⇩j, simplified]
suspends⇩j [symmetric] split_suspends⇩j sharing_consistent_append)
from valid_program_history [OF j_bound ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from valid_reads [OF j_bound ts⇩s⇩b_j]
have reads_consis_j: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b›
j_bound ts⇩s⇩b_j this]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m_j]
have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
hence reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W'])"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_write_sops [OF j_bound ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis_j]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b
(ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
have acq_simp:
"acquired True (ys @ [Write⇩s⇩b True a' sop' v' A' L' R' W'])
(acquired True ?take_sb⇩j 𝒪⇩j) =
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R'"
by (simp add: acquired_append)
from flush_store_buffer_append [where sb="ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']" and sb'="zs'", simplified,
OF j_bound_ts' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts'_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="share ?drop_sb 𝒮"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs zs' @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs'" and
steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d⇧*
(?ts'[j:=(last_prog
(hd_prog p⇩j (Write⇩s⇩b True a' sop' v' A' L' R' W'# zs')) (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']),
is⇩j',
θ⇩j |` (dom θ⇩j - read_tmps zs'),
(), True, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R',ℛ⇩j')],
flush (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) (flush ?drop_sb m),
share (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) (share ?drop_sb 𝒮))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append outstanding_refs_append)
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b, θ⇩s⇩b,(),
𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
note conflict_computation = rtranclp_trans [OF steps_flush_sb steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have safe: "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from flush_unchanged_addresses [OF no_write]
have "flush (ys @ [Write⇩s⇩b True a' sop' v' A' L' R' W']) m a = m a".
with safe_delayedE [OF safe i_bound_ys ts_ys_i, simplified is⇩s⇩b] cond'
have a_unowned:
"∀j < length ?ts_ys. i≠j ⟶ (let (𝒪⇩j) = map owned ?ts_ys!j in a ∉ 𝒪⇩j)"
apply cases
apply (auto simp add: Let_def is⇩s⇩b sb)
done
from a_A' a_unowned [rule_format, of j] neq_i_j j_bound leq A'_R'
show False
by (auto simp add: Let_def)
next
assume "∃A L R W ys zs. ?drop_sb⇩j = ys @ Ghost⇩s⇩b A L R W# zs ∧ a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys"
with suspends⇩j
obtain ys zs' A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Ghost⇩s⇩b A' L' R' W'# zs'" (is "suspends⇩j=?suspends") and
a_A': "a ∈ A'" and
no_write: "a ∉ outstanding_refs is_Write⇩s⇩b (ys @ [Ghost⇩s⇩b A' L' R' W'])"
by (auto simp add: outstanding_refs_append)
from last_prog
have lp: "last_prog p⇩j suspends⇩j = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from sharing_consis [OF j_bound ts⇩s⇩b_j]
have sharing_consis_j: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
then have A'_R': "A' ∩ R' = {}"
by (simp add: sharing_consistent_append [of _ _ ?take_sb⇩j ?drop_sb⇩j, simplified]
suspends⇩j [symmetric] split_suspends⇩j sharing_consistent_append)
from valid_program_history [OF j_bound ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from valid_reads [OF j_bound ts⇩s⇩b_j]
have reads_consis_j: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b›
j_bound ts⇩s⇩b_j this]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m_j]
have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) suspends⇩j".
hence reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
(flush ?drop_sb m) (ys@[Ghost⇩s⇩b A' L' R' W'])"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_write_sops [OF j_bound ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys@[Ghost⇩s⇩b A' L' R' W']). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis_j]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b
(ys@[Ghost⇩s⇩b A' L' R' W']) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
have acq_simp:
"acquired True (ys @ [Ghost⇩s⇩b A' L' R' W'])
(acquired True ?take_sb⇩j 𝒪⇩j) =
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R'"
by (simp add: acquired_append)
from flush_store_buffer_append [where sb="ys@[Ghost⇩s⇩b A' L' R' W']" and sb'="zs'", simplified,
OF j_bound_ts' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts'_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="share ?drop_sb 𝒮"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs zs' @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs'" and
steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d⇧*
(?ts'[j:=(last_prog
(hd_prog p⇩j (Ghost⇩s⇩b A' L' R' W'# zs')) (ys@[Ghost⇩s⇩b A' L' R' W']),
is⇩j',
θ⇩j |` (dom θ⇩j - read_tmps zs'),
(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R',ℛ⇩j')],
flush (ys@[Ghost⇩s⇩b A' L' R' W']) (flush ?drop_sb m),
share (ys@[Ghost⇩s⇩b A' L' R' W']) (share ?drop_sb 𝒮))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append outstanding_refs_append)
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b, θ⇩s⇩b,(),
𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
note conflict_computation = rtranclp_trans [OF steps_flush_sb steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have safe: "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from flush_unchanged_addresses [OF no_write]
have "flush (ys @ [Ghost⇩s⇩b A' L' R' W']) m a = m a".
with safe_delayedE [OF safe i_bound_ys ts_ys_i, simplified is⇩s⇩b] cond'
have a_unowned:
"∀j < length ?ts_ys. i≠j ⟶ (let (𝒪⇩j) = map owned ?ts_ys!j in a ∉ 𝒪⇩j)"
apply cases
apply (auto simp add: Let_def is⇩s⇩b sb)
done
from a_A' a_unowned [rule_format, of j] neq_i_j j_bound leq A'_R'
show False
by (auto simp add: Let_def)
qed
then show False
by simp
qed
}
note a_notin_unforwarded_non_volatile_reads_drop = this
have A_unused_by_others:
"∀j<length (map 𝒪_sb ts⇩s⇩b). i ≠ j ⟶
(let (𝒪⇩j, sb⇩j) = map 𝒪_sb ts⇩s⇩b! j
in A ∩ (𝒪⇩j ∪ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j) = {})"
proof -
{
fix j 𝒪⇩j sb⇩j
assume j_bound: "j < length (map owned ts⇩s⇩b)"
assume neq_i_j: "i≠j"
assume ts⇩s⇩b_j: "(map 𝒪_sb ts⇩s⇩b)!j = (𝒪⇩j,sb⇩j)"
assume conflict: "A ∩ (𝒪⇩j ∪ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j) ≠ {}"
have False
proof -
from j_bound leq
have j_bound': "j < length (map owned ts)"
by auto
from j_bound have j_bound'': "j < length ts⇩s⇩b"
by auto
from j_bound' have j_bound''': "j < length ts"
by simp
from conflict obtain a' where
a_in: "a' ∈ A" and
conflict: "a' ∈ 𝒪⇩j ∨ a' ∈ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j"
by auto
from A_unowned_by_others [rule_format, OF _ neq_i_j] j_bound ts⇩s⇩b_j
have A_unshared_j: "A ∩ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) = {}"
by (auto simp add: Let_def)
from conflict
show ?thesis
proof
assume "a' ∈ 𝒪⇩j"
from all_shared_acquired_in [OF this] A_unshared_j a_in
have conflict: "a' ∈ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j"
by (auto)
with A_unowned_by_others [rule_format, OF _ neq_i_j] j_bound ts⇩s⇩b_j a_in
show False by auto
next
assume a_in_j: "a' ∈ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j"
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
from ts_sim [rule_format, OF j_bound''] ts⇩s⇩b_j j_bound''
obtain p⇩j suspends⇩j "is⇩s⇩b⇩j" 𝒟⇩s⇩b⇩j 𝒟⇩j ℛ⇩j θ⇩s⇩b⇩j "is⇩j" where
ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j,is⇩s⇩b⇩j, θ⇩s⇩b⇩j, sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)" and
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(), 𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
apply (cases "ts⇩s⇩b!j")
apply (force simp add: Let_def)
done
have "a' ∈ outstanding_refs is_volatile_Write⇩s⇩b suspends⇩j"
proof -
from a_in_j
have "a' ∈ outstanding_refs is_volatile_Write⇩s⇩b (?take_sb⇩j @ ?drop_sb⇩j)"
by simp
thus ?thesis
apply (simp only: outstanding_refs_append suspends⇩j)
apply (auto simp add: outstanding_refs_conv dest: set_takeWhileD)
done
qed
from split_volatile_Write⇩s⇩b_in_outstanding_refs [OF this]
obtain sop' v' ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Write⇩s⇩b True a' sop' v' A' L' R' W'# zs" (is "suspends⇩j = ?suspends")
by blast
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b›
j_bound''
ts⇩s⇩b_j this]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' ts⇩s⇩b_j]
have nvo_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
with non_volatile_owned_or_read_only_append [of False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have nvo_take_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j"
by auto
from a_unowned_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
have a_not_acq: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j"
by auto
from a_notin_unforwarded_non_volatile_reads_drop[OF j_bound'' ts⇩s⇩b_j neq_i_j]
have a_notin_unforwarded_reads: "a ∉ unforwarded_non_volatile_reads suspends⇩j {}"
by (simp add: suspends⇩j)
let ?ma="(m(a := f (θ⇩s⇩b(t↦m a))))"
from reads_consistent_mem_eq_on_unforwarded_non_volatile_reads [where W="{}"
and m'="?ma",simplified, OF _ subset_refl reads_consis_m_j]
a_notin_unforwarded_reads
have reads_consis_ma_j:
"reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?ma suspends⇩j"
by auto
from reads_consis_ma_j
have reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?ma (ys)"
by (simp add: split_suspends⇩j reads_consistent_append)
from direct_memop_step.RMWWrite [where cond=cond and θ=θ⇩s⇩b and m=m, OF cond' ]
have "(RMW a t (D, f) cond ret A L R W# is⇩s⇩b', θ⇩s⇩b, (), m,𝒟, 𝒪⇩s⇩b, ℛ⇩s⇩b, 𝒮) →
(is⇩s⇩b', θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))), (), ?ma, False, 𝒪⇩s⇩b ∪ A - R, Map.empty,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
have step_a: "(ts, m, 𝒮) ⇒⇩d
(ts[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))), (), False, 𝒪⇩s⇩b ∪ A - R,Map.empty)],
?ma,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
(is " _ ⇒⇩d (?ts_a, _, ?shared_a)").
from ts⇩j neq_i_j j_bound
have ts_a_j: "?ts_a!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(), 𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom (𝒮⇩s⇩b)) ℛ⇩j)"
by auto
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b (ys) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from j_bound' have j_bound_ts_a: "j < length ?ts_a" by auto
from flush_store_buffer_append [where sb="ys" and sb'="Write⇩s⇩b True a' sop' v' A' L' R' W'#zs", simplified,
OF j_bound_ts_a is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_a_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?shared_a"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "Write True a' sop' A' L' R' W'# instrs zs @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs" and
steps_ys: "(?ts_a, ?ma, ?shared_a) ⇒⇩d⇧*
(?ts_a[j:=(last_prog
(hd_prog p⇩j zs) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs),
(), 𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')],
flush ys (?ma), share ys (?shared_a))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append)
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Write⇩s⇩b True a' sop' v' A' L' R' W']) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Write True a' sop' A' L' R' W'#is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b',
θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))),(), False, 𝒪⇩s⇩b ∪ A - R,Map.empty)"
by simp
from j_bound_ts_a have j_bound_ys: "j < length ?ts_ys"
by auto
then have ts_ys_j: "?ts_ys!j = (last_prog (hd_prog p⇩j zs) ys, Write True a' sop' A' L' R' W'#is⇩j'', θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs), (), 𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by (clarsimp simp add: is⇩j')
note conflict_computation = r_rtranclp_rtranclp [OF step_a steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j]
have a_unowned:
"∀i < length ts. j≠i ⟶ (let (𝒪⇩i) = map owned ?ts_ys!i in a' ∉ 𝒪⇩i)"
apply cases
apply (auto simp add: Let_def)
done
from a_in a_unowned [rule_format, of i] neq_i_j i_bound' A_R
show False
by (auto simp add: Let_def)
qed
qed
}
thus ?thesis
by (auto simp add: Let_def)
qed
have A_unacquired_by_others:
"∀j<length (map 𝒪_sb ts⇩s⇩b). i ≠ j ⟶
(let (𝒪⇩j, sb⇩j) = map 𝒪_sb ts⇩s⇩b! j
in A ∩ all_acquired sb⇩j = {})"
proof -
{
fix j 𝒪⇩j sb⇩j
assume j_bound: "j < length (map owned ts⇩s⇩b)"
assume neq_i_j: "i≠j"
assume ts⇩s⇩b_j: "(map 𝒪_sb ts⇩s⇩b)!j = (𝒪⇩j,sb⇩j)"
assume conflict: "A ∩ all_acquired sb⇩j ≠ {}"
have False
proof -
from j_bound leq
have j_bound': "j < length (map owned ts)"
by auto
from j_bound have j_bound'': "j < length ts⇩s⇩b"
by auto
from j_bound' have j_bound''': "j < length ts"
by simp
from conflict obtain a' where
a'_in: "a' ∈ A" and
a'_in_j: "a' ∈ all_acquired sb⇩j"
by auto
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
from ts_sim [rule_format, OF j_bound''] ts⇩s⇩b_j j_bound''
obtain p⇩j suspends⇩j "is⇩s⇩b⇩j" θ⇩s⇩b⇩j 𝒟⇩s⇩b⇩j ℛ⇩j 𝒟⇩j "is⇩j" where
ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j,is⇩s⇩b⇩j,θ⇩s⇩b⇩j, sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)" and
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
apply (cases "ts⇩s⇩b!j")
apply (force simp add: Let_def)
done
from a'_in_j all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have "a' ∈ all_acquired ?take_sb⇩j ∨ a' ∈ all_acquired suspends⇩j"
by (auto simp add: suspends⇩j)
thus False
proof
assume "a' ∈ all_acquired ?take_sb⇩j"
with A_unowned_by_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound a'_in
show False
by (auto dest: all_acquired_unshared_acquired)
next
assume conflict_drop: "a' ∈ all_acquired suspends⇩j"
from split_all_acquired_in [OF conflict_drop]
show ?thesis
proof
assume "∃sop a'' v ys zs A L R W.
suspends⇩j = ys @ Write⇩s⇩b True a'' sop v A L R W# zs ∧ a' ∈ A"
then
obtain a'' sop' v' ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W'# zs" (is "suspends⇩j = ?suspends") and
a'_A': "a' ∈ A'"
by blast
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF
‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound''
ts⇩s⇩b_j this]
have reads_consis_m_j:
"reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' ts⇩s⇩b_j]
have nvo_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
with non_volatile_owned_or_read_only_append [of False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have nvo_take_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j"
by auto
from a_unowned_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
have a_not_acq: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j"
by auto
from a_notin_unforwarded_non_volatile_reads_drop[OF j_bound'' ts⇩s⇩b_j neq_i_j]
have a_notin_unforwarded_reads: "a ∉ unforwarded_non_volatile_reads suspends⇩j {}"
by (simp add: suspends⇩j)
let ?ma="(m(a := f (θ⇩s⇩b(t↦m a))))"
from reads_consistent_mem_eq_on_unforwarded_non_volatile_reads [where W="{}"
and m'="?ma",simplified, OF _ subset_refl reads_consis_m_j]
a_notin_unforwarded_reads
have reads_consis_ma_j:
"reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?ma suspends⇩j"
by auto
from reads_consis_ma_j
have reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?ma (ys)"
by (simp add: split_suspends⇩j reads_consistent_append)
from direct_memop_step.RMWWrite [where cond=cond and θ=θ⇩s⇩b and m=m, OF cond']
have "(RMW a t (D, f) cond ret A L R W# is⇩s⇩b',
θ⇩s⇩b, (), m, 𝒟, 𝒪⇩s⇩b, ℛ⇩s⇩b, 𝒮) →
(is⇩s⇩b',
θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))), (), ?ma, False, 𝒪⇩s⇩b ∪ A - R,Map.empty, 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF i_bound' ts_i [simplified sb, simplified] this]
have step_a: "(ts, m, 𝒮) ⇒⇩d
(ts[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))), (), False, 𝒪⇩s⇩b ∪ A - R,Map.empty)],
?ma,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
(is " _ ⇒⇩d (?ts_a, _, ?shared_a)").
from ts⇩j neq_i_j j_bound
have ts_a_j: "?ts_a!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by auto
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b (ys) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from j_bound' have j_bound_ts_a: "j < length ?ts_a" by auto
from flush_store_buffer_append [where sb="ys" and sb'="Write⇩s⇩b True a'' sop' v' A' L' R' W'#zs", simplified,
OF j_bound_ts_a is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_a_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?shared_a"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "Write True a'' sop' A' L' R' W'# instrs zs @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs" and
steps_ys: "(?ts_a, ?ma, ?shared_a) ⇒⇩d⇧*
(?ts_a[j:=(last_prog
(hd_prog p⇩j zs) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs),
(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')],
flush ys (?ma),
share ys (?shared_a))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append)
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Write⇩s⇩b True a'' sop' v' A' L' R' W']) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Write True a'' sop' A' L' R' W'#is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b',
θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))),(), False, 𝒪⇩s⇩b ∪ A - R,Map.empty)"
by simp
from j_bound_ts_a have j_bound_ys: "j < length ?ts_ys"
by auto
then have ts_ys_j: "?ts_ys!j = (last_prog (hd_prog p⇩j zs) ys, Write True a'' sop' A' L' R' W'#is⇩j'',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs), (),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by (clarsimp simp add: is⇩j')
note conflict_computation = r_rtranclp_rtranclp [OF step_a steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j]
have A'_unowned:
"∀i < length ?ts_ys. j≠i ⟶ (let (𝒪⇩i) = map owned ?ts_ys!i in A' ∩ 𝒪⇩i = {})"
apply cases
apply (fastforce simp add: Let_def is⇩s⇩b)+
done
from a'_in a'_A' A'_unowned [rule_format, of i] neq_i_j i_bound' A_R
show False
by (auto simp add: Let_def)
next
assume "∃A L R W ys zs. suspends⇩j = ys @ Ghost⇩s⇩b A L R W# zs ∧ a' ∈ A"
then
obtain ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Ghost⇩s⇩b A' L' R' W'# zs" (is "suspends⇩j = ?suspends") and
a'_A': "a' ∈ A'"
by blast
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF
‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound''
ts⇩s⇩b_j this]
have reads_consis_m_j:
"reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' ts⇩s⇩b_j]
have nvo_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
with non_volatile_owned_or_read_only_append [of False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have nvo_take_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j"
by auto
from a_unowned_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
have a_not_acq: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j"
by auto
from a_notin_unforwarded_non_volatile_reads_drop[OF j_bound'' ts⇩s⇩b_j neq_i_j]
have a_notin_unforwarded_reads: "a ∉ unforwarded_non_volatile_reads suspends⇩j {}"
by (simp add: suspends⇩j)
let ?ma="(m(a := f (θ⇩s⇩b(t↦m a))))"
from reads_consistent_mem_eq_on_unforwarded_non_volatile_reads [where W="{}"
and m'="?ma",simplified, OF _ subset_refl reads_consis_m_j]
a_notin_unforwarded_reads
have reads_consis_ma_j:
"reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?ma suspends⇩j"
by auto
from reads_consis_ma_j
have reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?ma (ys)"
by (simp add: split_suspends⇩j reads_consistent_append)
from direct_memop_step.RMWWrite [where cond=cond and θ=θ⇩s⇩b and m=m, OF cond']
have "(RMW a t (D, f) cond ret A L R W# is⇩s⇩b',
θ⇩s⇩b, (), m, 𝒟,𝒪⇩s⇩b, ℛ⇩s⇩b, 𝒮) →
(is⇩s⇩b',
θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))), (), ?ma, False, 𝒪⇩s⇩b ∪ A - R,Map.empty,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF i_bound' ts_i [simplified sb, simplified] this]
have step_a: "(ts, m, 𝒮) ⇒⇩d
(ts[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))), (), False, 𝒪⇩s⇩b ∪ A - R,Map.empty)],
?ma,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
(is " _ ⇒⇩d (?ts_a, _, ?shared_a)").
from ts⇩j neq_i_j j_bound
have ts_a_j: "?ts_a!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(), 𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by auto
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b (ys) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from j_bound' have j_bound_ts_a: "j < length ?ts_a" by auto
from flush_store_buffer_append [where sb="ys" and sb'="Ghost⇩s⇩b A' L' R' W'#zs", simplified,
OF j_bound_ts_a is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_a_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?shared_a"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "Ghost A' L' R' W'# instrs zs @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs" and
steps_ys: "(?ts_a, ?ma, ?shared_a) ⇒⇩d⇧*
(?ts_a[j:=(last_prog
(hd_prog p⇩j zs) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs),
(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')],
flush ys (?ma),
share ys (?shared_a))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append)
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Ghost⇩s⇩b A' L' R' W']) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Ghost A' L' R' W'#is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b',
θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))),(), False, 𝒪⇩s⇩b ∪ A - R,Map.empty)"
by simp
from j_bound_ts_a have j_bound_ys: "j < length ?ts_ys"
by auto
then have ts_ys_j: "?ts_ys!j = (last_prog (hd_prog p⇩j zs) ys, Ghost A' L' R' W'#is⇩j'',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps zs), (),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by (clarsimp simp add: is⇩j')
note conflict_computation = r_rtranclp_rtranclp [OF step_a steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j]
have A'_unowned:
"∀i < length ?ts_ys. j≠i ⟶ (let (𝒪⇩i) = map owned ?ts_ys!i in A' ∩ 𝒪⇩i = {})"
apply cases
apply (fastforce simp add: Let_def is⇩s⇩b)+
done
from a'_in a'_A' A'_unowned [rule_format, of i] neq_i_j i_bound' A_R
show False
by (auto simp add: Let_def)
qed
qed
qed
}
thus ?thesis
by (auto simp add: Let_def)
qed
{
fix j
fix p⇩j is⇩s⇩b⇩j 𝒪⇩j ℛ⇩j 𝒟⇩s⇩b⇩j θ⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume ts⇩s⇩b_j: "ts⇩s⇩b!j=(p⇩j,is⇩s⇩b⇩j,θ⇩j,sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)"
assume neq_i_j: "i≠j"
have "A ∩ read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) = {}"
proof -
{
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume conflict: "A ∩ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j ≠ {}"
have False
proof -
from conflict obtain a' where
a'_in: "a' ∈ A" and
a'_in_j: "a' ∈ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
from ts_sim [rule_format, OF j_bound] ts⇩s⇩b_j j_bound
obtain p⇩j suspends⇩j "is⇩s⇩b⇩j" 𝒟⇩s⇩b⇩j 𝒟⇩j θ⇩s⇩b⇩j "is⇩j" where
ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j,is⇩s⇩b⇩j, θ⇩s⇩b⇩j, sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)" and
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(), 𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
apply (cases "ts⇩s⇩b!j")
apply (clarsimp simp add: Let_def)
done
from split_in_read_only_reads [OF a'_in_j [simplified suspends⇩j [symmetric]]]
obtain t' v' ys zs where
split_suspends⇩j: "suspends⇩j = ys @ Read⇩s⇩b False a' t' v'# zs" (is "suspends⇩j = ?suspends") and
a'_unacq: "a' ∉ acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j)"
by blast
from valid_program_history [OF j_bound ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from valid_last_prog [OF j_bound ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b›
j_bound
ts⇩s⇩b_j this]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound ts⇩s⇩b_j]
have nvo_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
with non_volatile_owned_or_read_only_append [of False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have nvo_take_j: "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j"
by auto
from a_unowned_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
have a_not_acq: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j"
by auto
from a_notin_unforwarded_non_volatile_reads_drop[OF j_bound ts⇩s⇩b_j neq_i_j]
have a_notin_unforwarded_reads: "a ∉ unforwarded_non_volatile_reads suspends⇩j {}"
by (simp add: suspends⇩j)
let ?ma="(m(a := f (θ⇩s⇩b(t↦m a))))"
from reads_consistent_mem_eq_on_unforwarded_non_volatile_reads [where W="{}"
and m'="?ma",simplified, OF _ subset_refl reads_consis_m_j]
a_notin_unforwarded_reads
have reads_consis_ma_j:
"reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?ma suspends⇩j"
by auto
from reads_consis_ma_j
have reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?ma (ys)"
by (simp add: split_suspends⇩j reads_consistent_append)
from direct_memop_step.RMWWrite [where cond=cond and θ=θ⇩s⇩b and m=m, OF cond' ]
have "(RMW a t (D, f) cond ret A L R W# is⇩s⇩b', θ⇩s⇩b, (), m, 𝒟,𝒪⇩s⇩b,ℛ⇩s⇩b,𝒮) →
(is⇩s⇩b', θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))), (), ?ma, False, 𝒪⇩s⇩b ∪ A - R,Map.empty, 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
have step_a: "(ts, m, 𝒮) ⇒⇩d
(ts[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))), (), False, 𝒪⇩s⇩b ∪ A - R,Map.empty)],
?ma,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
(is " _ ⇒⇩d (?ts_a, _, ?shared_a)").
from ts⇩j neq_i_j j_bound
have ts_a_j: "?ts_a!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(), 𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by auto
from valid_write_sops [OF j_bound ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b (ys) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from j_bound leq have j_bound_ts_a: "j < length ?ts_a" by auto
from flush_store_buffer_append [where sb="ys" and sb'="Read⇩s⇩b False a' t' v'#zs", simplified,
OF j_bound_ts_a is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_a_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?shared_a"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "Read False a' t'# instrs zs @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs" and
steps_ys: "(?ts_a, ?ma, ?shared_a) ⇒⇩d⇧*
(?ts_a[j:=(last_prog
(hd_prog p⇩j zs) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - insert t' (read_tmps zs)),
(), 𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')],
flush ys (?ma),
share ys (?shared_a))"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append)
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Read⇩s⇩b False a' t' v']) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Read False a' t'#is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j
have ts_ys_i: "?ts_ys!i = (p⇩s⇩b, is⇩s⇩b',
θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t ↦ m a)))),(), False, 𝒪⇩s⇩b ∪ A - R,Map.empty)"
by simp
from j_bound_ts_a have j_bound_ys: "j < length ?ts_ys"
by auto
then have ts_ys_j: "?ts_ys!j = (last_prog (hd_prog p⇩j zs) ys, Read False a' t'#is⇩j'', θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - insert t' (read_tmps zs)), (), 𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by (clarsimp simp add: is⇩j')
note conflict_computation = r_rtranclp_rtranclp [OF step_a steps_ys]
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j]
have "a' ∈ acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∨
a' ∈ read_only (share ys (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))"
apply cases
apply (auto simp add: Let_def is⇩s⇩b)
done
with a'_unacq
have a'_ro: "a' ∈ read_only (share ys (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))"
by auto
from a'_in
have a'_not_ro: "a' ∉ read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs)
have "a' ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
proof -
{
assume a_notin: "a' ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
from weak_sharing_consis [OF j_bound ts⇩s⇩b_j]
have "weak_sharing_consistent 𝒪⇩j sb⇩j".
with weak_sharing_consistent_append [of 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have "weak_sharing_consistent (acquired True ?take_sb⇩j 𝒪⇩j) suspends⇩j"
by (auto simp add: suspends⇩j)
with split_suspends⇩j
have weak_consis: "weak_sharing_consistent (acquired True ?take_sb⇩j 𝒪⇩j) ys"
by (simp add: weak_sharing_consistent_append)
from all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have "all_acquired ys ⊆ all_acquired sb⇩j"
apply (clarsimp)
apply (clarsimp simp add: suspends⇩j [symmetric] split_suspends⇩j all_acquired_append)
done
with a_notin acquired_takeWhile_non_volatile_Write⇩s⇩b [of sb⇩j 𝒪⇩j]
all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have "a' ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j ∪ all_acquired ys"
by auto
from read_only_share_unowned [OF weak_consis this a'_ro]
have "a' ∈ read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)" .
with a'_not_ro have False
by auto
with a_notin read_only_share_unowned [OF weak_consis _ a'_ro]
all_acquired_takeWhile [of "(Not ∘ is_volatile_Write⇩s⇩b)" sb⇩j]
have "a' ∈ read_only (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: acquired_takeWhile_non_volatile_Write⇩s⇩b)
with a'_not_ro have False
by auto
}
thus ?thesis by blast
qed
moreover
from A_unacquired_by_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
have "A ∩ all_acquired sb⇩j = {}"
by (auto simp add: Let_def)
moreover
from A_unowned_by_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
have "A ∩ 𝒪⇩j = {}"
by (auto simp add: Let_def dest: all_shared_acquired_in)
moreover note a'_in
ultimately
show False
by auto
qed
}
thus ?thesis
by (auto simp add: Let_def)
qed
} note A_no_read_only_reads = this
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof
fix j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j p⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume ts⇩s⇩b'_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "non_volatile_owned_or_read_only False 𝒮⇩s⇩b' 𝒪⇩j sb⇩j"
proof (cases "j=i")
case True
have "non_volatile_owned_or_read_only False
(𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) (𝒪⇩s⇩b ∪ A - R) []"
by simp
then show ?thesis
using True i_bound ts⇩s⇩b'_j
by (auto simp add: ts⇩s⇩b' 𝒮⇩s⇩b' sb sb')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
with ts⇩s⇩b'_j False i_bound
have ts⇩s⇩b_j: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
note nvo = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' ts⇩s⇩b_j]
from read_only_unowned [OF i_bound ts⇩s⇩b_i] R_owned
have "R ∩ read_only 𝒮⇩s⇩b = {}"
by auto
with A_no_read_only_reads [OF j_bound' ts⇩s⇩b_j False [symmetric]] L_subset
have "∀a∈read_only_reads
(acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j).
a ∈ read_only 𝒮⇩s⇩b ⟶ a ∈ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs)
from non_volatile_owned_or_read_only_read_only_reads_eq' [OF nvo this]
have "non_volatile_owned_or_read_only False (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) 𝒪⇩j sb⇩j".
thus ?thesis by (simp add: 𝒮⇩s⇩b')
qed
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof (unfold_locales)
fix i⇩1 j p⇩1 "is⇩1" 𝒪⇩1 ℛ⇩1 𝒟⇩1 xs⇩1 sb⇩1 p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume i⇩1_bound: "i⇩1 < length ts⇩s⇩b'"
assume j_bound: "j < length ts⇩s⇩b'"
assume i⇩1_j: "i⇩1 ≠ j"
assume ts_i⇩1: "ts⇩s⇩b'!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
assume ts_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}"
proof (cases "i⇩1=i")
case True
with ts_i⇩1 i_bound show ?thesis
by (simp add: ts⇩s⇩b' sb' sb)
next
case False
note i⇩1_i = this
from i⇩1_bound have i⇩1_bound': "i⇩1 < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b' sb' sb)
hence i⇩1_bound'': "i⇩1 < length (map owned ts⇩s⇩b)"
by auto
from ts_i⇩1 False have ts_i⇩1': "ts⇩s⇩b!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
by (simp add: ts⇩s⇩b' sb' sb)
show ?thesis
proof (cases "j=i")
case True
from i_bound ts_j ts⇩s⇩b' True have sb⇩j: "sb⇩j=[]"
by (simp add: ts⇩s⇩b' sb')
from A_unused_by_others [rule_format, OF _ False [symmetric]] ts_i⇩1 i⇩1_bound''
False i⇩1_bound'
have "A ∩ (𝒪⇩1 ∪ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1) = {}"
by (auto simp add: Let_def ts⇩s⇩b' 𝒪⇩s⇩b' sb' owned_def)
moreover
from outstanding_volatile_writes_unowned_by_others
[OF i⇩1_bound' i_bound i⇩1_i ts_i⇩1' ts⇩s⇩b_i]
have "𝒪⇩s⇩b ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}" by (simp add: sb)
ultimately show ?thesis using ts_j True
by (auto simp add: i_bound ts⇩s⇩b' 𝒪⇩s⇩b' sb⇩j)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_j False have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from outstanding_volatile_writes_unowned_by_others
[OF i⇩1_bound' j_bound' i⇩1_j ts_i⇩1' ts_j']
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}" .
qed
qed
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof
fix n m
fix p⇩n "is⇩n" 𝒪⇩n ℛ⇩n 𝒟⇩n θ⇩n sb⇩n p⇩m "is⇩m" 𝒪⇩m ℛ⇩m 𝒟⇩m θ⇩m sb⇩m
assume n_bound: "n < length ts⇩s⇩b'"
and m_bound: "m < length ts⇩s⇩b'"
and neq_n_m: "n≠m"
and nth: "ts⇩s⇩b'!n = (p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
and mth: "ts⇩s⇩b'!m =(p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
from n_bound have n_bound': "n < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
from m_bound have m_bound': "m < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
show "(𝒪⇩m ∪ all_acquired sb⇩m) ∩
read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) 𝒪⇩n)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) =
{}"
proof (cases "m=i")
case True
with neq_n_m have neq_n_i: "n≠i"
by auto
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
note read_only_reads_unowned [OF n_bound' i_bound neq_n_i nth' ts⇩s⇩b_i]
moreover
note A_no_read_only_reads [OF n_bound' nth']
ultimately
show ?thesis
using True ts⇩s⇩b_i neq_n_i nth mth n_bound' m_bound'
by (auto simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb sb')
next
case False
note neq_m_i = this
with m_bound mth i_bound have mth': "ts⇩s⇩b!m = (p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
by (auto simp add: ts⇩s⇩b')
show ?thesis
proof (cases "n=i")
case True
with ts⇩s⇩b_i nth mth neq_m_i n_bound'
show ?thesis
by (auto simp add: ts⇩s⇩b' sb')
next
case False
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
from read_only_reads_unowned [OF n_bound' m_bound' neq_n_m nth' mth'] False neq_m_i
show ?thesis
by (clarsimp)
qed
qed
qed
next
show "ownership_distinct ts⇩s⇩b'"
proof (unfold_locales)
fix i⇩1 j p⇩1 "is⇩1" 𝒪⇩1 ℛ⇩1 𝒟⇩1 xs⇩1 sb⇩1 p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume i⇩1_bound: "i⇩1 < length ts⇩s⇩b'"
assume j_bound: "j < length ts⇩s⇩b'"
assume i⇩1_j: "i⇩1 ≠ j"
assume ts_i⇩1: "ts⇩s⇩b'!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
assume ts_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "(𝒪⇩1 ∪ all_acquired sb⇩1) ∩ (𝒪⇩j ∪ all_acquired sb⇩j)= {}"
proof (cases "i⇩1=i")
case True
with i⇩1_j have i_j: "i≠j"
by simp
from i_bound ts_i⇩1 ts⇩s⇩b' True have sb⇩1: "sb⇩1=[]"
by (simp add: ts⇩s⇩b' sb')
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
hence j_bound'': "j < length (map owned ts⇩s⇩b)"
by simp
from ts_j i_j have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from A_unused_by_others [rule_format, OF _ i_j] ts_j i_j j_bound'
have "A ∩ (𝒪⇩j ∪ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j) = {}"
by (auto simp add: Let_def ts⇩s⇩b' owned_def)
moreover
from A_unacquired_by_others [rule_format, OF _ i_j] ts_j i_j j_bound'
have "A ∩ all_acquired sb⇩j = {}"
by (auto simp add: Let_def ts⇩s⇩b')
moreover
from ownership_distinct [OF i_bound j_bound' i_j ts⇩s⇩b_i ts_j']
have "𝒪⇩s⇩b ∩ (𝒪⇩j ∪ all_acquired sb⇩j)= {}" by (simp add: sb)
ultimately show ?thesis using ts_i⇩1 True
by (auto simp add: i_bound ts⇩s⇩b' 𝒪⇩s⇩b' sb' sb⇩1)
next
case False
note i⇩1_i = this
from i⇩1_bound have i⇩1_bound': "i⇩1 < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
hence i⇩1_bound'': "i⇩1 < length (map owned ts⇩s⇩b)"
by simp
from ts_i⇩1 False have ts_i⇩1': "ts⇩s⇩b!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
by (simp add: ts⇩s⇩b')
show ?thesis
proof (cases "j=i")
case True
from A_unused_by_others [rule_format, OF _ False [symmetric]] ts_i⇩1
False i⇩1_bound'
have "A ∩ (𝒪⇩1 ∪ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1) = {}"
by (auto simp add: Let_def ts⇩s⇩b' owned_def)
moreover
from A_unacquired_by_others [rule_format, OF _ False [symmetric]] ts_i⇩1 False i⇩1_bound'
have "A ∩ all_acquired sb⇩1 = {}"
by (auto simp add: Let_def ts⇩s⇩b' owned_def)
moreover
from ownership_distinct [OF i⇩1_bound' i_bound i⇩1_i ts_i⇩1' ts⇩s⇩b_i]
have "(𝒪⇩1 ∪ all_acquired sb⇩1) ∩ 𝒪⇩s⇩b = {}" by (simp add: sb)
ultimately show ?thesis
using ts_j True
by (auto simp add: i_bound ts⇩s⇩b' 𝒪⇩s⇩b' sb')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_j False have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from ownership_distinct [OF i⇩1_bound' j_bound' i⇩1_j ts_i⇩1' ts_j']
show "(𝒪⇩1 ∪ all_acquired sb⇩1) ∩ (𝒪⇩j ∪ all_acquired sb⇩j) = {}" .
qed
qed
qed
qed
have valid_hist': "valid_history program_step ts⇩s⇩b'"
proof -
from valid_history [OF i_bound ts⇩s⇩b_i]
have "history_consistent (θ⇩s⇩b(t↦ret (m⇩s⇩b a) (f (θ⇩s⇩b(t↦m⇩s⇩b a))))) (hd_prog p⇩s⇩b []) []" by simp
from valid_history_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' θ⇩s⇩b' sb' sb)
qed
from valid_reads [OF i_bound ts⇩s⇩b_i]
have reads_consis: "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b sb" .
have valid_reads': "valid_reads m⇩s⇩b' ts⇩s⇩b'"
proof (unfold_locales)
fix j p⇩j "is⇩j" 𝒪⇩j ℛ⇩j 𝒟⇩j acq⇩j θ⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume ts_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "reads_consistent False 𝒪⇩j m⇩s⇩b' sb⇩j"
proof (cases "i=j")
case True
from reads_consis ts_j j_bound sb show ?thesis
by (clarsimp simp add: True m⇩s⇩b' Write⇩s⇩b ts⇩s⇩b' sb')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
moreover from ts_j False have ts_j': "ts⇩s⇩b ! j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
using j_bound by (simp add: ts⇩s⇩b')
ultimately have consis_m: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j"
by (rule valid_reads)
let ?m' = "(m⇩s⇩b(a := f (θ⇩s⇩b(t ↦ m⇩s⇩b a))))"
from a_unowned_others [rule_format, OF _ False] j_bound' ts_j'
obtain a_acq: "a ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j" and
a_unsh: "a ∉ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
by auto
with a_notin_unforwarded_non_volatile_reads_drop [OF j_bound' ts_j' False]
have "∀a∈acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j ∪
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ∪
unforwarded_non_volatile_reads (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) {}.
?m' a = m⇩s⇩b a"
by auto
from reads_consistent_mem_eq_on_unforwarded_non_volatile_reads_drop
[where W="{}",simplified, OF this _ _ consis_m]
acquired_reads_all_acquired' [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)" 𝒪⇩j]
have "reads_consistent False 𝒪⇩j (m⇩s⇩b(a := f (θ⇩s⇩b(t ↦ m⇩s⇩b a)))) sb⇩j"
by (auto simp del: fun_upd_apply)
thus ?thesis
by (simp add: m⇩s⇩b')
qed
qed
have valid_sharing': "valid_sharing (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_writes_unshared (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (unfold_locales)
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j acq⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "non_volatile_writes_unshared (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) sb⇩j"
proof (cases "i=j")
case True
with i_bound jth show ?thesis
by (simp add: ts⇩s⇩b' sb' sb)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
from jth False have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from outstanding_non_volatile_writes_unshared [OF j_bound' jth']
have unshared: "non_volatile_writes_unshared 𝒮⇩s⇩b sb⇩j".
have "∀a∈dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) - dom 𝒮⇩s⇩b. a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
proof -
{
fix a
assume a_in: "a ∈ dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) - dom 𝒮⇩s⇩b"
hence a_R: "a ∈ R"
by clarsimp
assume a_in_j: "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
have False
proof -
from non_volatile_owned_or_read_only_outstanding_non_volatile_writes [OF
outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' jth']]
a_in_j
have "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
moreover
with ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth'] a_R R_owned
show False
by blast
qed
}
thus ?thesis by blast
qed
from non_volatile_writes_unshared_no_outstanding_non_volatile_Write⇩s⇩b
[OF unshared this]
show ?thesis .
qed
qed
next
show "sharing_consis (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (unfold_locales)
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j acq⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "sharing_consistent (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) 𝒪⇩j sb⇩j"
proof (cases "i=j")
case True
with i_bound jth show ?thesis
by (simp add: ts⇩s⇩b' sb' sb)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
from jth False have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from sharing_consis [OF j_bound' jth']
have consis: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
have acq_cond: "all_acquired sb⇩j ∩ dom 𝒮⇩s⇩b - dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
proof -
{
fix a
assume a_acq: "a ∈ all_acquired sb⇩j"
assume "a ∈ dom 𝒮⇩s⇩b"
assume a_L: "a ∈ L"
have False
proof -
from A_unacquired_by_others [rule_format, of j,OF _ False] j_bound' jth'
have "A ∩ all_acquired sb⇩j = {}"
by auto
with a_acq a_L L_subset
show False
by blast
qed
}
thus ?thesis
by auto
qed
have uns_cond: "all_unshared sb⇩j ∩ dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) - dom 𝒮⇩s⇩b = {}"
proof -
{
fix a
assume a_uns: "a ∈ all_unshared sb⇩j"
assume "a ∉ L"
assume a_R: "a ∈ R"
have False
proof -
from unshared_acquired_or_owned [OF consis] a_uns
have "a ∈ all_acquired sb⇩j ∪ 𝒪⇩j" by auto
with ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth'] R_owned a_R
show False
by blast
qed
}
thus ?thesis
by auto
qed
from sharing_consistent_preservation [OF consis acq_cond uns_cond]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
qed
next
show "unowned_shared (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (unfold_locales)
show "- ⋃((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set ts⇩s⇩b') ⊆ dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L)"
proof -
have s: "⋃((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set ts⇩s⇩b') =
⋃((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set ts⇩s⇩b) ∪ A - R"
apply (unfold ts⇩s⇩b' 𝒪⇩s⇩b')
apply (rule acquire_release_ownership_nth_update [OF R_owned i_bound ts⇩s⇩b_i])
apply fact
done
note unowned_shared L_subset A_R
then
show ?thesis
apply (simp only: s)
apply auto
done
qed
qed
next
show "read_only_unowned (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j acq⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "𝒪⇩j ∩ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
proof (cases "i=j")
case True
from read_only_unowned [OF i_bound ts⇩s⇩b_i] R_owned A_R
have "(𝒪⇩s⇩b ∪ A - R) ∩ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by (auto simp add: in_read_only_convs )
with jth ts⇩s⇩b_i i_bound True
show ?thesis
by (auto simp add: 𝒪⇩s⇩b' ts⇩s⇩b')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
with False jth have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from read_only_unowned [OF j_bound' jth']
have "𝒪⇩j ∩ read_only 𝒮⇩s⇩b = {}".
moreover
from A_unowned_by_others [rule_format, OF _ False] j_bound' jth'
have "A ∩ 𝒪⇩j = {}"
by (auto dest: all_shared_acquired_in )
moreover
from ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth']
have "𝒪⇩s⇩b ∩ 𝒪⇩j = {}"
by auto
moreover note R_owned A_R
ultimately show ?thesis
by (fastforce simp add: in_read_only_convs split: if_split_asm)
qed
qed
next
show "no_outstanding_write_to_read_only_memory (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j acq⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "no_write_to_read_only_memory (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) sb⇩j"
proof (cases "i=j")
case True
with jth ts⇩s⇩b_i i_bound
show ?thesis
by (auto simp add: sb sb' ts⇩s⇩b')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
with False jth have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from no_outstanding_write_to_read_only_memory [OF j_bound' jth']
have nw: "no_write_to_read_only_memory 𝒮⇩s⇩b sb⇩j".
have "R ∩ outstanding_refs is_Write⇩s⇩b sb⇩j = {}"
proof -
note dist = ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth']
from non_volatile_owned_or_read_only_outstanding_non_volatile_writes
[OF outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' jth']]
dist
have "outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j ∩ 𝒪⇩s⇩b = {}"
by auto
moreover
from outstanding_volatile_writes_unowned_by_others [OF j_bound' i_bound
False [symmetric] jth' ts⇩s⇩b_i ]
have "outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ∩ 𝒪⇩s⇩b = {}"
by auto
ultimately have "outstanding_refs is_Write⇩s⇩b sb⇩j ∩ 𝒪⇩s⇩b = {}"
by (auto simp add: misc_outstanding_refs_convs)
with R_owned
show ?thesis by blast
qed
then
have "∀a∈outstanding_refs is_Write⇩s⇩b sb⇩j.
a ∈ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ⟶ a ∈ read_only 𝒮⇩s⇩b"
by (auto simp add: in_read_only_convs)
from no_write_to_read_only_memory_read_only_reads_eq [OF nw this]
show ?thesis .
qed
qed
qed
have tmps_distinct': "tmps_distinct ts⇩s⇩b'"
proof (intro_locales)
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_load_tmps is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b" split: instr.splits)
from load_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' "is⇩s⇩b")
next
from read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_read_tmps []" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
from read_tmps_distinct_nth_update [OF i_bound this]
show "read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
next
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ read_tmps [] = {}"
by (clarsimp)
from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have valid_sops': "valid_sops ts⇩s⇩b'"
proof -
from valid_store_sops [OF i_bound ts⇩s⇩b_i]
obtain
valid_store_sops': "∀sop∈store_sops is⇩s⇩b'. valid_sop sop"
by (auto simp add: "is⇩s⇩b" ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
from valid_sops_nth_update [OF i_bound _ valid_store_sops', where sb= "[]" ]
show ?thesis by (auto simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have valid_dd': "valid_data_dependency ts⇩s⇩b'"
proof -
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
obtain
dd_is: "data_dependency_consistent_instrs (dom θ⇩s⇩b') is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b" θ⇩s⇩b')
from load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ ⋃(fst ` write_sops []) = {}"
by (auto simp add: write_sops_append)
from valid_data_dependency_nth_update [OF i_bound dd_is this]
show ?thesis by (simp add: ts⇩s⇩b' sb' sb 𝒪⇩s⇩b')
qed
have load_tmps_fresh': "load_tmps_fresh ts⇩s⇩b'"
proof -
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have "load_tmps (RMW a t (D,f) cond ret A L R W # is⇩s⇩b') ∩ dom θ⇩s⇩b = {}"
by (simp add: "is⇩s⇩b")
moreover
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i] have "t ∉ load_tmps is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b")
ultimately have "load_tmps is⇩s⇩b' ∩ dom (θ⇩s⇩b(t ↦ ret (m⇩s⇩b a) (f (θ⇩s⇩b(t↦m⇩s⇩b a))))) = {}"
by auto
from load_tmps_fresh_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' θ⇩s⇩b')
qed
from enough_flushs_nth_update [OF i_bound, where sb="[]" ]
have enough_flushs': "enough_flushs ts⇩s⇩b'"
by (auto simp: ts⇩s⇩b' sb' sb)
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
have causal': "causal_program_history is⇩s⇩b' sb'"
by (simp add: "is⇩s⇩b" sb sb')
have "last_prog p⇩s⇩b sb' = p⇩s⇩b"
by (simp add: sb' sb)
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b' sb')
qed
from is_sim have "is": "is = RMW a t (D,f) cond ret A L R W # is⇩s⇩b'"
by (simp add: suspends sb "is⇩s⇩b")
from direct_memop_step.RMWWrite [where cond=cond and θ=θ⇩s⇩b and m=m, OF cond']
have "(RMW a t (D, f) cond ret A L R W # is⇩s⇩b', θ⇩s⇩b, (),m, 𝒟, 𝒪⇩s⇩b,ℛ⇩s⇩b, 𝒮) →
(is⇩s⇩b',θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t↦m a)))), (),
m(a := f (θ⇩s⇩b(t ↦ m a))), False, 𝒪⇩s⇩b ∪ A - R, Map.empty, 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
have "(ts, m, 𝒮) ⇒⇩d (ts[i := (p⇩s⇩b, is⇩s⇩b',θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t↦m a)))), (), False, 𝒪⇩s⇩b ∪ A - R,Map.empty)],
m(a := f (θ⇩s⇩b(t ↦ m a))),𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
moreover
have tmps_commute: "θ⇩s⇩b(t ↦ ret (m⇩s⇩b a) (f (θ⇩s⇩b(t↦m⇩s⇩b a)))) =
(θ⇩s⇩b |` (dom θ⇩s⇩b - {t}))(t ↦ ret (m⇩s⇩b a) (f (θ⇩s⇩b(t↦m⇩s⇩b a))))"
apply (rule ext)
apply (auto simp add: restrict_map_def domIff)
done
from a_unflushed ts⇩s⇩b_i sb
have a_unflushed':
"∀j < length ts⇩s⇩b.
(let (_,_,_,sb⇩j,_,_,_) = ts⇩s⇩b!j
in a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j))"
by auto
have all_shared_L: "∀i p is 𝒪 ℛ 𝒟 acq θ sb. i < length ts⇩s⇩b ⟶
ts⇩s⇩b ! i = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ L = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume x_shared: "x ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_L: "x ∈ L"
have False
proof (cases "i=j")
case True with x_shared ts⇩s⇩b_i jth show False by (simp add: sb)
next
case False
show False
proof -
from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
have "all_shared sb⇩j ⊆ all_acquired sb⇩j ∪ 𝒪⇩j".
moreover have "all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ⊆ all_shared sb⇩j"
using all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
by auto
moreover
from A_unacquired_by_others [rule_format, OF _ False] jth j_bound
have "A ∩ all_acquired sb⇩j = {}" by auto
moreover
from A_unowned_by_others [rule_format, OF _ False] jth j_bound
have "A ∩ 𝒪⇩j = {}"
by (auto dest: all_shared_acquired_in)
ultimately
show False
using L_subset x_L x_shared
by blast
qed
qed
}
thus ?thesis by blast
qed
have all_shared_A: "∀i p is 𝒪 ℛ 𝒟 θ sb. i < length ts⇩s⇩b ⟶
ts⇩s⇩b ! i = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ A = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume x_shared: "x ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_A: "x ∈ A"
have False
proof (cases "i=j")
case True with x_shared ts⇩s⇩b_i jth show False by (simp add: sb)
next
case False
show False
proof -
from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
have "all_shared sb⇩j ⊆ all_acquired sb⇩j ∪ 𝒪⇩j".
moreover have "all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ⊆ all_shared sb⇩j"
using all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
by auto
moreover
from A_unacquired_by_others [rule_format, OF _ False] jth j_bound
have "A ∩ all_acquired sb⇩j = {}" by auto
moreover
from A_unowned_by_others [rule_format, OF _ False] jth j_bound
have "A ∩ 𝒪⇩j = {}"
by (auto dest: all_shared_acquired_in)
ultimately
show False
using x_A x_shared
by blast
qed
qed
}
thus ?thesis by blast
qed
hence all_shared_L: "∀i p is 𝒪 ℛ 𝒟 θ sb. i < length ts⇩s⇩b ⟶
ts⇩s⇩b ! i = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ L = {}"
using L_subset by blast
have all_unshared_R: "∀i p is 𝒪 ℛ 𝒟 θ sb. i < length ts⇩s⇩b ⟶
ts⇩s⇩b ! i = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ R = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume x_unshared: "x ∈ all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_R: "x ∈ R"
have False
proof (cases "i=j")
case True with x_unshared ts⇩s⇩b_i jth show False by (simp add: sb)
next
case False
show False
proof -
from unshared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
have "all_unshared sb⇩j ⊆ all_acquired sb⇩j ∪ 𝒪⇩j".
moreover have "all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ⊆ all_unshared sb⇩j"
using all_unshared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
by auto
moreover
note ownership_distinct [OF i_bound j_bound False ts⇩s⇩b_i jth]
ultimately
show False
using R_owned x_R x_unshared
by blast
qed
qed
}
thus ?thesis by blast
qed
have all_acquired_R: "∀i p is 𝒪 ℛ 𝒟 θ sb. i < length ts⇩s⇩b ⟶
ts⇩s⇩b ! i = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_acquired (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ R = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume x_acq: "x ∈ all_acquired (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_R: "x ∈ R"
have False
proof (cases "i=j")
case True with x_acq ts⇩s⇩b_i jth show False by (simp add: sb)
next
case False
show False
proof -
from x_acq have "x ∈ all_acquired sb⇩j"
using all_acquired_append [of "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
"dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"]
by auto
moreover
note ownership_distinct [OF i_bound j_bound False ts⇩s⇩b_i jth]
ultimately
show False
using R_owned x_R
by blast
qed
qed
}
thus ?thesis by blast
qed
have all_shared_R: "∀i p is 𝒪 ℛ 𝒟 θ sb. i < length ts⇩s⇩b ⟶
ts⇩s⇩b ! i = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ R = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume x_shared: "x ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_R: "x ∈ R"
have False
proof (cases "i=j")
case True with x_shared ts⇩s⇩b_i jth show False by (simp add: sb)
next
case False
show False
proof -
from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
have "all_shared sb⇩j ⊆ all_acquired sb⇩j ∪ 𝒪⇩j".
moreover have "all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ⊆ all_shared sb⇩j"
using all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
by auto
moreover
note ownership_distinct [OF i_bound j_bound False ts⇩s⇩b_i jth]
ultimately
show False
using R_owned x_R x_shared
by blast
qed
qed
}
thus ?thesis by blast
qed
from share_all_until_volatile_write_commute [OF ‹ownership_distinct ts⇩s⇩b› ‹sharing_consis 𝒮⇩s⇩b ts⇩s⇩b›
all_shared_L all_shared_A all_acquired_R all_unshared_R all_shared_R]
have share_commute: "share_all_until_volatile_write ts⇩s⇩b 𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L =
share_all_until_volatile_write ts⇩s⇩b (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L)".
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume j_bound: "j < length ts⇩s⇩b"
assume neq: "i ≠ j"
have "release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)
(dom 𝒮⇩s⇩b ∪ R - L) ℛ⇩j
= release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)
(dom 𝒮⇩s⇩b) ℛ⇩j"
proof -
{
fix a
assume a_in: "a ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
have "(a ∈ (dom 𝒮⇩s⇩b ∪ R - L)) = (a ∈ dom 𝒮⇩s⇩b)"
proof -
from A_unowned_by_others [rule_format, OF j_bound neq ] jth
A_unacquired_by_others [rule_format, OF _ neq] j_bound
have A_dist: "A ∩ (𝒪⇩j ∪ all_acquired sb⇩j) = {}"
by (auto dest: all_shared_acquired_in)
from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]] a_in
all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
have a_in: "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
with ownership_distinct [OF i_bound j_bound neq ts⇩s⇩b_i jth]
have "a ∉ (𝒪⇩s⇩b ∪ all_acquired sb)" by auto
with A_dist R_owned A_R A_shared_owned L_subset a_in
obtain "a ∉ R" and "a ∉ L"
by fastforce
then show ?thesis by auto
qed
}
then
show ?thesis
apply -
apply (rule release_all_shared_exchange)
apply auto
done
qed
}
note release_commute = this
have "(ts⇩s⇩b',m⇩s⇩b(a := f (θ⇩s⇩b(t ↦ m⇩s⇩b a))),𝒮⇩s⇩b') ∼ (ts[i := (p⇩s⇩b,is⇩s⇩b',
θ⇩s⇩b(t ↦ ret (m a) (f (θ⇩s⇩b(t↦m a)))),(), False,𝒪⇩s⇩b ∪ A - R,Map.empty)],m(a := f (θ⇩s⇩b(t ↦ m a))),𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
apply (rule sim_config.intros)
apply (simp only: m_a )
apply (simp only: m)
apply (simp only: flush_all_until_volatile_write_update_other [OF a_unflushed', symmetric] ts⇩s⇩b')
apply (simp add: flush_all_until_volatile_nth_update_unused [OF i_bound ts⇩s⇩b_i, simplified sb] sb')
apply (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b' m
flush_all_until_volatile_nth_update_unused [OF i_bound ts⇩s⇩b_i, simplified sb])
using share_all_until_volatile_write_RMW_commute [OF i_bound ts⇩s⇩b_i [simplified is⇩s⇩b sb]]
apply (clarsimp simp add: 𝒮 ts⇩s⇩b' 𝒮⇩s⇩b' is⇩s⇩b 𝒪⇩s⇩b' ℛ⇩s⇩b' θ⇩s⇩b' sb' sb share_commute)
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim
apply (clarsimp simp add: Let_def nth_list_update
ts⇩s⇩b' sb' sb 𝒪⇩s⇩b' ℛ⇩s⇩b' 𝒮⇩s⇩b' θ⇩s⇩b' 𝒟⇩s⇩b' ex_not m_a
split: if_split_asm)
apply (rule conjI)
apply clarsimp
apply (rule tmps_commute)
apply clarsimp
apply (frule (2) release_commute)
apply clarsimp
apply fastforce
done
ultimately
show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' valid_sops'
valid_dd' load_tmps_fresh' enough_flushs'
valid_program_history' valid' m⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply)
next
case (SBHGhost A L R W)
then obtain
"is⇩s⇩b": "is⇩s⇩b = Ghost A L R W# is⇩s⇩b'" and
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b" and
ℛ⇩s⇩b': "ℛ⇩s⇩b'=ℛ⇩s⇩b" and
θ⇩s⇩b': "θ⇩s⇩b' = θ⇩s⇩b" and
𝒟⇩s⇩b': "𝒟⇩s⇩b'=𝒟⇩s⇩b" and
sb': "sb'=sb@[Ghost⇩s⇩b A L R W]" and
m⇩s⇩b': "m⇩s⇩b' = m⇩s⇩b" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b"
by auto
from safe_memop_flush_sb [simplified is⇩s⇩b] obtain
L_subset: "L ⊆ A" and
A_shared_owned: "A ⊆ dom (share ?drop_sb 𝒮) ∪ acquired True sb 𝒪⇩s⇩b" and
R_acq: "R ⊆ acquired True sb 𝒪⇩s⇩b" and
A_R: "A ∩ R = {}" and
A_unowned_by_others_ts:
"∀j<length (map owned ts). i≠j ⟶ (A ∩ (owned (ts!j) ∪ dom (released (ts!j))) = {})"
by cases auto
from A_unowned_by_others_ts ts_sim leq
have A_unowned_by_others:
"∀j<length ts⇩s⇩b. i≠j ⟶ (let (_,_,_,sb⇩j,_,𝒪⇩j,_) = ts⇩s⇩b!j
in A ∩ (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j ∪
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)) = {})"
apply (clarsimp simp add: Let_def)
subgoal for j
apply (drule_tac x=j in spec)
apply (force simp add: dom_release_takeWhile)
done
done
have A_unused_by_others:
"∀j<length (map 𝒪_sb ts⇩s⇩b). i ≠ j ⟶
(let (𝒪⇩j, sb⇩j) = map 𝒪_sb ts⇩s⇩b! j
in A ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j = {})"
proof -
{
fix j 𝒪⇩j sb⇩j
assume j_bound: "j < length (map owned ts⇩s⇩b)"
assume neq_i_j: "i≠j"
assume ts⇩s⇩b_j: "(map 𝒪_sb ts⇩s⇩b)!j = (𝒪⇩j,sb⇩j)"
assume conflict: "A ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {}"
have False
proof -
from j_bound leq
have j_bound': "j < length (map owned ts)"
by auto
from j_bound have j_bound'': "j < length ts⇩s⇩b"
by auto
from j_bound' have j_bound''': "j < length ts"
by simp
from conflict obtain a' where
a'_in: "a' ∈ A" and
a'_in_j: "a' ∈ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j"
by auto
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
from ts_sim [rule_format, OF j_bound''] ts⇩s⇩b_j j_bound''
obtain p⇩j suspends⇩j "is⇩s⇩b⇩j" θ⇩s⇩b⇩j 𝒟⇩s⇩b⇩j 𝒟⇩j ℛ⇩j "is⇩j" where
ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j,is⇩s⇩b⇩j,θ⇩s⇩b⇩j, sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)" and
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j, release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
apply (cases "ts⇩s⇩b!j")
apply (force simp add: Let_def)
done
have "a' ∈ outstanding_refs is_volatile_Write⇩s⇩b suspends⇩j"
proof -
from a'_in_j
have "a' ∈ outstanding_refs is_volatile_Write⇩s⇩b (?take_sb⇩j @ ?drop_sb⇩j)"
by simp
thus ?thesis
apply (simp only: outstanding_refs_append suspends⇩j)
apply (auto simp add: outstanding_refs_conv dest: set_takeWhileD)
done
qed
from split_volatile_Write⇩s⇩b_in_outstanding_refs [OF this]
obtain sop v ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Write⇩s⇩b True a' sop v A' L' R' W' # zs" (is "suspends⇩j = ?suspends")
by blast
from direct_memop_step.Ghost [where θ=θ⇩s⇩b and m="flush ?drop_sb m"]
have "(Ghost A L R W# is⇩s⇩b',
θ⇩s⇩b, (), flush ?drop_sb m, 𝒟⇩s⇩b,
acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b, share ?drop_sb 𝒮) →
(is⇩s⇩b', θ⇩s⇩b, (), flush ?drop_sb m, 𝒟⇩s⇩b,
acquired True sb 𝒪⇩s⇩b ∪ A - R,
augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b),
share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF
i_bound_ts' [simplified is⇩s⇩b] ts'_i [simplified is⇩s⇩b] this [simplified is⇩s⇩b]]
have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d
(?ts'[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b, (),𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b ∪ A - R,augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b))],
flush ?drop_sb m,share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
(is " _ ⇒⇩d (?ts_A, ?m_A, ?share_A)")
by (simp add: is⇩s⇩b)
from i_bound' have i_bound'': "i < length ?ts_A"
by simp
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts_A_j: "?ts_A!j = (hd_prog p⇩j (ys @ Write⇩s⇩b True a' sop v A' L' R' W' # zs), is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (ys @ Write⇩s⇩b True a' sop v A' L' R' W' # zs)), (), 𝒟⇩j,
acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by (simp add: split_suspends⇩j)
from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
by simp
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound''
ts⇩s⇩b_j reads_consis]
have reads_consis_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m]
have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
?m_A suspends⇩j".
hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A ys"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops ys. valid_sop sop"
apply (simp only: write_sops_append )
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop h_consis]
last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b ys = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from flush_store_buffer_append [
OF j_bound'''' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_A_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_m_A_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?share_A"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs (Write⇩s⇩b True a' sop v A' L' R' W' # zs) @ is⇩s⇩b⇩j =
is⇩j' @ prog_instrs (Write⇩s⇩b True a' sop v A' L' R' W' # zs)" and
steps_ys: "(?ts_A, ?m_A, ?share_A) ⇒⇩d⇧*
(?ts_A[j:= (last_prog (hd_prog p⇩j (Write⇩s⇩b True a' sop v A' L' R' W' # zs)) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Write⇩s⇩b True a' sop v A' L' R' W' # zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j') ],
flush ys ?m_A,
share ys ?share_A)"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto)
note conflict_computation = rtranclp_trans [OF rtranclp_r_rtranclp [OF steps_flush_sb, OF store_step] steps_ys]
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Write⇩s⇩b True a' sop v A' L' R' W']) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Write True a' sop A' L' R' W' #is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from j_bound'''
have j_bound_ys: "j < length ?ts_ys"
by auto
from j_bound_ys neq_i_j
have ts_ys_j: "?ts_ys!j=(last_prog (hd_prog p⇩j (Write⇩s⇩b True a' sop v A' L' R' W'# zs)) ys, is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Write⇩s⇩b True a' sop v A' L' R' W'# zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by auto
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j, simplified is⇩j']
have a_unowned:
"∀i < length ?ts_ys. j≠i ⟶ (let (𝒪⇩i) = map owned ?ts_ys!i in a' ∉ 𝒪⇩i)"
apply cases
apply (auto simp add: Let_def is⇩s⇩b)
done
from a'_in a_unowned [rule_format, of i] neq_i_j i_bound' A_R
show False
by (auto simp add: Let_def)
qed
}
thus ?thesis
by (auto simp add: Let_def)
qed
have A_unaquired_by_others:
"∀j<length (map 𝒪_sb ts⇩s⇩b). i ≠ j ⟶
(let (𝒪⇩j, sb⇩j) = map 𝒪_sb ts⇩s⇩b! j
in A ∩ all_acquired sb⇩j = {})"
proof -
{
fix j 𝒪⇩j sb⇩j
assume j_bound: "j < length (map owned ts⇩s⇩b)"
assume neq_i_j: "i≠j"
assume ts⇩s⇩b_j: "(map 𝒪_sb ts⇩s⇩b)!j = (𝒪⇩j,sb⇩j)"
assume conflict: "A ∩ all_acquired sb⇩j ≠ {}"
have False
proof -
from j_bound leq
have j_bound': "j < length (map owned ts)"
by auto
from j_bound have j_bound'': "j < length ts⇩s⇩b"
by auto
from j_bound' have j_bound''': "j < length ts"
by simp
from conflict obtain a' where
a'_in: "a' ∈ A" and
a'_in_j: "a' ∈ all_acquired sb⇩j"
by auto
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
from ts_sim [rule_format, OF j_bound''] ts⇩s⇩b_j j_bound''
obtain p⇩j suspends⇩j "is⇩s⇩b⇩j" θ⇩s⇩b⇩j 𝒟⇩s⇩b⇩j 𝒟⇩j ℛ⇩j "is⇩j" where
ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j,is⇩s⇩b⇩j, θ⇩s⇩b⇩j, sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)" and
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
apply (cases "ts⇩s⇩b!j")
apply (force simp add: Let_def)
done
from a'_in_j all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have "a' ∈ all_acquired ?take_sb⇩j ∨ a' ∈ all_acquired suspends⇩j"
by (auto simp add: suspends⇩j)
thus False
proof
assume "a' ∈ all_acquired ?take_sb⇩j"
with A_unowned_by_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound a'_in
show False
by (auto dest: all_acquired_unshared_acquired)
next
assume conflict_drop: "a' ∈ all_acquired suspends⇩j"
from split_all_acquired_in [OF conflict_drop]
show False
proof
assume "∃sop a'' v ys zs A L R W.
suspends⇩j = ys @ Write⇩s⇩b True a'' sop v A L R W# zs ∧ a' ∈ A"
then
obtain a'' sop' v' ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs"
(is "suspends⇩j = ?suspends") and
a'_A': "a' ∈ A'"
by auto
from direct_memop_step.Ghost [where θ=θ⇩s⇩b and m="flush ?drop_sb m"]
have "(Ghost A L R W# is⇩s⇩b',
θ⇩s⇩b, (), flush ?drop_sb m,𝒟⇩s⇩b,
acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b,share ?drop_sb 𝒮) →
(is⇩s⇩b', θ⇩s⇩b, (), flush ?drop_sb m, 𝒟⇩s⇩b,
acquired True sb 𝒪⇩s⇩b ∪ A - R,
augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b),
share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF
i_bound_ts' [simplified is⇩s⇩b] ts'_i [simplified is⇩s⇩b] this [simplified is⇩s⇩b]]
have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d
(?ts'[i := (p⇩s⇩b, is⇩s⇩b',θ⇩s⇩b, (),𝒟⇩s⇩b,
acquired True sb 𝒪⇩s⇩b ∪ A - R,
augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b))],
flush ?drop_sb m,share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
(is " _ ⇒⇩d (?ts_A, ?m_A, ?share_A)")
by (simp add: is⇩s⇩b)
from i_bound' have i_bound'': "i < length ?ts_A"
by simp
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts_A_j: "?ts_A!j = (hd_prog p⇩j (ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs), is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (ys @ Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs)), (), 𝒟⇩j,
acquired True ?take_sb⇩j 𝒪⇩j, release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by (simp add: split_suspends⇩j)
from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
by simp
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound''
ts⇩s⇩b_j reads_consis]
have reads_consis_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m]
have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
?m_A suspends⇩j".
hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A ys"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops ys. valid_sop sop"
apply (simp only: write_sops_append )
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop h_consis]
last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b ys = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from flush_store_buffer_append [
OF j_bound'''' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_A_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_m_A_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?share_A"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs (Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs) @ is⇩s⇩b⇩j =
is⇩j' @ prog_instrs (Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs)" and
steps_ys: "(?ts_A, ?m_A, ?share_A) ⇒⇩d⇧*
(?ts_A[j:= (last_prog (hd_prog p⇩j (Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs)) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Write⇩s⇩b True a'' sop' v' A' L' R' W' # zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j') ],
flush ys ?m_A,share ys ?share_A)"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto)
note conflict_computation = rtranclp_trans [OF rtranclp_r_rtranclp [OF steps_flush_sb, OF store_step] steps_ys]
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Write⇩s⇩b True a'' sop' v' A' L' R' W']) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Write True a'' sop' A' L' R' W'#is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from j_bound'''
have j_bound_ys: "j < length ?ts_ys"
by auto
from j_bound_ys neq_i_j
have ts_ys_j: "?ts_ys!j=(last_prog (hd_prog p⇩j (Write⇩s⇩b True a'' sop' v' A' L' R' W'# zs)) ys, is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Write⇩s⇩b True a'' sop' v' A' L' R' W'# zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by auto
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j, simplified is⇩j']
have A'_unowned:
"∀i < length ?ts_ys. j≠i ⟶ (let (𝒪⇩i) = map owned ?ts_ys!i in A' ∩ 𝒪⇩i = {})"
apply cases
apply (fastforce simp add: Let_def is⇩s⇩b)+
done
from a'_in a'_A' A'_unowned [rule_format, of i] neq_i_j i_bound' A_R
show False
by (auto simp add: Let_def)
next
assume "∃A L R W ys zs.
suspends⇩j = ys @ Ghost⇩s⇩b A L R W # zs ∧ a' ∈ A"
then
obtain ys zs A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Ghost⇩s⇩b A' L' R' W'# zs" (is "suspends⇩j = ?suspends") and
a'_A': "a' ∈ A'"
by auto
from direct_memop_step.Ghost [where θ=θ⇩s⇩b and m="flush ?drop_sb m"]
have "(Ghost A L R W# is⇩s⇩b',
θ⇩s⇩b, (), flush ?drop_sb m, 𝒟⇩s⇩b,
acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b, share ?drop_sb 𝒮) →
(is⇩s⇩b', θ⇩s⇩b, (), flush ?drop_sb m, 𝒟⇩s⇩b,
acquired True sb 𝒪⇩s⇩b ∪ A - R,
augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b),
share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF
i_bound_ts' [simplified is⇩s⇩b] ts'_i [simplified is⇩s⇩b] this [simplified is⇩s⇩b]]
have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d
(?ts'[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b, (), 𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b ∪ A - R,augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b))],
flush ?drop_sb m,share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
(is " _ ⇒⇩d (?ts_A, ?m_A, ?share_A)")
by (simp add: is⇩s⇩b)
from i_bound' have i_bound'': "i < length ?ts_A"
by simp
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts_A_j: "?ts_A!j = (hd_prog p⇩j (ys @ Ghost⇩s⇩b A' L' R' W'# zs), is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (ys @ Ghost⇩s⇩b A' L' R' W'# zs)), (),𝒟⇩j,
acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by (simp add: split_suspends⇩j)
from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
by simp
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound''
ts⇩s⇩b_j reads_consis]
have reads_consis_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m]
have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
?m_A suspends⇩j".
hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A ys"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops ys. valid_sop sop"
apply (simp only: write_sops_append )
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop h_consis]
last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b ys = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from flush_store_buffer_append [
OF j_bound'''' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_A_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_m_A_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?share_A"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs (Ghost⇩s⇩b A' L' R' W'# zs) @ is⇩s⇩b⇩j =
is⇩j' @ prog_instrs (Ghost⇩s⇩b A' L' R' W'# zs)" and
steps_ys: "(?ts_A, ?m_A, ?share_A) ⇒⇩d⇧*
(?ts_A[j:= (last_prog (hd_prog p⇩j (Ghost⇩s⇩b A' L' R' W'# zs)) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Ghost⇩s⇩b A' L' R' W'# zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j') ],
flush ys ?m_A, share ys ?share_A)"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto)
note conflict_computation = rtranclp_trans [OF rtranclp_r_rtranclp [OF steps_flush_sb, OF store_step] steps_ys]
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Ghost⇩s⇩b A' L' R' W']) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Ghost A' L' R' W'#is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from j_bound'''
have j_bound_ys: "j < length ?ts_ys"
by auto
from j_bound_ys neq_i_j
have ts_ys_j: "?ts_ys!j=(last_prog (hd_prog p⇩j (Ghost⇩s⇩b A' L' R' W'# zs)) ys, is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Write⇩s⇩b True a'' sop' v' A' L' R' W'# zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by auto
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j, simplified is⇩j']
have A'_unowned:
"∀i < length ?ts_ys. j≠i ⟶ (let (𝒪⇩i) = map owned ?ts_ys!i in A' ∩ 𝒪⇩i = {})"
apply cases
apply (fastforce simp add: Let_def is⇩s⇩b)+
done
from a'_in a'_A' A'_unowned [rule_format, of i] neq_i_j i_bound' A_R
show False
by (auto simp add: Let_def)
qed
qed
qed
}
thus ?thesis
by (auto simp add: Let_def)
qed
have A_no_read_only_reads_by_others:
"∀j<length (map 𝒪_sb ts⇩s⇩b). i ≠ j ⟶
(let (𝒪⇩j, sb⇩j) = map 𝒪_sb ts⇩s⇩b! j
in A ∩ read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) = {})"
proof -
{
fix j 𝒪⇩j sb⇩j
assume j_bound: "j < length (map owned ts⇩s⇩b)"
assume neq_i_j: "i≠j"
assume ts⇩s⇩b_j: "(map 𝒪_sb ts⇩s⇩b)!j = (𝒪⇩j,sb⇩j)"
let ?take_sb⇩j = "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
let ?drop_sb⇩j = "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume conflict: "A ∩ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j ≠ {}"
have False
proof -
from j_bound leq
have j_bound': "j < length (map owned ts)"
by auto
from j_bound have j_bound'': "j < length ts⇩s⇩b"
by auto
from j_bound' have j_bound''': "j < length ts"
by simp
from conflict obtain a' where
a'_in: "a' ∈ A" and
a'_in_j: "a' ∈ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
from ts_sim [rule_format, OF j_bound''] ts⇩s⇩b_j j_bound''
obtain p⇩j suspends⇩j "is⇩s⇩b⇩j" 𝒟⇩s⇩b⇩j 𝒟⇩j ℛ⇩j θ⇩s⇩b⇩j "is⇩j" where
ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j,is⇩s⇩b⇩j, θ⇩s⇩b⇩j, sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)" and
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps suspends⇩j),(), 𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
apply (cases "ts⇩s⇩b!j")
apply (force simp add: Let_def)
done
from split_in_read_only_reads [OF a'_in_j [simplified suspends⇩j [symmetric]]]
obtain t v ys zs where
split_suspends⇩j: "suspends⇩j = ys @ Read⇩s⇩b False a' t v# zs" (is "suspends⇩j = ?suspends") and
a'_unacq: "a' ∉ acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j)"
by blast
from direct_memop_step.Ghost [where θ=θ⇩s⇩b and m="flush ?drop_sb m"]
have "(Ghost A L R W# is⇩s⇩b',
θ⇩s⇩b, (), flush ?drop_sb m, 𝒟⇩s⇩b,
acquired True sb 𝒪⇩s⇩b, release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b, share ?drop_sb 𝒮) →
(is⇩s⇩b', θ⇩s⇩b, (), flush ?drop_sb m, 𝒟⇩s⇩b,
acquired True sb 𝒪⇩s⇩b ∪ A - R,
augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b),
share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF
i_bound_ts' [simplified is⇩s⇩b] ts'_i [simplified is⇩s⇩b] this [simplified is⇩s⇩b]]
have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) ⇒⇩d
(?ts'[i := (p⇩s⇩b, is⇩s⇩b', θ⇩s⇩b, (),𝒟⇩s⇩b, acquired True sb 𝒪⇩s⇩b ∪ A - R,augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b))],
flush ?drop_sb m,share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
(is " _ ⇒⇩d (?ts_A, ?m_A, ?share_A)")
by (simp add: is⇩s⇩b)
from i_bound' have i_bound'': "i < length ?ts_A"
by simp
from valid_program_history [OF j_bound'' ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from ts⇩j neq_i_j j_bound
have ts_A_j: "?ts_A!j = (hd_prog p⇩j (ys @ Read⇩s⇩b False a' t v# zs), is⇩j,
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (ys @ Read⇩s⇩b False a' t v# zs)), (),𝒟⇩j,
acquired True ?take_sb⇩j 𝒪⇩j,release ?take_sb⇩j (dom 𝒮⇩s⇩b) ℛ⇩j)"
by (simp add: split_suspends⇩j)
from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
by simp
from valid_last_prog [OF j_bound'' ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
then
have lp: "last_prog p⇩j ?suspends = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_reads [OF j_bound'' ts⇩s⇩b_j]
have reads_consis: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b› j_bound''
ts⇩s⇩b_j reads_consis]
have reads_consis_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
have "outstanding_refs is_Write⇩s⇩b ?drop_sb ∩ outstanding_refs is_non_volatile_Read⇩s⇩b suspends⇩j = {}"
by (simp add: suspends⇩j)
from reads_consistent_flush_independent [OF this reads_consis_m]
have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
?m_A suspends⇩j".
hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) ?m_A ys"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_history [OF j_bound'' ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from valid_write_sops [OF j_bound'' ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops ys. valid_sop sop"
apply (simp only: write_sops_append )
apply auto
done
from read_tmps_distinct [OF j_bound'' ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop h_consis]
last_prog_hd_prog
have hist_consis': "history_consistent θ⇩s⇩b⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b ys = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
from flush_store_buffer_append [
OF j_bound'''' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts_A_j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_m_A_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="?share_A"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs (Read⇩s⇩b False a' t v # zs) @ is⇩s⇩b⇩j =
is⇩j' @ prog_instrs (Read⇩s⇩b False a' t v # zs)" and
steps_ys: "(?ts_A, ?m_A, ?share_A) ⇒⇩d⇧*
(?ts_A[j:= (last_prog (hd_prog p⇩j (Ghost⇩s⇩b A' L' R' W'# zs)) ys,
is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Read⇩s⇩b False a' t v # zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j') ],
flush ys ?m_A,
share ys ?share_A)"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto)
note conflict_computation = rtranclp_trans [OF rtranclp_r_rtranclp [OF steps_flush_sb, OF store_step] steps_ys]
from cph
have "causal_program_history is⇩s⇩b⇩j ((ys @ [Read⇩s⇩b False a' t v]) @ zs)"
by simp
from causal_program_history_suffix [OF this]
have cph': "causal_program_history is⇩s⇩b⇩j zs".
interpret causal⇩j: causal_program_history "is⇩s⇩b⇩j" "zs" by (rule cph')
from causal⇩j.causal_program_history [of "[]", simplified, OF refl] is⇩j'
obtain is⇩j''
where is⇩j': "is⇩j' = Read False a' t#is⇩j''" and
is⇩j'': "instrs zs @ is⇩s⇩b⇩j = is⇩j'' @ prog_instrs zs"
by clarsimp
from j_bound'''
have j_bound_ys: "j < length ?ts_ys"
by auto
from j_bound_ys neq_i_j
have ts_ys_j: "?ts_ys!j=(last_prog (hd_prog p⇩j (Read⇩s⇩b False a' t v# zs)) ys, is⇩j',
θ⇩s⇩b⇩j |` (dom θ⇩s⇩b⇩j - read_tmps (Read⇩s⇩b False a' t v# zs)),(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {},
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j),ℛ⇩j')"
by auto
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
from safe_delayedE [OF this j_bound_ys ts_ys_j, simplified is⇩j']
have "a' ∈ acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∨
a' ∈ read_only (share ys (share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))"
apply cases
apply (auto simp add: Let_def is⇩s⇩b)
done
with a'_unacq
have a'_ro: "a' ∈ read_only (share ys (share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))"
by auto
from a'_in
have a'_not_ro: "a' ∉ read_only (share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs)
have "a' ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
proof -
{
assume a_notin: "a' ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
from weak_sharing_consis [OF j_bound'' ts⇩s⇩b_j]
have "weak_sharing_consistent 𝒪⇩j sb⇩j".
with weak_sharing_consistent_append [of 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
have "weak_sharing_consistent (acquired True ?take_sb⇩j 𝒪⇩j) suspends⇩j"
by (auto simp add: suspends⇩j)
with split_suspends⇩j
have weak_consis: "weak_sharing_consistent (acquired True ?take_sb⇩j 𝒪⇩j) ys"
by (simp add: weak_sharing_consistent_append)
from all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have "all_acquired ys ⊆ all_acquired sb⇩j"
apply (clarsimp)
apply (clarsimp simp add: suspends⇩j [symmetric] split_suspends⇩j all_acquired_append)
done
with a_notin acquired_takeWhile_non_volatile_Write⇩s⇩b [of sb⇩j 𝒪⇩j]
all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have "a' ∉ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j ∪ all_acquired ys"
by auto
from read_only_share_unowned [OF weak_consis this a'_ro]
have "a' ∈ read_only (share ?drop_sb 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)" .
with a'_not_ro have False
by auto
}
thus ?thesis by blast
qed
moreover
from A_unaquired_by_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
have "A ∩ all_acquired sb⇩j = {}"
by (auto simp add: Let_def)
moreover
from A_unowned_by_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
have "A ∩ 𝒪⇩j = {}"
by (auto simp add: Let_def dest: all_shared_acquired_in)
moreover note a'_in
ultimately
show False
by auto
qed
}
thus ?thesis
by (auto simp add: Let_def)
qed
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof -
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b (sb @ [Ghost⇩s⇩b A L R W]) "
by (auto simp add: non_volatile_owned_or_read_only_append)
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof (unfold_locales)
fix i⇩1 j p⇩1 "is⇩1" 𝒪⇩1 ℛ⇩1 𝒟⇩1 xs⇩1 sb⇩1 p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume i⇩1_bound: "i⇩1 < length ts⇩s⇩b'"
assume j_bound: "j < length ts⇩s⇩b'"
assume i⇩1_j: "i⇩1 ≠ j"
assume ts_i⇩1: "ts⇩s⇩b'!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
assume ts_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}"
proof (cases "i⇩1=i")
case True
with i⇩1_j have i_j: "i≠j"
by simp
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
hence j_bound'': "j < length (map owned ts⇩s⇩b)"
by simp
from ts_j i_j have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from outstanding_volatile_writes_unowned_by_others
[OF i_bound j_bound' i_j ts⇩s⇩b_i ts_j']
have "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb = {}".
with ts_i⇩1 True i_bound show ?thesis
by (clarsimp simp add: ts⇩s⇩b' sb' outstanding_refs_append
acquired_takeWhile_non_volatile_Write⇩s⇩b)
next
case False
note i⇩1_i = this
from i⇩1_bound have i⇩1_bound': "i⇩1 < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_i⇩1 False have ts_i⇩1': "ts⇩s⇩b!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
by (simp add: ts⇩s⇩b')
show ?thesis
proof (cases "j=i")
case True
from i⇩1_bound'
have i⇩1_bound'': "i⇩1 < length (map owned ts⇩s⇩b)"
by simp
from outstanding_volatile_writes_unowned_by_others
[OF i⇩1_bound' i_bound i⇩1_i ts_i⇩1' ts⇩s⇩b_i]
have "(𝒪⇩s⇩b ∪ all_acquired sb) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}".
moreover
from A_unused_by_others [rule_format, OF _ False [symmetric]] False ts_i⇩1 i⇩1_bound
have "A ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}"
by (auto simp add: Let_def ts⇩s⇩b')
ultimately
show ?thesis
using ts_j True ts⇩s⇩b'
by (auto simp add: i_bound ts⇩s⇩b' 𝒪⇩s⇩b' sb' all_acquired_append)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_j False have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from outstanding_volatile_writes_unowned_by_others
[OF i⇩1_bound' j_bound' i⇩1_j ts_i⇩1' ts_j']
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}" .
qed
qed
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof
fix n m
fix p⇩n "is⇩n" 𝒪⇩n ℛ⇩n 𝒟⇩n θ⇩n sb⇩n p⇩m "is⇩m" 𝒪⇩m ℛ⇩m 𝒟⇩m θ⇩m sb⇩m
assume n_bound: "n < length ts⇩s⇩b'"
and m_bound: "m < length ts⇩s⇩b'"
and neq_n_m: "n≠m"
and nth: "ts⇩s⇩b'!n = (p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
and mth: "ts⇩s⇩b'!m =(p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
from n_bound have n_bound': "n < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
from m_bound have m_bound': "m < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
show "(𝒪⇩m ∪ all_acquired sb⇩m) ∩
read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) 𝒪⇩n)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) =
{}"
proof (cases "m=i")
case True
with neq_n_m have neq_n_i: "n≠i"
by auto
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
note read_only_reads_unowned [OF n_bound' i_bound neq_n_i nth' ts⇩s⇩b_i]
moreover
from A_no_read_only_reads_by_others [rule_format, OF _ neq_n_i [symmetric]] n_bound' nth'
have "A ∩ read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) 𝒪⇩n)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) =
{}"
by auto
ultimately
show ?thesis
using True ts⇩s⇩b_i nth' mth n_bound' m_bound'
by (auto simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb' all_acquired_append)
next
case False
note neq_m_i = this
with m_bound mth i_bound have mth': "ts⇩s⇩b!m = (p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
by (auto simp add: ts⇩s⇩b')
show ?thesis
proof (cases "n=i")
case True
note read_only_reads_unowned [OF i_bound m_bound' neq_m_i [symmetric] ts⇩s⇩b_i mth']
then show ?thesis
using True neq_m_i ts⇩s⇩b_i nth mth n_bound' m_bound'
apply (case_tac "outstanding_refs (is_volatile_Write⇩s⇩b) sb = {}")
apply (clarsimp simp add: outstanding_vol_write_take_drop_appends
acquired_append read_only_reads_append ts⇩s⇩b' sb' 𝒪⇩s⇩b')+
done
next
case False
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
from read_only_reads_unowned [OF n_bound' m_bound' neq_n_m nth' mth'] False neq_m_i
show ?thesis
by (clarsimp)
qed
qed
qed
next
show "ownership_distinct ts⇩s⇩b'"
proof -
have "∀j<length ts⇩s⇩b. i ≠ j ⟶
(let (p⇩j, is⇩j,θ⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j,ℛ⇩j) = ts⇩s⇩b ! j
in (𝒪⇩s⇩b ∪ all_acquired sb') ∩ (𝒪⇩j ∪ all_acquired sb⇩j) = {})"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j
assume neq_i_j: "i ≠ j"
assume j_bound: "j < length ts⇩s⇩b"
assume ts⇩s⇩b_j: "ts⇩s⇩b ! j = (p⇩j, is⇩j, θ⇩j, sb⇩j, 𝒟⇩j, 𝒪⇩j,ℛ⇩j)"
have "(𝒪⇩s⇩b ∪ all_acquired sb') ∩ (𝒪⇩j ∪ all_acquired sb⇩j) = {}"
proof -
{
fix a'
assume a'_in_i: "a' ∈ (𝒪⇩s⇩b ∪ all_acquired sb')"
assume a'_in_j: "a' ∈ (𝒪⇩j ∪ all_acquired sb⇩j)"
have False
proof -
from a'_in_i have "a' ∈ (𝒪⇩s⇩b ∪ all_acquired sb) ∨ a' ∈ A"
by (simp add: sb' all_acquired_append)
then show False
proof
assume "a' ∈ (𝒪⇩s⇩b ∪ all_acquired sb)"
with ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j] a'_in_j
show ?thesis
by auto
next
assume "a' ∈ A"
moreover
have j_bound': "j < length (map owned ts⇩s⇩b)"
using j_bound by auto
from A_unowned_by_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
obtain "A ∩ acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j = {}" and
"A ∩ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) = {}"
by (auto simp add: Let_def)
moreover
from A_unaquired_by_others [rule_format, OF _ neq_i_j] ts⇩s⇩b_j j_bound
have "A ∩ all_acquired sb⇩j = {}"
by auto
ultimately
show ?thesis
using a'_in_j
by (auto dest: all_shared_acquired_in)
qed
qed
}
then show ?thesis by auto
qed
}
then show ?thesis by (fastforce simp add: Let_def)
qed
from ownership_distinct_nth_update [OF i_bound ts⇩s⇩b_i this]
show ?thesis by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb')
qed
qed
have valid_hist': "valid_history program_step ts⇩s⇩b'"
proof -
from valid_history [OF i_bound ts⇩s⇩b_i]
have "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b sb) sb".
with valid_write_sops [OF i_bound ts⇩s⇩b_i]
valid_implies_valid_prog_hd [OF i_bound ts⇩s⇩b_i valid]
have "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b (sb@[Ghost⇩s⇩b A L R W]))
(sb@ [Ghost⇩s⇩b A L R W])"
apply -
apply (rule history_consistent_appendI)
apply (auto simp add: hd_prog_append_Ghost⇩s⇩b)
done
from valid_history_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' θ⇩s⇩b')
qed
have valid_reads': "valid_reads m⇩s⇩b ts⇩s⇩b'"
proof -
from valid_reads [OF i_bound ts⇩s⇩b_i]
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b sb" .
from reads_consistent_snoc_Ghost⇩s⇩b [OF this]
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b (sb @ [Ghost⇩s⇩b A L R W])".
from valid_reads_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b')
qed
have valid_sharing': "valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
from outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i]
have "non_volatile_writes_unshared 𝒮⇩s⇩b (sb @ [Ghost⇩s⇩b A L R W])"
by (auto simp add: non_volatile_writes_unshared_append)
from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
show "outstanding_non_volatile_writes_unshared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' 𝒮⇩s⇩b')
next
from sharing_consis [OF i_bound ts⇩s⇩b_i]
have consis': "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b sb".
from A_shared_owned
have "A ⊆ dom (share ?drop_sb 𝒮) ∪ acquired True sb 𝒪⇩s⇩b"
by (simp add: sharing_consistent_append acquired_takeWhile_non_volatile_Write⇩s⇩b)
moreover have "dom (share ?drop_sb 𝒮) ⊆ dom 𝒮 ∪ dom (share sb 𝒮⇩s⇩b)"
proof
fix a'
assume a'_in: "a' ∈ dom (share ?drop_sb 𝒮)"
from share_unshared_in [OF a'_in]
show "a' ∈ dom 𝒮 ∪ dom (share sb 𝒮⇩s⇩b)"
proof
assume "a' ∈ dom (share ?drop_sb Map.empty)"
from share_mono_in [OF this] share_append [of ?take_sb ?drop_sb]
have "a' ∈ dom (share sb 𝒮⇩s⇩b)"
by auto
thus ?thesis
by simp
next
assume "a' ∈ dom 𝒮 ∧ a' ∉ all_unshared ?drop_sb"
thus ?thesis by auto
qed
qed
ultimately
have A_subset: "A ⊆ dom 𝒮 ∪ dom (share sb 𝒮⇩s⇩b) ∪ acquired True sb 𝒪⇩s⇩b"
by auto
have "A ⊆ dom (share sb 𝒮⇩s⇩b) ∪ acquired True sb 𝒪⇩s⇩b"
proof -
{
fix x
assume x_A: "x ∈ A"
have "x ∈ dom (share sb 𝒮⇩s⇩b) ∪ acquired True sb 𝒪⇩s⇩b"
proof -
{
assume "x ∈ dom 𝒮"
from share_all_until_volatile_write_share_acquired [OF ‹sharing_consis 𝒮⇩s⇩b ts⇩s⇩b›
i_bound ts⇩s⇩b_i this [simplified 𝒮]]
A_unowned_by_others x_A
have ?thesis
by (fastforce simp add: Let_def)
}
with A_subset show ?thesis using x_A by auto
qed
}
thus ?thesis by blast
qed
with consis' L_subset A_R R_acq
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b (sb @ [Ghost⇩s⇩b A L R W])"
by (simp add: sharing_consistent_append acquired_takeWhile_non_volatile_Write⇩s⇩b)
from sharing_consis_nth_update [OF i_bound this]
show "sharing_consis 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb' 𝒮⇩s⇩b')
next
from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound ts⇩s⇩b_i] ]
show "read_only_unowned 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' 𝒪⇩s⇩b')
next
from unowned_shared_nth_update [OF i_bound ts⇩s⇩b_i subset_refl]
show "unowned_shared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' sb' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
have "no_write_to_read_only_memory 𝒮⇩s⇩b (sb @ [Ghost⇩s⇩b A L R W])"
by (simp add: no_write_to_read_only_memory_append)
from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
show "no_outstanding_write_to_read_only_memory 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' sb')
qed
have tmps_distinct': "tmps_distinct ts⇩s⇩b'"
proof (intro_locales)
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_load_tmps is⇩s⇩b'" by (simp add: "is⇩s⇩b")
from load_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b')
next
from read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_read_tmps (sb @ [Ghost⇩s⇩b A L R W])"
by (auto simp add: distinct_read_tmps_append)
from read_tmps_distinct_nth_update [OF i_bound this]
show "read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb')
next
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ read_tmps (sb @ [Ghost⇩s⇩b A L R W]) ={}"
by (auto simp add: read_tmps_append "is⇩s⇩b")
from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b' sb')
qed
have valid_sops': "valid_sops ts⇩s⇩b'"
proof -
from valid_store_sops [OF i_bound ts⇩s⇩b_i]
obtain
valid_store_sops': "∀sop∈store_sops is⇩s⇩b'. valid_sop sop"
by (auto simp add: "is⇩s⇩b")
from valid_write_sops [OF i_bound ts⇩s⇩b_i]
have valid_write_sops': "∀sop∈write_sops (sb@ [Ghost⇩s⇩b A L R W]).
valid_sop sop"
by (auto simp add: write_sops_append)
from valid_sops_nth_update [OF i_bound valid_write_sops' valid_store_sops']
show ?thesis by (simp add: ts⇩s⇩b' sb')
qed
have valid_dd': "valid_data_dependency ts⇩s⇩b'"
proof -
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
obtain
dd_is: "data_dependency_consistent_instrs (dom θ⇩s⇩b') is⇩s⇩b'"
by (auto simp add: "is⇩s⇩b" θ⇩s⇩b')
from load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ ⋃(fst ` write_sops (sb@ [Ghost⇩s⇩b A L R W])) ={}"
by (auto simp add: write_sops_append "is⇩s⇩b")
from valid_data_dependency_nth_update [OF i_bound dd_is this]
show ?thesis by (simp add: ts⇩s⇩b' sb')
qed
have load_tmps_fresh': "load_tmps_fresh ts⇩s⇩b'"
proof -
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b' ∩ dom θ⇩s⇩b = {}"
by (auto simp add: "is⇩s⇩b")
from load_tmps_fresh_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' θ⇩s⇩b')
qed
have enough_flushs': "enough_flushs ts⇩s⇩b'"
proof -
from clean_no_outstanding_volatile_Write⇩s⇩b [OF i_bound ts⇩s⇩b_i]
have "¬ 𝒟⇩s⇩b ⟶ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Ghost⇩s⇩b A L R W])= {}"
by (auto simp add: outstanding_refs_append)
from enough_flushs_nth_update [OF i_bound this]
show ?thesis
by (simp add: ts⇩s⇩b' sb' 𝒟⇩s⇩b')
qed
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
from valid_program_history [OF i_bound ts⇩s⇩b_i]
have "causal_program_history is⇩s⇩b sb" .
then have causal': "causal_program_history is⇩s⇩b' (sb@[Ghost⇩s⇩b A L R W])"
by (auto simp: causal_program_history_Ghost "is⇩s⇩b")
from valid_last_prog [OF i_bound ts⇩s⇩b_i]
have "last_prog p⇩s⇩b sb = p⇩s⇩b".
hence "last_prog p⇩s⇩b (sb @ [Ghost⇩s⇩b A L R W]) = p⇩s⇩b"
by (simp add: last_prog_append_Ghost⇩s⇩b)
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b' sb')
qed
show ?thesis
proof (cases "outstanding_refs is_volatile_Write⇩s⇩b sb = {}")
case True
from True have flush_all: "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = sb"
by (auto simp add: outstanding_refs_conv)
from True have suspend_nothing: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = []"
by (auto simp add: outstanding_refs_conv)
hence suspends_empty: "suspends = []"
by (simp add: suspends)
from suspends_empty is_sim have "is": "is =Ghost A L R W# is⇩s⇩b'"
by (simp add: "is⇩s⇩b")
with suspends_empty ts_i
have ts_i: "ts!i = (p⇩s⇩b, Ghost A L R W# is⇩s⇩b',
θ⇩s⇩b,(), 𝒟, acquired True ?take_sb 𝒪⇩s⇩b,release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
from direct_memop_step.Ghost
have "(Ghost A L R W# is⇩s⇩b',
θ⇩s⇩b, (),m, 𝒟, acquired True ?take_sb 𝒪⇩s⇩b,
release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b, 𝒮) →
(is⇩s⇩b',
θ⇩s⇩b, (), m, 𝒟, acquired True ?take_sb 𝒪⇩s⇩b ∪ A - R,
augment_rels (dom 𝒮) R (release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b),
𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
have "(ts, m, 𝒮) ⇒⇩d
(ts[i := (p⇩s⇩b, is⇩s⇩b',
θ⇩s⇩b, (),𝒟, acquired True ?take_sb 𝒪⇩s⇩b ∪ A - R,
augment_rels (dom 𝒮) R (release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b))],
m,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
moreover
from suspend_nothing
have suspend_nothing': "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') = []"
by (simp add: sb')
have all_shared_A: "∀j p is 𝒪 ℛ 𝒟 θ sb. j < length ts⇩s⇩b ⟶ i ≠ j ⟶
ts⇩s⇩b ! j = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ A = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume neq_i_j: "i ≠ j"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j, θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume x_shared: "x ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_A: "x ∈ A"
have False
proof -
from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
have "all_shared sb⇩j ⊆ all_acquired sb⇩j ∪ 𝒪⇩j".
moreover have "all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ⊆ all_shared sb⇩j"
using all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
by auto
moreover
from A_unaquired_by_others [rule_format, OF _ neq_i_j] jth j_bound
have "A ∩ all_acquired sb⇩j = {}" by auto
moreover
from A_unowned_by_others [rule_format, OF _ neq_i_j] jth j_bound
have "A ∩ 𝒪⇩j = {}"
by (auto dest: all_shared_acquired_in)
ultimately
show False
using x_A x_shared
by blast
qed
}
thus ?thesis by blast
qed
hence all_shared_L: "∀j p is 𝒪 ℛ 𝒟 θ sb. j < length ts⇩s⇩b ⟶ i ≠ j ⟶
ts⇩s⇩b ! j = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ L = {}"
using L_subset by blast
have all_shared_A: "∀j p is 𝒪 ℛ 𝒟 θ sb. j < length ts⇩s⇩b ⟶ i ≠ j ⟶
ts⇩s⇩b ! j = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ A = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume neq_i_j: "i ≠ j"
assume x_shared: "x ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_A: "x ∈ A"
have False
proof -
from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
have "all_shared sb⇩j ⊆ all_acquired sb⇩j ∪ 𝒪⇩j".
moreover have "all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ⊆ all_shared sb⇩j"
using all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
by auto
moreover
from A_unaquired_by_others [rule_format, OF _ neq_i_j] jth j_bound
have "A ∩ all_acquired sb⇩j = {}" by auto
moreover
from A_unowned_by_others [rule_format, OF _ neq_i_j] jth j_bound
have "A ∩ 𝒪⇩j = {}"
by (auto dest: all_shared_acquired_in)
ultimately
show False
using x_A x_shared
by blast
qed
}
thus ?thesis by blast
qed
hence all_shared_L: "∀j p is 𝒪 ℛ 𝒟 θ sb. j < length ts⇩s⇩b ⟶ i ≠ j ⟶
ts⇩s⇩b ! j = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ L = {}"
using L_subset by blast
have all_unshared_R: "∀j p is 𝒪 ℛ 𝒟 θ sb. j < length ts⇩s⇩b ⟶ i ≠ j ⟶
ts⇩s⇩b ! j = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ R = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume neq_i_j: "i ≠ j"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume x_unshared: "x ∈ all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_R: "x ∈ R"
have False
proof -
from unshared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
have "all_unshared sb⇩j ⊆ all_acquired sb⇩j ∪ 𝒪⇩j".
moreover have "all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ⊆ all_unshared sb⇩j"
using all_unshared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
by auto
moreover
note ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
ultimately
show False
using R_acq x_R x_unshared acquired_all_acquired [of True sb 𝒪⇩s⇩b]
by blast
qed
}
thus ?thesis by blast
qed
have all_acquired_R: "∀j p is 𝒪 ℛ 𝒟 θ sb. j < length ts⇩s⇩b ⟶ i ≠ j ⟶
ts⇩s⇩b ! j = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_acquired (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ R = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume neq_i_j: "i ≠ j"
assume x_acq: "x ∈ all_acquired (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_R: "x ∈ R"
have False
proof -
from x_acq have "x ∈ all_acquired sb⇩j"
using all_acquired_append [of "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
"dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"]
by auto
moreover
note ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
ultimately
show False
using R_acq x_R acquired_all_acquired [of True sb 𝒪⇩s⇩b]
by blast
qed
}
thus ?thesis by blast
qed
have all_shared_R: "∀j p is 𝒪 ℛ 𝒟 θ sb. j < length ts⇩s⇩b ⟶ i ≠ j ⟶
ts⇩s⇩b ! j = (p, is, θ, sb, 𝒟, 𝒪,ℛ) ⟶
all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∩ R = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume neq_i_j: "i ≠ j"
assume x_shared: "x ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_R: "x ∈ R"
have False
proof -
from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
have "all_shared sb⇩j ⊆ all_acquired sb⇩j ∪ 𝒪⇩j".
moreover have "all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ⊆ all_shared sb⇩j"
using all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
by auto
moreover
note ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
ultimately
show False
using R_acq x_R x_shared acquired_all_acquired [of True sb 𝒪⇩s⇩b]
by blast
qed
}
thus ?thesis by blast
qed
note share_commute =
share_all_until_volatile_write_append_Ghost⇩s⇩b [OF True ‹ownership_distinct ts⇩s⇩b› ‹sharing_consis 𝒮⇩s⇩b ts⇩s⇩b›
i_bound ts⇩s⇩b_i all_shared_L all_shared_A all_acquired_R all_unshared_R all_shared_R]
from 𝒟
have 𝒟': "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Ghost⇩s⇩b A L R W]) ≠ {})"
by (auto simp: outstanding_refs_append)
have "∀a ∈ R. (a ∈ (dom (share sb 𝒮⇩s⇩b)) ) = (a ∈ dom 𝒮)"
proof -
{
fix a
assume a_R: "a ∈ R"
have "(a ∈ (dom (share sb 𝒮⇩s⇩b)) ) = (a ∈ dom 𝒮)"
proof -
from a_R R_acq acquired_all_acquired [of True sb 𝒪⇩s⇩b]
have "a ∈ 𝒪⇩s⇩b ∪ all_acquired sb"
by auto
from share_all_until_volatile_write_thread_local' [OF ownership_distinct_ts⇩s⇩b sharing_consis_ts⇩s⇩b i_bound ts⇩s⇩b_i this] suspend_nothing
show ?thesis by (auto simp add: domIff 𝒮)
qed
}
then show ?thesis by auto
qed
from augment_rels_shared_exchange [OF this]
have rel_commute:
"augment_rels (dom 𝒮) R (release sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b) =
release (sb @ [Ghost⇩s⇩b A L R W]) (dom 𝒮⇩s⇩b') ℛ⇩s⇩b"
by (clarsimp simp add: release_append 𝒮⇩s⇩b')
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b') ∼
(ts[i := (p⇩s⇩b,is⇩s⇩b',
θ⇩s⇩b,(), 𝒟, acquired True ?take_sb 𝒪⇩s⇩b ∪ A - R,
augment_rels (dom 𝒮) R (release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b))],
m,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
apply (rule sim_config.intros)
apply (simp add: m ts⇩s⇩b' 𝒪⇩s⇩b' sb' θ⇩s⇩b' flush_all_until_volatile_write_append_Ghost_commute [OF i_bound ts⇩s⇩b_i])
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b' θ⇩s⇩b' share_commute)
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim ts_i True 𝒟'
apply (clarsimp simp add: Let_def nth_list_update
outstanding_refs_conv ts⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b' 𝒮⇩s⇩b' θ⇩s⇩b' sb' 𝒟⇩s⇩b' suspend_nothing' flush_all rel_commute
acquired_append split: if_split_asm)
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct'
valid_sops'
valid_dd' load_tmps_fresh' enough_flushs'
valid_program_history' valid' m⇩s⇩b' 𝒮⇩s⇩b' ℛ⇩s⇩b'
by auto
next
case False
then obtain r where r_in: "r ∈ set sb" and volatile_r: "is_volatile_Write⇩s⇩b r"
by (auto simp add: outstanding_refs_conv)
from takeWhile_dropWhile_real_prefix
[OF r_in, of "(Not ∘ is_volatile_Write⇩s⇩b)", simplified, OF volatile_r]
obtain a' v' sb'' A'' L'' R'' W'' sop' where
sb_split: "sb = takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ Write⇩s⇩b True a' sop' v' A'' L'' R'' W''# sb''"
and
drop: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = Write⇩s⇩b True a' sop' v' A'' L'' R'' W''# sb''"
apply (auto)
subgoal for y ys
apply (case_tac y)
apply auto
done
done
from drop suspends have suspends: "suspends = Write⇩s⇩b True a' sop' v' A'' L'' R'' W''# sb''"
by simp
have "(ts, m, 𝒮) ⇒⇩d⇧* (ts, m, 𝒮)" by auto
moreover
have "Write⇩s⇩b True a' sop' v' A'' L'' R'' W''∈ set sb"
by (subst sb_split) auto
note drop_app = dropWhile_append1
[OF this, of "(Not ∘ is_volatile_Write⇩s⇩b)", simplified]
from takeWhile_append1 [where P="Not ∘ is_volatile_Write⇩s⇩b", OF r_in] volatile_r
have takeWhile_app:
"(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb @ [Ghost⇩s⇩b A L R W])) = (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
by simp
note share_commute = share_all_until_volatile_write_append_Ghost⇩s⇩b' [OF False i_bound ts⇩s⇩b_i]
from 𝒟
have 𝒟': "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Ghost⇩s⇩b A L R W]) ≠ {})"
by (auto simp: outstanding_refs_append)
have "(ts⇩s⇩b',m⇩s⇩b,𝒮⇩s⇩b') ∼ (ts,m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_all_until_volatile_write_append_Ghost_commute [OF i_bound ts⇩s⇩b_i] ts⇩s⇩b' 𝒪⇩s⇩b' θ⇩s⇩b' sb')
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' ts⇩s⇩b' sb' 𝒪⇩s⇩b' θ⇩s⇩b' share_commute)
using leq
apply (simp add: ts⇩s⇩b')
using i_bound i_bound' ts_sim ts_i is_sim 𝒟'
apply (clarsimp simp add: Let_def nth_list_update is_sim drop_app
read_tmps_append suspends
prog_instrs_append_Ghost⇩s⇩b instrs_append_Ghost⇩s⇩b hd_prog_append_Ghost⇩s⇩b
drop "is⇩s⇩b" ts⇩s⇩b' sb' 𝒪⇩s⇩b' ℛ⇩s⇩b' 𝒮⇩s⇩b' θ⇩s⇩b' 𝒟⇩s⇩b' takeWhile_app split: if_split_asm)
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' valid_dd'
valid_sops' load_tmps_fresh' enough_flushs'
valid_program_history' valid' m⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply )
qed
qed
next
case (StoreBuffer i p⇩s⇩b "is⇩s⇩b" θ⇩s⇩b sb 𝒟⇩s⇩b 𝒪⇩s⇩b ℛ⇩s⇩b sb' 𝒪⇩s⇩b' ℛ⇩s⇩b')
then obtain
ts⇩s⇩b': "ts⇩s⇩b' = ts⇩s⇩b[i := (p⇩s⇩b, is⇩s⇩b, θ⇩s⇩b, sb', 𝒟⇩s⇩b, 𝒪⇩s⇩b',ℛ⇩s⇩b')]" and
i_bound: "i < length ts⇩s⇩b" and
ts⇩s⇩b_i: "ts⇩s⇩b ! i = (p⇩s⇩b, is⇩s⇩b, θ⇩s⇩b,sb, 𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ⇩s⇩b)" and
flush: "(m⇩s⇩b,sb,𝒪⇩s⇩b,ℛ⇩s⇩b,𝒮⇩s⇩b) →⇩f
(m⇩s⇩b',sb',𝒪⇩s⇩b',ℛ⇩s⇩b',𝒮⇩s⇩b')"
by auto
from sim obtain
m: "m = flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b" and
𝒮: "𝒮 = share_all_until_volatile_write ts⇩s⇩b 𝒮⇩s⇩b" and
leq: "length ts⇩s⇩b = length ts" and
ts_sim: "∀i<length ts⇩s⇩b.
let (p, is⇩s⇩b, θ, sb,𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ) = ts⇩s⇩b ! i;
suspends = dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb
in ∃is 𝒟. instrs suspends @ is⇩s⇩b = is @ prog_instrs suspends ∧
𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b sb ≠ {}) ∧
ts ! i =
(hd_prog p suspends,
is,
θ |` (dom θ - read_tmps suspends), (),
𝒟,
acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪⇩s⇩b,
release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) (dom 𝒮⇩s⇩b) ℛ)"
by cases blast
from i_bound leq have i_bound': "i < length ts"
by auto
have split_sb: "sb = takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb"
(is "sb = ?take_sb@?drop_sb")
by simp
from ts_sim [rule_format, OF i_bound] ts⇩s⇩b_i obtain suspends "is" 𝒟 where
suspends: "suspends = dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb" and
is_sim: "instrs suspends @ is⇩s⇩b = is @ prog_instrs suspends" and
𝒟: "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b sb ≠ {})" and
ts_i: "ts ! i =
(hd_prog p⇩s⇩b suspends, is,
θ⇩s⇩b |` (dom θ⇩s⇩b - read_tmps suspends), (),𝒟, acquired True ?take_sb 𝒪⇩s⇩b,
release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by (auto simp add: Let_def)
from flush_step_preserves_valid [OF i_bound ts⇩s⇩b_i flush valid]
have valid': "valid ts⇩s⇩b'"
by (simp add: ts⇩s⇩b')
from flush obtain r where sb: "sb=r#sb'"
by (cases) auto
from valid_history [OF i_bound ts⇩s⇩b_i]
have "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b sb) sb".
then
have hist_consis': "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b sb') sb'"
by (auto simp add: sb intro: history_consistent_hd_prog
split: memref.splits option.splits)
from valid_history_nth_update [OF i_bound this]
have valid_hist': "valid_history program_step ts⇩s⇩b'" by (simp add: ts⇩s⇩b')
from read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have dist_sb': "distinct_read_tmps sb'"
by (simp add: sb split: memref.splits)
have tmps_distinct': "tmps_distinct ts⇩s⇩b'"
proof (intro_locales)
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_load_tmps is⇩s⇩b".
from load_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_distinct ts⇩s⇩b'"
by (simp add: ts⇩s⇩b')
next
from read_tmps_distinct_nth_update [OF i_bound dist_sb']
show "read_tmps_distinct ts⇩s⇩b'"
by (simp add: ts⇩s⇩b')
next
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b ∩ read_tmps sb' = {}"
by (auto simp add: sb split: memref.splits)
from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b')
qed
from load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b ∩ ⋃(fst ` write_sops sb') = {}"
by (auto simp add: sb split: memref.splits)
from valid_data_dependency_nth_update
[OF i_bound data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i] this]
have valid_dd': "valid_data_dependency ts⇩s⇩b'"
by (simp add: ts⇩s⇩b')
from valid_store_sops [OF i_bound ts⇩s⇩b_i] valid_write_sops [OF i_bound ts⇩s⇩b_i]
valid_sops_nth_update [OF i_bound]
have valid_sops': "valid_sops ts⇩s⇩b'"
by (cases r) (auto simp add: sb ts⇩s⇩b')
have load_tmps_fresh': "load_tmps_fresh ts⇩s⇩b'"
proof -
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
have "load_tmps is⇩s⇩b ∩ dom θ⇩s⇩b = {}".
from load_tmps_fresh_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b')
qed
have enough_flushs': "enough_flushs ts⇩s⇩b'"
proof -
from clean_no_outstanding_volatile_Write⇩s⇩b [OF i_bound ts⇩s⇩b_i]
have "¬ 𝒟⇩s⇩b ⟶ outstanding_refs is_volatile_Write⇩s⇩b sb' = {}"
by (auto simp add: sb split: if_split_asm)
from enough_flushs_nth_update [OF i_bound this]
show ?thesis
by (simp add: ts⇩s⇩b' sb)
qed
show ?thesis
proof (cases r)
case (Write⇩s⇩b volatile a sop v A L R W)
from flush this
have m⇩s⇩b': "m⇩s⇩b' = (m⇩s⇩b(a := v))"
by cases (auto simp add: sb)
have non_volatile_owned: "¬ volatile ⟶ a ∈ 𝒪⇩s⇩b"
proof (cases volatile)
case True thus ?thesis by simp
next
case False
with outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]
have "a ∈ 𝒪⇩s⇩b"
by (simp add: sb Write⇩s⇩b)
thus ?thesis by simp
qed
have a_unowned_by_others:
"∀j < length ts⇩s⇩b. i ≠ j ⟶ (let (_,_,_,sb⇩j,_,𝒪⇩j,_) = ts⇩s⇩b ! j in
a ∉ 𝒪⇩j ∪ all_acquired sb⇩j)"
proof (unfold Let_def, clarify del: notI)
fix j p⇩j "is⇩j" 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume neq: "i ≠ j"
assume ts_j: "ts⇩s⇩b ! j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "a ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
proof (cases volatile)
case True
from outstanding_volatile_writes_unowned_by_others [OF i_bound j_bound neq
ts⇩s⇩b_i ts_j]
show ?thesis
by (simp add: sb Write⇩s⇩b True)
next
case False
with non_volatile_owned
have "a ∈ 𝒪⇩s⇩b"
by simp
with ownership_distinct [OF i_bound j_bound neq ts⇩s⇩b_i ts_j]
show ?thesis
by blast
qed
qed
from valid_reads [OF i_bound ts⇩s⇩b_i]
have reads_consis: "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b sb" .
{
fix j
fix p⇩j is⇩s⇩b⇩j 𝒪⇩j ℛ⇩j 𝒟⇩s⇩b⇩j θ⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume ts⇩s⇩b_j: "ts⇩s⇩b!j=(p⇩j,is⇩s⇩b⇩j,θ⇩j,sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)"
assume neq_i_j: "i≠j"
have "a ∉ outstanding_refs is_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
proof
assume "a ∈ outstanding_refs is_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
hence "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
by (simp add: outstanding_refs_is_non_volatile_Write⇩s⇩b_takeWhile_conv)
hence "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
using outstanding_refs_append [of _ "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
by auto
with non_volatile_owned_or_read_only_outstanding_non_volatile_writes
[OF outstanding_non_volatile_refs_owned_or_read_only [OF j_bound ts⇩s⇩b_j]]
have "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
with a_unowned_by_others [rule_format, OF j_bound neq_i_j] ts⇩s⇩b_j
show False
by auto
qed
}
note a_notin_others = this
from a_notin_others
have a_notin_others':
"∀j < length ts⇩s⇩b. i ≠ j ⟶
(let (_,_,_,sb⇩j,_,_,_) = ts⇩s⇩b!j in a ∉ outstanding_refs is_Write⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j))"
by (fastforce simp add: Let_def)
obtain D f where sop: "sop=(D,f)" by (cases sop) auto
from valid_history [OF i_bound ts⇩s⇩b_i] sop sb Write⇩s⇩b
obtain D_tmps: "D ⊆ dom θ⇩s⇩b" and f_v: "f θ⇩s⇩b = v" and
D_sb': "D ∩ read_tmps sb' = {}"
by auto
let ?θ = "(θ⇩s⇩b |` (dom θ⇩s⇩b - read_tmps sb'))"
from D_tmps D_sb'
have D_tmps': "D ⊆ dom ?θ"
by auto
from valid_write_sops [OF i_bound ts⇩s⇩b_i, rule_format, of sop]
have "valid_sop sop"
by (auto simp add: sb Write⇩s⇩b)
from this [simplified sop]
interpret valid_sop "(D,f)" .
from D_tmps D_sb'
have "((dom θ⇩s⇩b - read_tmps sb') ∩ D) = D"
by blast
with valid_sop [OF refl D_tmps] valid_sop [OF refl D_tmps'] f_v
have f_v': "f ?θ = v"
by auto
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
from valid_program_history [OF i_bound ts⇩s⇩b_i]
have "causal_program_history is⇩s⇩b sb" .
then have causal': "causal_program_history is⇩s⇩b sb'"
by (simp add: sb Write⇩s⇩b causal_program_history_def)
from valid_last_prog [OF i_bound ts⇩s⇩b_i]
have "last_prog p⇩s⇩b sb = p⇩s⇩b".
hence "last_prog p⇩s⇩b sb' = p⇩s⇩b"
by (simp add: sb Write⇩s⇩b)
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
show ?thesis
proof (cases volatile)
case True
note volatile = this
from flush Write⇩s⇩b volatile
obtain
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b ∪ A - R" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'= 𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L" and
ℛ⇩s⇩b': "ℛ⇩s⇩b' = Map.empty"
by cases (auto simp add: sb)
from sharing_consis [OF i_bound ts⇩s⇩b_i]
obtain
A_shared_owned: "A ⊆ dom 𝒮⇩s⇩b ∪ 𝒪⇩s⇩b" and
L_subset: "L ⊆ A" and
A_R: "A ∩ R = {}" and
R_owned: "R ⊆ 𝒪⇩s⇩b"
by (clarsimp simp add: sb Write⇩s⇩b volatile)
from sb Write⇩s⇩b True have take_empty: "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = []"
by (auto simp add: outstanding_refs_conv)
from sb Write⇩s⇩b True have suspend_all: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = sb"
by (auto simp add: outstanding_refs_conv)
hence suspends_all: "suspends = sb"
by (simp add: suspends)
from is_sim
have is_sim: "Write True a (D, f) A L R W# instrs sb' @ is⇩s⇩b = is @ prog_instrs sb'"
by (simp add: True Write⇩s⇩b suspends_all sb sop)
from valid_program_history [OF i_bound ts⇩s⇩b_i]
interpret causal_program_history "is⇩s⇩b" sb .
from valid_last_prog [OF i_bound ts⇩s⇩b_i]
have last_prog: "last_prog p⇩s⇩b sb = p⇩s⇩b".
from causal_program_history [of "[Write⇩s⇩b True a (D, f) v A L R W]" sb'] is_sim
obtain is' where
"is": "is = Write True a (D, f) A L R W# is'" and
is'_sim: "instrs sb'@is⇩s⇩b = is' @ prog_instrs sb'"
by (auto simp add: sb Write⇩s⇩b volatile sop)
from causal_program_history have
causal_program_history_sb': "causal_program_history is⇩s⇩b sb'"
apply -
apply (rule causal_program_history.intro)
apply (auto simp add: sb Write⇩s⇩b)
done
from ts_i have ts_i: "ts ! i =
(hd_prog p⇩s⇩b sb', Write True a (D, f) A L R W# is', ?θ, (), 𝒟,acquired True ?take_sb 𝒪⇩s⇩b,
release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by (simp add: suspends_all sb Write⇩s⇩b "is")
let ?ts' = "ts[i := (hd_prog p⇩s⇩b sb', is', ?θ, (), True, acquired True ?take_sb 𝒪⇩s⇩b ∪ A - R,
Map.empty)]"
from i_bound' have ts'_i: "?ts'!i = (hd_prog p⇩s⇩b sb', is', ?θ, (),True, acquired True ?take_sb 𝒪⇩s⇩b ∪ A - R,Map.empty)"
by simp
from no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
have a_not_ro: "a ∉ read_only 𝒮⇩s⇩b"
by (clarsimp simp add: sb Write⇩s⇩b volatile)
{
fix j
fix p⇩j is⇩s⇩b⇩j 𝒪⇩j ℛ⇩j 𝒟⇩s⇩b⇩j θ⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume ts⇩s⇩b_j: "ts⇩s⇩b!j=(p⇩j,is⇩s⇩b⇩j,θ⇩j,sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)"
assume neq_i_j: "i≠j"
have "a ∉ unforwarded_non_volatile_reads (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) {}"
proof
let ?take_sb⇩j = "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
let ?drop_sb⇩j = "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
assume a_in: "a ∈ unforwarded_non_volatile_reads ?drop_sb⇩j {}"
from a_unowned_by_others [rule_format, OF j_bound neq_i_j] ts⇩s⇩b_j
obtain a_unowned: "a ∉ 𝒪⇩j" and a_unacq: "a ∉ all_acquired sb⇩j"
by auto
with all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j] acquired_takeWhile_non_volatile_Write⇩s⇩b [of sb⇩j 𝒪⇩j]
have a_unacq_take: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j"
by (auto simp add: )
note nvo_j = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound ts⇩s⇩b_j]
from non_volatile_owned_or_read_only_drop [OF nvo_j]
have nvo_drop_j: "non_volatile_owned_or_read_only True (share ?take_sb⇩j 𝒮⇩s⇩b)
(acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j" .
note consis_j = sharing_consis [OF j_bound ts⇩s⇩b_j]
with sharing_consistent_append [of 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j ?drop_sb⇩j]
obtain consis_take_j: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j ?take_sb⇩j" and
consis_drop_j: "sharing_consistent (share ?take_sb⇩j 𝒮⇩s⇩b)
(acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by auto
from in_unforwarded_non_volatile_reads_non_volatile_Read⇩s⇩b [OF a_in]
have a_in': "a ∈ outstanding_refs is_non_volatile_Read⇩s⇩b ?drop_sb⇩j".
note reads_consis_j = valid_reads [OF j_bound ts⇩s⇩b_j]
from reads_consistent_drop [OF this]
have reads_consis_drop_j:
"reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) (flush ?take_sb⇩j m⇩s⇩b) ?drop_sb⇩j".
from read_only_share_all_shared [of a ?take_sb⇩j 𝒮⇩s⇩b] a_not_ro
all_shared_acquired_or_owned [OF consis_take_j]
all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j] a_unowned a_unacq
have a_not_ro_j: "a ∉ read_only (share ?take_sb⇩j 𝒮⇩s⇩b)"
by auto
from ts_sim [rule_format, OF j_bound] ts⇩s⇩b_j j_bound
obtain suspends⇩j "is⇩j" 𝒟⇩j ℛ⇩j where
suspends⇩j: "suspends⇩j = ?drop_sb⇩j" and
is⇩j: "instrs suspends⇩j @ is⇩s⇩b⇩j = is⇩j @ prog_instrs suspends⇩j" and
𝒟⇩j: "𝒟⇩s⇩b⇩j = (𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ≠ {})" and
ts⇩j: "ts!j = (hd_prog p⇩j suspends⇩j, is⇩j,
θ⇩j |` (dom θ⇩j - read_tmps suspends⇩j),(),
𝒟⇩j, acquired True ?take_sb⇩j 𝒪⇩j,ℛ⇩j)"
by (auto simp: Let_def)
from valid_last_prog [OF j_bound ts⇩s⇩b_j] have last_prog: "last_prog p⇩j sb⇩j = p⇩j".
from j_bound i_bound' leq have j_bound_ts': "j < length ts"
by simp
from read_only_read_acquired_unforwarded_acquire_witness [OF nvo_drop_j consis_drop_j
a_not_ro_j a_unacq_take a_in]
have False
proof
assume "∃sop a' v ys zs A L R W.
?drop_sb⇩j = ys @ Write⇩s⇩b True a' sop v A L R W # zs ∧ a ∈ A ∧
a ∉ outstanding_refs is_Write⇩s⇩b ys ∧ a'≠a"
with suspends⇩j
obtain a' sop' v' ys zs' A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Write⇩s⇩b True a' sop' v' A' L' R' W'# zs'" (is "suspends⇩j=?suspends") and
a_A': "a ∈ A'" and
no_write: "a ∉ outstanding_refs is_Write⇩s⇩b (ys @ [Write⇩s⇩b True a' sop' v' A' L' R' W'])"
by (auto simp add: outstanding_refs_append)
from last_prog
have lp: "last_prog p⇩j suspends⇩j = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from sharing_consis [OF j_bound ts⇩s⇩b_j]
have sharing_consis_j: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
then have A'_R': "A' ∩ R' = {}"
by (simp add: sharing_consistent_append [of _ _ ?take_sb⇩j ?drop_sb⇩j, simplified]
suspends⇩j [symmetric] split_suspends⇩j sharing_consistent_append)
from valid_program_history [OF j_bound ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from valid_reads [OF j_bound ts⇩s⇩b_j]
have reads_consis_j: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b›
j_bound ts⇩s⇩b_j this]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
hence reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
m (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W'])"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_write_sops [OF j_bound ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis_j]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b
(ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
have acq_simp:
"acquired True (ys @ [Write⇩s⇩b True a' sop' v' A' L' R' W'])
(acquired True ?take_sb⇩j 𝒪⇩j) =
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R'"
by (simp add: acquired_append)
from flush_store_buffer_append [where sb="ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']" and sb'="zs'", simplified,
OF j_bound_ts' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j] ts⇩j [simplified split_suspends⇩j]
refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="𝒮"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs zs' @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs'" and
steps_ys: "(ts, m, 𝒮) ⇒⇩d⇧*
(ts[j:=(last_prog
(hd_prog p⇩j (Write⇩s⇩b True a' sop' v' A' L' R' W'# zs')) (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']),
is⇩j',
θ⇩j |` (dom θ⇩j - read_tmps zs'),
(), True, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R',ℛ⇩j')],
flush (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) m,
share (ys@[Write⇩s⇩b True a' sop' v' A' L' R' W']) 𝒮)"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append outstanding_refs_append)
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j ts_i
have ts_ys_i: "?ts_ys!i = (hd_prog p⇩s⇩b sb', Write True a (D, f) A L R W# is', ?θ, (), 𝒟,
acquired True ?take_sb 𝒪⇩s⇩b,release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
note conflict_computation = steps_ys
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have safe: "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
with safe_delayedE [OF safe i_bound_ys ts_ys_i]
have a_unowned:
"∀j < length ?ts_ys. i≠j ⟶ (let (𝒪⇩j) = map owned ?ts_ys!j in a ∉ 𝒪⇩j)"
apply cases
apply (auto simp add: Let_def sb)
done
from a_A' a_unowned [rule_format, of j] neq_i_j j_bound leq A'_R'
show False
by (auto simp add: Let_def)
next
assume "∃A L R W ys zs. ?drop_sb⇩j = ys @ Ghost⇩s⇩b A L R W# zs ∧ a ∈ A ∧ a ∉ outstanding_refs is_Write⇩s⇩b ys"
with suspends⇩j
obtain ys zs' A' L' R' W' where
split_suspends⇩j: "suspends⇩j = ys @ Ghost⇩s⇩b A' L' R' W'# zs'" (is "suspends⇩j=?suspends") and
a_A': "a ∈ A'" and
no_write: "a ∉ outstanding_refs is_Write⇩s⇩b (ys @ [Ghost⇩s⇩b A' L' R' W'])"
by (auto simp add: outstanding_refs_append)
from last_prog
have lp: "last_prog p⇩j suspends⇩j = p⇩j"
apply -
apply (rule last_prog_same_append [where sb="?take_sb⇩j"])
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from valid_program_history [OF j_bound ts⇩s⇩b_j]
have "causal_program_history is⇩s⇩b⇩j sb⇩j".
then have cph: "causal_program_history is⇩s⇩b⇩j ?suspends"
apply -
apply (rule causal_program_history_suffix [where sb="?take_sb⇩j"] )
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp add: split_suspends⇩j)
done
from valid_reads [OF j_bound ts⇩s⇩b_j]
have reads_consis_j: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j".
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b›
j_bound ts⇩s⇩b_j this]
have reads_consis_m_j: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j) m suspends⇩j"
by (simp add: m suspends⇩j)
hence reads_consis_ys: "reads_consistent True (acquired True ?take_sb⇩j 𝒪⇩j)
m (ys@[Ghost⇩s⇩b A' L' R' W'])"
by (simp add: split_suspends⇩j reads_consistent_append)
from valid_write_sops [OF j_bound ts⇩s⇩b_j]
have "∀sop∈write_sops (?take_sb⇩j@?suspends). valid_sop sop"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain valid_sops_take: "∀sop∈write_sops ?take_sb⇩j. valid_sop sop" and
valid_sops_drop: "∀sop∈write_sops (ys@[Ghost⇩s⇩b A' L' R' W']). valid_sop sop"
apply (simp only: write_sops_append)
apply auto
done
from read_tmps_distinct [OF j_bound ts⇩s⇩b_j]
have "distinct_read_tmps (?take_sb⇩j@suspends⇩j)"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
then obtain
read_tmps_take_drop: "read_tmps ?take_sb⇩j ∩ read_tmps suspends⇩j = {}" and
distinct_read_tmps_drop: "distinct_read_tmps suspends⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply (simp only: distinct_read_tmps_append)
done
from valid_history [OF j_bound ts⇩s⇩b_j]
have h_consis:
"history_consistent θ⇩j (hd_prog p⇩j (?take_sb⇩j@suspends⇩j)) (?take_sb⇩j@suspends⇩j)"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
apply simp
done
from sharing_consis [OF j_bound ts⇩s⇩b_j]
have sharing_consis_j: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
then have A'_R': "A' ∩ R' = {}"
by (simp add: sharing_consistent_append [of _ _ ?take_sb⇩j ?drop_sb⇩j, simplified]
suspends⇩j [symmetric] split_suspends⇩j sharing_consistent_append)
have last_prog_hd_prog: "last_prog (hd_prog p⇩j sb⇩j) ?take_sb⇩j = (hd_prog p⇩j suspends⇩j)"
proof -
from last_prog have "last_prog p⇩j (?take_sb⇩j@?drop_sb⇩j) = p⇩j"
by simp
from last_prog_hd_prog_append' [OF h_consis] this
have "last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j = hd_prog p⇩j suspends⇩j"
by (simp only: split_suspends⇩j [symmetric] suspends⇩j)
moreover
have "last_prog (hd_prog p⇩j (?take_sb⇩j @ suspends⇩j)) ?take_sb⇩j =
last_prog (hd_prog p⇩j suspends⇩j) ?take_sb⇩j"
apply (simp only: split_suspends⇩j [symmetric] suspends⇩j)
by (rule last_prog_hd_prog_append)
ultimately show ?thesis
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
qed
from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop
h_consis] last_prog_hd_prog
have hist_consis': "history_consistent θ⇩j (hd_prog p⇩j suspends⇩j) suspends⇩j"
by (simp add: split_suspends⇩j [symmetric] suspends⇩j)
from reads_consistent_drop_volatile_writes_no_volatile_reads
[OF reads_consis_j]
have no_vol_read: "outstanding_refs is_volatile_Read⇩s⇩b
(ys@[Ghost⇩s⇩b A' L' R' W']) = {}"
by (auto simp add: outstanding_refs_append suspends⇩j [symmetric]
split_suspends⇩j )
have acq_simp:
"acquired True (ys @ [Ghost⇩s⇩b A' L' R' W'])
(acquired True ?take_sb⇩j 𝒪⇩j) =
acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R'"
by (simp add: acquired_append)
from flush_store_buffer_append [where sb="ys@[Ghost⇩s⇩b A' L' R' W']" and sb'="zs'", simplified,
OF j_bound_ts' is⇩j [simplified split_suspends⇩j] cph [simplified suspends⇩j]
ts⇩j [simplified split_suspends⇩j] refl lp [simplified split_suspends⇩j] reads_consis_ys
hist_consis' [simplified split_suspends⇩j] valid_sops_drop
distinct_read_tmps_drop [simplified split_suspends⇩j]
no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_vol_read], where
𝒮="𝒮"]
obtain is⇩j' ℛ⇩j' where
is⇩j': "instrs zs' @ is⇩s⇩b⇩j = is⇩j' @ prog_instrs zs'" and
steps_ys: "(ts, m,𝒮) ⇒⇩d⇧*
(ts[j:=(last_prog
(hd_prog p⇩j (Ghost⇩s⇩b A' L' R' W'# zs')) (ys@[Ghost⇩s⇩b A' L' R' W']),
is⇩j',
θ⇩j |` (dom θ⇩j - read_tmps zs'),
(),
𝒟⇩j ∨ outstanding_refs is_volatile_Write⇩s⇩b ys ≠ {}, acquired True ys (acquired True ?take_sb⇩j 𝒪⇩j) ∪ A' - R',ℛ⇩j')],
flush (ys@[Ghost⇩s⇩b A' L' R' W']) m, share (ys@[Ghost⇩s⇩b A' L' R' W']) 𝒮)"
(is "(_,_,_) ⇒⇩d⇧* (?ts_ys,?m_ys,?shared_ys)")
by (auto simp add: acquired_append outstanding_refs_append)
from i_bound' have i_bound_ys: "i < length ?ts_ys"
by auto
from i_bound' neq_i_j ts_i
have ts_ys_i: "?ts_ys!i = (hd_prog p⇩s⇩b sb', Write True a (D, f) A L R W# is', ?θ, (), 𝒟,
acquired True ?take_sb 𝒪⇩s⇩b,release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by simp
note conflict_computation = steps_ys
from safe_reach_safe_rtrancl [OF safe_reach conflict_computation]
have safe: "safe_delayed (?ts_ys,?m_ys,?shared_ys)".
with safe_delayedE [OF safe i_bound_ys ts_ys_i]
have a_unowned:
"∀j < length ?ts_ys. i≠j ⟶ (let (𝒪⇩j) = map owned ?ts_ys!j in a ∉ 𝒪⇩j)"
apply cases
apply (auto simp add: Let_def sb)
done
from a_A' a_unowned [rule_format, of j] neq_i_j j_bound leq A'_R'
show False
by (auto simp add: Let_def)
qed
then show False
by simp
qed
}
note a_notin_unforwarded_non_volatile_reads_drop = this
have valid_reads': "valid_reads m⇩s⇩b' ts⇩s⇩b'"
proof (unfold_locales)
fix j p⇩j "is⇩j" 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume ts_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "reads_consistent False 𝒪⇩j m⇩s⇩b' sb⇩j"
proof (cases "i=j")
case True
from reads_consis ts_j j_bound sb show ?thesis
by (clarsimp simp add: True m⇩s⇩b' Write⇩s⇩b ts⇩s⇩b' 𝒪⇩s⇩b' volatile reads_consistent_pending_write_antimono)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
moreover from ts_j False have ts_j': "ts⇩s⇩b ! j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
using j_bound by (simp add: ts⇩s⇩b')
ultimately have consis_m: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j"
by (rule valid_reads)
from a_unowned_by_others [rule_format, OF j_bound' False] ts_j'
have a_unowned:"a ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
by simp
let ?take_sb⇩j = "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
let ?drop_sb⇩j = "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
from a_unowned acquired_reads_all_acquired [of True ?take_sb⇩j 𝒪⇩j]
all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have a_not_acq_reads: "a ∉ acquired_reads True ?take_sb⇩j 𝒪⇩j"
by auto
moreover
note a_unfw= a_notin_unforwarded_non_volatile_reads_drop [OF j_bound' ts_j' False]
ultimately
show ?thesis
using reads_consistent_mem_eq_on_unforwarded_non_volatile_reads_drop [where W="{}" and
A="unforwarded_non_volatile_reads ?drop_sb⇩j {} ∪ acquired_reads True ?take_sb⇩j 𝒪⇩j" and
m'= "(m⇩s⇩b(a:=v))", OF _ _ _ consis_m]
by (fastforce simp add: m⇩s⇩b')
qed
qed
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof
fix j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j p⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume ts⇩s⇩b'_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "non_volatile_owned_or_read_only False 𝒮⇩s⇩b' 𝒪⇩j sb⇩j"
proof (cases "j=i")
case True
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]
have "non_volatile_owned_or_read_only False
(𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) (𝒪⇩s⇩b ∪ A - R) sb'"
by (auto simp add: sb Write⇩s⇩b volatile non_volatile_owned_or_read_only_pending_write_antimono)
then show ?thesis
using True i_bound ts⇩s⇩b'_j
by (auto simp add: ts⇩s⇩b' 𝒮⇩s⇩b' sb 𝒪⇩s⇩b')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
with ts⇩s⇩b'_j False i_bound
have ts⇩s⇩b_j: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
note nvo = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' ts⇩s⇩b_j]
from read_only_unowned [OF i_bound ts⇩s⇩b_i] R_owned
have "R ∩ read_only 𝒮⇩s⇩b = {}"
by auto
with read_only_reads_unowned [OF j_bound' i_bound False ts⇩s⇩b_j ts⇩s⇩b_i] L_subset
have "∀a∈read_only_reads
(acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j).
a ∈ read_only 𝒮⇩s⇩b ⟶ a ∈ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (auto simp add: in_read_only_convs sb Write⇩s⇩b volatile)
from non_volatile_owned_or_read_only_read_only_reads_eq' [OF nvo this]
have "non_volatile_owned_or_read_only False (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) 𝒪⇩j sb⇩j".
thus ?thesis by (simp add: 𝒮⇩s⇩b')
qed
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof (unfold_locales)
fix i⇩1 j p⇩1 "is⇩1" 𝒪⇩1 ℛ⇩1 𝒟⇩1 xs⇩1 sb⇩1 p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume i⇩1_bound: "i⇩1 < length ts⇩s⇩b'"
assume j_bound: "j < length ts⇩s⇩b'"
assume i⇩1_j: "i⇩1 ≠ j"
assume ts_i⇩1: "ts⇩s⇩b'!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
assume ts_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j, xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}"
proof (cases "i⇩1=i")
case True
from i⇩1_j True have neq_i_j: "i≠j"
by auto
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_j neq_i_j have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from outstanding_volatile_writes_unowned_by_others [OF i_bound j_bound' neq_i_j
ts⇩s⇩b_i ts_j'] ts_i⇩1 i_bound ts⇩s⇩b_i True show ?thesis
by (clarsimp simp add: ts⇩s⇩b' sb Write⇩s⇩b volatile)
next
case False
note i⇩1_i = this
from i⇩1_bound have i⇩1_bound': "i⇩1 < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b' sb)
hence i⇩1_bound'': "i⇩1 < length (map owned ts⇩s⇩b)"
by auto
from ts_i⇩1 False have ts_i⇩1': "ts⇩s⇩b!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
by (simp add: ts⇩s⇩b' sb)
show ?thesis
proof (cases "j=i")
case True
from outstanding_volatile_writes_unowned_by_others [OF i⇩1_bound' i_bound i⇩1_i ts_i⇩1' ts⇩s⇩b_i ]
have "(𝒪⇩s⇩b ∪ all_acquired sb) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}".
then show ?thesis
using True i⇩1_i ts_j ts⇩s⇩b_i i_bound
by (auto simp add: sb Write⇩s⇩b volatile ts⇩s⇩b' 𝒪⇩s⇩b')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_j False have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from outstanding_volatile_writes_unowned_by_others
[OF i⇩1_bound' j_bound' i⇩1_j ts_i⇩1' ts_j']
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}" .
qed
qed
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof
fix n m
fix p⇩n "is⇩n" 𝒪⇩n ℛ⇩n 𝒟⇩n θ⇩n sb⇩n p⇩m "is⇩m" 𝒪⇩m ℛ⇩m 𝒟⇩m θ⇩m sb⇩m
assume n_bound: "n < length ts⇩s⇩b'"
and m_bound: "m < length ts⇩s⇩b'"
and neq_n_m: "n≠m"
and nth: "ts⇩s⇩b'!n = (p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
and mth: "ts⇩s⇩b'!m =(p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
from n_bound have n_bound': "n < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
from m_bound have m_bound': "m < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
show "(𝒪⇩m ∪ all_acquired sb⇩m) ∩
read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) 𝒪⇩n)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) =
{}"
proof (cases "m=i")
case True
with neq_n_m have neq_n_i: "n≠i"
by auto
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
note read_only_reads_unowned [OF n_bound' i_bound neq_n_i nth' ts⇩s⇩b_i]
then
show ?thesis
using True ts⇩s⇩b_i neq_n_i nth mth n_bound' m_bound' L_subset
by (auto simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb Write⇩s⇩b volatile)
next
case False
note neq_m_i = this
with m_bound mth i_bound have mth': "ts⇩s⇩b!m = (p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
by (auto simp add: ts⇩s⇩b')
show ?thesis
proof (cases "n=i")
case True
from read_only_reads_append [of "(𝒪⇩s⇩b ∪ A - R)" "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n)"]
have "read_only_reads
(acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) (𝒪⇩s⇩b ∪ A - R))
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) ⊆ read_only_reads (𝒪⇩s⇩b ∪ A - R) sb⇩n"
by auto
with ts⇩s⇩b_i nth mth neq_m_i n_bound' True
read_only_reads_unowned [OF i_bound m_bound' False [symmetric] ts⇩s⇩b_i mth']
show ?thesis
by (auto simp add: ts⇩s⇩b' sb 𝒪⇩s⇩b' Write⇩s⇩b volatile)
next
case False
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
from read_only_reads_unowned [OF n_bound' m_bound' neq_n_m nth' mth'] False neq_m_i
show ?thesis
by (clarsimp)
qed
qed
qed
next
show "ownership_distinct ts⇩s⇩b'"
proof (unfold_locales)
fix i⇩1 j p⇩1 "is⇩1" 𝒪⇩1 ℛ⇩1 𝒟⇩1 xs⇩1 sb⇩1 p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume i⇩1_bound: "i⇩1 < length ts⇩s⇩b'"
assume j_bound: "j < length ts⇩s⇩b'"
assume i⇩1_j: "i⇩1 ≠ j"
assume ts_i⇩1: "ts⇩s⇩b'!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
assume ts_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "(𝒪⇩1 ∪ all_acquired sb⇩1) ∩ (𝒪⇩j ∪ all_acquired sb⇩j)= {}"
proof (cases "i⇩1=i")
case True
with i⇩1_j have i_j: "i≠j"
by simp
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
hence j_bound'': "j < length (map owned ts⇩s⇩b)"
by simp
from ts_j i_j have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from ownership_distinct [OF i_bound j_bound' i_j ts⇩s⇩b_i ts_j']
show ?thesis
using ts⇩s⇩b_i True ts_i⇩1 i_bound 𝒪⇩s⇩b'
by (auto simp add: ts⇩s⇩b' sb Write⇩s⇩b volatile)
next
case False
note i⇩1_i = this
from i⇩1_bound have i⇩1_bound': "i⇩1 < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
hence i⇩1_bound'': "i⇩1 < length (map owned ts⇩s⇩b)"
by simp
from ts_i⇩1 False have ts_i⇩1': "ts⇩s⇩b!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
by (simp add: ts⇩s⇩b')
show ?thesis
proof (cases "j=i")
case True
from ownership_distinct [OF i⇩1_bound' i_bound i⇩1_i ts_i⇩1' ts⇩s⇩b_i]
show ?thesis
using ts⇩s⇩b_i True ts_j i_bound 𝒪⇩s⇩b'
by (auto simp add: ts⇩s⇩b' sb Write⇩s⇩b volatile)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_j False have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from ownership_distinct [OF i⇩1_bound' j_bound' i⇩1_j ts_i⇩1' ts_j']
show ?thesis .
qed
qed
qed
qed
have valid_sharing': "valid_sharing (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_writes_unshared (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (unfold_locales)
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j acq⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "non_volatile_writes_unshared (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) sb⇩j"
proof (cases "i=j")
case True
with outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i]
i_bound jth ts⇩s⇩b_i show ?thesis
by (clarsimp simp add: ts⇩s⇩b' sb Write⇩s⇩b volatile)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
from jth False have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from outstanding_non_volatile_writes_unshared [OF j_bound' jth']
have unshared: "non_volatile_writes_unshared 𝒮⇩s⇩b sb⇩j".
have "∀a∈dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) - dom 𝒮⇩s⇩b. a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
proof -
{
fix a
assume a_in: "a ∈ dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) - dom 𝒮⇩s⇩b"
hence a_R: "a ∈ R"
by clarsimp
assume a_in_j: "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
have False
proof -
from non_volatile_owned_or_read_only_outstanding_non_volatile_writes [OF
outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' jth']]
a_in_j
have "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
moreover
with ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth'] a_R R_owned
show False
by blast
qed
}
thus ?thesis by blast
qed
from non_volatile_writes_unshared_no_outstanding_non_volatile_Write⇩s⇩b
[OF unshared this]
show ?thesis .
qed
qed
next
show "sharing_consis (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (unfold_locales)
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "sharing_consistent (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) 𝒪⇩j sb⇩j"
proof (cases "i=j")
case True
with i_bound jth ts⇩s⇩b_i sharing_consis [OF i_bound ts⇩s⇩b_i]
show ?thesis
by (clarsimp simp add: ts⇩s⇩b' sb Write⇩s⇩b volatile 𝒪⇩s⇩b')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
from jth False have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from sharing_consis [OF j_bound' jth']
have consis: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
have acq_cond: "all_acquired sb⇩j ∩ dom 𝒮⇩s⇩b - dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
proof -
{
fix a
assume a_acq: "a ∈ all_acquired sb⇩j"
assume "a ∈ dom 𝒮⇩s⇩b"
assume a_L: "a ∈ L"
have False
proof -
from ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth']
have "A ∩ all_acquired sb⇩j = {}"
by (auto simp add: sb Write⇩s⇩b volatile)
with a_acq a_L L_subset
show False
by blast
qed
}
thus ?thesis
by auto
qed
have uns_cond: "all_unshared sb⇩j ∩ dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) - dom 𝒮⇩s⇩b = {}"
proof -
{
fix a
assume a_uns: "a ∈ all_unshared sb⇩j"
assume "a ∉ L"
assume a_R: "a ∈ R"
have False
proof -
from unshared_acquired_or_owned [OF consis] a_uns
have "a ∈ all_acquired sb⇩j ∪ 𝒪⇩j" by auto
with ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth'] R_owned a_R
show False
by blast
qed
}
thus ?thesis
by auto
qed
from sharing_consistent_preservation [OF consis acq_cond uns_cond]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
qed
next
show "read_only_unowned (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "𝒪⇩j ∩ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
proof (cases "i=j")
case True
from read_only_unowned [OF i_bound ts⇩s⇩b_i] R_owned A_R
have "(𝒪⇩s⇩b ∪ A - R) ∩ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by (auto simp add: in_read_only_convs )
with jth ts⇩s⇩b_i i_bound True
show ?thesis
by (auto simp add: 𝒪⇩s⇩b' ts⇩s⇩b')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
with False jth have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from read_only_unowned [OF j_bound' jth']
have "𝒪⇩j ∩ read_only 𝒮⇩s⇩b = {}".
moreover
from ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth'] R_owned
have "(𝒪⇩s⇩b ∪ A) ∩ 𝒪⇩j = {}"
by (auto simp add: sb Write⇩s⇩b volatile)
moreover note R_owned A_R
ultimately show ?thesis
by (fastforce simp add: in_read_only_convs split: if_split_asm)
qed
qed
next
show "unowned_shared (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (unfold_locales)
show "- ⋃((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set ts⇩s⇩b') ⊆ dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L)"
proof -
have s: "⋃((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set ts⇩s⇩b') =
⋃((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set ts⇩s⇩b) ∪ A - R"
apply (unfold ts⇩s⇩b' 𝒪⇩s⇩b')
apply (rule acquire_release_ownership_nth_update [OF R_owned i_bound ts⇩s⇩b_i])
apply (rule local.ownership_distinct_axioms)
done
note unowned_shared L_subset A_R
then
show ?thesis
apply (simp only: s)
apply auto
done
qed
qed
next
show "no_outstanding_write_to_read_only_memory (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j acq⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "no_write_to_read_only_memory (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) sb⇩j"
proof (cases "i=j")
case True
with jth ts⇩s⇩b_i i_bound no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
show ?thesis
by (auto simp add: sb ts⇩s⇩b' Write⇩s⇩b volatile)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
with False jth have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from no_outstanding_write_to_read_only_memory [OF j_bound' jth']
have nw: "no_write_to_read_only_memory 𝒮⇩s⇩b sb⇩j".
have "R ∩ outstanding_refs is_Write⇩s⇩b sb⇩j = {}"
proof -
note dist = ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth']
from non_volatile_owned_or_read_only_outstanding_non_volatile_writes
[OF outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' jth']]
dist
have "outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j ∩ 𝒪⇩s⇩b = {}"
by auto
moreover
from outstanding_volatile_writes_unowned_by_others [OF j_bound' i_bound
False [symmetric] jth' ts⇩s⇩b_i ]
have "outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ∩ 𝒪⇩s⇩b = {}"
by auto
ultimately have "outstanding_refs is_Write⇩s⇩b sb⇩j ∩ 𝒪⇩s⇩b = {}"
by (auto simp add: misc_outstanding_refs_convs)
with R_owned
show ?thesis by blast
qed
then
have "∀a∈outstanding_refs is_Write⇩s⇩b sb⇩j.
a ∈ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ⟶ a ∈ read_only 𝒮⇩s⇩b"
by (auto simp add: in_read_only_convs)
from no_write_to_read_only_memory_read_only_reads_eq [OF nw this]
show ?thesis .
qed
qed
qed
from direct_memop_step.WriteVolatile [OF]
have "(Write True a (D, f) A L R W# is',
?θ, (), m,𝒟, acquired True ?take_sb 𝒪⇩s⇩b, release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b,𝒮) →
(is', ?θ, (), m (a := v),True, acquired True ?take_sb 𝒪⇩s⇩b ∪ A - R, Map.empty,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (simp add: f_v' [symmetric])
from direct_computation.Memop [OF i_bound' ts_i this]
have store_step:
"(ts, m, 𝒮) ⇒⇩d (?ts', m(a := v),𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)".
have sb'_split:
"sb' = takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb' @
dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb'"
by simp
from reads_consis
have no_vol_reads: "outstanding_refs is_volatile_Read⇩s⇩b sb' = {}"
by (simp add: sb Write⇩s⇩b True)
hence "outstanding_refs is_volatile_Read⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')
= {}"
by (auto simp add: outstanding_refs_conv dest: set_takeWhileD)
moreover
have "outstanding_refs is_volatile_Write⇩s⇩b
(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') = {}"
proof -
have "∀r ∈ set (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb'). ¬ (is_volatile_Write⇩s⇩b r)"
by (auto dest: set_takeWhileD)
thus ?thesis
by (simp add: outstanding_refs_conv)
qed
ultimately
have no_volatile:
"outstanding_refs is_volatile (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') = {}"
by (auto simp add: outstanding_refs_conv is_volatile_split)
moreover
from no_vol_reads have "∀r ∈ set sb'. ¬ is_volatile_Read⇩s⇩b r"
by (fastforce simp add: outstanding_refs_conv is_volatile_Read⇩s⇩b_def
split: memref.splits)
hence "∀r ∈ set sb'. (Not ∘ is_volatile_Write⇩s⇩b) r = (Not ∘ is_volatile) r"
by (auto simp add: is_volatile_split)
hence takeWhile_eq: "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') =
(takeWhile (Not ∘ is_volatile) sb')"
apply -
apply (rule takeWhile_cong)
apply auto
done
from leq
have leq': "length ts⇩s⇩b = length ?ts'"
by simp
hence i_bound_ts': "i < length ?ts'" using i_bound by simp
from is'_sim
have is'_sim_split:
"instrs
(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb' @
dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') @ is⇩s⇩b =
is' @ prog_instrs (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb' @
dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"
by (simp add: sb'_split [symmetric])
from reads_consistent_flush_all_until_volatile_write [OF ‹valid_ownership_and_sharing 𝒮⇩s⇩b ts⇩s⇩b›
i_bound ts⇩s⇩b_i reads_consis]
have "reads_consistent True (acquired True ?take_sb 𝒪⇩s⇩b) m (Write⇩s⇩b True a (D,f) v A L R W#sb')"
by (simp add: m sb Write⇩s⇩b volatile)
hence "reads_consistent True (acquired True ?take_sb 𝒪⇩s⇩b ∪ A - R) (m(a:=v)) sb'"
by simp
from reads_consistent_takeWhile [OF this]
have r_consis': "reads_consistent True (acquired True ?take_sb 𝒪⇩s⇩b ∪ A - R) (m(a:=v))
(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')".
from last_prog have last_prog_sb': "last_prog p⇩s⇩b sb' = p⇩s⇩b"
by (simp add: sb Write⇩s⇩b )
from valid_write_sops [OF i_bound ts⇩s⇩b_i]
have "∀sop ∈ write_sops sb'. valid_sop sop"
by (auto simp add: sb Write⇩s⇩b)
hence valid_sop': "∀sop∈write_sops (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb').
valid_sop sop"
by (fastforce dest: set_takeWhileD simp add: in_write_sops_conv)
from no_volatile
have no_volatile_Read⇩s⇩b:
"outstanding_refs is_volatile_Read⇩s⇩b (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') =
{}"
by (auto simp add: outstanding_refs_conv is_volatile_Read⇩s⇩b_def split: memref.splits)
from flush_store_buffer_append [OF i_bound_ts' is'_sim_split, simplified,
OF causal_program_history_sb' ts'_i refl last_prog_sb' r_consis' hist_consis'
valid_sop' dist_sb' no_volatile_Read⇩s⇩b_volatile_reads_consistent [OF no_volatile_Read⇩s⇩b],
where 𝒮="(𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"]
obtain is'' where
is''_sim: "instrs (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') @ is⇩s⇩b =
is'' @ prog_instrs (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')" and
steps: "(?ts', m(a := v), 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) ⇒⇩d⇧*
(ts[i := (last_prog (hd_prog p⇩s⇩b (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb'))
(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb'),
is'',
θ⇩s⇩b |` (dom θ⇩s⇩b -
read_tmps (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')),
(), True, acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')
(acquired True ?take_sb 𝒪⇩s⇩b ∪ A - R),
release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')
(dom (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)) Map.empty)],
flush (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') (m(a := v)),
share (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))"
by (auto)
note sim_flush = r_rtranclp_rtranclp [OF store_step steps]
moreover
note flush_commute =
flush_flush_all_until_volatile_write_Write⇩s⇩b_volatile_commute [OF i_bound ts⇩s⇩b_i [simplified sb Write⇩s⇩b True]
outstanding_refs_is_Write⇩s⇩b_takeWhile_disj a_notin_others']
from last_prog_hd_prog_append' [where sb="(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"
and sb'="(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')",
simplified sb'_split [symmetric], OF hist_consis' last_prog_sb']
have last_prog_eq:
"last_prog (hd_prog p⇩s⇩b (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb'))
(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') =
hd_prog p⇩s⇩b (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')".
have take_emtpy: "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) (r#sb) = []"
by (simp add: Write⇩s⇩b True)
have dist_sb': "∀i p is 𝒪 ℛ 𝒟 θ sb.
i < length ts⇩s⇩b ⟶
ts⇩s⇩b ! i = (p, is, θ, sb, 𝒟, 𝒪, ℛ) ⟶
(all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∪
all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∪
all_acquired (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)) ∩
(all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') ∪
all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') ∪
all_acquired (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')) =
{}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j, θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume x_shared: "x ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ∪
all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ∪
all_acquired (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_sb': "x ∈ (all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') ∪
all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') ∪
all_acquired (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb'))"
have False
proof (cases "i=j")
case True with x_shared ts⇩s⇩b_i jth show False by (simp add: sb volatile Write⇩s⇩b)
next
case False
from x_shared all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
unshared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
all_unshared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
all_acquired_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
have "x ∈ all_acquired sb⇩j ∪ 𝒪⇩j "
by auto
moreover
from x_sb' all_shared_acquired_or_owned [OF sharing_consis [OF i_bound ts⇩s⇩b_i]]
unshared_acquired_or_owned [OF sharing_consis [OF i_bound ts⇩s⇩b_i]]
all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"]
all_unshared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"]
all_acquired_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"]
have "x ∈ all_acquired sb ∪ 𝒪⇩s⇩b"
by (auto simp add: sb Write⇩s⇩b volatile)
moreover
note ownership_distinct [OF i_bound j_bound False ts⇩s⇩b_i jth]
ultimately show False by blast
qed
}
thus ?thesis by blast
qed
have dist_R_L_A: "∀j p is 𝒪 ℛ 𝒟 θ sb.
j < length ts⇩s⇩b ⟶ i≠ j⟶
ts⇩s⇩b ! j = (p, is, θ, sb, 𝒟, 𝒪, ℛ) ⟶
(all_shared sb ∪ all_unshared sb ∪ all_acquired sb) ∩ (R ∪ L ∪ A) = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume neq_i_j: "i ≠ j"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j, θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume x_shared: "x ∈ all_shared sb⇩j ∪
all_unshared sb⇩j ∪
all_acquired sb⇩j"
assume x_R_L_A: "x ∈ R ∪ L ∪ A"
have False
proof -
from x_shared all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
unshared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
have "x ∈ all_acquired sb⇩j ∪ 𝒪⇩j "
by auto
moreover
from x_R_L_A R_owned L_subset
have "x ∈ all_acquired sb ∪ 𝒪⇩s⇩b"
by (auto simp add: sb Write⇩s⇩b volatile)
moreover
note ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
ultimately show False by blast
qed
}
thus ?thesis by blast
qed
from local.ownership_distinct_axioms have "ownership_distinct ts⇩s⇩b" .
from local.sharing_consis_axioms have "sharing_consis 𝒮⇩s⇩b ts⇩s⇩b".
note share_commute=
share_all_until_volatile_write_flush_commute [OF take_empty ‹ownership_distinct ts⇩s⇩b› ‹sharing_consis 𝒮⇩s⇩b ts⇩s⇩b› i_bound ts⇩s⇩b_i dist_sb' dist_R_L_A]
have rel_commute_empty:
"release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') (dom 𝒮 ∪ R - L) Map.empty =
release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') (dom 𝒮⇩s⇩b ∪ R - L) Map.empty"
proof -
{
fix a
assume a_in: "a ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"
have "(a ∈ (dom 𝒮 ∪ R - L)) = (a ∈ (dom 𝒮⇩s⇩b ∪ R - L))"
proof -
from all_shared_acquired_or_owned [OF sharing_consis [OF i_bound ts⇩s⇩b_i]] a_in
all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')" "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"]
have "a ∈ 𝒪⇩s⇩b ∪ all_acquired sb "
by (auto simp add: sb Write⇩s⇩b volatile)
from share_all_until_volatile_write_thread_local [OF ‹ownership_distinct ts⇩s⇩b› ‹sharing_consis 𝒮⇩s⇩b ts⇩s⇩b› i_bound ts⇩s⇩b_i this]
have "𝒮 a = 𝒮⇩s⇩b a"
by (auto simp add: sb Write⇩s⇩b volatile 𝒮)
then show ?thesis
by (auto simp add: domIff)
qed
}
then show ?thesis
apply -
apply (rule release_all_shared_exchange)
apply auto
done
qed
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume j_bound: "j < length ts⇩s⇩b"
assume neq: "i ≠ j"
have "release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)
(dom 𝒮⇩s⇩b ∪ R - L) ℛ⇩j
= release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)
(dom 𝒮⇩s⇩b) ℛ⇩j"
proof -
{
fix a
assume a_in: "a ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
have "(a ∈ (dom 𝒮⇩s⇩b ∪ R - L)) = (a ∈ dom 𝒮⇩s⇩b)"
proof -
from ownership_distinct [OF i_bound j_bound neq ts⇩s⇩b_i jth]
have A_dist: "A ∩ (𝒪⇩j ∪ all_acquired sb⇩j) = {}"
by (auto simp add: sb Write⇩s⇩b volatile)
from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]] a_in
all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
have a_in: "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
with ownership_distinct [OF i_bound j_bound neq ts⇩s⇩b_i jth]
have "a ∉ (𝒪⇩s⇩b ∪ all_acquired sb)" by auto
with A_dist R_owned A_R A_shared_owned L_subset a_in
obtain "a ∉ R" and "a ∉ L"
by fastforce
then show ?thesis by auto
qed
}
then
show ?thesis
apply -
apply (rule release_all_shared_exchange)
apply auto
done
qed
}
note release_commute = this
have "(ts⇩s⇩b [i := (p⇩s⇩b,is⇩s⇩b, θ⇩s⇩b, sb', 𝒟⇩s⇩b, 𝒪⇩s⇩b ∪ A - R,Map.empty)],m⇩s⇩b(a:=v),𝒮⇩s⇩b') ∼
(ts[i := (last_prog (hd_prog p⇩s⇩b (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb'))
(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb'),
is'',
θ⇩s⇩b |` (dom θ⇩s⇩b -
read_tmps (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')),
(),True, acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')
(acquired True ?take_sb 𝒪⇩s⇩b ∪ A - R),
release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')
(dom (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)) Map.empty)],
flush (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') (m(a := v)),
share (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') (𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L))"
apply (rule sim_config.intros)
apply (simp add: flush_commute m)
apply (clarsimp simp add: 𝒮⇩s⇩b' 𝒮 share_commute simp del: restrict_restrict)
using leq
apply simp
using i_bound i_bound' ts_sim 𝒟
apply (clarsimp simp add: Let_def nth_list_update is''_sim last_prog_eq sb Write⇩s⇩b volatile 𝒮⇩s⇩b'
rel_commute_empty
split: if_split_asm )
apply (rule conjI)
apply blast
apply clarsimp
apply (frule (2) release_commute)
apply clarsimp
apply fastforce
done
ultimately
show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct'
valid_dd' valid_sops' load_tmps_fresh' enough_flushs'
valid_program_history' valid'
m⇩s⇩b' 𝒮⇩s⇩b' ts⇩s⇩b'
by (auto simp del: fun_upd_apply simp add: 𝒪⇩s⇩b' ℛ⇩s⇩b' )
next
case False
note non_vol = this
from flush Write⇩s⇩b False
obtain
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b" and
ℛ⇩s⇩b': "ℛ⇩s⇩b' = ℛ⇩s⇩b"
by cases (auto simp add: sb)
from non_volatile_owned non_vol have a_owned: "a ∈ 𝒪⇩s⇩b"
by simp
{
fix j
fix p⇩j is⇩s⇩b⇩j 𝒪⇩j 𝒟⇩s⇩b⇩j θ⇩j ℛ⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b"
assume ts⇩s⇩b_j: "ts⇩s⇩b!j=(p⇩j,is⇩s⇩b⇩j,θ⇩j,sb⇩j,𝒟⇩s⇩b⇩j,𝒪⇩j,ℛ⇩j)"
assume neq_i_j: "i≠j"
have "a ∉ unforwarded_non_volatile_reads (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) {}"
proof
let ?take_sb⇩j = "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
let ?drop_sb⇩j = "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
assume a_in: "a ∈ unforwarded_non_volatile_reads ?drop_sb⇩j {}"
from a_unowned_by_others [rule_format, OF j_bound neq_i_j] ts⇩s⇩b_j
obtain a_unowned: "a ∉ 𝒪⇩j" and a_unacq: "a ∉ all_acquired sb⇩j"
by auto
with all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j] acquired_takeWhile_non_volatile_Write⇩s⇩b [of sb⇩j 𝒪⇩j]
have a_unacq_take: "a ∉ acquired True ?take_sb⇩j 𝒪⇩j"
by (auto )
note nvo_j = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound ts⇩s⇩b_j]
from non_volatile_owned_or_read_only_drop [OF nvo_j]
have nvo_drop_j: "non_volatile_owned_or_read_only True (share ?take_sb⇩j 𝒮⇩s⇩b)
(acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j" .
from in_unforwarded_non_volatile_reads_non_volatile_Read⇩s⇩b [OF a_in]
have a_in': "a ∈ outstanding_refs is_non_volatile_Read⇩s⇩b ?drop_sb⇩j".
from non_volatile_owned_or_read_only_outstanding_refs [OF nvo_drop_j] a_in'
have "a ∈ acquired True ?take_sb⇩j 𝒪⇩j ∪ all_acquired ?drop_sb⇩j ∪
read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
by (auto simp add: misc_outstanding_refs_convs)
moreover
from acquired_append [of True ?take_sb⇩j ?drop_sb⇩j 𝒪⇩j] acquired_all_acquired [of True ?take_sb⇩j 𝒪⇩j]
all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have "acquired True ?take_sb⇩j 𝒪⇩j ∪ all_acquired ?drop_sb⇩j ⊆ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
ultimately
have "a ∈ read_only_reads (acquired True ?take_sb⇩j 𝒪⇩j) ?drop_sb⇩j"
using a_owned ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i ts⇩s⇩b_j]
by auto
with read_only_reads_unowned [OF j_bound i_bound neq_i_j [symmetric] ts⇩s⇩b_j ts⇩s⇩b_i] a_owned
show False
by auto
qed
} note a_notin_unforwarded_non_volatile_reads_drop = this
have valid_reads': "valid_reads m⇩s⇩b' ts⇩s⇩b'"
proof (unfold_locales)
fix j p⇩j "is⇩j" 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume ts_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "reads_consistent False 𝒪⇩j m⇩s⇩b' sb⇩j"
proof (cases "i=j")
case True
from reads_consis ts_j j_bound sb show ?thesis
by (clarsimp simp add: True m⇩s⇩b' Write⇩s⇩b ts⇩s⇩b' 𝒪⇩s⇩b' False reads_consistent_pending_write_antimono)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
moreover from ts_j False have ts_j': "ts⇩s⇩b ! j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
using j_bound by (simp add: ts⇩s⇩b')
ultimately have consis_m: "reads_consistent False 𝒪⇩j m⇩s⇩b sb⇩j"
by (rule valid_reads)
from a_unowned_by_others [rule_format, OF j_bound' False] ts_j'
have a_unowned:"a ∉ 𝒪⇩j ∪ all_acquired sb⇩j"
by simp
let ?take_sb⇩j = "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
let ?drop_sb⇩j = "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j"
from a_unowned acquired_reads_all_acquired [of True ?take_sb⇩j 𝒪⇩j]
all_acquired_append [of ?take_sb⇩j ?drop_sb⇩j]
have a_not_acq_reads: "a ∉ acquired_reads True ?take_sb⇩j 𝒪⇩j"
by auto
moreover
note a_unfw= a_notin_unforwarded_non_volatile_reads_drop [OF j_bound' ts_j' False]
ultimately
show ?thesis
using reads_consistent_mem_eq_on_unforwarded_non_volatile_reads_drop [where W="{}" and
A="unforwarded_non_volatile_reads ?drop_sb⇩j {} ∪ acquired_reads True ?take_sb⇩j 𝒪⇩j" and
m'= "(m⇩s⇩b(a:=v))", OF _ _ _ consis_m]
by (fastforce simp add: m⇩s⇩b')
qed
qed
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof -
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i] sb
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b sb'"
by (auto simp add: Write⇩s⇩b False)
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' Write⇩s⇩b False 𝒪⇩s⇩b' 𝒮⇩s⇩b')
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof -
from sb
have out: "outstanding_refs is_volatile_Write⇩s⇩b sb' ⊆ outstanding_refs is_volatile_Write⇩s⇩b sb"
by (auto simp add: Write⇩s⇩b False)
have acq: "all_acquired sb' ⊆ all_acquired sb"
by (auto simp add: Write⇩s⇩b False sb)
from outstanding_volatile_writes_unowned_by_others_store_buffer
[OF i_bound ts⇩s⇩b_i out acq]
show ?thesis by (simp add: ts⇩s⇩b' Write⇩s⇩b False 𝒪⇩s⇩b')
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof -
have ro: "read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')
⊆ read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
by (auto simp add: sb Write⇩s⇩b non_vol)
have "𝒪⇩s⇩b ∪ all_acquired sb' ⊆ 𝒪⇩s⇩b ∪ all_acquired sb"
by (auto simp add: sb Write⇩s⇩b non_vol)
from read_only_reads_unowned_nth_update [OF i_bound ts⇩s⇩b_i ro this]
show ?thesis
by (simp add: ts⇩s⇩b' sb 𝒪⇩s⇩b')
qed
next
show "ownership_distinct ts⇩s⇩b'"
proof -
have acq: "all_acquired sb' ⊆ all_acquired sb"
by (auto simp add: Write⇩s⇩b False sb)
with ownership_distinct_instructions_read_value_store_buffer_independent
[OF i_bound ts⇩s⇩b_i]
show ?thesis by (simp add: ts⇩s⇩b' Write⇩s⇩b False 𝒪⇩s⇩b')
qed
qed
have valid_sharing': "valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
from outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i]
have "non_volatile_writes_unshared 𝒮⇩s⇩b sb'"
by (auto simp add: sb Write⇩s⇩b False)
from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
show "outstanding_non_volatile_writes_unshared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒮⇩s⇩b')
next
from sharing_consis [OF i_bound ts⇩s⇩b_i]
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b sb'"
by (auto simp add: sb Write⇩s⇩b False)
from sharing_consis_nth_update [OF i_bound this]
show "sharing_consis 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound ts⇩s⇩b_i] ]
show "read_only_unowned 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' 𝒪⇩s⇩b')
next
from unowned_shared_nth_update [OF i_bound ts⇩s⇩b_i subset_refl]
show "unowned_shared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
have "no_write_to_read_only_memory 𝒮⇩s⇩b sb'"
by (auto simp add: sb Write⇩s⇩b False)
from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
show "no_outstanding_write_to_read_only_memory 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' sb)
qed
from is_sim
obtain is_sim: "instrs (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') @ is⇩s⇩b =
is @ prog_instrs (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"
by (simp add: suspends sb Write⇩s⇩b False)
have "(ts,m,𝒮) ⇒⇩d⇧* (ts,m,𝒮)" by blast
moreover
note flush_commute =
flush_all_until_volatile_write_Write⇩s⇩b_non_volatile_commute [OF i_bound ts⇩s⇩b_i [simplified sb Write⇩s⇩b non_vol]
outstanding_refs_is_Write⇩s⇩b_takeWhile_disj a_notin_others']
note share_commute =
share_all_until_volatile_write_update_sb [of sb' sb, OF _ i_bound ts⇩s⇩b_i, simplified sb Write⇩s⇩b False, simplified]
have "(ts⇩s⇩b [i := (p⇩s⇩b,is⇩s⇩b,θ⇩s⇩b, sb', 𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ⇩s⇩b)], m⇩s⇩b(a:=v),𝒮⇩s⇩b') ∼
(ts,m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute)
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' share_commute)
using leq
apply simp
using i_bound i_bound' is_sim ts_i ts_sim 𝒟
apply (clarsimp simp add: Let_def nth_list_update suspends sb Write⇩s⇩b False 𝒮⇩s⇩b'
split: if_split_asm )
done
ultimately
show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' m⇩s⇩b'
valid_dd' valid_sops' load_tmps_fresh' enough_flushs' valid_program_history' valid'
ts⇩s⇩b' 𝒪⇩s⇩b' 𝒮⇩s⇩b' ℛ⇩s⇩b'
by (auto simp del: fun_upd_apply)
qed
next
case (Read⇩s⇩b volatile a t v)
from flush this obtain m⇩s⇩b': "m⇩s⇩b'=m⇩s⇩b" and
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b" and 𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b" and
ℛ⇩s⇩b': "ℛ⇩s⇩b'=ℛ⇩s⇩b"
by cases (auto simp add: sb)
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof -
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i] sb
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b sb'"
by (auto simp add: Read⇩s⇩b)
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' Read⇩s⇩b 𝒪⇩s⇩b' 𝒮⇩s⇩b')
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof -
from sb
have out: "outstanding_refs is_volatile_Write⇩s⇩b sb' ⊆ outstanding_refs is_volatile_Write⇩s⇩b sb"
by (auto simp add: Read⇩s⇩b)
have acq: "all_acquired sb' ⊆ all_acquired sb"
by (auto simp add: Read⇩s⇩b sb)
from outstanding_volatile_writes_unowned_by_others_store_buffer
[OF i_bound ts⇩s⇩b_i out acq]
show ?thesis by (simp add: ts⇩s⇩b' Read⇩s⇩b 𝒪⇩s⇩b')
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof -
have ro: "read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')
⊆ read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
by (auto simp add: sb Read⇩s⇩b)
have "𝒪⇩s⇩b ∪ all_acquired sb' ⊆ 𝒪⇩s⇩b ∪ all_acquired sb"
by (auto simp add: sb Read⇩s⇩b)
from read_only_reads_unowned_nth_update [OF i_bound ts⇩s⇩b_i ro this]
show ?thesis
by (simp add: ts⇩s⇩b' sb 𝒪⇩s⇩b')
qed
next
show "ownership_distinct ts⇩s⇩b'"
proof -
have acq: "all_acquired sb' ⊆ all_acquired sb"
by (auto simp add: Read⇩s⇩b sb)
with ownership_distinct_instructions_read_value_store_buffer_independent
[OF i_bound ts⇩s⇩b_i]
show ?thesis by (simp add: ts⇩s⇩b' Read⇩s⇩b 𝒪⇩s⇩b')
qed
qed
have valid_sharing': "valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
from outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i]
have "non_volatile_writes_unshared 𝒮⇩s⇩b sb'"
by (auto simp add: sb Read⇩s⇩b)
from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
show "outstanding_non_volatile_writes_unshared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒮⇩s⇩b')
next
from sharing_consis [OF i_bound ts⇩s⇩b_i]
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b sb'"
by (auto simp add: sb Read⇩s⇩b)
from sharing_consis_nth_update [OF i_bound this]
show "sharing_consis 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound ts⇩s⇩b_i] ]
show "read_only_unowned 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' 𝒪⇩s⇩b')
next
from unowned_shared_nth_update [OF i_bound ts⇩s⇩b_i subset_refl]
show "unowned_shared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
have "no_write_to_read_only_memory 𝒮⇩s⇩b sb'"
by (auto simp add: sb Read⇩s⇩b)
from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
show "no_outstanding_write_to_read_only_memory 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' sb)
qed
have valid_reads': "valid_reads m⇩s⇩b' ts⇩s⇩b'"
proof -
from valid_reads [OF i_bound ts⇩s⇩b_i]
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b sb'"
by (simp add: sb Read⇩s⇩b)
from valid_reads_nth_update [OF i_bound this]
show ?thesis by (simp add: m⇩s⇩b' ts⇩s⇩b' 𝒪⇩s⇩b')
qed
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
from valid_program_history [OF i_bound ts⇩s⇩b_i]
have "causal_program_history is⇩s⇩b sb" .
then have causal': "causal_program_history is⇩s⇩b sb'"
by (simp add: sb Read⇩s⇩b causal_program_history_def)
from valid_last_prog [OF i_bound ts⇩s⇩b_i]
have "last_prog p⇩s⇩b sb = p⇩s⇩b".
hence "last_prog p⇩s⇩b sb' = p⇩s⇩b"
by (simp add: sb Read⇩s⇩b)
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
from is_sim
have is_sim: "instrs (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') @ is⇩s⇩b =
is @ prog_instrs (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"
by (simp add: sb Read⇩s⇩b suspends)
from valid_history [OF i_bound ts⇩s⇩b_i]
have θ⇩s⇩b_v: "θ⇩s⇩b t = Some v"
by (simp add: history_consistent_access_last_read sb Read⇩s⇩b split:option.splits)
have "(ts,m,𝒮) ⇒⇩d⇧* (ts,m,𝒮)" by blast
moreover
note flush_commute= flush_all_until_volatile_write_Read⇩s⇩b_commute [OF i_bound ts⇩s⇩b_i [simplified sb Read⇩s⇩b]]
note share_commute =
share_all_until_volatile_write_update_sb [of sb' sb, OF _ i_bound ts⇩s⇩b_i, simplified sb Read⇩s⇩b, simplified]
have "(ts⇩s⇩b [i := (p⇩s⇩b,is⇩s⇩b, θ⇩s⇩b, sb',𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ⇩s⇩b')],m⇩s⇩b,𝒮⇩s⇩b') ∼ (ts,m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute)
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' share_commute)
using leq
apply simp
using i_bound i_bound' ts_sim ts_i is_sim 𝒟
apply (clarsimp simp add: Let_def nth_list_update sb suspends Read⇩s⇩b 𝒮⇩s⇩b' ℛ⇩s⇩b'
split: if_split_asm)
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' m⇩s⇩b'
valid_dd' valid_sops' load_tmps_fresh' enough_flushs' valid_sharing'
valid_program_history' valid'
ts⇩s⇩b' 𝒪⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply)
next
case (Prog⇩s⇩b p⇩1 p⇩2 mis)
from flush this obtain m⇩s⇩b': "m⇩s⇩b'=m⇩s⇩b" and
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b" and 𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b" and
ℛ⇩s⇩b': "ℛ⇩s⇩b'=ℛ⇩s⇩b"
by cases (auto simp add: sb)
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof -
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i] sb
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b sb'"
by (auto simp add: Prog⇩s⇩b)
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' Prog⇩s⇩b 𝒪⇩s⇩b' 𝒮⇩s⇩b')
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof -
from sb
have out: "outstanding_refs is_volatile_Write⇩s⇩b sb' ⊆ outstanding_refs is_volatile_Write⇩s⇩b sb"
by (auto simp add: Prog⇩s⇩b)
have acq: "all_acquired sb' ⊆ all_acquired sb"
by (auto simp add: Prog⇩s⇩b sb)
from outstanding_volatile_writes_unowned_by_others_store_buffer
[OF i_bound ts⇩s⇩b_i out acq]
show ?thesis by (simp add: ts⇩s⇩b' Prog⇩s⇩b 𝒪⇩s⇩b')
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof -
have ro: "read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')
⊆ read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
by (auto simp add: sb Prog⇩s⇩b)
have "𝒪⇩s⇩b ∪ all_acquired sb' ⊆ 𝒪⇩s⇩b ∪ all_acquired sb"
by (auto simp add: sb Prog⇩s⇩b)
from read_only_reads_unowned_nth_update [OF i_bound ts⇩s⇩b_i ro this]
show ?thesis
by (simp add: ts⇩s⇩b' sb 𝒪⇩s⇩b')
qed
next
show "ownership_distinct ts⇩s⇩b'"
proof -
have acq: "all_acquired sb' ⊆ all_acquired sb"
by (auto simp add: Prog⇩s⇩b sb)
with ownership_distinct_instructions_read_value_store_buffer_independent
[OF i_bound ts⇩s⇩b_i]
show ?thesis by (simp add: ts⇩s⇩b' Prog⇩s⇩b 𝒪⇩s⇩b')
qed
qed
have valid_sharing': "valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
from outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i]
have "non_volatile_writes_unshared 𝒮⇩s⇩b sb'"
by (auto simp add: sb Prog⇩s⇩b)
from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
show "outstanding_non_volatile_writes_unshared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒮⇩s⇩b')
next
from sharing_consis [OF i_bound ts⇩s⇩b_i]
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b sb'"
by (auto simp add: sb Prog⇩s⇩b)
from sharing_consis_nth_update [OF i_bound this]
show "sharing_consis 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound ts⇩s⇩b_i] ]
show "read_only_unowned 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' 𝒪⇩s⇩b')
next
from unowned_shared_nth_update [OF i_bound ts⇩s⇩b_i subset_refl]
show "unowned_shared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒪⇩s⇩b' 𝒮⇩s⇩b')
next
from no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
have "no_write_to_read_only_memory 𝒮⇩s⇩b sb'"
by (auto simp add: sb Prog⇩s⇩b)
from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
show "no_outstanding_write_to_read_only_memory 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b' sb)
qed
have valid_reads': "valid_reads m⇩s⇩b' ts⇩s⇩b'"
proof -
from valid_reads [OF i_bound ts⇩s⇩b_i]
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b sb'"
by (simp add: sb Prog⇩s⇩b)
from valid_reads_nth_update [OF i_bound this]
show ?thesis by (simp add: m⇩s⇩b' ts⇩s⇩b' 𝒪⇩s⇩b')
qed
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
from valid_program_history [OF i_bound ts⇩s⇩b_i]
have "causal_program_history is⇩s⇩b sb" .
then have causal': "causal_program_history is⇩s⇩b sb'"
by (simp add: sb Prog⇩s⇩b causal_program_history_def)
from valid_last_prog [OF i_bound ts⇩s⇩b_i]
have "last_prog p⇩s⇩b sb = p⇩s⇩b".
hence "last_prog p⇩2 sb' = p⇩s⇩b"
by (simp add: sb Prog⇩s⇩b)
from last_prog_to_last_prog_same [OF this]
have "last_prog p⇩s⇩b sb' = p⇩s⇩b".
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
from is_sim
have is_sim: "instrs (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') @ is⇩s⇩b =
is @ prog_instrs (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"
by (simp add: suspends sb Prog⇩s⇩b)
have "(ts,m,𝒮) ⇒⇩d⇧* (ts,m,𝒮)" by blast
moreover
note flush_commute = flush_all_until_volatile_write_Prog⇩s⇩b_commute [OF i_bound
ts⇩s⇩b_i [simplified sb Prog⇩s⇩b]]
note share_commute =
share_all_until_volatile_write_update_sb [of sb' sb, OF _ i_bound ts⇩s⇩b_i, simplified sb Prog⇩s⇩b, simplified]
have "(ts⇩s⇩b [i := (p⇩s⇩b,is⇩s⇩b, θ⇩s⇩b, sb',𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ⇩s⇩b)],m⇩s⇩b,𝒮⇩s⇩b') ∼ (ts,m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute)
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' share_commute)
using leq
apply simp
using i_bound i_bound' ts_sim ts_i is_sim 𝒟
apply (clarsimp simp add: Let_def nth_list_update sb suspends Prog⇩s⇩b ℛ⇩s⇩b' 𝒮⇩s⇩b'
split: if_split_asm)
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' m⇩s⇩b'
valid_dd' valid_sops' load_tmps_fresh' enough_flushs' valid_sharing'
valid_program_history' valid'
ts⇩s⇩b' 𝒮⇩s⇩b' 𝒪⇩s⇩b' ℛ⇩s⇩b' 𝒮⇩s⇩b'
by (auto simp del: fun_upd_apply)
next
case (Ghost⇩s⇩b A L R W)
from flush Ghost⇩s⇩b
obtain
𝒪⇩s⇩b': "𝒪⇩s⇩b'=𝒪⇩s⇩b ∪ A - R" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L" and
ℛ⇩s⇩b': "ℛ⇩s⇩b'= augment_rels (dom 𝒮⇩s⇩b) R ℛ⇩s⇩b" and
m⇩s⇩b': "m⇩s⇩b'=m⇩s⇩b"
by cases (auto simp add: sb)
from sharing_consis [OF i_bound ts⇩s⇩b_i]
obtain
A_shared_owned: "A ⊆ dom 𝒮⇩s⇩b ∪ 𝒪⇩s⇩b" and
L_subset: "L ⊆ A" and
A_R: "A ∩ R = {}" and
R_owned: "R ⊆ 𝒪⇩s⇩b"
by (clarsimp simp add: sb Ghost⇩s⇩b)
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof
fix j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j acq⇩j θ⇩j sb⇩j p⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume ts⇩s⇩b'_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "non_volatile_owned_or_read_only False 𝒮⇩s⇩b' 𝒪⇩j sb⇩j"
proof (cases "j=i")
case True
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]
have "non_volatile_owned_or_read_only False (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) (𝒪⇩s⇩b ∪ A - R) sb'"
by (auto simp add: sb Ghost⇩s⇩b non_volatile_owned_or_read_only_pending_write_antimono)
then show ?thesis
using True i_bound ts⇩s⇩b'_j
by (auto simp add: ts⇩s⇩b' 𝒮⇩s⇩b' sb 𝒪⇩s⇩b')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
with ts⇩s⇩b'_j False i_bound
have ts⇩s⇩b_j: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
note nvo = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' ts⇩s⇩b_j]
from read_only_unowned [OF i_bound ts⇩s⇩b_i] R_owned
have "R ∩ read_only 𝒮⇩s⇩b = {}"
by auto
with read_only_reads_unowned [OF j_bound' i_bound False ts⇩s⇩b_j ts⇩s⇩b_i] L_subset
have "∀a∈read_only_reads
(acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) 𝒪⇩j)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j).
a ∈ read_only 𝒮⇩s⇩b ⟶ a ∈ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L )"
by (auto simp add: in_read_only_convs sb Ghost⇩s⇩b)
from non_volatile_owned_or_read_only_read_only_reads_eq' [OF nvo this]
have "non_volatile_owned_or_read_only False (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) 𝒪⇩j sb⇩j".
thus ?thesis by (simp add: 𝒮⇩s⇩b')
qed
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof (unfold_locales)
fix i⇩1 j p⇩1 "is⇩1" 𝒪⇩1 ℛ⇩1 𝒟⇩1 xs⇩1 sb⇩1 p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume i⇩1_bound: "i⇩1 < length ts⇩s⇩b'"
assume j_bound: "j < length ts⇩s⇩b'"
assume i⇩1_j: "i⇩1 ≠ j"
assume ts_i⇩1: "ts⇩s⇩b'!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
assume ts_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}"
proof (cases "i⇩1=i")
case True
from i⇩1_j True have neq_i_j: "i≠j"
by auto
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_j neq_i_j have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from outstanding_volatile_writes_unowned_by_others [OF i_bound j_bound' neq_i_j
ts⇩s⇩b_i ts_j'] ts_i⇩1 i_bound ts⇩s⇩b_i True show ?thesis
by (clarsimp simp add: ts⇩s⇩b' sb Ghost⇩s⇩b)
next
case False
note i⇩1_i = this
from i⇩1_bound have i⇩1_bound': "i⇩1 < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b' sb)
hence i⇩1_bound'': "i⇩1 < length (map owned ts⇩s⇩b)"
by auto
from ts_i⇩1 False have ts_i⇩1': "ts⇩s⇩b!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
by (simp add: ts⇩s⇩b' sb)
show ?thesis
proof (cases "j=i")
case True
from outstanding_volatile_writes_unowned_by_others [OF i⇩1_bound' i_bound i⇩1_i ts_i⇩1' ts⇩s⇩b_i ]
have "(𝒪⇩s⇩b ∪ all_acquired sb) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}".
then show ?thesis
using True i⇩1_i ts_j ts⇩s⇩b_i i_bound
by (auto simp add: sb Ghost⇩s⇩b ts⇩s⇩b' 𝒪⇩s⇩b')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_j False have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from outstanding_volatile_writes_unowned_by_others
[OF i⇩1_bound' j_bound' i⇩1_j ts_i⇩1' ts_j']
show "(𝒪⇩j ∪ all_acquired sb⇩j) ∩ outstanding_refs is_volatile_Write⇩s⇩b sb⇩1 = {}" .
qed
qed
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof
fix n m
fix p⇩n "is⇩n" 𝒪⇩n ℛ⇩n 𝒟⇩n θ⇩n sb⇩n p⇩m "is⇩m" 𝒪⇩m ℛ⇩m 𝒟⇩m θ⇩m sb⇩m
assume n_bound: "n < length ts⇩s⇩b'"
and m_bound: "m < length ts⇩s⇩b'"
and neq_n_m: "n≠m"
and nth: "ts⇩s⇩b'!n = (p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
and mth: "ts⇩s⇩b'!m =(p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
from n_bound have n_bound': "n < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
from m_bound have m_bound': "m < length ts⇩s⇩b" by (simp add: ts⇩s⇩b')
show "(𝒪⇩m ∪ all_acquired sb⇩m) ∩
read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) 𝒪⇩n)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) =
{}"
proof (cases "m=i")
case True
with neq_n_m have neq_n_i: "n≠i"
by auto
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
note read_only_reads_unowned [OF n_bound' i_bound neq_n_i nth' ts⇩s⇩b_i]
then
show ?thesis
using True ts⇩s⇩b_i neq_n_i nth mth n_bound' m_bound' L_subset
by (auto simp add: ts⇩s⇩b' 𝒪⇩s⇩b' sb Ghost⇩s⇩b)
next
case False
note neq_m_i = this
with m_bound mth i_bound have mth': "ts⇩s⇩b!m = (p⇩m, is⇩m, θ⇩m, sb⇩m, 𝒟⇩m, 𝒪⇩m,ℛ⇩m)"
by (auto simp add: ts⇩s⇩b')
show ?thesis
proof (cases "n=i")
case True
from read_only_reads_append [of "(𝒪⇩s⇩b ∪ A - R)" "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n)"]
have "read_only_reads
(acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) (𝒪⇩s⇩b ∪ A - R))
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩n) ⊆ read_only_reads (𝒪⇩s⇩b ∪ A - R) sb⇩n"
by auto
with ts⇩s⇩b_i nth mth neq_m_i n_bound' True
read_only_reads_unowned [OF i_bound m_bound' False [symmetric] ts⇩s⇩b_i mth']
show ?thesis
by (auto simp add: ts⇩s⇩b' sb 𝒪⇩s⇩b' Ghost⇩s⇩b)
next
case False
with n_bound nth i_bound have nth': "ts⇩s⇩b!n =(p⇩n, is⇩n, θ⇩n, sb⇩n, 𝒟⇩n, 𝒪⇩n,ℛ⇩n)"
by (auto simp add: ts⇩s⇩b')
from read_only_reads_unowned [OF n_bound' m_bound' neq_n_m nth' mth'] False neq_m_i
show ?thesis
by (clarsimp)
qed
qed
qed
next
show "ownership_distinct ts⇩s⇩b'"
proof (unfold_locales)
fix i⇩1 j p⇩1 "is⇩1" 𝒪⇩1 ℛ⇩1 𝒟⇩1 xs⇩1 sb⇩1 p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume i⇩1_bound: "i⇩1 < length ts⇩s⇩b'"
assume j_bound: "j < length ts⇩s⇩b'"
assume i⇩1_j: "i⇩1 ≠ j"
assume ts_i⇩1: "ts⇩s⇩b'!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
assume ts_j: "ts⇩s⇩b'!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "(𝒪⇩1 ∪ all_acquired sb⇩1) ∩ (𝒪⇩j ∪ all_acquired sb⇩j)= {}"
proof (cases "i⇩1=i")
case True
with i⇩1_j have i_j: "i≠j"
by simp
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
hence j_bound'': "j < length (map owned ts⇩s⇩b)"
by simp
from ts_j i_j have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from ownership_distinct [OF i_bound j_bound' i_j ts⇩s⇩b_i ts_j']
show ?thesis
using ts⇩s⇩b_i True ts_i⇩1 i_bound 𝒪⇩s⇩b'
by (auto simp add: ts⇩s⇩b' sb Ghost⇩s⇩b)
next
case False
note i⇩1_i = this
from i⇩1_bound have i⇩1_bound': "i⇩1 < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
hence i⇩1_bound'': "i⇩1 < length (map owned ts⇩s⇩b)"
by simp
from ts_i⇩1 False have ts_i⇩1': "ts⇩s⇩b!i⇩1 = (p⇩1,is⇩1,xs⇩1,sb⇩1,𝒟⇩1,𝒪⇩1,ℛ⇩1)"
by (simp add: ts⇩s⇩b')
show ?thesis
proof (cases "j=i")
case True
from ownership_distinct [OF i⇩1_bound' i_bound i⇩1_i ts_i⇩1' ts⇩s⇩b_i]
show ?thesis
using ts⇩s⇩b_i True ts_j i_bound 𝒪⇩s⇩b'
by (auto simp add: ts⇩s⇩b' sb Ghost⇩s⇩b)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (simp add: ts⇩s⇩b')
from ts_j False have ts_j': "ts⇩s⇩b!j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (simp add: ts⇩s⇩b')
from ownership_distinct [OF i⇩1_bound' j_bound' i⇩1_j ts_i⇩1' ts_j']
show ?thesis .
qed
qed
qed
qed
have valid_sharing': "valid_sharing (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_writes_unshared (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (unfold_locales)
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j acq⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "non_volatile_writes_unshared (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) sb⇩j"
proof (cases "i=j")
case True
with outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i]
i_bound jth ts⇩s⇩b_i show ?thesis
by (clarsimp simp add: ts⇩s⇩b' sb Ghost⇩s⇩b)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
from jth False have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from j_bound jth i_bound False
have j: "non_volatile_writes_unshared 𝒮⇩s⇩b sb⇩j"
apply -
apply (rule outstanding_non_volatile_writes_unshared)
apply (auto simp add: ts⇩s⇩b')
done
from jth False have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from outstanding_non_volatile_writes_unshared [OF j_bound' jth']
have unshared: "non_volatile_writes_unshared 𝒮⇩s⇩b sb⇩j".
have "∀a∈dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) - dom 𝒮⇩s⇩b. a ∉ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
proof -
{
fix a
assume a_in: "a ∈ dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) - dom 𝒮⇩s⇩b"
hence a_R: "a ∈ R"
by clarsimp
assume a_in_j: "a ∈ outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j"
have False
proof -
from non_volatile_owned_or_read_only_outstanding_non_volatile_writes [OF
outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' jth']]
a_in_j
have "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
moreover
with ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth'] a_R R_owned
show False
by blast
qed
}
thus ?thesis by blast
qed
from non_volatile_writes_unshared_no_outstanding_non_volatile_Write⇩s⇩b
[OF unshared this]
show ?thesis .
qed
qed
next
show "sharing_consis (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (unfold_locales)
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j acq⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "sharing_consistent (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) 𝒪⇩j sb⇩j"
proof (cases "i=j")
case True
with i_bound jth ts⇩s⇩b_i sharing_consis [OF i_bound ts⇩s⇩b_i]
show ?thesis
by (clarsimp simp add: ts⇩s⇩b' sb Ghost⇩s⇩b 𝒪⇩s⇩b')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
from jth False have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from sharing_consis [OF j_bound' jth']
have consis: "sharing_consistent 𝒮⇩s⇩b 𝒪⇩j sb⇩j".
have acq_cond: "all_acquired sb⇩j ∩ dom 𝒮⇩s⇩b - dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
proof -
{
fix a
assume a_acq: "a ∈ all_acquired sb⇩j"
assume "a ∈ dom 𝒮⇩s⇩b"
assume a_L: "a ∈ L"
have False
proof -
from ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth']
have "A ∩ all_acquired sb⇩j = {}"
by (auto simp add: sb Ghost⇩s⇩b)
with a_acq a_L L_subset
show False
by blast
qed
}
thus ?thesis
by auto
qed
have uns_cond: "all_unshared sb⇩j ∩ dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) - dom 𝒮⇩s⇩b = {}"
proof -
{
fix a
assume a_uns: "a ∈ all_unshared sb⇩j"
assume "a ∉ L"
assume a_R: "a ∈ R"
have False
proof -
from unshared_acquired_or_owned [OF consis] a_uns
have "a ∈ all_acquired sb⇩j ∪ 𝒪⇩j" by auto
with ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth'] R_owned a_R
show False
by blast
qed
}
thus ?thesis
by auto
qed
from sharing_consistent_preservation [OF consis acq_cond uns_cond]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
qed
next
show "unowned_shared (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof (unfold_locales)
show "- ⋃((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set ts⇩s⇩b') ⊆ dom (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L)"
proof -
have s: "⋃((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set ts⇩s⇩b') =
⋃((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set ts⇩s⇩b) ∪ A - R"
apply (unfold ts⇩s⇩b' 𝒪⇩s⇩b')
apply (rule acquire_release_ownership_nth_update [OF R_owned i_bound ts⇩s⇩b_i])
apply (rule local.ownership_distinct_axioms)
done
note unowned_shared L_subset A_R
then
show ?thesis
apply (simp only: s)
apply auto
done
qed
qed
next
show "read_only_unowned (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "𝒪⇩j ∩ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
proof (cases "i=j")
case True
from read_only_unowned [OF i_bound ts⇩s⇩b_i]
have "(𝒪⇩s⇩b ∪ A - R ) ∩ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) = {}"
by (auto simp add: in_read_only_convs )
with jth ts⇩s⇩b_i i_bound True
show ?thesis
by (auto simp add: 𝒪⇩s⇩b' ts⇩s⇩b')
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
with False jth have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from read_only_unowned [OF j_bound' jth']
have "𝒪⇩j ∩ read_only 𝒮⇩s⇩b = {}".
moreover
from ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth'] R_owned
have "(𝒪⇩s⇩b ∪ A) ∩ 𝒪⇩j = {}"
by (auto simp add: sb Ghost⇩s⇩b)
moreover note R_owned A_R
ultimately show ?thesis
by (fastforce simp add: in_read_only_convs split: if_split_asm)
qed
qed
next
show "no_outstanding_write_to_read_only_memory (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ts⇩s⇩b'"
proof
fix j p⇩j "is⇩j" "𝒪⇩j" ℛ⇩j 𝒟⇩j xs⇩j sb⇩j
assume j_bound: "j < length ts⇩s⇩b'"
assume jth: "ts⇩s⇩b' ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
show "no_write_to_read_only_memory (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) sb⇩j"
proof (cases "i=j")
case True
with jth ts⇩s⇩b_i i_bound no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
show ?thesis
by (auto simp add: sb ts⇩s⇩b' Ghost⇩s⇩b)
next
case False
from j_bound have j_bound': "j < length ts⇩s⇩b"
by (auto simp add: ts⇩s⇩b')
with False jth have jth': "ts⇩s⇩b ! j = (p⇩j,is⇩j,xs⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
by (auto simp add: ts⇩s⇩b')
from no_outstanding_write_to_read_only_memory [OF j_bound' jth']
have nw: "no_write_to_read_only_memory 𝒮⇩s⇩b sb⇩j".
have "R ∩ outstanding_refs is_Write⇩s⇩b sb⇩j = {}"
proof -
note dist = ownership_distinct [OF i_bound j_bound' False ts⇩s⇩b_i jth']
from non_volatile_owned_or_read_only_outstanding_non_volatile_writes
[OF outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' jth']]
dist
have "outstanding_refs is_non_volatile_Write⇩s⇩b sb⇩j ∩ 𝒪⇩s⇩b = {}"
by auto
moreover
from outstanding_volatile_writes_unowned_by_others [OF j_bound' i_bound
False [symmetric] jth' ts⇩s⇩b_i ]
have "outstanding_refs is_volatile_Write⇩s⇩b sb⇩j ∩ 𝒪⇩s⇩b = {}"
by auto
ultimately have "outstanding_refs is_Write⇩s⇩b sb⇩j ∩ 𝒪⇩s⇩b = {}"
by (auto simp add: misc_outstanding_refs_convs)
with R_owned
show ?thesis by blast
qed
then
have "∀a∈outstanding_refs is_Write⇩s⇩b sb⇩j.
a ∈ read_only (𝒮⇩s⇩b ⊕⇘W⇙ R ⊖⇘A⇙ L) ⟶ a ∈ read_only 𝒮⇩s⇩b"
by (auto simp add: in_read_only_convs)
from no_write_to_read_only_memory_read_only_reads_eq [OF nw this]
show ?thesis .
qed
qed
qed
have valid_reads': "valid_reads m⇩s⇩b' ts⇩s⇩b'"
proof -
from valid_reads [OF i_bound ts⇩s⇩b_i]
have "reads_consistent False (𝒪⇩s⇩b ∪ A - R) m⇩s⇩b sb'"
by (simp add: sb Ghost⇩s⇩b)
from valid_reads_nth_update [OF i_bound this]
show ?thesis by (simp add: m⇩s⇩b' ts⇩s⇩b' 𝒪⇩s⇩b')
qed
have valid_program_history': "valid_program_history ts⇩s⇩b'"
proof -
from valid_program_history [OF i_bound ts⇩s⇩b_i]
have "causal_program_history is⇩s⇩b sb" .
then have causal': "causal_program_history is⇩s⇩b sb'"
by (simp add: sb Ghost⇩s⇩b causal_program_history_def)
from valid_last_prog [OF i_bound ts⇩s⇩b_i]
have "last_prog p⇩s⇩b sb = p⇩s⇩b".
hence "last_prog p⇩s⇩b sb' = p⇩s⇩b"
by (simp add: sb Ghost⇩s⇩b)
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
from is_sim
have is_sim: "instrs (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb') @ is⇩s⇩b =
is @ prog_instrs (dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb')"
by (simp add: sb Ghost⇩s⇩b suspends)
have "(ts,m,𝒮) ⇒⇩d⇧* (ts,m,𝒮)" by blast
moreover
note flush_commute =
flush_all_until_volatile_write_Ghost⇩s⇩b_commute [OF i_bound ts⇩s⇩b_i [simplified sb Ghost⇩s⇩b]]
have dist_R_L_A: "∀j p is 𝒪 ℛ 𝒟 θ sb.
j < length ts⇩s⇩b ⟶ i≠ j⟶
ts⇩s⇩b ! j = (p, is, θ, sb, 𝒟, 𝒪, ℛ) ⟶
(all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∪
all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) ∪
all_acquired (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)) ∩ (R ∪ L ∪ A) = {}"
proof -
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume j_bound: "j < length ts⇩s⇩b"
assume neq_i_j: "i ≠ j"
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j, θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume x_shared: "x ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ∪
all_unshared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j) ∪
all_acquired (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
assume x_R_L_A: "x ∈ R ∪ L ∪ A"
have False
proof -
from x_shared all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
unshared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)" "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
all_unshared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)" "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
all_acquired_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)" "(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
have "x ∈ all_acquired sb⇩j ∪ 𝒪⇩j "
by auto
moreover
from x_R_L_A R_owned L_subset
have "x ∈ all_acquired sb ∪ 𝒪⇩s⇩b"
by (auto simp add: sb Ghost⇩s⇩b)
moreover
note ownership_distinct [OF i_bound j_bound neq_i_j ts⇩s⇩b_i jth]
ultimately show False by blast
qed
}
thus ?thesis by blast
qed
{
fix j p⇩j is⇩j 𝒪⇩j ℛ⇩j 𝒟⇩j θ⇩j sb⇩j x
assume jth: "ts⇩s⇩b!j = (p⇩j,is⇩j,θ⇩j,sb⇩j,𝒟⇩j,𝒪⇩j,ℛ⇩j)"
assume j_bound: "j < length ts⇩s⇩b"
assume neq: "i ≠ j"
have "release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)
(dom 𝒮⇩s⇩b ∪ R - L) ℛ⇩j
= release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)
(dom 𝒮⇩s⇩b) ℛ⇩j"
proof -
{
fix a
assume a_in: "a ∈ all_shared (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
have "(a ∈ (dom 𝒮⇩s⇩b ∪ R - L)) = (a ∈ dom 𝒮⇩s⇩b)"
proof -
from ownership_distinct [OF i_bound j_bound neq ts⇩s⇩b_i jth]
have A_dist: "A ∩ (𝒪⇩j ∪ all_acquired sb⇩j) = {}"
by (auto simp add: sb Ghost⇩s⇩b)
from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]] a_in
all_shared_append [of "(takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb⇩j)"]
have a_in: "a ∈ 𝒪⇩j ∪ all_acquired sb⇩j"
by auto
with ownership_distinct [OF i_bound j_bound neq ts⇩s⇩b_i jth]
have "a ∉ (𝒪⇩s⇩b ∪ all_acquired sb)" by auto
with A_dist R_owned A_R A_shared_owned L_subset a_in
obtain "a ∉ R" and "a ∉ L"
by fastforce
then show ?thesis by auto
qed
}
then
show ?thesis
apply -
apply (rule release_all_shared_exchange)
apply auto
done
qed
}
note release_commute = this
from ownership_distinct_axioms have "ownership_distinct ts⇩s⇩b".
from sharing_consis_axioms have "sharing_consis 𝒮⇩s⇩b ts⇩s⇩b".
note share_commute = share_all_until_volatile_write_Ghost⇩s⇩b_commute [OF ‹ownership_distinct ts⇩s⇩b›
‹sharing_consis 𝒮⇩s⇩b ts⇩s⇩b› i_bound ts⇩s⇩b_i [simplified sb Ghost⇩s⇩b] dist_R_L_A]
have "(ts⇩s⇩b [i := (p⇩s⇩b,is⇩s⇩b, θ⇩s⇩b, sb', 𝒟⇩s⇩b, 𝒪⇩s⇩b ∪ A - R,augment_rels (dom 𝒮⇩s⇩b) R ℛ⇩s⇩b)],m⇩s⇩b,𝒮⇩s⇩b') ∼ (ts,m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute)
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' share_commute)
using leq
apply simp
using i_bound i_bound' ts_sim ts_i is_sim 𝒟
apply (clarsimp simp add: Let_def nth_list_update sb suspends Ghost⇩s⇩b ℛ⇩s⇩b' 𝒮⇩s⇩b'
split: if_split_asm)
apply (rule conjI)
apply fastforce
apply clarsimp
apply (frule (2) release_commute)
apply clarsimp
apply auto
done
ultimately
show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct'
valid_dd' valid_sops' load_tmps_fresh' enough_flushs'
valid_program_history' valid'
m⇩s⇩b' 𝒮⇩s⇩b' ts⇩s⇩b'
by (auto simp del: fun_upd_apply simp add: 𝒪⇩s⇩b' ℛ⇩s⇩b')
qed
next
case (Program i p⇩s⇩b "is⇩s⇩b" θ⇩s⇩b sb 𝒟⇩s⇩b 𝒪⇩s⇩b ℛ⇩s⇩b p⇩s⇩b' mis)
then obtain
ts⇩s⇩b': "ts⇩s⇩b' = ts⇩s⇩b[i := (p⇩s⇩b', is⇩s⇩b@mis, θ⇩s⇩b, sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis], 𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ⇩s⇩b)]" and
i_bound: "i < length ts⇩s⇩b" and
ts⇩s⇩b_i: "ts⇩s⇩b ! i = (p⇩s⇩b, is⇩s⇩b,θ⇩s⇩b,sb, 𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ⇩s⇩b)" and
prog: "θ⇩s⇩b⊢ p⇩s⇩b →⇩p (p⇩s⇩b',mis)" and
𝒮⇩s⇩b': "𝒮⇩s⇩b'=𝒮⇩s⇩b" and
m⇩s⇩b': "m⇩s⇩b'=m⇩s⇩b"
by auto
from sim obtain
m: "m = flush_all_until_volatile_write ts⇩s⇩b m⇩s⇩b" and
𝒮: "𝒮 = share_all_until_volatile_write ts⇩s⇩b 𝒮⇩s⇩b" and
leq: "length ts⇩s⇩b = length ts" and
ts_sim: "∀i<length ts⇩s⇩b.
let (p, is⇩s⇩b, θ, sb, 𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ) = ts⇩s⇩b ! i;
suspends = dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb
in ∃is 𝒟. instrs suspends @ is⇩s⇩b = is @ prog_instrs suspends ∧
𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b sb ≠ {}) ∧
ts ! i =
(hd_prog p suspends,
is,
θ |` (dom θ - read_tmps suspends), (),
𝒟,
acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪⇩s⇩b,
release (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) (dom 𝒮⇩s⇩b) ℛ)"
by cases blast
from i_bound leq have i_bound': "i < length ts"
by auto
have split_sb: "sb = takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb"
(is "sb = ?take_sb@?drop_sb")
by simp
from ts_sim [rule_format, OF i_bound] ts⇩s⇩b_i obtain suspends "is" 𝒟 where
suspends: "suspends = dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb" and
is_sim: "instrs suspends @ is⇩s⇩b = is @ prog_instrs suspends" and
𝒟: "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b sb ≠ {})" and
ts_i: "ts ! i =
(hd_prog p⇩s⇩b suspends, is,
θ⇩s⇩b |` (dom θ⇩s⇩b - read_tmps suspends), (), 𝒟, acquired True ?take_sb 𝒪⇩s⇩b,
release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by (auto simp add: Let_def)
from prog_step_preserves_valid [OF i_bound ts⇩s⇩b_i prog valid]
have valid': "valid ts⇩s⇩b'"
by (simp add: ts⇩s⇩b')
have valid_own': "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
show "outstanding_non_volatile_refs_owned_or_read_only 𝒮⇩s⇩b' ts⇩s⇩b'"
proof -
from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts⇩s⇩b_i]
have "non_volatile_owned_or_read_only False 𝒮⇩s⇩b 𝒪⇩s⇩b (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])"
by (auto simp add: non_volatile_owned_or_read_only_append)
from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b' 𝒮⇩s⇩b')
qed
next
show "outstanding_volatile_writes_unowned_by_others ts⇩s⇩b'"
proof -
have out: "outstanding_refs is_volatile_Write⇩s⇩b (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]) ⊆
outstanding_refs is_volatile_Write⇩s⇩b sb"
by (auto simp add: outstanding_refs_conv )
from outstanding_volatile_writes_unowned_by_others_store_buffer
[OF i_bound ts⇩s⇩b_i this]
show ?thesis by (simp add: ts⇩s⇩b' all_acquired_append)
qed
next
show "read_only_reads_unowned ts⇩s⇩b'"
proof -
have ro: "read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])) 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]))
⊆ read_only_reads (acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb) 𝒪⇩s⇩b)
(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb)"
apply (case_tac "outstanding_refs (is_volatile_Write⇩s⇩b) sb = {}")
apply (simp_all add: outstanding_vol_write_take_drop_appends
acquired_append read_only_reads_append )
done
have "𝒪⇩s⇩b ∪ all_acquired (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]) ⊆ 𝒪⇩s⇩b ∪ all_acquired sb"
by (auto simp add: all_acquired_append)
from read_only_reads_unowned_nth_update [OF i_bound ts⇩s⇩b_i ro this]
show ?thesis
by (simp add: ts⇩s⇩b' )
qed
next
show "ownership_distinct ts⇩s⇩b'"
proof -
from ownership_distinct_instructions_read_value_store_buffer_independent
[OF i_bound ts⇩s⇩b_i, where sb'="(sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])"]
show ?thesis by (simp add: ts⇩s⇩b' all_acquired_append)
qed
qed
from valid_last_prog [OF i_bound ts⇩s⇩b_i]
have last_prog: "last_prog p⇩s⇩b sb = p⇩s⇩b".
have valid_hist': "valid_history program_step ts⇩s⇩b'"
proof -
from valid_history [OF i_bound ts⇩s⇩b_i]
have "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b sb) sb".
from history_consistent_append_Prog⇩s⇩b [OF prog this last_prog]
have hist_consis': "history_consistent θ⇩s⇩b (hd_prog p⇩s⇩b' (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]))
(sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])".
from valid_history_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b')
qed
have valid_reads': "valid_reads m⇩s⇩b ts⇩s⇩b'"
proof -
from valid_reads [OF i_bound ts⇩s⇩b_i]
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b sb" .
from reads_consistent_snoc_Prog⇩s⇩b [OF this]
have "reads_consistent False 𝒪⇩s⇩b m⇩s⇩b (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])".
from valid_reads_nth_update [OF i_bound this]
show ?thesis by (simp add: ts⇩s⇩b')
qed
have valid_sharing': "valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'"
proof (intro_locales)
from outstanding_non_volatile_writes_unshared [OF i_bound ts⇩s⇩b_i]
have "non_volatile_writes_unshared 𝒮⇩s⇩b (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])"
by (auto simp add: non_volatile_writes_unshared_append)
from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
show "outstanding_non_volatile_writes_unshared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒮⇩s⇩b')
next
from sharing_consis [OF i_bound ts⇩s⇩b_i]
have "sharing_consistent 𝒮⇩s⇩b 𝒪⇩s⇩b (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])"
by (auto simp add: sharing_consistent_append)
from sharing_consis_nth_update [OF i_bound this]
show "sharing_consis 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒮⇩s⇩b')
next
from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound ts⇩s⇩b_i] ]
show "read_only_unowned 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b')
next
from unowned_shared_nth_update [OF i_bound ts⇩s⇩b_i subset_refl]
show "unowned_shared 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: ts⇩s⇩b' 𝒮⇩s⇩b')
next
from no_outstanding_write_to_read_only_memory [OF i_bound ts⇩s⇩b_i]
have "no_write_to_read_only_memory 𝒮⇩s⇩b (sb @ [Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])"
by (simp add: no_write_to_read_only_memory_append)
from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
show "no_outstanding_write_to_read_only_memory 𝒮⇩s⇩b' ts⇩s⇩b'"
by (simp add: 𝒮⇩s⇩b' ts⇩s⇩b')
qed
have tmps_distinct': "tmps_distinct ts⇩s⇩b'"
proof (intro_locales)
from load_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_load_tmps is⇩s⇩b".
with distinct_load_tmps_prog_step [OF i_bound ts⇩s⇩b_i prog valid]
have "distinct_load_tmps (is⇩s⇩b@mis)"
by (auto simp add: distinct_load_tmps_append)
from load_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_distinct ts⇩s⇩b'"
by (simp add: ts⇩s⇩b')
next
from read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
have "distinct_read_tmps (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])"
by (simp add: distinct_read_tmps_append)
from read_tmps_distinct_nth_update [OF i_bound this]
show "read_tmps_distinct ts⇩s⇩b'"
by (simp add: ts⇩s⇩b')
next
from load_tmps_read_tmps_distinct [OF i_bound ts⇩s⇩b_i]
distinct_load_tmps_prog_step [OF i_bound ts⇩s⇩b_i prog valid]
have "load_tmps (is⇩s⇩b@mis) ∩ read_tmps (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]) = {}"
by (auto simp add: read_tmps_append load_tmps_append)
from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
show "load_tmps_read_tmps_distinct ts⇩s⇩b'" by (simp add: ts⇩s⇩b')
qed
have valid_dd': "valid_data_dependency ts⇩s⇩b'"
proof -
from data_dependency_consistent_instrs [OF i_bound ts⇩s⇩b_i]
have "data_dependency_consistent_instrs (dom θ⇩s⇩b) is⇩s⇩b".
with valid_data_dependency_prog_step [OF i_bound ts⇩s⇩b_i prog valid]
load_tmps_write_tmps_distinct [OF i_bound ts⇩s⇩b_i]
obtain
"data_dependency_consistent_instrs (dom θ⇩s⇩b) (is⇩s⇩b@mis)"
"load_tmps (is⇩s⇩b@mis) ∩ ⋃(fst ` write_sops (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])) = {}"
by (force simp add: load_tmps_append data_dependency_consistent_instrs_append
write_sops_append)
from valid_data_dependency_nth_update [OF i_bound this]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
have load_tmps_fresh': "load_tmps_fresh ts⇩s⇩b'"
proof -
from load_tmps_fresh [OF i_bound ts⇩s⇩b_i]
load_tmps_fresh_prog_step [OF i_bound ts⇩s⇩b_i prog valid]
have "load_tmps (is⇩s⇩b@mis) ∩ dom θ⇩s⇩b = {}"
by (auto simp add: load_tmps_append)
from load_tmps_fresh_nth_update [OF i_bound this]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
have enough_flushs': "enough_flushs ts⇩s⇩b'"
proof -
from clean_no_outstanding_volatile_Write⇩s⇩b [OF i_bound ts⇩s⇩b_i]
have "¬ 𝒟⇩s⇩b ⟶ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]) = {}"
by (auto simp add: outstanding_refs_append)
from enough_flushs_nth_update [OF i_bound this]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
have valid_sops': "valid_sops ts⇩s⇩b'"
proof -
from valid_store_sops [OF i_bound ts⇩s⇩b_i] valid_sops_prog_step [OF prog]
valid_implies_valid_prog[OF i_bound ts⇩s⇩b_i valid]
have valid_store: "∀sop∈store_sops (is⇩s⇩b@mis). valid_sop sop"
by (auto simp add: store_sops_append)
from valid_write_sops [OF i_bound ts⇩s⇩b_i]
have "∀sop∈write_sops (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]). valid_sop sop"
by (auto simp add: write_sops_append)
from valid_sops_nth_update [OF i_bound this valid_store]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
have valid_program_history':"valid_program_history ts⇩s⇩b'"
proof -
from valid_program_history [OF i_bound ts⇩s⇩b_i]
have "causal_program_history is⇩s⇩b sb" .
from causal_program_history_Prog⇩s⇩b [OF this]
have causal': "causal_program_history (is⇩s⇩b@mis) (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])".
from last_prog_append_Prog⇩s⇩b
have "last_prog p⇩s⇩b' (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]) = p⇩s⇩b'".
from valid_program_history_nth_update [OF i_bound causal' this]
show ?thesis
by (simp add: ts⇩s⇩b')
qed
show ?thesis
proof (cases "outstanding_refs is_volatile_Write⇩s⇩b sb = {}")
case True
from True have flush_all: "takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = sb"
by (auto simp add: outstanding_refs_conv)
from True have suspend_nothing: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = []"
by (auto simp add: outstanding_refs_conv)
hence suspends_empty: "suspends = []"
by (simp add: suspends)
from suspends_empty is_sim have "is": "is = is⇩s⇩b"
by (simp)
from ts_i have ts_i: "ts ! i = (p⇩s⇩b, is⇩s⇩b, θ⇩s⇩b, (),
𝒟, acquired True ?take_sb 𝒪⇩s⇩b,release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)"
by (simp add: suspends_empty "is")
from direct_computation.Program [OF i_bound' ts_i prog]
have "(ts,m,𝒮) ⇒⇩d (ts[i := (p⇩s⇩b', is⇩s⇩b @ mis, θ⇩s⇩b, (),
𝒟, acquired True ?take_sb 𝒪⇩s⇩b,release ?take_sb (dom 𝒮⇩s⇩b) ℛ⇩s⇩b)], m, 𝒮)".
moreover
note flush_commute = flush_all_until_volatile_write_append_Prog⇩s⇩b_commute [OF i_bound ts⇩s⇩b_i]
from True
have suspend_nothing':
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb @ [Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])) = []"
by (auto simp add: outstanding_refs_conv)
note share_commute =
share_all_until_volatile_write_update_sb [OF share_append_Prog⇩s⇩b i_bound ts⇩s⇩b_i]
from 𝒟
have 𝒟': "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]) ≠ {})"
by (auto simp: outstanding_refs_append)
have "(ts⇩s⇩b [i := (p⇩s⇩b',is⇩s⇩b@mis, θ⇩s⇩b, sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis], 𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ⇩s⇩b)],
m⇩s⇩b,𝒮⇩s⇩b') ∼
(ts[i:=(p⇩s⇩b', is⇩s⇩b @ mis, θ⇩s⇩b, (), 𝒟,
acquired True (takeWhile (Not ∘ is_volatile_Write⇩s⇩b)
(sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])) 𝒪⇩s⇩b,
release (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]) (dom 𝒮⇩s⇩b) ℛ⇩s⇩b )],m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute)
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' share_commute)
using leq
apply simp
using i_bound i_bound' ts_sim ts_i 𝒟'
apply (clarsimp simp add: Let_def nth_list_update flush_all suspend_nothing' Prog⇩s⇩b 𝒮⇩s⇩b'
release_append_Prog⇩s⇩b release_append
split: if_split_asm)
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' m⇩s⇩b'
valid_dd' valid_sops' load_tmps_fresh' enough_flushs' valid_sharing'
valid_program_history' valid'
𝒮⇩s⇩b' ts⇩s⇩b'
by (auto simp del: fun_upd_apply simp add: acquired_append_Prog⇩s⇩b release_append_Prog⇩s⇩b release_append flush_all)
next
case False
then obtain r where r_in: "r ∈ set sb" and volatile_r: "is_volatile_Write⇩s⇩b r"
by (auto simp add: outstanding_refs_conv)
from takeWhile_dropWhile_real_prefix
[OF r_in, of "(Not ∘ is_volatile_Write⇩s⇩b)", simplified, OF volatile_r]
obtain a' v' sb'' sop' A' L' R' W' where
sb_split: "sb = takeWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
and
drop: "dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb = Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
apply (auto)
subgoal for y
apply (case_tac y)
apply auto
done
done
from drop suspends have suspends': "suspends = Write⇩s⇩b True a' sop' v' A' L' R' W'# sb''"
by simp
have "(ts, m, 𝒮) ⇒⇩d⇧* (ts, m, 𝒮)" by auto
moreover
note flush_commute= flush_all_until_volatile_write_append_Prog⇩s⇩b_commute [OF i_bound ts⇩s⇩b_i]
have "Write⇩s⇩b True a' sop' v' A' L' R' W' ∈ set sb"
by (subst sb_split) auto
from dropWhile_append1 [OF this, of "(Not ∘ is_volatile_Write⇩s⇩b)"]
have drop_app_comm:
"(dropWhile (Not ∘ is_volatile_Write⇩s⇩b) (sb @ [Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis])) =
dropWhile (Not ∘ is_volatile_Write⇩s⇩b) sb @ [Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]"
by simp
note share_commute =
share_all_until_volatile_write_update_sb [OF share_append_Prog⇩s⇩b i_bound ts⇩s⇩b_i]
from 𝒟
have 𝒟': "𝒟⇩s⇩b = (𝒟 ∨ outstanding_refs is_volatile_Write⇩s⇩b (sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis]) ≠ {})"
by (auto simp: outstanding_refs_append)
have "(ts⇩s⇩b [i := (p⇩s⇩b',is⇩s⇩b@mis,θ⇩s⇩b, sb@[Prog⇩s⇩b p⇩s⇩b p⇩s⇩b' mis], 𝒟⇩s⇩b, 𝒪⇩s⇩b,ℛ⇩s⇩b)],
m⇩s⇩b,𝒮⇩s⇩b') ∼
(ts,m,𝒮)"
apply (rule sim_config.intros)
apply (simp add: m flush_commute)
apply (clarsimp simp add: 𝒮 𝒮⇩s⇩b' share_commute)
using leq
apply simp
using i_bound i_bound' ts_sim ts_i is_sim suspends suspends' [simplified suspends] 𝒟'
apply (clarsimp simp add: Let_def nth_list_update Prog⇩s⇩b
drop_app_comm instrs_append prog_instrs_append
read_tmps_append hd_prog_append_Prog⇩s⇩b acquired_append_Prog⇩s⇩b release_append_Prog⇩s⇩b release_append 𝒮⇩s⇩b'
split: if_split_asm)
done
ultimately show ?thesis
using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' m⇩s⇩b'
valid_dd' valid_sops' load_tmps_fresh' enough_flushs' valid_sharing'
valid_program_history' valid'
𝒮⇩s⇩b' ts⇩s⇩b'
by (auto simp del: fun_upd_apply)
qed
qed
qed
theorem (in xvalid_program) concurrent_direct_steps_simulates_store_buffer_history_steps:
assumes step_sb: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ⇒⇩s⇩b⇩h⇧* (ts⇩s⇩b',m⇩s⇩b',𝒮⇩s⇩b')"
assumes valid_own: "valid_ownership 𝒮⇩s⇩b ts⇩s⇩b"
assumes valid_sb_reads: "valid_reads m⇩s⇩b ts⇩s⇩b"
assumes valid_hist: "valid_history program_step ts⇩s⇩b"
assumes valid_sharing: "valid_sharing 𝒮⇩s⇩b ts⇩s⇩b"
assumes tmps_distinct: "tmps_distinct ts⇩s⇩b"
assumes valid_sops: "valid_sops ts⇩s⇩b"
assumes valid_dd: "valid_data_dependency ts⇩s⇩b"
assumes load_tmps_fresh: "load_tmps_fresh ts⇩s⇩b"
assumes enough_flushs: "enough_flushs ts⇩s⇩b"
assumes valid_program_history: "valid_program_history ts⇩s⇩b"
assumes valid: "valid ts⇩s⇩b"
shows "⋀ts 𝒮 m. (ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m,𝒮) ⟹ safe_reach_direct safe_delayed (ts,m,𝒮) ⟹
valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b' ∧ valid_reads m⇩s⇩b' ts⇩s⇩b' ∧ valid_history program_step ts⇩s⇩b' ∧
valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b' ∧ tmps_distinct ts⇩s⇩b' ∧ valid_data_dependency ts⇩s⇩b' ∧
valid_sops ts⇩s⇩b' ∧ load_tmps_fresh ts⇩s⇩b' ∧ enough_flushs ts⇩s⇩b' ∧
valid_program_history ts⇩s⇩b' ∧ valid ts⇩s⇩b' ∧
(∃ts' m' 𝒮'. (ts,m,𝒮) ⇒⇩d⇧* (ts',m',𝒮') ∧ (ts⇩s⇩b',m⇩s⇩b',𝒮⇩s⇩b') ∼ (ts',m',𝒮'))"
using step_sb valid_own valid_sb_reads valid_hist valid_sharing tmps_distinct valid_sops
valid_dd load_tmps_fresh enough_flushs valid_program_history valid
proof (induct rule: converse_rtranclp_induct_sbh_steps)
case refl thus ?case
by auto
next
case (step ts⇩s⇩b m⇩s⇩b 𝒮⇩s⇩b ts⇩s⇩b'' m⇩s⇩b'' 𝒮⇩s⇩b'')
note first = ‹(ts⇩s⇩b, m⇩s⇩b, 𝒮⇩s⇩b) ⇒⇩s⇩b⇩h (ts⇩s⇩b'', m⇩s⇩b'', 𝒮⇩s⇩b'')›
note sim = ‹(ts⇩s⇩b, m⇩s⇩b, 𝒮⇩s⇩b) ∼ (ts, m, 𝒮)›
note safe_reach = ‹safe_reach_direct safe_delayed (ts, m, 𝒮)›
note valid_own = ‹valid_ownership 𝒮⇩s⇩b ts⇩s⇩b›
note valid_reads = ‹valid_reads m⇩s⇩b ts⇩s⇩b›
note valid_hist = ‹valid_history program_step ts⇩s⇩b›
note valid_sharing = ‹valid_sharing 𝒮⇩s⇩b ts⇩s⇩b›
note tmps_distinct = ‹tmps_distinct ts⇩s⇩b›
note valid_sops = ‹valid_sops ts⇩s⇩b›
note valid_dd = ‹valid_data_dependency ts⇩s⇩b›
note load_tmps_fresh = ‹load_tmps_fresh ts⇩s⇩b›
note enough_flushs = ‹enough_flushs ts⇩s⇩b›
note valid_prog_hist = ‹valid_program_history ts⇩s⇩b›
note valid = ‹valid ts⇩s⇩b›
from concurrent_direct_steps_simulates_store_buffer_history_step [OF first
valid_own valid_reads valid_hist valid_sharing tmps_distinct valid_sops valid_dd
load_tmps_fresh enough_flushs valid_prog_hist valid sim safe_reach]
obtain ts'' m'' 𝒮'' where
valid_own'': "valid_ownership 𝒮⇩s⇩b'' ts⇩s⇩b''" and
valid_reads'': "valid_reads m⇩s⇩b'' ts⇩s⇩b''" and
valid_hist'': "valid_history program_step ts⇩s⇩b''" and
valid_sharing'': "valid_sharing 𝒮⇩s⇩b'' ts⇩s⇩b''" and
tmps_dist'': "tmps_distinct ts⇩s⇩b''" and
valid_dd'': "valid_data_dependency ts⇩s⇩b''" and
valid_sops'': "valid_sops ts⇩s⇩b''" and
load_tmps_fresh'': "load_tmps_fresh ts⇩s⇩b''" and
enough_flushs'': "enough_flushs ts⇩s⇩b''" and
valid_prog_hist'': "valid_program_history ts⇩s⇩b''"and
valid'': "valid ts⇩s⇩b''" and
steps: "(ts, m, 𝒮) ⇒⇩d⇧* (ts'', m'', 𝒮'')" and
sim: "(ts⇩s⇩b'', m⇩s⇩b'',𝒮⇩s⇩b'') ∼ (ts'', m'',𝒮'')"
by blast
from step.hyps (3) [OF sim safe_reach_steps [OF safe_reach steps] valid_own'' valid_reads'' valid_hist'' valid_sharing''
tmps_dist'' valid_sops'' valid_dd'' load_tmps_fresh'' enough_flushs'' valid_prog_hist'' valid'' ]
obtain ts' m' 𝒮' where
valid: "valid_ownership 𝒮⇩s⇩b' ts⇩s⇩b'" "valid_reads m⇩s⇩b' ts⇩s⇩b'" "valid_history program_step ts⇩s⇩b'"
"valid_sharing 𝒮⇩s⇩b' ts⇩s⇩b'" "tmps_distinct ts⇩s⇩b'" "valid_data_dependency ts⇩s⇩b'"
"valid_sops ts⇩s⇩b'" "load_tmps_fresh ts⇩s⇩b'" "enough_flushs ts⇩s⇩b'"
"valid_program_history ts⇩s⇩b'" "valid ts⇩s⇩b'" and
last: "(ts'', m'', 𝒮'') ⇒⇩d⇧* (ts', m', 𝒮')" and
sim': "(ts⇩s⇩b', m⇩s⇩b',𝒮⇩s⇩b') ∼ (ts', m',𝒮')"
by blast
note steps also note last
finally show ?case
using valid sim'
by blast
qed
sublocale initial⇩s⇩b ⊆ tmps_distinct ..
locale xvalid_program_progress = program_progress + xvalid_program
theorem (in xvalid_program_progress) concurrent_direct_execution_simulates_store_buffer_history_execution:
assumes exec_sb: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ⇒⇩s⇩b⇩h⇧* (ts⇩s⇩b',m⇩s⇩b',𝒮⇩s⇩b')"
assumes init: "initial⇩s⇩b ts⇩s⇩b 𝒮⇩s⇩b"
assumes valid: "valid ts⇩s⇩b"
assumes sim: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m,𝒮)"
assumes safe: "safe_reach_direct safe_free_flowing (ts,m,𝒮)"
shows "∃ts' m' 𝒮'. (ts,m,𝒮) ⇒⇩d⇧* (ts',m',𝒮') ∧
(ts⇩s⇩b',m⇩s⇩b',𝒮⇩s⇩b') ∼ (ts',m',𝒮')"
proof -
from init interpret ini: initial⇩s⇩b ts⇩s⇩b 𝒮⇩s⇩b .
from safe_free_flowing_implies_safe_delayed' [OF init sim safe]
have safe_delayed: "safe_reach_direct safe_delayed (ts, m, 𝒮)".
from local.ini.valid_ownership_axioms have "valid_ownership 𝒮⇩s⇩b ts⇩s⇩b" .
from local.ini.valid_reads_axioms have "valid_reads m⇩s⇩b ts⇩s⇩b".
from local.ini.valid_history_axioms have "valid_history program_step ts⇩s⇩b".
from local.ini.valid_sharing_axioms have "valid_sharing 𝒮⇩s⇩b ts⇩s⇩b".
from local.ini.tmps_distinct_axioms have "tmps_distinct ts⇩s⇩b".
from local.ini.valid_sops_axioms have "valid_sops ts⇩s⇩b".
from local.ini.valid_data_dependency_axioms have "valid_data_dependency ts⇩s⇩b".
from local.ini.load_tmps_fresh_axioms have "load_tmps_fresh ts⇩s⇩b".
from local.ini.enough_flushs_axioms have "enough_flushs ts⇩s⇩b".
from local.ini.valid_program_history_axioms have "valid_program_history ts⇩s⇩b".
from concurrent_direct_steps_simulates_store_buffer_history_steps [OF exec_sb
‹valid_ownership 𝒮⇩s⇩b ts⇩s⇩b›
‹valid_reads m⇩s⇩b ts⇩s⇩b› ‹valid_history program_step ts⇩s⇩b›
‹valid_sharing 𝒮⇩s⇩b ts⇩s⇩b› ‹tmps_distinct ts⇩s⇩b› ‹valid_sops ts⇩s⇩b›
‹valid_data_dependency ts⇩s⇩b› ‹load_tmps_fresh ts⇩s⇩b› ‹enough_flushs ts⇩s⇩b›
‹valid_program_history ts⇩s⇩b› valid sim safe_delayed]
show ?thesis by auto
qed
lemma filter_is_Write⇩s⇩b_Cons_Write⇩s⇩b: "filter is_Write⇩s⇩b xs = Write⇩s⇩b volatile a sop v A L R W#ys
⟹ ∃rs rws. (∀r ∈ set rs. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r) ∧
xs=rs@Write⇩s⇩b volatile a sop v A L R W#rws ∧ ys=filter is_Write⇩s⇩b rws"
proof (induct xs)
case Nil thus ?case by simp
next
case (Cons x xs)
note feq = ‹filter is_Write⇩s⇩b (x#xs) = Write⇩s⇩b volatile a sop v A L R W# ys›
show ?case
proof (cases x)
case (Write⇩s⇩b volatile' a' sop' v' A' L' R' W')
with feq obtain "volatile'=volatile" "a'=a" "v'=v" "sop'=sop" "A'=A" "L'=L" "R'=R" "W'=W"
"ys = filter is_Write⇩s⇩b xs"
by auto
thus ?thesis
apply -
apply (rule_tac x="[]" in exI)
apply (rule_tac x="xs" in exI)
apply (simp add: Write⇩s⇩b)
done
next
case (Read⇩s⇩b volatile' a' t' v')
from feq have "filter is_Write⇩s⇩b xs = Write⇩s⇩b volatile a sop v A L R W#ys"
by (simp add: Read⇩s⇩b)
from Cons.hyps [OF this] obtain rs rws where
"∀r ∈ set rs. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r" and
"xs=rs @ Write⇩s⇩b volatile a sop v A L R W# rws" and
"ys=filter is_Write⇩s⇩b rws"
by clarsimp
then show ?thesis
apply -
apply (rule_tac x="Read⇩s⇩b volatile' a' t' v'#rs" in exI)
apply (rule_tac x="rws" in exI)
apply (simp add: Read⇩s⇩b)
done
next
case (Prog⇩s⇩b p⇩1 p⇩2 mis)
from feq have "filter is_Write⇩s⇩b xs = Write⇩s⇩b volatile a sop v A L R W#ys"
by (simp add: Prog⇩s⇩b)
from Cons.hyps [OF this] obtain rs rws where
"∀r ∈ set rs. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r" and
"xs=rs @ Write⇩s⇩b volatile a sop v A L R W# rws" and
"ys=filter is_Write⇩s⇩b rws"
by clarsimp
then show ?thesis
apply -
apply (rule_tac x="Prog⇩s⇩b p⇩1 p⇩2 mis#rs" in exI)
apply (rule_tac x="rws" in exI)
apply (simp add: Prog⇩s⇩b)
done
next
case (Ghost⇩s⇩b A' L' R' W')
from feq have "filter is_Write⇩s⇩b xs = Write⇩s⇩b volatile a sop v A L R W # ys"
by (simp add: Ghost⇩s⇩b)
from Cons.hyps [OF this] obtain rs rws where
"∀r ∈ set rs. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r" and
"xs=rs @ Write⇩s⇩b volatile a sop v A L R W# rws" and
"ys=filter is_Write⇩s⇩b rws"
by clarsimp
then show ?thesis
apply -
apply (rule_tac x="Ghost⇩s⇩b A' L' R' W'#rs" in exI)
apply (rule_tac x="rws" in exI)
apply (simp add: Ghost⇩s⇩b)
done
qed
qed
lemma filter_is_Write⇩s⇩b_empty: "filter is_Write⇩s⇩b xs = []
⟹ (∀r ∈ set xs. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r)"
proof (induct xs)
case Nil thus ?case by simp
next
case (Cons x xs)
note feq = ‹filter is_Write⇩s⇩b (x#xs) = []›
show ?case
proof (cases x)
case (Write⇩s⇩b volatile' a' v')
with feq have False
by simp
thus ?thesis ..
next
case (Read⇩s⇩b a' v')
from feq have "filter is_Write⇩s⇩b xs = []"
by (simp add: Read⇩s⇩b)
from Cons.hyps [OF this] obtain
"∀r ∈ set xs. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r"
by clarsimp
then show ?thesis
by (simp add: Read⇩s⇩b)
next
case (Prog⇩s⇩b p⇩2 p⇩2 mis)
from feq have "filter is_Write⇩s⇩b xs = []"
by (simp add: Prog⇩s⇩b)
from Cons.hyps [OF this] obtain
"∀r ∈ set xs. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r"
by clarsimp
then show ?thesis
by (simp add: Prog⇩s⇩b)
next
case (Ghost⇩s⇩b A' L' R' W')
from feq have "filter is_Write⇩s⇩b xs = []"
by (simp add: Ghost⇩s⇩b)
from Cons.hyps [OF this] obtain
"∀r ∈ set xs. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r"
by clarsimp
then show ?thesis
by (simp add: Ghost⇩s⇩b)
qed
qed
lemma flush_reads_program: "⋀𝒪 𝒮 ℛ .
∀r ∈ set sb. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r ⟹
∃𝒪' ℛ' 𝒮'. (m,sb,𝒪,ℛ,𝒮) →⇩f⇧* (m,[],𝒪',ℛ',𝒮')"
proof (induct sb)
case Nil thus ?case by auto
next
case (Cons x sb)
note ‹∀r∈set (x # sb). is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r›
then obtain x: "is_Read⇩s⇩b x ∨ is_Prog⇩s⇩b x ∨ is_Ghost⇩s⇩b x" and sb: "∀r∈set sb. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r"
by (cases x) auto
{
assume "is_Read⇩s⇩b x"
then obtain volatile a t v where x: "x=Read⇩s⇩b volatile a t v"
by (cases x) auto
have "(m,Read⇩s⇩b volatile a t v#sb,𝒪,ℛ,𝒮) →⇩f (m,sb,𝒪,ℛ,𝒮)"
by (rule Read⇩s⇩b)
also
from Cons.hyps [OF sb] obtain 𝒪' 𝒮' acq' ℛ'
where "(m, sb,𝒪,ℛ,𝒮) →⇩f⇧* (m, [],𝒪',ℛ',𝒮')" by blast
finally
have ?case
by (auto simp add: x)
}
moreover
{
assume "is_Prog⇩s⇩b x"
then obtain p⇩1 p⇩2 mis where x: "x=Prog⇩s⇩b p⇩1 p⇩2 mis"
by (cases x) auto
have "(m,Prog⇩s⇩b p⇩1 p⇩2 mis#sb,𝒪,ℛ,𝒮) →⇩f (m,sb,𝒪,ℛ,𝒮)"
by (rule Prog⇩s⇩b)
also
from Cons.hyps [OF sb] obtain 𝒪' ℛ' 𝒮' acq'
where "(m, sb,𝒪,ℛ,𝒮) →⇩f⇧* (m, [],𝒪',ℛ',𝒮')" by blast
finally
have ?case
by (auto simp add: x)
}
moreover
{
assume "is_Ghost⇩s⇩b x"
then obtain A L R W where x: "x=Ghost⇩s⇩b A L R W"
by (cases x) auto
have "(m,Ghost⇩s⇩b A L R W#sb,𝒪,ℛ,𝒮) →⇩f (m,sb,𝒪 ∪ A - R,augment_rels (dom 𝒮) R ℛ,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (rule Ghost)
also
from Cons.hyps [OF sb] obtain 𝒪' 𝒮' ℛ' acq'
where "(m, sb,𝒪 ∪ A - R ,augment_rels (dom 𝒮) R ℛ,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L) →⇩f⇧* (m, [],𝒪',ℛ',𝒮')" by blast
finally
have ?case
by (auto simp add: x)
}
ultimately show ?case
using x by blast
qed
lemma flush_progress: "∃m' 𝒪' 𝒮' ℛ'. (m,r#sb,𝒪,ℛ,𝒮) →⇩f (m',sb,𝒪',ℛ',𝒮')"
proof (cases r)
case (Write⇩s⇩b volatile a sop v A L R W)
from flush_step.Write⇩s⇩b [OF refl refl refl, of m volatile a sop v A L R W sb 𝒪 ℛ 𝒮]
show ?thesis
by (auto simp add: Write⇩s⇩b)
next
case (Read⇩s⇩b volatile a t v)
from flush_step.Read⇩s⇩b [of m volatile a t v sb 𝒪 ℛ 𝒮]
show ?thesis
by (auto simp add: Read⇩s⇩b)
next
case (Prog⇩s⇩b p⇩1 p⇩2 mis)
from flush_step.Prog⇩s⇩b [of m p⇩1 p⇩2 mis sb 𝒪 ℛ 𝒮]
show ?thesis
by (auto simp add: Prog⇩s⇩b)
next
case (Ghost⇩s⇩b A L R W)
from flush_step.Ghost [of m A L R W sb 𝒪 ℛ 𝒮]
show ?thesis
by (auto simp add: Ghost⇩s⇩b)
qed
lemma flush_empty:
assumes steps: "(m, sb,𝒪,ℛ, 𝒮) →⇩f⇧* (m', sb',𝒪',ℛ',𝒮')"
shows "sb=[] ⟹ m'=m ∧ sb'=[] ∧ 𝒪'=𝒪 ∧ ℛ'=ℛ ∧ 𝒮'=𝒮 "
using steps
apply (induct rule: converse_rtranclp_induct5)
apply (auto elim: flush_step.cases)
done
lemma flush_append:
assumes steps: "(m, sb,𝒪,ℛ,𝒮) →⇩f⇧* (m', sb',𝒪',ℛ',𝒮')"
shows "⋀xs. (m, sb@xs,𝒪,ℛ,𝒮) →⇩f⇧* (m', sb'@xs,𝒪',ℛ',𝒮')"
using steps
proof (induct rule: converse_rtranclp_induct5)
case refl thus ?case by auto
next
case (step m sb 𝒪 ℛ 𝒮 m'' sb'' 𝒪'' ℛ'' 𝒮'')
note first= ‹(m,sb,𝒪,ℛ,𝒮) →⇩f (m'',sb'',𝒪'',ℛ'',𝒮'')›
note rest = ‹(m'', sb'',𝒪'',ℛ'',𝒮'') →⇩f⇧* (m', sb',𝒪',ℛ',𝒮')›
from step.hyps (3) have append_rest: "(m'', sb''@xs,𝒪'',ℛ'',𝒮'') →⇩f⇧* (m', sb'@xs,𝒪',ℛ',𝒮')".
from first show ?case
proof (cases)
case (Write⇩s⇩b volatile A R W L a sop v)
then obtain sb: "sb=Write⇩s⇩b volatile a sop v A L R W#sb''" and m'': "m''=m(a:=v)" and
𝒪'': "𝒪''=(if volatile then 𝒪 ∪ A - R else 𝒪)" and
ℛ'': "ℛ''=(if volatile then Map.empty else ℛ)" and
𝒮'': "𝒮''=(if volatile then 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L else 𝒮)"
by auto
have "(m,Write⇩s⇩b volatile a sop v A L R W#sb''@xs,𝒪,ℛ,𝒮) →⇩f
(m(a:=v),sb''@xs,if volatile then 𝒪 ∪ A - R else 𝒪,if volatile then Map.empty else ℛ,
if volatile then 𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L else 𝒮)"
apply (rule flush_step.Write⇩s⇩b)
apply auto
done
hence "(m,sb@xs,𝒪,ℛ,𝒮) →⇩f (m'',sb''@xs,𝒪'',ℛ'',𝒮'')"
by (simp add: sb m'' 𝒪'' ℛ'' 𝒮'')
also note append_rest
finally show ?thesis .
next
case (Read⇩s⇩b volatile a t v)
then obtain sb: "sb=Read⇩s⇩b volatile a t v #sb''" and m'': "m''=m"
and 𝒪'': "𝒪''=𝒪" and 𝒮'': "𝒮''=𝒮" and ℛ'': "ℛ''=ℛ"
by auto
have "(m,Read⇩s⇩b volatile a t v#sb''@xs,𝒪,ℛ,𝒮) →⇩f (m,sb''@xs,𝒪,ℛ,𝒮)"
by (rule flush_step.Read⇩s⇩b)
hence "(m,sb@xs,𝒪,ℛ,𝒮) →⇩f (m'',sb''@xs,𝒪'',ℛ'',𝒮'')"
by (simp add: sb m'' 𝒪'' ℛ'' 𝒮'' )
also note append_rest
finally show ?thesis .
next
case (Prog⇩s⇩b p⇩1 p⇩2 mis)
then obtain sb: "sb=Prog⇩s⇩b p⇩1 p⇩2 mis#sb''" and m'': "m''=m"
and 𝒪'': "𝒪''=𝒪" and 𝒮'': "𝒮''=𝒮" and ℛ'': "ℛ''=ℛ"
by auto
have "(m,Prog⇩s⇩b p⇩1 p⇩2 mis#sb''@xs,𝒪,ℛ,𝒮) →⇩f (m,sb''@xs,𝒪,ℛ,𝒮)"
by (rule flush_step.Prog⇩s⇩b)
hence "(m,sb@xs,𝒪,ℛ,𝒮) →⇩f (m'',sb''@xs,𝒪'',ℛ'',𝒮'')"
by (simp add: sb m'' 𝒪'' ℛ'' 𝒮'' )
also note append_rest
finally show ?thesis .
next
case (Ghost A L R W)
then obtain sb: "sb=Ghost⇩s⇩b A L R W#sb''" and m'': "m''=m"
and 𝒪'': "𝒪''=𝒪 ∪ A - R" and 𝒮'': "𝒮''=𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L" and
ℛ'': "ℛ''=augment_rels (dom 𝒮) R ℛ"
by auto
have "(m,Ghost⇩s⇩b A L R W#sb''@xs,𝒪,ℛ,𝒮) →⇩f (m,sb''@xs,𝒪 ∪ A - R,augment_rels (dom 𝒮) R ℛ,𝒮 ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (rule flush_step.Ghost)
hence "(m,sb@xs,𝒪,ℛ,𝒮) →⇩f (m'',sb''@xs,𝒪'',ℛ'',𝒮'')"
by (simp add: sb m'' 𝒪'' ℛ'' 𝒮'' )
also note append_rest
finally show ?thesis .
qed
qed
lemmas store_buffer_step_induct =
store_buffer_step.induct [split_format (complete),
consumes 1, case_names SBWrite⇩s⇩b]
theorem flush_simulates_filter_writes:
assumes step: "(m,sb,𝒪,ℛ,𝒮) →⇩w (m',sb',𝒪',ℛ',𝒮')"
shows "⋀sb⇩h 𝒪⇩h ℛ⇩h 𝒮⇩h. sb=filter is_Write⇩s⇩b sb⇩h
⟹
∃sb⇩h' 𝒪⇩h' ℛ⇩h' 𝒮⇩h'. (m,sb⇩h,𝒪⇩h,ℛ⇩h,𝒮⇩h) →⇩f⇧* (m',sb⇩h',𝒪⇩h',ℛ⇩h',𝒮⇩h') ∧
sb'=filter is_Write⇩s⇩b sb⇩h' ∧ (sb'=[] ⟶ sb⇩h'=[])"
using step
proof (induct rule: store_buffer_step_induct)
case (SBWrite⇩s⇩b m volatile a D f v A L R W sb 𝒪 ℛ 𝒮)
note filter_Write⇩s⇩b = ‹Write⇩s⇩b volatile a (D,f) v A L R W# sb = filter is_Write⇩s⇩b sb⇩h›
from filter_is_Write⇩s⇩b_Cons_Write⇩s⇩b [OF filter_Write⇩s⇩b [symmetric]]
obtain rs rws where
rs_reads: "∀r∈set rs. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r" and
sb⇩h: "sb⇩h = rs @ Write⇩s⇩b volatile a (D,f) v A L R W# rws" and
sb: "sb = filter is_Write⇩s⇩b rws"
by blast
from flush_reads_program [OF rs_reads] obtain 𝒪⇩h' ℛ⇩h' 𝒮⇩h' acq⇩h'
where "(m, rs,𝒪⇩h,ℛ⇩h,𝒮⇩h) →⇩f⇧* (m, [],𝒪⇩h',ℛ⇩h',𝒮⇩h')" by blast
from flush_append [OF this]
have "(m, rs@Write⇩s⇩b volatile a (D,f) v A L R W# rws,𝒪⇩h,ℛ⇩h,𝒮⇩h) →⇩f⇧*
(m, Write⇩s⇩b volatile a (D,f) v A L R W# rws,𝒪⇩h',ℛ⇩h',𝒮⇩h')"
by simp
also
from flush_step.Write⇩s⇩b [OF refl refl refl, of m volatile a "(D,f)" v A L R W rws 𝒪⇩h' ℛ⇩h' 𝒮⇩h']
obtain 𝒪⇩h'' ℛ⇩h'' 𝒮⇩h''
where "(m, Write⇩s⇩b volatile a (D,f) v A L R W# rws,𝒪⇩h',ℛ⇩h',𝒮⇩h') →⇩f (m(a:=v), rws, 𝒪⇩h'',ℛ⇩h'',𝒮⇩h'')"
by auto
finally have steps: "(m, sb⇩h,𝒪⇩h,ℛ⇩h,𝒮⇩h) →⇩f⇧* (m(a:=v), rws,𝒪⇩h'',ℛ⇩h'',𝒮⇩h'')"
by (simp add: sb⇩h sb)
show ?case
proof (cases "sb")
case Cons
with steps sb show ?thesis
by fastforce
next
case Nil
from filter_is_Write⇩s⇩b_empty [OF sb [simplified Nil, symmetric]]
have "∀r∈set rws. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r".
from flush_reads_program [OF this] obtain 𝒪⇩h''' ℛ⇩h''' 𝒮⇩h''' acq⇩h'''
where "(m(a:=v), rws,𝒪⇩h'',ℛ⇩h'',𝒮⇩h'') →⇩f⇧* (m(a:=v), [],𝒪⇩h''',ℛ⇩h''',𝒮⇩h''')" by blast
with steps
have "(m, sb⇩h,𝒪⇩h,ℛ⇩h,𝒮⇩h) →⇩f⇧* (m(a:=v), [],𝒪⇩h''',ℛ⇩h''',𝒮⇩h''')" by force
with sb Nil show ?thesis by fastforce
qed
qed
lemma bufferd_val_filter_is_Write⇩s⇩b_eq_ext:
"buffered_val (filter is_Write⇩s⇩b sb) a = buffered_val sb a"
by (induct sb) (auto split: memref.splits)
lemma bufferd_val_filter_is_Write⇩s⇩b_eq:
"buffered_val (filter is_Write⇩s⇩b sb) = buffered_val sb"
by (rule ext) (rule bufferd_val_filter_is_Write⇩s⇩b_eq_ext)
lemma outstanding_refs_is_volatile_Write⇩s⇩b_filter_writes:
"outstanding_refs is_volatile_Write⇩s⇩b (filter is_Write⇩s⇩b xs) =
outstanding_refs is_volatile_Write⇩s⇩b xs"
by (induct xs) (auto simp add: is_volatile_Write⇩s⇩b_def split: memref.splits)
subsection ‹Simulation of Store Buffer Machine without History by Store Buffer Machine with History›
theorem (in valid_program) concurrent_history_steps_simulates_store_buffer_step:
assumes step_sb: "(ts,m,𝒮) ⇒⇩s⇩b (ts',m',𝒮')"
assumes sim: "ts ∼⇩h ts⇩h"
shows "∃ts⇩h' 𝒮⇩h'. (ts⇩h,m,𝒮⇩h) ⇒⇩s⇩b⇩h⇧* (ts⇩h',m',𝒮⇩h') ∧ ts' ∼⇩h ts⇩h'"
proof -
interpret sbh_computation:
computation sbh_memop_step flush_step program_step
"λp p' is sb. sb @ [Prog⇩s⇩b p p' is]" .
from step_sb
show ?thesis
proof (cases rule: concurrent_step_cases)
case (Memop i _ p "is" θ sb 𝒟 𝒪 ℛ _ _ is' θ' sb' _ 𝒟' 𝒪' ℛ')
then obtain
ts': "ts' = ts[i := (p, is', θ', sb', 𝒟', 𝒪',ℛ')]" and
i_bound: "i < length ts" and
ts_i: "ts ! i = (p, is, θ, sb, 𝒟, 𝒪,ℛ)" and
step_sb: "(is, θ, sb, m, 𝒟, 𝒪, ℛ,𝒮) →⇩s⇩b
(is', θ', sb', m', 𝒟', 𝒪', ℛ',𝒮')"
by auto
from sim obtain
lts_eq: "length ts = length ts⇩h" and
sim_loc: "∀i < length ts. (∃𝒪' 𝒟' ℛ'.
let (p,is, θ, sb,𝒟, 𝒪,ℛ) = ts⇩h!i in
ts!i=(p,is, θ, filter is_Write⇩s⇩b sb,𝒟',𝒪',ℛ') ∧
(filter is_Write⇩s⇩b sb = [] ⟶ sb=[]))"
by cases (auto)
from lts_eq i_bound have i_bound': "i < length ts⇩h"
by simp
from step_sb
show ?thesis
proof (cases)
case (SBReadBuffered a v volatile t)
then obtain
"is": "is = Read volatile a t#is'" and
𝒪': "𝒪'=𝒪" and
𝒮': "𝒮'=𝒮" and
ℛ': "ℛ'=ℛ" and
𝒟': "𝒟'=𝒟" and
m': "m'=m" and
θ': "θ'=θ(t↦v)" and
sb': "sb' = sb" and
buf_val: "buffered_val sb a = Some v"
by auto
from sim_loc [rule_format, OF i_bound] ts_i "is"
obtain sb⇩h 𝒪⇩h ℛ⇩h 𝒟⇩h where
ts⇩h_i: "ts⇩h!i = (p,Read volatile a t#is',θ,sb⇩h,𝒟⇩h,𝒪⇩h,ℛ⇩h)" and
sb: "sb = filter is_Write⇩s⇩b sb⇩h" and
sb_empty: "filter is_Write⇩s⇩b sb⇩h = [] ⟶ sb⇩h=[]"
by (auto simp add: Let_def)
from buf_val
have buf_val': "buffered_val sb⇩h a = Some v"
by (simp add: bufferd_val_filter_is_Write⇩s⇩b_eq sb)
let ?ts⇩h_i' = "(p, is', θ(t ↦ v), sb⇩h @ [Read⇩s⇩b volatile a t v], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)"
let ?ts⇩h' = "ts⇩h[i := ?ts⇩h_i']"
from sbh_memop_step.SBHReadBuffered [OF buf_val']
have "(Read volatile a t # is', θ, sb⇩h, m,𝒟⇩h, 𝒪⇩h, ℛ⇩h,𝒮⇩h) →⇩s⇩b⇩h
(is', θ(t ↦ v), sb⇩h@ [Read⇩s⇩b volatile a t v], m, 𝒟⇩h, 𝒪⇩h, ℛ⇩h, 𝒮⇩h)".
from sbh_computation.Memop [OF i_bound' ts⇩h_i this]
have step: "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h (?ts⇩h', m, 𝒮⇩h)".
from sb have sb: "sb = filter is_Write⇩s⇩b (sb⇩h @ [Read⇩s⇩b volatile a t v])"
by simp
show ?thesis
proof (cases "filter is_Write⇩s⇩b sb⇩h = []")
case False
have "ts [i := (p,is',θ(t ↦ v),sb,𝒟,𝒪,ℛ)] ∼⇩h ?ts⇩h'"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb sb_empty False
apply (auto simp add: Let_def nth_list_update)
done
with step show ?thesis
by (auto simp del: fun_upd_apply simp add: 𝒮' m' ts' 𝒪' θ' 𝒟' sb' ℛ')
next
case True
with sb_empty have empty: "sb⇩h=[]" by simp
from i_bound' have "?ts⇩h'!i = ?ts⇩h_i'"
by auto
from sbh_computation.StoreBuffer [OF _ this, simplified empty, simplified, OF _ flush_step.Read⇩s⇩b, of m 𝒮⇩h] i_bound'
have "(?ts⇩h', m, 𝒮⇩h)
⇒⇩s⇩b⇩h (ts⇩h[i := (p, is', θ(t ↦ v), [], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)], m, 𝒮⇩h)"
by (simp add: empty list_update_overwrite)
with step have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h⇧*
(ts⇩h[i := (p, is', θ(t ↦ v), [], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)], m,𝒮⇩h)"
by force
moreover
have "ts [i := (p,is',θ(t ↦ v),sb,𝒟,𝒪,ℛ)] ∼⇩h ts⇩h[i := (p, is', θ(t ↦ v), [], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)]"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb empty
apply (auto simp add: Let_def nth_list_update)
done
ultimately show ?thesis
by (auto simp del: fun_upd_apply simp add: 𝒮' m' ts' 𝒪' θ' 𝒟' sb' ℛ')
qed
next
case (SBReadUnbuffered a volatile t)
then obtain
"is": "is = Read volatile a t#is'" and
𝒪': "𝒪'=𝒪" and
ℛ': "ℛ'=ℛ" and
𝒮': "𝒮'=𝒮" and
𝒟': "𝒟'=𝒟" and
m': "m'=m" and
θ': "θ'=θ(t↦m a)" and
sb': "sb' = sb" and
buf: "buffered_val sb a = None"
by auto
from sim_loc [rule_format, OF i_bound] ts_i "is"
obtain sb⇩h 𝒪⇩h ℛ⇩h 𝒟⇩h where
ts⇩h_i: "ts⇩h!i = (p,Read volatile a t#is',θ,sb⇩h,𝒟⇩h,𝒪⇩h,ℛ⇩h)" and
sb: "sb = filter is_Write⇩s⇩b sb⇩h" and
sb_empty: "filter is_Write⇩s⇩b sb⇩h = [] ⟶ sb⇩h=[]"
by (auto simp add: Let_def)
from buf
have buf': "buffered_val sb⇩h a = None"
by (simp add: bufferd_val_filter_is_Write⇩s⇩b_eq sb)
let ?ts⇩h_i' = "(p, is', θ(t ↦ m a), sb⇩h @ [Read⇩s⇩b volatile a t (m a)], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)"
let ?ts⇩h' = "ts⇩h[i := ?ts⇩h_i']"
from sbh_memop_step.SBHReadUnbuffered [OF buf']
have "(Read volatile a t # is',θ, sb⇩h, m, 𝒟⇩h, 𝒪⇩h, ℛ⇩h,𝒮⇩h) →⇩s⇩b⇩h
(is', θ(t ↦ (m a)), sb⇩h@ [Read⇩s⇩b volatile a t (m a)], m,𝒟⇩h, 𝒪⇩h, ℛ⇩h,𝒮⇩h)".
from sbh_computation.Memop [OF i_bound' ts⇩h_i this]
have step: "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h
(?ts⇩h', m, 𝒮⇩h)".
moreover
from sb have sb: "sb = filter is_Write⇩s⇩b (sb⇩h @ [Read⇩s⇩b volatile a t (m a)])"
by simp
show ?thesis
proof (cases "filter is_Write⇩s⇩b sb⇩h = []")
case False
have "ts [i := (p,is',θ (t↦m a),sb,𝒟,𝒪,ℛ)] ∼⇩h ?ts⇩h'"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb sb_empty False
apply (auto simp add: Let_def nth_list_update)
done
with step show ?thesis
by (auto simp del: fun_upd_apply simp add: 𝒮' m' ts' 𝒪' ℛ' 𝒟' θ' sb')
next
case True
with sb_empty have empty: "sb⇩h=[]" by simp
from i_bound' have "?ts⇩h'!i = ?ts⇩h_i'"
by auto
from sbh_computation.StoreBuffer [OF _ this, simplified empty, simplified, OF _ flush_step.Read⇩s⇩b, of m 𝒮⇩h] i_bound'
have "(?ts⇩h', m, 𝒮⇩h)
⇒⇩s⇩b⇩h (ts⇩h[i := (p, is', θ(t ↦ (m a)), [], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)], m, 𝒮⇩h)"
by (simp add: empty)
with step have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h⇧*
(ts⇩h[i := (p, is', θ(t ↦ m a), [], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)], m, 𝒮⇩h)"
by force
moreover
have "ts [i := (p,is',θ(t ↦ m a),sb,𝒟,𝒪,ℛ)] ∼⇩h ts⇩h[i := (p, is', θ(t ↦ m a), [], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)]"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb empty
apply (auto simp add: Let_def nth_list_update)
done
ultimately show ?thesis
by (auto simp del: fun_upd_apply simp add: 𝒮' m' ts' 𝒪' θ' 𝒟' sb' ℛ')
qed
next
case (SBWriteNonVolatile a D f A L R W)
then obtain
"is": "is = Write False a (D, f) A L R W#is'" and
𝒪': "𝒪'=𝒪" and
ℛ': "ℛ'=ℛ" and
𝒮': "𝒮'=𝒮" and
𝒟': "𝒟'=𝒟" and
m': "m'=m" and
θ': "θ'=θ" and
sb': "sb' = sb@[Write⇩s⇩b False a (D, f) (f θ) A L R W]"
by auto
from sim_loc [rule_format, OF i_bound] ts_i
obtain sb⇩h 𝒪⇩h ℛ⇩h 𝒟⇩h where
ts⇩h_i: "ts⇩h!i = (p,Write False a (D,f) A L R W#is',θ,sb⇩h,𝒟⇩h,𝒪⇩h,ℛ⇩h)" and
sb: "sb = filter is_Write⇩s⇩b sb⇩h"
by (auto simp add: Let_def "is")
from sbh_memop_step.SBHWriteNonVolatile
have "(Write False a (D, f) A L R W# is',θ, sb⇩h, m, 𝒟⇩h, 𝒪⇩h, ℛ⇩h,𝒮⇩h) →⇩s⇩b⇩h
(is', θ, sb⇩h @ [Write⇩s⇩b False a (D, f) (f θ) A L R W], m,𝒟⇩h, 𝒪⇩h, ℛ⇩h,𝒮⇩h)".
from sbh_computation.Memop [OF i_bound' ts⇩h_i this]
have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h
(ts⇩h[i := (p, is',θ, sb⇩h @ [Write⇩s⇩b False a (D, f) (f θ) A L R W], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)],
m, 𝒮⇩h)".
moreover
have "ts [i := (p,is',θ,sb @ [Write⇩s⇩b False a (D,f) (f θ) A L R W],𝒟,𝒪,ℛ)] ∼⇩h
ts⇩h[i := (p,is',θ, sb⇩h @ [Write⇩s⇩b False a (D,f) (f θ) A L R W],𝒟⇩h, 𝒪⇩h,ℛ⇩h)]"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb
apply (auto simp add: Let_def nth_list_update)
done
ultimately show ?thesis
by (auto simp add: 𝒮' m' θ' 𝒪' ℛ' 𝒟' ts' sb')
next
case (SBWriteVolatile a D f A L R W)
then obtain
"is": "is = Write True a (D, f) A L R W#is'" and
𝒪': "𝒪'=𝒪" and
ℛ': "ℛ'=ℛ" and
𝒮': "𝒮'=𝒮" and
𝒟': "𝒟'=𝒟" and
m': "m'=m" and
θ': "θ'=θ" and
sb': "sb' = sb@[Write⇩s⇩b True a (D, f) (f θ) A L R W]"
by auto
from sim_loc [rule_format, OF i_bound] ts_i "is"
obtain sb⇩h 𝒪⇩h ℛ⇩h 𝒟⇩h where
ts⇩h_i: "ts⇩h!i = (p,Write True a (D,f) A L R W#is',θ,sb⇩h,𝒟⇩h,𝒪⇩h,ℛ⇩h)" and
sb: "sb = filter is_Write⇩s⇩b sb⇩h"
by (auto simp add: Let_def)
from sbh_computation.Memop [OF i_bound' ts⇩h_i SBHWriteVolatile
]
have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h
(ts⇩h[i := (p, is',θ, sb⇩h @ [Write⇩s⇩b True a (D, f) (f θ) A L R W], True, 𝒪⇩h,ℛ⇩h)],
m, 𝒮⇩h)".
moreover
have "ts [i := (p,is',θ,sb @ [Write⇩s⇩b True a (D,f) (f θ) A L R W],𝒟,𝒪,ℛ)] ∼⇩h
ts⇩h[i := (p,is', θ, sb⇩h @ [Write⇩s⇩b True a (D,f) (f θ) A L R W],True, 𝒪⇩h,ℛ⇩h)]"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb
apply (auto simp add: Let_def nth_list_update)
done
ultimately show ?thesis
by (auto simp add: ts' 𝒪' θ' m' sb' 𝒟' ℛ' 𝒮')
next
case SBFence
then obtain
"is": "is = Fence #is'" and
𝒪': "𝒪'=𝒪" and
ℛ': "ℛ'=ℛ" and
𝒮': "𝒮'=𝒮" and
𝒟': "𝒟'=𝒟" and
m': "m'=m" and
θ': "θ'=θ" and
sb: "sb = []" and
sb': "sb' = []"
by auto
from sim_loc [rule_format, OF i_bound] ts_i sb "is"
obtain sb⇩h 𝒪⇩h ℛ⇩h 𝒟⇩h where
ts⇩h_i: "ts⇩h!i = (p,Fence # is',θ,sb⇩h,𝒟⇩h,𝒪⇩h,ℛ⇩h)" and
sb: "[] = filter is_Write⇩s⇩b sb⇩h"
by (auto simp add: Let_def)
from filter_is_Write⇩s⇩b_empty [OF sb [symmetric]]
have "∀r ∈ set sb⇩h. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r".
from flush_reads_program [OF this] obtain 𝒪⇩h' 𝒮⇩h' ℛ⇩h'
where flsh: "(m, sb⇩h,𝒪⇩h,ℛ⇩h,𝒮⇩h) →⇩f⇧* (m, [],𝒪⇩h',ℛ⇩h',𝒮⇩h')" by blast
let ?ts⇩h' = "ts⇩h[i := (p,Fence # is', θ, [], 𝒟⇩h, 𝒪⇩h',ℛ⇩h')]"
from sbh_computation.store_buffer_steps [OF flsh i_bound' ts⇩h_i]
have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h⇧* (?ts⇩h', m, 𝒮⇩h')".
also
from i_bound' have i_bound'': "i < length ?ts⇩h'"
by auto
from i_bound' have ts⇩h'_i: "?ts⇩h'!i = (p,Fence#is',θ,[],𝒟⇩h,𝒪⇩h',ℛ⇩h')"
by simp
from sbh_computation.Memop [OF i_bound'' ts⇩h'_i SBHFence] i_bound'
have "(?ts⇩h', m, 𝒮⇩h') ⇒⇩s⇩b⇩h (ts⇩h[i := (p, is',θ, [], False, 𝒪⇩h',Map.empty)], m,𝒮⇩h')"
by (simp)
finally
have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h⇧* (ts⇩h[i := (p, is', θ, [],False, 𝒪⇩h',Map.empty)],m, 𝒮⇩h')".
moreover
have "ts [i := (p,is',θ,[],𝒟,𝒪,ℛ)] ∼⇩h ts⇩h[i := (p,is', θ, [],False, 𝒪⇩h',Map.empty)]"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb
apply (auto simp add: Let_def nth_list_update)
done
ultimately show ?thesis
by (auto simp add: ts' 𝒪' θ' m' sb' 𝒟' 𝒮' ℛ')
next
case (SBRMWReadOnly cond t a D f ret A L R W)
then obtain
"is": "is = RMW a t (D, f) cond ret A L R W#is'" and
𝒪': "𝒪'=𝒪" and
ℛ': "ℛ'=ℛ" and
𝒮': "𝒮'=𝒮" and
𝒟': "𝒟'=𝒟" and
m': "m'=m" and
θ': "θ'=θ(t ↦ m a)" and
sb: "sb=[]" and
sb': "sb' = []" and
cond: "¬ cond (θ(t ↦ m a))"
by auto
from sim_loc [rule_format, OF i_bound] ts_i sb "is"
obtain sb⇩h 𝒪⇩h ℛ⇩h 𝒟⇩h where
ts⇩h_i: "ts⇩h!i = (p,RMW a t (D, f) cond ret A L R W# is',θ,sb⇩h,𝒟⇩h,𝒪⇩h,ℛ⇩h)" and
sb: "[] = filter is_Write⇩s⇩b sb⇩h"
by (auto simp add: Let_def)
from filter_is_Write⇩s⇩b_empty [OF sb [symmetric]]
have "∀r ∈ set sb⇩h. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r".
from flush_reads_program [OF this] obtain 𝒪⇩h' 𝒮⇩h' ℛ⇩h'
where flsh: "(m, sb⇩h,𝒪⇩h,ℛ⇩h,𝒮⇩h) →⇩f⇧* (m, [],𝒪⇩h',ℛ⇩h',𝒮⇩h')" by blast
let ?ts⇩h' = "ts⇩h[i := (p,RMW a t (D, f) cond ret A L R W# is',θ, [], 𝒟⇩h, 𝒪⇩h',ℛ⇩h')]"
from sbh_computation.store_buffer_steps [OF flsh i_bound' ts⇩h_i]
have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h⇧* (?ts⇩h', m, 𝒮⇩h')".
also
from i_bound' have i_bound'': "i < length ?ts⇩h'"
by auto
from i_bound' have ts⇩h'_i: "?ts⇩h'!i = (p,RMW a t (D, f) cond ret A L R W#is',θ,[],𝒟⇩h,𝒪⇩h',ℛ⇩h')"
by simp
note step= SBHRMWReadOnly [where cond=cond and θ=θ and m=m, OF cond ]
from sbh_computation.Memop [OF i_bound'' ts⇩h'_i step ] i_bound'
have "(?ts⇩h', m, 𝒮⇩h') ⇒⇩s⇩b⇩h (ts⇩h[i := (p, is',θ(t↦m a), [], False, 𝒪⇩h',Map.empty)],m, 𝒮⇩h')"
by (simp)
finally
have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h⇧* (ts⇩h[i := (p, is',θ(t↦m a), [], False, 𝒪⇩h',Map.empty)],m, 𝒮⇩h')".
moreover
have "ts [i := (p,is',θ(t↦m a),[],𝒟,𝒪,ℛ)] ∼⇩h ts⇩h[i := (p,is', θ(t↦m a), [], False, 𝒪⇩h',Map.empty)]"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb
apply (auto simp add: Let_def nth_list_update)
done
ultimately show ?thesis
by (auto simp add: ts' 𝒪' θ' m' sb' 𝒟' 𝒮' ℛ')
next
case (SBRMWWrite cond t a D f ret A L R W)
then obtain
"is": "is = RMW a t (D, f) cond ret A L R W#is'" and
𝒪': "𝒪'=𝒪" and
ℛ': "ℛ'=ℛ" and
𝒮': "𝒮'=𝒮" and
𝒟': "𝒟'=𝒟" and
m': "m'=m(a := f (θ(t ↦ (m a))))" and
θ': "θ'=θ(t ↦ ret (m a) (f (θ(t ↦ (m a)))))" and
sb: "sb=[]" and
sb': "sb' = []" and
cond: "cond (θ(t ↦ m a))"
by auto
from sim_loc [rule_format, OF i_bound] ts_i sb "is"
obtain sb⇩h 𝒪⇩h ℛ⇩h 𝒟⇩h acq⇩h where
ts⇩h_i: "ts⇩h!i = (p,RMW a t (D, f) cond ret A L R W# is',θ,sb⇩h,𝒟⇩h,𝒪⇩h,ℛ⇩h)" and
sb: "[] = filter is_Write⇩s⇩b sb⇩h"
by (auto simp add: Let_def)
from filter_is_Write⇩s⇩b_empty [OF sb [symmetric]]
have "∀r ∈ set sb⇩h. is_Read⇩s⇩b r ∨ is_Prog⇩s⇩b r ∨ is_Ghost⇩s⇩b r".
from flush_reads_program [OF this] obtain 𝒪⇩h' 𝒮⇩h' ℛ⇩h'
where flsh: "(m, sb⇩h,𝒪⇩h,ℛ⇩h,𝒮⇩h) →⇩f⇧* (m, [],𝒪⇩h',ℛ⇩h',𝒮⇩h')" by blast
let ?ts⇩h' = "ts⇩h[i := (p,RMW a t (D, f) cond ret A L R W# is',θ, [], 𝒟⇩h, 𝒪⇩h',ℛ⇩h')]"
from sbh_computation.store_buffer_steps [OF flsh i_bound' ts⇩h_i]
have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h⇧* (?ts⇩h', m, 𝒮⇩h')".
also
from i_bound' have i_bound'': "i < length ?ts⇩h'"
by auto
from i_bound' have ts⇩h'_i: "?ts⇩h'!i = (p,RMW a t (D, f) cond ret A L R W#is',θ,[],𝒟⇩h,𝒪⇩h',ℛ⇩h')"
by simp
note step= SBHRMWWrite [where cond=cond and θ=θ and m=m, OF cond]
from sbh_computation.Memop [OF i_bound'' ts⇩h'_i step ] i_bound'
have "(?ts⇩h', m, 𝒮⇩h') ⇒⇩s⇩b⇩h (ts⇩h[i := (p, is',
θ(t ↦ ret (m a) (f (θ(t ↦ (m a))))), [], False, 𝒪⇩h' ∪ A - R,Map.empty)],
m(a := f (θ(t ↦ (m a)))),𝒮⇩h' ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (simp)
finally
have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h⇧* (ts⇩h[i := (p, is',
θ(t ↦ ret (m a) (f (θ(t ↦ (m a))))), [], False, 𝒪⇩h' ∪ A - R,Map.empty)],
m(a := f (θ(t ↦ (m a)))),𝒮⇩h' ⊕⇘W⇙ R ⊖⇘A⇙ L)".
moreover
have "ts [i := (p,is',θ(t ↦ ret (m a) (f (θ(t ↦ (m a))))),[],𝒟,𝒪,ℛ)] ∼⇩h
ts⇩h[i := (p,is',θ(t ↦ ret (m a) (f (θ(t ↦ (m a))))), [],False, 𝒪⇩h' ∪ A - R,Map.empty)]"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb
apply (auto simp add: Let_def nth_list_update)
done
ultimately show ?thesis
by (auto simp add: ts' 𝒪' θ' m' sb' 𝒟' 𝒮' ℛ')
next
case (SBGhost A L R W)
then obtain
"is": "is = Ghost A L R W#is'" and
𝒪': "𝒪'=𝒪" and
ℛ': "ℛ'=ℛ" and
𝒮': "𝒮'=𝒮" and
𝒟': "𝒟'=𝒟" and
m': "m'=m" and
θ': "θ'=θ" and
sb': "sb' = sb"
by auto
from sim_loc [rule_format, OF i_bound] ts_i "is"
obtain sb⇩h 𝒪⇩h ℛ⇩h 𝒟⇩h where
ts⇩h_i: "ts⇩h!i = (p,Ghost A L R W# is',θ,sb⇩h,𝒟⇩h,𝒪⇩h,ℛ⇩h)" and
sb: "sb = filter is_Write⇩s⇩b sb⇩h" and
sb_empty: "filter is_Write⇩s⇩b sb⇩h = [] ⟶ sb⇩h=[]"
by (auto simp add: Let_def)
let ?ts⇩h_i' = "(p, is', θ, sb⇩h@[Ghost⇩s⇩b A L R W],𝒟⇩h, 𝒪⇩h,ℛ⇩h)"
let ?ts⇩h' = "ts⇩h[i := ?ts⇩h_i']"
note step= SBHGhost
from sbh_computation.Memop [OF i_bound' ts⇩h_i step ] i_bound'
have step: "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h (?ts⇩h',m, 𝒮⇩h)"
by (simp)
from sb have sb: "sb = filter is_Write⇩s⇩b (sb⇩h @ [Ghost⇩s⇩b A L R W])"
by simp
show ?thesis
proof (cases "filter is_Write⇩s⇩b sb⇩h = []")
case False
have "ts [i := (p,is',θ,sb,𝒟,𝒪,ℛ)] ∼⇩h ?ts⇩h'"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb sb_empty False
apply (auto simp add: Let_def nth_list_update)
done
with step show ?thesis
by (auto simp del: fun_upd_apply simp add: 𝒮' m' ts' 𝒪' 𝒟' θ' sb' ℛ')
next
case True
with sb_empty have empty: "sb⇩h=[]" by simp
from i_bound' have "?ts⇩h'!i = ?ts⇩h_i'"
by auto
from sbh_computation.StoreBuffer [OF _ this, simplified empty, simplified, OF _ flush_step.Ghost, of m 𝒮⇩h] i_bound'
have "(?ts⇩h', m, 𝒮⇩h)
⇒⇩s⇩b⇩h (ts⇩h[i := (p, is', θ, [], 𝒟⇩h, 𝒪⇩h ∪ A - R,augment_rels (dom 𝒮⇩h) R ℛ⇩h)], m, 𝒮⇩h ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by (simp add: empty)
with step have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h⇧*
(ts⇩h[i := (p, is', θ, [], 𝒟⇩h, 𝒪⇩h ∪ A - R,augment_rels (dom 𝒮⇩h) R ℛ⇩h)], m, 𝒮⇩h ⊕⇘W⇙ R ⊖⇘A⇙ L)"
by force
moreover
have "ts [i := (p,is',θ,sb,𝒟,𝒪,ℛ)] ∼⇩h
ts⇩h[i := (p, is', θ, [], 𝒟⇩h, 𝒪⇩h ∪ A - R,augment_rels (dom 𝒮⇩h) R ℛ⇩h)]"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb empty
apply (auto simp add: Let_def nth_list_update)
done
ultimately show ?thesis
by (auto simp del: fun_upd_apply simp add: 𝒮' m' ts' 𝒪' θ' 𝒟' sb' ℛ')
qed
qed
next
case (Program i _ p "is" θ sb 𝒟 𝒪 ℛ p' "is'")
then obtain
ts': "ts' = ts[i := (p', is@is',θ, sb, 𝒟, 𝒪,ℛ)]" and
i_bound: "i < length ts" and
ts_i: "ts ! i = (p, is, θ,sb,𝒟, 𝒪,ℛ)" and
prog_step: "θ⊢ p →⇩p (p', is')" and
𝒮': "𝒮'=𝒮" and
m': "m'=m"
by auto
from sim obtain
lts_eq: "length ts = length ts⇩h" and
sim_loc: "∀i < length ts. (∃𝒪' 𝒟' ℛ'.
let (p,is,θ, sb, 𝒟, 𝒪,ℛ) = ts⇩h!i in ts!i=(p,is, θ, filter is_Write⇩s⇩b sb,𝒟',𝒪',ℛ') ∧
(filter is_Write⇩s⇩b sb = [] ⟶ sb = []))"
by cases auto
from sim_loc [rule_format, OF i_bound] ts_i
obtain sb⇩h 𝒪⇩h ℛ⇩h 𝒟⇩h acq⇩h where
ts⇩h_i: "ts⇩h!i = (p,is,θ,sb⇩h,𝒟⇩h,𝒪⇩h,ℛ⇩h)" and
sb: "sb = filter is_Write⇩s⇩b sb⇩h" and
sb_empty: "filter is_Write⇩s⇩b sb⇩h = [] ⟶ sb⇩h=[]"
by (auto simp add: Let_def)
from lts_eq i_bound have i_bound': "i < length ts⇩h"
by simp
let ?ts⇩h_i' = "(p', is @ is',θ, sb⇩h @ [Prog⇩s⇩b p p' is'], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)"
let ?ts⇩h' = "ts⇩h[i := ?ts⇩h_i']"
from sbh_computation.Program [OF i_bound' ts⇩h_i prog_step]
have step: "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h (?ts⇩h',m, 𝒮⇩h)".
show ?thesis
proof (cases "filter is_Write⇩s⇩b sb⇩h = []")
case False
have "ts[i := (p', is@is', θ, sb,𝒟, 𝒪,ℛ)] ∼⇩h ?ts⇩h'"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb False sb_empty
apply (auto simp add: Let_def nth_list_update)
done
with step show ?thesis
by (auto simp add: ts' 𝒮' m')
next
case True
with sb_empty have empty: "sb⇩h=[]" by simp
from i_bound' have "?ts⇩h'!i = ?ts⇩h_i'"
by auto
from sbh_computation.StoreBuffer [OF _ this, simplified empty, simplified, OF _ flush_step.Prog⇩s⇩b, of m 𝒮⇩h] i_bound'
have "(?ts⇩h', m, 𝒮⇩h)
⇒⇩s⇩b⇩h (ts⇩h[i := (p', is@is', θ, [], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)], m, 𝒮⇩h)"
by (simp add: empty)
with step have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h⇧*
(ts⇩h[i := (p', is@is', θ, [], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)], m, 𝒮⇩h) "
by force
moreover
have "ts[i := (p', is@is', θ, sb,𝒟, 𝒪,ℛ)] ∼⇩h ts⇩h[i := (p', is@is', θ, [], 𝒟⇩h, 𝒪⇩h,ℛ⇩h)]"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb empty
apply (auto simp add: Let_def nth_list_update)
done
ultimately show ?thesis
by (auto simp del: fun_upd_apply simp add: 𝒮' m' ts')
qed
next
case (StoreBuffer i _ p "is" θ sb 𝒟 𝒪 ℛ _ _ _ sb' 𝒪' ℛ')
then obtain
ts': "ts' = ts[i := (p, is,θ, sb', 𝒟, 𝒪',ℛ')]" and
i_bound: "i < length ts" and
ts_i: "ts ! i = (p, is,θ,sb, 𝒟, 𝒪,ℛ)" and
sb_step: "(m,sb,𝒪,ℛ,𝒮) →⇩w (m', sb',𝒪',ℛ',𝒮')"
by auto
from sim obtain
lts_eq: "length ts = length ts⇩h" and
sim_loc: "∀i < length ts. (∃𝒪' 𝒟' ℛ'.
let (p,is, θ, sb,𝒟, 𝒪,ℛ) = ts⇩h!i in ts!i=(p,is, θ, filter is_Write⇩s⇩b sb,𝒟',𝒪',ℛ') ∧
(filter is_Write⇩s⇩b sb = [] ⟶ sb=[]))"
by cases auto
from sim_loc [rule_format, OF i_bound] ts_i
obtain sb⇩h 𝒪⇩h ℛ⇩h 𝒟⇩h acq⇩h where
ts⇩h_i: "ts⇩h!i = (p,is,θ,sb⇩h,𝒟⇩h,𝒪⇩h,ℛ⇩h)" and
sb: "sb = filter is_Write⇩s⇩b sb⇩h" and
sb_empty: "filter is_Write⇩s⇩b sb⇩h = [] ⟶ sb⇩h=[]"
by (auto simp add: Let_def)
from lts_eq i_bound have i_bound': "i < length ts⇩h"
by simp
from flush_simulates_filter_writes [OF sb_step sb, of 𝒪⇩h ℛ⇩h 𝒮⇩h]
obtain sb⇩h' 𝒪⇩h' ℛ⇩h' 𝒮⇩h'
where flush': "(m, sb⇩h,𝒪⇩h,ℛ⇩h,𝒮⇩h) →⇩f⇧* (m', sb⇩h',𝒪⇩h',ℛ⇩h',𝒮⇩h')" and
sb': "sb' = filter is_Write⇩s⇩b sb⇩h'" and
sb'_empty: "filter is_Write⇩s⇩b sb⇩h' = [] ⟶ sb⇩h'=[]"
by auto
from sb_step obtain volatile a sop v A L R W where "sb=Write⇩s⇩b volatile a sop v A L R W#sb'"
by cases force
from sbh_computation.store_buffer_steps [OF flush' i_bound' ts⇩h_i]
have "(ts⇩h, m, 𝒮⇩h) ⇒⇩s⇩b⇩h⇧* (ts⇩h[i := (p, is, θ, sb⇩h',𝒟⇩h, 𝒪⇩h',ℛ⇩h')], m', 𝒮⇩h')".
moreover
have "ts[i := (p, is, θ, sb',𝒟, 𝒪',ℛ')] ∼⇩h
ts⇩h[i := (p, is, θ, sb⇩h',𝒟⇩h, 𝒪⇩h',ℛ⇩h')]"
apply (rule sim_history_config.intros)
using lts_eq
apply simp
using sim_loc i_bound i_bound' sb sb' sb'_empty
apply (auto simp add: Let_def nth_list_update)
done
ultimately show ?thesis
by (auto simp add: ts' )
qed
qed
theorem (in valid_program) concurrent_history_steps_simulates_store_buffer_steps:
assumes step_sb: "(ts,m,𝒮) ⇒⇩s⇩b⇧* (ts',m',𝒮')"
shows "⋀ts⇩h 𝒮⇩h. ts ∼⇩h ts⇩h ⟹ ∃ts⇩h' 𝒮⇩h'. (ts⇩h,m,𝒮⇩h) ⇒⇩s⇩b⇩h⇧* (ts⇩h',m',𝒮⇩h') ∧ ts' ∼⇩h ts⇩h'"
using step_sb
proof (induct rule: converse_rtranclp_induct_sbh_steps)
case refl thus ?case by auto
next
case (step ts m 𝒮 ts'' m'' 𝒮'' )
have first: "(ts,m,𝒮) ⇒⇩s⇩b (ts'',m'',𝒮'')" by fact
have sim: "ts ∼⇩h ts⇩h" by fact
from concurrent_history_steps_simulates_store_buffer_step [OF first sim, of 𝒮⇩h]
obtain ts⇩h'' 𝒮⇩h'' where
exec: "(ts⇩h,m,𝒮⇩h) ⇒⇩s⇩b⇩h⇧* (ts⇩h'',m'',𝒮⇩h'')" and sim: "ts'' ∼⇩h ts⇩h''"
by auto
from step.hyps (3) [OF sim, of 𝒮⇩h'']
obtain ts⇩h' 𝒮⇩h' where exec_rest: "(ts⇩h'',m'',𝒮⇩h'') ⇒⇩s⇩b⇩h⇧* (ts⇩h',m',𝒮⇩h')" and sim': "ts' ∼⇩h ts⇩h'"
by auto
note exec also note exec_rest
finally show ?case
using sim' by blast
qed
theorem (in xvalid_program_progress) concurrent_direct_execution_simulates_store_buffer_execution:
assumes exec_sb: "(ts⇩s⇩b,m⇩s⇩b,x) ⇒⇩s⇩b⇧* (ts⇩s⇩b',m⇩s⇩b',x')"
assumes init: "initial⇩s⇩b ts⇩s⇩b 𝒮⇩s⇩b"
assumes valid: "valid ts⇩s⇩b"
assumes sim: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m,𝒮)"
assumes safe: "safe_reach_direct safe_free_flowing (ts,m,𝒮)"
shows "∃ts⇩h' 𝒮⇩h' ts' m' 𝒮'.
(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ⇒⇩s⇩b⇩h⇧* (ts⇩h',m⇩s⇩b',𝒮⇩h') ∧
ts⇩s⇩b' ∼⇩h ts⇩h' ∧
(ts,m,𝒮) ⇒⇩d⇧* (ts',m',𝒮') ∧
(ts⇩h',m⇩s⇩b',𝒮⇩h') ∼ (ts',m',𝒮')"
proof -
from init interpret ini: initial⇩s⇩b ts⇩s⇩b 𝒮⇩s⇩b .
from concurrent_history_steps_simulates_store_buffer_steps [OF exec_sb ini.history_refl, of 𝒮⇩s⇩b]
obtain ts⇩h' 𝒮⇩h' where
sbh: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ⇒⇩s⇩b⇩h⇧* (ts⇩h',m⇩s⇩b',𝒮⇩h')" and
sim_sbh: "ts⇩s⇩b' ∼⇩h ts⇩h'"
by auto
from concurrent_direct_execution_simulates_store_buffer_history_execution [OF sbh init valid sim safe]
obtain ts' m' 𝒮' where
d: "(ts,m,𝒮) ⇒⇩d⇧* (ts',m',𝒮')" and
d_sim: "(ts⇩h',m⇩s⇩b',𝒮⇩h') ∼ (ts',m',𝒮')"
by auto
with sbh sim_sbh show ?thesis by blast
qed
inductive sim_direct_config::
"('p,'p store_buffer,'dirty,'owns,'rels) thread_config list ⇒ ('p,unit,bool,'owns','rels') thread_config list ⇒ bool"
(‹_ ∼⇩d _ › [60,60] 100)
where
"⟦length ts = length ts⇩d;
∀i < length ts.
(∃𝒪' 𝒟' ℛ'.
let (p,is, θ,sb,𝒟, 𝒪,ℛ) = ts⇩d!i in
ts!i=(p,is, θ, [] ,𝒟',𝒪',ℛ'))
⟧
⟹
ts ∼⇩d ts⇩d"
lemma empty_sb_sims:
assumes empty:
"∀i p is xs sb 𝒟 𝒪 ℛ. i < length ts⇩s⇩b ⟶ ts⇩s⇩b!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟶ sb=[]"
assumes sim_h: "ts⇩s⇩b ∼⇩h ts⇩h"
assumes sim_d: "(ts⇩h,m⇩h,𝒮⇩h) ∼ (ts,m,𝒮)"
shows "ts⇩s⇩b ∼⇩d ts ∧ m⇩h=m ∧ length ts⇩s⇩b = length ts"
proof-
from sim_h empty
have empty':
"∀i p is xs sb 𝒟 𝒪 ℛ. i < length ts⇩h ⟶ ts⇩h!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟶ sb=[]"
apply (cases)
apply clarsimp
subgoal for i
apply (drule_tac x=i in spec)
apply (auto simp add: Let_def)
done
done
from sim_h sim_config_emptyE [OF empty' sim_d]
show ?thesis
apply cases
apply clarsimp
apply (rule sim_direct_config.intros)
apply clarsimp
apply clarsimp
using empty'
subgoal for i
apply (drule_tac x=i in spec)
apply (drule_tac x=i in spec)
apply (drule_tac x=i in spec)
apply (auto simp add: Let_def)
done
done
qed
lemma empty_d_sims:
assumes sim: "ts⇩s⇩b ∼⇩d ts"
shows "∃ts⇩h. ts⇩s⇩b ∼⇩h ts⇩h ∧ (ts⇩h,m,𝒮) ∼ (ts,m,𝒮)"
proof -
from sim obtain
leq: "length ts⇩s⇩b = length ts" and
sim: "∀i < length ts⇩s⇩b.
(∃𝒪' 𝒟' ℛ'.
let (p,is, θ,sb,𝒟, 𝒪,ℛ) = ts!i in
ts⇩s⇩b!i=(p,is, θ, [] ,𝒟',𝒪',ℛ'))"
by cases auto
define ts⇩h where "ts⇩h ≡ map (λ(p,is, θ,sb,𝒟, 𝒪,ℛ). (p,is, θ,[]::'a memref list,𝒟, 𝒪,ℛ)) ts"
have "ts⇩s⇩b ∼⇩h ts⇩h"
apply (rule sim_history_config.intros)
using leq sim
apply (auto simp add: ts⇩h_def Let_def leq)
done
moreover
have empty:
"∀i p is xs sb 𝒟 𝒪 ℛ. i < length ts⇩h ⟶ ts⇩h!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟶ sb=[]"
apply (clarsimp simp add: ts⇩h_def Let_def)
subgoal for i
apply (case_tac "ts!i")
apply auto
done
done
have "(ts⇩h,m,𝒮) ∼ (ts,m,𝒮)"
apply (rule sim_config_emptyI [OF empty])
apply (clarsimp simp add: ts⇩h_def)
apply (clarsimp simp add: ts⇩h_def Let_def)
subgoal for i
apply (case_tac "ts!i")
apply auto
done
done
ultimately show ?thesis by blast
qed
theorem (in xvalid_program_progress) concurrent_direct_execution_simulates_store_buffer_execution_empty:
assumes exec_sb: "(ts⇩s⇩b,m⇩s⇩b,x) ⇒⇩s⇩b⇧* (ts⇩s⇩b',m⇩s⇩b',x')"
assumes init: "initial⇩s⇩b ts⇩s⇩b 𝒮⇩s⇩b"
assumes valid: "valid ts⇩s⇩b"
assumes empty:
"∀i p is xs sb 𝒟 𝒪 ℛ. i < length ts⇩s⇩b' ⟶ ts⇩s⇩b'!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟶ sb=[]"
assumes sim: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m,𝒮)"
assumes safe: "safe_reach_direct safe_free_flowing (ts,m,𝒮)"
shows "∃ts' 𝒮'.
(ts,m,𝒮) ⇒⇩d⇧* (ts',m⇩s⇩b',𝒮') ∧ ts⇩s⇩b' ∼⇩d ts'"
proof -
from concurrent_direct_execution_simulates_store_buffer_execution [OF exec_sb init valid sim safe]
obtain ts⇩h' 𝒮⇩h' ts' m' 𝒮' where
"(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ⇒⇩s⇩b⇩h⇧* (ts⇩h',m⇩s⇩b',𝒮⇩h')" and
sim_h: "ts⇩s⇩b' ∼⇩h ts⇩h'" and
exec: "(ts,m,𝒮) ⇒⇩d⇧* (ts',m',𝒮')" and
sim: "(ts⇩h',m⇩s⇩b',𝒮⇩h') ∼ (ts',m',𝒮')"
by auto
from empty_sb_sims [OF empty sim_h sim]
obtain "ts⇩s⇩b' ∼⇩d ts'" "m⇩s⇩b' = m'" "length ts⇩s⇩b' = length ts'"
by auto
thus ?thesis
using exec
by blast
qed
locale initial⇩d = simple_ownership_distinct + read_only_unowned + unowned_shared +
fixes valid
assumes empty_is: "⟦i < length ts; ts!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟧ ⟹ is=[]"
assumes empty_rels: "⟦i < length ts; ts!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟧ ⟹ ℛ=Map.empty"
assumes valid_init: "valid (map (λ(p,is, θ,sb,𝒟, 𝒪,ℛ). (p,is, θ,[],𝒟, 𝒪,ℛ)) ts)"
locale empty_store_buffers =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes empty_sb: "⟦i < length ts; ts!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟧ ⟹ sb=[]"
lemma initial_d_sb:
assumes init: "initial⇩d ts 𝒮 valid"
shows "initial⇩s⇩b (map (λ(p,is, θ,sb,𝒟, 𝒪,ℛ). (p,is, θ,[],𝒟, 𝒪,ℛ)) ts) 𝒮"
(is "initial⇩s⇩b ?map 𝒮")
proof -
from init interpret ini: initial⇩d ts 𝒮 .
show ?thesis
proof (intro_locales)
show "simple_ownership_distinct ?map"
apply (clarsimp simp add: simple_ownership_distinct_def)
subgoal for i j
apply (case_tac "ts!i")
apply (case_tac "ts!j")
apply (cut_tac i=i and j=j in ini.simple_ownership_distinct)
apply clarsimp
apply clarsimp
apply clarsimp
apply assumption
apply assumption
apply auto
done
done
next
show "read_only_unowned 𝒮 ?map"
apply (clarsimp simp add: read_only_unowned_def)
subgoal for i
apply (case_tac "ts!i")
apply (cut_tac i=i in ini.read_only_unowned)
apply clarsimp
apply assumption
apply auto
done
done
next
show "unowned_shared 𝒮 ?map"
apply (clarsimp simp add: unowned_shared_def')
apply (rule ini.unowned_shared')
apply clarsimp
subgoal for a i
apply (case_tac "ts!i")
apply auto
done
done
next
show "initial⇩s⇩b_axioms ?map"
apply (unfold_locales)
subgoal for i
apply (case_tac "ts!i")
apply simp
done
subgoal for i
apply (case_tac "ts!i")
apply clarsimp
apply (rule_tac i=i in ini.empty_is)
apply clarsimp
apply fastforce
done
subgoal for i
apply (case_tac "ts!i")
apply clarsimp
apply (rule_tac i=i in ini.empty_rels)
apply clarsimp
apply fastforce
done
done
qed
qed
theorem (in xvalid_program_progress) store_buffer_execution_result_sequential_consistent:
assumes exec_sb: "(ts⇩s⇩b,m,x) ⇒⇩s⇩b⇧* (ts⇩s⇩b',m',x')"
assumes empty': "empty_store_buffers ts⇩s⇩b'"
assumes sim: "ts⇩s⇩b ∼⇩d ts"
assumes init: "initial⇩d ts 𝒮 valid"
assumes safe: "safe_reach_direct safe_free_flowing (ts,m,𝒮)"
shows "∃ts' 𝒮'.
(ts,m,𝒮) ⇒⇩d⇧* (ts',m',𝒮') ∧ ts⇩s⇩b' ∼⇩d ts'"
proof -
from empty'
have empty':
"∀i p is xs sb 𝒟 𝒪 ℛ. i < length ts⇩s⇩b' ⟶ ts⇩s⇩b'!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟶ sb=[]"
by (auto simp add: empty_store_buffers_def)
define ts⇩h where "ts⇩h ≡ map (λ(p,is, θ,sb,𝒟, 𝒪,ℛ). (p,is, θ,[]::'a memref list,𝒟, 𝒪,ℛ)) ts"
from initial_d_sb [OF init]
have init_h:"initial⇩s⇩b ts⇩h 𝒮"
by (simp add: ts⇩h_def)
from initial⇩d.valid_init [OF init]
have valid_h: "valid ts⇩h"
by (simp add: ts⇩h_def)
from sim obtain
leq: "length ts⇩s⇩b = length ts" and
sim: "∀i < length ts⇩s⇩b.
(∃𝒪' 𝒟' ℛ'.
let (p,is, θ,sb,𝒟, 𝒪,ℛ) = ts!i in
ts⇩s⇩b!i=(p,is, θ, [] ,𝒟',𝒪',ℛ'))"
by cases auto
have sim_h: "ts⇩s⇩b ∼⇩h ts⇩h"
apply (rule sim_history_config.intros)
using leq sim
apply (auto simp add: ts⇩h_def Let_def leq)
done
from concurrent_history_steps_simulates_store_buffer_steps [OF exec_sb sim_h, of 𝒮]
obtain ts⇩h' 𝒮⇩h' where steps_h: "(ts⇩h,m,𝒮) ⇒⇩s⇩b⇩h⇧* (ts⇩h',m',𝒮⇩h')" and
sim_h': "ts⇩s⇩b' ∼⇩h ts⇩h'"
by auto
moreover
have empty:
"∀i p is xs sb 𝒟 𝒪 ℛ. i < length ts⇩h ⟶ ts⇩h!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟶ sb=[]"
apply (clarsimp simp add: ts⇩h_def Let_def)
subgoal for i
apply (case_tac "ts!i")
apply auto
done
done
have sim': "(ts⇩h,m,𝒮) ∼ (ts,m,𝒮)"
apply (rule sim_config_emptyI [OF empty])
apply (clarsimp simp add: ts⇩h_def)
apply (clarsimp simp add: ts⇩h_def Let_def)
subgoal for i
apply (case_tac "ts!i")
apply auto
done
done
from concurrent_direct_execution_simulates_store_buffer_history_execution [OF steps_h init_h valid_h sim' safe]
obtain ts' m'' 𝒮'' where steps: "(ts, m, 𝒮) ⇒⇩d⇧* (ts', m'', 𝒮'')"
and sim': "(ts⇩h', m', 𝒮⇩h') ∼ (ts', m'', 𝒮'')"
by blast
from empty_sb_sims [OF empty' sim_h' sim'] steps
show ?thesis
by auto
qed
locale initial⇩v = simple_ownership_distinct + read_only_unowned + unowned_shared +
fixes valid
assumes empty_is: "⟦i < length ts; ts!i=(p,is,xs,sb,𝒟,𝒪,ℛ)⟧ ⟹ is=[]"
assumes valid_init: "valid (map (λ(p,is, θ,sb,𝒟, 𝒪,ℛ). (p,is, θ,[],𝒟, 𝒪,Map.empty)) ts)"
theorem (in xvalid_program_progress) store_buffer_execution_result_sequential_consistent':
assumes exec_sb: "(ts⇩s⇩b,m,x) ⇒⇩s⇩b⇧* (ts⇩s⇩b',m',x')"
assumes empty': "empty_store_buffers ts⇩s⇩b'"
assumes sim: "ts⇩s⇩b ∼⇩d ts"
assumes init: "initial⇩v ts 𝒮 valid"
assumes safe: "safe_reach_virtual safe_free_flowing (ts,m,𝒮)"
shows "∃ts' 𝒮'.
(ts,m,𝒮) ⇒⇩v⇧* (ts',m',𝒮') ∧ ts⇩s⇩b' ∼⇩d ts'"
proof -
define ts⇩d where "ts⇩d == (map (λ(p,is, θ,sb,𝒟, 𝒪,ℛ'). (p,is, θ,sb,𝒟, 𝒪,Map.empty::rels)) ts)"
have rem_ts: "remove_rels ts⇩d = ts"
apply (rule nth_equalityI)
apply (simp add: ts⇩d_def remove_rels_def)
apply (clarsimp simp add: ts⇩d_def remove_rels_def)
subgoal for i
apply (case_tac "ts!i")
apply clarsimp
done
done
from sim
have sim': "ts⇩s⇩b ∼⇩d ts⇩d"
apply cases
apply (rule sim_direct_config.intros)
apply (auto simp add: ts⇩d_def)
done
have init': "initial⇩d ts⇩d 𝒮 valid"
proof (intro_locales)
from init have "simple_ownership_distinct ts"
by (simp add: initial⇩v_def)
then
show "simple_ownership_distinct ts⇩d"
apply (clarsimp simp add: ts⇩d_def simple_ownership_distinct_def)
subgoal for i j
apply (case_tac "ts!i")
apply (case_tac "ts!j")
apply force
done
done
next
from init have "read_only_unowned 𝒮 ts"
by (simp add: initial⇩v_def)
then show "read_only_unowned 𝒮 ts⇩d"
apply (clarsimp simp add: ts⇩d_def read_only_unowned_def)
subgoal for i
apply (case_tac "ts!i")
apply force
done
done
next
from init have "unowned_shared 𝒮 ts"
by (simp add: initial⇩v_def)
then
show "unowned_shared 𝒮 ts⇩d"
apply (clarsimp simp add: ts⇩d_def unowned_shared_def)
apply force
done
next
have eq: "((λ(p, is, θ, sb, 𝒟, 𝒪, ℛ). (p, is, θ, [], 𝒟, 𝒪, ℛ)) ∘
(λ(p, is, θ, sb, 𝒟, 𝒪, ℛ'). (p, is, θ, (), 𝒟, 𝒪, Map.empty)))
= (λ(p, is, θ, sb, 𝒟, 𝒪, ℛ'). (p, is, θ, [], 𝒟, 𝒪, Map.empty))"
apply (rule ext)
subgoal for x
apply (case_tac x)
apply auto
done
done
from init have "initial⇩v_axioms ts valid"
by (simp add: initial⇩v_def)
then
show "initial⇩d_axioms ts⇩d valid"
apply (clarsimp simp add: ts⇩d_def initial⇩v_axioms_def initial⇩d_axioms_def eq)
apply (rule conjI)
apply clarsimp
subgoal for i
apply (case_tac "ts!i")
apply force
done
apply clarsimp
subgoal for i
apply (case_tac "ts!i")
apply force
done
done
qed
{
fix ts⇩d' m' 𝒮'
assume exec: "(ts⇩d, m, 𝒮) ⇒⇩d⇧* (ts⇩d', m', 𝒮')"
have "safe_free_flowing (ts⇩d', m', 𝒮')"
proof -
from virtual_simulates_direct_steps [OF exec]
have exec_v: "(ts, m, 𝒮) ⇒⇩v⇧* (remove_rels ts⇩d', m', 𝒮')"
by (simp add: rem_ts)
have eq: "map (owned ∘
(λ(p, is, θ, sb, 𝒟, 𝒪, ℛ). (p, is, θ, (), 𝒟, 𝒪, ())))
ts⇩d' = map owned ts⇩d'"
by auto
from exec_v safe
have "safe_free_flowing (remove_rels ts⇩d', m', 𝒮')"
by (auto simp add: safe_reach_def)
then show ?thesis
by (auto simp add: safe_free_flowing_def remove_rels_def owned_def eq)
qed
}
hence safe': "safe_reach_direct safe_free_flowing (ts⇩d, m, 𝒮)"
by (simp add: safe_reach_def)
from store_buffer_execution_result_sequential_consistent [OF exec_sb empty' sim' init' safe']
obtain ts⇩d' 𝒮' where
exec_d: "(ts⇩d, m, 𝒮) ⇒⇩d⇧* (ts⇩d', m', 𝒮')" and sim_d: "ts⇩s⇩b' ∼⇩d ts⇩d'"
by blast
from virtual_simulates_direct_steps [OF exec_d]
have "(ts, m, 𝒮) ⇒⇩v⇧* (remove_rels ts⇩d', m', 𝒮')"
by (simp add: rem_ts)
moreover
from sim_d
have "ts⇩s⇩b' ∼⇩d remove_rels ts⇩d'"
apply (cases)
apply (rule sim_direct_config.intros)
apply (auto simp add: remove_rels_def)
done
ultimately show ?thesis
by auto
qed
subsection ‹Plug Together the Two Simulations›
corollary (in xvalid_program) concurrent_direct_steps_simulates_store_buffer_step:
assumes step_sb: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ⇒⇩s⇩b (ts⇩s⇩b',m⇩s⇩b',𝒮⇩s⇩b')"
assumes sim_h: "ts⇩s⇩b ∼⇩h ts⇩s⇩b⇩h"
assumes sim: "(ts⇩s⇩b⇩h,m⇩s⇩b,𝒮⇩s⇩b⇩h) ∼ (ts,m,𝒮) "
assumes valid_own: "valid_ownership 𝒮⇩s⇩b⇩h ts⇩s⇩b⇩h"
assumes valid_sb_reads: "valid_reads m⇩s⇩b ts⇩s⇩b⇩h"
assumes valid_hist: "valid_history program_step ts⇩s⇩b⇩h"
assumes valid_sharing: "valid_sharing 𝒮⇩s⇩b⇩h ts⇩s⇩b⇩h"
assumes tmps_distinct: "tmps_distinct ts⇩s⇩b⇩h"
assumes valid_sops: "valid_sops ts⇩s⇩b⇩h"
assumes valid_dd: "valid_data_dependency ts⇩s⇩b⇩h"
assumes load_tmps_fresh: "load_tmps_fresh ts⇩s⇩b⇩h"
assumes enough_flushs: "enough_flushs ts⇩s⇩b⇩h"
assumes valid_program_history: "valid_program_history ts⇩s⇩b⇩h"
assumes valid: "valid ts⇩s⇩b⇩h"
assumes safe_reach: "safe_reach_direct safe_delayed (ts,m,𝒮)"
shows "∃ts⇩s⇩b⇩h' 𝒮⇩s⇩b⇩h'.
(ts⇩s⇩b⇩h,m⇩s⇩b,𝒮⇩s⇩b⇩h) ⇒⇩s⇩b⇩h⇧* (ts⇩s⇩b⇩h',m⇩s⇩b',𝒮⇩s⇩b⇩h') ∧ ts⇩s⇩b' ∼⇩h ts⇩s⇩b⇩h' ∧
valid_ownership 𝒮⇩s⇩b⇩h' ts⇩s⇩b⇩h' ∧ valid_reads m⇩s⇩b' ts⇩s⇩b⇩h' ∧
valid_history program_step ts⇩s⇩b⇩h' ∧
valid_sharing 𝒮⇩s⇩b⇩h' ts⇩s⇩b⇩h' ∧ tmps_distinct ts⇩s⇩b⇩h' ∧ valid_data_dependency ts⇩s⇩b⇩h' ∧
valid_sops ts⇩s⇩b⇩h' ∧ load_tmps_fresh ts⇩s⇩b⇩h' ∧ enough_flushs ts⇩s⇩b⇩h' ∧
valid_program_history ts⇩s⇩b⇩h' ∧ valid ts⇩s⇩b⇩h' ∧
(∃ts' 𝒮' m'. (ts,m,𝒮) ⇒⇩d⇧* (ts',m',𝒮') ∧
(ts⇩s⇩b⇩h',m⇩s⇩b',𝒮⇩s⇩b⇩h') ∼ (ts',m',𝒮'))"
proof -
from concurrent_history_steps_simulates_store_buffer_step [OF step_sb sim_h]
obtain ts⇩s⇩b⇩h' 𝒮⇩s⇩b⇩h' where
steps_h: "(ts⇩s⇩b⇩h,m⇩s⇩b,𝒮⇩s⇩b⇩h) ⇒⇩s⇩b⇩h⇧* (ts⇩s⇩b⇩h',m⇩s⇩b',𝒮⇩s⇩b⇩h')" and
sim_h': "ts⇩s⇩b' ∼⇩h ts⇩s⇩b⇩h'"
by blast
moreover
from concurrent_direct_steps_simulates_store_buffer_history_steps [OF steps_h
valid_own valid_sb_reads valid_hist valid_sharing tmps_distinct valid_sops valid_dd
load_tmps_fresh enough_flushs valid_program_history valid sim safe_reach]
obtain m' ts' 𝒮' where
"(ts,m,𝒮) ⇒⇩d⇧* (ts',m',𝒮')" "(ts⇩s⇩b⇩h', m⇩s⇩b',𝒮⇩s⇩b⇩h') ∼ (ts', m', 𝒮')"
"valid_ownership 𝒮⇩s⇩b⇩h' ts⇩s⇩b⇩h'" "valid_reads m⇩s⇩b' ts⇩s⇩b⇩h'" "valid_history program_step ts⇩s⇩b⇩h'"
"valid_sharing 𝒮⇩s⇩b⇩h' ts⇩s⇩b⇩h'" "tmps_distinct ts⇩s⇩b⇩h'" "valid_data_dependency ts⇩s⇩b⇩h'"
"valid_sops ts⇩s⇩b⇩h'" "load_tmps_fresh ts⇩s⇩b⇩h'" "enough_flushs ts⇩s⇩b⇩h'"
"valid_program_history ts⇩s⇩b⇩h'" "valid ts⇩s⇩b⇩h'"
by blast
ultimately
show ?thesis
by blast
qed
lemma conj_commI: "P ∧ Q ⟹ Q ∧ P"
by simp
lemma def_to_eq: "P = Q ⟹ P ≡ Q"
by simp
context xvalid_program
begin
definition
"invariant ts 𝒮 m ≡
valid_ownership 𝒮 ts ∧ valid_reads m ts ∧ valid_history program_step ts ∧
valid_sharing 𝒮 ts ∧ tmps_distinct ts ∧ valid_data_dependency ts ∧
valid_sops ts ∧ load_tmps_fresh ts ∧ enough_flushs ts ∧ valid_program_history ts ∧
valid ts"
definition "ownership_inv ≡ valid_ownership"
definition "sharing_inv ≡ valid_sharing"
definition "temporaries_inv ts ≡ tmps_distinct ts ∧ load_tmps_fresh ts"
definition "history_inv ts m ≡ valid_history program_step ts ∧ valid_program_history ts ∧ valid_reads m ts"
definition "data_dependency_inv ts ≡ valid_data_dependency ts ∧ load_tmps_fresh ts ∧ valid_sops ts"
definition "barrier_inv ≡ enough_flushs"
lemma invariant_grouped_def: "invariant ts 𝒮 m ≡
ownership_inv 𝒮 ts ∧ sharing_inv 𝒮 ts ∧ temporaries_inv ts ∧ data_dependency_inv ts ∧ history_inv ts m ∧ barrier_inv ts ∧ valid ts"
apply (rule def_to_eq)
apply (auto simp add: ownership_inv_def sharing_inv_def barrier_inv_def temporaries_inv_def history_inv_def data_dependency_inv_def invariant_def)
done
theorem (in xvalid_program) simulation':
assumes step_sb: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ⇒⇩s⇩b⇩h (ts⇩s⇩b',m⇩s⇩b',𝒮⇩s⇩b')"
assumes sim: "(ts⇩s⇩b,m⇩s⇩b,𝒮⇩s⇩b) ∼ (ts,m,𝒮)"
assumes inv: "invariant ts⇩s⇩b 𝒮⇩s⇩b m⇩s⇩b"
assumes safe_reach: "safe_reach_direct safe_delayed (ts,m,𝒮)"
shows "invariant ts⇩s⇩b' 𝒮⇩s⇩b' m⇩s⇩b' ∧
(∃ts' 𝒮' m'. (ts,m,𝒮) ⇒⇩d⇧* (ts',m',𝒮') ∧ (ts⇩s⇩b',m⇩s⇩b',𝒮⇩s⇩b') ∼ (ts',m',𝒮'))"
using inv sim safe_reach
apply (unfold invariant_def)
apply (simp only: conj_assoc)
apply (rule "concurrent_direct_steps_simulates_store_buffer_history_step" [OF step_sb])
apply simp_all
done
lemmas (in xvalid_program) simulation = conj_commI [OF simulation']
end
end