Theory ReduceStoreBufferSimulation

(* Copyright (C) 2007--2010 Norbert Schirmer
 * All rights reserved, DFKI GmbH 
 *)
theory ReduceStoreBufferSimulation
imports ReduceStoreBuffer
begin

(* FIXME: a lot of theorems that now have sharing_consistent as precondition, may as well work with weak_sharing_consistent
 *)

locale initialsb = 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 initialsb  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 initialsb  outstanding_volatile_writes_unowned_by_others
proof
  fix i j pi isi 𝒪i i 𝒟i θi sbi pj isj 𝒪j j 𝒟j θj sbj
  assume i_bound: "i < length ts" and 
    j_bound: "j < length ts" and
    neq_i_j: "i  j" and
    ts_i: "ts ! i = (pi, isi, θi, sbi, 𝒟i, 𝒪i, i)" and
    ts_j: "ts ! j = (pj, isj, θj, sbj, 𝒟j, 𝒪j, j)" 
  show "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sbi = {}"
  using empty_sb [OF i_bound ts_i] empty_sb [OF j_bound ts_j] by auto
qed

sublocale initialsb  read_only_reads_unowned
proof
  fix i j pi isi 𝒪i i 𝒟i θi sbi pj isj 𝒪j j 𝒟j θj sbj
  assume i_bound: "i < length ts" and 
    j_bound: "j < length ts" and
    neq_i_j: "i  j" and
    ts_i: "ts ! i = (pi, isi, θi, sbi, 𝒟i, 𝒪i, i)" and
    ts_j: "ts ! j = (pj, isj, θj, sbj, 𝒟j, 𝒪j, j)" 
  show "(𝒪j  all_acquired sbj)  
     read_only_reads (acquired True 
                          (takeWhile (Not  is_volatile_Writesb) sbi) 𝒪i) 
                          (dropWhile (Not  is_volatile_Writesb) sbi) = {}"
  using empty_sb [OF i_bound ts_i] empty_sb [OF j_bound ts_j] by auto
qed

sublocale initialsb  ownership_distinct
proof
  fix i j pi isi 𝒪i i 𝒟i θi sbi pj isj 𝒪j j 𝒟j θj sbj
  assume i_bound: "i < length ts" and 
    j_bound: "j < length ts" and
    neq_i_j: "i  j" and
    ts_i: "ts ! i = (pi, isi, θi, sbi, 𝒟i, 𝒪i, i)" and
    ts_j: "ts ! j = (pj, isj, θj, sbj, 𝒟j, 𝒪j, j)" 
  show "(𝒪i  all_acquired sbi)  (𝒪j  all_acquired sbj) = {}"
  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 initialsb  valid_ownership ..

sublocale initialsb  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 initialsb  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 initialsb  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 initialsb  valid_sharing ..
sublocale initialsb  valid_ownership_and_sharing ..

sublocale initialsb  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 initialsb  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 initialsb  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 initialsb  load_tmps_read_tmps_distinct ..

sublocale initialsb  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 initialsb  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 initialsb  valid_sops ..

sublocale initialsb  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 initialsb  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 initialsb  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 initialsb  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 initialsb  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_Writesb sb = {}"
   using empty_sb [OF i_bound ts_i] by auto
qed

sublocale initialsb  valid_program_history
proof
   fix i "is" 𝒪  𝒟 θ sb p sb1 sb2
   assume i_bound: "i < length ts"
   assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)"
   assume sb: "sb=sb1@sb2"
   show "isa. instrs sb2 @ is = isa @ prog_instrs sb2"
   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 tssb msb;
    𝒮 = share_all_until_volatile_write tssb 𝒮sb;
    length tssb = length ts; 
    i < length tssb. 
           let (p, issb, θ, sb, 𝒟sb, 𝒪, ) = tssb!i;
               suspends = dropWhile (Not  is_volatile_Writesb) sb
            in  is 𝒟. instrs suspends @ issb = is @ prog_instrs suspends 
                    𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb sb  {}) 
                ts!i = (hd_prog p suspends, 
                        is,
                        θ |` (dom θ - read_tmps suspends),(),
                        𝒟,  
                        acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪,
                        release (takeWhile (Not  is_volatile_Writesb) sb) (dom 𝒮sb)  )
    
     
     (tssb,msb,𝒮sb)  (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 tsh; 
    i < length ts. 
         (𝒪' 𝒟' ℛ'.
           let (p,is, θ, sb,𝒟, 𝒪,) = tsh!i in 
                ts!i=(p,is, θ, filter is_Writesb sb,𝒟',𝒪',ℛ') 
                (filter is_Writesb sb = []  sb=[]))
    
     
     ts h tsh"

lemma (in initialsb) 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 tssb  tssb!i=(p,is,xs,sb,𝒟,𝒪,) sb=[]"
  assumes sim: "(tssb,msb,𝒮sb)  (ts,m,𝒮)"
  shows "𝒮 = 𝒮sb  m = msb  length ts = length tssb 
         (i < length tssb. 
           let (p, is, θ, sb, 𝒟, 𝒪, ) = tssb!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 tssb  tssb!i=(p,is,xs,sb,𝒟,𝒪,) sb=[]"
  assumes leq: "length ts = length tssb"
  assumes ts: "(i < length tssb. 
           let (p, is, θ, sb, 𝒟, 𝒪, ) = tssb!i
            in ts!i = (p, is, θ, (), 𝒟, 𝒪, ))"
  shows "(tssb,msb,𝒮sb)  (ts,msb,𝒮sb)"
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)   (xset ts'. P x) = (xset 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

(* FIXME: move up *)
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 c0"
  shows "safe_reach_upto n safe c0"
proof
  fix k c l
  assume k_n: "k  n"
  assume trace: "trace c 0 k"
  assume c_0: "c 0 = c0"
  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: "initialsb tssb 𝒮sb"
  assumes sim: "(tssb,msb,𝒮sb)  (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: initialsb tssb 𝒮sb .
  from sim obtain
   m: "m = flush_all_until_volatile_write tssb msb" and
   𝒮: "𝒮 = share_all_until_volatile_write tssb 𝒮sb" and
   leq: "length tssb = length ts" and
   t_sim: "i < length tssb. 
           let (p, issb, θ, sb, 𝒟sb, 𝒪, ) = tssb!i;
               suspends = dropWhile (Not  is_volatile_Writesb) sb
            in  is 𝒟. instrs suspends @ issb = is @ prog_instrs suspends 
                    𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb sb  {}) 
                ts!i = (hd_prog p suspends, 
                        is,
                        θ |` (dom θ - read_tmps suspends),(),
                        𝒟,  
                        acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪,
                        release (takeWhile (Not  is_volatile_Writesb) sb) (dom 𝒮sb)  )"
    by cases auto

  from ini.empty_sb  
  have shared_eq: "𝒮 = 𝒮sb"
    apply (simp only: 𝒮)
    apply (rule share_all_empty)
    apply force
    done
  have sd: "simple_ownership_distinct ts"
  proof 
    fix i j pi isi 𝒪i i 𝒟i θi sbi pj isj 𝒪j j 𝒟j θj sbj
    assume i_bound: "i < length ts" and 
      j_bound: "j < length ts" and
      neq_i_j: "i  j" and
      ts_i: "ts ! i = (pi, isi, θi, sbi, 𝒟i, 𝒪i, i)" and
      ts_j: "ts ! j = (pj, isj, θj, sbj, 𝒟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: "tssb!i = (pi,isi,θ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: "tssb!j = (pj,isj,θ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: "tssb!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 tssb)) = (((λ(_, _, _, _, _, 𝒪, _). 𝒪) ` 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: "tssb!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

(* FIXME: move up *)
lemma map_onws_sb_owned:"j. j < length ts  map 𝒪_sb ts ! j = (𝒪j,sbj)  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,sbj)  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

(* FIXME: substitutes in application below: read_only_read_acquired_unforwarded_witness*)
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 @ Writesb True a' sop v A L R W # zs  
     a  A  a  outstanding_refs is_Writesb ys  a'  a) 
(A L R W ys zs. sb = ys @ Ghostsb A L R W# zs  a  A  a  outstanding_refs is_Writesb ys)"
proof (induct sb)
  case Nil thus ?case by simp
next
  case (Cons x sb)
  show ?case
  proof (cases x)
    case (Writesb 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 (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  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 (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb" and 
	a_unforw: "a  unforwarded_non_volatile_reads sb (insert a' X)" 
	by (clarsimp simp add: Writesb 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: Writesb 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 (𝒮 ⊕⇘WR ⊖⇘AL)" 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 @ Writesb True a' sop v A L R W # zs 
                 a  A  a  outstanding_refs is_Writesb ys  a'  a) 
              (A L R W ys zs. sb = ys @ Ghostsb A L R W# zs  a  A  a  outstanding_refs is_Writesb 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 @ Writesb True a'' sop' v' A' L' R' W' # zs" and
            props: "a  A'" "a  outstanding_refs is_Writesb 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: Writesb volatile False a'_a)
	    done
	next
	  assume ?ghst
	  then obtain ys zs A' L' R' W'  where 
            sb: "sb = ys @ Ghostsb A' L' R' W'# zs" and
            props: "a  A'" "a  outstanding_refs is_Writesb 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: Writesb 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: Writesb 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 @ Writesb True a' sop v A L R W # zs 
                 a  A  a  outstanding_refs is_Writesb ys  a'  a) 
              (A L R W ys zs. sb = ys @ Ghostsb A L R W# zs  a  A  a  outstanding_refs is_Writesb 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 @ Writesb True a'' sop' v' A' L' R' W' # zs" and
            props: "a  A'" "a  outstanding_refs is_Writesb 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: Writesb False )
	    done
	next
	  assume ?ghst
	  then obtain ys zs A' L' R' W'  where 
            sb: "sb = ys @ Ghostsb A' L' R' W' # zs" and
            props: "a  A'" "a  outstanding_refs is_Writesb 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: Writesb False )
	    done
	qed
      qed
    next
    case (Readsb 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: Readsb 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 @ Writesb True a' sop v A L R W # zs 
                 a  A  a  outstanding_refs is_Writesb ys  a'  a) 
              (A L R W ys zs. sb = ys @ Ghostsb A L R W# zs  a  A  a  outstanding_refs is_Writesb 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 @ Writesb True a'' sop' v' A' L' R' W' # zs" and
        props: "a  A'" "a  outstanding_refs is_Writesb 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: Readsb)
	done
    next
      assume ?ghst
      then obtain ys zs A' L' R' W' where 
        sb: "sb = ys @ Ghostsb A' L' R' W'# zs" and
        props: "a  A'" "a  outstanding_refs is_Writesb 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: Readsb )
      done
    qed
  next
    case Progsb
    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: Progsb)
    from Cons.hyps [OF nvo' consis' a_nro' a_unowned a_unforw]
    have "(sop a' v ys zs A L R W.
                 sb = ys @ Writesb True a' sop v A L R W # zs 
                 a  A  a  outstanding_refs is_Writesb ys  a'  a) 
              (A L R W ys zs. sb = ys @ Ghostsb A L R W# zs  a  A  a  outstanding_refs is_Writesb 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 @ Writesb True a'' sop' v' A' L' R' W' # zs" and
        props: "a  A'" "a  outstanding_refs is_Writesb 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: Progsb)
	done
    next
      assume ?ghst
      then obtain ys zs A' L' R' W' where 
        sb: "sb = ys @ Ghostsb A' L' R' W'# zs" and
        props: "a  A'" "a  outstanding_refs is_Writesb 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: Progsb )
      done
    qed
  next
    case (Ghostsb A L R W)
    from Cons.prems obtain 
      nvo': "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  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 (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb" and 
      a_unforw: "a  unforwarded_non_volatile_reads sb X"
      by (clarsimp simp add: Ghostsb)

    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: Ghostsb 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 (𝒮 ⊕⇘WR ⊖⇘AL)" 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 @ Writesb True a' sop v A L R W # zs 
                 a  A  a  outstanding_refs is_Writesb ys  a'  a) 
              (A L R W ys zs. sb = ys @ Ghostsb A L R W# zs  a  A  a  outstanding_refs is_Writesb 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 @ Writesb True a'' sop' v' A' L' R' W' # zs" and
          props: "a  A'" "a  outstanding_refs is_Writesb 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: Ghostsb False )
	  done
      next
	assume ?ghst
	then obtain ys zs A' L' R' W'  where 
          sb: "sb = ys @ Ghostsb A' L' R' W'# zs" and
          props: "a  A'" "a  outstanding_refs is_Writesb 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: Ghostsb False )
	  done
	qed
      qed
    qed
  qed (* FIXME: indentation*)

(*
lemma release_take_drop:
"⋀ℛ S. release (dropWhile (Not ∘ is_volatile_Writesb) sb) S (release (takeWhile (Not ∘ is_volatile_Writesb) sb) S ℛ) =
           release sb S ℛ"
apply (induct sb)
apply clarsimp
apply (auto split:memref.splits)
apply fastforce*)

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 (Writesb 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: Writesb True )
      from shared_eq
      have shared_eq': "a𝒪  A - R  all_acquired sb. (𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
        by (auto simp add: augment_shared_def restrict_shared_def)
      from Cons.hyps [OF shared_eq' consis']
      have "release sb (dom (𝒮' ⊕⇘WR ⊖⇘AL)) Map.empty = release sb (dom (𝒮 ⊕⇘WR ⊖⇘AL)) Map.empty" .
      then show ?thesis
        by (auto  simp add: Writesb True domIff) 
    next
      case False with Cons show ?thesis
	by (auto simp add: Writesb)
    qed
  next
    case Readsb with Cons show ?thesis
      by auto
  next
    case Progsb with Cons show ?thesis
      by auto
  next
    case (Ghostsb 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: Ghostsb )
    from shared_eq
    have shared_eq': "a𝒪  A - R  all_acquired sb. (𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
      by (auto simp add: augment_shared_def restrict_shared_def)
    from shared_eq R_owns have "aR. (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 (𝒮' ⊕⇘WR ⊖⇘AL)) (augment_rels (dom 𝒮') R ) = 
            release sb (dom (𝒮 ⊕⇘WR ⊖⇘AL)) (augment_rels (dom 𝒮) R )" by simp
    then show ?thesis
      by (clarsimp  simp add: Ghostsb 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 (Writesb volatile a sop v A L R W)
    show ?thesis
    proof (cases volatile)
      case True
      with Writesb Cons.hyps [of "(𝒮 ⊕⇘WR ⊖⇘AL)"] 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 Writesb Cons show ?thesis by auto
    qed
  next
    case Readsb with Cons show ?thesis by auto
  next
    case Progsb with Cons show ?thesis by auto
  next
    case (Ghostsb A L R W)
    with Cons.hyps [of "(𝒮 ⊕⇘WR ⊖⇘AL)"] 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_Writesb) 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_Writesb) = (λa. ¬ is_volatile_Writesb a)"
    by (rule ext) auto


  let ?take_sb = "(takeWhile (Not  is_volatile_Writesb) sb)"
  let ?drop_sb = "(dropWhile (Not  is_volatile_Writesb) 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_Writesb) sb)) ` set ts))"
    have "a  all_shared (takeWhile (Not  is_volatile_Writesb) 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 (Writesb 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 (𝒮 ⊕⇘WR ⊖⇘AL))" and
        a_A_all: "a  𝒪  A  all_acquired sb"
	by (clarsimp simp add: Writesb True)

      from owns_ro A_R R_owns have owns_ro': "(𝒪  A - R)  read_only (𝒮 ⊕⇘WR ⊖⇘AL) = {}"
        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: Writesb volatile augment_shared_def restrict_shared_def read_only_def domIff
           split: if_split_asm)
    next 
      case False
      with Cons Writesb show ?thesis by (auto)
    qed
  next
    case Readsb with Cons show ?thesis by auto
  next
    case Progsb with Cons show ?thesis by auto
  next
    case (Ghostsb 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 (𝒮 ⊕⇘WR ⊖⇘AL))" and
      a_A_all: "a  𝒪  A  all_acquired sb"
      by (clarsimp simp add: Ghostsb)

    from owns_ro A_R R_owns have owns_ro': "(𝒪  A - R)  read_only (𝒮 ⊕⇘WR ⊖⇘AL) = {}"
      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: Ghostsb 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 (Writesb volatile a' sop v A L R W)
    show ?thesis
    proof (cases volatile)
      case False
      with Cons Writesb 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: Writesb True)
      from owns_ro A_R R_owns have owns_ro': "(𝒪  A - R)  read_only (𝒮 ⊕⇘WR ⊖⇘AL) = {}"
        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 (𝒮 ⊕⇘WR ⊖⇘AL)"
         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 (𝒮 ⊕⇘WR ⊖⇘AL))"
         by auto
       then show ?thesis
         by (auto simp add: Writesb True)
    qed
  next
    case Readsb with Cons show ?thesis by auto
  next
    case Progsb with Cons show ?thesis by auto
  next
    case (Ghostsb 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: Ghostsb)
    from owns_ro A_R R_owns have owns_ro': "(𝒪  A - R)  read_only (𝒮 ⊕⇘WR ⊖⇘AL) = {}"
      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 (𝒮 ⊕⇘WR ⊖⇘AL)"
      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 (𝒮 ⊕⇘WR ⊖⇘AL))"
      by auto
    then show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed

(*
lemma release_False_mono:
  "⋀S ℛ. ℛ a = Some False ⟹ 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 (Ghostsb 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: Ghostsb)
  next
    case Writesb with Cons show ?thesis apply auto
  next 
    case Readsb with Cons show ?thesis by auto
  next
    case Progsb with Cons show ?thesis by auto
  qed
qed
*)


lemma release_False_mono:
  "S .  a = Some False  outstanding_refs is_volatile_Writesb 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 (Ghostsb 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: Ghostsb)
  next
    case Writesb with Cons show ?thesis by auto
  next 
    case Readsb with Cons show ?thesis by auto
  next
    case Progsb with Cons show ?thesis by auto
  qed
qed

lemma release_False_mono_take:
  "S .  a = Some False  release (takeWhile (Not  is_volatile_Writesb) 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 (Ghostsb 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: Ghostsb)
  next
    case Writesb with Cons show ?thesis by auto
  next 
    case Readsb with Cons show ?thesis by auto
  next
    case Progsb 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_Writesb) = (λa. ¬ is_volatile_Writesb a)"
    by (rule ext) auto
  show ?case
  proof (cases x)
    case (Ghostsb A L R W)
    from Cons.prems obtain 
      dist: "read_only 𝒮  𝒪 = {}" and
      share: "𝒮 a  Some False" and
      share': "share sb (𝒮 ⊕⇘WR ⊖⇘AL) 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: Ghostsb aargh)
  
    from dist L_A A_R R_owns have dist':  "read_only (𝒮 ⊕⇘WR ⊖⇘AL)  (𝒪  A - R)= {}"
      by (auto simp add: in_read_only_convs)

    show ?thesis
    proof (cases "(𝒮 ⊕⇘WR ⊖⇘AL) a = Some False")
      case False
      from Cons.hyps [OF consis' dist' this share']
      show ?thesis by (auto simp add: Ghostsb)
    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: Ghostsb)
   qed
 next
   case (Writesb 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 (𝒮 ⊕⇘WR ⊖⇘AL) 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: Writesb True aargh)
  
     from dist L_A A_R R_owns have dist':  "read_only (𝒮 ⊕⇘WR ⊖⇘AL)  (𝒪  A - R)= {}"
       by (auto simp add: in_read_only_convs)

     show ?thesis
     proof (cases "(𝒮 ⊕⇘WR ⊖⇘AL) a = Some False")
       case False
       from Cons.hyps [OF consis' dist' this share']
       show ?thesis by (auto simp add: Writesb 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: Writesb volatile)
     qed 
   next
     case False
     with Cons show ?thesis by (auto simp add: Writesb)
   qed
 next
   case Readsb with Cons show ?thesis by (auto)
 next
   case Progsb with Cons show ?thesis by (auto)
 qed
qed

lemma shared_switch_release_False: 
  "𝒮 . 
     outstanding_refs is_volatile_Writesb 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_Writesb) = (λa. ¬ is_volatile_Writesb a)"
    by (rule ext) auto
  show ?case
  proof (cases x)
    case (Ghostsb A L R W)
    from Cons.prems obtain 
      a_notin: "a  dom 𝒮" and
      share: "a  dom (share sb (𝒮 ⊕⇘WR ⊖⇘AL))" and
      out': "outstanding_refs is_volatile_Writesb sb = {}"
      by (clarsimp simp add: Ghostsb aargh)
  
    show ?thesis
    proof (cases "a  R")
      case False
      with a_notin have "a  dom (𝒮 ⊕⇘WR ⊖⇘AL)"
        by auto
      from Cons.hyps [OF out' this share]
      show ?thesis
        by (auto simp add: Ghostsb)
    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: Ghostsb)
    qed
  next
    case Writesb with Cons show ?thesis by (clarsimp split: if_split_asm)
  next
    case Readsb with Cons show ?thesis by auto
  next 
    case Progsb with Cons show ?thesis by auto
  qed 
qed


lemma release_not_unshared_no_write:  
  "𝒮 . 
     outstanding_refs is_volatile_Writesb sb = {};     
  non_volatile_writes_unshared 𝒮 sb;
  release sb (dom 𝒮)  a   Some False;
  a  dom (share sb 𝒮)
   
    a  outstanding_refs is_non_volatile_Writesb 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_Writesb) = (λa. ¬ is_volatile_Writesb a)"
    by (rule ext) auto
  show ?case
  proof (cases x)
    case (Ghostsb A L R W)
    from Cons.prems obtain 
      share: "a  dom (share sb (𝒮 ⊕⇘WR ⊖⇘AL))" and
      rel: "release sb 
                (dom (𝒮 ⊕⇘WR ⊖⇘AL)) (augment_rels (dom 𝒮) R ) a   Some False" and
      out': "outstanding_refs is_volatile_Writesb sb = {}" and
      nvu: "non_volatile_writes_unshared (𝒮 ⊕⇘WR ⊖⇘AL) sb" 
      by (clarsimp simp add: Ghostsb )
  
    
    from Cons.hyps [OF out' nvu rel share]
    show ?thesis by (auto simp add: Ghostsb)
  next
    case (Writesb volatile a' sop v A L R W)
    show ?thesis
    proof (cases volatile)
      case True with Writesb 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_Writesb 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: Writesb False)
      from Cons.hyps [OF out' nvo rel a_dom]
      have a_notin_rest: "a  outstanding_refs is_non_volatile_Writesb sb".
      
      show ?thesis
      proof (cases "a'=a")
        case False with a_notin_rest
        show ?thesis by (clarsimp simp add: Writesb  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 Readsb with Cons show ?thesis by auto
  next 
    case Progsb 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_Writesb) sb)"
 assumes rel: "release (takeWhile (Not  is_volatile_Writesb) sb) (dom 𝒮)  a   Some False"
 assumes a_in: "a  dom (share (takeWhile (Not  is_volatile_Writesb) sb) 𝒮)"
 shows
    "a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sb)" 
using release_not_unshared_no_write[OF takeWhile_not_vol_write_outstanding_refs [of sb] nvw rel a_in]
by simp

(* FIXME: may replace the un-primed variants, similar for the following lemmas *)
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 (Writesb 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: Writesb 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 ⊕⇘WR ⊖⇘AL) = {}"
	  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 ⊕⇘WR ⊖⇘AL)"
	  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: Writesb True)
      qed
   next
      case False
      with Cons show ?thesis
	by (clarsimp simp add: Writesb False)
    qed
  next
    case Readsb with Cons show ?thesis by (clarsimp)
  next
    case Progsb with Cons show ?thesis by (clarsimp)
  next
    case (Ghostsb 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: Ghostsb)
    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 ⊕⇘WR ⊖⇘AL) = {}"
        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 ⊕⇘WR ⊖⇘AL)"
        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: Ghostsb)
    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_Writesb) sb) 𝒪 
            a  all_shared (takeWhile (Not  is_volatile_Writesb) 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_Writesb) = (λa. ¬ is_volatile_Writesb 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_Writesb) sb)"
  let ?drop_sb = "(dropWhile (Not  is_volatile_Writesb) 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 pj isj 𝒪j j 𝒟j θj sbj
    assume j_bound: "j < length ts"
    assume jth: "ts!j = (pj,isj,θj,sbj,𝒟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 sbj) = {}"
            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_Writesb) sb) 𝒪 
            a  all_shared (takeWhile (Not  is_volatile_Writesb) sb)) " and
    unacq_sb: "a  acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪" and
    unsh_sb: "a  all_shared (takeWhile (Not  is_volatile_Writesb) 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_Writesb) 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 (Writesb 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: Writesb 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: Writesb True)
     qed
   next
     case False with Cons.hyps [of X Y] Cons.prems show ?thesis by (auto simp add: Writesb)
   qed
 next
   case Readsb with Cons.hyps [of X Y] Cons.prems show ?thesis by (auto)
 next
   case Progsb with Cons.hyps [of X Y] Cons.prems show ?thesis by (auto)
 next
   case (Ghostsb 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: Ghostsb)
   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: Ghostsb)
   qed
 qed
qed

(* FIXME: could be strengthened to acquired True sb empty I suppose *)
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 (Writesb 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 ⊕⇘WR ⊖⇘AL))" 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: Writesb True)

      from owns_ro A_R R_owns have owns_ro': "(𝒪  A - R)  read_only (S ⊕⇘WR ⊖⇘AL) = {}"
	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 ⊕⇘WR ⊖⇘AL)  a  acquired True sb (𝒪  A - R)  a  all_shared sb)".

      have "a  read_only (share sb (Map.empty ⊕⇘WR ⊖⇘AL))  
           (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 ⊕⇘WR ⊖⇘AL)"
	    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 ⊕⇘WR ⊖⇘AL))".
	}
	moreover
	{
	  assume a_ro: "a  read_only (S ⊕⇘WR ⊖⇘AL)" 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 ⊕⇘WR ⊖⇘AL)"
	      by (auto simp add: in_read_only_convs split: if_split_asm)
	    
	    have "read_only (Map.empty ⊕⇘WR ⊖⇘AL)  read_only (S ⊕⇘WR ⊖⇘AL)"
	      by (auto simp add: in_read_only_convs)
	    with owns_ro'
	    have owns_ro_empty: "(𝒪  A - R)  read_only (Map.empty ⊕⇘WR ⊖⇘AL) = {}"
	      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 ⊕⇘WR ⊖⇘AL))".
	    thus ?thesis
	      by simp
	  qed
	}
	moreover note hyp
	ultimately show ?thesis by blast
      qed

      then show ?thesis
	by (clarsimp simp add: Writesb True)
    next
      case False with Cons show ?thesis
	by (auto simp add: Writesb)
    qed
  next
    case Readsb with Cons show ?thesis by auto
  next
    case Progsb with Cons show ?thesis by auto
  next
    case (Ghostsb A L R W)
    from Cons.prems
    obtain a_in: "a  read_only (share sb (S ⊕⇘WR ⊖⇘AL))" 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: Ghostsb)
    
    from owns_ro A_R R_owns have owns_ro': "(𝒪  A - R)  read_only (S ⊕⇘WR ⊖⇘AL) = {}"
      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 ⊕⇘WR ⊖⇘AL)  a  acquired True sb (𝒪  A - R)  a  all_shared sb)".

    have "a  read_only (share sb (Map.empty ⊕⇘WR ⊖⇘AL))  
           (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 ⊕⇘WR ⊖⇘AL)"
	  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 ⊕⇘WR ⊖⇘AL))".
      }
      moreover
      {
	assume a_ro: "a  read_only (S ⊕⇘WR ⊖⇘AL)" 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 ⊕⇘WR ⊖⇘AL)"
	    by (auto simp add: in_read_only_convs split: if_split_asm)
	    
	  have "read_only (Map.empty ⊕⇘WR ⊖⇘AL)  read_only (S ⊕⇘WR ⊖⇘AL)"
	    by (auto simp add: in_read_only_convs)
	  with owns_ro'
	  have owns_ro_empty: "(𝒪  A - R)  read_only (Map.empty ⊕⇘WR ⊖⇘AL) = {}"
	    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 ⊕⇘WR ⊖⇘AL))".
	  thus ?thesis
	    by simp
	qed
      }
      moreover note hyp
      ultimately show ?thesis by blast
    qed
    
    then show ?thesis
      by (clarsimp simp add: Ghostsb)
  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. ij  
            (let (_,_,_,sbj,_,𝒪,_) = ts!j in
            a  acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪 
            a  all_shared (takeWhile (Not  is_volatile_Writesb) sbj ))"
  assumes a_ro_share: "a  read_only (share sb 𝒮)"
  shows "a  read_only (share (dropWhile (Not  is_volatile_Writesb) 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]
  (* make similar version with acquired and all_shared instead of all_acquired *)
  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_Writesb) sb)"
  let ?drop_sb = "(dropWhile (Not  is_volatile_Writesb) 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 (_,_,_,sbj,_,𝒪,_) = ts!j in
            a  acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪 
            a  all_shared (takeWhile (Not  is_volatile_Writesb) sbj ))"
        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. ij  (ℛ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 (Writesb 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: Writesb volatile)
    next
      case False with Cons Writesb show ?thesis by auto
    qed
  next
    case Readsb with Cons show ?thesis by auto
  next
    case Progsb with Cons show ?thesis by auto
  next
    case (Ghostsb 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: Ghostsb)

    with Cons.hyps [of "(S'  R - L)" "(S  R - L)" "augment_rels S R "] Cons.prems
    show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed

lemma release_append_Progsb:
"S . (release (takeWhile (Not  is_volatile_Writesb) (sb @ [Progsb p1 p2 mis])) S ) = 
       (release  (takeWhile (Not  is_volatile_Writesb) 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: "(tssb,msb,𝒮sb) sbh (tssb',msb',𝒮sb')"
  assumes valid_own: "valid_ownership 𝒮sb tssb"
  assumes valid_sb_reads: "valid_reads msb tssb"
  assumes valid_hist: "valid_history program_step tssb"
  assumes valid_sharing: "valid_sharing 𝒮sb tssb"
  assumes tmps_distinct: "tmps_distinct tssb"
  assumes valid_sops: "valid_sops tssb"
  assumes valid_dd: "valid_data_dependency tssb"
  assumes load_tmps_fresh: "load_tmps_fresh tssb"
  assumes enough_flushs: "enough_flushs tssb"
  assumes valid_program_history: "valid_program_history tssb"
  assumes valid: "valid tssb"
  assumes sim: "(tssb,msb,𝒮sb)  (ts,m,𝒮)"
  assumes safe_reach: "safe_reach_direct safe_delayed (ts,m,𝒮)"
  shows "valid_ownership 𝒮sb' tssb'  valid_reads msb' tssb'  valid_history program_step tssb' 
         valid_sharing 𝒮sb' tssb'  tmps_distinct tssb'  valid_data_dependency tssb' 
         valid_sops tssb'  load_tmps_fresh tssb'  enough_flushs tssb' 
         valid_program_history tssb'  valid tssb' 
           (ts' 𝒮' m'. (ts,m,𝒮) d* (ts',m',𝒮')  
                     (tssb',msb',𝒮sb')  (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 @ [Progsb p p' is]" .
  interpret valid_ownership 𝒮sb tssb by fact
  interpret valid_reads msb tssb by fact
  interpret valid_history program_step tssb by fact
  interpret valid_sharing 𝒮sb tssb by fact
  interpret tmps_distinct tssb by fact
  interpret valid_sops tssb by fact
  interpret valid_data_dependency tssb by fact
  interpret load_tmps_fresh tssb by fact
  interpret enough_flushs tssb by fact
  interpret valid_program_history tssb by fact
  from valid_own valid_sharing
  have valid_own_sharing: "valid_ownership_and_sharing 𝒮sb tssb"
    by (simp add: valid_sharing_def valid_ownership_and_sharing_def)
  then
  interpret valid_ownership_and_sharing 𝒮sb tssb .

  from safe_reach_safe_refl [OF safe_reach]
  have safe: "safe_delayed (ts,m,𝒮)".

  from step_sb
  show ?thesis
  proof (cases)
    case (Memop i psb "issb" θsb sb  𝒟sb 𝒪sb sb  "issb'" θsb' sb'  𝒟sb' 𝒪sb' sb')
    then obtain 
      tssb': "tssb' = tssb[i := (psb, issb',θsb', sb', 𝒟sb', 𝒪sb',sb')]" and
      i_bound: "i < length tssb" and
      tssb_i: "tssb ! i = (psb, issb,θsb,sb, 𝒟sb, 𝒪sb,sb)" and
      sbh_step: "(issb, θsb, sb, msb, 𝒟sb, 𝒪sb, sb,𝒮sb) sbh 
                  (issb', θsb', sb', msb', 𝒟sb', 𝒪sb', sb', 𝒮sb')"
      by auto

    from sim obtain 
      m: "m = flush_all_until_volatile_write tssb msb" and
      𝒮: "𝒮 = share_all_until_volatile_write tssb 𝒮sb" and
      leq: "length tssb = length ts" and
      ts_sim: "i<length tssb.
           let (p, issb, θ, sb, 𝒟sb, 𝒪sb,) = tssb ! i;
               suspends = dropWhile (Not  is_volatile_Writesb) sb
           in  is 𝒟. instrs suspends @ issb = is @ prog_instrs suspends 
                    𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb sb  {}) 
                    ts ! i =
                   (hd_prog p suspends, 
                    is,
                    θ |` (dom θ - read_tmps suspends), (),
                    𝒟, 
                    acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪sb,
                    release (takeWhile (Not  is_volatile_Writesb) sb) (dom 𝒮sb) )"
      by cases blast

    from i_bound leq have i_bound': "i < length ts"
      by auto

    have split_sb: "sb = takeWhile (Not  is_volatile_Writesb) sb @ dropWhile (Not  is_volatile_Writesb) sb"
      (is "sb = ?take_sb@?drop_sb")
      by simp

    from ts_sim [rule_format, OF i_bound] tssb_i obtain suspends "is" 𝒟 where
      suspends: "suspends = dropWhile (Not  is_volatile_Writesb) sb" and
      is_sim: "instrs suspends @ issb = is @ prog_instrs suspends" and
      𝒟: "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb sb  {})" and
      ts_i: "ts ! i =
          (hd_prog psb suspends, is,
           θsb |` (dom θsb - read_tmps suspends), (), 𝒟, acquired True ?take_sb 𝒪sb,
            release ?take_sb (dom 𝒮sb) sb)"
      by (auto simp add: Let_def)

    from sbh_step_preserves_valid [OF i_bound tssb_i sbh_step valid]
    have valid': "valid tssb'"
      by (simp add: tssb')
    

    from 𝒟 have 𝒟sb: "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb ?drop_sb  {})"
      apply -
      apply (case_tac "outstanding_refs is_volatile_Writesb 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 := (psb, issb, θsb, (), 𝒟sb, acquired True sb 𝒪sb,
                         release sb (dom 𝒮sb) sb)]"
    have i_bound_ts': "i < length ?ts'"
      using i_bound'
      by auto
    hence ts'_i: "?ts'!i = (psb, issb, θsb, (), 
                     𝒟sb, acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb)"
      by simp 

    from local.sharing_consis_axioms
    have sharing_consis_tssb: "sharing_consis 𝒮sb tssb" .
    from sharing_consis [OF i_bound tssb_i]
    have sharing_consis_sb: "sharing_consistent 𝒮sb 𝒪sb sb".
    from sharing_consistent_weak_sharing_consistent [OF this]
    have weak_consis_sb: "weak_sharing_consistent 𝒪sb sb".
    from this weak_sharing_consistent_append [of 𝒪sb ?take_sb ?drop_sb]
    have weak_consis_drop:"weak_sharing_consistent (acquired True ?take_sb 𝒪sb) ?drop_sb"
      by auto
    from local.ownership_distinct_axioms
    have ownership_distinct_tssb: "ownership_distinct tssb" .
    have steps_flush_sb: "(ts,m,𝒮) d* (?ts', flush ?drop_sb m, share ?drop_sb 𝒮)"
    proof -
      from valid_reads [OF i_bound tssb_i]
      have reads_consis: "reads_consistent False 𝒪sb msb sb".
      from reads_consistent_drop_volatile_writes_no_volatile_reads [OF this]
      have no_vol_read: "outstanding_refs is_volatile_Readsb ?drop_sb = {}".
      from valid_program_history [OF i_bound tssb_i]
      have "causal_program_history issb sb".
      then have cph: "causal_program_history issb ?drop_sb"
	apply -
	apply (rule causal_program_history_suffix [where sb="?take_sb"] )
	apply (simp)
	done
      from valid_last_prog [OF i_bound tssb_i] have last_prog: "last_prog psb sb = psb".
      then
      have lp: "last_prog psb ?drop_sb = psb"
	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 
	tssb_i reads_consis]
      have reads_consis_m: "reads_consistent True (acquired True ?take_sb 𝒪sb) m ?drop_sb"
	by (simp add: m)
      
      from valid_history [OF i_bound tssb_i]
      have h_consis: "history_consistent θsb (hd_prog psb (?take_sb@?drop_sb)) (?take_sb@?drop_sb)"
	by (simp)
      
      have last_prog_hd_prog: "last_prog (hd_prog psb sb) ?take_sb = (hd_prog psb ?drop_sb)"
      proof -
	from last_prog_hd_prog_append' [OF h_consis] last_prog
	have "last_prog (hd_prog psb ?drop_sb) ?take_sb = hd_prog psb ?drop_sb"
	  by (simp)
	moreover 
	have "last_prog (hd_prog psb (?take_sb @ ?drop_sb)) ?take_sb = 
          last_prog (hd_prog psb ?drop_sb) ?take_sb"
	  by (rule last_prog_hd_prog_append)
	ultimately show ?thesis
	  by (simp)
      qed
       
      from valid_write_sops [OF i_bound tssb_i]
      have "sopwrite_sops (?take_sb@?drop_sb). valid_sop sop"
	by (simp)
      then obtain valid_sops_take: "sopwrite_sops ?take_sb. valid_sop sop" and
	valid_sops_drop: "sopwrite_sops ?drop_sb. valid_sop sop"
	apply (simp only: write_sops_append)
	apply auto
	done
	  
      from read_tmps_distinct [OF i_bound tssb_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 θsb (hd_prog psb ?drop_sb) ?drop_sb"
	by (simp add: last_prog_hd_prog)

      have rel_eq: "release ?drop_sb (dom 𝒮) (release  ?take_sb (dom 𝒮sb) sb) = 
                       release sb (dom 𝒮sb) sb"
      proof -
        from release_append [of ?take_sb ?drop_sb]
        have "release sb (dom 𝒮sb) sb =
                release ?drop_sb (dom (share ?take_sb 𝒮sb)) (release  ?take_sb (dom 𝒮sb) sb)"
          by simp
        also
        have dist: "ownership_distinct tssb" by fact
        have consis: "sharing_consis 𝒮sb tssb" by fact

        have "release ?drop_sb (dom (share ?take_sb 𝒮sb)) (release  ?take_sb (dom 𝒮sb) sb) =
              release ?drop_sb (dom 𝒮) (release  ?take_sb (dom 𝒮sb) sb) "
          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 tssb_i, symmetric])
          using acquired_all_acquired [of True ?take_sb 𝒪sb] 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_Readsb_volatile_reads_consistent [OF no_vol_read], of 𝒮]
      show ?thesis by (simp add: acquired_take_drop [where pending_write=True, simplified] 𝒟sb 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 
      (issb, θsb, flush ?drop_sb m, 𝒟sb,acquired True sb 𝒪sb,
        share ?drop_sb 𝒮) ".


    
    from acquired_takeWhile_non_volatile_Writesb 
    have acquired_take_sb: "acquired True ?take_sb 𝒪sb  𝒪sb  all_acquired ?take_sb ".

(* FIXME delete
    from share_takeWhile_non_volatile_Writesb
    have share_take_sb: "share ?take_sb 𝒮sb = 
      𝒮sb(all_acquired ?take_sb) all_unshared ?take_sb".

    from sharing_consis [OF i_bound tssb_i]
    have "sharing_consistent 𝒮sb 𝒪sb sb".

    with sharing_consistent_append [where xs="?take_sb" and ys="?drop_sb", of 𝒮sb 𝒪sb]
    have sharing_consis_drop_sb: 
      "sharing_consistent (share ?take_sb 𝒮sb) (acquired True ?take_sb 𝒪sb) ?drop_sb"
      by (simp add: acquired_take_sb share_takeWhile_non_volatile_Writesb)

    from read_only_takeWhile_dropWhile_share_all_until_volatile_write [OF i_bound tssb_i]
    have read_only_drop:
      "read_only (share ?drop_sb 𝒮) ⊆ read_only (share sb 𝒮sb)"
      by (simp add: 𝒮)
  *)  
    from sbh_step 
    show ?thesis
    proof (cases)
      case (SBHReadBuffered a v volatile t)
      then obtain 
	"issb": "issb = Read volatile a t # issb'" and
	𝒪sb': "𝒪sb'=𝒪sb" and 
	𝒟sb': "𝒟sb'=𝒟sb" and
	θsb': "θsb' = θsb(tv)" and
	sb': "sb'=sb@[Readsb volatile a t v]" and
	msb': "msb' = msb" and
	𝒮sb': "𝒮sb'=𝒮sb" and
        sb': "sb'=sb" and
	buf_v: "buffered_val sb a = Some v" 
	by auto


      from safe_memop_flush_sb [simplified issb]  
      obtain access_cond': "a  acquired True sb 𝒪sb  
	a  read_only (share ?drop_sb 𝒮)  
	(volatile  a  dom (share ?drop_sb 𝒮))" and
	volatile_clean: "volatile  ¬ 𝒟sb" and
        rels_cond: "j < length ts. ij  released (ts!j) a  Some False" and
        rels_nv_cond: "¬volatile  (j < length ts. ij  a  dom (released (ts!j)))"
	by cases auto

      from clean_no_outstanding_volatile_Writesb [OF i_bound tssb_i] volatile_clean
      have volatile_cond: "volatile  outstanding_refs is_volatile_Writesb sb ={}"
	by auto
      
      from buffered_val_witness [OF buf_v] obtain volatile' sop' A' L' R' W'
	where
	witness: "Writesb volatile' a sop' v A' L' R' W'  set sb"
	by auto

      (* FIXME: since this is the buffered-val case, there should be a simpler proof not involving simulation to an
         unsafe state. Then we would not have to repeat the proof.*)

      {
	fix j pj "issbj" 𝒪j j 𝒟sbj θsbj sbj
	assume j_bound: "j < length tssb"
	assume neq_i_j: "i  j"
	assume jth: "tssb!j = (pj,issbj, θsbj, sbj, 𝒟sbj, 𝒪j,j)"
	assume non_vol: "¬ volatile"
	have "a  𝒪j  all_acquired sbj"
	proof 
	  assume a_j: "a  𝒪j  all_acquired sbj"
	  let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	  let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"


          from ts_sim [rule_format, OF j_bound] jth
	  obtain suspendsj "isj" 𝒟j where
	    suspendsj: "suspendsj = ?drop_sbj" and
	    isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	    𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
	    tsj: "ts!j = (hd_prog pj suspendsj, isj, 
	    θsbj |` (dom θsbj - read_tmps suspendsj),(), 
	    𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	    by (auto simp add: Let_def)
	    

	  from a_j ownership_distinct [OF i_bound j_bound neq_i_j tssb_i jth]
	  have a_notin_sb: "a  𝒪sb  all_acquired sb"
	    by auto
	  with acquired_all_acquired [of True sb 𝒪sb]
	  have a_not_acq: "a  acquired True sb 𝒪sb" 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 𝒪sb]
            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] tsj
          have "a  dom (release ?take_sbj (dom (𝒮sb)) j)"
            by auto
          with dom_release_takeWhile [of sbj "(dom (𝒮sb))" j]
          obtain
            a_relsj: "a  dom j" and
            a_sharedj: "a  all_shared ?take_sbj"
            by auto
          
          
          have "a  ((λ(_, _, _, sb, _, _, _). all_shared (takeWhile (Not  is_volatile_Writesb) sb)) `
                 set tssb)"
          proof -
            {
              fix k pk "isk" θk sbk 𝒟k 𝒪k k 
              assume k_bound: "k < length tssb" 
              assume ts_k: "tssb ! k = (pk,isk,θk,sbk,𝒟k,𝒪k,k)" 
              assume a_in: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sbk)"
              have False
              proof (cases "k=j")
                case True with a_sharedj 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 sbk)" 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_Writesb) sbk" 
                  "dropWhile (Not  is_volatile_Writesb) sbk"] 
                  all_shared_append [of "takeWhile (Not  is_volatile_Writesb) sbk" 
                  "dropWhile (Not  is_volatile_Writesb) sbk"] 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 tssb 𝒮sb]
          have a_ro_sharedsb: "a  read_only 𝒮sb"
            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 tssb" by fact
	  have share_consis: "sharing_consis 𝒮sb tssb" by fact
	  from sharing_consistent_share_all_until_volatile_write [OF own_dist share_consis i_bound tssb_i]
	  have consis': "sharing_consistent 𝒮 (acquired True ?take_sb 𝒪sb) ?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_sbj 𝒮sb)"
            by (auto simp add: domIff 𝒮 read_only_def)
          from sharing_consis [OF j_bound jth]
          have "sharing_consistent 𝒮sb 𝒪j sbj".
          from sharing_consistent_weak_sharing_consistent [OF this] weak_sharing_consistent_append [of 𝒪j ?take_sbj ?drop_sbj]
          have weak_consis_drop:"weak_sharing_consistent 𝒪j ?take_sbj"
            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_sharedj
          have "a  all_acquired ?take_sbj"
            by auto
	  with a_j a_notin_owns_j
	  have a_drop: "a  all_acquired ?drop_sbj"
	    using all_acquired_append [of ?take_sbj ?drop_sbj]
	    by simp
	  

	  from i_bound j_bound leq have j_bound_ts': "j < length ?ts'"
	    by auto

	  note conflict_drop = a_drop [simplified suspendsj [symmetric]]
	  from split_all_acquired_in [OF conflict_drop]
	    (* FIXME: exract common parts *)
	  show False
	  proof 
	    assume "sop a' v ys zs A L R W. 
              (suspendsj = ys @ Writesb True a' sop v A L R W# zs)  a  A"
	    then 
	    obtain a' sop' v' ys zs A' L' R' W' where
	      split_suspendsj: "suspendsj = ys @ Writesb True a' sop' v' A' L' R' W'# zs" 
	      (is "suspendsj = ?suspends") and
		a_A': "a  A'"
	      by blast

	    from sharing_consis [OF j_bound jth]
	    have "sharing_consistent 𝒮sb 𝒪j sbj".
	    then have A'_R': "A'  R' = {}" 
	      by (simp add: sharing_consistent_append [of _ _ ?take_sbj ?drop_sbj, simplified] 
		suspendsj [symmetric] split_suspendsj sharing_consistent_append)
	    from valid_program_history [OF j_bound jth] 
	    have "causal_program_history issbj sbj".
	    then have cph: "causal_program_history issbj ?suspends"
	      apply -
	      apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply (simp add: split_suspendsj)
	      done

	    from tsj neq_i_j j_bound 
	    have ts'_j: "?ts'!j = (hd_prog pj suspendsj, isj,
	      θsbj |` (dom θsbj - read_tmps suspendsj),(), 
	      𝒟j, acquired True ?take_sbj 𝒪j, release ?take_sbj (dom 𝒮sb) j)"
	      by auto
	    from valid_last_prog [OF j_bound jth] have last_prog: "last_prog pj sbj = pj".
	    then
	    have lp: "last_prog pj suspendsj = pj"
	      apply -
	      apply (rule last_prog_same_append [where sb="?take_sbj"])
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	    from valid_reads [OF j_bound jth]
	    have reads_consis_j: "reads_consistent False 𝒪j msb sbj".
	    from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb j_bound 
	      jth reads_consis_j]
	    have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
	      by (simp add: m suspendsj)


	    from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j tssb_i jth]
	    have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
	      by (simp add: suspendsj)
	    from reads_consistent_flush_independent [OF this reads_consis_m_j]
	    have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	      (flush ?drop_sb m) suspendsj".
	    hence reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	      (flush ?drop_sb m) (ys@[Writesb True a' sop' v' A' L' R' W'])"
	      by (simp add: split_suspendsj reads_consistent_append)

	    from valid_write_sops [OF j_bound jth]
	    have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
	      valid_sops_drop: "sopwrite_sops (ys@[Writesb 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_sbj@suspendsj)"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain 
	      read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
	      distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
	      apply (simp only: split_suspendsj [symmetric] suspendsj) 
	      apply (simp only: distinct_read_tmps_append)
	      done

	    from valid_history [OF j_bound jth]
	    have h_consis: 
	      "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	    
	    have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	    proof -
	      from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		by simp
	      from last_prog_hd_prog_append' [OF h_consis] this
	      have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		by (simp only: split_suspendsj [symmetric] suspendsj) 
	      moreover 
	      have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		last_prog (hd_prog pj suspendsj) ?take_sbj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		by (rule last_prog_hd_prog_append)
	      ultimately show ?thesis
		by (simp add: split_suspendsj [symmetric] suspendsj) 
	    qed

	    from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop 
	      h_consis] last_prog_hd_prog
	    have hist_consis': "history_consistent θsbj (hd_prog pj suspendsj) suspendsj"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    from reads_consistent_drop_volatile_writes_no_volatile_reads  
	    [OF reads_consis_j] 
	    have no_vol_read: "outstanding_refs is_volatile_Readsb 
	      (ys@[Writesb True a' sop' v' A' L' R' W']) = {}"
	      by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )

	    have acq_simp:
	      "acquired True (ys @ [Writesb True a' sop' v' A' L' R' W']) 
              (acquired True ?take_sbj 𝒪j) = 
              acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R'"
	      by (simp add: acquired_append)

	    from flush_store_buffer_append [where sb="ys@[Writesb True a' sop' v' A' L' R' W']" and sb'="zs", simplified,
	      OF j_bound_ts' isj [simplified split_suspendsj] cph [simplified suspendsj]
	      ts'_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys 
	      hist_consis' [simplified split_suspendsj] valid_sops_drop 
	      distinct_read_tmps_drop [simplified split_suspendsj] 
	      no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="share ?drop_sb 𝒮"]

	    obtain isj' j' where
	      isj': "instrs zs @ issbj = isj' @ prog_instrs zs" and
	      steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮)  d* 
	      (?ts'[j:=(last_prog
              (hd_prog pj (Writesb True a' sop' v' A' L' R' W'# zs)) (ys@[Writesb True a' sop' v' A' L' R' W']),
              isj',
              θsbj |` (dom θsbj - read_tmps zs),
              (), True, acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R',j')],
              flush (ys@[Writesb True a' sop' v' A' L' R' W']) (flush ?drop_sb m),
              share (ys@[Writesb 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 = (psb, issb, θsb,(), 
	      𝒟sb, acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb)"
	      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 issb] non_vol a_not_acq
	    have "a  read_only (share (ys@[Writesb True a' sop' v' A' L' R' W']) (share ?drop_sb 𝒮))"
	      apply cases
	      apply (auto simp add: Let_def issb)
	      done

	    with a_A'
	    show False
	      by (simp add: share_append in_read_only_convs)
	  next
	    assume "A L R W ys zs. suspendsj = ys @ Ghostsb A L R W # zs  a  A"
	    then 
	    obtain A' L' R' W' ys zs where
	      split_suspendsj: "suspendsj = ys @ Ghostsb A' L' R' W'# zs" 
	      (is "suspendsj = ?suspends") and
		a_A': "a  A'"
	      by blast

	    from valid_program_history [OF j_bound jth] 
	    have "causal_program_history issbj sbj".
	    then have cph: "causal_program_history issbj ?suspends"
	      apply -
	      apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply (simp add: split_suspendsj)
	      done

	    from tsj neq_i_j j_bound 
	    have ts'_j: "?ts'!j = (hd_prog pj suspendsj, isj,
	      θsbj |` (dom θsbj - read_tmps suspendsj),(), 
	      𝒟j, acquired True ?take_sbj 𝒪j, release ?take_sbj (dom 𝒮sb) j)"
	      by auto
	    from valid_last_prog [OF j_bound jth] have last_prog: "last_prog pj sbj = pj".
	    then
	    have lp: "last_prog pj suspendsj = pj"
	      apply -
	      apply (rule last_prog_same_append [where sb="?take_sbj"])
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done


	    from valid_reads [OF j_bound jth]
	    have reads_consis_j: "reads_consistent False 𝒪j msb sbj".
	    from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb j_bound 
	      jth reads_consis_j]
	    have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
	      by (simp add: m suspendsj)

	    from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j tssb_i jth]
	    have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
	      by (simp add: suspendsj)
	    from reads_consistent_flush_independent [OF this reads_consis_m_j]
	    have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	      (flush ?drop_sb m) suspendsj".

	    hence reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  
	      (flush ?drop_sb m) (ys@[Ghostsb A' L' R' W'])"
	      by (simp add: split_suspendsj reads_consistent_append)
	    from valid_write_sops [OF j_bound jth]
	    have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
	      valid_sops_drop: "sopwrite_sops (ys@[Ghostsb 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_sbj@suspendsj)"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain 
	      read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
	      distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
	      apply (simp only: split_suspendsj [symmetric] suspendsj) 
	      apply (simp only: distinct_read_tmps_append)
	      done

	    from valid_history [OF j_bound jth]
	    have h_consis: 
	      "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	    
	    have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	    proof -
	      from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		by simp
	      from last_prog_hd_prog_append' [OF h_consis] this
	      have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		by (simp only: split_suspendsj [symmetric] suspendsj) 
	      moreover 
	      have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		last_prog (hd_prog pj suspendsj) ?take_sbj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		by (rule last_prog_hd_prog_append)
	      ultimately show ?thesis
		by (simp add: split_suspendsj [symmetric] suspendsj) 
	    qed

	    from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop 
	      h_consis] last_prog_hd_prog
	    have hist_consis': "history_consistent θsbj (hd_prog pj suspendsj) suspendsj"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    from reads_consistent_drop_volatile_writes_no_volatile_reads  
	    [OF reads_consis_j] 
	    have no_vol_read: "outstanding_refs is_volatile_Readsb 
	      (ys@[Ghostsb A' L' R' W']) = {}"
	      by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )

	    have acq_simp:
	      "acquired True (ys @ [Ghostsb A' L' R' W']) 
              (acquired True ?take_sbj 𝒪j) = 
              acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R'"
	      by (simp add: acquired_append)

	    from flush_store_buffer_append [where sb="ys@[Ghostsb A' L' R' W']" and sb'="zs", simplified,
	      OF j_bound_ts' isj [simplified split_suspendsj] cph [simplified suspendsj]
	      ts'_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys 
	      hist_consis' [simplified split_suspendsj] valid_sops_drop 
	      distinct_read_tmps_drop [simplified split_suspendsj] 
	      no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="share ?drop_sb 𝒮"]
	    obtain isj' j' where
	      isj': "instrs zs @ issbj = isj' @ prog_instrs zs" and
	      steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮)  d* 
	      (?ts'[j:=(last_prog
              (hd_prog pj (Ghostsb A' L' R' W'# zs)) (ys@[Ghostsb A' L' R' W']),
              isj',
              θsbj |` (dom θsbj - read_tmps zs),
              (),
              𝒟j  outstanding_refs is_volatile_Writesb (ys @ [Ghostsb A' L' R' W'])  {}, acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R',j')],
              flush (ys@[Ghostsb A' L' R' W']) (flush ?drop_sb m),
              share (ys@[Ghostsb 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 = (psb, issb,θsb,(), 
	      𝒟sb, acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb)"
	      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 issb] non_vol a_not_acq
	    have "a  read_only (share (ys@[Ghostsb A' L' R' W']) (share ?drop_sb 𝒮))"
	      apply cases
	      apply (auto simp add: Let_def issb)
	      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_Writesb) sb) 𝒮)"
        assume nv: "¬ volatile"
        have "a  read_only (share sb 𝒮sb)"
        proof (cases "a  𝒪sb  all_acquired sb")
          case True
          from share_all_until_volatile_write_thread_local' [OF ownership_distinct_tssb 
            sharing_consis_tssb i_bound tssb_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 𝒪sb] 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_Writesb) sb)) ` set tssb)"
          proof -
            {
              fix k pk "isk" θk sbk 𝒟k 𝒪k k 
              assume k_bound: "k < length tssb" 
              assume ts_k: "tssb ! k = (pk,isk,θk,sbk,𝒟k,𝒪k,k)" 
              assume a_in: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sbk)"
              have False
              proof (cases "k=i")
                case True with False tssb_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_Writesb) sbk" 
                  "dropWhile (Not  is_volatile_Writesb) sbk"] 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_Writesb) sbk) (dom (𝒮sb)) k)"
                  by (auto simp add: Let_def)
                with dom_release_takeWhile [of sbk "(dom (𝒮sb))" k]
                obtain
                  a_relsj: "a  dom k" and
                  a_sharedj: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sbk)"
                  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 tssb 𝒮sb] a_ro_shared
          have "a  read_only 𝒮sb"
            by (auto simp add: 𝒮)
          from read_only_share_unowned' [OF weak_consis_sb read_only_unowned [OF i_bound tssb_i] False  this]
          show ?thesis .
        qed
      } note non_vol_ro_reduction = this

      have valid_own': "valid_ownership 𝒮sb' tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	proof (cases volatile)
	  case False
	  from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i]
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb sb".
	  then

	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb (sb@[Readsb 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 tssb' sb' 𝒪sb' 𝒮sb')
	next
	  case True
	  from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i]  
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb sb".
	  then
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb (sb@[Readsb 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 tssb' sb' 𝒪sb' 𝒮sb')
	qed
      next
	show "outstanding_volatile_writes_unowned_by_others tssb'"
	proof -
	  have out: "outstanding_refs is_volatile_Writesb (sb @ [Readsb volatile a t v])  
            outstanding_refs is_volatile_Writesb sb"
	    by (auto simp add: outstanding_refs_append)
	  have "all_acquired (sb @ [Readsb 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 tssb_i out this]
	  show ?thesis by (simp add: tssb' sb' 𝒪sb')
	qed
      next
	show "read_only_reads_unowned tssb'"
	proof (cases volatile)
	  case True
	  have r: "read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) (sb @ [Readsb volatile a t v])) 𝒪sb)
                    (dropWhile (Not  is_volatile_Writesb) (sb @ [Readsb volatile a t v]))
                 read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪sb)
                    (dropWhile (Not  is_volatile_Writesb) sb)"
	    apply (case_tac "outstanding_refs (is_volatile_Writesb) sb = {}")
	    apply (simp_all add: outstanding_vol_write_take_drop_appends
	    acquired_append read_only_reads_append True)
	    done

	  have "𝒪sb  all_acquired (sb @ [Readsb volatile a t v])  𝒪sb  all_acquired sb"
	    by (simp add: all_acquired_append)

	 
	  from  read_only_reads_unowned_nth_update [OF i_bound tssb_i r this]
	  show ?thesis
	    by (simp add: tssb' 𝒪sb' sb')
	next
	  case False
	  show ?thesis
	  proof (unfold_locales)
	    fix n m
	    fix pn "isn" 𝒪n n 𝒟n θn sbn pm "ism" 𝒪m m 𝒟m θm sbm
	    assume n_bound: "n < length tssb'"
	    and m_bound: "m < length tssb'"
	    and neq_n_m: "nm"
	    and nth: "tssb'!n = (pn, isn, θn, sbn, 𝒟n, 𝒪n, n)"
	    and mth: "tssb'!m =(pm, ism, θm, sbm, 𝒟m, 𝒪m, m)"
	    from n_bound have n_bound': "n < length tssb" by (simp add: tssb')
	    from m_bound have m_bound': "m < length tssb" by (simp add: tssb')

	    have acq_eq: "(𝒪sb'  all_acquired sb') = (𝒪sb  all_acquired sb)"
	      by (simp add: all_acquired_append sb' 𝒪sb')	      

	    show "(𝒪m  all_acquired sbm) 
              read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sbn) 𝒪n)
              (dropWhile (Not  is_volatile_Writesb) sbn) =
              {}"
	    proof (cases "m=i")
	      case True
	      with neq_n_m have neq_n_i: "ni"
		by auto
		
	      with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n, n)"
		by (auto simp add: tssb')
	      note read_only_reads_unowned [OF n_bound' i_bound  neq_n_i nth' tssb_i]
	      moreover
	      note acq_eq
	      ultimately show ?thesis
		using True tssb_i nth mth n_bound' m_bound'
		by (simp add: tssb')
	    next
	      case False
	      note neq_m_i = this
	      with m_bound mth i_bound have mth': "tssb!m = (pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
		by (auto simp add: tssb')
	      show ?thesis
	      proof (cases "n=i")
		case True
		note read_only_reads_unowned [OF i_bound m_bound' neq_m_i [symmetric] tssb_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 tssb_i nth mth n_bound' m_bound' neq_m_i
		  apply (case_tac "outstanding_refs (is_volatile_Writesb) sb = {}")
		  apply (clarsimp simp add: outstanding_vol_write_take_drop_appends
		    acquired_append read_only_reads_append tssb' sb' 𝒪sb')+
		  done
	      next
		case False
		with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n, n)"
		  by (auto simp add: tssb')
		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 tssb'"
	proof -
	  have "all_acquired (sb @ [Readsb 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 tssb_i this]
	  show ?thesis by (simp add: tssb' sb' 𝒪sb')
	qed
      qed


      have valid_hist': "valid_history program_step tssb'"
      proof -
	from valid_history [OF i_bound tssb_i]
	have hcons: "history_consistent θsb (hd_prog psb sb) sb".
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i]
	have t_notin_reads: "t  read_tmps sb"
	  by (auto simp add: "issb")
	from load_tmps_write_tmps_distinct [OF i_bound tssb_i]
	have t_notin_writes: "t  (fst ` write_sops sb)"
	  by (auto simp add: "issb")

	from valid_write_sops [OF i_bound tssb_i]
	have valid_sops: "sop  write_sops sb. valid_sop sop"
	  by auto
	from load_tmps_fresh [OF i_bound tssb_i]
	have t_fresh: "t  dom θsb"
	  using "issb"
	  by simp
	have "history_consistent (θsb(tv)) 
	       (hd_prog psb (sb@ [Readsb volatile a t v])) (sb@ [Readsb volatile a t v])"
	  using t_notin_writes valid_sops t_fresh hcons
	  valid_implies_valid_prog_hd [OF i_bound tssb_i valid]
	  apply -
	  apply (rule history_consistent_appendI)
	  apply (auto simp add: hd_prog_append_Readsb)
	  done
	from valid_history_nth_update [OF i_bound this]
	show ?thesis
	  by (auto simp add: tssb' sb' 𝒪sb' θsb')
      qed

      from reads_consistent_buffered_snoc [OF buf_v valid_reads [OF i_bound tssb_i] 
	volatile_cond] 
      have reads_consis': "reads_consistent False 𝒪sb msb (sb @ [Readsb volatile a t v])"
	by (simp split: if_split_asm)

      from valid_reads_nth_update [OF i_bound this]
      have valid_reads': "valid_reads msb tssb'" by (simp add: tssb' sb' 𝒪sb')

      have valid_sharing': "valid_sharing 𝒮sb' tssb'"
      proof (intro_locales)	
	from outstanding_non_volatile_writes_unshared [OF i_bound tssb_i]
	have "non_volatile_writes_unshared 𝒮sb (sb @ [Readsb 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 𝒮sb' tssb'"
	  by (simp add: tssb' sb' 𝒮sb')
      next
	from sharing_consis [OF i_bound tssb_i]
	have "sharing_consistent 𝒮sb 𝒪sb sb".
	then
	have "sharing_consistent 𝒮sb 𝒪sb (sb @ [Readsb volatile a t v])"
	  by (simp add:  sharing_consistent_append)
	from sharing_consis_nth_update [OF i_bound this]
	show "sharing_consis 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒪sb' sb' 𝒮sb')
      next
	note read_only_unowned [OF i_bound tssb_i]
	from read_only_unowned_nth_update [OF i_bound this]
	show "read_only_unowned 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' sb' 𝒪sb')
      next
	from unowned_shared_nth_update [OF i_bound tssb_i subset_refl]
	show "unowned_shared 𝒮sb' tssb'" by (simp add: tssb' 𝒪sb' 𝒮sb')
      next
	from no_outstanding_write_to_read_only_memory [OF i_bound tssb_i]
	have "no_write_to_read_only_memory 𝒮sb sb".
	hence "no_write_to_read_only_memory 𝒮sb (sb@[Readsb 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 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒮sb' sb')
      qed

      have tmps_distinct': "tmps_distinct tssb'"
      proof (intro_locales)
	from load_tmps_distinct [OF i_bound tssb_i]
	have "distinct_load_tmps issb'"
	  by (auto split: instr.splits simp add: issb)
	from load_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_distinct tssb'" by (simp add: tssb')
      next
	from read_tmps_distinct [OF i_bound tssb_i]
	have "distinct_read_tmps sb".
	moreover
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i]
	have "t  read_tmps sb"
	  by (auto simp add: issb)
	ultimately have "distinct_read_tmps (sb @ [Readsb 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 tssb'" by (simp add: tssb' sb')
      next
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i] 
          load_tmps_distinct [OF i_bound tssb_i]
	have "load_tmps issb'  read_tmps (sb @ [Readsb volatile a t v]) = {}"
	  by (clarsimp simp add: read_tmps_append "issb")
	from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_read_tmps_distinct tssb'" by (simp add: tssb' sb')
      qed

      have valid_sops': "valid_sops tssb'"
      proof -
	from valid_store_sops [OF i_bound tssb_i]
	have valid_store_sops': "sopstore_sops issb'. valid_sop sop"
	  by (auto simp add: "issb")
	from valid_write_sops [OF i_bound tssb_i]
	have valid_write_sops': "sopwrite_sops (sb@ [Readsb 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: tssb' sb')
      qed

      have valid_dd': "valid_data_dependency tssb'"
      proof -
	from data_dependency_consistent_instrs [OF i_bound tssb_i]
	have dd_is: "data_dependency_consistent_instrs (dom θsb') issb'"
	  by (auto simp add: "issb" θsb')
	from load_tmps_write_tmps_distinct [OF i_bound tssb_i]
	have "load_tmps issb'  (fst ` write_sops (sb@ [Readsb volatile a t v])) = {}"
	  by (auto simp add: write_sops_append "issb")
	from valid_data_dependency_nth_update [OF i_bound dd_is this]
	show ?thesis by (simp add: tssb' sb')
      qed

      have load_tmps_fresh': "load_tmps_fresh tssb'"
      proof -
	from load_tmps_fresh [OF i_bound tssb_i] 
	have "load_tmps (Read volatile a t # issb')  dom θsb = {}"
	  by (simp add: "issb")
	moreover
	from load_tmps_distinct [OF i_bound tssb_i] have "t  load_tmps issb'"
	  by (auto simp add: "issb")
	ultimately have "load_tmps issb'  dom (θsb(t  v)) = {}"
	  by auto
	from load_tmps_fresh_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' θsb')
      qed

      have enough_flushs': "enough_flushs tssb'"
      proof -
	from clean_no_outstanding_volatile_Writesb [OF i_bound tssb_i]
	have "¬ 𝒟sb  outstanding_refs is_volatile_Writesb (sb@[Readsb 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: tssb' sb' 𝒟sb')
      qed
	
      have valid_program_history': "valid_program_history tssb'"
      proof -	
	from valid_program_history [OF i_bound tssb_i]
	have "causal_program_history issb sb" .
	then have causal': "causal_program_history issb' (sb@[Readsb volatile a t v])"
	  by (auto simp: causal_program_history_Read  "issb")
	from valid_last_prog [OF i_bound tssb_i]
	have "last_prog psb sb = psb".
	hence "last_prog psb (sb @ [Readsb volatile a t v]) = psb"
	  by (simp add: last_prog_append_Readsb)

	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb' sb')
      qed
      show ?thesis
      proof (cases "outstanding_refs is_volatile_Writesb sb = {}")
	case True 

	from True have flush_all: "takeWhile (Not  is_volatile_Writesb) sb = sb"
	  by (auto simp add: outstanding_refs_conv )

	from True have suspend_nothing: "dropWhile (Not  is_volatile_Writesb) 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 # issb'"
	  by (simp add: "issb")
	with suspends_empty ts_i 
	have ts_i: "ts!i = (psb, Read volatile a t # issb', θsb,(), 𝒟, acquired True ?take_sb 𝒪sb, release ?take_sb (dom 𝒮sb) sb)"
	  by simp

	from direct_memop_step.Read 
	have "(Read volatile a t # issb', θsb, (), m, 𝒟, acquired True ?take_sb 𝒪sb,
                release ?take_sb (dom 𝒮sb) sb, 𝒮)  
          (issb', θsb(t  m a), (), m, 𝒟, acquired True ?take_sb 𝒪sb,release ?take_sb (dom 𝒮sb) sb, 𝒮)".
	from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
	have "(ts, m, 𝒮) d (ts[i := (psb, issb',  θsb(t  m a), (),
               𝒟, acquired True ?take_sb 𝒪sb, release ?take_sb (dom 𝒮sb) sb)], m, 𝒮)" . 

	moreover
	
	from flush_all_until_volatile_write_Read_commute [OF i_bound tssb_i [simplified "issb"] ]
	have flush_commute: "flush_all_until_volatile_write
          (tssb[i := (psb,issb', 
               θsb(tv), sb @ [Readsb volatile a t v], 𝒟sb, 𝒪sb, sb)]) msb =
          flush_all_until_volatile_write tssb msb".

	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 tssb_i]]
	have a_owned: "a  𝒪sb  all_acquired sb  read_only_reads 𝒪sb sb"
	  by auto

	have "flush_all_until_volatile_write tssb msb a = v"
	proof - (* FIXME: Same proof as in Unbuffered case *)
          have "j < length tssb. i  j 
                  (let (_,_,_,sbj,_,_,_) = tssb!j 
                  in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))"
	  proof -
	    {
	      fix j pj "isj" 𝒪j j 𝒟j xsj sbj
	      assume j_bound: "j < length tssb"
	      assume neq_i_j: "i  j"
	      assume jth: "tssb!j = (pj,isj, xsj, sbj, 𝒟j, 𝒪j, j)"
	      have "a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	      proof 
		let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
		let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"
		assume a_in: "a  outstanding_refs is_non_volatile_Writesb ?take_sbj"
		with outstanding_refs_takeWhile [where P'= "Not  is_volatile_Writesb"]
		have a_in': "a  outstanding_refs is_non_volatile_Writesb sbj"
		  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 sbj"
		  by auto
		with ownership_distinct [OF i_bound j_bound neq_i_j tssb_i jth]
		have a_not_owns: "a  𝒪sb  all_acquired sb"
		  by blast
		

		from non_volatile_owned_or_read_only_append [of False 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
		  outstanding_non_volatile_refs_owned_or_read_only [OF j_bound jth]
		have "non_volatile_owned_or_read_only False 𝒮sb 𝒪j ?take_sbj"
		  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_sbj"
		  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_sbj (dom (𝒮sb)) j  a  Some False"
                  by (auto simp add: Let_def)

	
		{
		  assume "a  acquired True sb 𝒪sb"
		  with acquired_all_acquired_in [OF this] ownership_distinct [OF i_bound j_bound neq_i_j tssb_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 𝒪sb]
                  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_tssb sharing_consis_tssb j_bound jth j_owns]
                  have "a  read_only (share ?take_sbj 𝒮sb)"
                    by (auto simp add: read_only_def 𝒮)
                  hence a_dom: "a  dom  (share ?take_sbj 𝒮sb)"
                    by (auto simp add: read_only_def domIff)
                  from outstanding_non_volatile_writes_unshared [OF j_bound jth]
                  non_volatile_writes_unshared_append [of 𝒮sb ?take_sbj ?drop_sbj]
                  have nvw: "non_volatile_writes_unshared 𝒮sb ?take_sbj" 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 𝒮sb sbj".
		  with non_volatile_writes_unshared_append [of 𝒮sb "?take_sbj"
		  "?drop_sbj"]
		  have unshared_take: "non_volatile_writes_unshared 𝒮sb (takeWhile (Not  is_volatile_Writesb) sbj)" 
		    by clarsimp
		   
		  from valid_own have own_dist: "ownership_distinct tssb"
		    by (simp add: valid_ownership_def)
		  from valid_sharing have "sharing_consis 𝒮sb tssb"
		    by (simp add: valid_sharing_def)
		  from sharing_consistent_share_all_until_volatile_write [OF own_dist this i_bound tssb_i]
		  have sc: "sharing_consistent 𝒮 (acquired True ?take_sb 𝒪sb) ?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 𝒪sb" by auto
		  also from acquired_all_acquired all_acquired_takeWhile 
		  have "  dom 𝒮  (𝒪sb  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_tssb sharing_consis_tssb j_bound jth j_owns]
                  have a_dom: "a  dom  (share ?take_sbj 𝒮sb)"
                    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 tssb_i this]
	  show ?thesis
	    by (simp add: buf_v)
	qed


	hence m_a_v: "m a = v"
	  by (simp add: m)
	
	have tmps_commute: "θsb(t  v) = (θsb |` (dom θsb - {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_Writesb) sb') = []"
	  by (simp add: sb')

	from 𝒟
	have 𝒟': "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb (sb@[Readsb volatile a t v])   {})"
	  by (auto simp: outstanding_refs_append)

	have "(tssb',msb,𝒮sb')  (ts[i := (psb,issb',
                θsb(tm a),(),𝒟, acquired True ?take_sb 𝒪sb,
                release ?take_sb (dom 𝒮sb) sb)],m,𝒮)"
	  apply (rule sim_config.intros) 
	  apply    (simp add: m flush_commute tssb' 𝒪sb' θsb' sb' 𝒟sb' sb')
	  using   share_all_until_volatile_write_Read_commute [OF i_bound tssb_i [simplified issb]]
	  apply   (simp add: 𝒮 𝒮sb' tssb' sb' 𝒪sb' θsb' sb')
	  using  leq
	  apply  (simp add: tssb')
	  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  tssb' 𝒪sb' 𝒮sb' θsb' sb' sb' suspend_nothing' 
	    𝒟sb' 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'
	    msb' 𝒮sb' 𝒪sb'
	  by (auto simp del: fun_upd_apply )
      next
	case False

	then obtain r where r_in: "r  set sb" and volatile_r: "is_volatile_Writesb r"
	  by (auto simp add: outstanding_refs_conv)
	from takeWhile_dropWhile_real_prefix 
	[OF r_in, of  "(Not  is_volatile_Writesb)", simplified, OF volatile_r] 
	obtain a' v' sb'' sop' A' L' R' W' where
	  sb_split: "sb = takeWhile (Not  is_volatile_Writesb) sb @ Writesb True a' sop' v' A' L' R' W'# sb''" 
	  and
	  drop: "dropWhile (Not  is_volatile_Writesb) sb = Writesb 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 = Writesb 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 tssb_i 
	  [simplified "issb"] ]

	have flush_commute: "flush_all_until_volatile_write
             (tssb[i := (psb,issb', θsb(t  v), sb @ [Readsb volatile a t v], 𝒟sb, 𝒪sb, sb)]) msb =
             flush_all_until_volatile_write tssb msb".

	have "Writesb True a' sop' v' A' L' R' W' set sb"
	  by (subst sb_split) auto
	
	from dropWhile_append1 [OF this, of "(Not  is_volatile_Writesb)"]
	have drop_app_comm:
	  "(dropWhile (Not  is_volatile_Writesb) (sb @ [Readsb volatile a t v])) =
                dropWhile (Not  is_volatile_Writesb) sb @ [Readsb volatile a t v]"
	  by simp

	from load_tmps_fresh [OF i_bound tssb_i]
	have "t  dom θsb"
	  by (auto simp add: "issb")
	then have tmps_commute: 
	  "θsb |` (dom θsb - read_tmps sb'') =
          θsb |` (dom θsb - insert t (read_tmps sb''))"
	  apply -
	  apply (rule ext)
	  apply auto
	  done

	from 𝒟
	have 𝒟': "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb (sb@[Readsb volatile a t v])   {})"
	  by (auto simp: outstanding_refs_append)

	have "(tssb',msb,𝒮sb)  (ts,m,𝒮)"
	  apply (rule sim_config.intros) 
	  apply    (simp add: m flush_commute tssb' 𝒪sb' sb' θsb' sb' 𝒟sb' )
	  using   share_all_until_volatile_write_Read_commute [OF i_bound tssb_i [simplified issb]]
	  apply   (simp add: 𝒮 𝒮sb' tssb' sb' 𝒪sb' sb' θsb')
	  using  leq
	  apply  (simp add: tssb')
	  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_Readsb instrs_append_Readsb 
	    hd_prog_append_Readsb
	    drop "issb" tssb' sb' 𝒪sb' sb' θsb' 𝒟sb' 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' msb' 𝒮sb' 
	  by (auto simp del: fun_upd_apply )
      qed
    next
      case (SBHReadUnbuffered a volatile t)
      then obtain 
	"issb": "issb = Read volatile a t # issb'" and
	𝒪sb': "𝒪sb'=𝒪sb" and
        sb': "sb'=sb" and
	θsb': "θsb' = θsb(t(msb a))" and
	sb': "sb'=sb@[Readsb volatile a t (msb a)]" and
	msb': "msb' = msb" and
	𝒮sb': "𝒮sb'=𝒮sb" and 
	𝒟sb': "𝒟sb'=𝒟sb" and
	buf_None: "buffered_val sb a = None" 

	by auto


      from safe_memop_flush_sb [simplified issb]
      obtain access_cond': "a  acquired True sb 𝒪sb  
	a  read_only (share ?drop_sb 𝒮)  (volatile  a  dom (share ?drop_sb 𝒮))" and
	volatile_clean: "volatile  ¬ 𝒟sb" and
        rels_cond: "j < length ts. ij  released (ts!j) a  Some False" and
        rels_nv_cond: "¬volatile  (j < length ts. ij  a  dom (released (ts!j)))"
	by cases auto

      from clean_no_outstanding_volatile_Writesb [OF i_bound tssb_i] volatile_clean
      have volatile_cond: "volatile  outstanding_refs is_volatile_Writesb sb ={}"
	by auto

      {
	fix j pj "issbj" 𝒪j j 𝒟sbj θsbj sbj
	assume j_bound: "j < length tssb"
	assume neq_i_j: "i  j"
	assume jth: "tssb!j = (pj,issbj, θsbj, sbj, 𝒟sbj, 𝒪j,j)"
	assume non_vol: "¬ volatile"
	have "a  𝒪j  all_acquired sbj"
	proof 
	  assume a_j: "a  𝒪j  all_acquired sbj"
	  let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	  let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"


          from ts_sim [rule_format, OF j_bound] jth
	  obtain suspendsj "isj" 𝒟j where
	    suspendsj: "suspendsj = ?drop_sbj" and
	    isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	    𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
	    tsj: "ts!j = (hd_prog pj suspendsj, isj, 
	    θsbj |` (dom θsbj - read_tmps suspendsj),(), 
	    𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	    by (auto simp add: Let_def)
	    

	  from a_j ownership_distinct [OF i_bound j_bound neq_i_j tssb_i jth]
	  have a_notin_sb: "a  𝒪sb  all_acquired sb"
	    by auto
	  with acquired_all_acquired [of True sb 𝒪sb]
	  have a_not_acq: "a  acquired True sb 𝒪sb" 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 𝒪sb]
            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] tsj
          have "a  dom (release ?take_sbj (dom (𝒮sb)) j)"
            by auto
          with dom_release_takeWhile [of sbj "(dom (𝒮sb))" j]
          obtain
            a_relsj: "a  dom j" and
            a_sharedj: "a  all_shared ?take_sbj"
            by auto
          
          
          have "a  ((λ(_, _, _, sb, _, _, _). all_shared (takeWhile (Not  is_volatile_Writesb) sb)) `
                 set tssb)"
          proof -
            {
              fix k pk "isk" θk sbk 𝒟k 𝒪k k 
              assume k_bound: "k < length tssb" 
              assume ts_k: "tssb ! k = (pk,isk,θk,sbk,𝒟k,𝒪k,k)" 
              assume a_in: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sbk)"
              have False
              proof (cases "k=j")
                case True with a_sharedj 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 sbk)" 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_Writesb) sbk" 
                  "dropWhile (Not  is_volatile_Writesb) sbk"] 
                  all_shared_append [of "takeWhile (Not  is_volatile_Writesb) sbk" 
                  "dropWhile (Not  is_volatile_Writesb) sbk"] 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 tssb 𝒮sb]
          have a_ro_sharedsb: "a  read_only 𝒮sb"
            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 tssb" by fact
	  have share_consis: "sharing_consis 𝒮sb tssb" by fact
	  from sharing_consistent_share_all_until_volatile_write [OF own_dist share_consis i_bound tssb_i]
	  have consis': "sharing_consistent 𝒮 (acquired True ?take_sb 𝒪sb) ?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_sbj 𝒮sb)"
            by (auto simp add: domIff 𝒮 read_only_def)
          from sharing_consis [OF j_bound jth]
          have "sharing_consistent 𝒮sb 𝒪j sbj".
          from sharing_consistent_weak_sharing_consistent [OF this] weak_sharing_consistent_append [of 𝒪j ?take_sbj ?drop_sbj]
          have weak_consis_drop:"weak_sharing_consistent 𝒪j ?take_sbj"
            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_sharedj
          have "a  all_acquired ?take_sbj"
            by auto
	  with a_j a_notin_owns_j
	  have a_drop: "a  all_acquired ?drop_sbj"
	    using all_acquired_append [of ?take_sbj ?drop_sbj]
	    by simp
	  

	  from i_bound j_bound leq have j_bound_ts': "j < length ?ts'"
	    by auto

	  note conflict_drop = a_drop [simplified suspendsj [symmetric]]
	  from split_all_acquired_in [OF conflict_drop]
	    (* FIXME: exract common parts *)
	  show False
	  proof 
	    assume "sop a' v ys zs A L R W. 
              (suspendsj = ys @ Writesb True a' sop v A L R W# zs)  a  A"
	    then 
	    obtain a' sop' v' ys zs A' L' R' W' where
	      split_suspendsj: "suspendsj = ys @ Writesb True a' sop' v' A' L' R' W'# zs" 
	      (is "suspendsj = ?suspends") and
		a_A': "a  A'"
	      by blast

	    from sharing_consis [OF j_bound jth]
	    have "sharing_consistent 𝒮sb 𝒪j sbj".
	    then have A'_R': "A'  R' = {}" 
	      by (simp add: sharing_consistent_append [of _ _ ?take_sbj ?drop_sbj, simplified] 
		suspendsj [symmetric] split_suspendsj sharing_consistent_append)
	    from valid_program_history [OF j_bound jth] 
	    have "causal_program_history issbj sbj".
	    then have cph: "causal_program_history issbj ?suspends"
	      apply -
	      apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply (simp add: split_suspendsj)
	      done

	    from tsj neq_i_j j_bound 
	    have ts'_j: "?ts'!j = (hd_prog pj suspendsj, isj,
	      θsbj |` (dom θsbj - read_tmps suspendsj),(), 
	      𝒟j, acquired True ?take_sbj 𝒪j, release ?take_sbj (dom 𝒮sb) j)"
	      by auto
	    from valid_last_prog [OF j_bound jth] have last_prog: "last_prog pj sbj = pj".
	    then
	    have lp: "last_prog pj suspendsj = pj"
	      apply -
	      apply (rule last_prog_same_append [where sb="?take_sbj"])
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	    from valid_reads [OF j_bound jth]
	    have reads_consis_j: "reads_consistent False 𝒪j msb sbj".
	    from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb j_bound 
	      jth reads_consis_j]
	    have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
	      by (simp add: m suspendsj)


	    from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j tssb_i jth]
	    have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
	      by (simp add: suspendsj)
	    from reads_consistent_flush_independent [OF this reads_consis_m_j]
	    have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	      (flush ?drop_sb m) suspendsj".
	    hence reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	      (flush ?drop_sb m) (ys@[Writesb True a' sop' v' A' L' R' W'])"
	      by (simp add: split_suspendsj reads_consistent_append)

	    from valid_write_sops [OF j_bound jth]
	    have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
	      valid_sops_drop: "sopwrite_sops (ys@[Writesb 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_sbj@suspendsj)"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain 
	      read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
	      distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
	      apply (simp only: split_suspendsj [symmetric] suspendsj) 
	      apply (simp only: distinct_read_tmps_append)
	      done

	    from valid_history [OF j_bound jth]
	    have h_consis: 
	      "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	    
	    have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	    proof -
	      from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		by simp
	      from last_prog_hd_prog_append' [OF h_consis] this
	      have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		by (simp only: split_suspendsj [symmetric] suspendsj) 
	      moreover 
	      have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		last_prog (hd_prog pj suspendsj) ?take_sbj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		by (rule last_prog_hd_prog_append)
	      ultimately show ?thesis
		by (simp add: split_suspendsj [symmetric] suspendsj) 
	    qed

	    from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop 
	      h_consis] last_prog_hd_prog
	    have hist_consis': "history_consistent θsbj (hd_prog pj suspendsj) suspendsj"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    from reads_consistent_drop_volatile_writes_no_volatile_reads  
	    [OF reads_consis_j] 
	    have no_vol_read: "outstanding_refs is_volatile_Readsb 
	      (ys@[Writesb True a' sop' v' A' L' R' W']) = {}"
	      by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )

	    have acq_simp:
	      "acquired True (ys @ [Writesb True a' sop' v' A' L' R' W']) 
              (acquired True ?take_sbj 𝒪j) = 
              acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R'"
	      by (simp add: acquired_append)

	    from flush_store_buffer_append [where sb="ys@[Writesb True a' sop' v' A' L' R' W']" and sb'="zs", simplified,
	      OF j_bound_ts' isj [simplified split_suspendsj] cph [simplified suspendsj]
	      ts'_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys 
	      hist_consis' [simplified split_suspendsj] valid_sops_drop 
	      distinct_read_tmps_drop [simplified split_suspendsj] 
	      no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="share ?drop_sb 𝒮"]

	    obtain isj' j' where
	      isj': "instrs zs @ issbj = isj' @ prog_instrs zs" and
	      steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮)  d* 
	      (?ts'[j:=(last_prog
              (hd_prog pj (Writesb True a' sop' v' A' L' R' W'# zs)) (ys@[Writesb True a' sop' v' A' L' R' W']),
              isj',
              θsbj |` (dom θsbj - read_tmps zs),
              (), True, acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R',j')],
              flush (ys@[Writesb True a' sop' v' A' L' R' W']) (flush ?drop_sb m),
              share (ys@[Writesb 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 = (psb, issb, θsb,(), 
	      𝒟sb, acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb)"
	      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 issb] non_vol a_not_acq
	    have "a  read_only (share (ys@[Writesb True a' sop' v' A' L' R' W']) (share ?drop_sb 𝒮))"
	      apply cases
	      apply (auto simp add: Let_def issb)
	      done

	    with a_A'
	    show False
	      by (simp add: share_append in_read_only_convs)
	  next
	    assume "A L R W ys zs. suspendsj = ys @ Ghostsb A L R W # zs  a  A"
	    then 
	    obtain A' L' R' W' ys zs where
	      split_suspendsj: "suspendsj = ys @ Ghostsb A' L' R' W'# zs" 
	      (is "suspendsj = ?suspends") and
		a_A': "a  A'"
	      by blast

	    from valid_program_history [OF j_bound jth] 
	    have "causal_program_history issbj sbj".
	    then have cph: "causal_program_history issbj ?suspends"
	      apply -
	      apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply (simp add: split_suspendsj)
	      done

	    from tsj neq_i_j j_bound 
	    have ts'_j: "?ts'!j = (hd_prog pj suspendsj, isj,
	      θsbj |` (dom θsbj - read_tmps suspendsj),(), 
	      𝒟j, acquired True ?take_sbj 𝒪j, release ?take_sbj (dom 𝒮sb) j)"
	      by auto
	    from valid_last_prog [OF j_bound jth] have last_prog: "last_prog pj sbj = pj".
	    then
	    have lp: "last_prog pj suspendsj = pj"
	      apply -
	      apply (rule last_prog_same_append [where sb="?take_sbj"])
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done


	    from valid_reads [OF j_bound jth]
	    have reads_consis_j: "reads_consistent False 𝒪j msb sbj".
	    from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb j_bound 
	      jth reads_consis_j]
	    have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
	      by (simp add: m suspendsj)

	    from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j tssb_i jth]
	    have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
	      by (simp add: suspendsj)
	    from reads_consistent_flush_independent [OF this reads_consis_m_j]
	    have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	      (flush ?drop_sb m) suspendsj".

	    hence reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  
	      (flush ?drop_sb m) (ys@[Ghostsb A' L' R' W'])"
	      by (simp add: split_suspendsj reads_consistent_append)
	    from valid_write_sops [OF j_bound jth]
	    have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
	      valid_sops_drop: "sopwrite_sops (ys@[Ghostsb 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_sbj@suspendsj)"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain 
	      read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
	      distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
	      apply (simp only: split_suspendsj [symmetric] suspendsj) 
	      apply (simp only: distinct_read_tmps_append)
	      done

	    from valid_history [OF j_bound jth]
	    have h_consis: 
	      "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	    
	    have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	    proof -
	      from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		by simp
	      from last_prog_hd_prog_append' [OF h_consis] this
	      have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		by (simp only: split_suspendsj [symmetric] suspendsj) 
	      moreover 
	      have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		last_prog (hd_prog pj suspendsj) ?take_sbj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		by (rule last_prog_hd_prog_append)
	      ultimately show ?thesis
		by (simp add: split_suspendsj [symmetric] suspendsj) 
	    qed

	    from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop 
	      h_consis] last_prog_hd_prog
	    have hist_consis': "history_consistent θsbj (hd_prog pj suspendsj) suspendsj"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    from reads_consistent_drop_volatile_writes_no_volatile_reads  
	    [OF reads_consis_j] 
	    have no_vol_read: "outstanding_refs is_volatile_Readsb 
	      (ys@[Ghostsb A' L' R' W']) = {}"
	      by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )

	    have acq_simp:
	      "acquired True (ys @ [Ghostsb A' L' R' W']) 
              (acquired True ?take_sbj 𝒪j) = 
              acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R'"
	      by (simp add: acquired_append)

	    from flush_store_buffer_append [where sb="ys@[Ghostsb A' L' R' W']" and sb'="zs", simplified,
	      OF j_bound_ts' isj [simplified split_suspendsj] cph [simplified suspendsj]
	      ts'_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys 
	      hist_consis' [simplified split_suspendsj] valid_sops_drop 
	      distinct_read_tmps_drop [simplified split_suspendsj] 
	      no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="share ?drop_sb 𝒮"]
	    obtain isj' j' where
	      isj': "instrs zs @ issbj = isj' @ prog_instrs zs" and
	      steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮)  d* 
	      (?ts'[j:=(last_prog
              (hd_prog pj (Ghostsb A' L' R' W'# zs)) (ys@[Ghostsb A' L' R' W']),
              isj',
              θsbj |` (dom θsbj - read_tmps zs),
              (),
              𝒟j  outstanding_refs is_volatile_Writesb (ys @ [Ghostsb A' L' R' W'])  {}, acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R',j')],
              flush (ys@[Ghostsb A' L' R' W']) (flush ?drop_sb m),
              share (ys@[Ghostsb 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 = (psb, issb,θsb,(), 
	      𝒟sb, acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb)"
	      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 issb] non_vol a_not_acq
	    have "a  read_only (share (ys@[Ghostsb A' L' R' W']) (share ?drop_sb 𝒮))"
	      apply cases
	      apply (auto simp add: Let_def issb)
	      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_Writesb) sb) 𝒮)"
        assume nv: "¬ volatile"
        have "a  read_only (share sb 𝒮sb)"
        proof (cases "a  𝒪sb  all_acquired sb")
          case True
          from share_all_until_volatile_write_thread_local' [OF ownership_distinct_tssb 
            sharing_consis_tssb i_bound tssb_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 𝒪sb] 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_Writesb) sb)) ` set tssb)"
          proof -
            {
              fix k pk "isk" θk sbk 𝒟k 𝒪k k 
              assume k_bound: "k < length tssb" 
              assume ts_k: "tssb ! k = (pk,isk,θk,sbk,𝒟k,𝒪k,k)" 
              assume a_in: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sbk)"
              have False
              proof (cases "k=i")
                case True with False tssb_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_Writesb) sbk" 
                  "dropWhile (Not  is_volatile_Writesb) sbk"] 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_Writesb) sbk) (dom (𝒮sb)) k)"
                  by (auto simp add: Let_def)
                with dom_release_takeWhile [of sbk "(dom (𝒮sb))" k]
                obtain
                  a_relsj: "a  dom k" and
                  a_sharedj: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sbk)"
                  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 tssb 𝒮sb] a_ro_shared
          have "a  read_only 𝒮sb"
            by (auto simp add: 𝒮)

          from read_only_share_unowned' [OF weak_consis_sb read_only_unowned [OF i_bound tssb_i] False  this]
          show ?thesis .
        qed
      } note non_vol_ro_reduction = this

      have valid_own': "valid_ownership 𝒮sb' tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	proof (cases volatile)
	  case False
	  from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i]
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb sb".
	  then

	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb (sb@[Readsb False a t (msb 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 tssb' sb' 𝒪sb' 𝒮sb')
	next
	  case True
	  from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i]  
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb sb".
	  then
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb (sb@[Readsb True a t (msb 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 tssb' sb' 𝒪sb' 𝒮sb')
	qed
      next
	show "outstanding_volatile_writes_unowned_by_others tssb'"
	proof -
	  have out: "outstanding_refs is_volatile_Writesb (sb @ [Readsb volatile a t (msb a)])  
            outstanding_refs is_volatile_Writesb sb"
	    by (auto simp add: outstanding_refs_append)
	  have "all_acquired (sb @ [Readsb volatile a t (msb a)])  all_acquired sb"
	    by (auto simp add: all_acquired_append)
	  from outstanding_volatile_writes_unowned_by_others_store_buffer 
	  [OF i_bound tssb_i out this]
	  show ?thesis by (simp add: tssb' sb' 𝒪sb')
	qed
      next
	show "read_only_reads_unowned tssb'"
	proof (cases volatile)
	  case True
	  have r: "read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) 
	            (sb @ [Readsb volatile a t (msb a)])) 𝒪sb)
                    (dropWhile (Not  is_volatile_Writesb) (sb @ [Readsb volatile a t (msb a)]))
                 read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪sb)
                    (dropWhile (Not  is_volatile_Writesb) sb)"
	    apply (case_tac "outstanding_refs (is_volatile_Writesb) sb = {}")
	    apply (simp_all add: outstanding_vol_write_take_drop_appends
	    acquired_append read_only_reads_append True)
	    done

	  have "𝒪sb  all_acquired (sb @ [Readsb volatile a t (msb a)])  𝒪sb  all_acquired sb"
	    by (simp add: all_acquired_append)

	 
	  from  read_only_reads_unowned_nth_update [OF i_bound tssb_i r this]
	  show ?thesis
	    by (simp add: tssb' 𝒪sb' sb')
	next
	  case False
	  show ?thesis
	  proof (unfold_locales)
	    fix n m
	    fix pn "isn" 𝒪n n 𝒟n θn sbn pm "ism" 𝒪m m 𝒟m θm sbm
	    assume n_bound: "n < length tssb'"
	    and m_bound: "m < length tssb'"
	    and neq_n_m: "nm"
	    and nth: "tssb'!n = (pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
	    and mth: "tssb'!m =(pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
	    from n_bound have n_bound': "n < length tssb" by (simp add: tssb')
	    from m_bound have m_bound': "m < length tssb" by (simp add: tssb')

	    have acq_eq: "(𝒪sb'  all_acquired sb') = (𝒪sb  all_acquired sb)"
	      by (simp add: all_acquired_append sb' 𝒪sb')	      

	    show "(𝒪m  all_acquired sbm) 
              read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sbn) 𝒪n)
              (dropWhile (Not  is_volatile_Writesb) sbn) =
              {}"
	    proof (cases "m=i")
	      case True
	      with neq_n_m have neq_n_i: "ni"
		by auto
		
	      with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
		by (auto simp add: tssb')
	      note read_only_reads_unowned [OF n_bound' i_bound  neq_n_i nth' tssb_i]
	      moreover
	      note acq_eq
	      ultimately show ?thesis
		using True tssb_i nth mth n_bound' m_bound'
		by (simp add: tssb')
	    next
	      case False
	      note neq_m_i = this
	      with m_bound mth i_bound have mth': "tssb!m = (pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
		by (auto simp add: tssb')
	      show ?thesis
	      proof (cases "n=i")
		case True
		note read_only_reads_unowned [OF i_bound m_bound' neq_m_i [symmetric] tssb_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 tssb_i nth mth n_bound' m_bound' neq_m_i
		  apply (case_tac "outstanding_refs (is_volatile_Writesb) sb = {}")
		  apply (clarsimp simp add: outstanding_vol_write_take_drop_appends
		    acquired_append read_only_reads_append tssb' sb' 𝒪sb')+
		  done
	      next
		case False
		with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
		  by (auto simp add: tssb')
		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 tssb'"
	proof -
	  have "all_acquired (sb @ [Readsb volatile a t (msb a)])  all_acquired sb"
	    by (auto simp add: all_acquired_append)
	  from ownership_distinct_instructions_read_value_store_buffer_independent 
	  [OF i_bound tssb_i this]
	  show ?thesis by (simp add: tssb' sb' 𝒪sb')
	qed
      qed


      have valid_hist': "valid_history program_step tssb'"
      proof -
	from valid_history [OF i_bound tssb_i]
	have hcons: "history_consistent θsb (hd_prog psb sb) sb".
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i]
	have t_notin_reads: "t  read_tmps sb"
	  by (auto simp add: "issb")
	from load_tmps_write_tmps_distinct [OF i_bound tssb_i]
	have t_notin_writes: "t  (fst ` write_sops sb )"
	  by (auto simp add: "issb")

	from valid_write_sops [OF i_bound tssb_i]
	have valid_sops: "sop  write_sops sb. valid_sop sop"
	  by auto
	from load_tmps_fresh [OF i_bound tssb_i]
	have t_fresh: "t  dom θsb"
	  using "issb"
	  by simp

	from valid_implies_valid_prog_hd [OF i_bound tssb_i valid]
	have "history_consistent (θsb(tmsb a)) 
	       (hd_prog psb (sb@ [Readsb volatile a t (msb a)])) 
               (sb@ [Readsb volatile a t (msb a)])"
	  using t_notin_writes valid_sops t_fresh hcons
	  apply -
	  apply (rule history_consistent_appendI)
	  apply (auto simp add: hd_prog_append_Readsb)
	  done

	from valid_history_nth_update [OF i_bound this]
	show ?thesis
	  by (auto simp add: tssb' sb' 𝒪sb' θsb')
      qed

      from 
	reads_consistent_unbuffered_snoc [OF buf_None refl valid_reads [OF i_bound tssb_i] volatile_cond ]    
      have reads_consis': "reads_consistent False 𝒪sb msb (sb @ [Readsb volatile a t (msb a)])"
	by (simp split: if_split_asm)

      from valid_reads_nth_update [OF i_bound this]
      have valid_reads': "valid_reads msb tssb'" by (simp add: tssb' sb' 𝒪sb')

      have valid_sharing': "valid_sharing 𝒮sb' tssb'"
      proof (intro_locales)	
	from outstanding_non_volatile_writes_unshared [OF i_bound tssb_i]
	have "non_volatile_writes_unshared 𝒮sb (sb @ [Readsb volatile a t (msb 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 𝒮sb' tssb'"
	  by (simp add: tssb' sb' 𝒮sb')
      next
	from sharing_consis [OF i_bound tssb_i]
	have "sharing_consistent 𝒮sb 𝒪sb sb".
	then
	have "sharing_consistent 𝒮sb 𝒪sb (sb @ [Readsb volatile a t (msb a)])"
	  by (simp add:  sharing_consistent_append)
	from sharing_consis_nth_update [OF i_bound this]
	show "sharing_consis 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒪sb' sb' 𝒮sb')
      next
	note read_only_unowned [OF i_bound tssb_i]
	from read_only_unowned_nth_update [OF i_bound this]
	show "read_only_unowned 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' sb' 𝒪sb')
      next
	from unowned_shared_nth_update [OF i_bound tssb_i subset_refl]
	show "unowned_shared 𝒮sb' tssb'" by (simp add: tssb' 𝒪sb' 𝒮sb')
      next
	from no_outstanding_write_to_read_only_memory [OF i_bound tssb_i]
	have "no_write_to_read_only_memory 𝒮sb sb".
	hence "no_write_to_read_only_memory 𝒮sb (sb@[Readsb volatile a t (msb 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 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒮sb' sb')
      qed

      have tmps_distinct': "tmps_distinct tssb'"
      proof (intro_locales)
	from load_tmps_distinct [OF i_bound tssb_i]
	have "distinct_load_tmps issb'"
	  by (auto split: instr.splits simp add: issb)
	from load_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_distinct tssb'" by (simp add: tssb')
      next
	from read_tmps_distinct [OF i_bound tssb_i]
	have "distinct_read_tmps sb".
	moreover
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i]
	have "t  read_tmps sb"
	  by (auto simp add: issb)
	ultimately have "distinct_read_tmps (sb @ [Readsb volatile a t (msb a)])"
	  by (auto simp add: distinct_read_tmps_append)
	from read_tmps_distinct_nth_update [OF i_bound this]
	show "read_tmps_distinct tssb'" by (simp add: tssb' sb')
      next
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i] 
          load_tmps_distinct [OF i_bound tssb_i]
	have "load_tmps issb'  read_tmps (sb @ [Readsb volatile a t (msb a)]) = {}"
	  by (clarsimp simp add: read_tmps_append "issb")
	from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_read_tmps_distinct tssb'" by (simp add: tssb' sb')
      qed

      have valid_sops': "valid_sops tssb'"
      proof -
	from valid_store_sops [OF i_bound tssb_i]
	have valid_store_sops': "sopstore_sops issb'. valid_sop sop"
	  by (auto simp add: "issb")
	from valid_write_sops [OF i_bound tssb_i]
	have valid_write_sops': "sopwrite_sops (sb@ [Readsb volatile a t (msb 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: tssb' sb')
      qed

      have valid_dd': "valid_data_dependency tssb'"
      proof -
	from data_dependency_consistent_instrs [OF i_bound tssb_i]
	have dd_is: "data_dependency_consistent_instrs (dom θsb') issb'"
	  by (auto simp add: "issb" θsb')
	from load_tmps_write_tmps_distinct [OF i_bound tssb_i]
	have "load_tmps issb'  (fst ` write_sops (sb@ [Readsb volatile a t (msb a)])) = {}"
	  by (auto simp add: write_sops_append "issb")
	from valid_data_dependency_nth_update [OF i_bound dd_is this]
	show ?thesis by (simp add: tssb' sb')
      qed

      have load_tmps_fresh': "load_tmps_fresh tssb'"
      proof -
	from load_tmps_fresh [OF i_bound tssb_i] 
	have "load_tmps (Read volatile a t # issb')  dom θsb = {}"
	  by (simp add: "issb")
	moreover
	from load_tmps_distinct [OF i_bound tssb_i] have "t  load_tmps issb'"
	  by (auto simp add: "issb")
	ultimately have "load_tmps issb'  dom (θsb(t  (msb a))) = {}"
	  by auto
	from load_tmps_fresh_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' θsb')
      qed

      have enough_flushs': "enough_flushs tssb'"
      proof -
	from clean_no_outstanding_volatile_Writesb [OF i_bound tssb_i]
	have "¬ 𝒟sb  outstanding_refs is_volatile_Writesb (sb@[Readsb volatile a t (msb a)]) = {}"
	  by (auto simp add: outstanding_refs_append )
	from enough_flushs_nth_update [OF i_bound this]
	show ?thesis
	  by (simp add: tssb' sb' 𝒟sb')
      qed

      have valid_program_history': "valid_program_history tssb'"
      proof -	
	from valid_program_history [OF i_bound tssb_i]
	have "causal_program_history issb sb" .
	then have causal': "causal_program_history issb' (sb@[Readsb volatile a t (msb a)])"
	  by (auto simp: causal_program_history_Read  "issb")
	from valid_last_prog [OF i_bound tssb_i]
	have "last_prog psb sb = psb".
	hence "last_prog psb (sb @ [Readsb volatile a t (msb a)]) = psb"
	  by (simp add: last_prog_append_Readsb)
	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb' sb')
      qed

      show ?thesis
      proof (cases "outstanding_refs is_volatile_Writesb sb = {}")
	case True 

	from True have flush_all: "takeWhile (Not  is_volatile_Writesb) sb = sb"
	  by (auto simp add: outstanding_refs_conv )

	from True have suspend_nothing: "dropWhile (Not  is_volatile_Writesb) 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 # issb'"
	  by (simp add: "issb")
	with suspends_empty ts_i 
	have ts_i: "ts!i = (psb, Read volatile a t # issb', θsb,(), 
          𝒟, acquired True ?take_sb 𝒪sb, release ?take_sb (dom 𝒮sb) sb)"
	  by simp

	from direct_memop_step.Read
	have "(Read volatile a t # issb',θsb, (), m, 
            𝒟, acquired True ?take_sb 𝒪sb,release ?take_sb (dom 𝒮sb) sb,𝒮)  
          (issb', θsb(t  m a), (), m, 𝒟, acquired True ?take_sb 𝒪sb,
           release ?take_sb (dom 𝒮sb) sb, 𝒮)".
	from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
	have "(ts, m, 𝒮) d (ts[i := (psb, issb', θsb(t  m a), (), 
	   𝒟, acquired True ?take_sb 𝒪sb,release ?take_sb (dom 𝒮sb) sb)], m, 𝒮)".

	moreover
	
	from flush_all_until_volatile_write_Read_commute [OF i_bound tssb_i [simplified "issb"] ]
	have flush_commute: "flush_all_until_volatile_write
          (tssb[i := (psb,issb', θsb(tmsb a), sb @ [Readsb volatile a t (msb a)], 𝒟sb, 𝒪sb,sb)]) 
           msb =
          flush_all_until_volatile_write tssb msb".
	
	have "flush_all_until_volatile_write tssb msb a = msb a"
	proof -
          have "j < length tssb. i  j 
                  (let (_,_,_,sbj,_,_,_) = tssb!j 
                  in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))"
	  proof -
	    {
	      fix j pj "isj" 𝒪j j 𝒟j acqj xsj sbj
	      assume j_bound: "j < length tssb"
	      assume neq_i_j: "i  j"
	      assume jth: "tssb!j = (pj,isj, xsj, sbj, 𝒟j, 𝒪j, j)"
	      have "a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	      proof 
		let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
		let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"
		assume a_in: "a  outstanding_refs is_non_volatile_Writesb ?take_sbj"
		with outstanding_refs_takeWhile [where P'= "Not  is_volatile_Writesb"]
		have a_in': "a  outstanding_refs is_non_volatile_Writesb sbj"
		  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 sbj"
		  by auto
		with ownership_distinct [OF i_bound j_bound neq_i_j tssb_i jth]
		have a_not_owns: "a  𝒪sb  all_acquired sb"
		  by blast
		
		from non_volatile_owned_or_read_only_append [of False 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
		  outstanding_non_volatile_refs_owned_or_read_only [OF j_bound jth]
		have "non_volatile_owned_or_read_only False 𝒮sb 𝒪j ?take_sbj"
		  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_sbj"
		  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_sbj (dom (𝒮sb)) j  a  Some False"
                  by (auto simp add: Let_def)
		{
		  assume "a  acquired True sb 𝒪sb"
		  with acquired_all_acquired_in [OF this] ownership_distinct [OF i_bound j_bound neq_i_j tssb_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 𝒪sb]
                  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_tssb sharing_consis_tssb j_bound jth j_owns]
                  have "a  read_only (share ?take_sbj 𝒮sb)"
                    by (auto simp add: read_only_def 𝒮)
                  hence a_dom: "a  dom  (share ?take_sbj 𝒮sb)"
                    by (auto simp add: read_only_def domIff)
                  from outstanding_non_volatile_writes_unshared [OF j_bound jth]
                  non_volatile_writes_unshared_append [of 𝒮sb ?take_sbj ?drop_sbj]
                  have nvw: "non_volatile_writes_unshared 𝒮sb ?take_sbj" 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 𝒮sb sbj".
		  with non_volatile_writes_unshared_append [of 𝒮sb "(takeWhile (Not  is_volatile_Writesb) sbj)"
		  "(dropWhile (Not  is_volatile_Writesb) sbj)"]
		  have unshared_take: "non_volatile_writes_unshared 𝒮sb (takeWhile (Not  is_volatile_Writesb) sbj)" 
		    by clarsimp
		
		  from valid_own have own_dist: "ownership_distinct tssb"
		    by (simp add: valid_ownership_def)
		  from valid_sharing have "sharing_consis 𝒮sb tssb"
		    by (simp add: valid_sharing_def)
		  from sharing_consistent_share_all_until_volatile_write [OF own_dist this i_bound tssb_i]
		  have sc: "sharing_consistent 𝒮 (acquired True ?take_sb 𝒪sb) ?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 𝒪sb" by auto
		  also from acquired_all_acquired all_acquired_takeWhile 
		  have "  dom 𝒮  (𝒪sb  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_tssb sharing_consis_tssb j_bound jth j_owns]
                  have a_dom: "a  dom  (share ?take_sbj 𝒮sb)"
                    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 tssb_i this]
	  show ?thesis
	    by (simp add: buf_None)
	qed
	
	hence m_a: "m a = msb a"
	  by (simp add: m)
	
	have tmps_commute: "θsb(t  (msb a)) = 
	  (θsb |` (dom θsb - {t}))(t  (msb a))"
	  apply (rule ext)
	  apply (auto simp add: restrict_map_def domIff)
	  done

	from suspend_nothing
	have suspend_nothing': "(dropWhile (Not  is_volatile_Writesb) sb') = []"
	  by (simp add: sb')

	from 𝒟
	have 𝒟': "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb (sb@[Readsb volatile a t (msb a)])   {})"
	  by (auto simp: outstanding_refs_append)

	have "(tssb',msb,𝒮sb')  (ts[i := (psb,issb', θsb(tm a),(), 𝒟, acquired True ?take_sb 𝒪sb,release ?take_sb (dom 𝒮sb) sb)], m,𝒮)"
	  apply (rule sim_config.intros) 
	  apply    (simp add: m flush_commute tssb' 𝒪sb' sb' θsb' sb' 𝒟sb' )
	  using   share_all_until_volatile_write_Read_commute [OF i_bound tssb_i [simplified issb]]
	  apply   (simp add: 𝒮 𝒮sb' tssb' sb' 𝒪sb' sb' θsb')
	  using  leq
	  apply  (simp add: tssb')
	  using i_bound i_bound' ts_sim ts_i True 𝒟'
	  apply (clarsimp simp add: Let_def nth_list_update 
	    outstanding_refs_conv m_a  tssb' 𝒪sb' sb' 𝒮sb' θsb' sb' 𝒟sb' 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'
	    msb' 𝒮sb' 
	  by (auto simp del: fun_upd_apply )
      next
	case False

	then obtain r where r_in: "r  set sb" and volatile_r: "is_volatile_Writesb r"
	  by (auto simp add: outstanding_refs_conv)
	from takeWhile_dropWhile_real_prefix 
	[OF r_in, of  "(Not  is_volatile_Writesb)", simplified, OF volatile_r] 
	obtain a' v' sb'' sop' A' L' R' W' where
	  sb_split: "sb = takeWhile (Not  is_volatile_Writesb) sb @ Writesb True a' sop' v' A' L' R' W'# sb''" 
	  and
	  drop: "dropWhile (Not  is_volatile_Writesb) sb = Writesb 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 = Writesb 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 tssb_i 
	  [simplified "issb"] ]

	have "Writesb True a' sop' v' A' L' R' W' set sb"
	  by (subst sb_split) auto
	
	from dropWhile_append1 [OF this, of "(Not  is_volatile_Writesb)"]
	have drop_app_comm:
	  "(dropWhile (Not  is_volatile_Writesb) (sb @ [Readsb volatile a t (msb a)])) =
                dropWhile (Not  is_volatile_Writesb) sb @ [Readsb volatile a t (msb a)]"
	  by simp

	from load_tmps_fresh [OF i_bound tssb_i]
	have "t  dom θsb"
	  by (auto simp add: "issb")
	then have tmps_commute: 
	  "θsb |` (dom θsb - read_tmps sb'') =
          θsb |` (dom θsb - insert t (read_tmps sb''))"
	  apply -
	  apply (rule ext)
	  apply auto
	  done

	from 𝒟
	have 𝒟': "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb (sb@[Readsb volatile a t (msb a)])   {})"
	  by (auto simp: outstanding_refs_append)

	have "(tssb',msb,𝒮sb)  (ts,m,𝒮)"
	  apply (rule sim_config.intros) 
	  apply    (simp add: m flush_commute tssb' 𝒪sb' sb' θsb' sb')
	  using   share_all_until_volatile_write_Read_commute [OF i_bound tssb_i [simplified issb]]
	  apply   (simp add: 𝒮 𝒮sb' tssb' sb' 𝒪sb' sb' θsb')
	  using  leq
	  apply  (simp add: tssb')
	  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_Readsb instrs_append_Readsb 
	    hd_prog_append_Readsb
	    drop "issb" tssb' sb' 𝒪sb' sb' θsb' 𝒟sb' 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'
	    msb' 𝒮sb' 
	  by (auto simp del: fun_upd_apply )
      qed
    next
      case (SBHWriteNonVolatile a D f A L R W)
      then obtain 
	"issb": "issb = Write False a (D, f) A L R W# issb'" and
	𝒪sb': "𝒪sb'=𝒪sb" and
        sb': "sb'=sb" and
	θsb': "θsb' = θsb" and
	𝒟sb': "𝒟sb'=𝒟sb" and
	sb': "sb'=sb@[Writesb False a (D, f) (f θsb) A L R W]" and
	msb': "msb' = msb" and
	𝒮sb': "𝒮sb'=𝒮sb" 
	by auto


      from data_dependency_consistent_instrs [OF i_bound tssb_i]
      have D_tmps: "D  dom θsb" 
	by (simp add: issb)

      from safe_memop_flush_sb [simplified issb]
      obtain a_owned': "a  acquired True sb 𝒪sb" and a_unshared': "a  dom (share ?drop_sb 𝒮)" and
        rels_cond: "j < length ts. ij  a  dom (released (ts!j))"
      (* FIXME: rels_cond unused; maybe remove from safe_delayed *) 
	by cases auto

      from a_owned' acquired_all_acquired
      have a_owned'': "a  𝒪sb  all_acquired sb"
	by auto


      {
	fix j
	fix pj isj 𝒪j j 𝒟j θj sbj
	assume j_bound: "j < length tssb"
	assume tssb_j: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	assume neq_i_j: "i  j"
	have "a  𝒪j  all_acquired sbj"
	proof -
	  from ownership_distinct [OF i_bound j_bound neq_i_j tssb_i tssb_j] a_owned''
	  show ?thesis
	    by auto
	qed
      } note a_unowned_others = this
	  
	    
      have a_unshared: "a  dom (share sb 𝒮sb)"
      proof 
	assume a_share: "a  dom (share sb 𝒮sb)"
	from valid_sharing have "sharing_consis 𝒮sb tssb"
	  by (simp add: valid_sharing_def)
	from in_shared_sb_share_all_until_volatile_write [OF this i_bound tssb_i a_owned'' a_share]
	have "a  dom (share ?drop_sb 𝒮)"
	  by (simp add: 𝒮)
	with a_unshared'
	show False
	  by auto
      qed

(*
      from acquired_owns_shared [OF sharing_consis_drop_sb]
      have "acquired True ?drop_sb 𝒪 ⊆ 𝒪 ∪ 𝒮".
      moreover
      from share_owns_shared [OF sharing_consis_drop_sb]
      have "share ?drop_sb 𝒮 ⊆ 𝒪 ∪ 𝒮".
*)
(*
      obtain a_owned: "a ∈ 𝒪" and a_unshared: "a ∉ 𝒮" 
	by cases auto
*)
      have valid_own': "valid_ownership 𝒮sb' tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	proof -
	  from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i]  
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb sb".
	  
	  with a_owned' 
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb (sb @ [Writesb False a (D,f) (f θsb) 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: tssb' "issb" sb' 𝒪sb' 𝒮sb')
	qed
      next
	show "outstanding_volatile_writes_unowned_by_others tssb'"
	proof -
	  have "outstanding_refs is_volatile_Writesb (sb @ [Writesb False a (D,f) (f θsb) A L R W])  
	    outstanding_refs is_volatile_Writesb sb"
	    by (auto simp add: outstanding_refs_append)
	  from outstanding_volatile_writes_unowned_by_others_store_buffer 
	  [OF i_bound tssb_i this]
	  show ?thesis by (simp add: tssb' "issb" sb' 𝒪sb' all_acquired_append)
	qed
      next
	show "read_only_reads_unowned tssb'"
	proof -
	  have r: "read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) 
	    (sb @ [Writesb False a (D,f) (f θsb) A L R W])) 𝒪sb)
            (dropWhile (Not  is_volatile_Writesb) (sb @ [Writesb False a (D,f) (f θsb) A L R W]))
             
	    read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪sb)
            (dropWhile (Not  is_volatile_Writesb) sb)"
	    apply (case_tac "outstanding_refs (is_volatile_Writesb) sb = {}")
	    apply (simp_all add: outstanding_vol_write_take_drop_appends
	    acquired_append read_only_reads_append )
	    done
	  have "𝒪sb  all_acquired (sb @ [Writesb False a (D,f) (f θsb) A L R W])  𝒪sb  all_acquired sb"
	    by (simp add: all_acquired_append)
	  

	  from read_only_reads_unowned_nth_update [OF i_bound tssb_i r this]
	  show ?thesis
	    by (simp add: tssb' 𝒪sb' sb')
	qed 
      next
	show "ownership_distinct tssb'"
	proof -
	  from ownership_distinct_instructions_read_value_store_buffer_independent 
	  [OF i_bound tssb_i]
	  show ?thesis by (simp add: tssb' "issb" sb' 𝒪sb' all_acquired_append)
	qed
      qed

      have valid_hist': "valid_history program_step tssb'"
      proof -
	from valid_history [OF i_bound tssb_i]
	have "history_consistent θsb (hd_prog psb sb) sb".
	with valid_write_sops [OF i_bound tssb_i] D_tmps 
	  valid_implies_valid_prog_hd [OF i_bound tssb_i valid]
	have "history_consistent θsb (hd_prog psb (sb@[Writesb False a (D,f) (f θsb) A L R W])) 
	       (sb@ [Writesb False a (D,f) (f θsb) A L R W])"
	  apply -
	  apply (rule history_consistent_appendI)
	  apply (auto simp add: hd_prog_append_Writesb)
	  done
	from valid_history_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' "issb" sb' 𝒪sb' θsb')
      qed

      have valid_reads': "valid_reads msb tssb'"
      proof -
	from valid_reads [OF i_bound tssb_i]
	have "reads_consistent False 𝒪sb msb sb" .
	from reads_consistent_snoc_Writesb [OF this]
	have "reads_consistent False 𝒪sb msb (sb @ [Writesb False a (D,f) (f θsb) A L R W])".
	from valid_reads_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' "issb" sb' 𝒪sb' θsb')
      qed

      have valid_sharing': "valid_sharing 𝒮sb' tssb'"
      proof (intro_locales)	
	from outstanding_non_volatile_writes_unshared [OF i_bound tssb_i] a_unshared
	have "non_volatile_writes_unshared 𝒮sb
	      (sb @ [Writesb False a (D,f) (f θsb) 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 𝒮sb' tssb'"
	  by (simp add: tssb' "issb" sb' 𝒪sb' θsb' 𝒮sb')
      next
	from sharing_consis [OF i_bound tssb_i]
	have "sharing_consistent 𝒮sb 𝒪sb sb".
	then
	have "sharing_consistent 𝒮sb 𝒪sb (sb @ [Writesb False a (D,f) (f θsb) A L R W])"
	  by (simp add:  sharing_consistent_append)
	from sharing_consis_nth_update [OF i_bound this]
	show "sharing_consis 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒪sb' sb' 𝒮sb')
      next
	from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound tssb_i] ]
	show "read_only_unowned 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' 𝒪sb')
      next
	from unowned_shared_nth_update [OF i_bound tssb_i subset_refl]
	show "unowned_shared 𝒮sb' tssb'"
	  by (simp add: tssb' "issb" sb' 𝒪sb' θsb' 𝒮sb')
      next
	from a_unshared
	have "a  read_only (share sb 𝒮sb)"
	  by (auto simp add: read_only_def dom_def)
	with no_outstanding_write_to_read_only_memory [OF i_bound tssb_i] 

	have "no_write_to_read_only_memory 𝒮sb (sb @ [Writesb False a (D,f) (f θsb) 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 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' sb')
      qed

      have tmps_distinct': "tmps_distinct tssb'"
      proof (intro_locales)
	from load_tmps_distinct [OF i_bound tssb_i]
	have "distinct_load_tmps issb'" 
	  by (auto split: instr.splits simp add: "issb")
	from load_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_distinct tssb'"	  
	  by (simp add: tssb' "issb" sb' 𝒪sb' θsb')      
      next
	from read_tmps_distinct [OF i_bound tssb_i]
	have "distinct_read_tmps sb".
	hence "distinct_read_tmps (sb @ [Writesb False a (D,f) (f θsb) 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 tssb'"
	  by (simp add: tssb' "issb" sb' 𝒪sb' θsb')      
      next
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i] 
          load_tmps_distinct [OF i_bound tssb_i]
	have "load_tmps issb'  read_tmps (sb @ [Writesb False a (D,f) (f θsb) A L R W]) = {}"
	  by (clarsimp simp add: read_tmps_append "issb")
	from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_read_tmps_distinct tssb'" 
	  by (simp add: tssb' "issb" sb' 𝒪sb' θsb')      
      qed

      have valid_sops': "valid_sops tssb'"
      proof -
	from valid_store_sops [OF i_bound tssb_i]
	obtain valid_Df: "valid_sop (D,f)" and  
	  valid_store_sops': "sopstore_sops issb'. valid_sop sop"
	  by (auto simp add: "issb")
	from valid_Df valid_write_sops [OF i_bound tssb_i]
	have valid_write_sops': "sopwrite_sops (sb@ [Writesb False a (D, f) (f θsb) 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: tssb' "issb" sb' 𝒪sb' θsb')      
      qed

      have valid_dd': "valid_data_dependency tssb'"
      proof -
	from data_dependency_consistent_instrs [OF i_bound tssb_i]
	obtain D_indep: "D  load_tmps issb' = {}" and 
	  dd_is: "data_dependency_consistent_instrs (dom θsb') issb'"
	  by (auto simp add: "issb" θsb')
	from load_tmps_write_tmps_distinct [OF i_bound tssb_i] D_indep
	have "load_tmps issb'  
	      (fst ` write_sops (sb@ [Writesb False a (D, f) (f θsb) A L R W])) = {}"
	  by (auto simp add: write_sops_append "issb")
	from valid_data_dependency_nth_update [OF i_bound dd_is this]
	show ?thesis 	  
	  by (simp add: tssb' "issb" sb' 𝒪sb' θsb')      
      qed

      have load_tmps_fresh': "load_tmps_fresh tssb'"
      proof -
	from load_tmps_fresh [OF i_bound tssb_i] 
	have "load_tmps issb'  dom θsb = {}"
	  by (auto simp add: "issb")
	from load_tmps_fresh_nth_update [OF i_bound this]
	show ?thesis 	  
	  by (simp add: tssb' "issb" sb' 𝒪sb' θsb')      
      qed

      have enough_flushs': "enough_flushs tssb'"
      proof -
	from clean_no_outstanding_volatile_Writesb [OF i_bound tssb_i]
	have "¬ 𝒟sb  outstanding_refs is_volatile_Writesb (sb@[Writesb False a (D,f) (f θsb) 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: tssb' sb' 𝒟sb')
      qed


      have valid_program_history': "valid_program_history tssb'"
      proof -	
	from valid_program_history [OF i_bound tssb_i]
	have "causal_program_history issb sb" .
	then have causal': "causal_program_history issb' (sb@[Writesb False a (D,f) (f θsb) A L R W])"
	  by (auto simp: causal_program_history_Write  "issb")
	from valid_last_prog [OF i_bound tssb_i]
	have "last_prog psb sb = psb".
	hence "last_prog psb (sb @ [Writesb False a (D,f) (f θsb) A L R W]) = psb"
	  by (simp add: last_prog_append_Writesb)
	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb' sb')
      qed
      

      from valid_store_sops [OF i_bound tssb_i, rule_format]
      have "valid_sop (D,f)" by (auto simp add: "issb")
      then interpret valid_sop "(D,f)" .

      show ?thesis
      proof (cases "outstanding_refs is_volatile_Writesb sb = {}")
	case True
      
	from True have flush_all: "takeWhile (Not  is_volatile_Writesb) sb = sb"
	  by (auto simp add: outstanding_refs_conv)
      
	from True have suspend_nothing: "dropWhile (Not  is_volatile_Writesb) 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# issb'"
	  by (simp add: "issb")
	with suspends_empty ts_i 
	have ts_i: "ts!i = (psb, Write False a (D,f) A L R W# issb',
                     θsb,(),
                     𝒟, acquired True ?take_sb 𝒪sb,release ?take_sb (dom (𝒮sb)) sb)"
	  by simp

	from direct_memop_step.WriteNonVolatile [OF ]
	have "(Write False a (D, f) A L R W# issb', 
	  θsb, (),m,𝒟,acquired True ?take_sb 𝒪sb ,release ?take_sb (dom (𝒮sb)) sb, 𝒮)  
               (issb',
                  θsb, (), m(a := f θsb), 𝒟, acquired True ?take_sb 𝒪sb,
                  release ?take_sb (dom (𝒮sb)) sb, 𝒮)".
	from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
	have "(ts, m, 𝒮) d 
              (ts[i := (psb, issb', θsb, (),𝒟, acquired True ?take_sb 𝒪sb,
                  release ?take_sb (dom (𝒮sb)) sb)], 
	       m(a := f θsb),𝒮)".

	moreover


	have "j<length tssb. i  j 
          (let (_,_, _, sbj,_,_,_) = tssb ! j
          in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))"
	proof -
	  {
	    fix j pj "isj" 𝒪j j 𝒟j acqj xsj sbj
	    assume j_bound: "j < length tssb"
	    assume neq_i_j: "i  j"
	    assume jth: "tssb!j = (pj,isj, xsj, sbj, 𝒟j, 𝒪j,j)"
	    have "a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	    proof 
	      assume a_in: "a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	      hence "a  outstanding_refs is_non_volatile_Writesb sbj"
		using outstanding_refs_append [of is_non_volatile_Writesb "(takeWhile (Not  is_volatile_Writesb) sbj)"
		"(dropWhile (Not  is_volatile_Writesb) sbj)"]
		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 sbj"
		by auto

	      from j_owns a_owned'' ownership_distinct [OF i_bound j_bound neq_i_j tssb_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 tssb_i this]

	from suspend_nothing
	have suspend_nothing': "(dropWhile (Not  is_volatile_Writesb) sb') = []"
	  by (simp add: sb')

	from 𝒟
	have 𝒟': "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb 
	  (sb@[Writesb False a (D,f) (f θsb) A L R W])   {})"
	  by (auto simp: outstanding_refs_append)

	have "(tssb',msb,𝒮sb')  
	   (ts[i := (psb,issb', θsb,(),𝒟, acquired True ?take_sb 𝒪sb,
                     release ?take_sb (dom (𝒮sb)) sb)], 
                m(a:=f θsb),𝒮)"
	  apply (rule sim_config.intros) 
	  apply    (simp add: m flush_commute tssb' 𝒪sb' sb' sb' θsb' 𝒟sb' )
	  using  share_all_until_volatile_write_Write_commute 
	          [OF i_bound tssb_i [simplified issb]]
	  apply   (clarsimp simp add: 𝒮 𝒮sb' tssb' sb' 𝒪sb' sb' θsb')
	  using  leq
	  apply  (simp add: tssb')
	  using i_bound i_bound' ts_sim ts_i True 𝒟'
	  apply (clarsimp simp add: Let_def nth_list_update 
	    outstanding_refs_conv tssb' 𝒪sb' sb' 𝒮sb' θsb' sb' 𝒟sb' 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' msb' 𝒮sb' 
	  by (auto simp del: fun_upd_apply)
      next

	case False

	then obtain r where r_in: "r  set sb" and volatile_r: "is_volatile_Writesb r"
	  by (auto simp add: outstanding_refs_conv)
	from takeWhile_dropWhile_real_prefix 
	[OF r_in, of  "(Not  is_volatile_Writesb)", simplified, OF volatile_r] 
	obtain a' v' sb'' sop' A' L' R' W' where
	  sb_split: "sb = takeWhile (Not  is_volatile_Writesb) sb @ Writesb True a' sop' v' A' L' R' W'# sb''" 
	  and
	  drop: "dropWhile (Not  is_volatile_Writesb) sb = Writesb 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 = Writesb 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 tssb_i]

	have "Writesb 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_Writesb)", simplified]

	from 𝒟
	have 𝒟': "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb (sb@[Writesb False a (D,f) (f θsb) A L R W])   {})"
	  by (auto simp: outstanding_refs_append)


	have "(tssb',msb,𝒮sb')  (ts,m,𝒮)"
	  apply (rule sim_config.intros) 
	  apply    (simp add: m flush_commute tssb' 𝒪sb' sb' θsb' sb')
	  using   share_all_until_volatile_write_Write_commute 
	           [OF i_bound tssb_i [simplified issb]]
	  apply   (clarsimp simp add: 𝒮 𝒮sb' tssb' sb' 𝒪sb' sb' θsb')
	  using  leq
	  apply  (simp add: tssb')
	  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_Writesb instrs_append_Writesb hd_prog_append_Writesb
	    drop "issb" tssb' sb' 𝒪sb' sb' 𝒮sb' 
            θsb' 𝒟sb' 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' msb' 𝒮sb' 
	  by (auto simp del: fun_upd_apply )
      qed
    next	
      case (SBHWriteVolatile a D f A L R W)
      then obtain 
	"issb": "issb = Write True a (D, f) A L R W# issb'" and
	𝒪sb': "𝒪sb'=𝒪sb" and
        sb': "sb'=sb" and
	θsb': "θsb' = θsb" and
	𝒟sb': "𝒟sb'=True" and
	sb': "sb'=sb@[Writesb True a (D, f) (f θsb) A L R W]" and
	msb': "msb' = msb" and
	𝒮sb': "𝒮sb'=𝒮sb" 
	by auto

      from data_dependency_consistent_instrs [OF i_bound tssb_i]
      have D_subset: "D  dom θsb" 
	by (simp add: issb)

      from safe_memop_flush_sb [simplified issb] 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 𝒪sb" and
	R_acq: "R  acquired True sb 𝒪sb" and
	A_R: "A  R = {}" and
        A_unowned_by_others_ts:  
	"j<length (map owned ts). ij  (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 tssb. i  j  
          (let (_,_,_,sbj,_,𝒪j,_) = tssb!j in 
	    a  acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j 
            a  all_shared (takeWhile (Not  is_volatile_Writesb) sbj))" 
  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_others_ts ts_sim leq
      have a_unowned_others:
        "∀j<length (map owns_sb tssb). i ≠ j ⟶ 
          (let (𝒪j,sbj) = (map owns_sb tssb)!j in 
	    a ∉ acquired True (takeWhile (Not ∘ is_volatile_Writesb) sbj) 𝒪j ∧
            a ∉ all_shared (takeWhile (Not ∘ is_volatile_Writesb) sbj))" 
	apply (clarsimp simp add: Let_def)
	apply (drule_tac x=j in spec)
	apply (auto simp add: dom_release_takeWhile)
	done
*)
(*
      from a_unowned_others
      have a_unacquired_others:
        "∀j<length tssb. i ≠ j ⟶ 
          (let (_,_,_,sbj,_,_) = tssb!j in 
	    a ∉ all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sbj))" 
	by (auto simp add: acquired_takeWhile_non_volatile_Writesb)
*)
      have a_not_ro: "a  read_only (share sb 𝒮sb)"
      proof 
	assume a: "a  read_only (share sb 𝒮sb)"
	from local.read_only_unowned_axioms have "read_only_unowned 𝒮sb tssb".
        from in_read_only_share_all_until_volatile_write' [OF ownership_distinct_tssb sharing_consis_tssb
          read_only_unowned 𝒮sb tssb i_bound tssb_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 tssb. ij  (let (_,_,_,sbj,_,𝒪j,_) = tssb!j 
	  in A  (acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j 
                  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)) = {})" 
  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 tssb). i  j  
        (let (𝒪j,sbj) = (map 𝒪_sb tssb)!j in a  all_acquired sbj)" 
      proof -
	{
	  fix j 𝒪j sbj
	  assume j_bound: "j < length (map owned tssb)"
	  assume neq_i_j: "ij"
	  assume tssb_j: "(map 𝒪_sb tssb)!j = (𝒪j,sbj)"
	  assume conflict: "a  all_acquired sbj"
	  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 tssb"
	      by auto
	    from j_bound' have j_bound''': "j < length ts"
	      by simp

	    let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	    let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"

	    from ts_sim [rule_format, OF  j_bound''] tssb_j j_bound''
            
	    obtain pj suspendsj "issbj" j 𝒟sbj 𝒟j θsbj "isj" where
		  tssb_j: "tssb ! j = (pj,issbj, θsbj, sbj, 𝒟sbj,𝒪j,j)"  and
		  suspendsj: "suspendsj = dropWhile (Not  is_volatile_Writesb) sbj" and
		  isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	          𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
		  tsj: "ts!j = (hd_prog pj suspendsj, isj,
                               θsbj |` (dom θsbj - read_tmps suspendsj),(), 
	                       𝒟j, 
                               acquired True ?take_sbj 𝒪j,
                               release ?take_sbj (dom 𝒮sb) j)"
	      apply (cases "tssb!j")
	      apply (force simp add: Let_def)
	      done

            
	    from a_unowned_others [rule_format,OF _ neq_i_j] tssb_j j_bound
	    obtain a_unacq: "a  acquired True ?take_sbj 𝒪j" and a_not_shared: "a  all_shared ?take_sbj"
	      by auto
            have conflict_drop: "a  all_acquired suspendsj"
            proof (rule ccontr)
              assume "a  all_acquired suspendsj"
              with all_acquired_append [of ?take_sbj ?drop_sbj] conflict
              have "a  all_acquired ?take_sbj"
                by (auto simp add: suspendsj)
              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

	    (* FIXME: extract common intermediate steps of both cases*)
	    from split_all_acquired_in [OF conflict_drop]
	    show ?thesis
	    proof
	      assume "sop a' v ys zs A L R W. 
                suspendsj = ys @ Writesb True a' sop v A L R W# zs  a  A"
	      then 
	      obtain a' sop' v' ys zs A' L' R' W' where
		split_suspendsj: "suspendsj = ys @ Writesb True a' sop' v' A' L' R' W'# zs" 
		(is "suspendsj = ?suspends") and
		a_A': "a  A'"
		by blast

	      from sharing_consis [OF j_bound'' tssb_j]
	      have sharing_consis_j: "sharing_consistent 𝒮sb 𝒪j sbj".
	      then have A'_R': "A'  R' = {}" 
		by (simp add: sharing_consistent_append [of _ _ ?take_sbj ?drop_sbj, simplified] 
		  suspendsj [symmetric] split_suspendsj sharing_consistent_append)
	      from valid_program_history [OF j_bound'' tssb_j] 
	      have "causal_program_history issbj sbj".
	      then have cph: "causal_program_history issbj ?suspends"
		apply -
		apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply (simp add: split_suspendsj)
		done

	      from tsj neq_i_j j_bound 
	      have ts'_j: "?ts'!j = (hd_prog pj suspendsj, isj,
		θsbj |` (dom θsbj - read_tmps suspendsj),(), 
		𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
		by auto
	      from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
	      then
	      have lp: "last_prog pj suspendsj = pj"
		apply -
		apply (rule last_prog_same_append [where sb="?take_sbj"])
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done


	      from valid_reads [OF j_bound'' tssb_j]
	      have reads_consis_j: "reads_consistent False 𝒪j msb sbj".

	      from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb 
		j_bound'' tssb_j this]
	      have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		by (simp add: m suspendsj)

	      from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j tssb_i tssb_j]
	      have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
		by (simp add: suspendsj)
	      from reads_consistent_flush_independent [OF this reads_consis_m_j]
	      have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
		(flush ?drop_sb m) suspendsj".

	      hence reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  
		(flush ?drop_sb m) (ys@[Writesb True a' sop' v' A' L' R' W'])"
		by (simp add: split_suspendsj reads_consistent_append)

	      from valid_write_sops [OF j_bound'' tssb_j]
	      have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		valid_sops_drop: "sopwrite_sops (ys@[Writesb 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'' tssb_j]
	      have "distinct_read_tmps (?take_sbj@suspendsj)"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain 
		read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
		distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		apply (simp only: distinct_read_tmps_append)
		done

	      from valid_history [OF j_bound'' tssb_j]
	      have h_consis: 
		"history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done
	    
	      have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	      proof -
		from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		  by simp
		from last_prog_hd_prog_append' [OF h_consis] this
		have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		  by (simp only: split_suspendsj [symmetric] suspendsj) 
		moreover 
		have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		  last_prog (hd_prog pj suspendsj) ?take_sbj"
		  apply (simp only: split_suspendsj [symmetric] suspendsj) 
		  by (rule last_prog_hd_prog_append)
		ultimately show ?thesis
		  by (simp add: split_suspendsj [symmetric] suspendsj) 
	      qed

	      from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop 
		h_consis] last_prog_hd_prog
	      have hist_consis': "history_consistent θsbj (hd_prog pj suspendsj) suspendsj"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      from reads_consistent_drop_volatile_writes_no_volatile_reads  
	      [OF reads_consis_j] 
	      have no_vol_read: "outstanding_refs is_volatile_Readsb 
		(ys@[Writesb True a' sop' v' A' L' R' W']) = {}"
		by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		  split_suspendsj )

	      have acq_simp:
		"acquired True (ys @ [Writesb True a' sop' v' A' L' R' W']) 
                    (acquired True ?take_sbj 𝒪j) = 
                 acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R'"
		by (simp add: acquired_append)

	      from flush_store_buffer_append [where sb="ys@[Writesb True a' sop' v' A' L' R' W']" and sb'="zs", simplified,
	      OF j_bound_ts' isj [simplified split_suspendsj] cph [simplified suspendsj]
	      ts'_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys 
	      hist_consis' [simplified split_suspendsj] valid_sops_drop 
	      distinct_read_tmps_drop [simplified split_suspendsj] 
		no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="share ?drop_sb 𝒮"]
	      obtain isj' j' where
		isj': "instrs zs @ issbj = isj' @ prog_instrs zs" and
		steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮)  d* 
		  (?ts'[j:=(last_prog
                              (hd_prog pj (Writesb True a' sop' v' A' L' R' W'# zs)) (ys@[Writesb True a' sop' v' A' L' R' W']),
                             isj',
                             θsbj |` (dom θsbj - read_tmps zs),
                              (), True, acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R',j')],
                    flush (ys@[Writesb True a' sop' v' A' L' R' W']) (flush ?drop_sb m),
                    share (ys@[Writesb 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 = (psb, issb, θsb,(), 
		𝒟sb, acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb)"
		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 issb] 
	      have a_unowned: 
		"j < length ?ts_ys. ij  (let (𝒪j) = map owned ?ts_ys!j in a  𝒪j)"
		apply cases
		apply (auto simp add: Let_def issb)
		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. suspendsj = ys @ Ghostsb A L R W# zs  a  A"
	      then 
	      obtain A' L' R' W' ys zs where
		split_suspendsj: "suspendsj = ys @ Ghostsb A' L' R' W'# zs" 
		(is "suspendsj = ?suspends") and
		  a_A': "a  A'"
		by blast

	      from sharing_consis [OF j_bound'' tssb_j]
	      have sharing_consis_j: "sharing_consistent 𝒮sb 𝒪j sbj".
	      then have A'_R': "A'  R' = {}" 
		by (simp add: sharing_consistent_append [of _ _ ?take_sbj ?drop_sbj, simplified] 
		  suspendsj [symmetric] split_suspendsj sharing_consistent_append)
	      from valid_program_history [OF j_bound'' tssb_j] 
	      have "causal_program_history issbj sbj".
	      then have cph: "causal_program_history issbj ?suspends"
		apply -
		apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply (simp add: split_suspendsj)
		done

	      from tsj neq_i_j j_bound 
	      have ts'_j: "?ts'!j = (hd_prog pj suspendsj, isj,
		θsbj |` (dom θsbj - read_tmps suspendsj),(), 
		𝒟j, acquired True ?take_sbj 𝒪j, release ?take_sbj (dom 𝒮sb) j)"
		by auto
	      from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
	      then
	      have lp: "last_prog pj suspendsj = pj"
		apply -
		apply (rule last_prog_same_append [where sb="?take_sbj"])
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done


	      from valid_reads [OF j_bound'' tssb_j]
	      have reads_consis_j: "reads_consistent False 𝒪j msb sbj".
	      from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb 
		j_bound'' tssb_j this]
	      have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		by (simp add: m suspendsj)

	      from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j tssb_i tssb_j]
	      have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
		by (simp add: suspendsj)
	      from reads_consistent_flush_independent [OF this reads_consis_m_j]
	      have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
		(flush ?drop_sb m) suspendsj".

	      hence reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  
		(flush ?drop_sb m) (ys@[Ghostsb A' L' R' W'])"
		by (simp add: split_suspendsj reads_consistent_append)

	      from valid_write_sops [OF j_bound'' tssb_j]
	      have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		valid_sops_drop: "sopwrite_sops (ys@[Ghostsb A' L' R' W']). valid_sop sop"
		apply (simp only: write_sops_append)
		apply auto
		done

	      from read_tmps_distinct [OF j_bound'' tssb_j]
	      have "distinct_read_tmps (?take_sbj@suspendsj)"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain 
		read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
		distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		apply (simp only: distinct_read_tmps_append)
		done

	      from valid_history [OF j_bound'' tssb_j]
	      have h_consis: 
		"history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done
	    
	      have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	      proof -
		from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		  by simp
		from last_prog_hd_prog_append' [OF h_consis] this
		have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		  by (simp only: split_suspendsj [symmetric] suspendsj) 
		moreover 
		have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		  last_prog (hd_prog pj suspendsj) ?take_sbj"
		  apply (simp only: split_suspendsj [symmetric] suspendsj) 
		  by (rule last_prog_hd_prog_append)
		ultimately show ?thesis
		  by (simp add: split_suspendsj [symmetric] suspendsj) 
	      qed

	      from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop 
		h_consis] last_prog_hd_prog
	      have hist_consis': "history_consistent θsbj (hd_prog pj suspendsj) suspendsj"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      from reads_consistent_drop_volatile_writes_no_volatile_reads  
	      [OF reads_consis_j] 
	      have no_vol_read: "outstanding_refs is_volatile_Readsb 
		(ys@[Ghostsb A' L' R' W']) = {}"
		by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		  split_suspendsj )

	      have acq_simp:
		"acquired True (ys @ [Ghostsb A' L' R' W']) 
                    (acquired True ?take_sbj 𝒪j) = 
                 acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R'"
		by (simp add: acquired_append)

	      from flush_store_buffer_append [where sb="ys@[Ghostsb A' L' R' W']" and sb'="zs", simplified,
	      OF j_bound_ts' isj [simplified split_suspendsj] cph [simplified suspendsj]
	      ts'_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys 
	      hist_consis' [simplified split_suspendsj] valid_sops_drop 
	      distinct_read_tmps_drop [simplified split_suspendsj] 
	      no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="share ?drop_sb 𝒮"]
	      obtain isj' j'  where
		isj': "instrs zs @ issbj = isj' @ prog_instrs zs" and
		steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮)  d* 
		  (?ts'[j:=(last_prog
                              (hd_prog pj (Ghostsb A' L' R' W'# zs)) (ys@[Ghostsb A' L' R' W']),
                             isj',
                             θsbj |` (dom θsbj - read_tmps zs),
                              (),
                             𝒟j  outstanding_refs is_volatile_Writesb (ys @ [Ghostsb A' L' R' W'])  {}, acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R',j')],
                    flush (ys@[Ghostsb A' L' R' W']) (flush ?drop_sb m),
                    share (ys@[Ghostsb 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 = (psb, issb,θsb,(), 
		𝒟sb, acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb)"
		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 issb] 
	      have a_unowned: 
		"j < length ?ts_ys. ij  (let (𝒪j) = map owned ?ts_ys!j in a  𝒪j)"
		apply cases
		apply (auto simp add: Let_def issb)
		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 tssb). i  j 
             (let (𝒪j, sbj) = map 𝒪_sb tssb! j
             in A  outstanding_refs is_volatile_Writesb sbj = {})"
      proof -
	{
	  fix j 𝒪j sbj
	  assume j_bound: "j < length (map owned tssb)"
	  assume neq_i_j: "ij"
	  assume tssb_j: "(map 𝒪_sb tssb)!j = (𝒪j,sbj)"
	  assume conflict: "A  outstanding_refs is_volatile_Writesb sbj  {}"
	  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 tssb"
	      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_Writesb sbj"
	      by auto

	    let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	    let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"

	    from ts_sim [rule_format, OF  j_bound''] tssb_j j_bound''
	    obtain pj suspendsj "issbj" 𝒟sbj 𝒟j j θsbj "isj" where
	      tssb_j: "tssb ! j = (pj,issbj, θsbj, sbj,𝒟sbj,𝒪j,j)" and
	      suspendsj: "suspendsj = ?drop_sbj" and
	      isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	      𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
	      tsj: "ts!j = (hd_prog pj suspendsj, isj,
	             θsbj |` (dom θsbj - read_tmps suspendsj),(), 𝒟j, 
                     acquired True ?take_sbj 𝒪j,
                     release ?take_sbj (dom 𝒮sb) j)"
	      apply (cases "tssb!j")
	      apply (force simp add: Let_def)
	      done
	      
	    have "a'  outstanding_refs is_volatile_Writesb suspendsj"
	    proof -	
	      from a'_in_j 
	      have "a'  outstanding_refs is_volatile_Writesb (?take_sbj @ ?drop_sbj)"
		by simp
	      thus ?thesis
		apply (simp only: outstanding_refs_append suspendsj)
		apply (auto simp add: outstanding_refs_conv dest: set_takeWhileD)
		done
	    qed
		
	    from split_volatile_Writesb_in_outstanding_refs [OF this]
	    obtain sop v ys zs A' L' R' W' where
	      split_suspendsj: "suspendsj = ys @ Writesb True a' sop v A' L' R' W'# zs" (is "suspendsj = ?suspends")
	      by blast
	    
	    from direct_memop_step.WriteVolatile [where  θ=θsb and m="flush ?drop_sb m"]
	    have "(Write True a (D, f) A L R W# issb', 
                       θsb, (), flush ?drop_sb m,𝒟sb,acquired True sb 𝒪sb,
                        release sb (dom 𝒮sb) sb, 
                        share ?drop_sb 𝒮)  
                    (issb', θsb, (), (flush ?drop_sb m)(a := f θsb), True, acquired True sb 𝒪sb  A - R, Map.empty,
                      share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)".

	    from direct_computation.concurrent_step.Memop [OF 
	      i_bound_ts' [simplified issb] ts'_i [simplified issb] this [simplified issb]] 
	    have store_step: "(?ts', flush ?drop_sb m,share ?drop_sb 𝒮 ) d 
                    (?ts'[i := (psb, issb', θsb, (), 
                         True, acquired True sb 𝒪sb  A - R,Map.empty)], 
                         (flush ?drop_sb m)(a := f θsb), share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL )"
		  (is " _ d (?ts_A, ?m_A, ?share_A)")
	     by (simp add: issb)
	      
	       
	   from i_bound' have i_bound'': "i < length ?ts_A"
	     by simp

	   from valid_program_history [OF j_bound'' tssb_j] 
	   have "causal_program_history issbj sbj".
	   then have cph: "causal_program_history issbj ?suspends"
	     apply -
	     apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
	     apply (simp only: split_suspendsj [symmetric] suspendsj)
	     apply (simp add: split_suspendsj)
	     done
	   
	   from tsj neq_i_j j_bound 
	   have ts_A_j: "?ts_A!j = (hd_prog pj (ys @ Writesb True a' sop v A' L' R' W'# zs), isj,
	     θsbj |` (dom θsbj - read_tmps (ys @ Writesb True a' sop v A' L' R' W'# zs)), (), 𝒟j, 
	     acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	     by (simp add: split_suspendsj)


	   from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
	     by simp

	   from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
	   then
	   have lp: "last_prog pj ?suspends = pj"
	     apply -
	     apply (rule last_prog_same_append [where sb="?take_sbj"])
	     apply (simp only: split_suspendsj [symmetric] suspendsj)
	     apply simp
	     done

	   from valid_reads [OF j_bound'' tssb_j]
	   have reads_consis: "reads_consistent False 𝒪j msb sbj".

	   from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb j_bound''
	     tssb_j reads_consis]
	   have reads_consis_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
	     by (simp add: m suspendsj)

	   from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j tssb_i tssb_j]
	   have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
	     by (simp add: suspendsj)
	   from reads_consistent_flush_independent [OF this reads_consis_m]
	   have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	     (flush ?drop_sb m) suspendsj".

	   from a_unowned_others [rule_format, OF _ neq_i_j] j_bound tssb_j
	   obtain a_notin_owns_j: "a  acquired True ?take_sbj 𝒪j" and a_unshared: "a  all_shared ?take_sbj"
	     by auto
	   from a_not_acquired_others [rule_format, OF _ neq_i_j] j_bound tssb_j
	   have a_not_acquired_j: "a  all_acquired sbj"
	     by auto
	   
	   from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' tssb_j]
	   have nvo_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j sbj".
	   
	   (* FIXME: make this a lemma, duplicated some times below *)
	   have a_no_non_vol_read: "a  outstanding_refs is_non_volatile_Readsb ?drop_sbj"
	   proof 
	     assume a_in_nvr:"a  outstanding_refs is_non_volatile_Readsb ?drop_sbj"

	     from reads_consistent_drop [OF reads_consis]
	     have rc: "reads_consistent True (acquired True ?take_sbj 𝒪j) (flush ?take_sbj msb) ?drop_sbj".

	     from non_volatile_owned_or_read_only_drop [OF nvo_j]
	     have nvo_j_drop: "non_volatile_owned_or_read_only True (share ?take_sbj 𝒮sb)
	       (acquired True ?take_sbj 𝒪j)
	       ?drop_sbj"
	       by simp

	     from outstanding_refs_non_volatile_Readsb_all_acquired [OF rc this a_in_nvr]

	     have a_owns_acq_ror: 
	       "a  𝒪j  all_acquired sbj  read_only_reads (acquired True ?take_sbj 𝒪j) ?drop_sbj"
	       by (auto dest!: acquired_all_acquired_in all_acquired_takeWhile_dropWhile_in
		 simp add: acquired_takeWhile_non_volatile_Writesb)

	     have a_unowned_j: "a  𝒪j  all_acquired sbj"
             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_sbj 𝒪j) ?drop_sbj"
	       by auto

	     with read_only_reads_unowned [OF j_bound'' i_bound neq_i_j [symmetric] tssb_j tssb_i]
	     have a_unowned_sb: "a  𝒪sb  all_acquired sb"
	       by auto
	       
	     from sharing_consis [OF j_bound'' tssb_j] sharing_consistent_append [of 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
	     have consis_j_drop: "sharing_consistent (share ?take_sbj 𝒮sb) (acquired True ?take_sbj 𝒪j) ?drop_sbj"
	       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_sbj ?drop_sbj] acquired_takeWhile_non_volatile_Writesb [of sbj 𝒪j]
	     have "a  read_only (share ?take_sbj 𝒮sb)"
	       by (auto simp add: )
             from read_only_share_all_shared [OF this] a_unshared
	     have "a  read_only 𝒮sb"
	       by fastforce
	      
	     from read_only_unacquired_share [OF read_only_unowned [OF i_bound tssb_i] 
	       weak_sharing_consis [OF i_bound tssb_i] this] a_unowned_sb
	     have "a  read_only (share sb 𝒮sb)"
	       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_sbj 𝒪j) ?m_A suspendsj"
	     by (auto simp add: suspendsj)


	   hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j) ?m_A ys"
	     by (simp add: split_suspendsj reads_consistent_append)

	   from valid_history [OF j_bound'' tssb_j]
	   have h_consis: 
	     "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
	     apply (simp only: split_suspendsj [symmetric] suspendsj)
	     apply simp
	     done

	   have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	   proof -
	     from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
	       by simp
	     from last_prog_hd_prog_append' [OF h_consis] this
	     have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
	       by (simp only: split_suspendsj [symmetric] suspendsj) 
	     moreover 
	     have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		  last_prog (hd_prog pj suspendsj) ?take_sbj"
	       apply (simp only: split_suspendsj [symmetric] suspendsj) 
	       by (rule last_prog_hd_prog_append)
	     ultimately show ?thesis
	       by (simp add: split_suspendsj [symmetric] suspendsj) 
	   qed

	   from valid_write_sops [OF j_bound'' tssb_j]
	   have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
	     by (simp add: split_suspendsj [symmetric] suspendsj)
	   then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
	     valid_sops_drop: "sopwrite_sops ys. valid_sop sop"
	     apply (simp only: write_sops_append )
	     apply auto
	     done

	   from read_tmps_distinct [OF j_bound'' tssb_j]
	   have "distinct_read_tmps (?take_sbj@suspendsj)"
	     by (simp add: split_suspendsj [symmetric] suspendsj)
	   then obtain 
	     read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
	     distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
	     apply (simp only: split_suspendsj [symmetric] suspendsj) 
	     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 θsbj (hd_prog pj suspendsj) suspendsj"
	     by (simp add: split_suspendsj [symmetric] suspendsj) 
	    from reads_consistent_drop_volatile_writes_no_volatile_reads  
	    [OF reads_consis] 
	    have no_vol_read: "outstanding_refs is_volatile_Readsb ys = {}"
	      by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )
	   
	    from flush_store_buffer_append [
	      OF j_bound''''  isj [simplified split_suspendsj] cph [simplified suspendsj]
	      ts_A_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_m_A_ys
	      hist_consis' [simplified split_suspendsj] valid_sops_drop distinct_read_tmps_drop [simplified split_suspendsj] 
	      no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="?share_A"]
	    obtain isj' j' where
	      isj': "instrs (Writesb True a' sop v A' L' R' W'# zs) @ issbj = 
	            isj' @ prog_instrs (Writesb 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 pj (Writesb True a' sop v A' L' R' W'# zs)) ys,
                           isj',
                           θsbj |` (dom θsbj - read_tmps (Writesb True a' sop v A' L' R' W' # zs)),(),
                           𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Writesb True a' sop v A' L' R' W']) @ zs)"
	      by simp
	    from causal_program_history_suffix [OF this]
	    have cph': "causal_program_history issbj zs".	      
	    interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

	    from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
	    obtain isj''
	      where isj': "isj' = Write True a' sop A' L' R' W'#isj''" and
	      isj'': "instrs zs @ issbj = isj'' @ 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 pj (Writesb True a' sop v A' L' R' W'# zs)) ys, isj',
                 θsbj |` (dom θsbj - read_tmps (Writesb True a' sop v A' L' R' W'# zs)),(),
	         𝒟j  outstanding_refs is_volatile_Writesb ys  {},
                 acquired True ys (acquired True ?take_sbj 𝒪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 isj']
	    have a_unowned: 
		"i < length ?ts_ys. ji  (let (𝒪i) = map owned ?ts_ys!i in a'  𝒪i)"
	      apply cases
	      apply (auto simp add: Let_def issb)
	      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 tssb). i  j 
             (let (𝒪j, sbj) = map 𝒪_sb tssb! j
             in A  all_acquired sbj = {})"
      proof -
	{
	  fix j 𝒪j sbj
	  assume j_bound: "j < length (map owned tssb)"
	  assume neq_i_j: "ij"
	  assume tssb_j: "(map 𝒪_sb tssb)!j = (𝒪j,sbj)"
	  assume conflict: "A  all_acquired sbj  {}"
	  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 tssb"
	      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 sbj"
	      by auto

	    let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	    let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"

	    from ts_sim [rule_format, OF  j_bound''] tssb_j j_bound''
	    obtain pj suspendsj "issbj" 𝒟sbj 𝒟j j θsbj "isj" where
	      tssb_j: "tssb ! j = (pj,issbj, θsbj, sbj,𝒟sbj,𝒪j,j)" and
	      suspendsj: "suspendsj = ?drop_sbj" and
	      isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	      𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
	      tsj: "ts!j = (hd_prog pj suspendsj, isj, 
	                   θsbj |` (dom θsbj - read_tmps suspendsj),(), 
	                   𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	      apply (cases "tssb!j")
	      apply (force simp add: Let_def)
	      done

	    from a'_in_j all_acquired_append [of ?take_sbj ?drop_sbj]
	    have "a'  all_acquired ?take_sbj  a'  all_acquired suspendsj"
	      by (auto simp add: suspendsj)
	    thus False
	    proof 
	      assume "a'  all_acquired ?take_sbj"
	      with A_unowned_by_others [rule_format, OF _ neq_i_j] tssb_j j_bound a'_in
	      show False
		by (auto dest: all_acquired_unshared_acquired)
	    next
	      assume conflict_drop: "a'  all_acquired suspendsj"
	      from split_all_acquired_in [OF conflict_drop]
	      (* FIXME: exract common parts *)
	      show False
	      proof 
		assume "sop a'' v ys zs A L R W. 
                         suspendsj = ys @ Writesb True a'' sop v A L R W# zs  a'  A" 
	        then
		obtain a'' sop' v' ys zs A' L' R' W' where
		  split_suspendsj: "suspendsj = ys @ Writesb True a'' sop' v' A' L' R' W'# zs" 
		    (is "suspendsj = ?suspends") and
		  a'_A': "a'  A'"
		 by auto
	    
	       from direct_memop_step.WriteVolatile [where  θ=θsb and m="flush ?drop_sb m"]
	       have "(Write True a (D, f) A L R W # issb', 
                         θsb, (), flush ?drop_sb m ,𝒟sb, acquired True sb 𝒪sb,
                         release sb (dom 𝒮sb) sb, 
                         share ?drop_sb 𝒮)  
                    (issb', θsb, (), (flush ?drop_sb m)(a := f θsb), True, acquired True sb 𝒪sb  A - R,Map.empty, 
                      share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)".

	       from direct_computation.concurrent_step.Memop [OF 
		 i_bound_ts' [simplified issb] ts'_i [simplified issb] this [simplified issb]] 

	       have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) d 
                   (?ts'[i := (psb, issb',
		        θsb, (),True, acquired True sb 𝒪sb  A - R,Map.empty)], 
                         (flush ?drop_sb m)(a := f θsb),share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)"
		   (is " _ d (?ts_A, ?m_A, ?share_A)")
		 by (simp add: issb)
	      
	       
	       from i_bound' have i_bound'': "i < length ?ts_A"
		 by simp

	       from valid_program_history [OF j_bound'' tssb_j] 
	       have "causal_program_history issbj sbj".
	       then have cph: "causal_program_history issbj ?suspends"
		 apply -
		 apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply (simp add: split_suspendsj)
		 done
	       
	       from tsj neq_i_j j_bound 
	       have ts_A_j: "?ts_A!j = (hd_prog pj (ys @ Writesb True a'' sop' v' A' L' R' W'# zs), isj, 
		   θsbj |` (dom θsbj - read_tmps (ys @ Writesb True a'' sop' v' A' L' R' W'# zs)), (), 𝒟j, 
		   acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
		 by (simp add: split_suspendsj)


	       from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
		 by simp

	       from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
	       then
	       have lp: "last_prog pj ?suspends = pj"
		 apply -
		 apply (rule last_prog_same_append [where sb="?take_sbj"])
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply simp
		 done
	       
	       from valid_reads [OF j_bound'' tssb_j]
	       have reads_consis: "reads_consistent False 𝒪j msb sbj".
	       
	       from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb 
		 j_bound''
		 tssb_j reads_consis]
	       have reads_consis_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		 by (simp add: m suspendsj)

	       from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j tssb_i tssb_j]
	       have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
		 by (simp add: suspendsj)
	       from reads_consistent_flush_independent [OF this reads_consis_m]
	       have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
		 (flush ?drop_sb m) suspendsj".
       
	       from a_unowned_others [rule_format, OF _ neq_i_j] j_bound tssb_j
	       obtain a_notin_owns_j: "a  acquired True ?take_sbj 𝒪j" and a_unshared: "a  all_shared ?take_sbj"
	         by auto
	       from a_not_acquired_others [rule_format, OF _ neq_i_j] j_bound tssb_j
	       have a_not_acquired_j: "a  all_acquired sbj"
		 by auto
	       
	       from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' tssb_j]
	       have nvo_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j sbj".

	       have a_no_non_vol_read: "a  outstanding_refs is_non_volatile_Readsb ?drop_sbj"
	       proof 
		 assume a_in_nvr:"a  outstanding_refs is_non_volatile_Readsb ?drop_sbj"
		 
		 from reads_consistent_drop [OF reads_consis]
		 have rc: "reads_consistent True (acquired True ?take_sbj 𝒪j) (flush ?take_sbj msb) ?drop_sbj".
		 
		 from non_volatile_owned_or_read_only_drop [OF nvo_j]
		 have nvo_j_drop: "non_volatile_owned_or_read_only True (share ?take_sbj 𝒮sb)
		   (acquired True ?take_sbj 𝒪j)
		   ?drop_sbj"
		   by simp
		 
		 from outstanding_refs_non_volatile_Readsb_all_acquired [OF rc this a_in_nvr]

		 have a_owns_acq_ror: 
		   "a  𝒪j  all_acquired sbj  read_only_reads (acquired True ?take_sbj 𝒪j) ?drop_sbj"
		   by (auto dest!: acquired_all_acquired_in all_acquired_takeWhile_dropWhile_in
		     simp add: acquired_takeWhile_non_volatile_Writesb)
		 have a_unowned_j: "a  𝒪j  all_acquired sbj"
                 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_sbj 𝒪j) ?drop_sbj"
		   by auto
		 
		 with read_only_reads_unowned [OF j_bound'' i_bound neq_i_j [symmetric] tssb_j tssb_i]
		 have a_unowned_sb: "a  𝒪sb  all_acquired sb"
		   by auto
		 
		 from sharing_consis [OF j_bound'' tssb_j] sharing_consistent_append [of 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
		 have consis_j_drop: "sharing_consistent (share ?take_sbj 𝒮sb) (acquired True ?take_sbj 𝒪j) ?drop_sbj"
		   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_sbj ?drop_sbj] acquired_takeWhile_non_volatile_Writesb [of sbj 𝒪j]
		 have "a  read_only (share ?take_sbj 𝒮sb)"
		   by (auto)
		 from read_only_share_all_shared [OF this] a_unshared
		 have "a  read_only 𝒮sb"
		   by fastforce
	      
		 from read_only_unacquired_share [OF read_only_unowned [OF i_bound tssb_i] 
		   weak_sharing_consis [OF i_bound tssb_i] this] a_unowned_sb
		 have "a  read_only (share sb 𝒮sb)"
		   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_sbj 𝒪j) ?m_A suspendsj"
		 by (auto simp add: suspendsj)
	       
	       hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j) ?m_A ys"
		 by (simp add: split_suspendsj reads_consistent_append)

	       
	       from valid_history [OF j_bound'' tssb_j]
	       have h_consis: 
		 "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply simp
		 done
	       
	       have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	       proof -
		 from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		   by simp
		 from last_prog_hd_prog_append' [OF h_consis] this
		 have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		   by (simp only: split_suspendsj [symmetric] suspendsj) 
		 moreover 
		 have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		   last_prog (hd_prog pj suspendsj) ?take_sbj"
		   apply (simp only: split_suspendsj [symmetric] suspendsj) 
		   by (rule last_prog_hd_prog_append)
		 ultimately show ?thesis
		   by (simp add: split_suspendsj [symmetric] suspendsj) 
	       qed
	       
	       from valid_write_sops [OF j_bound'' tssb_j]
	       have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		 by (simp add: split_suspendsj [symmetric] suspendsj)
	       then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		 valid_sops_drop: "sopwrite_sops ys. valid_sop sop"
		 apply (simp only: write_sops_append )
		 apply auto
		 done
	       
	       from read_tmps_distinct [OF j_bound'' tssb_j]
	       have "distinct_read_tmps (?take_sbj@suspendsj)"
		 by (simp add: split_suspendsj [symmetric] suspendsj)
	       then obtain 
		 read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
		 distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		 apply (simp only: split_suspendsj [symmetric] suspendsj) 
		 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 θsbj (hd_prog pj suspendsj) suspendsj"
		 by (simp add: split_suspendsj [symmetric] suspendsj) 
	       from reads_consistent_drop_volatile_writes_no_volatile_reads  
	       [OF reads_consis] 
	       have no_vol_read: "outstanding_refs is_volatile_Readsb ys = {}"
		 by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		   split_suspendsj )
	       
	       from flush_store_buffer_append [
		 OF j_bound''''  isj [simplified split_suspendsj] cph [simplified suspendsj]
		 ts_A_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_m_A_ys
		 hist_consis' [simplified split_suspendsj] valid_sops_drop distinct_read_tmps_drop [simplified split_suspendsj] 
		 no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
		 𝒮="?share_A"]
	       obtain isj' j' where
		 isj': "instrs (Writesb True a'' sop' v' A' L' R' W' # zs) @ issbj = 
	            isj' @ prog_instrs (Writesb 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 pj (Writesb True a'' sop' v' A' L' R' W' # zs)) ys,
                           isj',
                           θsbj |` (dom θsbj - read_tmps (Writesb True a'' sop' v' A' L' R' W' # zs)),(),
		           𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Writesb True a'' sop' v' A' L' R' W']) @ zs)"
		 by simp
	       from causal_program_history_suffix [OF this]
	       have cph': "causal_program_history issbj zs".	      
	       interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

	       from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
	       obtain isj''
		 where isj': "isj' = Write True a'' sop' A' L' R' W'#isj''" and
		 isj'': "instrs zs @ issbj = isj'' @ 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 pj (Writesb True a'' sop' v' A' L' R' W'# zs)) ys, isj',
                 θsbj |` (dom θsbj - read_tmps (Writesb True a'' sop' v' A' L' R' W'# zs)),(),𝒟j  outstanding_refs is_volatile_Writesb ys  {},
                 acquired True ys (acquired True ?take_sbj 𝒪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 isj']
	       have A'_unowned: 
		 "i < length ?ts_ys. ji  (let (𝒪i) = map owned ?ts_ys!i in A'   𝒪i = {})"
		 apply cases
		 apply (fastforce simp add: Let_def issb)+
		 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. 
                 suspendsj = ys @ Ghostsb A L R W # zs  a'  A" 
	       then
	       obtain ys zs A' L' R' W' where
		  split_suspendsj: "suspendsj = ys @ Ghostsb A' L' R' W'# zs" (is "suspendsj = ?suspends") and
		 a'_A': "a'  A'"
		 by auto
		 
	       from direct_memop_step.WriteVolatile [where  θ=θsb and m="flush ?drop_sb m"]
	       have "(Write True a (D, f) A L R W# issb', 
                          θsb, (), flush ?drop_sb m,𝒟sb,acquired True sb 𝒪sb, 
                          release sb (dom 𝒮sb) sb,
                         share ?drop_sb 𝒮)  
                    (issb', θsb, (), (flush ?drop_sb m)(a := f θsb), True, acquired True sb 𝒪sb  A - R, Map.empty, 
                      share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)".

	       from direct_computation.concurrent_step.Memop [OF 
		 i_bound_ts' [simplified issb] ts'_i [simplified issb] this [simplified issb]] 
	       have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) d 
                   (?ts'[i := (psb, issb', 
		         θsb, (), True, acquired True sb 𝒪sb  A - R,Map.empty)], 
                         (flush ?drop_sb m)(a := f θsb),share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)"
		   (is " _ d (?ts_A, ?m_A, ?share_A)")
		 by (simp add: issb)
	      
	       
	       from i_bound' have i_bound'': "i < length ?ts_A"
		 by simp

	       from valid_program_history [OF j_bound'' tssb_j] 
	       have "causal_program_history issbj sbj".
	       then have cph: "causal_program_history issbj ?suspends"
		 apply -
		 apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply (simp add: split_suspendsj)
		 done
	       
	       from tsj neq_i_j j_bound 
	       have ts_A_j: "?ts_A!j = (hd_prog pj (ys @ Ghostsb A' L' R' W'# zs), isj, 
		 θsbj |` (dom θsbj - read_tmps (ys @ Ghostsb A' L' R' W'# zs)), (),𝒟j, 
		 acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
		 by (simp add: split_suspendsj)


	       from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
		 by simp
	       
	       from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
	       then
	       have lp: "last_prog pj ?suspends = pj"
		 apply -
		 apply (rule last_prog_same_append [where sb="?take_sbj"])
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply simp
		 done
	       
	       from valid_reads [OF j_bound'' tssb_j]
	       have reads_consis: "reads_consistent False 𝒪j msb sbj".
	       
	       from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb 
		 j_bound''
		 tssb_j reads_consis]
	       have reads_consis_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		 by (simp add: m suspendsj)
	       
	       from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j tssb_i tssb_j]
	       have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
		 by (simp add: suspendsj)
	       from reads_consistent_flush_independent [OF this reads_consis_m]
	       have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
		 (flush ?drop_sb m) suspendsj".       

	       from a_unowned_others [rule_format, OF _ neq_i_j] j_bound tssb_j
	       obtain a_notin_owns_j: "a  acquired True ?take_sbj 𝒪j" and a_unshared: "a  all_shared ?take_sbj"
	         by auto
	       from a_not_acquired_others [rule_format, OF _ neq_i_j] j_bound tssb_j
	       have a_not_acquired_j: "a  all_acquired sbj"
		 by auto
	       
	       from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' tssb_j]
	       have nvo_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j sbj".

	       have a_no_non_vol_read: "a  outstanding_refs is_non_volatile_Readsb ?drop_sbj"
	       proof 
		 assume a_in_nvr:"a  outstanding_refs is_non_volatile_Readsb ?drop_sbj"
		 
		 from reads_consistent_drop [OF reads_consis]
		 have rc: "reads_consistent True (acquired True ?take_sbj 𝒪j) (flush ?take_sbj msb) ?drop_sbj".

		 from non_volatile_owned_or_read_only_drop [OF nvo_j]
		 have nvo_j_drop: "non_volatile_owned_or_read_only True (share ?take_sbj 𝒮sb)
		   (acquired True ?take_sbj 𝒪j)
		   ?drop_sbj"
		   by simp
		 
		 from outstanding_refs_non_volatile_Readsb_all_acquired [OF rc this a_in_nvr]
		 
		 have a_owns_acq_ror: 
		   "a  𝒪j  all_acquired sbj  read_only_reads (acquired True ?take_sbj 𝒪j) ?drop_sbj"
		   by (auto dest!: acquired_all_acquired_in all_acquired_takeWhile_dropWhile_in
		     simp add: acquired_takeWhile_non_volatile_Writesb)
		 
		 have a_unowned_j: "a  𝒪j  all_acquired sbj"
                 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_sbj 𝒪j) ?drop_sbj"
		   by auto
		 
		 with read_only_reads_unowned [OF j_bound'' i_bound neq_i_j [symmetric] tssb_j tssb_i]
		 have a_unowned_sb: "a  𝒪sb  all_acquired sb"
		   by auto
		 
		 from sharing_consis [OF j_bound'' tssb_j] sharing_consistent_append [of 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
		 have consis_j_drop: "sharing_consistent (share ?take_sbj 𝒮sb) (acquired True ?take_sbj 𝒪j) ?drop_sbj"
		   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_sbj ?drop_sbj] acquired_takeWhile_non_volatile_Writesb [of sbj 𝒪j]
		 have "a  read_only (share ?take_sbj 𝒮sb)"
		   by (auto)
		 from read_only_share_all_shared [OF this] a_unshared
		 have "a  read_only 𝒮sb"
		   by fastforce
		 
		 from read_only_unacquired_share [OF read_only_unowned [OF i_bound tssb_i] 
		   weak_sharing_consis [OF i_bound tssb_i] this] a_unowned_sb
		 have "a  read_only (share sb 𝒮sb)"
		   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_sbj 𝒪j) ?m_A suspendsj"
		 by (auto simp add: suspendsj)
	       

	       hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j) ?m_A ys"
		 by (simp add: split_suspendsj reads_consistent_append)       

	       from valid_history [OF j_bound'' tssb_j]
	       have h_consis: 
		 "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply simp
		 done
	       
	       have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	       proof -
		 from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		   by simp
		 from last_prog_hd_prog_append' [OF h_consis] this
		 have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		   by (simp only: split_suspendsj [symmetric] suspendsj) 
		 moreover 
		 have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		   last_prog (hd_prog pj suspendsj) ?take_sbj"
		   apply (simp only: split_suspendsj [symmetric] suspendsj) 
		   by (rule last_prog_hd_prog_append)
		 ultimately show ?thesis
		   by (simp add: split_suspendsj [symmetric] suspendsj) 
	       qed
	       
	       from valid_write_sops [OF j_bound'' tssb_j]
	       have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		 by (simp add: split_suspendsj [symmetric] suspendsj)
	       then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		   valid_sops_drop: "sopwrite_sops ys. valid_sop sop"
		 apply (simp only: write_sops_append )
		 apply auto
		 done
	       
	       from read_tmps_distinct [OF j_bound'' tssb_j]
	       have "distinct_read_tmps (?take_sbj@suspendsj)"
		 by (simp add: split_suspendsj [symmetric] suspendsj)
	       then obtain 
		 read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
		 distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		 apply (simp only: split_suspendsj [symmetric] suspendsj) 
		 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 θsbj (hd_prog pj suspendsj) suspendsj"
		 by (simp add: split_suspendsj [symmetric] suspendsj) 
	       from reads_consistent_drop_volatile_writes_no_volatile_reads  
	       [OF reads_consis] 
	       have no_vol_read: "outstanding_refs is_volatile_Readsb ys = {}"
		 by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		   split_suspendsj )
	   
	       from flush_store_buffer_append [
		 OF j_bound''''  isj [simplified split_suspendsj] cph [simplified suspendsj]
		 ts_A_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_m_A_ys
		 hist_consis' [simplified split_suspendsj] valid_sops_drop distinct_read_tmps_drop [simplified split_suspendsj] 
		 no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
		 𝒮="?share_A"]
	       obtain isj' j' where
		 isj': "instrs (Ghostsb A' L' R' W' # zs) @ issbj = 
	            isj' @ prog_instrs (Ghostsb A' L' R' W'# zs)" and
		 steps_ys: "(?ts_A, ?m_A, ?share_A)  d* 
		(?ts_A[j:= (last_prog (hd_prog pj (Ghostsb A' L' R' W'# zs)) ys,
                           isj',
                           θsbj |` (dom θsbj - read_tmps (Ghostsb A' L' R' W'# zs)),(),
		           𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Ghostsb A' L' R' W']) @ zs)"
		 by simp
	       from causal_program_history_suffix [OF this]
	       have cph': "causal_program_history issbj zs".	      
	       interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

	       from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
	       obtain isj''
		 where isj': "isj' = Ghost A' L' R' W'#isj''" and
		 isj'': "instrs zs @ issbj = isj'' @ 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 pj (Ghostsb A' L' R' W'# zs)) ys, isj',
                 θsbj |` (dom θsbj - read_tmps (Writesb True a'' sop' v' A' L' R' W'# zs)),(),𝒟j  outstanding_refs is_volatile_Writesb ys  {},
                 acquired True ys (acquired True ?take_sbj 𝒪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 isj']
	       have A'_unowned: 
		 "i < length ?ts_ys. ji  (let (𝒪i) = map owned ?ts_ys!i in A'   𝒪i = {})"
		 apply cases
		 apply (fastforce simp add: Let_def issb)+
		 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 tssb). i  j 
             (let (𝒪j, sbj) = map 𝒪_sb tssb! j
             in A  read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j)
	             (dropWhile (Not  is_volatile_Writesb) sbj) = {})"
      proof -
	{
	  fix j 𝒪j sbj
	  assume j_bound: "j < length (map 𝒪_sb tssb)"
	  assume neq_i_j: "ij"
	  assume tssb_j: "(map 𝒪_sb tssb)!j = (𝒪j,sbj)"
	  let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	  let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"

	  assume conflict: "A  read_only_reads (acquired True ?take_sbj 𝒪j) ?drop_sbj   {}"
	  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 tssb"
	      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_sbj 𝒪j) ?drop_sbj"
	      by auto


	    from ts_sim [rule_format, OF  j_bound''] tssb_j j_bound''
	    obtain pj suspendsj "issbj" 𝒟sbj 𝒟j j θsbj "isj" where
	      tssb_j: "tssb ! j = (pj,issbj, θsbj, sbj,𝒟sbj,𝒪j,j)" and
	      suspendsj: "suspendsj = ?drop_sbj" and
	      isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	      𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
	      tsj: "ts!j = (hd_prog pj suspendsj, isj,
	             θsbj |` (dom θsbj - read_tmps suspendsj),(), 𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	      apply (cases "tssb!j")
	      apply (force simp add: Let_def)
	      done
	      

	    from split_in_read_only_reads [OF a'_in_j [simplified suspendsj [symmetric]]]
	    obtain t v ys zs where
	      split_suspendsj: "suspendsj = ys @ Readsb False a' t v# zs" (is "suspendsj = ?suspends") and
	      a'_unacq: "a'  acquired True ys (acquired True ?take_sbj 𝒪j)"
	      by blast
	    
	    from direct_memop_step.WriteVolatile [where  θ=θsb and m="flush ?drop_sb m"]
	    have "(Write True a (D, f) A L R W# issb', 
                  θsb, (), flush ?drop_sb m, 𝒟sb,acquired True sb 𝒪sb, 
                  release sb (dom 𝒮sb) sb, share ?drop_sb 𝒮)  
                    (issb', θsb, (), (flush ?drop_sb m)(a := f θsb), True, acquired True sb 𝒪sb  A - R,Map.empty, 
                      share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)".

	    from direct_computation.concurrent_step.Memop [OF 
	      i_bound_ts' [simplified issb] ts'_i [simplified issb] this [simplified issb]] 
	    have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) d 
                    (?ts'[i := (psb, issb', θsb, (), 
                         True, acquired True sb 𝒪sb  A - R,Map.empty)], 
                         (flush ?drop_sb m)(a := f θsb),share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)"
		  (is " _ d (?ts_A, ?m_A, ?share_A)")
	     by (simp add: issb)
	      
	       
	   from i_bound' have i_bound'': "i < length ?ts_A"
	     by simp

	   from valid_program_history [OF j_bound'' tssb_j] 
	   have "causal_program_history issbj sbj".
	   then have cph: "causal_program_history issbj ?suspends"
	     apply -
	     apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
	     apply (simp only: split_suspendsj [symmetric] suspendsj)
	     apply (simp add: split_suspendsj)
	     done
	   
	   from tsj neq_i_j j_bound 
	   have ts_A_j: "?ts_A!j = (hd_prog pj (ys @ Readsb False a' t v# zs), isj,
	     θsbj |` (dom θsbj - read_tmps (ys @ Readsb False a' t v# zs)), (), 𝒟j, 
	     acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	     by (simp add: split_suspendsj)


	   from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
	     by simp

	   from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
	   then
	   have lp: "last_prog pj ?suspends = pj"
	     apply -
	     apply (rule last_prog_same_append [where sb="?take_sbj"])
	     apply (simp only: split_suspendsj [symmetric] suspendsj)
	     apply simp
	     done

	   from valid_reads [OF j_bound'' tssb_j]
	   have reads_consis: "reads_consistent False 𝒪j msb sbj".

	   from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb j_bound''
	     tssb_j reads_consis]
	   have reads_consis_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
	     by (simp add: m suspendsj)

	   from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j tssb_i tssb_j]
	   have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
	     by (simp add: suspendsj)
	   from reads_consistent_flush_independent [OF this reads_consis_m]
	   have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	     (flush ?drop_sb m) suspendsj".

	   from a_unowned_others [rule_format, OF j_bound'' neq_i_j ] j_bound tssb_j
	   obtain a_notin_owns_j: "a  acquired True ?take_sbj 𝒪j" and a_unshared: "a  all_shared ?take_sbj"
	     by auto
	   from a_not_acquired_others [rule_format, OF j_bound neq_i_j] j_bound tssb_j
	   have a_not_acquired_j: "a  all_acquired sbj"
	     by auto
	   
	   from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' tssb_j]
	   have nvo_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j sbj".
	   
	   have a_no_non_vol_read: "a  outstanding_refs is_non_volatile_Readsb ?drop_sbj"
	   proof 
	     assume a_in_nvr:"a  outstanding_refs is_non_volatile_Readsb ?drop_sbj"

	     from reads_consistent_drop [OF reads_consis]
	     have rc: "reads_consistent True (acquired True ?take_sbj 𝒪j) (flush ?take_sbj msb) ?drop_sbj".

	     from non_volatile_owned_or_read_only_drop [OF nvo_j]
	     have nvo_j_drop: "non_volatile_owned_or_read_only True (share ?take_sbj 𝒮sb)
	       (acquired True ?take_sbj 𝒪j)
	       ?drop_sbj"
	       by simp

	     from outstanding_refs_non_volatile_Readsb_all_acquired [OF rc this a_in_nvr]

	     have a_owns_acq_ror: 
	       "a  𝒪j  all_acquired sbj  read_only_reads (acquired True ?take_sbj 𝒪j) ?drop_sbj"
	       by (auto dest!: acquired_all_acquired_in all_acquired_takeWhile_dropWhile_in
		 simp add: acquired_takeWhile_non_volatile_Writesb)
             
	     have a_unowned_j: "a  𝒪j  all_acquired sbj"
             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_sbj 𝒪j) ?drop_sbj"
	       by auto

	     with read_only_reads_unowned [OF j_bound'' i_bound neq_i_j [symmetric] tssb_j tssb_i]
	     have a_unowned_sb: "a  𝒪sb  all_acquired sb"
	       by auto
	       
	     from sharing_consis [OF j_bound'' tssb_j] sharing_consistent_append [of 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
	     have consis_j_drop: "sharing_consistent (share ?take_sbj 𝒮sb) (acquired True ?take_sbj 𝒪j) ?drop_sbj"
	       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_sbj ?drop_sbj] acquired_takeWhile_non_volatile_Writesb [of sbj 𝒪j]
	     have "a  read_only (share ?take_sbj 𝒮sb)"
	       by (auto)
	     from read_only_share_all_shared [OF this] a_unshared
	     have "a  read_only 𝒮sb"
	       by fastforce
	      
	     from read_only_unacquired_share [OF read_only_unowned [OF i_bound tssb_i] 
	       weak_sharing_consis [OF i_bound tssb_i] this] a_unowned_sb
	     have "a  read_only (share sb 𝒮sb)"
	       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_sbj 𝒪j) ?m_A suspendsj"
	     by (auto simp add: suspendsj)


	   hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j) ?m_A ys"
	     by (simp add: split_suspendsj reads_consistent_append)

	   from valid_history [OF j_bound'' tssb_j]
	   have h_consis: 
	     "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
	     apply (simp only: split_suspendsj [symmetric] suspendsj)
	     apply simp
	     done

	   have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	   proof -
	     from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
	       by simp
	     from last_prog_hd_prog_append' [OF h_consis] this
	     have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
	       by (simp only: split_suspendsj [symmetric] suspendsj) 
	     moreover 
	     have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		  last_prog (hd_prog pj suspendsj) ?take_sbj"
	       apply (simp only: split_suspendsj [symmetric] suspendsj) 
	       by (rule last_prog_hd_prog_append)
	     ultimately show ?thesis
	       by (simp add: split_suspendsj [symmetric] suspendsj) 
	   qed

	   from valid_write_sops [OF j_bound'' tssb_j]
	   have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
	     by (simp add: split_suspendsj [symmetric] suspendsj)
	   then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
	     valid_sops_drop: "sopwrite_sops ys. valid_sop sop"
	     apply (simp only: write_sops_append )
	     apply auto
	     done

	   from read_tmps_distinct [OF j_bound'' tssb_j]
	   have "distinct_read_tmps (?take_sbj@suspendsj)"
	     by (simp add: split_suspendsj [symmetric] suspendsj)
	   then obtain 
	     read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
	     distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
	     apply (simp only: split_suspendsj [symmetric] suspendsj) 
	     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 θsbj (hd_prog pj suspendsj) suspendsj"
	     by (simp add: split_suspendsj [symmetric] suspendsj) 
	    from reads_consistent_drop_volatile_writes_no_volatile_reads  
	    [OF reads_consis] 
	    have no_vol_read: "outstanding_refs is_volatile_Readsb ys = {}"
	      by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )
	   
	    from flush_store_buffer_append [
	      OF j_bound''''  isj [simplified split_suspendsj] cph [simplified suspendsj]
	      ts_A_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_m_A_ys
	      hist_consis' [simplified split_suspendsj] valid_sops_drop distinct_read_tmps_drop [simplified split_suspendsj] 
	      no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="?share_A"]
	    obtain isj' j' where
	      isj': "instrs (Readsb False a' t v# zs) @ issbj = 
	            isj' @ prog_instrs (Readsb False a' t v# zs)" and
	      steps_ys: "(?ts_A, ?m_A, ?share_A)  d* 
		(?ts_A[j:= (last_prog (hd_prog pj (Readsb False a' t v# zs)) ys,
                           isj',
                           θsbj |` (dom θsbj - read_tmps (Readsb False a' t v# zs)),(),
                           𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Readsb False a' t v]) @ zs)"
	      by simp
	    from causal_program_history_suffix [OF this]
	    have cph': "causal_program_history issbj zs".	      
	    interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

	    from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
	    obtain isj''
	      where isj': "isj' = Read False a' t#isj''" and
	      isj'': "instrs zs @ issbj = isj'' @ 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 pj (Readsb False a' t v# zs)) ys, isj',
                 θsbj |` (dom θsbj - read_tmps (Readsb False a' t v# zs)),(),
	         𝒟j  outstanding_refs is_volatile_Writesb ys  {},
                 acquired True ys (acquired True ?take_sbj 𝒪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 isj']
	    have "a'  acquired True ys (acquired True ?take_sbj 𝒪j) 
                  a'  read_only (share ys (share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL))"
	      apply cases
	      apply (auto simp add: Let_def issb)
	      done
	    with a'_unacq
	    have a'_ro: "a'  read_only (share ys (share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL))"
	      by auto
	    from a'_in
	    have a'_not_ro: "a'  read_only (share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)"
	      by (auto simp add:  in_read_only_convs)

	    have "a'  𝒪j  all_acquired sbj"
	    proof -
	      {
		assume a_notin: "a'  𝒪j  all_acquired sbj"
		from weak_sharing_consis [OF j_bound'' tssb_j]
		have "weak_sharing_consistent 𝒪j sbj".
		with weak_sharing_consistent_append [of 𝒪j ?take_sbj ?drop_sbj]
		have "weak_sharing_consistent (acquired True ?take_sbj 𝒪j) suspendsj"
		  by (auto simp add: suspendsj)
                
		with split_suspendsj
		have weak_consis: "weak_sharing_consistent (acquired True ?take_sbj 𝒪j) ys"
		  by (simp add: weak_sharing_consistent_append)
		from all_acquired_append [of ?take_sbj ?drop_sbj]
		have "all_acquired ys  all_acquired sbj"
		  apply (clarsimp)
		  apply (clarsimp simp add: suspendsj [symmetric] split_suspendsj all_acquired_append)
		  done

                with a_notin acquired_takeWhile_non_volatile_Writesb [of sbj 𝒪j]
                  all_acquired_append [of ?take_sbj ?drop_sbj]
		have "a'  acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j  all_acquired ys"
                  by auto
                
		from read_only_share_unowned [OF weak_consis this a'_ro] 
		have "a'  read_only (share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)" .
		  
		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] tssb_j j_bound
	    have "A  all_acquired sbj = {}"
	      by (auto simp add: Let_def)
	    moreover
	    from A_unowned_by_others [rule_format, OF j_bound'' neq_i_j] tssb_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 𝒮sb' tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	proof -
	  from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i] 
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb (sb @ [Writesb True a (D,f) (f θsb) 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: tssb' sb' 𝒪sb' 𝒮sb')
	qed
      next
	show "outstanding_volatile_writes_unowned_by_others tssb'"
	proof (unfold_locales)
	  fix i1 j p1 "is1" 𝒪1 1 𝒟1 xs1 sb1 pj "isj" "𝒪j" j 𝒟j xsj sbj
	  assume i1_bound: "i1 < length tssb'"
	  assume j_bound: "j < length tssb'"
	  assume i1_j: "i1  j"
	  assume ts_i1: "tssb'!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	  assume ts_j: "tssb'!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb1 = {}"
	  proof (cases "i1=i")
	    case True
	    with i1_j have i_j: "ij" 
	      by simp
	    
	    from j_bound have j_bound': "j < length tssb"
	      by (simp add: tssb')
	    hence j_bound'': "j < length (map owned tssb)"
	      by simp
	    from ts_j i_j have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (simp add: tssb')
	    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_Writesb) sbj) 𝒪j" and
              a_unshared: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)"
	      by (auto simp add: Let_def tssb')
	    from a_not_acquired_others [rule_format, OF _ i_j] i_j ts_j j_bound
	    have a_notin_acq: "a  all_acquired sbj"
	      by (auto simp add: Let_def tssb')
	    from outstanding_volatile_writes_unowned_by_others 
	    [OF i_bound j_bound' i_j tssb_i ts_j']
	    have "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb = {}".
	    with ts_i1 a_notin_j a_unshared a_notin_acq True i_bound show ?thesis
	      by (auto simp add: tssb' sb' outstanding_refs_append 
		acquired_takeWhile_non_volatile_Writesb dest: all_shared_acquired_in)
	  next
	    case False
	    note i1_i = this
	    from i1_bound have i1_bound': "i1 < length tssb"
	      by (simp add: tssb')
	    from ts_i1 False have ts_i1': "tssb!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	      by (simp add: tssb')
	    show ?thesis
	    proof (cases "j=i")
	      case True

	      from i1_bound'
	      have i1_bound'': "i1 < length (map owned tssb)"
		by simp

	      from outstanding_volatile_writes_unowned_by_others 
	      [OF i1_bound' i_bound i1_i ts_i1' tssb_i]
	      have "(𝒪sb  all_acquired sb)  outstanding_refs is_volatile_Writesb sb1 = {}".
	      moreover
	      from A_unused_by_others [rule_format, OF _ False [symmetric]] False ts_i1 i1_bound
	      have "A  outstanding_refs is_volatile_Writesb sb1 = {}"
		by (auto simp add: Let_def tssb')
	      
	      ultimately
	      show ?thesis
		using ts_j True tssb' 
		by (auto simp add: i_bound tssb' 𝒪sb' sb' all_acquired_append)
	    next
	      case False
	      from j_bound have j_bound': "j < length tssb"
		by (simp add: tssb')
	      from ts_j False have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (simp add: tssb')
	      from outstanding_volatile_writes_unowned_by_others 
              [OF i1_bound' j_bound' i1_j ts_i1' ts_j']
	      show "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb1 = {}" .
	    qed
	  qed
	qed
      next
	show "ownership_distinct tssb'"
	proof -
	  have "j<length tssb. i  j 
	    (let (pj, isj, θj, sbj, 𝒟j, 𝒪j,j) = tssb ! j
	      in (𝒪sb  all_acquired sb')  (𝒪j  all_acquired sbj) = {})"
	  proof -
	    {
	      fix j pj isj 𝒪j j 𝒟j acqj θj sbj
	      assume neq_i_j: "i  j"
	      assume j_bound: "j < length tssb"
	      assume tssb_j: "tssb ! j = (pj, isj, θj, sbj, 𝒟j, 𝒪j,j)"
	      have "(𝒪sb  all_acquired sb')  (𝒪j  all_acquired sbj) = {}"
	      proof -
		{
		  fix a'
		  assume a'_in_i: "a'  (𝒪sb  all_acquired sb')"
		  assume a'_in_j: "a'  (𝒪j  all_acquired sbj)"
		  have False
		  proof -
		    from a'_in_i have "a'  (𝒪sb  all_acquired sb)  a'  A"
		      by (simp add: sb' all_acquired_append)
		    then show False
		    proof 
		      assume "a'  (𝒪sb  all_acquired sb)"
		      with ownership_distinct [OF i_bound j_bound neq_i_j tssb_i tssb_j] a'_in_j
		      show ?thesis
			by auto
		    next
		      assume "a'  A"
		      moreover
		      have j_bound': "j < length (map owned tssb)"
			using j_bound by auto
		      from A_unowned_by_others [rule_format, OF _ neq_i_j] tssb_j j_bound
		      obtain "A  acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j = {}" and
                             "A  all_shared (takeWhile (Not  is_volatile_Writesb) sbj) = {}"
			by (auto simp add: Let_def)
		      moreover
		      from A_unaquired_by_others [rule_format, OF _ neq_i_j] tssb_j j_bound
		      have "A  all_acquired sbj = {}"
			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 tssb_i this]
	  show ?thesis by (simp add: tssb' 𝒪sb' sb')
	qed
      next
	show "read_only_reads_unowned tssb'"
	proof 
	  fix n m
	  fix pn "isn" 𝒪n n 𝒟n θn sbn pm "ism" 𝒪m m 𝒟m θm sbm
	  assume n_bound: "n < length tssb'"
	    and m_bound: "m < length tssb'"
	    and neq_n_m: "nm"
	    and nth: "tssb'!n = (pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
	    and mth: "tssb'!m =(pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
	  from n_bound have n_bound': "n < length tssb" by (simp add: tssb')
	  from m_bound have m_bound': "m < length tssb" by (simp add: tssb')
	  
	  show "(𝒪m  all_acquired sbm) 
            read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sbn) 𝒪n)
            (dropWhile (Not  is_volatile_Writesb) sbn) =
            {}"
	  proof (cases "m=i")
	    case True
	    with neq_n_m have neq_n_i: "ni"
	      by auto
	    with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
	      by (auto simp add: tssb')
	    note read_only_reads_unowned [OF n_bound' i_bound  neq_n_i nth' tssb_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_Writesb) sbn) 𝒪n)
              (dropWhile (Not  is_volatile_Writesb) sbn) =
              {}"
	      by auto
	    ultimately 
	    show ?thesis
	      using True tssb_i nth' mth n_bound' m_bound'
	      by (auto simp add: tssb' 𝒪sb' sb' all_acquired_append)
	  next
	    case False
	    note neq_m_i = this
	    with m_bound mth i_bound have mth': "tssb!m = (pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
	      by (auto simp add: tssb')
	    show ?thesis
	    proof (cases "n=i")
	      case True
	      note read_only_reads_unowned [OF i_bound m_bound' neq_m_i [symmetric] tssb_i mth']
	      then show ?thesis
		using True neq_m_i tssb_i nth mth n_bound' m_bound'
		apply (case_tac "outstanding_refs (is_volatile_Writesb) sb = {}")
		apply (clarsimp simp add: outstanding_vol_write_take_drop_appends
		  acquired_append read_only_reads_append tssb' sb' 𝒪sb')+
		done
	    next
	      case False
	      with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
		by (auto simp add: tssb')
	      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 tssb'"
      proof -
	from valid_history [OF i_bound tssb_i]
	have "history_consistent θsb (hd_prog psb sb) sb".
	with valid_write_sops [OF i_bound tssb_i] D_subset 
	  valid_implies_valid_prog_hd [OF i_bound tssb_i valid]
	have "history_consistent θsb (hd_prog psb (sb@[Writesb True a (D,f) (f θsb) A L R W])) 
	         (sb@ [Writesb True a (D,f) (f θsb) A L R W])"
	  apply -
	  apply (rule history_consistent_appendI)
	  apply (auto simp add: hd_prog_append_Writesb)
	  done
	from valid_history_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' θsb')
      qed

      have valid_reads': "valid_reads msb tssb'"
      proof -
	from valid_reads [OF i_bound tssb_i]
	have "reads_consistent False 𝒪sb msb sb" .
	from reads_consistent_snoc_Writesb [OF this]
	have "reads_consistent False 𝒪sb msb (sb @ [Writesb True a (D,f) (f θsb) A L R W])".
	from valid_reads_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' 𝒪sb')
      qed

      have valid_sharing': "valid_sharing 𝒮sb' tssb'"
      proof (intro_locales)	
	from outstanding_non_volatile_writes_unshared [OF i_bound tssb_i] 
	have "non_volatile_writes_unshared 𝒮sb (sb @ [Writesb True a (D,f) (f θsb) 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 𝒮sb' tssb'"
	  by (simp add: tssb' sb' 𝒮sb')
      next
	from sharing_consis [OF i_bound tssb_i]
	have consis': "sharing_consistent 𝒮sb 𝒪sb sb".
	from  A_shared_owned
	have "A  dom (share ?drop_sb 𝒮)  acquired True sb 𝒪sb"
	  by (simp add:  sharing_consistent_append  acquired_takeWhile_non_volatile_Writesb)
	moreover have "dom (share ?drop_sb 𝒮)  dom 𝒮  dom (share sb 𝒮sb)"
	proof
	  fix a'
	  assume a'_in: "a'  dom (share ?drop_sb 𝒮)" 
	  from share_unshared_in [OF a'_in]
	  show "a'  dom 𝒮  dom (share sb 𝒮sb)"
	  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 𝒮sb)"
	      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 𝒮sb)  acquired True sb 𝒪sb"
	  by auto

        with A_unowned_by_others 
        
        have "A  dom (share sb 𝒮sb)  acquired True sb 𝒪sb"
        proof -
          {
            fix x
            assume x_A: "x  A"
            have "x  dom (share sb 𝒮sb)  acquired True sb 𝒪sb"
            proof -
              {
                assume "x  dom 𝒮"
                
                from share_all_until_volatile_write_share_acquired [OF sharing_consis 𝒮sb tssb 
                  i_bound tssb_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 𝒮sb 𝒪sb (sb @ [Writesb True a (D,f) (f θsb) A L R W])"
	  by (simp add:  sharing_consistent_append  acquired_takeWhile_non_volatile_Writesb)
	from sharing_consis_nth_update [OF i_bound this]
	show "sharing_consis 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒪sb' sb' 𝒮sb')
      next
	from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound tssb_i] ]
	show "read_only_unowned 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' 𝒪sb')
      next
	from unowned_shared_nth_update [OF i_bound tssb_i subset_refl]
	show "unowned_shared 𝒮sb' tssb'"
	  by (simp add: tssb' sb' 𝒪sb' 𝒮sb')
      next
	from a_not_ro no_outstanding_write_to_read_only_memory [OF i_bound tssb_i] 
	have "no_write_to_read_only_memory 𝒮sb (sb @ [Writesb True a (D,f) (f θsb) 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 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' sb')
      qed

      have tmps_distinct': "tmps_distinct tssb'"
      proof (intro_locales)
	from load_tmps_distinct [OF i_bound tssb_i]
	have "distinct_load_tmps issb'" by (simp add: "issb")
	from load_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_distinct tssb'" by (simp add: tssb')
      next
	from read_tmps_distinct [OF i_bound tssb_i]
	have "distinct_read_tmps (sb @ [Writesb True a (D, f) (f θsb) 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 tssb'" by (simp add: tssb' sb')
      next
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i]
	have "load_tmps issb'  read_tmps (sb @ [Writesb True a (D, f) (f θsb) A L R W]) ={}"
	  by (auto simp add: read_tmps_append "issb")
	from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_read_tmps_distinct tssb'" by (simp add: tssb' sb')
      qed
      
      have valid_sops': "valid_sops tssb'"
      proof -
	from valid_store_sops [OF i_bound tssb_i]
	obtain valid_Df: "valid_sop (D,f)" and  
	  valid_store_sops': "sopstore_sops issb'. valid_sop sop"
	  by (auto simp add: "issb")
	from valid_Df valid_write_sops [OF i_bound tssb_i]
	have valid_write_sops': "sopwrite_sops (sb@ [Writesb True a (D, f) (f θsb) 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: tssb' sb')
      qed

      have valid_dd': "valid_data_dependency tssb'"
      proof -
	from data_dependency_consistent_instrs [OF i_bound tssb_i]
	obtain D_indep: "D  load_tmps issb' = {}" and 
	  dd_is: "data_dependency_consistent_instrs (dom θsb') issb'"
	  by (auto simp add: "issb" θsb')
	from load_tmps_write_tmps_distinct [OF i_bound tssb_i] D_indep
	have "load_tmps issb'  (fst ` write_sops (sb@ [Writesb True a (D, f) (f θsb) A L R W])) ={}"
	  by (auto simp add: write_sops_append "issb")
	from valid_data_dependency_nth_update [OF i_bound dd_is this]
	show ?thesis by (simp add: tssb' sb')
      qed

      have load_tmps_fresh': "load_tmps_fresh tssb'"
      proof -
	from load_tmps_fresh [OF i_bound tssb_i] 
	have "load_tmps issb'  dom θsb = {}"
	  by (auto simp add: "issb")
	from load_tmps_fresh_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' θsb')
      qed

      have enough_flushs': "enough_flushs tssb'"
      proof -
	from clean_no_outstanding_volatile_Writesb [OF i_bound tssb_i]
	have "¬ True  outstanding_refs is_volatile_Writesb (sb@[Writesb True a (D,f) (f θsb) 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: tssb' sb' 𝒟sb')
      qed

      have valid_program_history': "valid_program_history tssb'"
      proof -	
	from valid_program_history [OF i_bound tssb_i]
	have "causal_program_history issb sb" .
	then have causal': "causal_program_history issb' (sb@[Writesb True a (D,f) (f θsb) A L R W])"
	  by (auto simp: causal_program_history_Write  "issb")
	from valid_last_prog [OF i_bound tssb_i]
	have "last_prog psb sb = psb".
	hence "last_prog psb (sb @ [Writesb True a (D,f) (f θsb) A L R W]) = psb"
	  by (simp add: last_prog_append_Writesb)
	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb' sb')
      qed

      show ?thesis
      proof (cases "outstanding_refs is_volatile_Writesb sb = {}")
	case True
	
	from True have flush_all: "takeWhile (Not  is_volatile_Writesb) sb = sb"
	  by (auto simp add: outstanding_refs_conv)
	
	from True have suspend_nothing: "dropWhile (Not  is_volatile_Writesb) 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# issb'"
	  by (simp add: "issb")
	with suspends_empty ts_i 
	have ts_i: "ts!i = (psb, Write True a (D,f) A L R W# issb', θsb,(),𝒟, acquired True ?take_sb 𝒪sb, release ?take_sb (dom 𝒮sb) sb)"
	  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 tssb_i]


	from True 
	have drop_app: "dropWhile (Not  is_volatile_Writesb) 
	  (sb@[Writesb True a (D,f) (f θsb) A L R W]) =
          [Writesb True a (D,f) (f θsb) A L R W]"
	  by (auto simp add: outstanding_refs_conv)

	have "(tssb',msb,𝒮sb')  (ts,m,𝒮)"
	  apply (rule sim_config.intros) 
	  apply    (simp add: m flush_commute tssb' θsb' 𝒪sb' sb' sb') 	  
	  using  share_all_until_volatile_write_Write_commute 
	          [OF i_bound tssb_i [simplified issb]]
	  apply   (clarsimp simp add: 𝒮 𝒮sb' tssb' sb' 𝒪sb' sb' θsb')
	  using  leq
	  apply  (simp add: tssb')
	  using i_bound i_bound' ts_sim ts_i 
	  apply (clarsimp simp add: Let_def nth_list_update drop_app (* drop*) 
	    tssb' sb' 𝒪sb' sb' 𝒮sb' θsb' 𝒟sb'  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' msb' 𝒮sb'
	  by auto
      next
	case False

	then obtain r where r_in: "r  set sb" and volatile_r: "is_volatile_Writesb r"
	  by (auto simp add: outstanding_refs_conv)
	from takeWhile_dropWhile_real_prefix 
	[OF r_in, of  "(Not  is_volatile_Writesb)", simplified, OF volatile_r] 
	obtain a' v' sb'' A'' L'' R'' W'' sop' where
	  sb_split: "sb = takeWhile (Not  is_volatile_Writesb) sb @ Writesb True a' sop' v' A'' L'' R'' W''# sb''" 
	  and
	  drop: "dropWhile (Not  is_volatile_Writesb) sb = Writesb 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 = Writesb 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 tssb_i]

	have "Writesb 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_Writesb)", simplified]

	have "(tssb',msb,𝒮sb')  (ts,m,𝒮)"
	  apply (rule sim_config.intros) 
	  apply    (simp add: m flush_commute tssb' 𝒪sb' sb' θsb' sb')
	  using  share_all_until_volatile_write_Write_commute 
	          [OF i_bound tssb_i [simplified issb]]
	  apply   (clarsimp simp add: 𝒮 𝒮sb' tssb' sb' 𝒪sb' sb' θsb')
	  using  leq
	  apply  (simp add: tssb')
	  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_Writesb instrs_append_Writesb hd_prog_append_Writesb
	    drop "issb" tssb' sb' 𝒪sb' 𝒮sb' sb' θsb' 𝒟sb' 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' msb' 𝒮sb' 
	  by (auto simp del: fun_upd_apply )
      qed
    next
      case SBHFence
      then obtain 
	"issb": "issb = Fence # issb'" and
	sb: "sb=[]" and
	𝒪sb': "𝒪sb'=𝒪sb" and
        sb': "sb'=Map.empty" and
	θsb': "θsb' = θsb" and
	𝒟sb': "¬ 𝒟sb'" and
	sb': "sb'=sb" and
	msb': "msb' = msb" and
	𝒮sb': "𝒮sb'=𝒮sb" 
	by auto

      have valid_own': "valid_ownership 𝒮sb' tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	proof -
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb []"
	    by simp
	  from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
	  show ?thesis by (simp add: tssb' sb' sb 𝒪sb' 𝒮sb')
	qed
      next
	from outstanding_volatile_writes_unowned_by_others_store_buffer 
	[OF i_bound tssb_i subset_refl]
	show "outstanding_volatile_writes_unowned_by_others tssb'" 
	  by (simp add: tssb' sb' sb 𝒪sb')
      next
	from read_only_reads_unowned_nth_update [OF i_bound tssb_i, of "[]" 𝒪sb]
	show "read_only_reads_unowned tssb'"
	  by (simp add: tssb' sb' sb 𝒪sb')
      next
	from ownership_distinct_instructions_read_value_store_buffer_independent 
	[OF i_bound tssb_i]
	show "ownership_distinct tssb'"
	  by (simp add: tssb' sb' sb 𝒪sb')
      qed


      have valid_hist': "valid_history program_step tssb'"
      proof -
	from valid_history [OF i_bound tssb_i] 
	have "history_consistent θsb (hd_prog psb []) []" by simp
	from valid_history_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' sb 𝒪sb' θsb')
      qed
      
      have valid_reads': "valid_reads msb tssb'"
      proof -
	have "reads_consistent False 𝒪sb msb []" by simp
	from valid_reads_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' sb 𝒪sb')
      qed


      have valid_sharing': "valid_sharing 𝒮sb' tssb'"
      proof (intro_locales)
	have "non_volatile_writes_unshared 𝒮sb []" 
	  by (simp add: sb)
	from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
	show "outstanding_non_volatile_writes_unshared 𝒮sb' tssb'"
	  by (simp add: tssb' sb sb' 𝒮sb')
      next
	have "sharing_consistent 𝒮sb 𝒪sb []" by simp
	from sharing_consis_nth_update [OF i_bound this]
	show "sharing_consis 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒪sb' sb' sb 𝒮sb')
      next
	from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound tssb_i] ]
	show "read_only_unowned 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' 𝒪sb')
      next
	from unowned_shared_nth_update [OF i_bound tssb_i subset_refl]
	show "unowned_shared 𝒮sb' tssb'" by (simp add: tssb' sb' sb 𝒪sb' 𝒮sb') 
      next
	from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound, of "[]"]
	show "no_outstanding_write_to_read_only_memory 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' sb' sb)
      qed

      have tmps_distinct': "tmps_distinct tssb'"
      proof (intro_locales)
	from load_tmps_distinct [OF i_bound tssb_i]
	have "distinct_load_tmps issb'" 
	  by (auto simp add: "issb" split: instr.splits)
	from load_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_distinct tssb'" by (simp add: tssb' sb' sb 𝒪sb' "issb")
      next
	from read_tmps_distinct [OF i_bound tssb_i]
	have "distinct_read_tmps []" by (simp add: tssb' sb' sb 𝒪sb')
	from read_tmps_distinct_nth_update [OF i_bound this]
	show "read_tmps_distinct tssb'" by (simp add: tssb' sb' sb 𝒪sb')
      next
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i] 
          load_tmps_distinct [OF i_bound tssb_i]
	have "load_tmps issb'  read_tmps [] = {}"
	  by (clarsimp)
	from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_read_tmps_distinct tssb'"  by (simp add: tssb' sb' sb 𝒪sb')
      qed

      have valid_sops': "valid_sops tssb'"
      proof -
	from valid_store_sops [OF i_bound tssb_i]
	obtain 
	  valid_store_sops': "sopstore_sops issb'. valid_sop sop"
	  by (auto simp add: "issb" tssb' sb' sb 𝒪sb')

	from valid_sops_nth_update [OF i_bound  _ valid_store_sops', where sb= "[]" ]
	show ?thesis by (auto simp add: tssb' sb' sb 𝒪sb')
      qed

      have valid_dd': "valid_data_dependency tssb'"
      proof -
	from data_dependency_consistent_instrs [OF i_bound tssb_i]
	obtain 
	  dd_is: "data_dependency_consistent_instrs (dom θsb') issb'"
	  by (auto simp add: "issb" θsb')
	from load_tmps_write_tmps_distinct [OF i_bound tssb_i] 
	have "load_tmps issb'  (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: tssb' sb' sb 𝒪sb')
      qed

      have load_tmps_fresh': "load_tmps_fresh tssb'"
      proof -
	from load_tmps_fresh [OF i_bound tssb_i] 
	have "load_tmps issb'  dom θsb = {}"
	  by (auto simp add: "issb")
	from load_tmps_fresh_nth_update [OF i_bound this]
	show ?thesis by (simp add: "issb" tssb' sb' sb θsb')
      qed


      from enough_flushs_nth_update [OF i_bound, where sb="[]" ]
      have enough_flushs': "enough_flushs tssb'"
	by (auto simp add: tssb' sb' sb)

      have valid_program_history': "valid_program_history tssb'"
      proof -	
	have causal': "causal_program_history issb' sb'"
	  by (simp add: "issb" sb sb')
	have "last_prog psb sb' = psb"
	  by (simp add: sb' sb)
	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb' sb')
      qed

      from is_sim have "is": "is = Fence # issb'"
	by (simp add: suspends sb "issb")
      with ts_i 
      have ts_i: "ts!i = (psb, Fence # issb', θsb,(), 𝒟, acquired True ?take_sb 𝒪sb, release ?take_sb (dom 𝒮sb) sb)"
	by (simp add: suspends sb)

      from direct_memop_step.Fence 
      have "(Fence # issb',
	    θsb, (),m,𝒟, acquired True ?take_sb 𝒪sb, release ?take_sb (dom 𝒮sb) sb, 𝒮)  
        (issb', θsb, (), m, False, acquired True ?take_sb 𝒪sb, Map.empty, 𝒮)".
      from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this] 
      have "(ts, m, 𝒮) d 
        (ts[i := (psb, issb', θsb, (), False, acquired True ?take_sb 𝒪sb,Map.empty)], m, 𝒮)".

      moreover

      have "(tssb',msb,𝒮sb')  (ts[i := (psb,issb', θsb,(), False,acquired True ?take_sb 𝒪sb,Map.empty)],m,𝒮)"
	apply (rule sim_config.intros) 
	apply    (simp add: tssb' sb' 𝒪sb' sb' 𝒮sb' m 
	  flush_all_until_volatile_nth_update_unused [OF i_bound tssb_i])
	using   share_all_until_volatile_write_Fence_commute 
	           [OF i_bound tssb_i [simplified issb sb]]
	apply  (clarsimp simp add: 𝒮 tssb' 𝒮sb' issb 𝒪sb' sb' θsb' sb' sb)
	using  leq
	apply  (simp add: tssb')
	using i_bound i_bound' ts_sim 
	apply (clarsimp simp add: Let_def nth_list_update 
	  tssb' sb' sb 𝒪sb' sb' 𝒮sb' 𝒟sb' ex_not  θsb'
	  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' msb' 𝒮sb' 
	by (auto simp del: fun_upd_apply)
    next	
      case (SBHRMWReadOnly cond t a D f ret A L R W)
      then obtain 
	"issb": "issb = RMW a t (D,f) cond ret A L R W # issb'" and
	cond: "¬ (cond (θsb(tmsb a)))" and
	𝒪sb': "𝒪sb'=𝒪sb" and
        sb': "sb'=Map.empty" and
	θsb': "θsb' = θsb(tmsb a)" and
	𝒟sb': "¬ 𝒟sb'" and
	sb: "sb=[]" and
	sb': "sb'=[]" and
	msb': "msb' = msb" and
	𝒮sb': "𝒮sb'=𝒮sb" 
	by auto

      from safe_RMW_common  [OF safe_memop_flush_sb [simplified issb]]
      obtain access_cond: "a  𝒪sb  a  dom 𝒮" and
      rels_cond: " j < length ts. ij  released (ts!j) a  Some False"
        by (auto simp add: 𝒮 sb)
	

      have valid_own': "valid_ownership 𝒮sb' tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	proof -
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb []"
	    by simp
	  from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
	  show ?thesis by (simp add: tssb' sb' sb 𝒪sb' 𝒮sb')
	qed
      next
	from outstanding_volatile_writes_unowned_by_others_store_buffer 
	[OF i_bound tssb_i subset_refl]
	show "outstanding_volatile_writes_unowned_by_others tssb'" 
	  by (simp add: tssb' sb' sb 𝒪sb' 𝒮sb')
      next
	from read_only_reads_unowned_nth_update [OF i_bound tssb_i, of "[]" 𝒪sb]
	show "read_only_reads_unowned tssb'"
	  by (simp add: tssb' sb' sb 𝒪sb')
      next
	from ownership_distinct_instructions_read_value_store_buffer_independent 
	[OF i_bound tssb_i]
	show "ownership_distinct tssb'"
	  by (simp add: tssb' sb' sb 𝒪sb')
      qed


      have valid_hist': "valid_history program_step tssb'"
      proof -
	from valid_history [OF i_bound tssb_i] 
	have "history_consistent (θsb(tmsb a)) (hd_prog psb []) []" by simp
	from valid_history_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' sb 𝒪sb' θsb')
      qed
      
      have valid_reads': "valid_reads msb tssb'"
      proof -
	have "reads_consistent False 𝒪sb msb []" by simp
	from valid_reads_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' sb 𝒪sb')
      qed


      have valid_sharing': "valid_sharing 𝒮sb' tssb'"
      proof (intro_locales)
	from outstanding_non_volatile_writes_unshared [OF i_bound tssb_i]
	have "non_volatile_writes_unshared 𝒮sb []" 
	  by (simp add: sb)
	from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
	show "outstanding_non_volatile_writes_unshared 𝒮sb' tssb'"
	  by (simp add: tssb' sb sb' 𝒮sb')
      next
	have "sharing_consistent 𝒮sb 𝒪sb []" by simp
	from sharing_consis_nth_update [OF i_bound this]
	show "sharing_consis 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒪sb' sb' sb 𝒮sb')
      next
	from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound tssb_i] ]
	show "read_only_unowned 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' 𝒪sb')
      next
	from unowned_shared_nth_update [OF i_bound tssb_i subset_refl]
	show "unowned_shared 𝒮sb' tssb'" by (simp add: tssb' sb' sb 𝒪sb' 𝒮sb')
      next 
	from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound, of "[]"]
	show "no_outstanding_write_to_read_only_memory 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' sb' sb)
      qed

      have tmps_distinct': "tmps_distinct tssb'"
      proof (intro_locales)
	from load_tmps_distinct [OF i_bound tssb_i]
	have "distinct_load_tmps issb'" 
	  by (auto simp add: "issb" split: instr.splits)
	from load_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_distinct tssb'" by (simp add: tssb' sb' sb 𝒪sb' "issb")
      next
	from read_tmps_distinct [OF i_bound tssb_i]
	have "distinct_read_tmps []" by (simp add: tssb' sb' sb 𝒪sb')
	from read_tmps_distinct_nth_update [OF i_bound this]
	show "read_tmps_distinct tssb'" by (simp add: tssb' sb' sb 𝒪sb')
      next
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i] 
          load_tmps_distinct [OF i_bound tssb_i]
	have "load_tmps issb'  read_tmps [] = {}"
	  by (clarsimp)
	from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_read_tmps_distinct tssb'"  by (simp add: tssb' sb' sb 𝒪sb')
      qed

      have valid_sops': "valid_sops tssb'"
      proof -
	from valid_store_sops [OF i_bound tssb_i]
	obtain 
	  valid_store_sops': "sopstore_sops issb'. valid_sop sop"
	  by (auto simp add: "issb" tssb' sb' sb 𝒪sb')

	from valid_sops_nth_update [OF i_bound  _ valid_store_sops', where sb= "[]" ]
	show ?thesis by (auto simp add: tssb' sb' sb 𝒪sb')
      qed

      have valid_dd': "valid_data_dependency tssb'"
      proof -
	from data_dependency_consistent_instrs [OF i_bound tssb_i]
	obtain 
	  dd_is: "data_dependency_consistent_instrs (dom θsb') issb'"
	  by (auto simp add: "issb" θsb')
	from load_tmps_write_tmps_distinct [OF i_bound tssb_i] 
	have "load_tmps issb'  (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: tssb' sb' sb 𝒪sb')
      qed


      have load_tmps_fresh': "load_tmps_fresh tssb'"
      proof -
	from load_tmps_fresh [OF i_bound tssb_i] 
	have "load_tmps (RMW a t (D,f) cond ret A L R W# issb')  dom θsb = {}"
	  by (simp add: "issb")
	moreover
	from load_tmps_distinct [OF i_bound tssb_i] have "t  load_tmps issb'"
	  by (auto simp add: "issb")
	ultimately have "load_tmps issb'  dom (θsb(t  msb a)) = {}"
	  by auto
	from load_tmps_fresh_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' θsb')
      qed

      from enough_flushs_nth_update [OF i_bound, where sb="[]" ]
      have enough_flushs': "enough_flushs tssb'"
	by (auto simp add: tssb' sb' sb)

      have valid_program_history': "valid_program_history tssb'"
      proof -	
	have causal': "causal_program_history issb' sb'"
	  by (simp add: "issb" sb sb')
	have "last_prog psb sb' = psb"
	  by (simp add: sb' sb)
	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb' sb')
      qed

      from is_sim have "is": "is = RMW a t (D,f) cond ret A L R W# issb'"
	by (simp add: suspends sb "issb")
      with ts_i 
      have ts_i: "ts!i = (psb, RMW a t (D,f) cond ret A L R W# issb', θsb,(), 
	𝒟, acquired True ?take_sb 𝒪sb, release ?take_sb (dom 𝒮sb) sb)"
	by (simp add: suspends sb)
      
	
      have "flush_all_until_volatile_write tssb msb a = msb a"
      proof -
        have "j < length tssb. i  j 
          (let (_,_,_,sbj,_,_,_) = tssb!j 
          in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))"
	proof -
	  {
	    fix j pj "isj" 𝒪j j 𝒟j xsj sbj
	    assume j_bound: "j < length tssb"
	    assume neq_i_j: "i  j"
	    assume jth: "tssb!j = (pj,isj, xsj, sbj, 𝒟j, 𝒪j,j)"
	    have "a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	    proof 
	      let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	      let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"
	      assume a_in: "a  outstanding_refs is_non_volatile_Writesb ?take_sbj"
	      with outstanding_refs_takeWhile [where P'= "Not  is_volatile_Writesb"]
	      have a_in': "a  outstanding_refs is_non_volatile_Writesb sbj"
		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 sbj"
		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_sbj (dom (𝒮sb)) j  a  Some False"
                by (auto simp add: Let_def)
	      from access_cond
	      show False
	      proof 
		assume "a  𝒪sb"
		with ownership_distinct [OF i_bound j_bound neq_i_j tssb_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_tssb sharing_consis_tssb j_bound jth j_owns]
                have a_dom: "a  dom  (share ?take_sbj 𝒮sb)"
                  by (auto simp add: 𝒮 domIff)
		from outstanding_non_volatile_writes_unshared [OF j_bound jth]
		have "non_volatile_writes_unshared 𝒮sb sbj".
		with non_volatile_writes_unshared_append [of 𝒮sb "(takeWhile (Not  is_volatile_Writesb) sbj)"
		  "(dropWhile (Not  is_volatile_Writesb) sbj)"]
		have unshared_take: "non_volatile_writes_unshared 𝒮sb (takeWhile (Not  is_volatile_Writesb) sbj)" 
		  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 tssb_i this]
	show ?thesis
	  by (simp add: sb)
      qed
      
      hence m_a: "m a = msb a"
	by (simp add: m)

      from cond have cond': "¬ cond (θsb(t  m a))"
	by (simp add: m_a)

      from direct_memop_step.RMWReadOnly [where cond=cond and θ=θsb and m=m, OF cond']
      have "(RMW a t (D, f) cond ret A L R W # issb',
             θsb, (),m, 𝒟, 𝒪sb, sb, 𝒮)  
            (issb', θsb(t  m a), (), m, False, 𝒪sb, Map.empty, 𝒮)".

      from direct_computation.concurrent_step.Memop [OF i_bound' ts_i [simplified sb, simplified] this]
      have "(ts, m, 𝒮) d (ts[i := (psb, issb',
	       θsb(t  m a), (), False, 𝒪sb,Map.empty)], m, 𝒮)".

      moreover
      
      have tmps_commute: "θsb(t  (msb a)) = 
	(θsb |` (dom θsb - {t}))(t  (msb a))"
	apply (rule ext)
	apply (auto simp add: restrict_map_def domIff)
	done

      have "(tssb',msb,𝒮sb')  (ts[i := (psb,issb', θsb(t  m a),(), False,𝒪sb,Map.empty)],m,𝒮)"
	apply (rule sim_config.intros)
	apply    (simp add: tssb' sb' 𝒪sb' sb' m 
	  flush_all_until_volatile_nth_update_unused [OF i_bound tssb_i, simplified sb])
	using  share_all_until_volatile_write_RMW_commute [OF i_bound tssb_i [simplified issb sb]]
	apply  (clarsimp simp add: 𝒮 tssb' 𝒮sb' issb 𝒪sb' θsb' sb' sb)
	using  leq
	apply  (simp add: tssb')
	using i_bound i_bound' ts_sim
	apply (clarsimp simp add: Let_def nth_list_update 
	  tssb' sb' sb 𝒪sb' sb' 𝒮sb' θsb' 𝒟sb' 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' msb' 𝒮sb' 
	by (auto simp del: fun_upd_apply)
    next
      case (SBHRMWWrite cond t a D f ret A L R W) 
      then obtain 
	"issb": "issb = RMW a t (D,f) cond ret A L R W # issb'" and
	cond: "(cond (θsb(tmsb a)))" and
	𝒪sb': "𝒪sb'=𝒪sb  A - R" and
        sb': "sb'=Map.empty" and
	𝒟sb': "¬ 𝒟sb'" and
	θsb': "θsb' = θsb(tret (msb a) (f (θsb(tmsb a))))" and
	sb: "sb=[]" and
	sb': "sb'=[]" and
	msb': "msb' = msb(a := f (θsb(tmsb a)))" and
	𝒮sb': "𝒮sb'=𝒮sb ⊕⇘WR ⊖⇘AL" 
	by auto

      from data_dependency_consistent_instrs [OF i_bound tssb_i]
      have D_subset: "D  dom θsb" 
	by (simp add: issb)

      from is_sim have "is": "is = RMW a t (D,f) cond ret A L R W # issb'"
	by (simp add: suspends sb "issb")
      with ts_i 
      have ts_i: "ts!i = (psb, RMW a t (D,f) cond ret A L R W # issb', θsb,(), 𝒟, 𝒪sb,sb)"
	by (simp add: suspends sb)
      
      from safe_RMW_common  [OF safe_memop_flush_sb [simplified issb]]
      obtain access_cond: "a  𝒪sb  a  dom 𝒮" and
      rels_cond: " j < length ts. ij  released (ts!j) a  Some False"
        by (auto simp add: 𝒮 sb)



      have a_unflushed: 
	"j < length tssb. i  j 
                  (let (_,_,_,sbj,_,_,_) = tssb!j 
                  in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))"
      proof -
	{
	  fix j pj "isj" 𝒪j j 𝒟j xsj sbj
	  assume j_bound: "j < length tssb"
	  assume neq_i_j: "i  j"
	  assume jth: "tssb!j = (pj,isj, xsj, sbj, 𝒟j, 𝒪j, j)"
	  have "a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	  proof 
	    let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	    let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"
	    assume a_in: "a  outstanding_refs is_non_volatile_Writesb ?take_sbj"
	    with outstanding_refs_takeWhile [where P'= "Not  is_volatile_Writesb"]
	    have a_in': "a  outstanding_refs is_non_volatile_Writesb sbj"
	      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 sbj"
	      by auto
	    with ownership_distinct [OF i_bound j_bound neq_i_j tssb_i jth]
	    have a_not_owns: "a  𝒪sb  all_acquired sb"
	      by blast
	    assume a_in: "a  outstanding_refs is_non_volatile_Writesb 
		(takeWhile (Not  is_volatile_Writesb) sbj)"
	    with outstanding_refs_takeWhile [where P'= "Not  is_volatile_Writesb"]
	    have a_in': "a  outstanding_refs is_non_volatile_Writesb sbj"
	      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_sbj (dom (𝒮sb)) j  a  Some False"
              by (auto simp add: Let_def)
	    from access_cond
	    show False
	    proof 
	      assume "a  𝒪sb"
	      with ownership_distinct [OF i_bound j_bound neq_i_j tssb_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_tssb sharing_consis_tssb j_bound jth j_owns]
              have a_dom: "a  dom  (share ?take_sbj 𝒮sb)"
                by (auto simp add: 𝒮 domIff)
	      from outstanding_non_volatile_writes_unshared [OF j_bound jth]
	      have "non_volatile_writes_unshared 𝒮sb sbj".
	      with non_volatile_writes_unshared_append [of 𝒮sb "(takeWhile (Not  is_volatile_Writesb) sbj)"
		"(dropWhile (Not  is_volatile_Writesb) sbj)"]
	      have unshared_take: "non_volatile_writes_unshared 𝒮sb (takeWhile (Not  is_volatile_Writesb) sbj)" 
	        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 tssb msb a = msb a"
      proof -
	from flush_all_until_volatile_write_buffered_val_conv 
	[OF _ i_bound tssb_i a_unflushed]
	show ?thesis
	  by (simp add: sb)
      qed     
      
      hence m_a: "m a = msb a"
	by (simp add: m)

      from cond have cond': "cond (θsb(t  m a))"
	by (simp add: m_a)


      from safe_memop_flush_sb [simplified issb] cond'
      obtain 
	L_subset: "L  A" and
	A_shared_owned: "A  dom 𝒮  𝒪sb" and
	R_owned: "R  𝒪sb" 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 tssb. i  j  
          (let (_,_,_,sbj,_,𝒪j,_) = tssb!j in 
	    a  acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j 
            a  all_shared (takeWhile (Not  is_volatile_Writesb) sbj))" 
  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 tssb. ij  (let (_,_,_,sbj,_,𝒪j,_) = tssb!j 
	  in A  (acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j 
                  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)) = {})" 
  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 𝒮sb"
      proof 
	assume a: "a  read_only (𝒮sb)"
	  from local.read_only_unowned_axioms have "read_only_unowned 𝒮sb tssb". 
        from in_read_only_share_all_until_volatile_write' [OF ownership_distinct_tssb sharing_consis_tssb
          read_only_unowned 𝒮sb tssb i_bound tssb_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 pj issbj 𝒪j j 𝒟sbj θj sbj
	assume j_bound: "j < length tssb"
	assume tssb_j: "tssb!j=(pj,issbj,θj,sbj,𝒟sbj,𝒪j,j)"
	assume neq_i_j: "ij"
	have "a  unforwarded_non_volatile_reads (dropWhile (Not  is_volatile_Writesb) sbj) {}"
	proof 
	  let ?take_sbj = "takeWhile (Not  is_volatile_Writesb) sbj"
	  let ?drop_sbj = "dropWhile (Not  is_volatile_Writesb) sbj"
	  assume a_in: "a   unforwarded_non_volatile_reads ?drop_sbj {}"

	  from a_unowned_others [rule_format, OF _ neq_i_j] tssb_j j_bound
	  obtain a_unacq_take: "a  acquired True ?take_sbj 𝒪j" and a_not_shared: "a  all_shared ?take_sbj"
	    by auto
(*
	  then obtain
	    a_unowned: "a ∉ 𝒪j" and a_unacq_take': "a ∉ all_acquired ?take_sbj"
	    by (auto simp add: acquired_takeWhile_non_volatile_Writesb)
*)
	  note nvo_j = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound tssb_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_sbj 𝒮sb)
	                    (acquired True ?take_sbj 𝒪j) ?drop_sbj" .

	  note consis_j = sharing_consis [OF j_bound tssb_j]
	  with sharing_consistent_append [of 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
	  obtain consis_take_j: "sharing_consistent 𝒮sb 𝒪j ?take_sbj" and
	    consis_drop_j: "sharing_consistent (share ?take_sbj 𝒮sb)
	      (acquired True ?take_sbj 𝒪j) ?drop_sbj"
	    by auto

	  from in_unforwarded_non_volatile_reads_non_volatile_Readsb [OF a_in]
	  have a_in': "a  outstanding_refs is_non_volatile_Readsb ?drop_sbj".

	  note reads_consis_j = valid_reads [OF j_bound tssb_j]
	  from reads_consistent_drop [OF this]
	  have reads_consis_drop_j:
	    "reads_consistent True (acquired True ?take_sbj 𝒪j) (flush ?take_sbj msb) ?drop_sbj".

        
          from read_only_share_all_shared [of a ?take_sbj 𝒮sb] a_not_ro' a_not_shared
          have a_not_ro_j: "a  read_only (share ?take_sbj 𝒮sb)"
            by auto
          



	  from ts_sim [rule_format, OF j_bound] tssb_j j_bound
	  obtain suspendsj "isj" 𝒟j where
	    suspendsj: "suspendsj = ?drop_sbj" and
	    isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	    𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
	    tsj: "ts!j = (hd_prog pj suspendsj, isj, 
	    θj |` (dom θj - read_tmps suspendsj),(),   
	    𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	    by (auto simp: Let_def)

	  from tsj neq_i_j j_bound 
	  have ts'_j: "?ts'!j = (hd_prog pj suspendsj, isj,
	    θj |` (dom θj - read_tmps suspendsj),(), 
	    𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	    by auto

	  from valid_last_prog [OF j_bound tssb_j] have last_prog: "last_prog pj sbj = pj".

	  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_sbj= ys @ Writesb True a' sop v A L R W # zs  a  A  
                a  outstanding_refs is_Writesb ys  a'a"
	    with suspendsj
	    obtain a' sop' v' ys zs' A' L' R' W' where
		split_suspendsj: "suspendsj = ys @ Writesb True a' sop' v' A' L' R' W'# zs'" (is "suspendsj=?suspends") and
		a_A': "a  A'" and
		no_write: "a  outstanding_refs is_Writesb (ys @ [Writesb True a' sop' v' A' L' R' W'])"
	        by(auto simp add: outstanding_refs_append )
		
	    from last_prog
	    have lp: "last_prog pj suspendsj = pj"
	      apply -
	      apply (rule last_prog_same_append [where sb="?take_sbj"])
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	    
	    from sharing_consis [OF j_bound tssb_j]
	    have sharing_consis_j: "sharing_consistent 𝒮sb 𝒪j sbj".
	    then have A'_R': "A'  R' = {}" 
	      by (simp add: sharing_consistent_append [of _ _ ?take_sbj ?drop_sbj, simplified] 
		  suspendsj [symmetric] split_suspendsj sharing_consistent_append)	  

	    from valid_program_history [OF j_bound tssb_j] 
	    have "causal_program_history issbj sbj".
	    then have cph: "causal_program_history issbj ?suspends"
	      apply -
	      apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply (simp add: split_suspendsj)
	      done
	    
	    from valid_reads [OF j_bound tssb_j]
	    have reads_consis_j: "reads_consistent False 𝒪j msb sbj".
	    
	   from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb 
	      j_bound tssb_j this]
	   have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
	    by (simp add: m suspendsj)
	    
	   from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j tssb_i tssb_j]
	   have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
	     by (simp add: suspendsj)
	   from reads_consistent_flush_independent [OF this reads_consis_m_j]
	   have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	        (flush ?drop_sb m) suspendsj".

	   hence reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  
	        (flush ?drop_sb m) (ys@[Writesb True a' sop' v' A' L' R' W'])"
	     by (simp add: split_suspendsj reads_consistent_append)

  	   from valid_write_sops [OF j_bound tssb_j]
	   have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
	     by (simp add: split_suspendsj [symmetric] suspendsj)
	   then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		valid_sops_drop: "sopwrite_sops (ys@[Writesb 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 tssb_j]
	   have "distinct_read_tmps (?take_sbj@suspendsj)"
	     by (simp add: split_suspendsj [symmetric] suspendsj)
	   then obtain 
	     read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
	     distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
	     apply (simp only: split_suspendsj [symmetric] suspendsj) 
	     apply (simp only: distinct_read_tmps_append)
	     done
	    
	   from valid_history [OF j_bound tssb_j]
	   have h_consis: 
	      "history_consistent θj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	    
	   have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	   proof -
	     from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
	       by simp
	     from last_prog_hd_prog_append' [OF h_consis] this
	     have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
	       by (simp only: split_suspendsj [symmetric] suspendsj) 
	     moreover 
	     have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
	       last_prog (hd_prog pj suspendsj) ?take_sbj"
	       apply (simp only: split_suspendsj [symmetric] suspendsj) 
	       by (rule last_prog_hd_prog_append)
	     ultimately show ?thesis
	       by (simp add: split_suspendsj [symmetric] suspendsj) 
	   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 pj suspendsj) suspendsj"
	     by (simp add: split_suspendsj [symmetric] suspendsj)
	   from reads_consistent_drop_volatile_writes_no_volatile_reads  
	   [OF reads_consis_j] 
	   have no_vol_read: "outstanding_refs is_volatile_Readsb 
	      (ys@[Writesb True a' sop' v' A' L' R' W']) = {}"
	     by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )
	    
	   have acq_simp:
	      "acquired True (ys @ [Writesb True a' sop' v' A' L' R' W']) 
              (acquired True ?take_sbj 𝒪j) = 
              acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R'"
	     by (simp add: acquired_append)

	   from flush_store_buffer_append [where sb="ys@[Writesb True a' sop' v' A' L' R' W']" and sb'="zs'", simplified,
	      OF j_bound_ts' isj [simplified split_suspendsj] cph [simplified suspendsj]
	      ts'_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys 
	      hist_consis' [simplified split_suspendsj] valid_sops_drop 
	      distinct_read_tmps_drop [simplified split_suspendsj] 
	      no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="share ?drop_sb 𝒮"]
	    
	   obtain isj' j' where
	      isj': "instrs zs' @ issbj = isj' @ prog_instrs zs'" and
	      steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮)  d* 
		  (?ts'[j:=(last_prog
                              (hd_prog pj (Writesb True a' sop' v' A' L' R' W'# zs')) (ys@[Writesb True a' sop' v' A' L' R' W']),
                             isj',
                             θj |` (dom θj - read_tmps zs'),
                              (), True, acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R',j')],
                    flush (ys@[Writesb True a' sop' v' A' L' R' W']) (flush ?drop_sb m),
                    share (ys@[Writesb 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 = (psb, issb, θsb,(), 
	      𝒟sb, acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb)"
	     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 @ [Writesb True a' sop' v' A' L' R' W']) m a = m a".
	    
	   with safe_delayedE [OF safe i_bound_ys ts_ys_i, simplified issb] cond'
	   have a_unowned: 
	      
	      "j < length ?ts_ys. ij  (let (𝒪j) = map owned ?ts_ys!j in a  𝒪j)"
	     apply cases
	     apply (auto simp add: Let_def issb 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_sbj = ys @ Ghostsb A L R W# zs  a  A  a  outstanding_refs is_Writesb ys"
	    with  suspendsj 
	    obtain ys zs' A' L' R' W' where
	      split_suspendsj: "suspendsj = ys @ Ghostsb A' L' R' W'# zs'" (is "suspendsj=?suspends") and
	      a_A': "a  A'" and
	      no_write: "a  outstanding_refs is_Writesb (ys @ [Ghostsb A' L' R' W'])"
	      by (auto simp add: outstanding_refs_append)

	    from last_prog
	    have lp: "last_prog pj suspendsj = pj"
	      apply -
	      apply (rule last_prog_same_append [where sb="?take_sbj"])
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	    from sharing_consis [OF j_bound tssb_j]
	    have sharing_consis_j: "sharing_consistent 𝒮sb 𝒪j sbj".
	    then have A'_R': "A'  R' = {}" 
	      by (simp add: sharing_consistent_append [of _ _ ?take_sbj ?drop_sbj, simplified] 
		  suspendsj [symmetric] split_suspendsj sharing_consistent_append)	  


	    from valid_program_history [OF j_bound tssb_j] 
	    have "causal_program_history issbj sbj".
	    then have cph: "causal_program_history issbj ?suspends"
	      apply -
	      apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply (simp add: split_suspendsj)
	      done
	    
	    from valid_reads [OF j_bound tssb_j]
	    have reads_consis_j: "reads_consistent False 𝒪j msb sbj".
	    
	    from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb 
	      j_bound tssb_j this]
	    have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
	      by (simp add: m suspendsj)
	    
	    from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound neq_i_j tssb_i tssb_j]
	    have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
	      by (simp add: suspendsj)
	    from reads_consistent_flush_independent [OF this reads_consis_m_j]
	    have reads_consis_flush_suspend: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	      (flush ?drop_sb m) suspendsj".

	    hence reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  
	      (flush ?drop_sb m) (ys@[Ghostsb A' L' R' W'])"
	      by (simp add: split_suspendsj reads_consistent_append)

	    from valid_write_sops [OF j_bound tssb_j]
	    have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
	      valid_sops_drop: "sopwrite_sops (ys@[Ghostsb A' L' R' W']). valid_sop sop"
	      apply (simp only: write_sops_append)
	      apply auto
	      done
	    
	    from read_tmps_distinct [OF j_bound tssb_j]
	    have "distinct_read_tmps (?take_sbj@suspendsj)"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain 
	      read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
	      distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
	      apply (simp only: split_suspendsj [symmetric] suspendsj) 
	      apply (simp only: distinct_read_tmps_append)
	      done
	    
	    from valid_history [OF j_bound tssb_j]
	    have h_consis: 
	      "history_consistent θj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	    
	    have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	    proof -
	      from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		by simp
	      from last_prog_hd_prog_append' [OF h_consis] this
	      have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		by (simp only: split_suspendsj [symmetric] suspendsj) 
	      moreover 
	      have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		last_prog (hd_prog pj suspendsj) ?take_sbj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		by (rule last_prog_hd_prog_append)
	      ultimately show ?thesis
		by (simp add: split_suspendsj [symmetric] suspendsj) 
	    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 pj suspendsj) suspendsj"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    from reads_consistent_drop_volatile_writes_no_volatile_reads  
	    [OF reads_consis_j] 
	    have no_vol_read: "outstanding_refs is_volatile_Readsb 
	      (ys@[Ghostsb A' L' R' W']) = {}"
	      by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )
	    
	    have acq_simp:
	      "acquired True (ys @ [Ghostsb A' L' R' W']) 
              (acquired True ?take_sbj 𝒪j) = 
              acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R'"
	      by (simp add: acquired_append)

	    from flush_store_buffer_append [where sb="ys@[Ghostsb A' L' R' W']" and sb'="zs'", simplified,
	      OF j_bound_ts' isj [simplified split_suspendsj] cph [simplified suspendsj]
	      ts'_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys 
	      hist_consis' [simplified split_suspendsj] valid_sops_drop 
	      distinct_read_tmps_drop [simplified split_suspendsj] 
	      no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="share ?drop_sb 𝒮"]
	    
	    obtain isj' j' where
	      isj': "instrs zs' @ issbj = isj' @ prog_instrs zs'" and
	      steps_ys: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮)  d* 
		  (?ts'[j:=(last_prog
                              (hd_prog pj (Ghostsb A' L' R' W'# zs')) (ys@[Ghostsb A' L' R' W']),
                             isj',
                             θj |` (dom θj - read_tmps zs'),
                              (), 
	                     𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R',j')],
                    flush (ys@[Ghostsb A' L' R' W']) (flush ?drop_sb m),
                    share (ys@[Ghostsb 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 = (psb, issb, θsb,(), 
	      𝒟sb, acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb)"
	      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 @ [Ghostsb A' L' R' W']) m a = m a".
	    
	    with safe_delayedE [OF safe i_bound_ys ts_ys_i, simplified issb] cond'
	    have a_unowned: 
	      
	      "j < length ?ts_ys. ij  (let (𝒪j) = map owned ?ts_ys!j in a  𝒪j)"
	      apply cases
	      apply (auto simp add: Let_def issb 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

      (* FIXME: split in to theorems, one for A ∩ 𝒪j and  one for
	 A ∩ outstanding_refs…    *) 
      have A_unused_by_others:
	  "j<length (map 𝒪_sb tssb). i  j 
             (let (𝒪j, sbj) = map 𝒪_sb tssb! j
             in A   (𝒪j  outstanding_refs is_volatile_Writesb sbj) = {})"
      proof -
	{
	  fix j 𝒪j sbj
	  assume j_bound: "j < length (map owned tssb)"
	  assume neq_i_j: "ij"
	  assume tssb_j: "(map 𝒪_sb tssb)!j = (𝒪j,sbj)"
	  assume conflict: "A  (𝒪j  outstanding_refs is_volatile_Writesb sbj)  {}"
	  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 tssb"
	      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_Writesb sbj"
	      by auto
            from A_unowned_by_others [rule_format, OF _ neq_i_j] j_bound  tssb_j
            have A_unshared_j: "A  all_shared (takeWhile (Not  is_volatile_Writesb) sbj) = {}"
              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_Writesb) sbj) 𝒪j"
	        by (auto)
              with A_unowned_by_others [rule_format, OF _ neq_i_j] j_bound  tssb_j a_in
              show False by auto
	    next
	      assume a_in_j: "a'  outstanding_refs is_volatile_Writesb sbj"

	      let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	      let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"

	      from ts_sim [rule_format, OF  j_bound''] tssb_j j_bound''
	      obtain pj suspendsj "issbj" 𝒟sbj 𝒟j j θsbj "isj" where
		  tssb_j: "tssb ! j = (pj,issbj, θsbj, sbj,𝒟sbj,𝒪j,j)" and
		  suspendsj: "suspendsj = ?drop_sbj" and
		  𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
		  isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
		  tsj: "ts!j = (hd_prog pj suspendsj, isj,
		       θsbj |` (dom θsbj - read_tmps suspendsj),(), 𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
		apply (cases "tssb!j")
		apply (force simp add: Let_def)
		done
	      


	      have "a'  outstanding_refs is_volatile_Writesb suspendsj"
	      proof -	
		from a_in_j 
		have "a'  outstanding_refs is_volatile_Writesb (?take_sbj @ ?drop_sbj)"
		  by simp
		thus ?thesis
		  apply (simp only: outstanding_refs_append suspendsj)
		  apply (auto simp add: outstanding_refs_conv dest: set_takeWhileD)
		  done
	      qed
		 
	      from split_volatile_Writesb_in_outstanding_refs [OF this]
	      obtain sop' v' ys zs A' L' R' W' where
		split_suspendsj: "suspendsj = ys @ Writesb True a' sop' v' A' L' R' W'# zs" (is "suspendsj = ?suspends")
		by blast
	      


	      from valid_program_history [OF j_bound'' tssb_j] 
	      have "causal_program_history issbj sbj".
	      then have cph: "causal_program_history issbj ?suspends"
		apply -
		apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply (simp add: split_suspendsj)
		done

	      from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
	      then
	      have lp: "last_prog pj ?suspends = pj"
		apply -
		apply (rule last_prog_same_append [where sb="?take_sbj"])
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done

	      from valid_reads [OF j_bound'' tssb_j]
	      have reads_consis: "reads_consistent False 𝒪j msb sbj".
	      from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb 
		j_bound''
		tssb_j this]
	      have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		by (simp add: m suspendsj)

	      from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' tssb_j]
	      have nvo_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j sbj".
	      with non_volatile_owned_or_read_only_append [of False 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
	      have nvo_take_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j ?take_sbj"
		by auto

	      from a_unowned_others [rule_format, OF _ neq_i_j] tssb_j j_bound
	      have a_not_acq: "a  acquired True ?take_sbj 𝒪j"
		by auto

	      from a_notin_unforwarded_non_volatile_reads_drop[OF j_bound'' tssb_j neq_i_j]
	      have a_notin_unforwarded_reads: "a  unforwarded_non_volatile_reads suspendsj {}"
		by (simp add: suspendsj)
		
	      let ?ma="(m(a := f (θsb(tm 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_sbj 𝒪j) ?ma suspendsj"
		by auto

	      from reads_consis_ma_j 
	      have reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j) ?ma (ys)"
		by (simp add: split_suspendsj reads_consistent_append)

	      from direct_memop_step.RMWWrite [where cond=cond and θ=θsb and m=m, OF cond' ]
	      have "(RMW a t (D, f) cond ret A L R W# issb',  θsb, (), m,𝒟, 𝒪sb, sb, 𝒮)  
                    (issb', θsb(t  ret (m a) (f (θsb(t  m a)))), (), ?ma, False, 𝒪sb  A - R, Map.empty,𝒮 ⊕⇘WR ⊖⇘AL)".
	      from direct_computation.concurrent_step.Memop [OF i_bound' ts_i  this] 
	      have step_a: "(ts, m, 𝒮) d 
                    (ts[i := (psb, issb', θsb(t  ret (m a) (f (θsb(t  m a)))), (), False, 𝒪sb  A - R,Map.empty)], 
                      ?ma,𝒮 ⊕⇘WR ⊖⇘AL)"
		(is " _ d (?ts_a, _, ?shared_a)").

	      from tsj neq_i_j j_bound 

	      have ts_a_j: "?ts_a!j = (hd_prog pj suspendsj, isj,
		θsbj |` (dom θsbj - read_tmps suspendsj),(), 𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom (𝒮sb)) j)"
		by auto


	      from valid_write_sops [OF j_bound'' tssb_j]
	      have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		valid_sops_drop: "sopwrite_sops (ys). valid_sop sop"
		apply (simp only: write_sops_append)
		apply auto
		done

	      from read_tmps_distinct [OF j_bound'' tssb_j]
	      have "distinct_read_tmps (?take_sbj@suspendsj)"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain 
		read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
	      distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		apply (simp only: distinct_read_tmps_append)
		done
	    
	      from valid_history [OF j_bound'' tssb_j]
	      have h_consis: 
		"history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done
	    
	      have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	      proof -
		from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		  by simp
	      from last_prog_hd_prog_append' [OF h_consis] this
	      have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		by (simp only: split_suspendsj [symmetric] suspendsj) 
	      moreover 
	      have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		last_prog (hd_prog pj suspendsj) ?take_sbj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		by (rule last_prog_hd_prog_append)
	      ultimately show ?thesis
		by (simp add: split_suspendsj [symmetric] suspendsj) 
	    qed

	    from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop 
	      h_consis] last_prog_hd_prog
	    have hist_consis': "history_consistent θsbj (hd_prog pj suspendsj) suspendsj"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    from reads_consistent_drop_volatile_writes_no_volatile_reads  
	    [OF reads_consis] 
	    have no_vol_read: "outstanding_refs is_volatile_Readsb (ys) = {}"
	      by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )
	    from j_bound' have j_bound_ts_a: "j < length ?ts_a" by auto

	    
	    from flush_store_buffer_append [where sb="ys" and sb'="Writesb True a' sop' v' A' L' R' W'#zs", simplified,
	    OF j_bound_ts_a isj [simplified split_suspendsj] cph [simplified suspendsj]
ts_a_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys
 	      hist_consis' [simplified split_suspendsj] valid_sops_drop 
	      distinct_read_tmps_drop [simplified split_suspendsj] 
              no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="?shared_a"]

	    obtain isj' j' where
	      isj': "Write True a' sop' A' L' R' W'# instrs zs @ issbj = isj' @ prog_instrs zs" and
	      steps_ys: "(?ts_a, ?ma, ?shared_a)  d* 
		(?ts_a[j:=(last_prog
                            (hd_prog pj zs) ys,
              isj',
                           θsbj |` (dom θsbj - read_tmps zs),
                            (), 𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Writesb True a' sop' v' A' L' R' W']) @ zs)"
	      by simp
	    from causal_program_history_suffix [OF this]
	    have cph': "causal_program_history issbj zs".	      
	    interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

	    from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
	    obtain isj''
	      where isj': "isj' = Write True a' sop' A' L' R' W'#isj''" and
	      isj'': "instrs zs @ issbj = isj'' @ 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 = (psb, issb', 
	      θsb(t  ret (m a) (f (θsb(t  m a)))),(), False, 𝒪sb  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 pj zs) ys, Write True a' sop' A' L' R' W'#isj'', θsbj |` (dom θsbj - read_tmps zs), (), 𝒟j  outstanding_refs is_volatile_Writesb ys  {}, 
	      acquired True ys (acquired True ?take_sbj 𝒪j),j')"
	      by (clarsimp simp add: isj')
	    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. ji  (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 tssb). i  j 
             (let (𝒪j, sbj) = map 𝒪_sb tssb! j
             in A  all_acquired sbj = {})"
      proof -
	{
	  fix j 𝒪j sbj
	  assume j_bound: "j < length (map owned tssb)"
	  assume neq_i_j: "ij"
	  assume tssb_j: "(map 𝒪_sb tssb)!j = (𝒪j,sbj)"
	  assume conflict: "A  all_acquired sbj  {}"
	  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 tssb"
	      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 sbj"
	      by auto

	    let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	    let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"

	    from ts_sim [rule_format, OF  j_bound''] tssb_j j_bound''
	    obtain pj suspendsj "issbj" θsbj 𝒟sbj j 𝒟j "isj" where
	      tssb_j: "tssb ! j = (pj,issbj,θsbj, sbj,𝒟sbj,𝒪j,j)" and
	      suspendsj: "suspendsj = ?drop_sbj" and
	      isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	      𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
	      tsj: "ts!j = (hd_prog pj suspendsj, isj, 
	             θsbj |` (dom θsbj - read_tmps suspendsj),(),   
	             𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	      apply (cases "tssb!j")
	      apply (force simp add: Let_def)
	      done
	      

	    from a'_in_j all_acquired_append [of ?take_sbj ?drop_sbj]
	    have "a'  all_acquired ?take_sbj  a'  all_acquired suspendsj"
	      by (auto simp add: suspendsj)
	    thus False
	    proof 
	      assume "a'  all_acquired ?take_sbj"
	      with A_unowned_by_others [rule_format, OF _ neq_i_j] tssb_j j_bound a'_in
	      show False
		by (auto dest: all_acquired_unshared_acquired)
	    next
	      assume conflict_drop: "a'  all_acquired suspendsj"
	      
	      from split_all_acquired_in [OF conflict_drop]
	      show ?thesis
	      proof
		assume "sop a'' v ys zs A L R W. 
                  suspendsj = ys @ Writesb True a'' sop v A L R W# zs  a'  A"
		then 
		obtain a'' sop' v' ys zs A' L' R' W' where
		  split_suspendsj: "suspendsj = ys @ Writesb True a'' sop' v' A' L' R' W'# zs" (is "suspendsj = ?suspends") and
		  a'_A': "a'  A'"
		  by blast
	    

		from valid_program_history [OF j_bound'' tssb_j] 
		have "causal_program_history issbj sbj".
		then have cph: "causal_program_history issbj ?suspends"
		  apply -
		  apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		  apply (simp only: split_suspendsj [symmetric] suspendsj)
		  apply (simp add: split_suspendsj)
		  done
		
		from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
		then
		have lp: "last_prog pj ?suspends = pj"
		  apply -
		  apply (rule last_prog_same_append [where sb="?take_sbj"])
		  apply (simp only: split_suspendsj [symmetric] suspendsj)
		  apply simp
		  done
		
		from valid_reads [OF j_bound'' tssb_j]
		have reads_consis: "reads_consistent False 𝒪j msb sbj".
		from reads_consistent_flush_all_until_volatile_write [OF 
		  valid_ownership_and_sharing 𝒮sb tssb  j_bound''
		  tssb_j this]
		have reads_consis_m_j: 
		  "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		  by (simp add: m suspendsj)
		
		from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' tssb_j]
		have nvo_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j sbj".
		with non_volatile_owned_or_read_only_append [of False 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
		have nvo_take_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j ?take_sbj"
		  by auto
		
		from a_unowned_others [rule_format, OF _ neq_i_j] tssb_j j_bound
		have a_not_acq: "a  acquired True ?take_sbj 𝒪j"
		  by auto
	      
		from a_notin_unforwarded_non_volatile_reads_drop[OF j_bound'' tssb_j neq_i_j]
		have a_notin_unforwarded_reads: "a  unforwarded_non_volatile_reads suspendsj {}"
		  by (simp add: suspendsj)    

		let ?ma="(m(a := f (θsb(tm 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_sbj 𝒪j) ?ma suspendsj"
		  by auto


		from reads_consis_ma_j 
		have reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j) ?ma (ys)"
		  by (simp add: split_suspendsj reads_consistent_append)

	    
		from direct_memop_step.RMWWrite [where cond=cond and θ=θsb and m=m, OF cond']
		have "(RMW a t (D, f) cond ret A L R W# issb',
		        θsb, (), m, 𝒟, 𝒪sb, sb, 𝒮)  
                    (issb', 
		       θsb(t  ret (m a) (f (θsb(t  m a)))), (), ?ma, False, 𝒪sb  A - R,Map.empty, 𝒮 ⊕⇘WR ⊖⇘AL)".
		from direct_computation.concurrent_step.Memop [OF i_bound' ts_i [simplified sb, simplified] this]
		have step_a: "(ts, m, 𝒮) d 
                    (ts[i := (psb, issb', θsb(t  ret (m a) (f (θsb(t  m a)))), (), False, 𝒪sb  A - R,Map.empty)], 
                       ?ma,𝒮 ⊕⇘WR ⊖⇘AL)"
		  (is " _ d (?ts_a, _, ?shared_a)").
	      

		from tsj neq_i_j j_bound 

		have ts_a_j: "?ts_a!j = (hd_prog pj suspendsj, isj,
		  θsbj |` (dom θsbj - read_tmps suspendsj),(), 
		  𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
		  by auto
	    
		
		from valid_write_sops [OF j_bound'' tssb_j]
		have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		  by (simp add: split_suspendsj [symmetric] suspendsj)
		then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		  valid_sops_drop: "sopwrite_sops (ys). valid_sop sop"
		  apply (simp only: write_sops_append)
		  apply auto
		  done
	    
		from read_tmps_distinct [OF j_bound'' tssb_j]
		have "distinct_read_tmps (?take_sbj@suspendsj)"
		  by (simp add: split_suspendsj [symmetric] suspendsj)
		then obtain 
		  read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
		  distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		  apply (simp only: split_suspendsj [symmetric] suspendsj) 
		  apply (simp only: distinct_read_tmps_append)
		  done
	    
		from valid_history [OF j_bound'' tssb_j]
		have h_consis: 
		  "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		  apply (simp only: split_suspendsj [symmetric] suspendsj)
		  apply simp
		  done
	    
		have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
		proof -
		  from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		    by simp
		  from last_prog_hd_prog_append' [OF h_consis] this
		  have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		    by (simp only: split_suspendsj [symmetric] suspendsj) 
		  moreover 
		  have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		    last_prog (hd_prog pj suspendsj) ?take_sbj"
		    apply (simp only: split_suspendsj [symmetric] suspendsj) 
		    by (rule last_prog_hd_prog_append)
		  ultimately show ?thesis
		    by (simp add: split_suspendsj [symmetric] suspendsj) 
		qed
	    
		from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop 
		  h_consis] last_prog_hd_prog
		have hist_consis': "history_consistent θsbj (hd_prog pj suspendsj) suspendsj"
		  by (simp add: split_suspendsj [symmetric] suspendsj)
		from reads_consistent_drop_volatile_writes_no_volatile_reads  
		[OF reads_consis] 
		have no_vol_read: "outstanding_refs is_volatile_Readsb (ys) = {}"
		  by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		    split_suspendsj )
		from j_bound' have j_bound_ts_a: "j < length ?ts_a" by auto
	    
		from flush_store_buffer_append [where sb="ys" and sb'="Writesb True a'' sop' v' A' L' R' W'#zs", simplified,
		  OF j_bound_ts_a isj [simplified split_suspendsj] cph [simplified suspendsj]
		  ts_a_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys
 		  hist_consis' [simplified split_suspendsj] valid_sops_drop 
		  distinct_read_tmps_drop [simplified split_suspendsj] 
		  no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
		  𝒮="?shared_a"]
	    
		obtain isj' j' where
		  isj': "Write True a'' sop' A' L' R' W'# instrs zs @ issbj = isj' @ prog_instrs zs" and
		  steps_ys: "(?ts_a, ?ma, ?shared_a)  d* 
		  (?ts_a[j:=(last_prog
                            (hd_prog pj zs) ys,
		             isj',
                             θsbj |` (dom θsbj - read_tmps zs),
                             (),
		             𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Writesb True a'' sop' v' A' L' R' W']) @ zs)"
		  by simp
		from causal_program_history_suffix [OF this]
		have cph': "causal_program_history issbj zs".	      
		interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

		from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
		obtain isj''
		  where isj': "isj' = Write True a'' sop' A' L' R' W'#isj''" and
		  isj'': "instrs zs @ issbj = isj'' @ 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 = (psb, issb',
		  θsb(t  ret (m a) (f (θsb(t  m a)))),(), False, 𝒪sb  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 pj zs) ys, Write True a'' sop' A' L' R' W'#isj'',
		  θsbj |` (dom θsbj - read_tmps zs), (), 
		  𝒟j  outstanding_refs is_volatile_Writesb ys  {}, 
		  acquired True ys (acquired True ?take_sbj 𝒪j),j')"
		  by (clarsimp simp add: isj')
		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. ji  (let (𝒪i) = map owned ?ts_ys!i in A'   𝒪i = {})"
		  apply cases
		  apply (fastforce simp add: Let_def issb)+
		  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. suspendsj = ys @ Ghostsb A L R W# zs  a'  A"
		then 
		obtain ys zs A' L' R' W' where
		  split_suspendsj: "suspendsj = ys @ Ghostsb A' L' R' W'# zs" (is "suspendsj = ?suspends") and
		  a'_A': "a'  A'"
		  by blast
	    

		from valid_program_history [OF j_bound'' tssb_j] 
		have "causal_program_history issbj sbj".
		then have cph: "causal_program_history issbj ?suspends"
		  apply -
		  apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		  apply (simp only: split_suspendsj [symmetric] suspendsj)
		  apply (simp add: split_suspendsj)
		  done
		
		from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
		then
		have lp: "last_prog pj ?suspends = pj"
		  apply -
		  apply (rule last_prog_same_append [where sb="?take_sbj"])
		  apply (simp only: split_suspendsj [symmetric] suspendsj)
		  apply simp
		  done
		
		
		from valid_reads [OF j_bound'' tssb_j]
		have reads_consis: "reads_consistent False 𝒪j msb sbj".
		from reads_consistent_flush_all_until_volatile_write [OF 
		  valid_ownership_and_sharing 𝒮sb tssb  j_bound''
		  tssb_j this]
		have reads_consis_m_j: 
		  "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		  by (simp add: m suspendsj)
		
		from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound'' tssb_j]
		have nvo_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j sbj".
		with non_volatile_owned_or_read_only_append [of False 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
		have nvo_take_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j ?take_sbj"
		  by auto
		
		from a_unowned_others [rule_format, OF _ neq_i_j] tssb_j j_bound
		have a_not_acq: "a  acquired True ?take_sbj 𝒪j"
		  by auto
	      
		from a_notin_unforwarded_non_volatile_reads_drop[OF j_bound'' tssb_j neq_i_j]
		have a_notin_unforwarded_reads: "a  unforwarded_non_volatile_reads suspendsj {}"
		  by (simp add: suspendsj)    

		let ?ma="(m(a := f (θsb(tm 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_sbj 𝒪j) ?ma suspendsj"
		  by auto


		from reads_consis_ma_j 
		have reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j) ?ma (ys)"
		  by (simp add: split_suspendsj reads_consistent_append)
	    
		from direct_memop_step.RMWWrite [where cond=cond and θ=θsb and m=m, OF cond']
		have "(RMW a t (D, f) cond ret A L R W# issb', 
		        θsb, (), m, 𝒟,𝒪sb,  sb, 𝒮)  
                    (issb', 
                        θsb(t  ret (m a) (f (θsb(t  m a)))), (), ?ma, False, 𝒪sb  A - R,Map.empty,𝒮 ⊕⇘WR ⊖⇘AL)".
		from direct_computation.concurrent_step.Memop [OF i_bound' ts_i [simplified sb, simplified] this]
		have step_a: "(ts, m, 𝒮) d 
                    (ts[i := (psb, issb', θsb(t  ret (m a) (f (θsb(t  m a)))), (), False, 𝒪sb  A - R,Map.empty)], 
                      ?ma,𝒮 ⊕⇘WR ⊖⇘AL)"
		  (is " _ d (?ts_a, _, ?shared_a)").
	      

		from tsj neq_i_j j_bound 

		have ts_a_j: "?ts_a!j = (hd_prog pj suspendsj, isj,
		  θsbj |` (dom θsbj - read_tmps suspendsj),(), 𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
		  by auto
	    
		
		from valid_write_sops [OF j_bound'' tssb_j]
		have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		  by (simp add: split_suspendsj [symmetric] suspendsj)
		then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		  valid_sops_drop: "sopwrite_sops (ys). valid_sop sop"
		  apply (simp only: write_sops_append)
		  apply auto
		  done
	    
		from read_tmps_distinct [OF j_bound'' tssb_j]
		have "distinct_read_tmps (?take_sbj@suspendsj)"
		  by (simp add: split_suspendsj [symmetric] suspendsj)
		then obtain 
		  read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
		  distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		  apply (simp only: split_suspendsj [symmetric] suspendsj) 
		  apply (simp only: distinct_read_tmps_append)
		  done
	    
		from valid_history [OF j_bound'' tssb_j]
		have h_consis: 
		  "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		  apply (simp only: split_suspendsj [symmetric] suspendsj)
		  apply simp
		  done
	    
		have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
		proof -
		  from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		    by simp
		  from last_prog_hd_prog_append' [OF h_consis] this
		  have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		    by (simp only: split_suspendsj [symmetric] suspendsj) 
		  moreover 
		  have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		    last_prog (hd_prog pj suspendsj) ?take_sbj"
		    apply (simp only: split_suspendsj [symmetric] suspendsj) 
		    by (rule last_prog_hd_prog_append)
		  ultimately show ?thesis
		    by (simp add: split_suspendsj [symmetric] suspendsj) 
		qed
	    
		from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop 
		  h_consis] last_prog_hd_prog
		have hist_consis': "history_consistent θsbj (hd_prog pj suspendsj) suspendsj"
		  by (simp add: split_suspendsj [symmetric] suspendsj)
		from reads_consistent_drop_volatile_writes_no_volatile_reads  
		[OF reads_consis] 
		have no_vol_read: "outstanding_refs is_volatile_Readsb (ys) = {}"
		  by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		    split_suspendsj )
		from j_bound' have j_bound_ts_a: "j < length ?ts_a" by auto
	    
		from flush_store_buffer_append [where sb="ys" and sb'="Ghostsb A' L' R' W'#zs", simplified,
		  OF j_bound_ts_a isj [simplified split_suspendsj] cph [simplified suspendsj]
		  ts_a_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys
 		  hist_consis' [simplified split_suspendsj] valid_sops_drop 
		  distinct_read_tmps_drop [simplified split_suspendsj] 
		  no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
		  𝒮="?shared_a"]
	    
		obtain isj' j' where
		  isj': "Ghost A' L' R' W'# instrs zs @ issbj = isj' @ prog_instrs zs" and
		  steps_ys: "(?ts_a, ?ma, ?shared_a)  d* 
		  (?ts_a[j:=(last_prog
                            (hd_prog pj zs) ys,
		             isj',
                             θsbj |` (dom θsbj - read_tmps zs),
                             (),
		             𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Ghostsb A' L' R' W']) @ zs)"
		 by simp
	       from causal_program_history_suffix [OF this]
	       have cph': "causal_program_history issbj zs".	      
	       interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

	       from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
	       obtain isj''
		 where isj': "isj' = Ghost A' L' R' W'#isj''" and
		 isj'': "instrs zs @ issbj = isj'' @ 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 = (psb, issb',
		 θsb(t  ret (m a) (f (θsb(t  m a)))),(), False, 𝒪sb  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 pj zs) ys, Ghost A' L' R' W'#isj'',
		 θsbj |` (dom θsbj - read_tmps zs), (), 
		 𝒟j  outstanding_refs is_volatile_Writesb ys  {}, 
		 acquired True ys (acquired True ?take_sbj 𝒪j),j')"
		 by (clarsimp simp add: isj')
	       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. ji  (let (𝒪i) = map owned ?ts_ys!i in A'   𝒪i = {})"
		 apply cases
		 apply (fastforce simp add: Let_def issb)+
		 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 pj issbj 𝒪j j 𝒟sbj θj sbj
	assume j_bound: "j < length tssb"
	assume tssb_j: "tssb!j=(pj,issbj,θj,sbj,𝒟sbj,𝒪j,j)"
	assume neq_i_j: "ij"
	have "A  read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j) 
	           (dropWhile (Not  is_volatile_Writesb) sbj) = {}"
	proof -
	  {
	    let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	    let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"
	    
	    assume conflict: "A  read_only_reads (acquired True ?take_sbj 𝒪j) ?drop_sbj  {}"
	    have False
	    proof -
	      from conflict obtain a' where
		a'_in: "a'  A" and
		a'_in_j: "a'  read_only_reads (acquired True ?take_sbj 𝒪j) ?drop_sbj"
		by auto
	      
	      
	      from ts_sim [rule_format, OF  j_bound] tssb_j j_bound
	      obtain pj suspendsj "issbj" 𝒟sbj 𝒟j θsbj "isj" where
		tssb_j: "tssb ! j = (pj,issbj, θsbj, sbj,𝒟sbj,𝒪j,j)" and
		suspendsj: "suspendsj = ?drop_sbj" and
		isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
		𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
		tsj: "ts!j = (hd_prog pj suspendsj, isj,
	        θsbj |` (dom θsbj - read_tmps suspendsj),(), 𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
		apply (cases "tssb!j")
		apply (clarsimp simp add: Let_def)
		done
	      from split_in_read_only_reads [OF a'_in_j [simplified suspendsj [symmetric]]]
	      obtain t' v' ys zs where
		split_suspendsj: "suspendsj = ys @ Readsb False a' t' v'# zs" (is "suspendsj = ?suspends") and
		a'_unacq: "a'  acquired True ys (acquired True ?take_sbj 𝒪j)"
		by blast

	      from valid_program_history [OF j_bound tssb_j] 
	      have "causal_program_history issbj sbj".
	      then have cph: "causal_program_history issbj ?suspends"
		apply -
		apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply (simp add: split_suspendsj)
		done

	      from valid_last_prog [OF j_bound tssb_j] have last_prog: "last_prog pj sbj = pj".
	      then
	      have lp: "last_prog pj ?suspends = pj"
		apply -
		apply (rule last_prog_same_append [where sb="?take_sbj"])
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done

	      from valid_reads [OF j_bound tssb_j]
	      have reads_consis: "reads_consistent False 𝒪j msb sbj".
	      from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb 
		j_bound
		tssb_j this]
	      have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		by (simp add: m suspendsj)

	      from outstanding_non_volatile_refs_owned_or_read_only [OF j_bound tssb_j]
	      have nvo_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j sbj".
	      with non_volatile_owned_or_read_only_append [of False 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
	      have nvo_take_j: "non_volatile_owned_or_read_only False 𝒮sb 𝒪j ?take_sbj"
		by auto

	      from a_unowned_others [rule_format, OF _ neq_i_j] tssb_j j_bound
	      have a_not_acq: "a  acquired True ?take_sbj 𝒪j"
		by auto

	      from a_notin_unforwarded_non_volatile_reads_drop[OF j_bound tssb_j neq_i_j]
	      have a_notin_unforwarded_reads: "a  unforwarded_non_volatile_reads suspendsj {}"
		by (simp add: suspendsj)
	      
	      let ?ma="(m(a := f (θsb(tm 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_sbj 𝒪j) ?ma suspendsj"
		by auto

	      from reads_consis_ma_j 
	      have reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j) ?ma (ys)"
		by (simp add: split_suspendsj reads_consistent_append)

	      from direct_memop_step.RMWWrite [where cond=cond and θ=θsb and m=m, OF cond' ]
	      have "(RMW a t (D, f) cond ret A L R W# issb', θsb, (), m, 𝒟,𝒪sb,sb,𝒮)  
                (issb', θsb(t  ret (m a) (f (θsb(t  m a)))), (), ?ma, False, 𝒪sb  A - R,Map.empty, 𝒮 ⊕⇘WR ⊖⇘AL)".
	      from direct_computation.concurrent_step.Memop [OF i_bound' ts_i  this] 
	      have step_a: "(ts, m, 𝒮) d 
                (ts[i := (psb, issb', θsb(t  ret (m a) (f (θsb(t  m a)))), (), False, 𝒪sb  A - R,Map.empty)], 
                ?ma,𝒮 ⊕⇘WR ⊖⇘AL)"
		(is " _ d (?ts_a, _, ?shared_a)").

	      from tsj neq_i_j j_bound 

	      have ts_a_j: "?ts_a!j = (hd_prog pj suspendsj, isj,
		θsbj |` (dom θsbj - read_tmps suspendsj),(), 𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
		by auto


	      from valid_write_sops [OF j_bound tssb_j]
	      have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		valid_sops_drop: "sopwrite_sops (ys). valid_sop sop"
		apply (simp only: write_sops_append)
		apply auto
		done

	      from read_tmps_distinct [OF j_bound tssb_j]
	      have "distinct_read_tmps (?take_sbj@suspendsj)"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain 
		read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
		distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		apply (simp only: distinct_read_tmps_append)
		done
	      
	      from valid_history [OF j_bound tssb_j]
	      have h_consis: 
		"history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done
	      
	      have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	      proof -
		from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		  by simp
		from last_prog_hd_prog_append' [OF h_consis] this
		have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		  by (simp only: split_suspendsj [symmetric] suspendsj) 
		moreover 
		have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		  last_prog (hd_prog pj suspendsj) ?take_sbj"
		  apply (simp only: split_suspendsj [symmetric] suspendsj) 
		  by (rule last_prog_hd_prog_append)
		ultimately show ?thesis
		  by (simp add: split_suspendsj [symmetric] suspendsj) 
	      qed

	      from history_consistent_appendD [OF valid_sops_take read_tmps_take_drop 
		h_consis] last_prog_hd_prog
	      have hist_consis': "history_consistent θsbj (hd_prog pj suspendsj) suspendsj"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      from reads_consistent_drop_volatile_writes_no_volatile_reads  
	      [OF reads_consis] 
	      have no_vol_read: "outstanding_refs is_volatile_Readsb (ys) = {}"
		by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		  split_suspendsj )

	      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'="Readsb False a' t' v'#zs", simplified,
		OF j_bound_ts_a isj [simplified split_suspendsj] cph [simplified suspendsj]
		ts_a_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys
 		hist_consis' [simplified split_suspendsj] valid_sops_drop 
		distinct_read_tmps_drop [simplified split_suspendsj] 
		no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
		𝒮="?shared_a"]

	      obtain isj' j' where
		isj': "Read False a' t'# instrs zs @ issbj = isj' @ prog_instrs zs" and
		steps_ys: "(?ts_a, ?ma, ?shared_a)  d* 
		(?ts_a[j:=(last_prog
                (hd_prog pj zs) ys,
		isj',
                θsbj |` (dom θsbj - insert t' (read_tmps zs)),
                (), 𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Readsb False a' t' v']) @ zs)"
		by simp
	      from causal_program_history_suffix [OF this]
	      have cph': "causal_program_history issbj zs".	      
	      interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

	      from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
	      obtain isj''
		where isj': "isj' = Read False a' t'#isj''" and
		isj'': "instrs zs @ issbj = isj'' @ 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 = (psb, issb', 
		θsb(t  ret (m a) (f (θsb(t  m a)))),(), False, 𝒪sb  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 pj zs) ys, Read False a' t'#isj'', θsbj |` (dom θsbj - insert t' (read_tmps zs)), (), 𝒟j  outstanding_refs is_volatile_Writesb ys  {}, 
		acquired True ys (acquired True ?take_sbj 𝒪j),j')"
		by (clarsimp simp add: isj')
	      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_sbj 𝒪j) 
		a'  read_only (share ys (𝒮 ⊕⇘WR ⊖⇘AL))"
		apply cases
		apply (auto simp add: Let_def issb)
		done

	      with a'_unacq
	      have a'_ro: "a'  read_only (share ys (𝒮 ⊕⇘WR ⊖⇘AL))"
		by auto
	      from a'_in
	      have a'_not_ro: "a'  read_only (𝒮 ⊕⇘WR ⊖⇘AL)"
		by (auto simp add:  in_read_only_convs)

	      have "a'  𝒪j  all_acquired sbj"
	      proof -
		{
		  assume a_notin: "a'  𝒪j  all_acquired sbj"
		  from weak_sharing_consis [OF j_bound tssb_j]
		  have "weak_sharing_consistent 𝒪j sbj".
		  with weak_sharing_consistent_append [of 𝒪j ?take_sbj ?drop_sbj]
		  have "weak_sharing_consistent (acquired True ?take_sbj 𝒪j) suspendsj"
		    by (auto simp add: suspendsj)
		  with split_suspendsj
		  have weak_consis: "weak_sharing_consistent (acquired True ?take_sbj 𝒪j) ys"
		    by (simp add: weak_sharing_consistent_append)
		  from all_acquired_append [of ?take_sbj ?drop_sbj]
		  have "all_acquired ys  all_acquired sbj"
		    apply (clarsimp)
		    apply (clarsimp simp add: suspendsj [symmetric] split_suspendsj all_acquired_append)
		    done
	          with a_notin acquired_takeWhile_non_volatile_Writesb [of sbj 𝒪j]
                    all_acquired_append [of ?take_sbj ?drop_sbj]
		  have "a'  acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j  all_acquired ys"
                    by auto
                
		  from read_only_share_unowned [OF weak_consis this a'_ro] 
		  have "a'  read_only (𝒮 ⊕⇘WR ⊖⇘AL)" .
		  
		  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_Writesb)" sbj]

		  have "a'  read_only (𝒮 ⊕⇘WR ⊖⇘AL)"
		    by (auto simp add: acquired_takeWhile_non_volatile_Writesb)
		  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] tssb_j j_bound
	      have "A  all_acquired sbj = {}"
		by (auto simp add: Let_def)
	      moreover
	      from A_unowned_by_others [rule_format, OF _ neq_i_j] tssb_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 𝒮sb' tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	proof 
	  fix j isj 𝒪j j 𝒟j θj sbj pj
	  assume j_bound: "j < length tssb'"
	  assume tssb'_j: "tssb'!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	  show "non_volatile_owned_or_read_only False 𝒮sb' 𝒪j sbj"
	  proof (cases "j=i")
	    case True
	    have "non_volatile_owned_or_read_only False 
	      (𝒮sb ⊕⇘WR ⊖⇘AL) (𝒪sb  A - R) []"
	      by simp
	    then show ?thesis
	      using True i_bound tssb'_j
	      by (auto simp add: tssb' 𝒮sb' sb sb')
	  next
	    case False
	    from j_bound have j_bound': "j < length tssb"
	      by (auto simp add: tssb')
	    with tssb'_j False i_bound 
	    have tssb_j: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	      by (auto simp add: tssb')


	    note nvo = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' tssb_j]

	    from read_only_unowned [OF i_bound tssb_i] R_owned
	    have "R  read_only 𝒮sb = {}"
	      by auto
	    with A_no_read_only_reads [OF j_bound' tssb_j False [symmetric]] L_subset
	    have "aread_only_reads
	      (acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j)
	      (dropWhile (Not  is_volatile_Writesb) sbj).
		a  read_only 𝒮sb  a  read_only (𝒮sb ⊕⇘WR ⊖⇘AL)"
	      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 (𝒮sb ⊕⇘WR ⊖⇘AL) 𝒪j sbj".
	    thus ?thesis by (simp add: 𝒮sb')
	  qed
	qed
      next
	show "outstanding_volatile_writes_unowned_by_others tssb'"
	proof (unfold_locales)
	  fix i1 j p1 "is1" 𝒪1 1 𝒟1 xs1 sb1 pj "isj" "𝒪j" j 𝒟j xsj sbj
	  assume i1_bound: "i1 < length tssb'"
	  assume j_bound: "j < length tssb'"
	  assume i1_j: "i1  j"
	  assume ts_i1: "tssb'!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	  assume ts_j: "tssb'!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb1 = {}"
	  proof (cases "i1=i")
	    case True
	    with ts_i1 i_bound show ?thesis
	      by (simp add: tssb' sb' sb)
	  next
	    case False
	    note i1_i = this
	    from i1_bound have i1_bound': "i1 < length tssb"
	      by (simp add: tssb' sb' sb)
	    hence i1_bound'': "i1 < length (map owned tssb)"
	      by auto
	    from ts_i1 False have ts_i1': "tssb!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	      by (simp add: tssb' sb' sb)
	    show ?thesis
	    proof (cases "j=i")
	      case True

	      from i_bound ts_j tssb' True have sbj: "sbj=[]"
		by (simp add: tssb' sb')
	      from A_unused_by_others [rule_format, OF _ False [symmetric]] ts_i1 i1_bound''
		False i1_bound'
	      have "A  (𝒪1  outstanding_refs is_volatile_Writesb sb1) = {}"
		by (auto simp add: Let_def tssb' 𝒪sb' sb' owned_def)
	      moreover
	      from outstanding_volatile_writes_unowned_by_others 
	      [OF i1_bound' i_bound i1_i ts_i1' tssb_i]
	      have "𝒪sb  outstanding_refs is_volatile_Writesb sb1 = {}" by (simp add: sb)
	      
	      ultimately show ?thesis using ts_j True 
		by (auto simp add: i_bound tssb' 𝒪sb' sbj)
	    next
	      case False
	      from j_bound have j_bound': "j < length tssb"
		by (simp add: tssb')
	      from ts_j False have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (simp add: tssb')
	      from outstanding_volatile_writes_unowned_by_others 
              [OF i1_bound' j_bound' i1_j ts_i1' ts_j']
	      show "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb1 = {}" .
	    qed
	  qed
	qed
      next
	show "read_only_reads_unowned tssb'"
	proof 
	  fix n m
	  fix pn "isn" 𝒪n n 𝒟n θn sbn pm "ism" 𝒪m m 𝒟m θm sbm
	  assume n_bound: "n < length tssb'"
	    and m_bound: "m < length tssb'"
	    and neq_n_m: "nm"
	    and nth: "tssb'!n = (pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
	    and mth: "tssb'!m =(pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
	  from n_bound have n_bound': "n < length tssb" by (simp add: tssb')
	  from m_bound have m_bound': "m < length tssb" by (simp add: tssb')
	  show "(𝒪m  all_acquired sbm) 
            read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sbn) 𝒪n)
            (dropWhile (Not  is_volatile_Writesb) sbn) =
            {}"
	  proof (cases "m=i")
	    case True
	    with neq_n_m have neq_n_i: "ni"
	      by auto
	    
	    with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
	      by (auto simp add: tssb')
	    note read_only_reads_unowned [OF n_bound' i_bound  neq_n_i nth' tssb_i]
	    moreover
	    note A_no_read_only_reads [OF n_bound' nth']
	    ultimately
	    show ?thesis
	      using True tssb_i neq_n_i nth mth n_bound' m_bound'
	      by (auto simp add: tssb' 𝒪sb' sb sb')
	  next
	    case False
	    note neq_m_i = this
	    with m_bound mth i_bound have mth': "tssb!m = (pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
	      by (auto simp add: tssb')
	    show ?thesis
	    proof (cases "n=i")
	      case True
	      with tssb_i nth mth neq_m_i n_bound' 
	      show ?thesis
		by (auto simp add: tssb'  sb')
	    next
	      case False
	      with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
		by (auto simp add: tssb')
	      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 tssb'"
	proof (unfold_locales)
	  fix i1 j p1 "is1" 𝒪1 1 𝒟1 xs1 sb1 pj "isj" "𝒪j" j 𝒟j xsj sbj
	  assume i1_bound: "i1 < length tssb'"
	  assume j_bound: "j < length tssb'"
	  assume i1_j: "i1  j"
	  assume ts_i1: "tssb'!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	  assume ts_j: "tssb'!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "(𝒪1  all_acquired sb1)  (𝒪j  all_acquired sbj)= {}"
	  proof (cases "i1=i")
	    case True
	    with i1_j have i_j: "ij" 
	      by simp
	    
	    from i_bound ts_i1 tssb' True have sb1: "sb1=[]"
	      by (simp add: tssb' sb')
	    from j_bound have j_bound': "j < length tssb"
	      by (simp add: tssb')
	    hence j_bound'': "j < length (map owned tssb)"
	      by simp	    
	    from ts_j i_j have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (simp add: tssb')
	    
	    from A_unused_by_others [rule_format, OF _ i_j] ts_j i_j j_bound'
	    have "A  (𝒪j  outstanding_refs is_volatile_Writesb sbj) = {}"
	      by (auto simp add: Let_def tssb' owned_def)
	    moreover
	    from A_unacquired_by_others [rule_format, OF _ i_j] ts_j i_j j_bound'
	    have "A  all_acquired sbj = {}"
	      by (auto simp add: Let_def tssb')
(*
	    from a_not_acquired_others [rule_format, OF j_bound'' i_j] ts_j i_j j_bound'
	    have "a ∉ all_acquired sbj"
	      by (auto simp add: Let_def tssb')
*)
	    moreover
	    from ownership_distinct [OF i_bound j_bound' i_j tssb_i ts_j']
	    have "𝒪sb  (𝒪j  all_acquired sbj)= {}" by (simp add: sb)
	    ultimately show ?thesis using ts_i1 True
	      by (auto simp add: i_bound tssb' 𝒪sb' sb' sb1)
	  next
	    case False
	    note i1_i = this
	    from i1_bound have i1_bound': "i1 < length tssb"
	      by (simp add: tssb')
	    hence i1_bound'': "i1 < length (map owned tssb)"
	      by simp	    
	    from ts_i1 False have ts_i1': "tssb!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	      by (simp add: tssb')
	    show ?thesis
	    proof (cases "j=i")
	      case True
	      from A_unused_by_others [rule_format, OF _ False [symmetric]] ts_i1  
		False i1_bound'
	      have "A  (𝒪1  outstanding_refs is_volatile_Writesb sb1) = {}"
		by (auto simp add: Let_def tssb' owned_def)
	      moreover
	      from A_unacquired_by_others [rule_format, OF _ False [symmetric]] ts_i1  False i1_bound'
	      have "A  all_acquired sb1 = {}"
		by (auto simp add: Let_def tssb' owned_def)
	      moreover
	      from ownership_distinct [OF i1_bound' i_bound i1_i ts_i1' tssb_i]
	      have "(𝒪1  all_acquired sb1)  𝒪sb = {}" by (simp add: sb)
	      ultimately show ?thesis
		using ts_j True
		by (auto simp add: i_bound tssb' 𝒪sb' sb')
	    next
	      case False
	      from j_bound have j_bound': "j < length tssb"
		by (simp add: tssb')
	      from ts_j False have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (simp add: tssb')
	      from ownership_distinct [OF i1_bound' j_bound' i1_j ts_i1' ts_j']
	      show "(𝒪1  all_acquired sb1)  (𝒪j  all_acquired sbj) = {}" .
	    qed
	  qed
	qed
      qed

      have valid_hist': "valid_history program_step tssb'"
      proof -
	from valid_history [OF i_bound tssb_i] 
	have "history_consistent (θsb(tret (msb a) (f (θsb(tmsb a))))) (hd_prog psb []) []" by simp
	from valid_history_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' θsb' sb' sb)
      qed

      from valid_reads [OF i_bound tssb_i]
      have reads_consis: "reads_consistent False 𝒪sb msb sb" .

      have valid_reads': "valid_reads msb' tssb'"
      proof (unfold_locales)
	fix j pj "isj" 𝒪j j 𝒟j acqj θj sbj
	assume j_bound: "j < length tssb'"
	assume ts_j: "tssb'!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	show "reads_consistent False 𝒪j msb' sbj"
	proof (cases "i=j")
	  case True
	  from reads_consis ts_j j_bound sb show ?thesis
	    by (clarsimp simp add: True  msb' Writesb tssb' sb')
	next
	  case False
	  from j_bound have j_bound':  "j < length tssb"
	    by (simp add: tssb')
	  moreover from ts_j False have ts_j': "tssb ! j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	    using j_bound by (simp add: tssb')
	  ultimately have consis_m: "reads_consistent False 𝒪j msb sbj"
	    by (rule valid_reads)
	  let ?m' = "(msb(a := f (θsb(t  msb a))))"
	  from a_unowned_others [rule_format, OF _ False] j_bound' ts_j'
          obtain a_acq: "a  acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j" and
          a_unsh: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)"
            by auto
          with a_notin_unforwarded_non_volatile_reads_drop [OF j_bound' ts_j' False]
	  have "aacquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j  
                   all_shared (takeWhile (Not  is_volatile_Writesb) sbj)  
	           unforwarded_non_volatile_reads (dropWhile (Not  is_volatile_Writesb) sbj) {}. 
	    ?m' a = msb 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_Writesb) sbj)" 𝒪j]
	  have "reads_consistent False 𝒪j (msb(a := f (θsb(t  msb a)))) sbj"
	    by (auto simp del: fun_upd_apply)
	  thus ?thesis 
	    by (simp add: msb')
	qed
      qed

      have valid_sharing': "valid_sharing (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
      proof (intro_locales)	
	show "outstanding_non_volatile_writes_unshared (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	proof (unfold_locales)
	  fix j pj "isj" "𝒪j" j 𝒟j acqj xsj sbj
	  assume j_bound: "j < length tssb'"
	  assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "non_volatile_writes_unshared (𝒮sb ⊕⇘WR ⊖⇘AL)  sbj"
	  proof (cases "i=j")
	    case True
	    with i_bound jth show ?thesis
	      by (simp add: tssb' sb' sb)
	  next
	    case False
	    from j_bound have j_bound': "j < length tssb"
	      by (auto simp add: tssb')
	    from jth False have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (auto simp add: tssb')
	    from outstanding_non_volatile_writes_unshared [OF j_bound' jth']
	    have unshared: "non_volatile_writes_unshared 𝒮sb sbj".
	    have "adom (𝒮sb ⊕⇘WR ⊖⇘AL) - dom 𝒮sb. a  outstanding_refs is_non_volatile_Writesb sbj"
	    proof -
	      { 
		fix a 
		assume a_in: "a  dom (𝒮sb ⊕⇘WR ⊖⇘AL) - dom 𝒮sb"
		hence a_R: "a  R"
		  by clarsimp
		assume a_in_j: "a  outstanding_refs is_non_volatile_Writesb sbj"
		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 sbj"
		    by auto
		  moreover
		  with ownership_distinct [OF i_bound j_bound' False tssb_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_Writesb 
	    [OF unshared this]
	    show ?thesis .
	  qed
	qed
      next
	show "sharing_consis (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	proof (unfold_locales)  
	  fix j pj "isj" "𝒪j" j 𝒟j acqj xsj sbj
	  assume j_bound: "j < length tssb'"
	  assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "sharing_consistent (𝒮sb ⊕⇘WR ⊖⇘AL) 𝒪j sbj"
	  proof (cases "i=j")
	    case True
	    with i_bound jth show ?thesis
	      by (simp add: tssb' sb' sb)
	  next
	    case False
	    from j_bound have j_bound': "j < length tssb"
	      by (auto simp add: tssb')
	    from jth False have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (auto simp add: tssb')
	    from sharing_consis [OF j_bound' jth']
	    have consis: "sharing_consistent 𝒮sb 𝒪j sbj".

	    have acq_cond: "all_acquired sbj  dom 𝒮sb - dom (𝒮sb ⊕⇘WR ⊖⇘AL) = {}"
	    proof -
	      {
		fix a
		assume a_acq: "a  all_acquired sbj" 
		assume "a  dom 𝒮sb" 
		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 sbj = {}"
		    by auto
		  with a_acq a_L L_subset
		  show False
		    by blast
		qed
	      }
	      thus ?thesis
		by auto
	    qed
	    have uns_cond: "all_unshared sbj  dom (𝒮sb ⊕⇘WR ⊖⇘AL) - dom 𝒮sb = {}"
	    proof -
	      {
		fix a
		assume a_uns: "a  all_unshared sbj" 
		assume "a  L" 
		assume a_R:  "a  R"
		have False
		proof -
		  from unshared_acquired_or_owned [OF consis] a_uns
		  have "a  all_acquired sbj  𝒪j" by auto
		  with ownership_distinct [OF i_bound j_bound' False tssb_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: tssb')
	  qed
	qed
      next
	show "unowned_shared (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	proof (unfold_locales)
	  show "- ((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set tssb')  dom (𝒮sb ⊕⇘WR ⊖⇘AL)"
	  proof -

	    have s: "((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set tssb') =
              ((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set tssb)  A - R"
	      
	      apply (unfold tssb' 𝒪sb') 
	      apply (rule acquire_release_ownership_nth_update [OF R_owned i_bound tssb_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 (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	proof 
	  fix j pj "isj" "𝒪j" j 𝒟j acqj xsj sbj
	  assume j_bound: "j < length tssb'"
	  assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "𝒪j  read_only (𝒮sb ⊕⇘WR ⊖⇘AL) = {}"
	  proof (cases "i=j")
	    case True
	    from read_only_unowned [OF i_bound tssb_i] R_owned  A_R 
	    have "(𝒪sb  A - R)  read_only (𝒮sb ⊕⇘WR ⊖⇘AL) = {}"
	      by (auto simp add: in_read_only_convs )
	    with jth tssb_i i_bound True
	    show ?thesis
	      by (auto simp add: 𝒪sb' tssb')
	  next
	    case False
	    from j_bound have j_bound': "j < length tssb"
	      by (auto simp add: tssb')
	    with False jth have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (auto simp add: tssb')
	    from read_only_unowned [OF j_bound' jth']
	    have "𝒪j  read_only 𝒮sb = {}".
	    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 tssb_i jth']
	    have "𝒪sb  𝒪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 (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	proof 
	  fix j pj "isj" "𝒪j" j 𝒟j acqj xsj sbj
	  assume j_bound: "j < length tssb'"
	  assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "no_write_to_read_only_memory (𝒮sb ⊕⇘WR ⊖⇘AL) sbj"
	  proof (cases "i=j")
	    case True
	    with jth tssb_i i_bound 
	    show ?thesis
	      by (auto simp add: sb sb' tssb')
	  next
	    case False
	    from j_bound have j_bound': "j < length tssb"
	      by (auto simp add: tssb')
	    with False jth have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (auto simp add: tssb')
	    from no_outstanding_write_to_read_only_memory [OF j_bound' jth']
	    have nw: "no_write_to_read_only_memory 𝒮sb sbj".
	    have "R  outstanding_refs is_Writesb sbj = {}"
	    proof -
	      note dist = ownership_distinct [OF i_bound j_bound' False tssb_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_Writesb sbj  𝒪sb = {}"
		by auto
	      moreover
	      from outstanding_volatile_writes_unowned_by_others [OF j_bound'  i_bound 
		False [symmetric] jth' tssb_i ]
	      have "outstanding_refs is_volatile_Writesb sbj  𝒪sb = {}" 
		by auto
	      ultimately have "outstanding_refs is_Writesb sbj  𝒪sb = {}" 
		by (auto simp add: misc_outstanding_refs_convs)
	      with R_owned
	      show ?thesis by blast
	    qed
	    then
	    have "aoutstanding_refs is_Writesb sbj.
	      a  read_only (𝒮sb ⊕⇘WR ⊖⇘AL)  a  read_only 𝒮sb"
	      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 tssb'"
      proof (intro_locales)
	from load_tmps_distinct [OF i_bound tssb_i]
	have "distinct_load_tmps issb'" 
	  by (auto simp add: "issb" split: instr.splits)
	from load_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_distinct tssb'" by (simp add: tssb' sb' sb 𝒪sb' "issb")
      next
	from read_tmps_distinct [OF i_bound tssb_i]
	have "distinct_read_tmps []" by (simp add: tssb' sb' sb 𝒪sb')
	from read_tmps_distinct_nth_update [OF i_bound this]
	show "read_tmps_distinct tssb'" by (simp add: tssb' sb' sb 𝒪sb')
      next
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i] 
          load_tmps_distinct [OF i_bound tssb_i]
	have "load_tmps issb'  read_tmps [] = {}"
	  by (clarsimp)
	from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_read_tmps_distinct tssb'"  by (simp add: tssb' sb' sb 𝒪sb')
      qed

      have valid_sops': "valid_sops tssb'"
      proof -
	from valid_store_sops [OF i_bound tssb_i]
	obtain 
	  valid_store_sops': "sopstore_sops issb'. valid_sop sop"
	  by (auto simp add: "issb" tssb' sb' sb 𝒪sb')

	from valid_sops_nth_update [OF i_bound  _ valid_store_sops', where sb= "[]" ]
	show ?thesis by (auto simp add: tssb' sb' sb 𝒪sb')
      qed

      have valid_dd': "valid_data_dependency tssb'"
      proof -
	from data_dependency_consistent_instrs [OF i_bound tssb_i]
	obtain 
	  dd_is: "data_dependency_consistent_instrs (dom θsb')  issb'"
	  by (auto simp add: "issb" θsb')
	from load_tmps_write_tmps_distinct [OF i_bound tssb_i] 
	have "load_tmps issb'  (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: tssb' sb' sb 𝒪sb')
      qed

      have load_tmps_fresh': "load_tmps_fresh tssb'"
      proof -
	from load_tmps_fresh [OF i_bound tssb_i] 
	have "load_tmps (RMW a t (D,f) cond ret A L R W # issb')  dom θsb = {}"
	  by (simp add: "issb")
	moreover
	from load_tmps_distinct [OF i_bound tssb_i] have "t  load_tmps issb'"
	  by (auto simp add: "issb")
	ultimately have "load_tmps issb'  dom (θsb(t  ret (msb a) (f (θsb(tmsb a))))) = {}"
	  by auto
	from load_tmps_fresh_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' θsb')
      qed

      from enough_flushs_nth_update [OF i_bound, where sb="[]" ]
      have enough_flushs': "enough_flushs tssb'"
	by (auto simp: tssb' sb' sb)


      have valid_program_history': "valid_program_history tssb'"
      proof -	
	have causal': "causal_program_history issb' sb'"
	  by (simp add: "issb" sb sb')
	have "last_prog psb sb' = psb"
	  by (simp add: sb' sb)
	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb' sb')
      qed

      from is_sim have "is": "is = RMW a t (D,f) cond ret A L R W # issb'"
	by (simp add: suspends sb "issb")

      from direct_memop_step.RMWWrite [where cond=cond and θ=θsb and m=m, OF cond']
      have "(RMW a t (D, f) cond ret A L R W # issb', θsb, (),m, 𝒟, 𝒪sb,sb, 𝒮)  
            (issb',θsb(t  ret (m a) (f (θsb(tm a)))), (), 
             m(a := f (θsb(t  m a))), False, 𝒪sb  A - R, Map.empty, 𝒮 ⊕⇘WR ⊖⇘AL)".

      from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
      have "(ts, m, 𝒮) d (ts[i := (psb, issb',θsb(t  ret (m a) (f (θsb(tm a)))), (), False, 𝒪sb  A - R,Map.empty)], 
             m(a := f (θsb(t  m a))),𝒮 ⊕⇘WR ⊖⇘AL)".

      moreover 

      have tmps_commute: "θsb(t  ret (msb a) (f (θsb(tmsb a)))) = 
	(θsb |` (dom θsb - {t}))(t  ret (msb a) (f (θsb(tmsb a))))"
	apply (rule ext)
	apply (auto simp add: restrict_map_def domIff)
	done

	 
      from a_unflushed tssb_i sb
      have a_unflushed':
	"j < length tssb. 
                  (let (_,_,_,sbj,_,_,_) = tssb!j 
                  in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))"
	by auto

      have all_shared_L: "i p is 𝒪  𝒟 acq θ sb. i < length tssb 
            tssb ! i = (p, is, θ, sb, 𝒟, 𝒪,) 
            all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {}"
      proof -
	{
	  fix j pj isj 𝒪j j 𝒟j θj sbj x
	  assume j_bound: "j < length tssb"
	  assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	  assume x_shared: "x  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)"
	  assume x_L: "x  L"
	  have False
	  proof (cases "i=j")
	    case True with x_shared tssb_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 sbj  all_acquired sbj  𝒪j".

	      moreover have "all_shared (takeWhile (Not  is_volatile_Writesb) sbj)  all_shared sbj"
		using all_shared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
		  "(dropWhile (Not  is_volatile_Writesb) sbj)"]
		by auto
	      moreover
	      from A_unacquired_by_others [rule_format, OF _ False] jth j_bound
	      have "A  all_acquired sbj = {}" 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 tssb 
            tssb ! i = (p, is, θ, sb, 𝒟, 𝒪,) 
            all_shared (takeWhile (Not  is_volatile_Writesb) sb)  A = {}"
      proof -
	{
	  fix j pj isj 𝒪j j 𝒟j θj sbj x
	  assume j_bound: "j < length tssb"
	  assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	  assume x_shared: "x  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)"
	  assume x_A: "x  A"
	  have False
	  proof (cases "i=j")
	    case True with x_shared tssb_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 sbj  all_acquired sbj  𝒪j".

	      moreover have "all_shared (takeWhile (Not  is_volatile_Writesb) sbj)  all_shared sbj"
		using all_shared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
		  "(dropWhile (Not  is_volatile_Writesb) sbj)"]
		by auto
	      moreover
	      from A_unacquired_by_others [rule_format, OF _ False] jth j_bound
	      have "A  all_acquired sbj = {}" 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 tssb 
            tssb ! i = (p, is, θ, sb, 𝒟, 𝒪,) 
            all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {}"
	using L_subset by blast

     have all_unshared_R: "i p is 𝒪  𝒟 θ sb. i < length tssb 
            tssb ! i = (p, is, θ, sb, 𝒟, 𝒪,) 
            all_unshared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}"
      proof -
	{
	  fix j pj isj 𝒪j j 𝒟j θj sbj x
	  assume j_bound: "j < length tssb"
	  assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	  assume x_unshared: "x  all_unshared (takeWhile (Not  is_volatile_Writesb) sbj)"
	  assume x_R: "x  R"
	  have False
	  proof (cases "i=j")
	    case True with x_unshared tssb_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 sbj  all_acquired sbj  𝒪j".

	      moreover have "all_unshared (takeWhile (Not  is_volatile_Writesb) sbj)  all_unshared sbj"
		using all_unshared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
		  "(dropWhile (Not  is_volatile_Writesb) sbj)"]
		by auto
	      moreover

	      note ownership_distinct [OF i_bound j_bound False tssb_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 tssb 
            tssb ! i = (p, is, θ, sb, 𝒟, 𝒪,) 
            all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  R = {}"
      proof -
	{
	  fix j pj isj 𝒪j j 𝒟j θj sbj x
	  assume j_bound: "j < length tssb"
	  assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	  assume x_acq: "x  all_acquired (takeWhile (Not  is_volatile_Writesb) sbj)"
	  assume x_R: "x  R"
	  have False
	  proof (cases "i=j")
	    case True with x_acq tssb_i jth show False by (simp add: sb)
	  next
	    case False
	    show False
	    proof -

	      from x_acq have "x  all_acquired sbj"
		using all_acquired_append [of "takeWhile (Not  is_volatile_Writesb) sbj" 
		  "dropWhile (Not  is_volatile_Writesb) sbj"]
		by auto
	      moreover
	      note ownership_distinct [OF i_bound j_bound False tssb_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 tssb 
            tssb ! i = (p, is, θ, sb, 𝒟, 𝒪,) 
            all_shared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}"
      proof -
	{
	  fix j pj isj 𝒪j j 𝒟j θj sbj x
	  assume j_bound: "j < length tssb"
	  assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	  assume x_shared: "x  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)"
	  assume x_R: "x  R"
	  have False
	  proof (cases "i=j")
	    case True with x_shared tssb_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 sbj  all_acquired sbj  𝒪j".

	      moreover have "all_shared (takeWhile (Not  is_volatile_Writesb) sbj)  all_shared sbj"
		using all_shared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
		  "(dropWhile (Not  is_volatile_Writesb) sbj)"]
		by auto
	      moreover
	      note ownership_distinct [OF i_bound j_bound False tssb_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 tssb sharing_consis 𝒮sb tssb 
	all_shared_L all_shared_A all_acquired_R all_unshared_R all_shared_R]
      have share_commute: "share_all_until_volatile_write tssb 𝒮sb ⊕⇘WR ⊖⇘AL =
        share_all_until_volatile_write tssb (𝒮sb ⊕⇘WR ⊖⇘AL)".

      {
	fix j pj isj 𝒪j j 𝒟j θj sbj x
	assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	assume j_bound: "j < length tssb"
        assume neq: "i  j" 

        have "release (takeWhile (Not  is_volatile_Writesb) sbj)
                            (dom 𝒮sb  R - L) j
              = release (takeWhile (Not  is_volatile_Writesb) sbj)
                            (dom 𝒮sb) j"
        proof -
          {
            fix a
            assume a_in: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)"
            have "(a  (dom 𝒮sb  R - L)) = (a  dom 𝒮sb)"
            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 sbj) = {}"
                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_Writesb) sbj)" 
                "(dropWhile (Not  is_volatile_Writesb) sbj)"]
              have a_in: "a  𝒪j  all_acquired sbj"
                by auto
              with ownership_distinct [OF i_bound j_bound neq  tssb_i jth]
              have "a  (𝒪sb  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 "(tssb',msb(a := f (θsb(t  msb a))),𝒮sb')  (ts[i := (psb,issb',
            θsb(t  ret (m a) (f (θsb(tm a)))),(), False,𝒪sb  A - R,Map.empty)],m(a := f (θsb(t  m a))),𝒮 ⊕⇘WR ⊖⇘AL)"
	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] tssb')
	apply    (simp add: flush_all_until_volatile_nth_update_unused [OF i_bound tssb_i, simplified sb] sb')
	apply    (simp add: tssb' sb' 𝒪sb' m 
	  flush_all_until_volatile_nth_update_unused [OF i_bound tssb_i, simplified sb])
	using  share_all_until_volatile_write_RMW_commute [OF i_bound tssb_i [simplified issb sb]]
	apply  (clarsimp simp add: 𝒮 tssb' 𝒮sb' issb 𝒪sb' sb' θsb' sb' sb share_commute)
	using  leq
	apply  (simp add: tssb')
	using i_bound i_bound' ts_sim
	apply (clarsimp simp add: Let_def nth_list_update 
	  tssb' sb' sb 𝒪sb' sb' 𝒮sb' θsb' 𝒟sb'  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' msb' 𝒮sb' 
	by (auto simp del: fun_upd_apply)	
    next
      case (SBHGhost A L R W)
      then obtain 
	"issb": "issb = Ghost A L R W# issb'" and
	𝒪sb': "𝒪sb'=𝒪sb" and
        sb': "sb'=sb" and
	θsb': "θsb' = θsb" and
	𝒟sb': "𝒟sb'=𝒟sb" and
	sb': "sb'=sb@[Ghostsb A L R W]" and
	msb': "msb' = msb" and
	𝒮sb': "𝒮sb'=𝒮sb" 
	by auto

      from safe_memop_flush_sb [simplified issb] obtain      
        L_subset: "L  A" and
	A_shared_owned: "A  dom (share ?drop_sb 𝒮)  acquired True sb 𝒪sb" and
	R_acq: "R  acquired True sb 𝒪sb" and
	A_R: "A  R = {}" and
        A_unowned_by_others_ts:  
	"j<length (map owned ts). ij  (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 tssb. ij  (let (_,_,_,sbj,_,𝒪j,_) = tssb!j 
	  in A  (acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j 
                  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)) = {})" 
  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 tssb). i  j 
             (let (𝒪j, sbj) = map 𝒪_sb tssb! j
             in A  outstanding_refs is_volatile_Writesb sbj = {})"
      proof -
	{
	  fix j 𝒪j sbj
	  assume j_bound: "j < length (map owned tssb)"
	  assume neq_i_j: "ij"
	  assume tssb_j: "(map 𝒪_sb tssb)!j = (𝒪j,sbj)"
	  assume conflict: "A  outstanding_refs is_volatile_Writesb sbj  {}"
	  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 tssb"
	      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_Writesb sbj"
	      by auto

	    let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	    let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"

	    from ts_sim [rule_format, OF  j_bound''] tssb_j j_bound''
	    obtain pj suspendsj "issbj" θsbj 𝒟sbj 𝒟j j "isj" where
	      tssb_j: "tssb ! j = (pj,issbj,θsbj, sbj,𝒟sbj,𝒪j,j)" and
	      suspendsj: "suspendsj = ?drop_sbj" and
	      𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
	      isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	      tsj: "ts!j = (hd_prog pj suspendsj, isj,
	             θsbj |` (dom θsbj - read_tmps suspendsj),(), 
	             𝒟j, acquired True ?take_sbj 𝒪j, release ?take_sbj (dom 𝒮sb) j)"
	      apply (cases "tssb!j")
	      apply (force simp add: Let_def)
	      done
	      
	    have "a'  outstanding_refs is_volatile_Writesb suspendsj"
	    proof -	
	      from a'_in_j 
	      have "a'  outstanding_refs is_volatile_Writesb (?take_sbj @ ?drop_sbj)"
		by simp
	      thus ?thesis
		apply (simp only: outstanding_refs_append suspendsj)
		apply (auto simp add: outstanding_refs_conv dest: set_takeWhileD)
		done
	    qed
		
	    from split_volatile_Writesb_in_outstanding_refs [OF this]
	    obtain sop v ys zs A' L' R' W' where
	      split_suspendsj: "suspendsj = ys @ Writesb True a' sop v A' L' R' W' # zs" (is "suspendsj = ?suspends")
	      by blast
	    
	    from direct_memop_step.Ghost [where  θ=θsb and m="flush ?drop_sb m"]
	    have "(Ghost A L R W# issb', 
                       θsb, (), flush ?drop_sb m, 𝒟sb, 
                       acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb, share ?drop_sb 𝒮)  
                    (issb', θsb, (), flush ?drop_sb m, 𝒟sb, 
                      acquired True sb 𝒪sb  A - R, 
                      augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮sb) sb),
                      share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)".
	   
	    from direct_computation.concurrent_step.Memop [OF 
	      i_bound_ts' [simplified issb] ts'_i [simplified issb] this [simplified issb]] 
	    have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) d 
                    (?ts'[i := (psb, issb', θsb, (),𝒟sb, acquired True sb 𝒪sb  A - R,augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮sb) sb))], 
                         flush ?drop_sb m,share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)"
		  (is " _ d (?ts_A, ?m_A, ?share_A)")
	     by (simp add: issb)
	      
	       
	   from i_bound' have i_bound'': "i < length ?ts_A"
	     by simp

	   from valid_program_history [OF j_bound'' tssb_j] 
	   have "causal_program_history issbj sbj".
	   then have cph: "causal_program_history issbj ?suspends"
	     apply -
	     apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
	     apply (simp only: split_suspendsj [symmetric] suspendsj)
	     apply (simp add: split_suspendsj)
	     done
	   
	   from tsj neq_i_j j_bound 
	   have ts_A_j: "?ts_A!j = (hd_prog pj (ys @ Writesb True a' sop v A' L' R' W' # zs), isj,
	     θsbj |` (dom θsbj - read_tmps (ys @ Writesb True a' sop v A' L' R' W' # zs)), (), 𝒟j, 
	     acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	     by (simp add: split_suspendsj)


	   from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
	     by simp

	   from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
	   then
	   have lp: "last_prog pj ?suspends = pj"
	     apply -
	     apply (rule last_prog_same_append [where sb="?take_sbj"])
	     apply (simp only: split_suspendsj [symmetric] suspendsj)
	     apply simp
	     done

	   from valid_reads [OF j_bound'' tssb_j]
	   have reads_consis: "reads_consistent False 𝒪j msb sbj".

	   from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb j_bound''
	     tssb_j reads_consis]
	   have reads_consis_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
	     by (simp add: m suspendsj)

	   from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j tssb_i tssb_j]
	   have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
	     by (simp add: suspendsj)
	   from reads_consistent_flush_independent [OF this reads_consis_m]
	   have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	     ?m_A suspendsj".
	   hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  ?m_A ys"
	     by (simp add: split_suspendsj reads_consistent_append)


	   from valid_history [OF j_bound'' tssb_j]
	   have h_consis: 
	     "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
	     apply (simp only: split_suspendsj [symmetric] suspendsj)
	     apply simp
	     done

	   have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	   proof -
	     from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
	       by simp
	     from last_prog_hd_prog_append' [OF h_consis] this
	     have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
	       by (simp only: split_suspendsj [symmetric] suspendsj) 
	     moreover 
	     have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		  last_prog (hd_prog pj suspendsj) ?take_sbj"
	       apply (simp only: split_suspendsj [symmetric] suspendsj) 
	       by (rule last_prog_hd_prog_append)
	     ultimately show ?thesis
	       by (simp add: split_suspendsj [symmetric] suspendsj) 
	   qed

	   from valid_write_sops [OF j_bound'' tssb_j]
	   have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
	     by (simp add: split_suspendsj [symmetric] suspendsj)
	   then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
	     valid_sops_drop: "sopwrite_sops ys. valid_sop sop"
	     apply (simp only: write_sops_append )
	     apply auto
	     done

	   from read_tmps_distinct [OF j_bound'' tssb_j]
	   have "distinct_read_tmps (?take_sbj@suspendsj)"
	     by (simp add: split_suspendsj [symmetric] suspendsj)
	   then obtain 
	     read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
	     distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
	     apply (simp only: split_suspendsj [symmetric] suspendsj) 
	     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 θsbj (hd_prog pj suspendsj) suspendsj"
	     by (simp add: split_suspendsj [symmetric] suspendsj) 
	    from reads_consistent_drop_volatile_writes_no_volatile_reads  
	    [OF reads_consis] 
	    have no_vol_read: "outstanding_refs is_volatile_Readsb ys = {}"
	      by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )
	   
	    from flush_store_buffer_append [
	      OF j_bound''''  isj [simplified split_suspendsj] cph [simplified suspendsj]
	      ts_A_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_m_A_ys
	      hist_consis' [simplified split_suspendsj] valid_sops_drop distinct_read_tmps_drop [simplified split_suspendsj] 
	      no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="?share_A"]
	    obtain isj' j' where
	      isj': "instrs (Writesb True a' sop v A' L' R' W' # zs) @ issbj = 
	            isj' @ prog_instrs (Writesb 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 pj (Writesb True a' sop v A' L' R' W' # zs)) ys,
                           isj',
                           θsbj |` (dom θsbj - read_tmps (Writesb True a' sop v A' L' R' W' # zs)),(),
                           𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Writesb True a' sop v A' L' R' W']) @ zs)"
	      by simp
	    from causal_program_history_suffix [OF this]
	    have cph': "causal_program_history issbj zs".	      
	    interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

	    from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
	    obtain isj''
	      where isj': "isj' = Write True a' sop A' L' R' W' #isj''" and
	      isj'': "instrs zs @ issbj = isj'' @ 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 pj (Writesb True a' sop v A' L' R' W'# zs)) ys, isj',
                 θsbj |` (dom θsbj - read_tmps (Writesb True a' sop v A' L' R' W'# zs)),(),
                 𝒟j  outstanding_refs is_volatile_Writesb ys  {},
                 acquired True ys (acquired True ?take_sbj 𝒪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 isj']
	    have a_unowned: 
		"i < length ?ts_ys. ji  (let (𝒪i) = map owned ?ts_ys!i in a'  𝒪i)"
	      apply cases
	      apply (auto simp add: Let_def issb)
	      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 tssb). i  j 
             (let (𝒪j, sbj) = map 𝒪_sb tssb! j
             in A  all_acquired sbj = {})"
      proof -
	{
	  fix j 𝒪j sbj
	  assume j_bound: "j < length (map owned tssb)"
	  assume neq_i_j: "ij"
	  assume tssb_j: "(map 𝒪_sb tssb)!j = (𝒪j,sbj)"
	  assume conflict: "A  all_acquired sbj  {}"
	  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 tssb"
	      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 sbj"
	      by auto

	    let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	    let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"

	    from ts_sim [rule_format, OF  j_bound''] tssb_j j_bound''
	    obtain pj suspendsj "issbj" θsbj 𝒟sbj 𝒟j j "isj" where
	      tssb_j: "tssb ! j = (pj,issbj, θsbj, sbj,𝒟sbj,𝒪j,j)" and
	      suspendsj: "suspendsj = ?drop_sbj" and
	      𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
	      isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	      tsj: "ts!j = (hd_prog pj suspendsj, isj,
	                   θsbj |` (dom θsbj - read_tmps suspendsj),(), 
                           𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	      apply (cases "tssb!j")
	      apply (force simp add: Let_def)
	      done

	    from a'_in_j all_acquired_append [of ?take_sbj ?drop_sbj]
	    have "a'  all_acquired ?take_sbj  a'  all_acquired suspendsj"
	      by (auto simp add: suspendsj)
	    thus False
	    proof 
	      assume "a'  all_acquired ?take_sbj"
	      with A_unowned_by_others [rule_format, OF _ neq_i_j] tssb_j j_bound a'_in
	      show False
		by (auto dest: all_acquired_unshared_acquired)
	    next
	      assume conflict_drop: "a'  all_acquired suspendsj"
	      from split_all_acquired_in [OF conflict_drop]
	      (* FIXME: exract common parts *)
	      show False
	      proof 
		assume "sop a'' v ys zs A L R W. 
                         suspendsj = ys @ Writesb True a'' sop v A L R W# zs  a'  A" 
	        then
		obtain a'' sop' v' ys zs A' L' R' W' where
		  split_suspendsj: "suspendsj = ys @ Writesb True a'' sop' v' A' L' R' W' # zs" 
		    (is "suspendsj = ?suspends") and
		  a'_A': "a'  A'"
		 by auto
	    
	       from direct_memop_step.Ghost [where  θ=θsb and m="flush ?drop_sb m"]
	       have "(Ghost A L R W# issb', 
                         θsb, (), flush ?drop_sb m,𝒟sb, 
                         acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb,share ?drop_sb 𝒮)  
                    (issb', θsb, (), flush ?drop_sb m, 𝒟sb, 
                      acquired True sb 𝒪sb  A - R, 
                      augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮sb) sb),
                      share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)".

	       from direct_computation.concurrent_step.Memop [OF 
		 i_bound_ts' [simplified issb] ts'_i [simplified issb] this [simplified issb]] 
	       have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) d 
                   (?ts'[i := (psb, issb',θsb, (),𝒟sb, 
                         acquired True sb 𝒪sb  A - R,
                         augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮sb) sb))], 
          
                         flush ?drop_sb m,share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)"
		   (is " _ d (?ts_A, ?m_A, ?share_A)")
		 by (simp add: issb)
	      
	       
	       from i_bound' have i_bound'': "i < length ?ts_A"
		 by simp

	       from valid_program_history [OF j_bound'' tssb_j] 
	       have "causal_program_history issbj sbj".
	       then have cph: "causal_program_history issbj ?suspends"
		 apply -
		 apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply (simp add: split_suspendsj)
		 done
	       
	       from tsj neq_i_j j_bound 
	       have ts_A_j: "?ts_A!j = (hd_prog pj (ys @ Writesb True a'' sop' v' A' L' R' W' # zs), isj,
		   θsbj |` (dom θsbj - read_tmps (ys @ Writesb True a'' sop' v' A' L' R' W' # zs)), (), 𝒟j, 
		   acquired True ?take_sbj 𝒪j, release ?take_sbj (dom 𝒮sb) j)"
		 by (simp add: split_suspendsj)


	       from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
		 by simp

	       from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
	       then
	       have lp: "last_prog pj ?suspends = pj"
		 apply -
		 apply (rule last_prog_same_append [where sb="?take_sbj"])
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply simp
		 done
	       

	       from valid_reads [OF j_bound'' tssb_j]
	       have reads_consis: "reads_consistent False 𝒪j msb sbj".
	       
	       from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb j_bound''
		 tssb_j reads_consis]
	       have reads_consis_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		 by (simp add: m suspendsj)

	       from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j tssb_i tssb_j]
	       have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
		 by (simp add: suspendsj)
	       from reads_consistent_flush_independent [OF this reads_consis_m]
	       have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
		 ?m_A suspendsj".
	       hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  ?m_A ys"
		 by (simp add: split_suspendsj reads_consistent_append)       
	       
	       from valid_history [OF j_bound'' tssb_j]
	       have h_consis: 
		 "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply simp
		 done
	       
	       have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	       proof -
		 from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		   by simp
		 from last_prog_hd_prog_append' [OF h_consis] this
		 have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		   by (simp only: split_suspendsj [symmetric] suspendsj) 
		 moreover 
		 have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		   last_prog (hd_prog pj suspendsj) ?take_sbj"
		   apply (simp only: split_suspendsj [symmetric] suspendsj) 
		   by (rule last_prog_hd_prog_append)
		 ultimately show ?thesis
		   by (simp add: split_suspendsj [symmetric] suspendsj) 
	       qed
	       
	       from valid_write_sops [OF j_bound'' tssb_j]
	       have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		 by (simp add: split_suspendsj [symmetric] suspendsj)
	       then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		 valid_sops_drop: "sopwrite_sops ys. valid_sop sop"
		 apply (simp only: write_sops_append )
		 apply auto
		 done
	       
	       from read_tmps_distinct [OF j_bound'' tssb_j]
	       have "distinct_read_tmps (?take_sbj@suspendsj)"
		 by (simp add: split_suspendsj [symmetric] suspendsj)
	       then obtain 
		 read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
		 distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		 apply (simp only: split_suspendsj [symmetric] suspendsj) 
		 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 θsbj (hd_prog pj suspendsj) suspendsj"
		 by (simp add: split_suspendsj [symmetric] suspendsj) 
	       from reads_consistent_drop_volatile_writes_no_volatile_reads  
	       [OF reads_consis] 
	       have no_vol_read: "outstanding_refs is_volatile_Readsb ys = {}"
		 by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		   split_suspendsj )
	       
	       from flush_store_buffer_append [
		 OF j_bound''''  isj [simplified split_suspendsj] cph [simplified suspendsj]
		 ts_A_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_m_A_ys
		 hist_consis' [simplified split_suspendsj] valid_sops_drop distinct_read_tmps_drop [simplified split_suspendsj] 
		 no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
		 𝒮="?share_A"]
	       obtain isj' j' where
		 isj': "instrs (Writesb True a'' sop' v' A' L' R' W' # zs) @ issbj = 
	            isj' @ prog_instrs (Writesb 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 pj (Writesb True a'' sop' v' A' L' R' W' # zs)) ys,
                           isj',
                           θsbj |` (dom θsbj - read_tmps (Writesb True a'' sop' v' A' L' R' W' # zs)),(),
                           𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Writesb True a'' sop' v' A' L' R' W']) @ zs)"
		 by simp
	       from causal_program_history_suffix [OF this]
	       have cph': "causal_program_history issbj zs".	      
	       interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

	       from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
	       obtain isj''
		 where isj': "isj' = Write True a'' sop' A' L' R' W'#isj''" and
		 isj'': "instrs zs @ issbj = isj'' @ 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 pj (Writesb True a'' sop' v' A' L' R' W'# zs)) ys, isj',
                 θsbj |` (dom θsbj - read_tmps (Writesb True a'' sop' v' A' L' R' W'# zs)),(), 
		 𝒟j  outstanding_refs is_volatile_Writesb ys  {},
                 acquired True ys (acquired True ?take_sbj 𝒪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 isj']
	       have A'_unowned: 
		 "i < length ?ts_ys. ji  (let (𝒪i) = map owned ?ts_ys!i in A'   𝒪i = {})"
		 apply cases
		 apply (fastforce simp add: Let_def issb)+
		 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. 
                 suspendsj = ys @ Ghostsb A L R W # zs  a'  A" 
	       then
	       obtain ys zs A' L' R' W' where
		  split_suspendsj: "suspendsj = ys @ Ghostsb A' L' R' W'# zs" (is "suspendsj = ?suspends") and
		 a'_A': "a'  A'"
		 by auto
		 
	       from direct_memop_step.Ghost [where  θ=θsb and m="flush ?drop_sb m"]
	       have "(Ghost A L R W# issb', 
                       θsb, (), flush ?drop_sb m, 𝒟sb, 
                       acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb, share ?drop_sb 𝒮)  
                    (issb', θsb, (), flush ?drop_sb m, 𝒟sb, 
                      acquired True sb 𝒪sb  A - R, 
                      augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮sb) sb),
                      share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)".

	       from direct_computation.concurrent_step.Memop [OF 
		 i_bound_ts' [simplified issb] ts'_i [simplified issb] this [simplified issb]] 
	       have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) d 
                   (?ts'[i := (psb, issb', θsb, (), 𝒟sb, acquired True sb 𝒪sb  A - R,augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮sb) sb))], 
                         flush ?drop_sb m,share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)"
		   (is " _ d (?ts_A, ?m_A, ?share_A)")
		 by (simp add: issb)
	      
	       
	       from i_bound' have i_bound'': "i < length ?ts_A"
		 by simp

	       from valid_program_history [OF j_bound'' tssb_j] 
	       have "causal_program_history issbj sbj".
	       then have cph: "causal_program_history issbj ?suspends"
		 apply -
		 apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply (simp add: split_suspendsj)
		 done
	       
	       from tsj neq_i_j j_bound 
	       have ts_A_j: "?ts_A!j = (hd_prog pj (ys @ Ghostsb A' L' R' W'# zs), isj,
		 θsbj |` (dom θsbj - read_tmps (ys @ Ghostsb A' L' R' W'# zs)), (),𝒟j, 
		 acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
		 by (simp add: split_suspendsj)


	       from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
		 by simp
	       
	       from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
	       then
	       have lp: "last_prog pj ?suspends = pj"
		 apply -
		 apply (rule last_prog_same_append [where sb="?take_sbj"])
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply simp
		 done
	       from valid_reads [OF j_bound'' tssb_j]
	       have reads_consis: "reads_consistent False 𝒪j msb sbj".
	       
	       from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb j_bound''
		 tssb_j reads_consis]
	       have reads_consis_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		 by (simp add: m suspendsj)

	       from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j tssb_i tssb_j]
	       have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
		 by (simp add: suspendsj)
	       from reads_consistent_flush_independent [OF this reads_consis_m]
	       have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
		 ?m_A suspendsj".
	       hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  ?m_A ys"
		 by (simp add: split_suspendsj reads_consistent_append)


	       from valid_history [OF j_bound'' tssb_j]
	       have h_consis: 
		 "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		 apply (simp only: split_suspendsj [symmetric] suspendsj)
		 apply simp
		 done
	       
	       have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	       proof -
		 from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		   by simp
		 from last_prog_hd_prog_append' [OF h_consis] this
		 have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		   by (simp only: split_suspendsj [symmetric] suspendsj) 
		 moreover 
		 have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		   last_prog (hd_prog pj suspendsj) ?take_sbj"
		   apply (simp only: split_suspendsj [symmetric] suspendsj) 
		   by (rule last_prog_hd_prog_append)
		 ultimately show ?thesis
		   by (simp add: split_suspendsj [symmetric] suspendsj) 
	       qed
	       
	       from valid_write_sops [OF j_bound'' tssb_j]
	       have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		 by (simp add: split_suspendsj [symmetric] suspendsj)
	       then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		   valid_sops_drop: "sopwrite_sops ys. valid_sop sop"
		 apply (simp only: write_sops_append )
		 apply auto
		 done
	       
	       from read_tmps_distinct [OF j_bound'' tssb_j]
	       have "distinct_read_tmps (?take_sbj@suspendsj)"
		 by (simp add: split_suspendsj [symmetric] suspendsj)
	       then obtain 
		 read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
		 distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		 apply (simp only: split_suspendsj [symmetric] suspendsj) 
		 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 θsbj (hd_prog pj suspendsj) suspendsj"
		 by (simp add: split_suspendsj [symmetric] suspendsj) 
	       from reads_consistent_drop_volatile_writes_no_volatile_reads  
	       [OF reads_consis] 
	       have no_vol_read: "outstanding_refs is_volatile_Readsb ys = {}"
		 by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		   split_suspendsj )
	   
	       from flush_store_buffer_append [
		 OF j_bound''''  isj [simplified split_suspendsj] cph [simplified suspendsj]
		 ts_A_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_m_A_ys
		 hist_consis' [simplified split_suspendsj] valid_sops_drop distinct_read_tmps_drop [simplified split_suspendsj] 
		 no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
		 𝒮="?share_A"]
	       obtain isj' j' where
		 isj': "instrs (Ghostsb A' L' R' W'# zs) @ issbj = 
	            isj' @ prog_instrs (Ghostsb A' L' R' W'# zs)" and
		 steps_ys: "(?ts_A, ?m_A, ?share_A)  d* 
		(?ts_A[j:= (last_prog (hd_prog pj (Ghostsb A' L' R' W'# zs)) ys,
                           isj',
                           θsbj |` (dom θsbj - read_tmps (Ghostsb A' L' R' W'# zs)),(),
		           𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Ghostsb A' L' R' W']) @ zs)"
		 by simp
	       from causal_program_history_suffix [OF this]
	       have cph': "causal_program_history issbj zs".	      
	       interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

	       from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
	       obtain isj''
		 where isj': "isj' = Ghost A' L' R' W'#isj''" and
		 isj'': "instrs zs @ issbj = isj'' @ 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 pj (Ghostsb A' L' R' W'# zs)) ys, isj',
                 θsbj |` (dom θsbj - read_tmps (Writesb True a'' sop' v' A' L' R' W'# zs)),(),
		 𝒟j  outstanding_refs is_volatile_Writesb ys  {},
                 acquired True ys (acquired True ?take_sbj 𝒪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 isj']
	       have A'_unowned: 
		 "i < length ?ts_ys. ji  (let (𝒪i) = map owned ?ts_ys!i in A'   𝒪i = {})"
		 apply cases
		 apply (fastforce simp add: Let_def issb)+
		 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 tssb). i  j 
             (let (𝒪j, sbj) = map 𝒪_sb tssb! j
             in A  read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j)
	             (dropWhile (Not  is_volatile_Writesb) sbj) = {})"
      proof -
	{
	  fix j 𝒪j sbj
	  assume j_bound: "j < length (map owned tssb)"
	  assume neq_i_j: "ij"
	  assume tssb_j: "(map 𝒪_sb tssb)!j = (𝒪j,sbj)"
	  let ?take_sbj = "(takeWhile (Not  is_volatile_Writesb) sbj)"
	  let ?drop_sbj = "(dropWhile (Not  is_volatile_Writesb) sbj)"

	  assume conflict: "A  read_only_reads (acquired True ?take_sbj 𝒪j) ?drop_sbj   {}"
	  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 tssb"
	      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_sbj 𝒪j) ?drop_sbj"
	      by auto


	    from ts_sim [rule_format, OF  j_bound''] tssb_j j_bound''
	    obtain pj suspendsj "issbj" 𝒟sbj 𝒟j j θsbj "isj" where
	      tssb_j: "tssb ! j = (pj,issbj, θsbj, sbj,𝒟sbj,𝒪j,j)" and
	      suspendsj: "suspendsj = ?drop_sbj" and
	      isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	      𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
	      tsj: "ts!j = (hd_prog pj suspendsj, isj,
	             θsbj |` (dom θsbj - read_tmps suspendsj),(), 𝒟j, acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	      apply (cases "tssb!j")
	      apply (force simp add: Let_def)
	      done
	      

	    from split_in_read_only_reads [OF a'_in_j [simplified suspendsj [symmetric]]]
	    obtain t v ys zs where
	      split_suspendsj: "suspendsj = ys @ Readsb False a' t v# zs" (is "suspendsj = ?suspends") and
	      a'_unacq: "a'  acquired True ys (acquired True ?take_sbj 𝒪j)"
	      by blast

	    
	    from direct_memop_step.Ghost [where  θ=θsb and m="flush ?drop_sb m"]
	    have "(Ghost A L R W# issb', 
                       θsb, (), flush ?drop_sb m, 𝒟sb, 
                       acquired True sb 𝒪sb, release sb (dom 𝒮sb) sb, share ?drop_sb 𝒮)  
                    (issb', θsb, (), flush ?drop_sb m, 𝒟sb, 
                      acquired True sb 𝒪sb  A - R, 
                      augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮sb) sb),
                      share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)".

	    from direct_computation.concurrent_step.Memop [OF 
		 i_bound_ts' [simplified issb] ts'_i [simplified issb] this [simplified issb]] 
	    have store_step: "(?ts', flush ?drop_sb m, share ?drop_sb 𝒮) d 
                    (?ts'[i := (psb, issb', θsb, (),𝒟sb, acquired True sb 𝒪sb  A - R,augment_rels (dom (share ?drop_sb 𝒮)) R (release sb (dom 𝒮sb) sb))], 
                         flush ?drop_sb m,share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)"
		  (is " _ d (?ts_A, ?m_A, ?share_A)")
	     by (simp add: issb)
	    
	    from i_bound' have i_bound'': "i < length ?ts_A"
	      by simp

	    from valid_program_history [OF j_bound'' tssb_j] 
	    have "causal_program_history issbj sbj".
	    then have cph: "causal_program_history issbj ?suspends"
	      apply -
	      apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply (simp add: split_suspendsj)
	      done
	       
	    from tsj neq_i_j j_bound 
	    have ts_A_j: "?ts_A!j = (hd_prog pj (ys @ Readsb False a' t v# zs), isj,
	      θsbj |` (dom θsbj - read_tmps (ys @ Readsb False a' t v# zs)), (),𝒟j, 
	      acquired True ?take_sbj 𝒪j,release ?take_sbj (dom 𝒮sb) j)"
	      by (simp add: split_suspendsj)
	    

	    from j_bound''' i_bound' neq_i_j have j_bound'''': "j < length ?ts_A"
	      by simp
	       
	    from valid_last_prog [OF j_bound'' tssb_j] have last_prog: "last_prog pj sbj = pj".
	    then
	    have lp: "last_prog pj ?suspends = pj"
	      apply -
	      apply (rule last_prog_same_append [where sb="?take_sbj"])
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	    from valid_reads [OF j_bound'' tssb_j]
	    have reads_consis: "reads_consistent False 𝒪j msb sbj".
	    
	    from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb j_bound''
		 tssb_j reads_consis]
	    have reads_consis_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
	      by (simp add: m suspendsj)
	    
	    from outstanding_non_write_non_vol_reads_drop_disj [OF i_bound j_bound'' neq_i_j tssb_i tssb_j]
	    have "outstanding_refs is_Writesb ?drop_sb  outstanding_refs is_non_volatile_Readsb suspendsj = {}"
	      by (simp add: suspendsj)
	    from reads_consistent_flush_independent [OF this reads_consis_m]
	    have reads_consis_flush_m: "reads_consistent True (acquired True ?take_sbj 𝒪j) 
	      ?m_A suspendsj".
	    hence reads_consis_m_A_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  ?m_A ys"
	      by (simp add: split_suspendsj reads_consistent_append)
	    

	    from valid_history [OF j_bound'' tssb_j]
	    have h_consis: 
	      "history_consistent θsbj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
	      apply (simp only: split_suspendsj [symmetric] suspendsj)
	      apply simp
	      done
	       
	    have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	    proof -
	      from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		by simp
	      from last_prog_hd_prog_append' [OF h_consis] this
	      have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		by (simp only: split_suspendsj [symmetric] suspendsj) 
	      moreover 
	      have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		last_prog (hd_prog pj suspendsj) ?take_sbj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		by (rule last_prog_hd_prog_append)
	      ultimately show ?thesis
		by (simp add: split_suspendsj [symmetric] suspendsj) 
	    qed
	    
	    from valid_write_sops [OF j_bound'' tssb_j]
	    have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		   valid_sops_drop: "sopwrite_sops ys. valid_sop sop"
	      apply (simp only: write_sops_append )
	      apply auto
	      done
	    
	    from read_tmps_distinct [OF j_bound'' tssb_j]
	    have "distinct_read_tmps (?take_sbj@suspendsj)"
	      by (simp add: split_suspendsj [symmetric] suspendsj)
	    then obtain 
		 read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
	      distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
	      apply (simp only: split_suspendsj [symmetric] suspendsj) 
	      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 θsbj (hd_prog pj suspendsj) suspendsj"
	      by (simp add: split_suspendsj [symmetric] suspendsj) 
	    from reads_consistent_drop_volatile_writes_no_volatile_reads  
	    [OF reads_consis] 
	    have no_vol_read: "outstanding_refs is_volatile_Readsb ys = {}"
	      by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )
	   
	    from flush_store_buffer_append [
		 OF j_bound''''  isj [simplified split_suspendsj] cph [simplified suspendsj]
		 ts_A_j [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_m_A_ys
		 hist_consis' [simplified split_suspendsj] valid_sops_drop distinct_read_tmps_drop [simplified split_suspendsj] 
		 no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
		 𝒮="?share_A"]
	    obtain isj' j' where
		 isj': "instrs (Readsb False a' t v # zs) @ issbj = 
	            isj' @ prog_instrs (Readsb False a' t v # zs)" and
		 steps_ys: "(?ts_A, ?m_A, ?share_A)  d* 
		(?ts_A[j:= (last_prog (hd_prog pj (Ghostsb A' L' R' W'# zs)) ys,
                           isj',
                           θsbj |` (dom θsbj - read_tmps (Readsb False a' t v # zs)),(),
		           𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪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 issbj ((ys @ [Readsb False a' t v]) @ zs)"
	      by simp
	    from causal_program_history_suffix [OF this]
	    have cph': "causal_program_history issbj zs".	      
	    interpret causalj: causal_program_history "issbj" "zs" by (rule cph')

	    from causalj.causal_program_history [of "[]", simplified, OF refl] isj'   
	    obtain isj''
	      where isj': "isj' = Read False a' t#isj''" and
	      isj'': "instrs zs @ issbj = isj'' @ 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 pj (Readsb False a' t v# zs)) ys, isj',
                 θsbj |` (dom θsbj - read_tmps (Readsb False a' t v# zs)),(),
	         𝒟j  outstanding_refs is_volatile_Writesb ys  {},
                 acquired True ys (acquired True ?take_sbj 𝒪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 isj']
	    have "a'  acquired True ys (acquired True ?take_sbj 𝒪j) 
                  a'  read_only (share ys (share ?drop_sb 𝒮 ⊕⇘WR  ⊖⇘AL))"
	      apply cases
	      apply (auto simp add: Let_def issb)
	      done
	    with a'_unacq
	    have a'_ro: "a'  read_only (share ys (share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL))"
	      by auto
	    from a'_in
	    have a'_not_ro: "a'  read_only (share ?drop_sb 𝒮 ⊕⇘WR  ⊖⇘AL)"
	      by (auto simp add:  in_read_only_convs)

	    have "a'  𝒪j  all_acquired sbj"
	    proof -
	      {
		assume a_notin: "a'  𝒪j  all_acquired sbj"
		from weak_sharing_consis [OF j_bound'' tssb_j]
		have "weak_sharing_consistent 𝒪j sbj".
		with weak_sharing_consistent_append [of 𝒪j ?take_sbj ?drop_sbj]
		have "weak_sharing_consistent (acquired True ?take_sbj 𝒪j) suspendsj"
		  by (auto simp add: suspendsj)
		with split_suspendsj
		have weak_consis: "weak_sharing_consistent (acquired True ?take_sbj 𝒪j) ys"
		  by (simp add: weak_sharing_consistent_append)
		from all_acquired_append [of ?take_sbj ?drop_sbj]
		have "all_acquired ys  all_acquired sbj"
		  apply (clarsimp)
		  apply (clarsimp simp add: suspendsj [symmetric] split_suspendsj all_acquired_append)
		  done
		with a_notin acquired_takeWhile_non_volatile_Writesb [of sbj 𝒪j]
                  all_acquired_append [of ?take_sbj ?drop_sbj]
		have "a'  acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j  all_acquired ys"
                  by auto
                
		from read_only_share_unowned [OF weak_consis this a'_ro] 
		have "a'  read_only (share ?drop_sb 𝒮 ⊕⇘WR ⊖⇘AL)" .
		  
		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] tssb_j j_bound
	    have "A  all_acquired sbj = {}"
	      by (auto simp add: Let_def)
	    moreover
	    from A_unowned_by_others [rule_format, OF _ neq_i_j] tssb_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 𝒮sb' tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	proof -
	  from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i] 
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb (sb @ [Ghostsb 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: tssb' sb' 𝒪sb' 𝒮sb')
	qed
      next
	show "outstanding_volatile_writes_unowned_by_others tssb'"
	proof (unfold_locales)
	  fix i1 j p1 "is1" 𝒪1 1 𝒟1 xs1 sb1 pj "isj" "𝒪j" j 𝒟j xsj sbj
	  assume i1_bound: "i1 < length tssb'"
	  assume j_bound: "j < length tssb'"
	  assume i1_j: "i1  j"
	  assume ts_i1: "tssb'!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	  assume ts_j: "tssb'!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb1 = {}"
	  proof (cases "i1=i")
	    case True
	    with i1_j have i_j: "ij" 
	      by simp
	    
	    from j_bound have j_bound': "j < length tssb"
	      by (simp add: tssb')
	    hence j_bound'': "j < length (map owned tssb)"
	      by simp
	    from ts_j i_j have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (simp add: tssb')

	    from outstanding_volatile_writes_unowned_by_others 
	    [OF i_bound j_bound' i_j tssb_i ts_j']
	    have "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb = {}".
	    with ts_i1 True i_bound show ?thesis
	      by (clarsimp simp add: tssb' sb' outstanding_refs_append 
		acquired_takeWhile_non_volatile_Writesb)
	  next
	    case False
	    note i1_i = this
	    from i1_bound have i1_bound': "i1 < length tssb"
	      by (simp add: tssb')
	    from ts_i1 False have ts_i1': "tssb!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	      by (simp add: tssb')
	    show ?thesis
	    proof (cases "j=i")
	      case True

	      from i1_bound'
	      have i1_bound'': "i1 < length (map owned tssb)"
		by simp

	      from outstanding_volatile_writes_unowned_by_others 
	      [OF i1_bound' i_bound i1_i ts_i1' tssb_i]
	      have "(𝒪sb  all_acquired sb)  outstanding_refs is_volatile_Writesb sb1 = {}".
	      moreover
	      from A_unused_by_others [rule_format, OF _ False [symmetric]] False ts_i1 i1_bound
	      have "A  outstanding_refs is_volatile_Writesb sb1 = {}"
		by (auto simp add: Let_def tssb')
	      
	      ultimately
	      show ?thesis
		using ts_j True tssb' 
		by (auto simp add: i_bound tssb' 𝒪sb' sb' all_acquired_append)
	    next
	      case False
	      from j_bound have j_bound': "j < length tssb"
		by (simp add: tssb')
	      from ts_j False have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (simp add: tssb')
	      from outstanding_volatile_writes_unowned_by_others 
              [OF i1_bound' j_bound' i1_j ts_i1' ts_j']
	      show "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb1 = {}" .
	    qed
	  qed
	qed
      next
	show "read_only_reads_unowned tssb'"
	proof 
	  fix n m
	  fix pn "isn" 𝒪n n 𝒟n θn sbn pm "ism" 𝒪m m 𝒟m θm sbm
	  assume n_bound: "n < length tssb'"
	    and m_bound: "m < length tssb'"
	    and neq_n_m: "nm"
	    and nth: "tssb'!n = (pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
	    and mth: "tssb'!m =(pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
	  from n_bound have n_bound': "n < length tssb" by (simp add: tssb')
	  from m_bound have m_bound': "m < length tssb" by (simp add: tssb')
	  show "(𝒪m  all_acquired sbm) 
            read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sbn) 𝒪n)
            (dropWhile (Not  is_volatile_Writesb) sbn) =
            {}"
	  proof (cases "m=i")
	    case True
	    with neq_n_m have neq_n_i: "ni"
	      by auto
	    with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
	      by (auto simp add: tssb')
	    note read_only_reads_unowned [OF n_bound' i_bound  neq_n_i nth' tssb_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_Writesb) sbn) 𝒪n)
              (dropWhile (Not  is_volatile_Writesb) sbn) =
              {}"
	      by auto
	    ultimately 
	    show ?thesis
	      using True tssb_i nth' mth n_bound' m_bound'
	      by (auto simp add: tssb' 𝒪sb' sb' all_acquired_append)
	  next
	    case False
	    note neq_m_i = this
	    with m_bound mth i_bound have mth': "tssb!m = (pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
	      by (auto simp add: tssb')
	    show ?thesis
	    proof (cases "n=i")
	      case True
	      note read_only_reads_unowned [OF i_bound m_bound' neq_m_i [symmetric] tssb_i mth']
	      then show ?thesis
		using True neq_m_i tssb_i nth mth n_bound' m_bound'
		apply (case_tac "outstanding_refs (is_volatile_Writesb) sb = {}")
		apply (clarsimp simp add: outstanding_vol_write_take_drop_appends
		  acquired_append read_only_reads_append tssb' sb' 𝒪sb')+
		done
	    next
	      case False
	      with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
		by (auto simp add: tssb')
	      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 tssb'"
	proof -
	  have "j<length tssb. i  j 
	    (let (pj, isj,θj, sbj, 𝒟j, 𝒪j,j) = tssb ! j
	      in (𝒪sb  all_acquired sb')  (𝒪j  all_acquired sbj) = {})"
	  proof -
	    {
	      fix j pj isj 𝒪j j 𝒟j θj sbj
	      assume neq_i_j: "i  j"
	      assume j_bound: "j < length tssb"
	      assume tssb_j: "tssb ! j = (pj, isj, θj, sbj, 𝒟j, 𝒪j,j)"
	      have "(𝒪sb  all_acquired sb')  (𝒪j  all_acquired sbj) = {}"
	      proof -
		{
		  fix a'
		  assume a'_in_i: "a'  (𝒪sb  all_acquired sb')"
		  assume a'_in_j: "a'  (𝒪j  all_acquired sbj)"
		  have False
		  proof -
		    from a'_in_i have "a'  (𝒪sb  all_acquired sb)  a'  A"
		      by (simp add: sb' all_acquired_append)
		    then show False
		    proof 
		      assume "a'  (𝒪sb  all_acquired sb)"
		      with ownership_distinct [OF i_bound j_bound neq_i_j tssb_i tssb_j] a'_in_j
		      show ?thesis
			by auto
		    next
		      assume "a'  A"
		      moreover
		      have j_bound': "j < length (map owned tssb)"
			using j_bound by auto
		      from A_unowned_by_others [rule_format, OF _ neq_i_j] tssb_j j_bound
		      obtain "A  acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j = {}" and
                             "A  all_shared (takeWhile (Not  is_volatile_Writesb) sbj) = {}"
			by (auto simp add: Let_def)
		      moreover
		      from A_unaquired_by_others [rule_format, OF _ neq_i_j] tssb_j j_bound
		      have "A  all_acquired sbj = {}"
			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 tssb_i this]
	  show ?thesis by (simp add: tssb' 𝒪sb' sb')
	qed
      qed

      have valid_hist': "valid_history program_step tssb'"
      proof -
	from valid_history [OF i_bound tssb_i]
	have "history_consistent θsb (hd_prog psb sb) sb".
	with valid_write_sops [OF i_bound tssb_i] 
	  valid_implies_valid_prog_hd [OF i_bound tssb_i valid]
	have "history_consistent θsb (hd_prog psb (sb@[Ghostsb A L R W])) 
	         (sb@ [Ghostsb A L R W])"
	  apply -
	  apply (rule history_consistent_appendI)
	  apply (auto simp add: hd_prog_append_Ghostsb)
	  done
	from valid_history_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' θsb')
      qed

      have valid_reads': "valid_reads msb tssb'"
      proof -
	from valid_reads [OF i_bound tssb_i]
	have "reads_consistent False 𝒪sb msb sb" .
	from reads_consistent_snoc_Ghostsb [OF this]
	have "reads_consistent False 𝒪sb msb (sb @ [Ghostsb A L R W])".
	from valid_reads_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' sb' 𝒪sb') 
      qed

      have valid_sharing': "valid_sharing 𝒮sb' tssb'"
      proof (intro_locales)	
	from outstanding_non_volatile_writes_unshared [OF i_bound tssb_i] 
	have "non_volatile_writes_unshared 𝒮sb (sb @ [Ghostsb 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 𝒮sb' tssb'"
	  by (simp add: tssb' sb' 𝒮sb')
      next
	from sharing_consis [OF i_bound tssb_i]
	have consis': "sharing_consistent 𝒮sb 𝒪sb sb".
	from  A_shared_owned
	have "A  dom (share ?drop_sb 𝒮)  acquired True sb 𝒪sb"
	  by (simp add:  sharing_consistent_append  acquired_takeWhile_non_volatile_Writesb)
	moreover have "dom (share ?drop_sb 𝒮)  dom 𝒮  dom (share sb 𝒮sb)"
	proof
	  fix a'
	  assume a'_in: "a'  dom (share ?drop_sb 𝒮)" 
	  from share_unshared_in [OF a'_in]
	  show "a'  dom 𝒮  dom (share sb 𝒮sb)"
	  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 𝒮sb)"
	      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 𝒮sb)  acquired True sb 𝒪sb"
	  by auto
        have "A  dom (share sb 𝒮sb)  acquired True sb 𝒪sb"
        proof -
          {
            fix x
            assume x_A: "x  A"
            have "x  dom (share sb 𝒮sb)  acquired True sb 𝒪sb"
            proof -
              {
                assume "x  dom 𝒮"
                
                from share_all_until_volatile_write_share_acquired [OF sharing_consis 𝒮sb tssb 
                  i_bound tssb_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 𝒮sb 𝒪sb (sb @ [Ghostsb A L R W])"
	  by (simp add:  sharing_consistent_append  acquired_takeWhile_non_volatile_Writesb)
	from sharing_consis_nth_update [OF i_bound this]
	show "sharing_consis 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒪sb' sb' 𝒮sb')

      next
	from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound tssb_i] ]
	show "read_only_unowned 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' 𝒪sb')
      next
	from unowned_shared_nth_update [OF i_bound tssb_i subset_refl]
	show "unowned_shared 𝒮sb' tssb'"
	  by (simp add: tssb' sb' 𝒪sb' 𝒮sb')
      next
	from no_outstanding_write_to_read_only_memory [OF i_bound tssb_i] 
	have "no_write_to_read_only_memory 𝒮sb (sb @ [Ghostsb 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 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' sb')
      qed

      have tmps_distinct': "tmps_distinct tssb'"
      proof (intro_locales)
	from load_tmps_distinct [OF i_bound tssb_i]
	have "distinct_load_tmps issb'" by (simp add: "issb")
	from load_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_distinct tssb'" by (simp add: tssb')
      next
	from read_tmps_distinct [OF i_bound tssb_i]
	have "distinct_read_tmps (sb @ [Ghostsb 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 tssb'" by (simp add: tssb' sb')
      next
	from load_tmps_read_tmps_distinct [OF i_bound tssb_i]
	have "load_tmps issb'  read_tmps (sb @ [Ghostsb A L R W]) ={}"
	  by (auto simp add: read_tmps_append "issb")
	from load_tmps_read_tmps_distinct_nth_update [OF i_bound this]
	show "load_tmps_read_tmps_distinct tssb'" by (simp add: tssb' sb')
      qed
      
      have valid_sops': "valid_sops tssb'"
      proof -
	from valid_store_sops [OF i_bound tssb_i]
	obtain 
	  valid_store_sops': "sopstore_sops issb'. valid_sop sop"
	  by (auto simp add: "issb")
	from valid_write_sops [OF i_bound tssb_i]
	have valid_write_sops': "sopwrite_sops (sb@ [Ghostsb 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: tssb' sb')
      qed

      have valid_dd': "valid_data_dependency tssb'"
      proof -
	from data_dependency_consistent_instrs [OF i_bound tssb_i]
	obtain 
	  dd_is: "data_dependency_consistent_instrs (dom θsb') issb'"
	  by (auto simp add: "issb" θsb')
	from load_tmps_write_tmps_distinct [OF i_bound tssb_i] 
	have "load_tmps issb'  (fst ` write_sops (sb@ [Ghostsb A L R W])) ={}"
	  by (auto simp add: write_sops_append "issb")
	from valid_data_dependency_nth_update [OF i_bound dd_is this]
	show ?thesis by (simp add: tssb' sb')
      qed

      have load_tmps_fresh': "load_tmps_fresh tssb'"
      proof -
	from load_tmps_fresh [OF i_bound tssb_i] 
	have "load_tmps issb'  dom θsb = {}"
	  by (auto simp add: "issb")
	from load_tmps_fresh_nth_update [OF i_bound this]
	show ?thesis by (simp add: tssb' θsb')
      qed

      have enough_flushs': "enough_flushs tssb'"
      proof -
	from clean_no_outstanding_volatile_Writesb [OF i_bound tssb_i]
	have "¬ 𝒟sb  outstanding_refs is_volatile_Writesb (sb@[Ghostsb 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: tssb' sb' 𝒟sb')
      qed
	

      have valid_program_history': "valid_program_history tssb'"
      proof -	
	from valid_program_history [OF i_bound tssb_i]
	have "causal_program_history issb sb" .
	then have causal': "causal_program_history issb' (sb@[Ghostsb A L R W])"
	  by (auto simp: causal_program_history_Ghost  "issb")
	from valid_last_prog [OF i_bound tssb_i]
	have "last_prog psb sb = psb".
	hence "last_prog psb (sb @ [Ghostsb A L R W]) = psb"
	  by (simp add: last_prog_append_Ghostsb)
	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb' sb')
      qed

      show ?thesis
      proof (cases "outstanding_refs is_volatile_Writesb sb = {}")
	case True

	from True have flush_all: "takeWhile (Not  is_volatile_Writesb) sb = sb"
	  by (auto simp add: outstanding_refs_conv)
	
	from True have suspend_nothing: "dropWhile (Not  is_volatile_Writesb) 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# issb'"
	  by (simp add: "issb")

	with suspends_empty ts_i 
	have ts_i: "ts!i = (psb, Ghost A L R W# issb',
	  θsb,(), 𝒟, acquired True ?take_sb 𝒪sb,release ?take_sb (dom 𝒮sb) sb)"
	  by simp

	from direct_memop_step.Ghost 	
	have "(Ghost A L R W# issb', 
	  θsb, (),m, 𝒟, acquired True ?take_sb 𝒪sb, 
               release ?take_sb (dom 𝒮sb) sb, 𝒮)  
               (issb', 
	  θsb, (), m, 𝒟, acquired True ?take_sb 𝒪sb  A - R, 
           augment_rels (dom 𝒮) R (release ?take_sb (dom 𝒮sb) sb),
           𝒮 ⊕⇘WR ⊖⇘AL)".
	from direct_computation.concurrent_step.Memop [OF i_bound' ts_i this]
	have "(ts, m, 𝒮) d 
              (ts[i := (psb, issb', 
	          θsb, (),𝒟, acquired True ?take_sb 𝒪sb  A - R,
                  augment_rels (dom 𝒮) R (release ?take_sb (dom 𝒮sb) sb))], 
	       m,𝒮 ⊕⇘WR ⊖⇘AL)".

	moreover

	from suspend_nothing
	have suspend_nothing': "(dropWhile (Not  is_volatile_Writesb) sb') = []"
	  by (simp add: sb')


	have all_shared_A: "j p is 𝒪  𝒟 θ sb. j < length tssb  i  j 
	  tssb ! j = (p, is, θ, sb, 𝒟, 𝒪,) 
	  all_shared (takeWhile (Not  is_volatile_Writesb) sb)  A = {}"
	proof -
	  {
	    fix j pj isj 𝒪j j 𝒟j θj sbj x
	    assume j_bound: "j < length tssb"
	    assume neq_i_j: "i  j"
	    assume jth: "tssb!j = (pj,isj, θj,sbj,𝒟j,𝒪j,j)"
	    assume x_shared: "x  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)"
	    assume x_A: "x  A"
	    have False
	    proof -
	      from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
	      have "all_shared sbj  all_acquired sbj  𝒪j".

	      moreover have "all_shared (takeWhile (Not  is_volatile_Writesb) sbj)  all_shared sbj"
		using all_shared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
		  "(dropWhile (Not  is_volatile_Writesb) sbj)"]
		by auto
	      moreover

	      from A_unaquired_by_others [rule_format, OF _ neq_i_j] jth j_bound
	      have "A  all_acquired sbj = {}" 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 tssb  i  j 
	  tssb ! j = (p, is, θ, sb, 𝒟, 𝒪,) 
	  all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {}"
	  using L_subset by blast

        have all_shared_A: "j p is 𝒪  𝒟 θ sb. j < length tssb  i  j 
            tssb ! j = (p, is, θ, sb, 𝒟, 𝒪,) 
            all_shared (takeWhile (Not  is_volatile_Writesb) sb)  A = {}"
        proof -
	  {
	    fix j pj isj 𝒪j j 𝒟j θj sbj x
	    assume j_bound: "j < length tssb"
	    assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
            assume neq_i_j: "i  j" 
	    assume x_shared: "x  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)"
	    assume x_A: "x  A"
	    have False
	    proof -
	      from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
	      have "all_shared sbj  all_acquired sbj  𝒪j".

	      moreover have "all_shared (takeWhile (Not  is_volatile_Writesb) sbj)  all_shared sbj"
		using all_shared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
		  "(dropWhile (Not  is_volatile_Writesb) sbj)"]
		by auto
	      moreover
	      from A_unaquired_by_others [rule_format, OF _ neq_i_j] jth j_bound
	      have "A  all_acquired sbj = {}" 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 tssb  i  j 
            tssb ! j = (p, is, θ, sb, 𝒟, 𝒪,) 
            all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {}"
	  using L_subset by blast

        have all_unshared_R: "j p is 𝒪  𝒟 θ sb. j < length tssb  i  j 
            tssb ! j = (p, is, θ, sb, 𝒟, 𝒪,) 
            all_unshared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}"
        proof -
	  {
	    fix j pj isj 𝒪j j 𝒟j θj sbj x
	    assume j_bound: "j < length tssb"
            assume neq_i_j: "i  j"
	    assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	    assume x_unshared: "x  all_unshared (takeWhile (Not  is_volatile_Writesb) sbj)"
	    assume x_R: "x  R"
	    have False
	    proof -
	      from unshared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
	      have "all_unshared sbj  all_acquired sbj  𝒪j".

	      moreover have "all_unshared (takeWhile (Not  is_volatile_Writesb) sbj)  all_unshared sbj"
		using all_unshared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
		  "(dropWhile (Not  is_volatile_Writesb) sbj)"]
		by auto
	      moreover

	      note ownership_distinct [OF i_bound j_bound neq_i_j tssb_i jth]

	      ultimately
	      show False
		using  R_acq x_R x_unshared acquired_all_acquired [of True sb 𝒪sb]
                by blast
	    qed
	  }
	  thus ?thesis by blast
        qed

        have all_acquired_R: "j p is 𝒪  𝒟 θ sb. j < length tssb  i  j 
            tssb ! j = (p, is, θ, sb, 𝒟, 𝒪,) 
            all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  R = {}"
        proof -
	  {
	    fix j pj isj 𝒪j j 𝒟j θj sbj x
	    assume j_bound: "j < length tssb"
	    assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
            assume neq_i_j: "i  j" 
	    assume x_acq: "x  all_acquired (takeWhile (Not  is_volatile_Writesb) sbj)"
	    assume x_R: "x  R"
            have False
	    proof -

	      from x_acq have "x  all_acquired sbj"
		using all_acquired_append [of "takeWhile (Not  is_volatile_Writesb) sbj" 
		  "dropWhile (Not  is_volatile_Writesb) sbj"]
		by auto
	      moreover
	      note ownership_distinct [OF i_bound j_bound neq_i_j tssb_i jth]
	      ultimately
	      show False
		using  R_acq x_R acquired_all_acquired [of True sb 𝒪sb]
		by blast
	    qed
	  }
	  thus ?thesis by blast
        qed

        have all_shared_R: "j p is 𝒪  𝒟 θ sb. j < length tssb  i  j  
            tssb ! j = (p, is, θ, sb, 𝒟, 𝒪,) 
            all_shared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}"
        proof -
	  {
	    fix j pj isj 𝒪j j 𝒟j θj sbj x
	    assume j_bound: "j < length tssb"
	    assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
            assume neq_i_j: "i  j" 
	    assume x_shared: "x  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)"
	    assume x_R: "x  R"
	    have False
	    proof -
	      from all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]]
	      have "all_shared sbj  all_acquired sbj  𝒪j".

	      moreover have "all_shared (takeWhile (Not  is_volatile_Writesb) sbj)  all_shared sbj"
		using all_shared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
		  "(dropWhile (Not  is_volatile_Writesb) sbj)"]
		by auto
	      moreover
	      note ownership_distinct [OF i_bound j_bound neq_i_j tssb_i jth]
	      ultimately 
	      show False
		using R_acq x_R x_shared acquired_all_acquired [of True sb 𝒪sb]
		by blast
	    qed
	  }
	  thus ?thesis by blast
        qed

	note share_commute = 
	  share_all_until_volatile_write_append_Ghostsb [OF True ownership_distinct tssb sharing_consis 𝒮sb tssb
	  i_bound tssb_i all_shared_L all_shared_A all_acquired_R all_unshared_R all_shared_R]
        
	from 𝒟
	have 𝒟': "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb (sb@[Ghostsb A L R W])   {})"
	  by (auto simp: outstanding_refs_append)


        have "a  R. (a  (dom (share sb 𝒮sb)) ) = (a  dom 𝒮)"
        proof -
          {
            fix a
            assume a_R: "a  R"
            have "(a  (dom (share sb 𝒮sb)) ) = (a  dom 𝒮)"
            proof -
              from a_R R_acq acquired_all_acquired [of True sb 𝒪sb]
              have "a  𝒪sb  all_acquired sb"
                by auto
              
              
              from  share_all_until_volatile_write_thread_local' [OF ownership_distinct_tssb sharing_consis_tssb i_bound tssb_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 𝒮sb) sb) =
           release (sb @ [Ghostsb A L R W]) (dom 𝒮sb') sb"
          by (clarsimp simp add: release_append 𝒮sb')

	have "(tssb',msb,𝒮sb')  
	   (ts[i := (psb,issb', 
	       θsb,(), 𝒟, acquired True ?take_sb 𝒪sb  A - R,
                augment_rels (dom 𝒮) R (release ?take_sb (dom 𝒮sb) sb))], 
                 m,𝒮 ⊕⇘WR ⊖⇘AL)"
	  apply (rule sim_config.intros) 
	  apply    (simp add: m tssb' 𝒪sb' sb' θsb' flush_all_until_volatile_write_append_Ghost_commute [OF i_bound tssb_i])
	  apply   (clarsimp simp add: 𝒮 𝒮sb' tssb' sb' 𝒪sb' θsb' share_commute)
	  using  leq
	  apply  (simp add: tssb')
	  using i_bound i_bound' ts_sim ts_i True 𝒟'
	  apply (clarsimp simp add: Let_def nth_list_update 
	    outstanding_refs_conv tssb' 𝒪sb' sb' 𝒮sb' θsb' sb'  𝒟sb' 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' msb' 𝒮sb' sb'
	  by auto
      next
	case False

	then obtain r where r_in: "r  set sb" and volatile_r: "is_volatile_Writesb r"
	  by (auto simp add: outstanding_refs_conv)
	from takeWhile_dropWhile_real_prefix 
	[OF r_in, of  "(Not  is_volatile_Writesb)", simplified, OF volatile_r] 
	obtain a' v' sb'' A'' L'' R'' W'' sop' where
	  sb_split: "sb = takeWhile (Not  is_volatile_Writesb) sb @ Writesb True a' sop' v' A'' L'' R'' W''# sb''" 
	  and
	  drop: "dropWhile (Not  is_volatile_Writesb) sb = Writesb 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 = Writesb True a' sop' v' A'' L'' R'' W''# sb''"
	  by simp

	have "(ts, m, 𝒮) d* (ts, m, 𝒮)" by auto
	moreover

	have "Writesb 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_Writesb)", simplified]

	from takeWhile_append1 [where P="Not  is_volatile_Writesb", OF r_in] volatile_r
	have takeWhile_app: 
	  "(takeWhile (Not  is_volatile_Writesb) (sb @ [Ghostsb A L R W])) = (takeWhile (Not  is_volatile_Writesb) sb)"
	  by simp

	note share_commute = share_all_until_volatile_write_append_Ghostsb' [OF False i_bound tssb_i]
	
	from 𝒟
	have 𝒟': "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb (sb@[Ghostsb A L R W])   {})"
	  by (auto simp: outstanding_refs_append)


	have "(tssb',msb,𝒮sb')  (ts,m,𝒮)"
	  apply (rule sim_config.intros) 
	  apply    (simp add: m flush_all_until_volatile_write_append_Ghost_commute [OF i_bound tssb_i] tssb' 𝒪sb' θsb' sb')
	  apply   (clarsimp simp add: 𝒮 𝒮sb' tssb' sb' 𝒪sb' θsb' share_commute)
	  using  leq
	  apply  (simp add: tssb')
	  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_Ghostsb instrs_append_Ghostsb hd_prog_append_Ghostsb
	    drop "issb" tssb' sb' 𝒪sb' sb' 𝒮sb' θsb' 𝒟sb' 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' msb' 𝒮sb' 
	  by (auto simp del: fun_upd_apply )
      qed	
    qed
  next
    case (StoreBuffer i psb "issb" θsb sb 𝒟sb 𝒪sb  sb sb' 𝒪sb' sb')
    then obtain 
      
      tssb': "tssb' = tssb[i := (psb, issb, θsb, sb', 𝒟sb, 𝒪sb',sb')]" and
      i_bound: "i < length tssb" and
      tssb_i: "tssb ! i = (psb, issb, θsb,sb, 𝒟sb, 𝒪sb,sb)" and
      flush: "(msb,sb,𝒪sb,sb,𝒮sb) f 
              (msb',sb',𝒪sb',sb',𝒮sb')" 
      by auto

    from sim obtain 
      m: "m = flush_all_until_volatile_write tssb msb" and
      𝒮: "𝒮 = share_all_until_volatile_write tssb 𝒮sb" and
      leq: "length tssb = length ts" and
      ts_sim: "i<length tssb.
           let (p, issb, θ, sb,𝒟sb, 𝒪sb,) = tssb ! i;
               suspends = dropWhile (Not  is_volatile_Writesb) sb
           in  is 𝒟. instrs suspends @ issb = is @ prog_instrs suspends  
                          𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb sb  {}) 
               ts ! i =
                   (hd_prog p suspends, 
                    is,
                    θ |` (dom θ - read_tmps suspends), (),
                    𝒟, 
                    acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪sb,
                    release (takeWhile (Not  is_volatile_Writesb) sb) (dom 𝒮sb) )"
      by cases blast

    from i_bound leq have i_bound': "i < length ts"
      by auto


    have split_sb: "sb = takeWhile (Not  is_volatile_Writesb) sb @ dropWhile (Not  is_volatile_Writesb) sb"
      (is "sb = ?take_sb@?drop_sb")
      by simp

    from ts_sim [rule_format, OF i_bound] tssb_i obtain suspends "is" 𝒟 where
      suspends: "suspends = dropWhile (Not  is_volatile_Writesb) sb" and
      is_sim: "instrs suspends @ issb = is @ prog_instrs suspends" and
      𝒟: "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb sb  {})" and
      ts_i: "ts ! i =
          (hd_prog psb suspends, is,
           θsb |` (dom θsb - read_tmps suspends), (),𝒟, acquired True ?take_sb 𝒪sb,
           release ?take_sb (dom 𝒮sb) sb)"
      by (auto simp add: Let_def)

    from flush_step_preserves_valid [OF i_bound tssb_i flush valid]
    have valid': "valid tssb'" 
      by (simp add: tssb')

    from flush obtain r where sb: "sb=r#sb'"
      by (cases) auto

    from valid_history [OF i_bound tssb_i]
    have "history_consistent θsb (hd_prog psb sb) sb".
    then
    have hist_consis': "history_consistent θsb (hd_prog psb 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 tssb'" by (simp add: tssb')

    from read_tmps_distinct [OF i_bound tssb_i]
    have dist_sb': "distinct_read_tmps sb'"
      by (simp add: sb split: memref.splits)

    have tmps_distinct': "tmps_distinct tssb'"
    proof (intro_locales)
      from load_tmps_distinct [OF i_bound tssb_i]
      have "distinct_load_tmps issb".
	
      from load_tmps_distinct_nth_update [OF i_bound this]
      show "load_tmps_distinct tssb'"
	by (simp add: tssb')
    next
      from read_tmps_distinct_nth_update [OF i_bound dist_sb']
      show "read_tmps_distinct tssb'"
	by (simp add: tssb')
    next
      from load_tmps_read_tmps_distinct [OF i_bound tssb_i]
      have "load_tmps issb  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 tssb'" by (simp add: tssb')
    qed

    from load_tmps_write_tmps_distinct [OF i_bound tssb_i]
    have "load_tmps issb  (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 tssb_i] this]
    have valid_dd': "valid_data_dependency tssb'"
      by (simp add: tssb')

    from valid_store_sops [OF i_bound tssb_i] valid_write_sops [OF i_bound tssb_i] 
    valid_sops_nth_update [OF i_bound]
    have valid_sops': "valid_sops tssb'"
      by (cases r) (auto simp add: sb tssb')
    
    have load_tmps_fresh': "load_tmps_fresh tssb'"
    proof -
      from load_tmps_fresh [OF i_bound tssb_i] 
      have "load_tmps issb  dom θsb = {}".
      from load_tmps_fresh_nth_update [OF i_bound this]
      show ?thesis by (simp add: tssb')
    qed

    have enough_flushs': "enough_flushs tssb'"
    proof -
      from clean_no_outstanding_volatile_Writesb [OF i_bound tssb_i]
      have "¬ 𝒟sb  outstanding_refs is_volatile_Writesb sb' = {}"
	by (auto simp add: sb split: if_split_asm)
      from enough_flushs_nth_update [OF i_bound this]
      show ?thesis
	by (simp add: tssb' sb)
    qed

    show ?thesis
    proof (cases r)
      case (Writesb volatile a sop v A L R W)
      from flush this
      have msb': "msb' = (msb(a := v))"
	by cases (auto simp add: sb)

      have non_volatile_owned: "¬ volatile  a  𝒪sb"
      proof (cases volatile)
	case True thus ?thesis by simp
      next
	case False
	with outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i] 
	have "a  𝒪sb"
	  by (simp add: sb Writesb)
	thus ?thesis by simp
      qed

      have a_unowned_by_others:
	"j < length tssb. i  j  (let (_,_,_,sbj,_,𝒪j,_) = tssb ! j in 
	a  𝒪j  all_acquired sbj)"
      proof (unfold Let_def, clarify del: notI)
	fix j pj "isj" 𝒪j j 𝒟j θj sbj
	assume j_bound: "j < length tssb"
	assume neq: "i  j"
	assume ts_j: "tssb ! j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	show "a  𝒪j  all_acquired sbj"
	proof (cases volatile)
	  case True
	  from outstanding_volatile_writes_unowned_by_others [OF i_bound j_bound neq 
           tssb_i ts_j] 
	  show ?thesis 
	    by (simp add: sb Writesb True)
	next
	  case False
	  with non_volatile_owned
	  have "a  𝒪sb"
	    by simp
	  with ownership_distinct [OF i_bound j_bound neq tssb_i ts_j]
	  show ?thesis
	    by blast
	qed
      qed

      from valid_reads [OF i_bound tssb_i]
      have reads_consis: "reads_consistent False 𝒪sb msb sb" .


      {
	fix j 
	fix pj issbj 𝒪j j 𝒟sbj θj sbj
	assume j_bound: "j < length tssb"
	assume tssb_j: "tssb!j=(pj,issbj,θj,sbj,𝒟sbj,𝒪j,j)"
	assume neq_i_j: "ij"
	have "a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	proof 
	  assume "a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	  hence "a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	    by (simp add: outstanding_refs_is_non_volatile_Writesb_takeWhile_conv)
	  hence "a  outstanding_refs is_non_volatile_Writesb sbj"
	    using outstanding_refs_append [of _ "(takeWhile (Not  is_volatile_Writesb) sbj)" 
	      "(dropWhile (Not  is_volatile_Writesb) sbj)"]
	    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 tssb_j]]
	  have "a  𝒪j  all_acquired sbj"
	    by auto
	  with a_unowned_by_others [rule_format, OF j_bound neq_i_j]  tssb_j
	  show False
	    by auto
	qed
      }
      note a_notin_others = this

	
      from a_notin_others
      have a_notin_others': 
	"j < length tssb. i  j 
        (let (_,_,_,sbj,_,_,_) = tssb!j in a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))"
	by (fastforce simp add: Let_def)
      

      
      obtain D f where sop: "sop=(D,f)" by (cases sop) auto
      from valid_history [OF i_bound tssb_i] sop sb Writesb
      obtain D_tmps: "D  dom θsb" and f_v: "f θsb = v" and
	 D_sb': "D  read_tmps sb' = {}"
	by auto
      let  = "(θsb |` (dom θsb - read_tmps sb'))"
      from D_tmps D_sb'
      have D_tmps': "D  dom "
	by auto
      from valid_write_sops [OF i_bound tssb_i, rule_format, of sop]
      have "valid_sop sop"
	by (auto simp add: sb Writesb)
      from this [simplified sop]
      interpret valid_sop "(D,f)" .
      from D_tmps D_sb' 
      have "((dom θsb - 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 tssb'"
      proof -	
	from valid_program_history [OF i_bound tssb_i]
	have "causal_program_history issb sb" .
	then have causal': "causal_program_history issb sb'"
	  by (simp add: sb Writesb causal_program_history_def)

	from valid_last_prog [OF i_bound tssb_i]
	have "last_prog psb sb = psb".
	hence "last_prog psb sb' = psb"
	  by (simp add: sb Writesb)

	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb')
      qed



      show ?thesis
      proof (cases volatile)
	case True
	note volatile = this
	from flush Writesb volatile
	obtain 
	  𝒪sb': "𝒪sb'=𝒪sb  A - R" and
	  𝒮sb': "𝒮sb'= 𝒮sb ⊕⇘WR ⊖⇘AL" and
          sb': "sb' = Map.empty"
	  by cases (auto  simp add: sb)

	from sharing_consis [OF i_bound tssb_i] 
	obtain 
	  A_shared_owned: "A  dom 𝒮sb  𝒪sb" and
	  L_subset: "L  A" and
	  A_R: "A  R = {}" and
	  R_owned: "R  𝒪sb"
	  by (clarsimp simp add: sb Writesb volatile)


	 
	from sb Writesb True have take_empty: "takeWhile (Not  is_volatile_Writesb) sb = []"
	  by (auto simp add: outstanding_refs_conv)
	
	from sb Writesb True have suspend_all: "dropWhile (Not  is_volatile_Writesb) 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' @ issb = is @ prog_instrs sb'"
	  by (simp add: True Writesb suspends_all sb sop)

	from valid_program_history [OF i_bound tssb_i]
	interpret causal_program_history "issb" sb .
	from valid_last_prog [OF i_bound tssb_i]
	have last_prog: "last_prog psb sb = psb".

	from causal_program_history [of "[Writesb 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'@issb = is' @ prog_instrs sb'" 
	  by (auto simp add: sb Writesb volatile sop)
	  
	from causal_program_history have
	  causal_program_history_sb': "causal_program_history issb sb'"
	  apply -
	  apply (rule causal_program_history.intro)
	  apply (auto simp add: sb Writesb)
	  done

	from ts_i have ts_i: "ts ! i =
          (hd_prog psb sb', Write True a (D, f) A L R W# is',  , (), 𝒟,acquired True ?take_sb 𝒪sb,
           release ?take_sb (dom 𝒮sb) sb)"	
	  by (simp add: suspends_all sb Writesb "is")

	let ?ts' = "ts[i := (hd_prog psb sb', is', , (), True, acquired True ?take_sb 𝒪sb  A - R,
                       Map.empty)]"

	from i_bound' have ts'_i: "?ts'!i = (hd_prog psb sb', is', , (),True, acquired True ?take_sb 𝒪sb  A - R,Map.empty)" 
	  by simp

	from no_outstanding_write_to_read_only_memory [OF i_bound tssb_i]
	have a_not_ro: "a  read_only 𝒮sb"
	  by (clarsimp simp add: sb Writesb volatile)

	{
	  fix j 
	  fix pj issbj 𝒪j j 𝒟sbj θj sbj
	  assume j_bound: "j < length tssb"
	  assume tssb_j: "tssb!j=(pj,issbj,θj,sbj,𝒟sbj,𝒪j,j)"
	  assume neq_i_j: "ij"
	  have "a  unforwarded_non_volatile_reads (dropWhile (Not  is_volatile_Writesb) sbj) {}"
	  proof 
	    let ?take_sbj = "takeWhile (Not  is_volatile_Writesb) sbj"
	    let ?drop_sbj = "dropWhile (Not  is_volatile_Writesb) sbj"
	    assume a_in: "a   unforwarded_non_volatile_reads ?drop_sbj {}"
	    
	    from a_unowned_by_others [rule_format, OF j_bound neq_i_j] tssb_j 
	    obtain a_unowned: "a  𝒪j" and a_unacq: "a  all_acquired sbj"
	      by auto
	    with all_acquired_append [of ?take_sbj ?drop_sbj] acquired_takeWhile_non_volatile_Writesb [of sbj 𝒪j]
	    have a_unacq_take: "a  acquired True ?take_sbj 𝒪j"
	      by (auto simp add: )

	    note nvo_j = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound tssb_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_sbj 𝒮sb)
	      (acquired True ?take_sbj 𝒪j) ?drop_sbj" .

	    note consis_j = sharing_consis [OF j_bound tssb_j]
	    with sharing_consistent_append [of 𝒮sb 𝒪j ?take_sbj ?drop_sbj]
	    obtain consis_take_j: "sharing_consistent 𝒮sb 𝒪j ?take_sbj" and
	      consis_drop_j: "sharing_consistent (share ?take_sbj 𝒮sb)
	      (acquired True ?take_sbj 𝒪j) ?drop_sbj"
	      by auto

	    from in_unforwarded_non_volatile_reads_non_volatile_Readsb [OF a_in]
	    have a_in': "a  outstanding_refs is_non_volatile_Readsb ?drop_sbj".
	    
	    note reads_consis_j = valid_reads [OF j_bound tssb_j]
	    from reads_consistent_drop [OF this]
	    have reads_consis_drop_j:
	      "reads_consistent True (acquired True ?take_sbj 𝒪j) (flush ?take_sbj msb) ?drop_sbj".
	    

            
            from read_only_share_all_shared [of a ?take_sbj 𝒮sb] a_not_ro 
              all_shared_acquired_or_owned [OF consis_take_j]
              all_acquired_append [of ?take_sbj ?drop_sbj] a_unowned a_unacq
	    have a_not_ro_j: "a  read_only (share ?take_sbj 𝒮sb)"
              by auto
	    
	  
	    from ts_sim [rule_format, OF j_bound] tssb_j j_bound
	    obtain suspendsj "isj" 𝒟j j where
	      suspendsj: "suspendsj = ?drop_sbj" and
	      isj: "instrs suspendsj @ issbj = isj @ prog_instrs suspendsj" and
	      𝒟j: "𝒟sbj = (𝒟j  outstanding_refs is_volatile_Writesb sbj  {})" and
	      tsj: "ts!j = (hd_prog pj suspendsj, isj, 
	      θj |` (dom θj - read_tmps suspendsj),(),   
	      𝒟j, acquired True ?take_sbj 𝒪j,j)"
	      by (auto simp: Let_def)

	    from valid_last_prog [OF j_bound tssb_j] have last_prog: "last_prog pj sbj = pj".
	    

	    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_sbj = ys @ Writesb True a' sop v A L R W # zs  a  A  
		a  outstanding_refs is_Writesb ys  a'a"
	      with suspendsj 
	      obtain a' sop' v' ys zs' A' L' R' W' where
		split_suspendsj: "suspendsj = ys @ Writesb True a' sop' v' A' L' R' W'# zs'" (is "suspendsj=?suspends") and
		a_A': "a  A'" and
		no_write: "a  outstanding_refs is_Writesb (ys @ [Writesb True a' sop' v' A' L' R' W'])"
		by (auto simp add: outstanding_refs_append)

	      from last_prog
	      have lp: "last_prog pj suspendsj = pj"
		apply -
		apply (rule last_prog_same_append [where sb="?take_sbj"])
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done

	      from sharing_consis [OF j_bound tssb_j]
	      have sharing_consis_j: "sharing_consistent 𝒮sb 𝒪j sbj".
	      then have A'_R': "A'  R' = {}" 
		by (simp add: sharing_consistent_append [of _ _ ?take_sbj ?drop_sbj, simplified] 
		  suspendsj [symmetric] split_suspendsj sharing_consistent_append)	  
	      
	      from valid_program_history [OF j_bound tssb_j] 
	      have "causal_program_history issbj sbj".
	      then have cph: "causal_program_history issbj ?suspends"
		apply -
		apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply (simp add: split_suspendsj)
		done
	      
	      from valid_reads [OF j_bound tssb_j]
	      have reads_consis_j: "reads_consistent False 𝒪j msb sbj".
	      
	      from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb 
		j_bound tssb_j this]
	      have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		by (simp add: m suspendsj)
	    
	      hence reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  
		m (ys@[Writesb True a' sop' v' A' L' R' W'])"
		by (simp add: split_suspendsj reads_consistent_append)

	      from valid_write_sops [OF j_bound tssb_j]
	      have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		valid_sops_drop: "sopwrite_sops (ys@[Writesb 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 tssb_j]
	      have "distinct_read_tmps (?take_sbj@suspendsj)"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain 
		read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
		distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		apply (simp only: distinct_read_tmps_append)
		done
	    
	      from valid_history [OF j_bound tssb_j]
	      have h_consis: 
		"history_consistent θj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done
	      
	      have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	      proof -
		from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		  by simp
		from last_prog_hd_prog_append' [OF h_consis] this
		have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		  by (simp only: split_suspendsj [symmetric] suspendsj) 
		moreover 
		have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		  last_prog (hd_prog pj suspendsj) ?take_sbj"
		  apply (simp only: split_suspendsj [symmetric] suspendsj) 
		  by (rule last_prog_hd_prog_append)
		ultimately show ?thesis
		  by (simp add: split_suspendsj [symmetric] suspendsj) 
	      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 pj suspendsj) suspendsj"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      from reads_consistent_drop_volatile_writes_no_volatile_reads  
	      [OF reads_consis_j] 
	      have no_vol_read: "outstanding_refs is_volatile_Readsb 
		(ys@[Writesb True a' sop' v' A' L' R' W']) = {}"
		by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		split_suspendsj )
	    
	      have acq_simp:
		"acquired True (ys @ [Writesb True a' sop' v' A' L' R' W']) 
		(acquired True ?take_sbj 𝒪j) = 
		acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R'"
		by (simp add: acquired_append)

	      from flush_store_buffer_append [where sb="ys@[Writesb True a' sop' v' A' L' R' W']" and sb'="zs'", simplified,
	      OF j_bound_ts' isj [simplified split_suspendsj] cph [simplified suspendsj]  tsj [simplified split_suspendsj]
	      refl lp [simplified split_suspendsj] reads_consis_ys 	      
	      hist_consis' [simplified split_suspendsj] valid_sops_drop 
	      distinct_read_tmps_drop [simplified split_suspendsj] 
	      no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
	      𝒮="𝒮"]
	    
	      obtain isj' j' where
		isj': "instrs zs' @ issbj = isj' @ prog_instrs zs'" and
		steps_ys: "(ts, m, 𝒮)  d* 
		  (ts[j:=(last_prog
                              (hd_prog pj (Writesb True a' sop' v' A' L' R' W'# zs')) (ys@[Writesb True a' sop' v' A' L' R' W']),
                             isj',
                             θj |` (dom θj - read_tmps zs'),
                              (), True, acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R',j')],
                    flush (ys@[Writesb True a' sop' v' A' L' R' W']) m,
                    share (ys@[Writesb 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 psb sb', Write True a (D, f) A L R W# is', , (), 𝒟, 
		acquired True ?take_sb 𝒪sb,release ?take_sb (dom 𝒮sb) sb)"
		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. ij  (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_sbj = ys @ Ghostsb A L R W# zs  a  A  a  outstanding_refs is_Writesb ys"
	      with suspendsj 
	      obtain ys zs' A' L' R' W' where
		split_suspendsj: "suspendsj = ys @ Ghostsb A' L' R' W'# zs'" (is "suspendsj=?suspends") and
		a_A': "a  A'" and
		no_write: "a  outstanding_refs is_Writesb (ys @ [Ghostsb A' L' R' W'])"
		by (auto simp add: outstanding_refs_append)

	      from last_prog
	      have lp: "last_prog pj suspendsj = pj"
		apply -
		apply (rule last_prog_same_append [where sb="?take_sbj"])
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done


	      from valid_program_history [OF j_bound tssb_j] 
	      have "causal_program_history issbj sbj".
	      then have cph: "causal_program_history issbj ?suspends"
		apply -
		apply (rule causal_program_history_suffix [where sb="?take_sbj"] )
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply (simp add: split_suspendsj)
		done
	    
	      from valid_reads [OF j_bound tssb_j]
	      have reads_consis_j: "reads_consistent False 𝒪j msb sbj".
	    
	      from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb 
		j_bound tssb_j this]
	      have reads_consis_m_j: "reads_consistent True (acquired True ?take_sbj 𝒪j) m suspendsj"
		by (simp add: m suspendsj)
	    
	    
	      hence reads_consis_ys: "reads_consistent True (acquired True ?take_sbj 𝒪j)  
		m (ys@[Ghostsb A' L' R' W'])"
		by (simp add: split_suspendsj reads_consistent_append)

	      from valid_write_sops [OF j_bound tssb_j]
	      have "sopwrite_sops (?take_sbj@?suspends). valid_sop sop"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain valid_sops_take: "sopwrite_sops ?take_sbj. valid_sop sop" and
		valid_sops_drop: "sopwrite_sops (ys@[Ghostsb A' L' R' W']). valid_sop sop"
		apply (simp only: write_sops_append)
		apply auto
		done
	      
	      from read_tmps_distinct [OF j_bound tssb_j]
	      have "distinct_read_tmps (?take_sbj@suspendsj)"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      then obtain 
		read_tmps_take_drop: "read_tmps ?take_sbj  read_tmps suspendsj = {}" and
		distinct_read_tmps_drop: "distinct_read_tmps suspendsj"
		apply (simp only: split_suspendsj [symmetric] suspendsj) 
		apply (simp only: distinct_read_tmps_append)
		done
	    
	      from valid_history [OF j_bound tssb_j]
	      have h_consis: 
		"history_consistent θj (hd_prog pj (?take_sbj@suspendsj)) (?take_sbj@suspendsj)"
		apply (simp only: split_suspendsj [symmetric] suspendsj)
		apply simp
		done
	    
	      from sharing_consis [OF j_bound tssb_j]
	      have sharing_consis_j: "sharing_consistent 𝒮sb 𝒪j sbj".
	      then have A'_R': "A'  R' = {}" 
		by (simp add: sharing_consistent_append [of _ _ ?take_sbj ?drop_sbj, simplified] 
		  suspendsj [symmetric] split_suspendsj sharing_consistent_append)	  

	      have last_prog_hd_prog: "last_prog (hd_prog pj sbj) ?take_sbj = (hd_prog pj suspendsj)"
	      proof -
		from last_prog have "last_prog pj (?take_sbj@?drop_sbj) = pj"
		  by simp
		from last_prog_hd_prog_append' [OF h_consis] this
		have "last_prog (hd_prog pj suspendsj) ?take_sbj = hd_prog pj suspendsj"
		  by (simp only: split_suspendsj [symmetric] suspendsj) 
		moreover 
		have "last_prog (hd_prog pj (?take_sbj @ suspendsj)) ?take_sbj = 
		  last_prog (hd_prog pj suspendsj) ?take_sbj"
		  apply (simp only: split_suspendsj [symmetric] suspendsj) 
		  by (rule last_prog_hd_prog_append)
		ultimately show ?thesis
		  by (simp add: split_suspendsj [symmetric] suspendsj) 
	      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 pj suspendsj) suspendsj"
		by (simp add: split_suspendsj [symmetric] suspendsj)
	      from reads_consistent_drop_volatile_writes_no_volatile_reads  
	      [OF reads_consis_j] 
	      have no_vol_read: "outstanding_refs is_volatile_Readsb 
	      (ys@[Ghostsb A' L' R' W']) = {}"
		by (auto simp add: outstanding_refs_append suspendsj [symmetric] 
		  split_suspendsj )
	    
	      have acq_simp:
		"acquired True (ys @ [Ghostsb A' L' R' W']) 
		(acquired True ?take_sbj 𝒪j) = 
		acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R'"
		by (simp add: acquired_append)
	    
	      from flush_store_buffer_append [where sb="ys@[Ghostsb A' L' R' W']" and sb'="zs'", simplified,
		OF j_bound_ts' isj [simplified split_suspendsj] cph [simplified suspendsj]
		tsj [simplified split_suspendsj] refl lp [simplified split_suspendsj] reads_consis_ys 
		hist_consis' [simplified split_suspendsj] valid_sops_drop 
		distinct_read_tmps_drop [simplified split_suspendsj] 
		no_volatile_Readsb_volatile_reads_consistent [OF no_vol_read], where
		𝒮="𝒮"]
	      
	      obtain isj' j' where
		isj': "instrs zs' @ issbj = isj' @ prog_instrs zs'" and
		steps_ys: "(ts, m,𝒮)  d* 
		  (ts[j:=(last_prog
                              (hd_prog pj (Ghostsb A' L' R' W'# zs')) (ys@[Ghostsb A' L' R' W']),
                             isj',
                             θj |` (dom θj - read_tmps zs'),
                              (), 
	                     𝒟j  outstanding_refs is_volatile_Writesb ys  {}, acquired True ys (acquired True ?take_sbj 𝒪j)  A' - R',j')],
                    flush (ys@[Ghostsb A' L' R' W']) m,                    share (ys@[Ghostsb 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 psb sb', Write True a (D, f) A L R W# is', , (), 𝒟, 
		acquired True ?take_sb 𝒪sb,release ?take_sb (dom 𝒮sb) sb)"
		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. ij  (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 msb' tssb'"
	proof (unfold_locales)
	  fix j pj "isj" 𝒪j j 𝒟j θj sbj
	  assume j_bound: "j < length tssb'"
	  assume ts_j: "tssb'!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	  show "reads_consistent False 𝒪j msb' sbj"
	  proof (cases "i=j")
	    case True
	    from reads_consis ts_j j_bound sb show ?thesis
	      by (clarsimp simp add: True  msb' Writesb tssb' 𝒪sb' volatile reads_consistent_pending_write_antimono)
	  next
	    case False
	    from j_bound have j_bound':  "j < length tssb"
	      by (simp add: tssb')
	    moreover from ts_j False have ts_j': "tssb ! j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	      using j_bound by (simp add: tssb')
	    ultimately have consis_m: "reads_consistent False 𝒪j msb sbj"
	      by (rule valid_reads)
	    from a_unowned_by_others [rule_format, OF j_bound' False] ts_j'
	    have a_unowned:"a  𝒪j  all_acquired sbj"
	      by simp

	    let ?take_sbj = "takeWhile (Not  is_volatile_Writesb) sbj"
	    let ?drop_sbj = "dropWhile (Not  is_volatile_Writesb) sbj"

	    from a_unowned acquired_reads_all_acquired [of True ?take_sbj 𝒪j]
	    all_acquired_append [of ?take_sbj ?drop_sbj]
	    have a_not_acq_reads: "a  acquired_reads True ?take_sbj 𝒪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_sbj {}  acquired_reads True ?take_sbj 𝒪j" and
		m'= "(msb(a:=v))", OF _ _ _ consis_m]
	      by (fastforce simp add: msb')
	  qed
	qed
       

	have valid_own': "valid_ownership 𝒮sb' tssb'"
	proof (intro_locales)
	  show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	  proof 
	    fix j isj 𝒪j j 𝒟j θj sbj pj
	    assume j_bound: "j < length tssb'"
	    assume tssb'_j: "tssb'!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	    show "non_volatile_owned_or_read_only False 𝒮sb' 𝒪j sbj"
	    proof (cases "j=i")
	      case True
	      from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i]
	      have "non_volatile_owned_or_read_only False 
	        (𝒮sb ⊕⇘WR ⊖⇘AL) (𝒪sb  A - R) sb'"
		by (auto simp add: sb Writesb volatile non_volatile_owned_or_read_only_pending_write_antimono)
	      then show ?thesis
		using True i_bound tssb'_j
		by (auto simp add: tssb' 𝒮sb' sb 𝒪sb')
	    next
	      case False
	      from j_bound have j_bound': "j < length tssb"
		by (auto simp add: tssb')
	      with tssb'_j False i_bound 
	      have tssb_j: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
		by (auto simp add: tssb')
	      
	      
	      note nvo = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' tssb_j]

	      from read_only_unowned [OF i_bound tssb_i] R_owned
	      have "R  read_only 𝒮sb = {}"
		by auto
	      with read_only_reads_unowned [OF j_bound' i_bound False tssb_j tssb_i] L_subset
	      have "aread_only_reads
	      (acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j)
		(dropWhile (Not  is_volatile_Writesb) sbj).
		a  read_only 𝒮sb  a  read_only (𝒮sb ⊕⇘WR ⊖⇘AL)"
		by (auto simp add: in_read_only_convs sb Writesb volatile)
	      from non_volatile_owned_or_read_only_read_only_reads_eq' [OF nvo this]
	      have "non_volatile_owned_or_read_only False (𝒮sb ⊕⇘WR ⊖⇘AL) 𝒪j sbj".
	      thus ?thesis by (simp add: 𝒮sb')
	    qed
	  qed
	next
	  show "outstanding_volatile_writes_unowned_by_others tssb'"
	  proof (unfold_locales)
	    fix i1 j p1 "is1" 𝒪1 1 𝒟1 xs1 sb1 pj "isj" "𝒪j" j 𝒟j xsj sbj
	    assume i1_bound: "i1 < length tssb'"
	    assume j_bound: "j < length tssb'"
	    assume i1_j: "i1  j"
	    assume ts_i1: "tssb'!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	    assume ts_j: "tssb'!j = (pj,isj, xsj,sbj,𝒟j,𝒪j,j)"
	    show "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb1 = {}"
	    proof (cases "i1=i")
	      case True
	      from i1_j True have neq_i_j: "ij"
		by auto
	      from j_bound have j_bound': "j < length tssb"
		by (simp add: tssb')
	      from ts_j neq_i_j have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (simp add: tssb')
	      from outstanding_volatile_writes_unowned_by_others [OF i_bound j_bound' neq_i_j
		tssb_i ts_j'] ts_i1 i_bound tssb_i True show ?thesis
		by (clarsimp simp add: tssb' sb Writesb volatile)
	    next
	      case False
	      note i1_i = this
	      from i1_bound have i1_bound': "i1 < length tssb"
		by (simp add: tssb' sb)
	      hence i1_bound'': "i1 < length (map owned tssb)"
		by auto
	      from ts_i1 False have ts_i1': "tssb!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
		by (simp add: tssb' sb)
	      show ?thesis
	      proof (cases "j=i")
		case True
		from outstanding_volatile_writes_unowned_by_others [OF i1_bound' i_bound  i1_i  ts_i1' tssb_i ]
		have "(𝒪sb  all_acquired sb)  outstanding_refs is_volatile_Writesb sb1 = {}".
		then show ?thesis
		  using True i1_i ts_j tssb_i i_bound
		  by (auto simp add: sb Writesb volatile tssb' 𝒪sb')
	      next
		case False
		from j_bound have j_bound': "j < length tssb"
		  by (simp add: tssb')
		from ts_j False have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		  by (simp add: tssb')
		from outstanding_volatile_writes_unowned_by_others 
		[OF i1_bound' j_bound' i1_j ts_i1' ts_j']
		show "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb1 = {}" .
	      qed
	    qed
	  qed
	next
	  show "read_only_reads_unowned tssb'"
	  proof 
	    fix n m
	    fix pn "isn" 𝒪n n 𝒟n θn sbn pm "ism" 𝒪m m 𝒟m θm sbm
	    assume n_bound: "n < length tssb'"
	      and m_bound: "m < length tssb'"
	      and neq_n_m: "nm"
	      and nth: "tssb'!n = (pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
	      and mth: "tssb'!m =(pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
	    from n_bound have n_bound': "n < length tssb" by (simp add: tssb')
	    from m_bound have m_bound': "m < length tssb" by (simp add: tssb')
	    show "(𝒪m  all_acquired sbm) 
              read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sbn) 𝒪n)
              (dropWhile (Not  is_volatile_Writesb) sbn) =
              {}"
	    proof (cases "m=i")
	      case True
	      with neq_n_m have neq_n_i: "ni"
		by auto
	      
	      with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
		by (auto simp add: tssb')
	      note read_only_reads_unowned [OF n_bound' i_bound  neq_n_i nth' tssb_i]
	      then
	      show ?thesis
		using True tssb_i neq_n_i nth mth n_bound' m_bound' L_subset
		by (auto simp add: tssb' 𝒪sb' sb Writesb volatile)
	    next
	      case False
	      note neq_m_i = this
	      with m_bound mth i_bound have mth': "tssb!m = (pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
		by (auto simp add: tssb')
	      show ?thesis
	      proof (cases "n=i")
		case True
		from read_only_reads_append [of "(𝒪sb  A - R)" "(takeWhile (Not  is_volatile_Writesb) sbn)"
		  "(dropWhile (Not  is_volatile_Writesb) sbn)"]
		have "read_only_reads
                  (acquired True (takeWhile (Not  is_volatile_Writesb) sbn) (𝒪sb  A - R))
                  (dropWhile (Not  is_volatile_Writesb) sbn)  read_only_reads (𝒪sb  A - R) sbn"
		  by auto
		
		with tssb_i nth mth neq_m_i n_bound' True
		  read_only_reads_unowned [OF i_bound m_bound' False [symmetric] tssb_i mth']
		show ?thesis
		  by (auto simp add: tssb'  sb 𝒪sb' Writesb volatile)
	      next
		case False
		with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
		  by (auto simp add: tssb')
		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 tssb'"
	  proof (unfold_locales)
	    fix i1 j p1 "is1" 𝒪1 1 𝒟1 xs1 sb1 pj "isj" "𝒪j" j 𝒟j xsj sbj
	    assume i1_bound: "i1 < length tssb'"
	    assume j_bound: "j < length tssb'"
	    assume i1_j: "i1  j"
	    assume ts_i1: "tssb'!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	    assume ts_j: "tssb'!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	    show "(𝒪1  all_acquired sb1)  (𝒪j  all_acquired sbj)= {}"
	    proof (cases "i1=i")
	      case True
	      with i1_j have i_j: "ij" 
		by simp
	      
	      from j_bound have j_bound': "j < length tssb"
		by (simp add: tssb')
	      hence j_bound'': "j < length (map owned tssb)"
		by simp	    
	      from ts_j i_j have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (simp add: tssb')
	      
	      from ownership_distinct [OF i_bound j_bound' i_j tssb_i ts_j']
	      show ?thesis
		using tssb_i True ts_i1 i_bound 𝒪sb'
		by (auto simp add: tssb' sb Writesb volatile)
	    next
	      case False
	      note i1_i = this
	      from i1_bound have i1_bound': "i1 < length tssb"
		by (simp add: tssb')
	      hence i1_bound'': "i1 < length (map owned tssb)"
		by simp	    
	      from ts_i1 False have ts_i1': "tssb!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
		by (simp add: tssb')
	      show ?thesis
	      proof (cases "j=i")
		case True
		from ownership_distinct [OF i1_bound' i_bound  i1_i ts_i1' tssb_i]
		show ?thesis
		  using tssb_i True ts_j i_bound 𝒪sb'
		  by (auto simp add: tssb' sb Writesb volatile)
	      next
		case False
		from j_bound have j_bound': "j < length tssb"
		  by (simp add: tssb')
		from ts_j False have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		  by (simp add: tssb')
		from ownership_distinct [OF i1_bound' j_bound' i1_j ts_i1' ts_j']
		show ?thesis .
	      qed
	    qed
	  qed
	qed

	have valid_sharing': "valid_sharing (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	proof (intro_locales)	
	  show "outstanding_non_volatile_writes_unshared (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	  proof (unfold_locales)
	    fix j pj "isj" "𝒪j" j 𝒟j acqj xsj sbj
	    assume j_bound: "j < length tssb'"
	    assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	    show "non_volatile_writes_unshared (𝒮sb ⊕⇘WR ⊖⇘AL)  sbj"
	    proof (cases "i=j")
	      case True
	      with outstanding_non_volatile_writes_unshared [OF i_bound tssb_i]
		i_bound jth tssb_i show ?thesis
		by (clarsimp simp add: tssb' sb Writesb volatile)
	    next
	      case False
	      from j_bound have j_bound': "j < length tssb"
		by (auto simp add: tssb')
	      from jth False have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (auto simp add: tssb')
	      from outstanding_non_volatile_writes_unshared [OF j_bound' jth']
	      have unshared: "non_volatile_writes_unshared 𝒮sb sbj".
	      
	      have "adom (𝒮sb ⊕⇘WR ⊖⇘AL) - dom 𝒮sb. a  outstanding_refs is_non_volatile_Writesb sbj"
	      proof -
		{
		  fix a 
		  assume a_in: "a  dom (𝒮sb ⊕⇘WR ⊖⇘AL) - dom 𝒮sb"
		  hence a_R: "a  R"
		    by clarsimp
		  assume a_in_j: "a  outstanding_refs is_non_volatile_Writesb sbj"
		  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 sbj"
		      by auto

		    moreover
		    with ownership_distinct [OF i_bound j_bound' False tssb_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_Writesb 
	      [OF unshared this]
	      show ?thesis .
	    qed
	  qed
	next
	  show "sharing_consis (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	  proof (unfold_locales)  
	    fix j pj "isj" "𝒪j" j 𝒟j xsj sbj
	    assume j_bound: "j < length tssb'"
	    assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	    show "sharing_consistent (𝒮sb ⊕⇘WR ⊖⇘AL) 𝒪j sbj"
	    proof (cases "i=j")
	      case True
	      with i_bound jth tssb_i sharing_consis [OF i_bound tssb_i]
	      show ?thesis
		by (clarsimp simp add: tssb' sb Writesb volatile 𝒪sb')
	    next
	      case False
	      from j_bound have j_bound': "j < length tssb"
		by (auto simp add: tssb')
	      from jth False have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (auto simp add: tssb')
	      from sharing_consis [OF j_bound' jth']
	      have consis: "sharing_consistent 𝒮sb 𝒪j sbj".
	      
	      have acq_cond: "all_acquired sbj  dom 𝒮sb - dom (𝒮sb ⊕⇘WR ⊖⇘AL) = {}"
	      proof -
		{
		  fix a
		  assume a_acq: "a  all_acquired sbj" 
		  assume "a  dom 𝒮sb" 
		  assume a_L: "a  L"
		  have False
		  proof -
		    from ownership_distinct [OF i_bound j_bound' False tssb_i jth']
		    have "A  all_acquired sbj = {}"
		      by (auto simp add: sb Writesb volatile)
		    with a_acq a_L L_subset
		    show False
		      by blast
		  qed
		}
		thus ?thesis
		  by auto
	      qed
	      have uns_cond: "all_unshared sbj  dom (𝒮sb ⊕⇘WR ⊖⇘AL) - dom 𝒮sb = {}"
	      proof -
		{
		  fix a
		  assume a_uns: "a  all_unshared sbj" 
		  assume "a  L" 
		  assume a_R:  "a  R"
		  have False
		  proof -
		    from unshared_acquired_or_owned [OF consis] a_uns
		    have "a  all_acquired sbj  𝒪j" by auto
		    with ownership_distinct [OF i_bound j_bound' False tssb_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: tssb')
	    qed
	  qed
	next
	  show "read_only_unowned (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	  proof 
	    fix j pj "isj" "𝒪j" j  𝒟j xsj sbj
	    assume j_bound: "j < length tssb'"
	    assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	    show "𝒪j  read_only (𝒮sb ⊕⇘WR ⊖⇘AL) = {}"
	    proof (cases "i=j")
	      case True
	      from read_only_unowned [OF i_bound tssb_i] R_owned  A_R 
	      have "(𝒪sb  A - R)  read_only (𝒮sb ⊕⇘WR ⊖⇘AL) = {}"
		by (auto simp add: in_read_only_convs )
	      with jth tssb_i i_bound True
	      show ?thesis
		by (auto simp add: 𝒪sb' tssb')
	    next
	      case False
	      from j_bound have j_bound': "j < length tssb"
		by (auto simp add: tssb')
	      with False jth have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (auto simp add: tssb')
	      from read_only_unowned [OF j_bound' jth']
	      have "𝒪j  read_only 𝒮sb = {}".
	      moreover
	      from ownership_distinct [OF i_bound j_bound' False tssb_i jth'] R_owned
	      have "(𝒪sb  A)  𝒪j = {}"
		by (auto simp add: sb Writesb 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 (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	  proof (unfold_locales)
	    show "- ((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set tssb')  dom (𝒮sb ⊕⇘WR ⊖⇘AL)"
	    proof -
	      
	      have s: "((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set tssb') =
              ((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set tssb)  A - R"
	      
		apply (unfold tssb' 𝒪sb') 
		apply (rule acquire_release_ownership_nth_update [OF R_owned i_bound tssb_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 (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	  proof 
	    fix j pj "isj" "𝒪j" j 𝒟j acqj xsj sbj
	    assume j_bound: "j < length tssb'"
	    assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	    show "no_write_to_read_only_memory (𝒮sb ⊕⇘WR ⊖⇘AL) sbj"
	    proof (cases "i=j")
	      case True
	      with jth tssb_i i_bound no_outstanding_write_to_read_only_memory [OF i_bound tssb_i]
	      show ?thesis
		by (auto simp add: sb tssb' Writesb volatile)
	    next
	      case False
	      from j_bound have j_bound': "j < length tssb"
		by (auto simp add: tssb')
	      with False jth have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (auto simp add: tssb')
	      from no_outstanding_write_to_read_only_memory [OF j_bound' jth']
	      have nw: "no_write_to_read_only_memory 𝒮sb sbj".
	      have "R  outstanding_refs is_Writesb sbj = {}"
	      proof -
		note dist = ownership_distinct [OF i_bound j_bound' False tssb_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_Writesb sbj  𝒪sb = {}"
		  by auto
		moreover
		from outstanding_volatile_writes_unowned_by_others [OF j_bound'  i_bound 
		  False [symmetric] jth' tssb_i ]
		have "outstanding_refs is_volatile_Writesb sbj  𝒪sb = {}" 
		  by auto
		ultimately have "outstanding_refs is_Writesb sbj  𝒪sb = {}" 
		  by (auto simp add: misc_outstanding_refs_convs)
		with R_owned
		show ?thesis by blast
	      qed
	      then
	      have "aoutstanding_refs is_Writesb sbj.
		a  read_only (𝒮sb ⊕⇘WR ⊖⇘AL)  a  read_only 𝒮sb"
		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 𝒪sb, release ?take_sb (dom 𝒮sb) sb,𝒮)  
          (is', , (), m (a := v),True, acquired True ?take_sb 𝒪sb  A - R, Map.empty,𝒮 ⊕⇘WR ⊖⇘AL)"
	  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),𝒮 ⊕⇘WR ⊖⇘AL)".	

	have sb'_split: 
	  "sb' = takeWhile (Not  is_volatile_Writesb) sb' @ 
                    dropWhile (Not  is_volatile_Writesb) sb'"
	  by simp

	from reads_consis 
	have no_vol_reads: "outstanding_refs is_volatile_Readsb sb' = {}"
	  by (simp add: sb Writesb True)
	hence "outstanding_refs is_volatile_Readsb (takeWhile (Not  is_volatile_Writesb) sb') 
	  = {}"
	  by (auto simp add: outstanding_refs_conv dest: set_takeWhileD)
	moreover 
	have "outstanding_refs is_volatile_Writesb 
           (takeWhile (Not  is_volatile_Writesb) sb') = {}"
	proof -
	  have "r  set (takeWhile (Not  is_volatile_Writesb) sb'). ¬ (is_volatile_Writesb 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_Writesb) sb') = {}"
	  by (auto simp add: outstanding_refs_conv is_volatile_split)

	moreover

	from no_vol_reads have "r  set sb'. ¬ is_volatile_Readsb r"
	  by (fastforce simp add: outstanding_refs_conv is_volatile_Readsb_def 
	    split: memref.splits)
	hence "r  set sb'. (Not  is_volatile_Writesb) r = (Not  is_volatile) r"
	  by (auto simp add: is_volatile_split)

	hence takeWhile_eq: "(takeWhile (Not  is_volatile_Writesb) sb') =
              (takeWhile (Not  is_volatile) sb')" 
	  apply -
	  apply (rule takeWhile_cong)
	  apply auto
	  done

	from leq
	have leq': "length tssb = 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_Writesb) sb' @ 
                 dropWhile (Not  is_volatile_Writesb) sb') @ issb = 
              is' @ prog_instrs (takeWhile (Not  is_volatile_Writesb) sb' @ 
                                 dropWhile (Not  is_volatile_Writesb) sb')"
	  by (simp add: sb'_split [symmetric])

	from reads_consistent_flush_all_until_volatile_write [OF valid_ownership_and_sharing 𝒮sb tssb
	i_bound tssb_i reads_consis]
	have "reads_consistent True (acquired True ?take_sb 𝒪sb) m (Writesb True a (D,f) v A L R W#sb')"
	  by (simp add: m sb Writesb volatile)

	hence "reads_consistent True (acquired True ?take_sb 𝒪sb  A - R) (m(a:=v)) sb'"
	  by simp
	from reads_consistent_takeWhile [OF this]
	have r_consis': "reads_consistent True (acquired True ?take_sb 𝒪sb  A - R) (m(a:=v)) 
	       (takeWhile (Not  is_volatile_Writesb) sb')".


	
	from last_prog have last_prog_sb': "last_prog psb sb' = psb"
	  by (simp add: sb Writesb )


	from valid_write_sops  [OF i_bound tssb_i]
	have "sop  write_sops sb'. valid_sop sop"
	  by (auto simp add: sb Writesb)
	hence valid_sop': "sopwrite_sops (takeWhile (Not  is_volatile_Writesb) sb').
	        valid_sop sop"
	  by (fastforce dest: set_takeWhileD simp add: in_write_sops_conv)

	from no_volatile
	have no_volatile_Readsb:
	  "outstanding_refs is_volatile_Readsb (takeWhile (Not  is_volatile_Writesb) sb') =
              {}"
	  by (auto simp add: outstanding_refs_conv is_volatile_Readsb_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_Readsb_volatile_reads_consistent [OF no_volatile_Readsb], 
	  where 𝒮="(𝒮 ⊕⇘WR ⊖⇘AL)"]

	  
	obtain is'' where
	  is''_sim: "instrs (dropWhile (Not  is_volatile_Writesb) sb') @ issb =
                      is'' @ prog_instrs (dropWhile (Not  is_volatile_Writesb) sb')" and

          steps: "(?ts', m(a := v), 𝒮 ⊕⇘WR ⊖⇘AL) d*
                   (ts[i := (last_prog (hd_prog psb (dropWhile (Not  is_volatile_Writesb) sb'))
                            (takeWhile (Not  is_volatile_Writesb) sb'),
                         is'',
                         θsb |` (dom θsb -
                                    read_tmps (dropWhile (Not  is_volatile_Writesb) sb')),
                         (), True, acquired True (takeWhile (Not  is_volatile_Writesb) sb')
                                (acquired True ?take_sb 𝒪sb  A - R),
                                release (takeWhile (Not  is_volatile_Writesb) sb')
                                   (dom (𝒮 ⊕⇘WR ⊖⇘AL)) Map.empty)],
	           flush (takeWhile (Not  is_volatile_Writesb) sb') (m(a := v)),
                   share (takeWhile (Not  is_volatile_Writesb) sb') (𝒮 ⊕⇘WR ⊖⇘AL))"


	  by (auto)

	note sim_flush = r_rtranclp_rtranclp [OF store_step steps]

	moreover
	note flush_commute =
	  flush_flush_all_until_volatile_write_Writesb_volatile_commute [OF i_bound tssb_i [simplified sb Writesb True] 
	outstanding_refs_is_Writesb_takeWhile_disj a_notin_others']

	from last_prog_hd_prog_append' [where sb="(takeWhile (Not  is_volatile_Writesb) sb')"
          and sb'="(dropWhile (Not  is_volatile_Writesb) sb')",
	  simplified sb'_split [symmetric], OF hist_consis' last_prog_sb']
	have last_prog_eq: 
	  "last_prog (hd_prog psb (dropWhile (Not  is_volatile_Writesb) sb'))
                 (takeWhile (Not  is_volatile_Writesb) sb') =
              hd_prog psb (dropWhile (Not  is_volatile_Writesb) sb')".

	have take_emtpy: "takeWhile (Not  is_volatile_Writesb) (r#sb) = []"
	  by (simp add: Writesb True)


        have dist_sb': "i p is 𝒪  𝒟 θ sb.
          i < length tssb 
          tssb ! i = (p, is, θ, sb, 𝒟, 𝒪, ) 
          (all_shared (takeWhile (Not  is_volatile_Writesb) sb) 
          all_unshared (takeWhile (Not  is_volatile_Writesb) sb) 
          all_acquired (takeWhile (Not  is_volatile_Writesb) sb)) 
          (all_shared (takeWhile (Not  is_volatile_Writesb) sb') 
          all_unshared (takeWhile (Not  is_volatile_Writesb) sb') 
          all_acquired (takeWhile (Not  is_volatile_Writesb) sb')) =
          {}"
        proof -
          {
            fix j pj isj 𝒪j j 𝒟j θj sbj x
	    assume j_bound: "j < length tssb"
	    assume jth: "tssb!j = (pj,isj, θj,sbj,𝒟j,𝒪j,j)"
	    assume x_shared: "x  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)  
                                 all_unshared (takeWhile (Not  is_volatile_Writesb) sbj)  
                                 all_acquired (takeWhile (Not  is_volatile_Writesb) sbj)"
            assume x_sb': "x  (all_shared (takeWhile (Not  is_volatile_Writesb) sb') 
                        all_unshared (takeWhile (Not  is_volatile_Writesb) sb') 
                        all_acquired (takeWhile (Not  is_volatile_Writesb) sb'))"
            have False
            proof (cases "i=j")
              case True with x_shared tssb_i jth show False by (simp add: sb volatile Writesb)
            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_Writesb) sbj)" 
		"(dropWhile (Not  is_volatile_Writesb) sbj)"]
                all_unshared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
		"(dropWhile (Not  is_volatile_Writesb) sbj)"]
                all_acquired_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
		"(dropWhile (Not  is_volatile_Writesb) sbj)"]
              have "x  all_acquired sbj  𝒪j "
                by auto
              moreover
              from x_sb' all_shared_acquired_or_owned [OF sharing_consis [OF i_bound tssb_i]]
                unshared_acquired_or_owned [OF sharing_consis [OF i_bound tssb_i]]
                all_shared_append [of "(takeWhile (Not  is_volatile_Writesb) sb')" 
		"(dropWhile (Not  is_volatile_Writesb) sb')"]
                all_unshared_append [of "(takeWhile (Not  is_volatile_Writesb) sb')" 
		"(dropWhile (Not  is_volatile_Writesb) sb')"]
                all_acquired_append [of "(takeWhile (Not  is_volatile_Writesb) sb')" 
		"(dropWhile (Not  is_volatile_Writesb) sb')"]
              have "x  all_acquired sb  𝒪sb"
                by (auto simp add: sb Writesb volatile)
              moreover
              note ownership_distinct [OF i_bound j_bound False tssb_i jth]
              ultimately show False by blast
            qed
          }
          thus ?thesis by blast
        qed

        have dist_R_L_A: "j p is 𝒪  𝒟 θ sb.
          j < length tssb  i j
          tssb ! j = (p, is, θ, sb, 𝒟, 𝒪, ) 
          (all_shared sb  all_unshared sb  all_acquired sb)  (R  L  A) = {}"
        proof -
          {
            fix j pj isj 𝒪j j 𝒟j θj sbj x
	    assume j_bound: "j < length tssb"
            assume neq_i_j: "i  j"
	    assume jth: "tssb!j = (pj,isj, θj,sbj,𝒟j,𝒪j,j)"
	    assume x_shared: "x  all_shared sbj  
                                 all_unshared sbj  
                                 all_acquired  sbj"
            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 sbj  𝒪j "
                by auto
              moreover
              from x_R_L_A R_owned L_subset
              have "x  all_acquired sb  𝒪sb"
                by (auto simp add: sb Writesb volatile)
              moreover
              note ownership_distinct [OF i_bound j_bound neq_i_j tssb_i jth]
              ultimately show False by blast
            qed
          }
          thus ?thesis by blast
        qed
        from local.ownership_distinct_axioms have "ownership_distinct tssb" .
        from local.sharing_consis_axioms have "sharing_consis 𝒮sb tssb".
        note share_commute=
          share_all_until_volatile_write_flush_commute [OF take_empty ownership_distinct tssb sharing_consis 𝒮sb tssb i_bound tssb_i dist_sb' dist_R_L_A]
        
        have rel_commute_empty:
          "release (takeWhile (Not  is_volatile_Writesb) sb') (dom 𝒮  R - L) Map.empty =
                 release (takeWhile (Not  is_volatile_Writesb) sb') (dom 𝒮sb  R - L) Map.empty"
        proof -
          {
            fix a
            assume a_in: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sb')"
            have "(a  (dom 𝒮  R - L)) = (a  (dom 𝒮sb  R - L))"
            proof -
              
              from all_shared_acquired_or_owned [OF sharing_consis [OF i_bound tssb_i]] a_in
                all_shared_append [of "(takeWhile (Not  is_volatile_Writesb) sb')" "(dropWhile (Not  is_volatile_Writesb) sb')"]
              have "a  𝒪sb  all_acquired sb  " 
                by (auto simp add: sb Writesb volatile)
              from share_all_until_volatile_write_thread_local [OF ownership_distinct tssb sharing_consis 𝒮sb tssb i_bound tssb_i this]
              have "𝒮 a = 𝒮sb a"
                by (auto simp add: sb Writesb 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 pj isj 𝒪j j 𝒟j θj sbj x
	  assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	  assume j_bound: "j < length tssb"
          assume neq: "i  j" 
          have "release (takeWhile (Not  is_volatile_Writesb) sbj)
                            (dom 𝒮sb  R - L) j
              = release (takeWhile (Not  is_volatile_Writesb) sbj)
                            (dom 𝒮sb) j"
          proof -
            {
              fix a
              assume a_in: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)"
              have "(a  (dom 𝒮sb  R - L)) = (a  dom 𝒮sb)"
              proof -
                from ownership_distinct [OF i_bound j_bound neq  tssb_i jth]
                
                have A_dist: "A  (𝒪j  all_acquired sbj) = {}"
                  by (auto simp add: sb Writesb volatile)
              
                from  all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]] a_in
                  all_shared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
                  "(dropWhile (Not  is_volatile_Writesb) sbj)"]
                have a_in: "a  𝒪j  all_acquired sbj"
                  by auto
                with ownership_distinct [OF i_bound j_bound neq  tssb_i jth]
                have "a  (𝒪sb  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 "(tssb [i := (psb,issb, θsb, sb', 𝒟sb, 𝒪sb  A - R,Map.empty)],msb(a:=v),𝒮sb')  
	      (ts[i := (last_prog (hd_prog psb (dropWhile (Not  is_volatile_Writesb) sb'))
                            (takeWhile (Not  is_volatile_Writesb) sb'),
                         is'',
                         θsb |` (dom θsb -
                                    read_tmps (dropWhile (Not  is_volatile_Writesb) sb')),
                         (),True, acquired True (takeWhile (Not  is_volatile_Writesb) sb')
                                (acquired True ?take_sb 𝒪sb  A - R),
                             release (takeWhile (Not  is_volatile_Writesb) sb')
                                   (dom (𝒮 ⊕⇘WR ⊖⇘AL)) Map.empty)],
               flush (takeWhile (Not  is_volatile_Writesb) sb') (m(a := v)),
               share (takeWhile (Not  is_volatile_Writesb) sb') (𝒮 ⊕⇘WR ⊖⇘AL))"
	  apply (rule sim_config.intros) 
	  apply    (simp add: flush_commute m)
	  apply   (clarsimp simp add: 𝒮sb' 𝒮 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 Writesb volatile  𝒮sb'
            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'
	    msb' 𝒮sb' tssb' 
	  by (auto simp del: fun_upd_apply simp add: 𝒪sb' sb' )

      next

	case False
	note non_vol = this
	
	from flush Writesb False
	obtain 
	  𝒪sb': "𝒪sb'=𝒪sb" and
	  𝒮sb': "𝒮sb'=𝒮sb" and
          sb': "sb' = sb"
	  by cases (auto  simp add: sb)


	from non_volatile_owned non_vol have a_owned: "a  𝒪sb"
	  by simp

	{
	  fix j 
	  fix pj issbj 𝒪j 𝒟sbj θj j sbj
	  assume j_bound: "j < length tssb"
	  assume tssb_j: "tssb!j=(pj,issbj,θj,sbj,𝒟sbj,𝒪j,j)"
	  assume neq_i_j: "ij"
	  have "a  unforwarded_non_volatile_reads (dropWhile (Not  is_volatile_Writesb) sbj) {}"
	  proof 
	    let ?take_sbj = "takeWhile (Not  is_volatile_Writesb) sbj"
	    let ?drop_sbj = "dropWhile (Not  is_volatile_Writesb) sbj"
	    assume a_in: "a   unforwarded_non_volatile_reads ?drop_sbj {}"
	    
	    from a_unowned_by_others [rule_format, OF j_bound neq_i_j] tssb_j 
	    obtain a_unowned: "a  𝒪j" and a_unacq: "a  all_acquired sbj"
	      by auto
	    with all_acquired_append [of ?take_sbj ?drop_sbj] acquired_takeWhile_non_volatile_Writesb [of sbj 𝒪j]
	    have a_unacq_take: "a  acquired True ?take_sbj 𝒪j"
	      by (auto )

	    note nvo_j = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound tssb_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_sbj 𝒮sb)
	      (acquired True ?take_sbj 𝒪j) ?drop_sbj" .
	    from in_unforwarded_non_volatile_reads_non_volatile_Readsb [OF a_in]
	    have a_in': "a  outstanding_refs is_non_volatile_Readsb ?drop_sbj".

	    from non_volatile_owned_or_read_only_outstanding_refs [OF nvo_drop_j] a_in'
	    have "a  acquired True ?take_sbj 𝒪j  all_acquired ?drop_sbj   
	      read_only_reads (acquired True ?take_sbj 𝒪j) ?drop_sbj"
	      by (auto simp add: misc_outstanding_refs_convs)
	    
	    moreover 
	    from acquired_append [of True ?take_sbj ?drop_sbj 𝒪j] acquired_all_acquired [of True ?take_sbj 𝒪j]
	      all_acquired_append [of ?take_sbj ?drop_sbj]
	    have "acquired True ?take_sbj 𝒪j  all_acquired ?drop_sbj  𝒪j  all_acquired sbj"
	      by auto
	    ultimately 
	    have "a  read_only_reads (acquired True ?take_sbj 𝒪j) ?drop_sbj"
	      using a_owned ownership_distinct [OF i_bound j_bound neq_i_j tssb_i tssb_j]
	      by auto
	    
	    with read_only_reads_unowned [OF j_bound i_bound neq_i_j [symmetric] tssb_j tssb_i] a_owned
	    show False
	      by auto
	  qed
	} note a_notin_unforwarded_non_volatile_reads_drop = this
	    

	(* FIXME: the same proof as in volatile, case. depends on a_notin_unforwarded_non_volatile_reads_drop *)
	have valid_reads': "valid_reads msb' tssb'"
	proof (unfold_locales)
	  fix j pj "isj" 𝒪j j 𝒟j θj sbj
	  assume j_bound: "j < length tssb'"
	  assume ts_j: "tssb'!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	  show "reads_consistent False 𝒪j msb' sbj"
	  proof (cases "i=j")
	    case True
	    from reads_consis ts_j j_bound sb show ?thesis
	      by (clarsimp simp add: True  msb' Writesb tssb' 𝒪sb' False reads_consistent_pending_write_antimono)
	  next
	    case False
	    from j_bound have j_bound':  "j < length tssb"
	      by (simp add: tssb')
	    moreover from ts_j False have ts_j': "tssb ! j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	      using j_bound by (simp add: tssb')
	    ultimately have consis_m: "reads_consistent False 𝒪j msb sbj"
	      by (rule valid_reads)
	    from a_unowned_by_others [rule_format, OF j_bound' False] ts_j'
	    have a_unowned:"a  𝒪j  all_acquired sbj"
	      by simp

	    let ?take_sbj = "takeWhile (Not  is_volatile_Writesb) sbj"
	    let ?drop_sbj = "dropWhile (Not  is_volatile_Writesb) sbj"

	    from a_unowned acquired_reads_all_acquired [of True ?take_sbj 𝒪j]
	    all_acquired_append [of ?take_sbj ?drop_sbj]
	    have a_not_acq_reads: "a  acquired_reads True ?take_sbj 𝒪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_sbj {}  acquired_reads True ?take_sbj 𝒪j" and
		m'= "(msb(a:=v))", OF _ _ _ consis_m]
	      by (fastforce simp add: msb')
	  qed
	qed

	have valid_own': "valid_ownership 𝒮sb' tssb'"
	proof (intro_locales)
	  show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	  proof -
	    from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i] sb
	    have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb sb'"
	      by (auto simp add: Writesb False)
	    from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
	    show ?thesis by (simp add: tssb' Writesb False 𝒪sb' 𝒮sb')
	  qed
	next
	  show "outstanding_volatile_writes_unowned_by_others tssb'"
	  proof -
	    from sb 
	    have out: "outstanding_refs is_volatile_Writesb sb'  outstanding_refs is_volatile_Writesb sb"
	      by (auto simp add: Writesb False)
	    have acq: "all_acquired sb'  all_acquired sb"
	      by (auto simp add: Writesb False sb)
	    from outstanding_volatile_writes_unowned_by_others_store_buffer 
	    [OF i_bound tssb_i out acq]
	    show ?thesis by (simp add: tssb' Writesb False 𝒪sb')
	  qed
	next
	  show "read_only_reads_unowned tssb'"
	  proof -
	    have ro: "read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb') 𝒪sb)
	      (dropWhile (Not  is_volatile_Writesb) sb')
	       read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪sb)
	      (dropWhile (Not  is_volatile_Writesb) sb)"
	      by (auto simp add: sb Writesb non_vol)
	    have "𝒪sb  all_acquired sb'  𝒪sb  all_acquired sb"
	      by (auto simp add: sb Writesb non_vol)
	    from read_only_reads_unowned_nth_update [OF i_bound tssb_i ro this] 
	    show ?thesis
	      by (simp add: tssb' sb 𝒪sb')
	  qed
	next
	  show "ownership_distinct tssb'"
	  proof -
	    have acq: "all_acquired sb'  all_acquired sb"
	      by (auto simp add: Writesb False sb)
	    with ownership_distinct_instructions_read_value_store_buffer_independent 
	    [OF i_bound tssb_i]
	    show ?thesis by (simp add: tssb' Writesb False 𝒪sb')
	  qed
	qed

	have valid_sharing': "valid_sharing 𝒮sb' tssb'"
	proof (intro_locales)	
	  from outstanding_non_volatile_writes_unshared [OF i_bound tssb_i]
	  have "non_volatile_writes_unshared 𝒮sb sb'"
	    by (auto simp add: sb Writesb False)
	  from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
	  show "outstanding_non_volatile_writes_unshared 𝒮sb' tssb'"
	    by (simp add: tssb' 𝒮sb')
	next
	  from sharing_consis [OF i_bound tssb_i]
	  have "sharing_consistent 𝒮sb 𝒪sb sb'"
	    by (auto simp add: sb Writesb False)
	  from sharing_consis_nth_update [OF i_bound this]
	  show "sharing_consis 𝒮sb' tssb'"
	    by (simp add: tssb' 𝒪sb' 𝒮sb')
	next
	  from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound tssb_i] ]
	  show "read_only_unowned 𝒮sb' tssb'"
	    by (simp add: 𝒮sb' tssb' 𝒪sb')
	next
	  from unowned_shared_nth_update [OF i_bound tssb_i subset_refl]
	  show "unowned_shared 𝒮sb' tssb'"
	    by (simp add: tssb' 𝒪sb' 𝒮sb')
	next
	  from no_outstanding_write_to_read_only_memory [OF i_bound tssb_i] 
	  have "no_write_to_read_only_memory 𝒮sb sb'"
	    by (auto simp add: sb Writesb False)
	  from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
	  show "no_outstanding_write_to_read_only_memory 𝒮sb' tssb'"
	    by (simp add: 𝒮sb' tssb' sb)
	qed

	from is_sim
	obtain is_sim: "instrs (dropWhile (Not  is_volatile_Writesb) sb') @ issb =
                 is @ prog_instrs (dropWhile (Not  is_volatile_Writesb) sb')"
	  by (simp add: suspends sb Writesb False)

	have "(ts,m,𝒮) d* (ts,m,𝒮)" by blast

	moreover


	note flush_commute =
	  flush_all_until_volatile_write_Writesb_non_volatile_commute [OF i_bound tssb_i [simplified sb Writesb non_vol] 
	outstanding_refs_is_Writesb_takeWhile_disj a_notin_others']

	note share_commute = 
	  share_all_until_volatile_write_update_sb [of sb' sb, OF _ i_bound tssb_i, simplified sb Writesb False, simplified]
	have "(tssb [i := (psb,issb,θsb, sb', 𝒟sb, 𝒪sb,sb)], msb(a:=v),𝒮sb')  
                (ts,m,𝒮)"
	  apply (rule sim_config.intros) 
	  apply    (simp add: m flush_commute)
	  apply   (clarsimp simp add: 𝒮 𝒮sb' 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 Writesb False 𝒮sb'
	     split: if_split_asm )
	  done	

	ultimately
	show ?thesis
	  using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' msb'
	   valid_dd' valid_sops' load_tmps_fresh' enough_flushs' valid_program_history' valid'
	    tssb' 𝒪sb' 𝒮sb' sb'
	  by (auto simp del: fun_upd_apply)
      qed
    next
      case (Readsb volatile a t v)
      from flush this obtain msb': "msb'=msb" and 
	𝒪sb': "𝒪sb'=𝒪sb" and 𝒮sb': "𝒮sb'=𝒮sb" and
        sb': "sb'=sb"
	by cases (auto simp add: sb)

      have valid_own': "valid_ownership 𝒮sb' tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	proof -
	  from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i] sb
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb sb'"
	    by (auto simp add: Readsb)
	  from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
	  show ?thesis by (simp add: tssb' Readsb 𝒪sb' 𝒮sb')
	qed
      next
	show "outstanding_volatile_writes_unowned_by_others tssb'"
	proof -
	  from sb 
	  have out: "outstanding_refs is_volatile_Writesb sb'  outstanding_refs is_volatile_Writesb sb"
	    by (auto simp add: Readsb)
	  have acq: "all_acquired sb'  all_acquired sb"
	    by (auto simp add: Readsb sb)
	  from outstanding_volatile_writes_unowned_by_others_store_buffer 
	  [OF i_bound tssb_i out acq]
	  show ?thesis by (simp add: tssb' Readsb 𝒪sb')
	qed
      next
	show "read_only_reads_unowned tssb'"
	proof -
	  have ro: "read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb') 𝒪sb)
	    (dropWhile (Not  is_volatile_Writesb) sb')
	     read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪sb)
	    (dropWhile (Not  is_volatile_Writesb) sb)"
	    by (auto simp add: sb Readsb)
	  have "𝒪sb  all_acquired sb'  𝒪sb  all_acquired sb"
	    by (auto simp add: sb Readsb)
	  from read_only_reads_unowned_nth_update [OF i_bound tssb_i ro this] 
	  show ?thesis
	    by (simp add: tssb' sb 𝒪sb')
	qed
      next
	show "ownership_distinct tssb'"
	proof -
	  have acq: "all_acquired sb'  all_acquired sb"
	    by (auto simp add: Readsb sb)
	  with ownership_distinct_instructions_read_value_store_buffer_independent 
	  [OF i_bound tssb_i]
	  show ?thesis by (simp add: tssb' Readsb 𝒪sb')
	qed
      qed

      have valid_sharing': "valid_sharing 𝒮sb' tssb'"
      proof (intro_locales)	
	from outstanding_non_volatile_writes_unshared [OF i_bound tssb_i]
	have "non_volatile_writes_unshared 𝒮sb sb'"
	  by (auto simp add: sb Readsb)
	from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
	show "outstanding_non_volatile_writes_unshared 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒮sb')
      next
	from sharing_consis [OF i_bound tssb_i]
	have "sharing_consistent 𝒮sb 𝒪sb sb'"
	  by (auto simp add: sb Readsb)
	from sharing_consis_nth_update [OF i_bound this]
	show "sharing_consis 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒪sb' 𝒮sb')
      next
	from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound tssb_i] ]
	show "read_only_unowned 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' 𝒪sb')
      next
	from unowned_shared_nth_update [OF i_bound tssb_i subset_refl]
	show "unowned_shared 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒪sb' 𝒮sb')
      next
	from no_outstanding_write_to_read_only_memory [OF i_bound tssb_i] 
	have "no_write_to_read_only_memory 𝒮sb sb'"
	  by (auto simp add: sb Readsb)
	from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
	show "no_outstanding_write_to_read_only_memory 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' sb)
      qed	

      have valid_reads': "valid_reads msb' tssb'"
      proof -
	from valid_reads [OF i_bound tssb_i]
	have "reads_consistent False 𝒪sb msb sb'"
	  by (simp add: sb Readsb)
	from valid_reads_nth_update [OF i_bound this]
	show ?thesis by (simp add: msb' tssb' 𝒪sb')
      qed

      have valid_program_history': "valid_program_history tssb'"
      proof -	
	from valid_program_history [OF i_bound tssb_i]
	have "causal_program_history issb sb" .
	then have causal': "causal_program_history issb sb'"
	  by (simp add: sb Readsb causal_program_history_def)

	from valid_last_prog [OF i_bound tssb_i]
	have "last_prog psb sb = psb".
	hence "last_prog psb sb' = psb"
	  by (simp add: sb Readsb)

	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb')
      qed

      from is_sim
      have is_sim: "instrs (dropWhile (Not  is_volatile_Writesb) sb') @ issb =
	         is @ prog_instrs (dropWhile (Not  is_volatile_Writesb) sb')"
	by (simp add: sb Readsb suspends)

      from valid_history [OF i_bound tssb_i]
      have θsb_v: "θsb t = Some v"
	by (simp add: history_consistent_access_last_read sb Readsb split:option.splits)

      have "(ts,m,𝒮) d* (ts,m,𝒮)" by blast

      moreover

      note flush_commute= flush_all_until_volatile_write_Readsb_commute [OF i_bound tssb_i [simplified sb Readsb]]
 
      note share_commute = 
	  share_all_until_volatile_write_update_sb [of sb' sb, OF _ i_bound tssb_i, simplified sb Readsb, simplified]
      have "(tssb [i := (psb,issb, θsb, sb',𝒟sb, 𝒪sb,sb')],msb,𝒮sb')  (ts,m,𝒮)"
	apply (rule sim_config.intros) 
	apply    (simp add: m flush_commute)
	apply   (clarsimp simp add: 𝒮 𝒮sb' 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 Readsb 𝒮sb' sb'
	   split: if_split_asm)
	done	

      ultimately show ?thesis
	using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' msb'
	  valid_dd' valid_sops' load_tmps_fresh' enough_flushs' valid_sharing' 
	  valid_program_history' valid'
	  tssb' 𝒪sb' 𝒮sb' 
	by (auto simp del: fun_upd_apply)
    next
      case (Progsb p1 p2 mis)
      from flush this obtain msb': "msb'=msb" and 
	𝒪sb': "𝒪sb'=𝒪sb" and 𝒮sb': "𝒮sb'=𝒮sb" and
        sb': "sb'=sb"
	by cases (auto simp add: sb)

      have valid_own': "valid_ownership 𝒮sb' tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	proof -
	  from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i] sb
	  have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb sb'"
	    by (auto simp add: Progsb)
	  from outstanding_non_volatile_refs_owned_or_read_only_nth_update [OF i_bound this]
	  show ?thesis by (simp add: tssb' Progsb 𝒪sb' 𝒮sb')
	qed
      next
	show "outstanding_volatile_writes_unowned_by_others tssb'"
	proof -
	  from sb 
	  have out: "outstanding_refs is_volatile_Writesb sb'  outstanding_refs is_volatile_Writesb sb"
	    by (auto simp add: Progsb)
	  have acq: "all_acquired sb'  all_acquired sb"
	    by (auto simp add: Progsb sb)
	  from outstanding_volatile_writes_unowned_by_others_store_buffer 
	  [OF i_bound tssb_i out acq]
	  show ?thesis by (simp add: tssb' Progsb 𝒪sb')
	qed
      next
	show "read_only_reads_unowned tssb'"
	proof -
	  have ro: "read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb') 𝒪sb)
	    (dropWhile (Not  is_volatile_Writesb) sb')
	       read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪sb)
	    (dropWhile (Not  is_volatile_Writesb) sb)"
	    by (auto simp add: sb Progsb)
	  have "𝒪sb  all_acquired sb'  𝒪sb  all_acquired sb"
	    by (auto simp add: sb Progsb)
	  from read_only_reads_unowned_nth_update [OF i_bound tssb_i ro this] 
	  show ?thesis
	    by (simp add: tssb' sb 𝒪sb')
	qed
      next
	  show "ownership_distinct tssb'"
	  proof -
	  have acq: "all_acquired sb'  all_acquired sb"
	    by (auto simp add: Progsb sb)
	  with ownership_distinct_instructions_read_value_store_buffer_independent 
	  [OF i_bound tssb_i]
	  show ?thesis by (simp add: tssb' Progsb 𝒪sb')
	qed
      qed

      have valid_sharing': "valid_sharing 𝒮sb' tssb'"
      proof (intro_locales)	
	from outstanding_non_volatile_writes_unshared [OF i_bound tssb_i]
	have "non_volatile_writes_unshared 𝒮sb sb'"
	  by (auto simp add: sb Progsb)
	from outstanding_non_volatile_writes_unshared_nth_update [OF i_bound this]
	show "outstanding_non_volatile_writes_unshared 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒮sb')
      next
	from sharing_consis [OF i_bound tssb_i]
	have "sharing_consistent 𝒮sb 𝒪sb sb'"
	  by (auto simp add: sb Progsb)
	from sharing_consis_nth_update [OF i_bound this]
	show "sharing_consis 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒪sb' 𝒮sb')
      next
	from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound tssb_i] ]
	show "read_only_unowned 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' 𝒪sb')
      next
	from unowned_shared_nth_update [OF i_bound tssb_i subset_refl]
	show "unowned_shared 𝒮sb' tssb'"
	  by (simp add: tssb' 𝒪sb' 𝒮sb')
      next
	from no_outstanding_write_to_read_only_memory [OF i_bound tssb_i] 
	have "no_write_to_read_only_memory 𝒮sb sb'"
	  by (auto simp add: sb Progsb)
	from no_outstanding_write_to_read_only_memory_nth_update [OF i_bound this]
	show "no_outstanding_write_to_read_only_memory 𝒮sb' tssb'"
	  by (simp add: 𝒮sb' tssb' sb)
      qed
      
      have valid_reads': "valid_reads msb' tssb'"
      proof -
	from valid_reads [OF i_bound tssb_i]
	have "reads_consistent False 𝒪sb msb sb'"
	  by (simp add: sb Progsb)
	from valid_reads_nth_update [OF i_bound this]
	show ?thesis by (simp add: msb' tssb' 𝒪sb')
      qed

      have valid_program_history': "valid_program_history tssb'"
      proof -	
	from valid_program_history [OF i_bound tssb_i]
	have "causal_program_history issb sb" .
	then have causal': "causal_program_history issb sb'"
	  by (simp add: sb Progsb causal_program_history_def)

	from valid_last_prog [OF i_bound tssb_i]
	have "last_prog psb sb = psb".
	hence "last_prog p2 sb' = psb"
	  by (simp add: sb Progsb)
	from last_prog_to_last_prog_same [OF this]
	have "last_prog psb sb' = psb".

	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb')
      qed

      from is_sim
      have is_sim: "instrs (dropWhile (Not  is_volatile_Writesb) sb') @ issb =
	is @ prog_instrs (dropWhile (Not  is_volatile_Writesb) sb')"
	by (simp add: suspends sb Progsb)

      have "(ts,m,𝒮) d* (ts,m,𝒮)" by blast

      moreover

      note flush_commute = flush_all_until_volatile_write_Progsb_commute [OF i_bound 
	tssb_i [simplified sb Progsb]]

      note share_commute = 
	  share_all_until_volatile_write_update_sb [of sb' sb, OF _ i_bound tssb_i, simplified sb Progsb, simplified]

      have "(tssb [i := (psb,issb, θsb, sb',𝒟sb, 𝒪sb,sb)],msb,𝒮sb')  (ts,m,𝒮)"
	apply (rule sim_config.intros) 
	apply    (simp add: m flush_commute)
	apply   (clarsimp simp add: 𝒮 𝒮sb' 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 Progsb sb' 𝒮sb'
	   split: if_split_asm)
	done	
      ultimately show ?thesis
	using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' msb'
	  valid_dd' valid_sops' load_tmps_fresh' enough_flushs' valid_sharing' 
	  valid_program_history' valid' 
	  tssb' 𝒮sb' 𝒪sb' sb' 𝒮sb'
	by (auto simp del: fun_upd_apply)
    next
      case (Ghostsb A L R W)
      from flush Ghostsb
      obtain 
	𝒪sb': "𝒪sb'=𝒪sb  A - R" and
	𝒮sb': "𝒮sb'=𝒮sb ⊕⇘WR ⊖⇘AL" and
        sb': "sb'= augment_rels (dom 𝒮sb) R sb" and
	msb': "msb'=msb" 
	by cases (auto simp add: sb)

      from sharing_consis [OF i_bound tssb_i] 
      obtain 
	A_shared_owned: "A  dom 𝒮sb  𝒪sb" and
	L_subset: "L  A" and
	A_R: "A  R = {}" and
	R_owned: "R  𝒪sb"
	by (clarsimp simp add: sb Ghostsb)

      have valid_own': "valid_ownership 𝒮sb' tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
	proof 
	  fix j isj 𝒪j j 𝒟j acqj θj sbj pj
	  assume j_bound: "j < length tssb'"
	  assume tssb'_j: "tssb'!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	  show "non_volatile_owned_or_read_only False 𝒮sb' 𝒪j sbj"
	  proof (cases "j=i")
	    case True
	    from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i]
	    have "non_volatile_owned_or_read_only False (𝒮sb ⊕⇘WR ⊖⇘AL) (𝒪sb  A - R) sb'"
	      by (auto simp add: sb Ghostsb non_volatile_owned_or_read_only_pending_write_antimono)
	    then show ?thesis
	      using True i_bound tssb'_j
	      by (auto simp add: tssb' 𝒮sb' sb 𝒪sb')
	  next
	    case False
	    from j_bound have j_bound': "j < length tssb"
	      by (auto simp add: tssb')
	    with tssb'_j False i_bound 
	    have tssb_j: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	      by (auto simp add: tssb')
	      
	      
	    note nvo = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound' tssb_j]

	    from read_only_unowned [OF i_bound tssb_i] R_owned
	    have "R  read_only 𝒮sb = {}"
	      by auto

	    with read_only_reads_unowned [OF j_bound' i_bound False tssb_j tssb_i] L_subset
	    have "aread_only_reads
	      (acquired True (takeWhile (Not  is_volatile_Writesb) sbj) 𝒪j)
		(dropWhile (Not  is_volatile_Writesb) sbj).
		a  read_only 𝒮sb  a  read_only (𝒮sb ⊕⇘WR ⊖⇘AL )"
	      by (auto simp add: in_read_only_convs sb Ghostsb)
	    from non_volatile_owned_or_read_only_read_only_reads_eq' [OF nvo this]
	    have "non_volatile_owned_or_read_only False (𝒮sb ⊕⇘WR ⊖⇘AL) 𝒪j sbj".
	    thus ?thesis by (simp add: 𝒮sb')
	  qed
	qed
      next
	show "outstanding_volatile_writes_unowned_by_others tssb'"
	proof (unfold_locales)
	  fix i1 j p1 "is1" 𝒪1 1 𝒟1 xs1 sb1 pj "isj" "𝒪j" j 𝒟j xsj sbj
	  assume i1_bound: "i1 < length tssb'"
	  assume j_bound: "j < length tssb'"
	  assume i1_j: "i1  j"
	  assume ts_i1: "tssb'!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	  assume ts_j: "tssb'!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb1 = {}"
	  proof (cases "i1=i")
	    case True
	    from i1_j True have neq_i_j: "ij"
	      by auto
	    from j_bound have j_bound': "j < length tssb"
	      by (simp add: tssb')
	    from ts_j neq_i_j have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (simp add: tssb')
	    from outstanding_volatile_writes_unowned_by_others [OF i_bound j_bound' neq_i_j
	      tssb_i ts_j'] ts_i1 i_bound tssb_i True show ?thesis
	      by (clarsimp simp add: tssb' sb Ghostsb)
	  next
	    case False
	    note i1_i = this
	    from i1_bound have i1_bound': "i1 < length tssb"
	      by (simp add: tssb' sb)
	    hence i1_bound'': "i1 < length (map owned tssb)"
	      by auto
	    from ts_i1 False have ts_i1': "tssb!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	      by (simp add: tssb' sb)
	    show ?thesis
	    proof (cases "j=i")
	      case True
	      from outstanding_volatile_writes_unowned_by_others [OF i1_bound' i_bound  i1_i  ts_i1' tssb_i ]
	      have "(𝒪sb  all_acquired sb)  outstanding_refs is_volatile_Writesb sb1 = {}".
	      then show ?thesis
		using True i1_i ts_j tssb_i i_bound
		by (auto simp add: sb Ghostsb tssb' 𝒪sb')
	    next
	      case False
	      from j_bound have j_bound': "j < length tssb"
		by (simp add: tssb')
	      from ts_j False have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (simp add: tssb')
	      from outstanding_volatile_writes_unowned_by_others 
	      [OF i1_bound' j_bound' i1_j ts_i1' ts_j']
	      show "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sb1 = {}" .
	    qed
	  qed
	qed
      next
	show "read_only_reads_unowned tssb'"
	proof 
	  fix n m
	  fix pn "isn" 𝒪n n 𝒟n θn sbn pm "ism" 𝒪m m 𝒟m θm sbm
	  assume n_bound: "n < length tssb'"
	    and m_bound: "m < length tssb'"
	    and neq_n_m: "nm"
	    and nth: "tssb'!n = (pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
	    and mth: "tssb'!m =(pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
	  from n_bound have n_bound': "n < length tssb" by (simp add: tssb')
	  from m_bound have m_bound': "m < length tssb" by (simp add: tssb')
	  show "(𝒪m  all_acquired sbm) 
            read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sbn) 𝒪n)
            (dropWhile (Not  is_volatile_Writesb) sbn) =
            {}"
	  proof (cases "m=i")
	    case True
	    with neq_n_m have neq_n_i: "ni"
	      by auto
	    
	    with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
	      by (auto simp add: tssb')
	    note read_only_reads_unowned [OF n_bound' i_bound  neq_n_i nth' tssb_i]
	    then
	    show ?thesis
	      using True tssb_i neq_n_i nth mth n_bound' m_bound' L_subset
	      by (auto simp add: tssb' 𝒪sb' sb Ghostsb)
	  next
	    case False
	    note neq_m_i = this
	    with m_bound mth i_bound have mth': "tssb!m = (pm, ism, θm, sbm, 𝒟m, 𝒪m,m)"
	      by (auto simp add: tssb')
	    show ?thesis
	    proof (cases "n=i")
	      case True
	      from read_only_reads_append [of "(𝒪sb  A - R)" "(takeWhile (Not  is_volatile_Writesb) sbn)"
		"(dropWhile (Not  is_volatile_Writesb) sbn)"]
	      have "read_only_reads
                (acquired True (takeWhile (Not  is_volatile_Writesb) sbn) (𝒪sb  A - R))
                (dropWhile (Not  is_volatile_Writesb) sbn)  read_only_reads (𝒪sb  A - R) sbn"
		by auto
		
	      with tssb_i nth mth neq_m_i n_bound' True
		read_only_reads_unowned [OF i_bound m_bound' False [symmetric] tssb_i mth']
	      show ?thesis
		by (auto simp add: tssb'  sb 𝒪sb' Ghostsb)
	    next
	      case False
	      with n_bound nth i_bound have nth': "tssb!n =(pn, isn, θn, sbn, 𝒟n, 𝒪n,n)"
		by (auto simp add: tssb')
	      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 tssb'"
	proof (unfold_locales)
	  fix i1 j p1 "is1" 𝒪1 1 𝒟1 xs1 sb1 pj "isj" "𝒪j" j 𝒟j xsj sbj
	  assume i1_bound: "i1 < length tssb'"
	  assume j_bound: "j < length tssb'"
	  assume i1_j: "i1  j"
	  assume ts_i1: "tssb'!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	  assume ts_j: "tssb'!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "(𝒪1  all_acquired sb1)  (𝒪j  all_acquired sbj)= {}"
	  proof (cases "i1=i")
	    case True
	    with i1_j have i_j: "ij" 
	      by simp
	    
	    from j_bound have j_bound': "j < length tssb"
	      by (simp add: tssb')
	    hence j_bound'': "j < length (map owned tssb)"
	      by simp	    
	    from ts_j i_j have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (simp add: tssb')
	    
	    from ownership_distinct [OF i_bound j_bound' i_j tssb_i ts_j']
	    show ?thesis
	      using tssb_i True ts_i1 i_bound 𝒪sb'
	      by (auto simp add: tssb' sb Ghostsb)
	  next
	    case False
	    note i1_i = this
	    from i1_bound have i1_bound': "i1 < length tssb"
	      by (simp add: tssb')
	    hence i1_bound'': "i1 < length (map owned tssb)"
	      by simp	    
	    from ts_i1 False have ts_i1': "tssb!i1 = (p1,is1,xs1,sb1,𝒟1,𝒪1,1)"
	      by (simp add: tssb')
	    show ?thesis
	    proof (cases "j=i")
	      case True
	      from ownership_distinct [OF i1_bound' i_bound  i1_i ts_i1' tssb_i]
	      show ?thesis
		using tssb_i True ts_j i_bound 𝒪sb'
		by (auto simp add: tssb' sb Ghostsb)
	    next
	      case False
	      from j_bound have j_bound': "j < length tssb"
		by (simp add: tssb')
	      from ts_j False have ts_j': "tssb!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
		by (simp add: tssb')
	      from ownership_distinct [OF i1_bound' j_bound' i1_j ts_i1' ts_j']
	      show ?thesis .
	    qed
	  qed
	qed
      qed

      have valid_sharing': "valid_sharing (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
      proof (intro_locales)
	show "outstanding_non_volatile_writes_unshared (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	proof (unfold_locales)
	  fix j pj "isj" "𝒪j" j 𝒟j acqj xsj sbj
	  assume j_bound: "j < length tssb'"
	  assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "non_volatile_writes_unshared (𝒮sb ⊕⇘WR ⊖⇘AL)  sbj"
	  proof (cases "i=j")
	    case True
	    with outstanding_non_volatile_writes_unshared [OF i_bound tssb_i]
	      i_bound jth tssb_i show ?thesis
	      by (clarsimp simp add: tssb' sb Ghostsb)
	  next
	    case False
	    from j_bound have j_bound': "j < length tssb"
	      by (auto simp add: tssb')
	    from jth False have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (auto simp add: tssb')
	    from j_bound jth i_bound False
	    have j: "non_volatile_writes_unshared 𝒮sb sbj"
	      apply -
	      apply (rule outstanding_non_volatile_writes_unshared)
	      apply (auto simp add: tssb')
	      done
	    from jth False have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (auto simp add: tssb')
	    from outstanding_non_volatile_writes_unshared [OF j_bound' jth']
	    have unshared: "non_volatile_writes_unshared 𝒮sb sbj".
	      
	    have "adom (𝒮sb ⊕⇘WR ⊖⇘AL) - dom 𝒮sb. a  outstanding_refs is_non_volatile_Writesb sbj"
	    proof -
	      {
		fix a 
		assume a_in: "a  dom (𝒮sb ⊕⇘WR ⊖⇘AL) - dom 𝒮sb"
		hence a_R: "a  R"
		  by clarsimp
		assume a_in_j: "a  outstanding_refs is_non_volatile_Writesb sbj"
		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 sbj"
		    by auto

		  moreover
		  with ownership_distinct [OF i_bound j_bound' False tssb_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_Writesb 
	      [OF unshared this]
	    show ?thesis .
	  qed
	qed
      next
	show "sharing_consis (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	proof (unfold_locales)  
	  fix j pj "isj" "𝒪j" j 𝒟j acqj xsj sbj
	  assume j_bound: "j < length tssb'"
	  assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "sharing_consistent (𝒮sb ⊕⇘WR ⊖⇘AL) 𝒪j sbj"
	  proof (cases "i=j")
	    case True
	    with i_bound jth tssb_i sharing_consis [OF i_bound tssb_i]
	    show ?thesis
	      by (clarsimp simp add: tssb' sb Ghostsb 𝒪sb')
	  next
	    case False
	    from j_bound have j_bound': "j < length tssb"
	      by (auto simp add: tssb')
	    from jth False have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (auto simp add: tssb')
	    from sharing_consis [OF j_bound' jth']
	    have consis: "sharing_consistent 𝒮sb 𝒪j sbj".
	    
	    have acq_cond: "all_acquired sbj  dom 𝒮sb - dom (𝒮sb ⊕⇘WR ⊖⇘AL) = {}"
	    proof -
	      {
		fix a
		assume a_acq: "a  all_acquired sbj" 
		assume "a  dom 𝒮sb" 
		assume a_L: "a  L"
		have False
		proof -
		  from ownership_distinct [OF i_bound j_bound' False tssb_i jth']
		  have "A  all_acquired sbj = {}"
		    by (auto simp add: sb Ghostsb)
		  with a_acq a_L L_subset
		  show False
		    by blast
		qed
	      }
	      thus ?thesis
		by auto
	    qed

	    have uns_cond: "all_unshared sbj  dom (𝒮sb ⊕⇘WR ⊖⇘AL) - dom 𝒮sb = {}"
	    proof -
	      {
		fix a
		assume a_uns: "a  all_unshared sbj" 
		assume "a  L" 
		assume a_R:  "a  R"
		have False
	        proof -
		  from unshared_acquired_or_owned [OF consis] a_uns
		  have "a  all_acquired sbj  𝒪j" by auto
		  with ownership_distinct [OF i_bound j_bound' False tssb_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: tssb')
	  qed
	qed
      next
	show "unowned_shared (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	proof (unfold_locales)
	  show "- ((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set tssb')  dom (𝒮sb ⊕⇘WR ⊖⇘AL)"
	  proof -
	    
	    have s: "((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set tssb') =
              ((λ(_,_, _, _,_, 𝒪,_). 𝒪) ` set tssb)  A - R"
	      
	      apply (unfold tssb' 𝒪sb') 
	      apply (rule acquire_release_ownership_nth_update [OF R_owned i_bound tssb_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 (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	proof 
	  fix j pj "isj" "𝒪j" j 𝒟j xsj sbj
	  assume j_bound: "j < length tssb'"
	  assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "𝒪j  read_only (𝒮sb ⊕⇘WR ⊖⇘AL) = {}"
	  proof (cases "i=j")
	    case True
	    from read_only_unowned [OF i_bound tssb_i] 
	    have "(𝒪sb  A - R )  read_only (𝒮sb ⊕⇘WR ⊖⇘AL) = {}"
	      by (auto simp add: in_read_only_convs )
	    with jth tssb_i i_bound True
	    show ?thesis
	      by (auto simp add: 𝒪sb' tssb')
	  next
	    case False
	    from j_bound have j_bound': "j < length tssb"
	      by (auto simp add: tssb')
	    with False jth have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (auto simp add: tssb')
	    from read_only_unowned [OF j_bound' jth']
	    have "𝒪j  read_only 𝒮sb = {}".
	    moreover
	    from ownership_distinct [OF i_bound j_bound' False tssb_i jth'] R_owned
	    have "(𝒪sb  A)  𝒪j = {}"
	      by (auto simp add: sb Ghostsb)
	    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 (𝒮sb ⊕⇘WR ⊖⇘AL) tssb'"
	proof 
	  fix j pj "isj" "𝒪j" j 𝒟j xsj sbj
	  assume j_bound: "j < length tssb'"
	  assume jth: "tssb' ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	  show "no_write_to_read_only_memory (𝒮sb ⊕⇘WR ⊖⇘AL) sbj"
	  proof (cases "i=j")
	    case True
	    with jth tssb_i i_bound no_outstanding_write_to_read_only_memory [OF i_bound tssb_i]
	    show ?thesis
	      by (auto simp add: sb tssb' Ghostsb)
	  next
	    case False
	    from j_bound have j_bound': "j < length tssb"
	      by (auto simp add: tssb')
	    with False jth have jth': "tssb ! j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)"
	      by (auto simp add: tssb')
	    from no_outstanding_write_to_read_only_memory [OF j_bound' jth']
	    have nw: "no_write_to_read_only_memory 𝒮sb sbj".

	    have "R  outstanding_refs is_Writesb sbj = {}"
	    proof -
	      note dist = ownership_distinct [OF i_bound j_bound' False tssb_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_Writesb sbj  𝒪sb = {}"
	        by auto
	      moreover
	      from outstanding_volatile_writes_unowned_by_others [OF j_bound'  i_bound 
		False [symmetric] jth' tssb_i ]
	      have "outstanding_refs is_volatile_Writesb sbj  𝒪sb = {}" 
	        by auto
	      ultimately have "outstanding_refs is_Writesb sbj  𝒪sb = {}" 
	        by (auto simp add: misc_outstanding_refs_convs)
	      with R_owned
	      show ?thesis by blast
	    qed
	    then
	    have "aoutstanding_refs is_Writesb sbj.
	      a  read_only (𝒮sb ⊕⇘WR ⊖⇘AL)  a  read_only 𝒮sb"
	      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 msb' tssb'"
      proof -
	from valid_reads [OF i_bound tssb_i]
	have "reads_consistent False (𝒪sb  A - R) msb sb'"
	  by (simp add: sb Ghostsb)
	from valid_reads_nth_update [OF i_bound this]
	show ?thesis by (simp add: msb' tssb' 𝒪sb')
      qed
      
      have valid_program_history': "valid_program_history tssb'"
      proof -	
	from valid_program_history [OF i_bound tssb_i]
	have "causal_program_history issb sb" .
	then have causal': "causal_program_history issb sb'"
	  by (simp add: sb Ghostsb causal_program_history_def)

	from valid_last_prog [OF i_bound tssb_i]
	have "last_prog psb sb = psb".
	hence "last_prog psb sb' = psb"
	  by (simp add: sb Ghostsb)

	from valid_program_history_nth_update [OF i_bound causal' this]
	show ?thesis
	  by (simp add: tssb')
      qed

      from is_sim
      have is_sim: "instrs (dropWhile (Not  is_volatile_Writesb) sb') @ issb =
	         is @ prog_instrs (dropWhile (Not  is_volatile_Writesb) sb')"
	by (simp add: sb Ghostsb suspends)


      have "(ts,m,𝒮) d* (ts,m,𝒮)" by blast
      moreover

      note flush_commute =
	flush_all_until_volatile_write_Ghostsb_commute [OF i_bound tssb_i [simplified sb Ghostsb]]

      have dist_R_L_A: "j p is 𝒪  𝒟 θ sb.
        j < length tssb  i j
        tssb ! j = (p, is, θ, sb, 𝒟, 𝒪, ) 
        (all_shared (takeWhile (Not  is_volatile_Writesb) sb)  
         all_unshared (takeWhile (Not  is_volatile_Writesb) sb)  
         all_acquired (takeWhile (Not  is_volatile_Writesb) sb))  (R  L  A) = {}"
      proof -
        {
          fix j pj isj 𝒪j j 𝒟j θj sbj x
	  assume j_bound: "j < length tssb"
          assume neq_i_j: "i  j"
	  assume jth: "tssb!j = (pj,isj, θj,sbj,𝒟j,𝒪j,j)"
	  assume x_shared: "x  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)  
                                 all_unshared (takeWhile (Not  is_volatile_Writesb) sbj)  
                                 all_acquired  (takeWhile (Not  is_volatile_Writesb) sbj)"
          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_Writesb) sbj)" "(dropWhile (Not  is_volatile_Writesb) sbj)"]
              all_unshared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" "(dropWhile (Not  is_volatile_Writesb) sbj)"]
              all_acquired_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" "(dropWhile (Not  is_volatile_Writesb) sbj)"]
            have "x  all_acquired sbj  𝒪j "
              by auto
            moreover
            from x_R_L_A R_owned L_subset
            have "x  all_acquired sb  𝒪sb"
              by (auto simp add: sb Ghostsb)
            moreover
            note ownership_distinct [OF i_bound j_bound neq_i_j tssb_i jth]
            ultimately show False by blast
          qed
        }
        thus ?thesis by blast
      qed

      {
	fix j pj isj 𝒪j j 𝒟j θj sbj x
	assume jth: "tssb!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	assume j_bound: "j < length tssb"
        assume neq: "i  j" 
        have "release (takeWhile (Not  is_volatile_Writesb) sbj)
                            (dom 𝒮sb  R - L) j
              = release (takeWhile (Not  is_volatile_Writesb) sbj)
                            (dom 𝒮sb) j"
        proof -
          {
            fix a
            assume a_in: "a  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)"
            have "(a  (dom 𝒮sb  R - L)) = (a  dom 𝒮sb)"
            proof -
              from ownership_distinct [OF i_bound j_bound neq  tssb_i jth]
                
              have A_dist: "A  (𝒪j  all_acquired sbj) = {}"
                by (auto simp add: sb Ghostsb)
              
              from  all_shared_acquired_or_owned [OF sharing_consis [OF j_bound jth]] a_in
                all_shared_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" 
                "(dropWhile (Not  is_volatile_Writesb) sbj)"]
              have a_in: "a  𝒪j  all_acquired sbj"
                by auto
              with ownership_distinct [OF i_bound j_bound neq  tssb_i jth]
              have "a  (𝒪sb  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 tssb".
      from sharing_consis_axioms have "sharing_consis 𝒮sb tssb".
      note share_commute = share_all_until_volatile_write_Ghostsb_commute [OF ownership_distinct tssb 
	sharing_consis 𝒮sb tssb i_bound tssb_i [simplified sb Ghostsb] dist_R_L_A]
      
      have "(tssb [i := (psb,issb, θsb, sb', 𝒟sb, 𝒪sb  A - R,augment_rels (dom 𝒮sb) R sb)],msb,𝒮sb')  (ts,m,𝒮)"
	apply (rule sim_config.intros) 
	apply    (simp add: m flush_commute)
	apply   (clarsimp simp add: 𝒮 𝒮sb' 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 Ghostsb sb' 𝒮sb'
	   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'
	  msb' 𝒮sb' tssb' 
	by (auto simp del: fun_upd_apply simp add: 𝒪sb' sb')
    qed
 next
    case (Program i psb "issb" θsb sb 𝒟sb 𝒪sb sb psb' mis)
    then obtain
      tssb': "tssb' = tssb[i := (psb', issb@mis, θsb, sb@[Progsb psb psb' mis], 𝒟sb, 𝒪sb,sb)]" and
      i_bound: "i < length tssb" and
      tssb_i: "tssb ! i = (psb, issb,θsb,sb, 𝒟sb, 𝒪sb,sb)" and
      prog: "θsb psb p (psb',mis)" and
      𝒮sb': "𝒮sb'=𝒮sb" and
      msb': "msb'=msb"
      by auto

    from sim obtain 
      m: "m = flush_all_until_volatile_write tssb msb" and
      𝒮: "𝒮 = share_all_until_volatile_write tssb 𝒮sb" and
      leq: "length tssb = length ts" and
      ts_sim: "i<length tssb.
           let (p, issb, θ, sb, 𝒟sb, 𝒪sb,) = tssb ! i;
               suspends = dropWhile (Not  is_volatile_Writesb) sb
           in  is 𝒟. instrs suspends @ issb = is @ prog_instrs suspends  
                          𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb sb  {}) 
               ts ! i =
                   (hd_prog p suspends, 
                    is,
                    θ |` (dom θ - read_tmps suspends), (),
                    𝒟, 
                    acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪sb,
                    release (takeWhile (Not  is_volatile_Writesb) sb) (dom 𝒮sb) )"
      by cases blast

    from i_bound leq have i_bound': "i < length ts"
      by auto

    have split_sb: "sb = takeWhile (Not  is_volatile_Writesb) sb @ dropWhile (Not  is_volatile_Writesb) sb"
      (is "sb = ?take_sb@?drop_sb")
      by simp

    from ts_sim [rule_format, OF i_bound] tssb_i obtain suspends "is" 𝒟 where
      suspends: "suspends = dropWhile (Not  is_volatile_Writesb) sb" and
      is_sim: "instrs suspends @ issb = is @ prog_instrs suspends" and
      𝒟: "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb sb  {})" and
      ts_i: "ts ! i =
          (hd_prog psb suspends, is,
           θsb |` (dom θsb - read_tmps suspends), (), 𝒟, acquired True ?take_sb 𝒪sb,
           release ?take_sb (dom 𝒮sb) sb)"
      by (auto simp add: Let_def)

    from prog_step_preserves_valid [OF i_bound tssb_i prog valid]
    have valid': "valid tssb'"
      by (simp add: tssb')

    have valid_own': "valid_ownership 𝒮sb' tssb'"
    proof (intro_locales)
      show "outstanding_non_volatile_refs_owned_or_read_only 𝒮sb' tssb'"
      proof -
	from outstanding_non_volatile_refs_owned_or_read_only [OF i_bound tssb_i] 
	have "non_volatile_owned_or_read_only False 𝒮sb 𝒪sb (sb@[Progsb psb psb' 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: tssb' 𝒮sb')
      qed
    next
      show "outstanding_volatile_writes_unowned_by_others tssb'"
      proof -
	have out: "outstanding_refs is_volatile_Writesb (sb@[Progsb psb psb' mis])  
	      outstanding_refs is_volatile_Writesb sb"
	  by (auto simp add: outstanding_refs_conv )
	from outstanding_volatile_writes_unowned_by_others_store_buffer 
	[OF i_bound tssb_i this]
	show ?thesis by (simp add: tssb' all_acquired_append)
      qed
    next
      show "read_only_reads_unowned tssb'"
      proof -
	have ro: "read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) (sb@[Progsb psb psb' mis])) 𝒪sb)
	  (dropWhile (Not  is_volatile_Writesb) (sb@[Progsb psb psb' mis]))
	   read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪sb)
	  (dropWhile (Not  is_volatile_Writesb) sb)"
	  apply (case_tac "outstanding_refs (is_volatile_Writesb) sb = {}")
	  apply (simp_all add: outstanding_vol_write_take_drop_appends
	    acquired_append read_only_reads_append )
	  done
	have "𝒪sb  all_acquired (sb@[Progsb psb psb' mis])  𝒪sb  all_acquired sb"
	  by (auto simp add: all_acquired_append)
	from read_only_reads_unowned_nth_update [OF i_bound tssb_i ro this] 
	show ?thesis
	  by (simp add: tssb' )
      qed
    next
      show "ownership_distinct tssb'"
      proof -
	from ownership_distinct_instructions_read_value_store_buffer_independent 
	[OF i_bound tssb_i, where sb'="(sb@[Progsb psb psb' mis])"]
	show ?thesis by (simp add: tssb' all_acquired_append)
      qed
    qed

    from valid_last_prog [OF i_bound tssb_i]
    have last_prog: "last_prog psb sb = psb".
    
    have valid_hist': "valid_history program_step tssb'"
    proof -
      from valid_history [OF i_bound tssb_i]
      have "history_consistent θsb (hd_prog psb sb) sb".
      from history_consistent_append_Progsb [OF prog this last_prog]
      have hist_consis': "history_consistent θsb (hd_prog psb' (sb@[Progsb psb psb' mis])) 
        (sb@[Progsb psb psb' mis])".
      from valid_history_nth_update [OF i_bound this]
      show ?thesis by (simp add: tssb')
    qed


    have valid_reads': "valid_reads msb tssb'"
    proof -
      from valid_reads [OF i_bound tssb_i]
      have "reads_consistent False 𝒪sb msb sb" .
      from reads_consistent_snoc_Progsb [OF this] 
      have "reads_consistent False 𝒪sb msb  (sb@[Progsb psb psb' mis])".
      from valid_reads_nth_update [OF i_bound this]
      show ?thesis by (simp add: tssb')
    qed

    have valid_sharing': "valid_sharing 𝒮sb' tssb'"
    proof (intro_locales)	
      from outstanding_non_volatile_writes_unshared [OF i_bound tssb_i] 
      have "non_volatile_writes_unshared 𝒮sb (sb@[Progsb psb psb' 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 𝒮sb' tssb'"
	by (simp add: tssb' 𝒮sb')
    next
      from sharing_consis [OF i_bound tssb_i]
      have "sharing_consistent 𝒮sb 𝒪sb (sb@[Progsb psb psb' mis])"
	by (auto simp add: sharing_consistent_append)
      from sharing_consis_nth_update [OF i_bound this]
      show "sharing_consis 𝒮sb' tssb'"
	by (simp add: tssb' 𝒮sb')
    next
      from read_only_unowned_nth_update [OF i_bound read_only_unowned [OF i_bound tssb_i] ]
      show "read_only_unowned 𝒮sb' tssb'"
	by (simp add: 𝒮sb' tssb')
    next
      from unowned_shared_nth_update [OF i_bound tssb_i subset_refl]
      show "unowned_shared 𝒮sb' tssb'"
	by (simp add: tssb' 𝒮sb')
    next
      from no_outstanding_write_to_read_only_memory [OF i_bound tssb_i] 

      have "no_write_to_read_only_memory 𝒮sb (sb @ [Progsb psb psb' 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 𝒮sb' tssb'"
	by (simp add: 𝒮sb' tssb')
    qed

    have tmps_distinct': "tmps_distinct tssb'"
    proof (intro_locales)
      from load_tmps_distinct [OF i_bound tssb_i]
      have "distinct_load_tmps issb".
      with distinct_load_tmps_prog_step [OF i_bound tssb_i prog valid] 
      have "distinct_load_tmps (issb@mis)" 
	by (auto simp add: distinct_load_tmps_append)
	
      from load_tmps_distinct_nth_update [OF i_bound this]
      show "load_tmps_distinct tssb'"
	by (simp add: tssb')
    next
      from read_tmps_distinct [OF i_bound tssb_i]
      have "distinct_read_tmps (sb@[Progsb psb psb' mis])"
	by (simp add: distinct_read_tmps_append)
      from read_tmps_distinct_nth_update [OF i_bound this]
      show "read_tmps_distinct tssb'"
	by (simp add: tssb')
    next
      from load_tmps_read_tmps_distinct [OF i_bound tssb_i]
	distinct_load_tmps_prog_step [OF i_bound tssb_i prog valid] 
      have "load_tmps (issb@mis)  read_tmps (sb@[Progsb psb psb' 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 tssb'" by (simp add: tssb')
    qed

    have valid_dd': "valid_data_dependency tssb'"
    proof -
      from data_dependency_consistent_instrs [OF i_bound tssb_i]
      have "data_dependency_consistent_instrs (dom θsb) issb".
      with valid_data_dependency_prog_step [OF i_bound tssb_i prog valid]  
	   load_tmps_write_tmps_distinct [OF i_bound tssb_i]
      obtain
	"data_dependency_consistent_instrs (dom θsb) (issb@mis)"
	"load_tmps (issb@mis)  (fst ` write_sops (sb@[Progsb psb psb' 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: tssb')
    qed

    have load_tmps_fresh': "load_tmps_fresh tssb'"
    proof -
      
      from load_tmps_fresh [OF i_bound tssb_i] 
      load_tmps_fresh_prog_step [OF i_bound tssb_i prog valid]
      have "load_tmps (issb@mis)  dom θsb = {}"
	by (auto simp add: load_tmps_append)
      from load_tmps_fresh_nth_update [OF i_bound this]
      show ?thesis
	by (simp add:  tssb')
    qed

    have enough_flushs': "enough_flushs tssb'"
    proof -
      from clean_no_outstanding_volatile_Writesb [OF i_bound tssb_i]
      have "¬ 𝒟sb  outstanding_refs is_volatile_Writesb (sb@[Progsb psb psb' mis]) = {}"
	by (auto simp add: outstanding_refs_append)

      from enough_flushs_nth_update [OF i_bound this]
      show ?thesis
	by (simp add: tssb')
    qed

    have valid_sops': "valid_sops tssb'"
    proof -
      from valid_store_sops [OF i_bound tssb_i] valid_sops_prog_step [OF prog] 
	valid_implies_valid_prog[OF i_bound tssb_i valid]
      have valid_store: "sopstore_sops (issb@mis). valid_sop sop"
	by (auto simp add: store_sops_append)

      from valid_write_sops [OF i_bound tssb_i]
      have "sopwrite_sops (sb@[Progsb psb psb' 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: tssb')
    qed

    have valid_program_history':"valid_program_history tssb'"
    proof -	
      from valid_program_history [OF i_bound tssb_i]
      have "causal_program_history issb sb" .
      from causal_program_history_Progsb [OF this]
      have causal': "causal_program_history (issb@mis) (sb@[Progsb psb psb' mis])".
      from last_prog_append_Progsb
      have "last_prog psb' (sb@[Progsb psb psb' mis]) = psb'".
      from valid_program_history_nth_update [OF i_bound causal' this]
      show ?thesis
	by (simp add: tssb')
    qed

    show ?thesis
    proof (cases "outstanding_refs is_volatile_Writesb sb = {}")
      case True
      from True have flush_all: "takeWhile (Not  is_volatile_Writesb) sb = sb"
	by (auto simp add: outstanding_refs_conv)
      
      from True have suspend_nothing: "dropWhile (Not  is_volatile_Writesb) sb = []"
	by (auto simp add: outstanding_refs_conv)

      hence suspends_empty: "suspends = []"
	by (simp add: suspends)

      from suspends_empty is_sim have "is": "is = issb"
	by (simp)

      from ts_i have ts_i: "ts ! i = (psb, issb, θsb, (), 
	𝒟, acquired True ?take_sb 𝒪sb,release ?take_sb (dom 𝒮sb) sb)"
	by (simp add: suspends_empty "is")

      from direct_computation.Program [OF i_bound' ts_i prog]
      have "(ts,m,𝒮) d (ts[i := (psb', issb @ mis, θsb, (),
	𝒟, acquired True ?take_sb 𝒪sb,release ?take_sb (dom 𝒮sb) sb)], m, 𝒮)".
    
      moreover

      note flush_commute = flush_all_until_volatile_write_append_Progsb_commute [OF i_bound tssb_i]

      from True
      have suspend_nothing':
	"(dropWhile (Not  is_volatile_Writesb) (sb @ [Progsb psb psb' mis])) = []"
	by (auto simp add: outstanding_refs_conv)

      note share_commute =
	share_all_until_volatile_write_update_sb [OF share_append_Progsb i_bound tssb_i]

      from 𝒟
      have 𝒟': "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb (sb@[Progsb psb psb' mis])   {})"
	by (auto simp: outstanding_refs_append)

      have "(tssb [i := (psb',issb@mis, θsb, sb@[Progsb psb psb' mis], 𝒟sb, 𝒪sb,sb)],
              msb,𝒮sb')  
              (ts[i:=(psb', issb @ mis, θsb, (), 𝒟, 
                  acquired True (takeWhile (Not  is_volatile_Writesb) 
                    (sb@[Progsb psb psb' mis])) 𝒪sb,
                   release (sb@[Progsb psb psb' mis])  (dom 𝒮sb) sb )],m,𝒮)"
	apply (rule sim_config.intros) 
	apply    (simp add: m flush_commute)
	apply   (clarsimp simp add: 𝒮 𝒮sb' 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' Progsb 𝒮sb' 
          release_append_Progsb release_append
	   split: if_split_asm)
	done	
      ultimately show ?thesis
	using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' msb'
	  valid_dd' valid_sops' load_tmps_fresh' enough_flushs' valid_sharing' 
	  valid_program_history' valid'  
	  𝒮sb' tssb'  
	by (auto simp del: fun_upd_apply simp add: acquired_append_Progsb release_append_Progsb release_append flush_all) 
    next
      case False

      then obtain r where r_in: "r  set sb" and volatile_r: "is_volatile_Writesb r"
	by (auto simp add: outstanding_refs_conv)
      from takeWhile_dropWhile_real_prefix 
      [OF r_in, of  "(Not  is_volatile_Writesb)", simplified, OF volatile_r] 
      obtain a' v' sb'' sop' A' L' R' W' where
	sb_split: "sb = takeWhile (Not  is_volatile_Writesb) sb @ Writesb True a' sop' v' A' L' R' W'# sb''" 
	and
	drop: "dropWhile (Not  is_volatile_Writesb) sb = Writesb 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 = Writesb 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_Progsb_commute [OF i_bound tssb_i]

      have "Writesb True a' sop' v' A' L' R' W'  set sb"
	by (subst sb_split) auto

      from dropWhile_append1 [OF this, of "(Not  is_volatile_Writesb)"]
      have drop_app_comm:
	  "(dropWhile (Not  is_volatile_Writesb) (sb @ [Progsb psb psb' mis])) =
                dropWhile (Not  is_volatile_Writesb) sb @ [Progsb psb psb' mis]"
	by simp

      note share_commute =
	share_all_until_volatile_write_update_sb [OF share_append_Progsb i_bound tssb_i]

      from 𝒟
      have 𝒟': "𝒟sb = (𝒟  outstanding_refs is_volatile_Writesb (sb@[Progsb psb psb' mis])   {})"
	by (auto simp: outstanding_refs_append)
      have "(tssb [i := (psb',issb@mis,θsb, sb@[Progsb psb psb' mis], 𝒟sb, 𝒪sb,sb)],
              msb,𝒮sb')  
              (ts,m,𝒮)"
	apply (rule sim_config.intros) 
	apply    (simp add: m flush_commute)
	apply   (clarsimp  simp add: 𝒮 𝒮sb' 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 Progsb
	  drop_app_comm instrs_append prog_instrs_append  
	  read_tmps_append hd_prog_append_Progsb acquired_append_Progsb release_append_Progsb release_append 𝒮sb'
	   split: if_split_asm)
	done	

      ultimately show ?thesis
	using valid_own' valid_hist' valid_reads' valid_sharing' tmps_distinct' msb'
	  valid_dd' valid_sops' load_tmps_fresh' enough_flushs' valid_sharing' 
	  valid_program_history' valid'
	  𝒮sb' tssb' 
	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: "(tssb,msb,𝒮sb) sbh* (tssb',msb',𝒮sb')"
  assumes valid_own: "valid_ownership 𝒮sb tssb"
  assumes valid_sb_reads: "valid_reads msb tssb"
  assumes valid_hist: "valid_history program_step tssb"
  assumes valid_sharing: "valid_sharing 𝒮sb tssb"
  assumes tmps_distinct: "tmps_distinct tssb"
  assumes valid_sops: "valid_sops tssb"
  assumes valid_dd: "valid_data_dependency tssb"
  assumes load_tmps_fresh: "load_tmps_fresh tssb"
  assumes enough_flushs: "enough_flushs tssb"
  assumes valid_program_history: "valid_program_history tssb"
  assumes valid: "valid tssb"
  shows "ts 𝒮 m. (tssb,msb,𝒮sb)  (ts,m,𝒮)  safe_reach_direct safe_delayed (ts,m,𝒮) 
         valid_ownership 𝒮sb' tssb'  valid_reads msb' tssb'  valid_history program_step tssb' 
         valid_sharing 𝒮sb' tssb'  tmps_distinct tssb'  valid_data_dependency tssb' 
         valid_sops tssb'  load_tmps_fresh tssb'  enough_flushs tssb' 
         valid_program_history tssb'  valid tssb' 
           (ts' m' 𝒮'. (ts,m,𝒮) d* (ts',m',𝒮')  (tssb',msb',𝒮sb')  (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 tssb  msb 𝒮sb tssb''  msb'' 𝒮sb'')
  note first = (tssb, msb, 𝒮sb) sbh (tssb'', msb'', 𝒮sb'')
  note sim = (tssb, msb, 𝒮sb)  (ts, m, 𝒮)
  note safe_reach = safe_reach_direct safe_delayed (ts, m, 𝒮)
  note valid_own = valid_ownership 𝒮sb tssb
  note valid_reads = valid_reads msb tssb
  note valid_hist = valid_history program_step tssb
  note valid_sharing = valid_sharing 𝒮sb tssb
  note tmps_distinct = tmps_distinct tssb
  note valid_sops = valid_sops tssb
  note valid_dd = valid_data_dependency tssb
  note load_tmps_fresh = load_tmps_fresh tssb
  note enough_flushs = enough_flushs tssb
  note valid_prog_hist = valid_program_history tssb
  note valid = valid tssb
  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 𝒮sb'' tssb''" and
    valid_reads'': "valid_reads msb'' tssb''" and
    valid_hist'': "valid_history program_step tssb''" and
    valid_sharing'': "valid_sharing 𝒮sb'' tssb''" and
    tmps_dist'': "tmps_distinct tssb''" and
    valid_dd'': "valid_data_dependency tssb''" and
    valid_sops'': "valid_sops tssb''" and
    load_tmps_fresh'': "load_tmps_fresh tssb''" and
    enough_flushs'': "enough_flushs tssb''" and
    valid_prog_hist'': "valid_program_history tssb''"and
    valid'': "valid tssb''" and
    steps: "(ts, m, 𝒮) d* (ts'', m'', 𝒮'')" and
    sim: "(tssb'', msb'',𝒮sb'')  (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 𝒮sb' tssb'" "valid_reads msb' tssb'" "valid_history program_step tssb'"
    "valid_sharing 𝒮sb' tssb'" "tmps_distinct tssb'" "valid_data_dependency tssb'"
    "valid_sops tssb'" "load_tmps_fresh tssb'" "enough_flushs tssb'"
    "valid_program_history tssb'" "valid tssb'" and
    last: "(ts'', m'', 𝒮'') d* (ts', m', 𝒮')" and
    sim': "(tssb', msb',𝒮sb')  (ts', m',𝒮')"
    by blast

  note steps also note last
  finally show ?case
    using valid sim'
    by blast
qed

(* FIXME: move up *)
sublocale initialsb  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: "(tssb,msb,𝒮sb) sbh* (tssb',msb',𝒮sb')"
assumes init: "initialsb tssb 𝒮sb"
assumes valid: "valid tssb" (* FIXME: move into initial ?*)
assumes sim: "(tssb,msb,𝒮sb)  (ts,m,𝒮)"
assumes safe: "safe_reach_direct safe_free_flowing (ts,m,𝒮)"
shows "ts' m' 𝒮'. (ts,m,𝒮) d* (ts',m',𝒮')  
                (tssb',msb',𝒮sb')  (ts',m',𝒮')"
proof -
  from init interpret ini: initialsb tssb 𝒮sb .
  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 𝒮sb tssb" .
  from local.ini.valid_reads_axioms have "valid_reads msb tssb".
  from local.ini.valid_history_axioms have "valid_history program_step tssb".
  from local.ini.valid_sharing_axioms have "valid_sharing 𝒮sb tssb".
  from local.ini.tmps_distinct_axioms have "tmps_distinct tssb".
  from local.ini.valid_sops_axioms have "valid_sops tssb".
  from local.ini.valid_data_dependency_axioms have "valid_data_dependency tssb".  
  from local.ini.load_tmps_fresh_axioms have "load_tmps_fresh tssb".
  from local.ini.enough_flushs_axioms have "enough_flushs tssb".
  from local.ini.valid_program_history_axioms have "valid_program_history tssb".
  from concurrent_direct_steps_simulates_store_buffer_history_steps [OF exec_sb 
    valid_ownership 𝒮sb tssb
    valid_reads msb tssb valid_history program_step tssb
    valid_sharing 𝒮sb tssb tmps_distinct tssb valid_sops tssb
    valid_data_dependency tssb load_tmps_fresh tssb enough_flushs tssb
   valid_program_history tssb valid sim safe_delayed]
  show ?thesis by auto
qed





lemma filter_is_Writesb_Cons_Writesb: "filter is_Writesb xs = Writesb volatile a sop v A L R W#ys
       rs rws. (r  set rs. is_Readsb r  is_Progsb r  is_Ghostsb r)  
              xs=rs@Writesb volatile a sop v A L R W#rws  ys=filter is_Writesb rws"
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x xs)
  note feq = filter is_Writesb (x#xs) = Writesb volatile a sop v A L R W# ys
  show ?case
  proof (cases x)
    case (Writesb 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_Writesb xs"
      by auto
    thus ?thesis
      apply -
      apply (rule_tac x="[]" in exI)
      apply (rule_tac x="xs" in exI)
      apply (simp add: Writesb)
      done
  next
    case (Readsb volatile' a' t' v')
    from feq have "filter is_Writesb xs = Writesb volatile a sop v A L R W#ys"
      by (simp add: Readsb)
    from Cons.hyps [OF this] obtain rs rws where
      "r  set rs. is_Readsb r  is_Progsb r  is_Ghostsb r" and
      "xs=rs @ Writesb volatile a sop v A L R W# rws" and
      "ys=filter is_Writesb rws"
      by clarsimp
    then show ?thesis
      apply -
      apply (rule_tac x="Readsb volatile' a' t' v'#rs" in exI)
      apply (rule_tac x="rws" in exI)
      apply (simp add: Readsb)
      done
  next
    case (Progsb p1 p2 mis)
    from feq have "filter is_Writesb xs = Writesb volatile a sop v A L R W#ys"
      by (simp add: Progsb)
    from Cons.hyps [OF this] obtain rs rws where
      "r  set rs. is_Readsb r  is_Progsb r  is_Ghostsb r" and
      "xs=rs @ Writesb volatile a sop v A L R W# rws" and
      "ys=filter is_Writesb rws"
      by clarsimp
    then show ?thesis
      apply -
      apply (rule_tac x="Progsb p1 p2 mis#rs" in exI)
      apply (rule_tac x="rws" in exI)
      apply (simp add: Progsb)
      done
    next
    case (Ghostsb A' L' R' W')
    from feq have "filter is_Writesb xs = Writesb volatile a sop v A L R W # ys"
      by (simp add: Ghostsb)
    from Cons.hyps [OF this] obtain rs rws where
      "r  set rs. is_Readsb r  is_Progsb r  is_Ghostsb r" and
      "xs=rs @ Writesb volatile a sop v A L R W# rws" and
      "ys=filter is_Writesb rws"
      by clarsimp
    then show ?thesis
      apply -
      apply (rule_tac x="Ghostsb A' L' R' W'#rs" in exI)
      apply (rule_tac x="rws" in exI)
      apply (simp add: Ghostsb)
      done
  qed
qed

lemma filter_is_Writesb_empty: "filter is_Writesb xs = []
       (r  set xs. is_Readsb r  is_Progsb r  is_Ghostsb r)"
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x xs)
  note feq = filter is_Writesb (x#xs) = []
  show ?case
  proof (cases x)
    case (Writesb volatile' a' v')
    with feq have False
      by simp
    thus ?thesis ..
  next
    case (Readsb a' v')
    from feq have "filter is_Writesb xs = []"
      by (simp add: Readsb)
    from Cons.hyps [OF this] obtain 
      "r  set xs. is_Readsb r  is_Progsb r  is_Ghostsb r" 
      by clarsimp
    then show ?thesis
      by (simp add: Readsb)
  next
    case (Progsb p2 p2 mis)
    from feq have "filter is_Writesb xs = []"
      by (simp add: Progsb)
    from Cons.hyps [OF this] obtain 
      "r  set xs. is_Readsb r  is_Progsb r  is_Ghostsb r" 
      by clarsimp
    then show ?thesis
      by (simp add: Progsb)
  next
    case (Ghostsb A' L' R' W')
    from feq have "filter is_Writesb xs = []"
      by (simp add: Ghostsb)
    from Cons.hyps [OF this] obtain 
      "r  set xs. is_Readsb r  is_Progsb r  is_Ghostsb r" 
      by clarsimp
    then show ?thesis
      by (simp add: Ghostsb)
  qed
qed

lemma flush_reads_program: "𝒪 𝒮  .
  r  set sb. is_Readsb r  is_Progsb r  is_Ghostsb r  
𝒪' ℛ' 𝒮'. (m,sb,𝒪,,𝒮) f* (m,[],𝒪',ℛ',𝒮')"      
proof (induct sb)
  case Nil thus ?case by auto
next
  case (Cons x sb)
  note rset (x # sb). is_Readsb r  is_Progsb r  is_Ghostsb r
  then obtain  x: "is_Readsb x  is_Progsb x  is_Ghostsb x" and sb: "rset sb. is_Readsb r  is_Progsb r  is_Ghostsb r"
    by (cases x) auto

  
  {
    assume "is_Readsb x"
    then obtain volatile a t v where x: "x=Readsb volatile a t v"
      by (cases x) auto
    
    have "(m,Readsb volatile a t v#sb,𝒪,,𝒮) f (m,sb,𝒪,,𝒮)"
      by (rule Readsb)
    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_Progsb x"
    then obtain p1 p2 mis  where x: "x=Progsb p1 p2 mis"
      by (cases x) auto
    
    have "(m,Progsb p1 p2 mis#sb,𝒪,,𝒮) f (m,sb,𝒪,,𝒮)"
      by (rule Progsb)
    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_Ghostsb x"
    then obtain A L R W where x: "x=Ghostsb A L R W"
      by (cases x) auto
    
    have "(m,Ghostsb A L R W#sb,𝒪,,𝒮) f (m,sb,𝒪  A - R,augment_rels (dom 𝒮) R ,𝒮 ⊕⇘WR ⊖⇘AL)"
      by (rule Ghost)
    also
    from Cons.hyps [OF sb] obtain 𝒪' 𝒮' ℛ' acq'
    where "(m, sb,𝒪  A - R ,augment_rels (dom 𝒮) R ,𝒮 ⊕⇘WR  ⊖⇘AL) 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 (Writesb volatile a sop v A L R W)
  from flush_step.Writesb  [OF refl refl refl, of m volatile a sop v A L R W sb 𝒪  𝒮]
  show ?thesis
    by (auto simp add: Writesb)
next
  case (Readsb volatile a t v)
  from flush_step.Readsb [of m volatile a t v sb 𝒪  𝒮]
  show ?thesis
    by (auto simp add: Readsb)
next
  case (Progsb p1 p2 mis)
  from flush_step.Progsb [of m p1 p2 mis sb 𝒪  𝒮]  
  show ?thesis
    by (auto simp add: Progsb)
next
  case (Ghostsb A L R W)
  from flush_step.Ghost [of m A L R W sb 𝒪  𝒮]
  show ?thesis
    by (auto simp add: Ghostsb)
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 (Writesb volatile A R W L a sop v)
    then obtain sb: "sb=Writesb 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 𝒮 ⊕⇘WR ⊖⇘AL else 𝒮)"
      by auto
    have "(m,Writesb 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 𝒮 ⊕⇘WR ⊖⇘AL else 𝒮)"
      apply (rule flush_step.Writesb)
      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 (Readsb volatile a t v)
    then obtain sb: "sb=Readsb volatile a t v #sb''" and m'': "m''=m" 
      and 𝒪'': "𝒪''=𝒪" and 𝒮'': "𝒮''=𝒮" and ℛ'': "ℛ''=" 
      by auto
    have "(m,Readsb volatile a t v#sb''@xs,𝒪,,𝒮) f (m,sb''@xs,𝒪,,𝒮)"
      by (rule flush_step.Readsb)
    hence "(m,sb@xs,𝒪,,𝒮) f (m'',sb''@xs,𝒪'',ℛ'',𝒮'')"
      by (simp add: sb m'' 𝒪'' ℛ'' 𝒮'' )
    also note append_rest
    finally show ?thesis .
  next
    case (Progsb p1 p2 mis)
    then obtain sb: "sb=Progsb p1 p2 mis#sb''" and m'': "m''=m"
      and 𝒪'': "𝒪''=𝒪" and 𝒮'': "𝒮''=𝒮" and ℛ'': "ℛ''=" 
      by auto
    have "(m,Progsb p1 p2 mis#sb''@xs,𝒪,,𝒮) f (m,sb''@xs,𝒪,,𝒮)"
      by (rule flush_step.Progsb)
    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=Ghostsb A L R W#sb''" and m'': "m''=m"
      and 𝒪'': "𝒪''=𝒪  A - R" and 𝒮'': "𝒮''=𝒮 ⊕⇘WR ⊖⇘AL" and 
      ℛ'': "ℛ''=augment_rels (dom 𝒮) R "
      by auto
    have "(m,Ghostsb A L R W#sb''@xs,𝒪,,𝒮) f (m,sb''@xs,𝒪  A - R,augment_rels (dom 𝒮) R ,𝒮 ⊕⇘WR ⊖⇘AL)"
      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
(*
theorem flush_simulates_filter_writes:
  assumes step: "(m,sb,𝒪,ℛ,𝒮) →f (m',sb',𝒪',ℛ',𝒮')"
  shows "⋀sbh 𝒪hh 𝒮h. sb=filter is_Writesb sbh 
          ⟹ 
          ∃sbh' 𝒪h' ℛh' 𝒮h'. (m,sbh,𝒪h,ℛh,𝒮h) →f* (m',sbh',𝒪h',ℛh',𝒮h') ∧ 
  sb'=filter is_Writesb sbh'"
using step
proof (induct rule: flush_step_induct)
  case (Writesb 𝒪' volatile 𝒪 A R 𝒮' 𝒮 W L ℛ' ℛ  m a D f v sb)
  note filter_Writesb = `Writesb volatile a (D,f) v A L R W# sb = filter is_Writesb sbh`
  note 𝒪' = `𝒪' = (if volatile then 𝒪 ∪ A - R else 𝒪)`
  note ℛ' = `ℛ'= (if volatile then empty else ℛ)`
  note 𝒮' = `𝒮' = (if volatile then 𝒮 ⊕W R ⊖A L else 𝒮)`
  from filter_is_Writesb_Cons_Writesb [OF filter_Writesb [symmetric]]
  obtain rs rws where
    rs_reads: "∀r∈set rs. is_Readsb r ∨ is_Progsb r ∨ is_Ghostsb r" and
    sbh: "sbh = rs @ Writesb volatile a (D,f) v A L R W# rws" and 
    sb: "sb = filter is_Writesb rws"
    by blast

  from flush_reads_program [OF rs_reads] obtain 𝒪h' ℛh' 𝒮h' acqh'
  where "(m, rs,𝒪h,ℛh,𝒮h) →f* (m, [],𝒪h',ℛh',𝒮h')" by blast
  from flush_append [OF this]
  have "(m, rs@Writesb volatile a (D,f) v A L R W# rws,𝒪h,ℛh,𝒮h) →f* 
       (m, Writesb volatile a (D,f) v A L R W# rws,𝒪h',ℛh',𝒮h')"
    by simp
  also
  from flush_step.Writesb [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, Writesb 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 "(m, sbh,𝒪h,ℛh,𝒮h) →f* (m(a:=v), rws,𝒪h'',ℛh'',𝒮h'')"
    by (simp add: sbh sb)
  with sb show ?case
    by blast
next
  case (Readsb m volatile a t v sb) 
  note `Readsb volatile a t v # sb = filter is_Writesb sbh`
  from this [symmetric]
  have r_in: "Readsb volatile a t v ∈ set (filter is_Writesb sbh)"
    by auto
  have "∀r ∈ set (filter is_Writesb sbh). is_Writesb r"
    by auto
  from this [rule_format, OF r_in]
  have False by simp
  thus ?case ..
next
  case (Progsb m p1 p2 mis sb)
  note `Progsb p1 p2 mis # sb = filter is_Writesb sbh`
  from this [symmetric]
  have r_in: "Progsb p1 p2 mis ∈ set (filter is_Writesb sbh)"
    by auto
  have "∀r ∈ set (filter is_Writesb sbh). is_Writesb r"
    by auto
  from this [rule_format, OF r_in]
  have False by simp
  thus ?case ..
next
  case (Ghost m A L R W sb)
  note `Ghostsb A L R W# sb = filter is_Writesb sbh`
  from this [symmetric]
  have r_in: "Ghostsb A L R W∈ set (filter is_Writesb sbh)"
    by auto
  have "∀r ∈ set (filter is_Writesb sbh). is_Writesb r"
    by auto
  from this [rule_format, OF r_in]
  have False by simp
  thus ?case ..
qed
*)
(* FIXME: move up *)
lemmas store_buffer_step_induct =  
  store_buffer_step.induct [split_format (complete),
  consumes 1, case_names SBWritesb]
theorem flush_simulates_filter_writes:
  assumes step: "(m,sb,𝒪,,𝒮) w (m',sb',𝒪',ℛ',𝒮')"
  shows "sbh 𝒪h h 𝒮h. sb=filter is_Writesb sbh 
           
          sbh' 𝒪h' h' 𝒮h'. (m,sbh,𝒪h,h,𝒮h) f* (m',sbh',𝒪h',h',𝒮h')  
  sb'=filter is_Writesb sbh'  (sb'=[]  sbh'=[])"
using step
proof (induct rule: store_buffer_step_induct)
  case (SBWritesb m volatile a D f v A L R W sb 𝒪  𝒮)
  note filter_Writesb = Writesb volatile a (D,f) v A L R W# sb = filter is_Writesb sbh
  
  from filter_is_Writesb_Cons_Writesb [OF filter_Writesb [symmetric]]
  obtain rs rws where
    rs_reads: "rset rs. is_Readsb r  is_Progsb r  is_Ghostsb r" and
    sbh: "sbh = rs @ Writesb volatile a (D,f) v A L R W# rws" and 
    sb: "sb = filter is_Writesb rws"
    by blast

  from flush_reads_program [OF rs_reads] obtain 𝒪h' h' 𝒮h' acqh'
  where "(m, rs,𝒪h,h,𝒮h) f* (m, [],𝒪h',h',𝒮h')" by blast
  from flush_append [OF this]
  have "(m, rs@Writesb volatile a (D,f) v A L R W# rws,𝒪h,h,𝒮h) f* 
       (m, Writesb volatile a (D,f) v A L R W# rws,𝒪h',h',𝒮h')"
    by simp
  also
  from flush_step.Writesb [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, Writesb 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, sbh,𝒪h,h,𝒮h) f* (m(a:=v), rws,𝒪h'',h'',𝒮h'')"
    by (simp add: sbh sb)
  show ?case
  proof (cases "sb")
    case Cons
    with steps sb show ?thesis 
      by fastforce
  next  
    case Nil
    from filter_is_Writesb_empty [OF sb [simplified Nil, symmetric]]
    have "rset rws. is_Readsb r  is_Progsb r  is_Ghostsb r".
    from flush_reads_program [OF this] obtain 𝒪h''' h''' 𝒮h''' acqh'''
      where "(m(a:=v), rws,𝒪h'',h'',𝒮h'') f* (m(a:=v), [],𝒪h''',h''',𝒮h''')" by blast
    with steps 
    have "(m, sbh,𝒪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_Writesb_eq_ext:
  "buffered_val (filter is_Writesb sb) a = buffered_val sb a" 
  by (induct sb) (auto split: memref.splits)

lemma bufferd_val_filter_is_Writesb_eq:
  "buffered_val (filter is_Writesb sb) = buffered_val sb"
  by (rule ext) (rule bufferd_val_filter_is_Writesb_eq_ext)

lemma outstanding_refs_is_volatile_Writesb_filter_writes: 
  "outstanding_refs is_volatile_Writesb (filter is_Writesb xs) = 
   outstanding_refs is_volatile_Writesb xs"
  by (induct xs) (auto simp add: is_volatile_Writesb_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,𝒮) sb (ts',m',𝒮')"
  assumes sim: "ts h tsh"
  shows "tsh' 𝒮h'. (tsh,m,𝒮h) sbh* (tsh',m',𝒮h')  ts' h tsh'"
proof -
  interpret sbh_computation: 
    computation sbh_memop_step flush_step program_step 
       "λp p' is sb. sb @ [Progsb 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, 𝒟, 𝒪, ,𝒮) sb 
                                (is', θ', sb', m', 𝒟', 𝒪', ℛ',𝒮')"
      by auto
  
    from sim obtain 
      lts_eq: "length ts = length tsh" and
      sim_loc: "i < length ts. (𝒪' 𝒟' ℛ'.
            let (p,is, θ, sb,𝒟, 𝒪,) = tsh!i in 
             ts!i=(p,is, θ, filter is_Writesb sb,𝒟',𝒪',ℛ') 
             (filter is_Writesb sb = []  sb=[]))"
      by cases (auto)

    from lts_eq i_bound have i_bound': "i < length tsh"
      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
	θ': "θ'=θ(tv)" 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 sbh 𝒪h h 𝒟h where 
	tsh_i: "tsh!i = (p,Read volatile a t#is',θ,sbh,𝒟h,𝒪h,h)" and
	sb: "sb = filter is_Writesb sbh" and
        sb_empty: "filter is_Writesb sbh = []  sbh=[]"
	by (auto simp add: Let_def)

      from buf_val
      have buf_val': "buffered_val sbh a = Some v"
	by (simp add: bufferd_val_filter_is_Writesb_eq sb)

      let ?tsh_i' = "(p, is', θ(t  v), sbh @ [Readsb volatile a t v], 𝒟h, 𝒪h,h)"
      let ?tsh' = "tsh[i := ?tsh_i']" 
      from sbh_memop_step.SBHReadBuffered [OF buf_val'] 
      have "(Read volatile a t # is', θ, sbh, m,𝒟h, 𝒪h, h,𝒮h) sbh 
            (is', θ(t  v), sbh@ [Readsb volatile a t v], m, 𝒟h, 𝒪h, h, 𝒮h)".
      from sbh_computation.Memop [OF i_bound' tsh_i this] 
      have step: "(tsh, m, 𝒮h) sbh (?tsh', m, 𝒮h)".

      from sb have sb: "sb = filter is_Writesb (sbh @ [Readsb volatile a t v])"
	by simp

      show ?thesis
      proof (cases "filter is_Writesb sbh = []")
        case False

        have "ts [i := (p,is',θ(t  v),sb,𝒟,𝒪,)] h ?tsh'"
          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: "sbh=[]" by simp
        from i_bound' have "?tsh'!i = ?tsh_i'"
          by auto

        
        from sbh_computation.StoreBuffer [OF _ this, simplified empty, simplified, OF _ flush_step.Readsb, of m 𝒮h] i_bound'
        have "(?tsh', m, 𝒮h)
              sbh  (tsh[i := (p, is', θ(t  v), [], 𝒟h, 𝒪h,h)], m, 𝒮h)"
          by (simp add: empty list_update_overwrite)
        with step have "(tsh, m, 𝒮h) sbh*
              (tsh[i := (p, is', θ(t  v), [], 𝒟h, 𝒪h,h)], m,𝒮h)"
          by force
        moreover
        have "ts [i := (p,is',θ(t  v),sb,𝒟,𝒪,)] h tsh[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
	θ': "θ'=θ(tm 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 sbh 𝒪h h 𝒟h where 
	tsh_i: "tsh!i = (p,Read volatile a t#is',θ,sbh,𝒟h,𝒪h,h)" and
	sb: "sb = filter is_Writesb sbh" and
        sb_empty: "filter is_Writesb sbh = []  sbh=[]"
	by (auto simp add: Let_def)

      from buf
      have buf': "buffered_val sbh a = None"
	by (simp add: bufferd_val_filter_is_Writesb_eq sb)

      let ?tsh_i' = "(p, is', θ(t  m a), sbh @ [Readsb volatile a t (m a)], 𝒟h, 𝒪h,h)"
      let ?tsh' = "tsh[i := ?tsh_i']" 

      from sbh_memop_step.SBHReadUnbuffered [OF buf']
      have "(Read volatile a t # is',θ, sbh, m, 𝒟h, 𝒪h, h,𝒮h) sbh 
            (is', θ(t  (m a)), sbh@ [Readsb volatile a t (m a)], m,𝒟h, 𝒪h, h,𝒮h)".
      from sbh_computation.Memop [OF i_bound' tsh_i this] 
      have step: "(tsh, m, 𝒮h) sbh 
            (?tsh', m, 𝒮h)".
      moreover 
      from sb have sb: "sb = filter is_Writesb (sbh @ [Readsb volatile a t (m a)])"
	by simp
    
      show ?thesis
      proof (cases "filter is_Writesb sbh = []")
        case False
        have "ts [i := (p,is',θ (tm a),sb,𝒟,𝒪,)] h ?tsh'"
	  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: "sbh=[]" by simp
        from i_bound' have "?tsh'!i = ?tsh_i'"
          by auto

        
        from sbh_computation.StoreBuffer [OF _ this, simplified empty, simplified, OF _ flush_step.Readsb, of m 𝒮h] i_bound'
        have "(?tsh', m, 𝒮h)
              sbh  (tsh[i := (p, is', θ(t  (m a)), [], 𝒟h, 𝒪h,h)], m, 𝒮h)"
          by (simp add: empty)
        with step have "(tsh, m, 𝒮h) sbh*
              (tsh[i := (p, is', θ(t  m a), [], 𝒟h, 𝒪h,h)], m, 𝒮h)"
          by force
        moreover
        have "ts [i := (p,is',θ(t  m a),sb,𝒟,𝒪,)] h tsh[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@[Writesb False a (D, f) (f θ) A L R W]" 
	by auto

      from sim_loc [rule_format, OF i_bound] ts_i
      obtain sbh 𝒪h h 𝒟h where 
	tsh_i: "tsh!i = (p,Write False a (D,f) A L R W#is',θ,sbh,𝒟h,𝒪h,h)" and
	sb: "sb = filter is_Writesb sbh" 
	by (auto simp add: Let_def "is") 

      from sbh_memop_step.SBHWriteNonVolatile 
      have "(Write False a (D, f) A L R W# is',θ, sbh, m, 𝒟h, 𝒪h, h,𝒮h) sbh 
               (is', θ, sbh @ [Writesb False a (D, f) (f θ) A L R W], m,𝒟h, 𝒪h, h,𝒮h)".
      from sbh_computation.Memop [OF i_bound' tsh_i this] 
      have "(tsh, m, 𝒮h) sbh 
            (tsh[i := (p, is',θ, sbh @ [Writesb False a (D, f) (f θ) A L R W], 𝒟h, 𝒪h,h)],
             m, 𝒮h)".
      moreover
      have "ts [i := (p,is',θ,sb @ [Writesb False a (D,f) (f θ) A L R W],𝒟,𝒪,)] h 
            tsh[i := (p,is',θ, sbh @ [Writesb 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@[Writesb True a (D, f) (f θ) A L R W]" 
	by auto

      from sim_loc [rule_format, OF i_bound] ts_i "is"
      obtain sbh 𝒪h h 𝒟h where 
	tsh_i: "tsh!i = (p,Write True a (D,f) A L R W#is',θ,sbh,𝒟h,𝒪h,h)" and
	sb: "sb = filter is_Writesb sbh"
	by (auto simp add: Let_def)

      from sbh_computation.Memop [OF i_bound' tsh_i SBHWriteVolatile 
	]
      have "(tsh, m, 𝒮h) sbh 
            (tsh[i := (p, is',θ, sbh @ [Writesb True a (D, f) (f θ) A L R W], True, 𝒪h,h)],
                m, 𝒮h)".

      moreover
      have "ts [i := (p,is',θ,sb @ [Writesb True a (D,f) (f θ) A L R W],𝒟,𝒪,)] h 
            tsh[i := (p,is', θ, sbh @ [Writesb 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 sbh 𝒪h h 𝒟h where 
	tsh_i: "tsh!i = (p,Fence # is',θ,sbh,𝒟h,𝒪h,h)" and
	sb: "[] = filter is_Writesb sbh"
	by (auto simp add: Let_def)


      from filter_is_Writesb_empty [OF sb [symmetric]]
      have "r  set sbh. is_Readsb r  is_Progsb r  is_Ghostsb r".

      from flush_reads_program [OF this] obtain 𝒪h' 𝒮h'  h'
      where flsh: "(m, sbh,𝒪h,h,𝒮h) f* (m, [],𝒪h',h',𝒮h')" by blast

      let ?tsh' = "tsh[i := (p,Fence # is', θ, [], 𝒟h, 𝒪h',h')]"
      from sbh_computation.store_buffer_steps [OF flsh i_bound' tsh_i]
      have "(tsh, m, 𝒮h) sbh* (?tsh', m, 𝒮h')".

      also

      from i_bound' have i_bound'': "i < length ?tsh'"
	by auto

      from i_bound' have tsh'_i: "?tsh'!i = (p,Fence#is',θ,[],𝒟h,𝒪h',h')"
	by simp
      from sbh_computation.Memop [OF i_bound'' tsh'_i SBHFence] i_bound'
      have "(?tsh', m, 𝒮h') sbh (tsh[i := (p, is',θ, [], False, 𝒪h',Map.empty)], m,𝒮h')"
	by (simp)
      finally
      have "(tsh, m, 𝒮h) sbh* (tsh[i := (p, is', θ, [],False, 𝒪h',Map.empty)],m, 𝒮h')".

      moreover
    
      have "ts [i := (p,is',θ,[],𝒟,𝒪,)] h tsh[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 sbh 𝒪h h 𝒟h where 
	tsh_i: "tsh!i = (p,RMW a t (D, f) cond ret A L R W# is',θ,sbh,𝒟h,𝒪h,h)" and
	sb: "[] = filter is_Writesb sbh"
	by (auto simp add: Let_def)



      from filter_is_Writesb_empty [OF sb [symmetric]]
      have "r  set sbh. is_Readsb r  is_Progsb r  is_Ghostsb r".

      from flush_reads_program [OF this] obtain 𝒪h' 𝒮h' h'
      where flsh: "(m, sbh,𝒪h,h,𝒮h) f* (m, [],𝒪h',h',𝒮h')" by blast

      let ?tsh' = "tsh[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' tsh_i]
      have "(tsh, m, 𝒮h) sbh* (?tsh', m, 𝒮h')".

      also

      from i_bound' have i_bound'': "i < length ?tsh'"
	by auto

      from i_bound' have tsh'_i: "?tsh'!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'' tsh'_i step ] i_bound' 
      have "(?tsh', m, 𝒮h') sbh (tsh[i := (p, is',θ(tm a), [], False, 𝒪h',Map.empty)],m, 𝒮h')"
	by (simp)
      finally
      have "(tsh, m, 𝒮h) sbh* (tsh[i := (p, is',θ(tm a), [], False, 𝒪h',Map.empty)],m, 𝒮h')".

      moreover
    
      have "ts [i := (p,is',θ(tm a),[],𝒟,𝒪,)] h tsh[i := (p,is', θ(tm 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 sbh 𝒪h h 𝒟h acqh where 
	tsh_i: "tsh!i = (p,RMW a t (D, f) cond ret A L R W# is',θ,sbh,𝒟h,𝒪h,h)" and
	sb: "[] = filter is_Writesb sbh"
	by (auto simp add: Let_def)

      from filter_is_Writesb_empty [OF sb [symmetric]]
      have "r  set sbh. is_Readsb r  is_Progsb r  is_Ghostsb r".

      from flush_reads_program [OF this] obtain 𝒪h' 𝒮h' h'
      where flsh: "(m, sbh,𝒪h,h,𝒮h) f* (m, [],𝒪h',h',𝒮h')" by blast

      let ?tsh' = "tsh[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' tsh_i]
      have "(tsh, m, 𝒮h) sbh* (?tsh', m, 𝒮h')".

      also

      from i_bound' have i_bound'': "i < length ?tsh'"
	by auto

      from i_bound' have tsh'_i: "?tsh'!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'' tsh'_i step ] i_bound' 
      have "(?tsh', m, 𝒮h') sbh (tsh[i := (p, is',
	     θ(t  ret (m a) (f (θ(t  (m a))))), [], False, 𝒪h'  A - R,Map.empty)],
	     m(a := f (θ(t  (m a)))),𝒮h' ⊕⇘WR ⊖⇘AL)"
	by (simp)
      finally
      have "(tsh, m, 𝒮h) sbh* (tsh[i := (p, is',
	     θ(t  ret (m a) (f (θ(t  (m a))))), [], False, 𝒪h'  A - R,Map.empty)],
            m(a := f (θ(t  (m a)))),𝒮h' ⊕⇘WR ⊖⇘AL)".

      moreover
    
      have "ts [i := (p,is',θ(t  ret (m a) (f (θ(t  (m a))))),[],𝒟,𝒪,)] h 
            tsh[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 sbh 𝒪h h 𝒟h where 
	tsh_i: "tsh!i = (p,Ghost A L R W# is',θ,sbh,𝒟h,𝒪h,h)" and
	sb: "sb = filter is_Writesb sbh" and
        sb_empty: "filter is_Writesb sbh = []  sbh=[]"
	by (auto simp add: Let_def)

      let ?tsh_i' = "(p, is', θ, sbh@[Ghostsb A L R W],𝒟h, 𝒪h,h)"
      let ?tsh' = "tsh[i := ?tsh_i']" 
      note step= SBHGhost  
      from sbh_computation.Memop [OF i_bound' tsh_i step ] i_bound' 
      have step: "(tsh, m, 𝒮h) sbh (?tsh',m, 𝒮h)"
	by (simp)

      from sb have sb: "sb = filter is_Writesb (sbh @ [Ghostsb A L R W])"
	by simp

      show ?thesis
      proof (cases "filter is_Writesb sbh = []")
        case False

        have "ts [i := (p,is',θ,sb,𝒟,𝒪,)] h ?tsh'"
	  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: "sbh=[]" by simp
        from i_bound' have "?tsh'!i = ?tsh_i'"
          by auto
        from sbh_computation.StoreBuffer [OF _ this, simplified empty, simplified, OF _ flush_step.Ghost, of m 𝒮h] i_bound'
        have "(?tsh', m, 𝒮h)
              sbh  (tsh[i := (p, is', θ, [], 𝒟h, 𝒪h  A - R,augment_rels (dom 𝒮h) R h)], m, 𝒮h ⊕⇘WR ⊖⇘AL)"
          by (simp add: empty)
        with step have "(tsh, m, 𝒮h) sbh*
              (tsh[i := (p, is', θ, [], 𝒟h, 𝒪h  A - R,augment_rels (dom 𝒮h) R h)], m, 𝒮h ⊕⇘WR ⊖⇘AL)"
          by force
        moreover
        have "ts [i := (p,is',θ,sb,𝒟,𝒪,)] h 
                 tsh[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 tsh" and
      sim_loc: "i < length ts. (𝒪' 𝒟' ℛ'. 
            let (p,is,θ, sb, 𝒟, 𝒪,) = tsh!i in ts!i=(p,is, θ, filter is_Writesb sb,𝒟',𝒪',ℛ') 
                (filter is_Writesb sb = []  sb = []))"
      by cases auto

    from sim_loc [rule_format, OF i_bound] ts_i 
      obtain sbh 𝒪h h 𝒟h acqh where 
	tsh_i: "tsh!i = (p,is,θ,sbh,𝒟h,𝒪h,h)" and
	sb: "sb = filter is_Writesb sbh" and
        sb_empty: "filter is_Writesb sbh = []  sbh=[]"
	by (auto simp add: Let_def)

    from lts_eq i_bound have i_bound': "i < length tsh"
      by simp
   
    let ?tsh_i' = "(p', is @ is',θ, sbh @ [Progsb p p' is'], 𝒟h, 𝒪h,h)"
      let ?tsh' = "tsh[i := ?tsh_i']" 
    from sbh_computation.Program [OF i_bound' tsh_i prog_step]
    have step: "(tsh, m, 𝒮h) sbh (?tsh',m, 𝒮h)".
    
    show ?thesis
    proof (cases "filter is_Writesb sbh = []")
      case False
      have "ts[i := (p', is@is', θ, sb,𝒟, 𝒪,)] h ?tsh'"
        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: "sbh=[]" by simp
      from i_bound' have "?tsh'!i = ?tsh_i'"
        by auto
      
      from sbh_computation.StoreBuffer [OF _ this, simplified empty, simplified, OF _ flush_step.Progsb, of m 𝒮h] i_bound'
      have "(?tsh', m, 𝒮h)
              sbh  (tsh[i := (p', is@is', θ, [], 𝒟h, 𝒪h,h)], m, 𝒮h)"
        by (simp add: empty)
      with step have "(tsh, m, 𝒮h) sbh*
           (tsh[i := (p', is@is', θ, [], 𝒟h, 𝒪h,h)], m, 𝒮h) "
        by force
      moreover
      have "ts[i := (p', is@is', θ, sb,𝒟, 𝒪,)] h tsh[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 tsh" and
      sim_loc: "i < length ts. (𝒪' 𝒟' ℛ'. 
            let (p,is, θ, sb,𝒟, 𝒪,) = tsh!i in ts!i=(p,is, θ, filter is_Writesb sb,𝒟',𝒪',ℛ') 
                (filter is_Writesb sb = []  sb=[]))"
      by cases auto

    from sim_loc [rule_format, OF i_bound] ts_i 
      obtain sbh 𝒪h h 𝒟h acqh where 
	tsh_i: "tsh!i = (p,is,θ,sbh,𝒟h,𝒪h,h)" and
	sb: "sb = filter is_Writesb sbh" and        
        sb_empty: "filter is_Writesb sbh = []  sbh=[]"
	by (auto simp add: Let_def)

    from lts_eq i_bound have i_bound': "i < length tsh"
      by simp

    from flush_simulates_filter_writes [OF sb_step sb, of 𝒪h h 𝒮h] 
    obtain sbh' 𝒪h' h' 𝒮h' 
      where flush': "(m, sbh,𝒪h,h,𝒮h) f* (m', sbh',𝒪h',h',𝒮h')" and 
      sb': "sb' = filter is_Writesb sbh'" and
      sb'_empty: "filter is_Writesb sbh' = []  sbh'=[]"
      by auto

    from sb_step obtain volatile a sop v A L R W where "sb=Writesb volatile a sop v A L R W#sb'"
      by cases force
    from sbh_computation.store_buffer_steps [OF flush' i_bound' tsh_i]
    have "(tsh, m, 𝒮h) sbh* (tsh[i := (p, is, θ, sbh',𝒟h, 𝒪h',h')], m', 𝒮h')".
    
    moreover
    have "ts[i := (p, is, θ, sb',𝒟, 𝒪',ℛ')] h 
          tsh[i := (p, is, θ, sbh',𝒟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,𝒮) sb*  (ts',m',𝒮')"
  shows "tsh 𝒮h. ts h tsh  tsh' 𝒮h'. (tsh,m,𝒮h) sbh* (tsh',m',𝒮h')  ts' h tsh'"
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,𝒮) sb  (ts'',m'',𝒮'')" by fact
  have sim: "ts h tsh" by fact
  from concurrent_history_steps_simulates_store_buffer_step [OF first sim, of 𝒮h]
  obtain tsh'' 𝒮h'' where
    exec: "(tsh,m,𝒮h) sbh* (tsh'',m'',𝒮h'')" and  sim: "ts'' h tsh''"
    by auto
  from step.hyps (3) [OF sim, of 𝒮h'']
  obtain tsh' 𝒮h' where exec_rest: "(tsh'',m'',𝒮h'')  sbh* (tsh',m',𝒮h')" and sim': "ts' h tsh'"
    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: "(tssb,msb,x) sb* (tssb',msb',x')"
assumes init: "initialsb tssb 𝒮sb"
assumes valid: "valid tssb" (* FIXME: move into initial ?*)
assumes sim: "(tssb,msb,𝒮sb)  (ts,m,𝒮)"
assumes safe: "safe_reach_direct safe_free_flowing (ts,m,𝒮)"
shows "tsh' 𝒮h' ts' m' 𝒮'. 
          (tssb,msb,𝒮sb) sbh* (tsh',msb',𝒮h') 
               tssb' h tsh' 
          (ts,m,𝒮) d* (ts',m',𝒮')  
                (tsh',msb',𝒮h')  (ts',m',𝒮')"
proof -
  from init interpret ini: initialsb tssb 𝒮sb .
  from concurrent_history_steps_simulates_store_buffer_steps [OF exec_sb ini.history_refl, of 𝒮sb]
  obtain tsh' 𝒮h' where
    sbh: "(tssb,msb,𝒮sb) sbh* (tsh',msb',𝒮h')" and
    sim_sbh: "tssb' h tsh'"
    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: "(tsh',msb',𝒮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 tsd; 
    i < length ts. 
         (𝒪' 𝒟' ℛ'.
           let (p,is, θ,sb,𝒟, 𝒪,) = tsd!i in 
                ts!i=(p,is, θ, [] ,𝒟',𝒪',ℛ'))
    
     
     ts d tsd"

lemma empty_sb_sims: 
assumes empty:
  "i p is xs sb 𝒟 𝒪 . i < length tssb  tssb!i=(p,is,xs,sb,𝒟,𝒪,) sb=[]"
assumes sim_h: "tssb h tsh"
assumes sim_d: "(tsh,mh,𝒮h)  (ts,m,𝒮)"
shows "tssb d ts  mh=m  length tssb = length ts"
proof-
  from sim_h empty
  have empty':
  "i p is xs sb 𝒟 𝒪 . i < length tsh  tsh!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: "tssb d ts"
shows "tsh. tssb h tsh  (tsh,m,𝒮)  (ts,m,𝒮)"
proof -
  from sim obtain
    leq: "length tssb = length ts" and
    sim: "i < length tssb. 
         (𝒪' 𝒟' ℛ'.
           let (p,is, θ,sb,𝒟, 𝒪,) = ts!i in 
                tssb!i=(p,is, θ, [] ,𝒟',𝒪',ℛ'))"
    by cases auto
  define tsh where "tsh  map (λ(p,is, θ,sb,𝒟, 𝒪,). (p,is, θ,[]::'a memref list,𝒟, 𝒪,)) ts" 
  have "tssb h tsh"
    apply (rule sim_history_config.intros)
    using leq sim
    apply (auto simp add: tsh_def Let_def leq)
    done
  moreover
  have empty:
  "i p is xs sb 𝒟 𝒪 . i < length tsh  tsh!i=(p,is,xs,sb,𝒟,𝒪,) sb=[]"
    apply (clarsimp simp add: tsh_def Let_def)
    subgoal for i
    apply (case_tac "ts!i")
    apply auto
    done
    done
  
  have "(tsh,m,𝒮)  (ts,m,𝒮)"
    apply (rule sim_config_emptyI [OF empty])
    apply  (clarsimp simp add: tsh_def)
    apply (clarsimp simp add: tsh_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: "(tssb,msb,x) sb* (tssb',msb',x')"
assumes init: "initialsb tssb 𝒮sb"
assumes valid: "valid tssb" (* FIXME: move into initial ?*)
assumes empty: 
  "i p is xs sb 𝒟 𝒪 . i < length tssb'  tssb'!i=(p,is,xs,sb,𝒟,𝒪,) sb=[]"
assumes sim: "(tssb,msb,𝒮sb)  (ts,m,𝒮)"
assumes safe: "safe_reach_direct safe_free_flowing (ts,m,𝒮)"
shows "ts' 𝒮'. 
          (ts,m,𝒮) d* (ts',msb',𝒮')  tssb' d ts'"
proof -
  from concurrent_direct_execution_simulates_store_buffer_execution [OF exec_sb init valid sim safe]
  obtain tsh' 𝒮h' ts' m' 𝒮' where
    "(tssb,msb,𝒮sb) sbh* (tsh',msb',𝒮h')" and
    sim_h: "tssb' h tsh'" and
    exec: "(ts,m,𝒮) d* (ts',m',𝒮')" and
    sim: "(tsh',msb',𝒮h')  (ts',m',𝒮')"
    by auto
  from empty_sb_sims [OF empty sim_h sim]
  obtain "tssb' d ts'" "msb' = m'" "length tssb' = length ts'"
    by auto
  thus ?thesis
    using exec
    by blast
qed

locale initiald = 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: "initiald ts 𝒮 valid"
  shows "initialsb (map (λ(p,is, θ,sb,𝒟, 𝒪,). (p,is, θ,[],𝒟, 𝒪,)) ts) 𝒮" 
         (is "initialsb ?map 𝒮")
proof -
  from init interpret ini: initiald 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 "initialsb_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: "(tssb,m,x) sb* (tssb',m',x')"
assumes empty': "empty_store_buffers tssb'"
assumes sim: "tssb d ts"
assumes init: "initiald ts 𝒮 valid"
assumes safe: "safe_reach_direct safe_free_flowing (ts,m,𝒮)"
shows "ts' 𝒮'. 
          (ts,m,𝒮) d* (ts',m',𝒮')  tssb' d ts'"
proof -
  from empty'
  have empty': 
  "i p is xs sb 𝒟 𝒪 . i < length tssb'  tssb'!i=(p,is,xs,sb,𝒟,𝒪,) sb=[]"
    by (auto simp add: empty_store_buffers_def)

  define tsh where "tsh  map (λ(p,is, θ,sb,𝒟, 𝒪,). (p,is, θ,[]::'a memref list,𝒟, 𝒪,)) ts" 
  from initial_d_sb [OF init]
  have init_h:"initialsb tsh 𝒮"
    by (simp add: tsh_def)
  from initiald.valid_init [OF init]
  have valid_h: "valid tsh"
    by (simp add: tsh_def)
  from sim obtain
    leq: "length tssb = length ts" and
    sim: "i < length tssb. 
         (𝒪' 𝒟' ℛ'.
           let (p,is, θ,sb,𝒟, 𝒪,) = ts!i in 
                tssb!i=(p,is, θ, [] ,𝒟',𝒪',ℛ'))"
    by cases auto
  have sim_h: "tssb h tsh"
    apply (rule sim_history_config.intros)
    using leq sim
    apply (auto simp add: tsh_def Let_def leq)
    done

  from concurrent_history_steps_simulates_store_buffer_steps [OF exec_sb sim_h, of 𝒮]
  obtain tsh' 𝒮h' where steps_h: "(tsh,m,𝒮) sbh* (tsh',m',𝒮h')" and
     sim_h': "tssb' h tsh'"
    by auto

  moreover
  have empty:
  "i p is xs sb 𝒟 𝒪 . i < length tsh  tsh!i=(p,is,xs,sb,𝒟,𝒪,) sb=[]"
    apply (clarsimp simp add: tsh_def Let_def)
    subgoal for i
    apply (case_tac "ts!i")
    apply auto
    done
    done
  
  have sim': "(tsh,m,𝒮)  (ts,m,𝒮)"
    apply (rule sim_config_emptyI [OF empty])
    apply  (clarsimp simp add: tsh_def)
    apply (clarsimp simp add: tsh_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': "(tsh', m', 𝒮h')  (ts', m'', 𝒮'')"
    by blast
  from empty_sb_sims [OF empty' sim_h' sim'] steps
  show ?thesis
    by auto
qed


locale initialv = 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)"

(*
term "initialv"
context xvalid_program_progress
begin
term "safe_reach safe_free_flowing (ts,m,𝒮)"
theorem (in xvalid_program_progress) store_buffer_execution_result_sequential_consistent':
assumes exec_sb: "(tssb,m,x) ⇒sb* (tssb',m',x')"
assumes sim: "tssbd ts"

assumes safe: " safe_reach safe_free_flowing (ts,m,𝒮)"
shows "∃ts' 𝒮'. 
          (ts,m,𝒮) ⇒v* (ts',m',𝒮') "
*)



theorem (in xvalid_program_progress) store_buffer_execution_result_sequential_consistent':
assumes exec_sb: "(tssb,m,x) sb* (tssb',m',x')"
assumes empty': "empty_store_buffers tssb'"
assumes sim: "tssb d ts"
assumes init: "initialv ts 𝒮 valid"
assumes safe: "safe_reach_virtual safe_free_flowing (ts,m,𝒮)"
shows "ts' 𝒮'. 
          (ts,m,𝒮) v* (ts',m',𝒮')  tssb' d ts'"
proof -
  define tsd where "tsd == (map (λ(p,is, θ,sb,𝒟, 𝒪,ℛ'). (p,is, θ,sb,𝒟, 𝒪,Map.empty::rels)) ts)"
  have rem_ts: "remove_rels tsd = ts"
    apply (rule nth_equalityI)
    apply  (simp add: tsd_def remove_rels_def)
    apply (clarsimp simp add: tsd_def remove_rels_def)
    subgoal for i
    apply (case_tac "ts!i")
    apply clarsimp
    done
    done
  from sim
  have sim': "tssb d tsd"
    apply cases
    apply (rule sim_direct_config.intros)
    apply (auto simp add: tsd_def)
    done
  
  have init': "initiald tsd 𝒮 valid"
  proof (intro_locales)
    from init have "simple_ownership_distinct ts"
      by (simp add: initialv_def)
    then
    show "simple_ownership_distinct tsd"
      apply (clarsimp simp add: tsd_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: initialv_def)
    then show "read_only_unowned 𝒮 tsd"
      apply (clarsimp simp add: tsd_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: initialv_def)
    then 
    show "unowned_shared 𝒮 tsd"
      apply (clarsimp simp add: tsd_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 "initialv_axioms ts valid"
      by (simp add: initialv_def)
     
    then
    show "initiald_axioms tsd valid"
      apply (clarsimp simp add: tsd_def initialv_axioms_def initiald_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 tsd' m' 𝒮'
    assume exec: "(tsd, m, 𝒮) d* (tsd', m', 𝒮')"
    have "safe_free_flowing (tsd', m', 𝒮')"
    proof -
      from virtual_simulates_direct_steps [OF exec]
      have exec_v: "(ts, m, 𝒮) v* (remove_rels tsd', m', 𝒮')"
        by (simp add: rem_ts)
      have eq: "map (owned 
                    (λ(p, is, θ, sb, 𝒟, 𝒪, ). (p, is, θ, (), 𝒟, 𝒪, ())))
                tsd' = map owned tsd'"
        by auto
      from exec_v safe
      have "safe_free_flowing (remove_rels tsd', 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 (tsd, m, 𝒮)"
    by (simp add: safe_reach_def)
          
  from store_buffer_execution_result_sequential_consistent [OF exec_sb empty' sim' init' safe'] 
  obtain tsd' 𝒮' where
     exec_d: "(tsd, m, 𝒮) d* (tsd', m', 𝒮')"  and sim_d: "tssb' d tsd'"
    by blast

  from virtual_simulates_direct_steps [OF exec_d]
  have "(ts, m, 𝒮) v* (remove_rels tsd', m', 𝒮')"
    by (simp add: rem_ts)
  moreover
  from sim_d
  have "tssb' d remove_rels tsd'"
    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: "(tssb,msb,𝒮sb) sb (tssb',msb',𝒮sb')"
  assumes sim_h: "tssb h tssbh"
  assumes sim: "(tssbh,msb,𝒮sbh)  (ts,m,𝒮) "
  assumes valid_own: "valid_ownership 𝒮sbh tssbh"
  assumes valid_sb_reads: "valid_reads msb tssbh"
  assumes valid_hist: "valid_history program_step tssbh"
  assumes valid_sharing: "valid_sharing 𝒮sbh tssbh"
  assumes tmps_distinct: "tmps_distinct tssbh"
  assumes valid_sops: "valid_sops tssbh"
  assumes valid_dd: "valid_data_dependency tssbh"
  assumes load_tmps_fresh: "load_tmps_fresh tssbh"
  assumes enough_flushs: "enough_flushs tssbh"
  assumes valid_program_history: "valid_program_history tssbh"
  assumes valid: "valid tssbh"
  assumes safe_reach: "safe_reach_direct safe_delayed (ts,m,𝒮)"
  shows "tssbh' 𝒮sbh'. 
         (tssbh,msb,𝒮sbh) sbh* (tssbh',msb',𝒮sbh')  tssb' h tssbh' 
         valid_ownership 𝒮sbh' tssbh'  valid_reads msb' tssbh'  
         valid_history program_step tssbh' 
         valid_sharing 𝒮sbh' tssbh'  tmps_distinct tssbh'  valid_data_dependency tssbh' 
         valid_sops tssbh'  load_tmps_fresh tssbh'  enough_flushs tssbh' 
         valid_program_history tssbh'  valid tssbh' 
           (ts' 𝒮' m'. (ts,m,𝒮) d* (ts',m',𝒮')  
            (tssbh',msb',𝒮sbh')  (ts',m',𝒮'))"
proof -
  from concurrent_history_steps_simulates_store_buffer_step [OF step_sb sim_h]
  obtain tssbh' 𝒮sbh' where
    steps_h: "(tssbh,msb,𝒮sbh) sbh* (tssbh',msb',𝒮sbh')" and
    sim_h': "tssb' h tssbh'"
    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',𝒮')" "(tssbh', msb',𝒮sbh')  (ts', m', 𝒮')"
    "valid_ownership 𝒮sbh' tssbh'" "valid_reads msb' tssbh'" "valid_history program_step tssbh'"
    "valid_sharing 𝒮sbh' tssbh'" "tmps_distinct tssbh'" "valid_data_dependency tssbh'"
    "valid_sops tssbh'" "load_tmps_fresh tssbh'" "enough_flushs tssbh'"
    "valid_program_history tssbh'" "valid tssbh'"
    by blast
  ultimately
  show ?thesis
    by blast
qed

(* ******************* Some tuned version for presentations ********************************** *)

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: "(tssb,msb,𝒮sb) sbh (tssb',msb',𝒮sb')"
  assumes sim: "(tssb,msb,𝒮sb)  (ts,m,𝒮)"
  assumes inv: "invariant tssb 𝒮sb msb"
  assumes safe_reach: "safe_reach_direct safe_delayed (ts,m,𝒮)"
  shows "invariant tssb' 𝒮sb' msb' 
           (ts' 𝒮' m'. (ts,m,𝒮) d* (ts',m',𝒮')  (tssb',msb',𝒮sb')  (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