Theory ReduceStoreBuffer

(* Copyright (C) 2007--2010 Norbert Schirmer
 * All rights reserved, DFKI GmbH 
 *)
theory ReduceStoreBuffer
imports Main
begin
(*
Basic access policy:
shared + owned by some thread (only the owner is allowed to write): 
 - owner: LMS, volatile writes and arbitrary reads; 
 - others: volatile reads

shared + not owned by any thread:
 - read/write: LMS, volatile writes and volatile reads
 - read-only: arbitrary reads

The core argument in the simulation proof is to show that all reads can be executed in the
virtual machine, preserving the values they have seen in the store-buffer machine.

 * Non volatile reads *

non-volatile means that the tread knows everything about the evolution of the value. It is thread local
or read only. Writes in other store buffers may not interfere with non-volatile reads.
Complicated case:
i: [Load (read-only) a (m a)]
j: [Acq {a}, Write a v]

This may only be sceduled such that i comes first and then j. If Acq {a} already happend, safety of
the Load would be violated.
So when i is poised to Load, and j has already executed on the virtual machine we have a safety
violation.
On the other hand, if i has already executed on the virtual machine ("m a" has ben seen) and the
Acq of j is fine.

Not that read-only loads (before the first volatile store) in the store buffer may become stale before 
they hit memory.
 
 * volatile reads *

The value seen can depend on writes of other threads. That is why the flushing policy enforces the
store-buffer to be empty, if there is an outstanding volatile write in the store-buffer.
The volatile reads must be executed in the virtual machine at the time
they are issued to the store-buffer, to ensure that both machines see the same value.
It is perfectly valid that volatile reads in the store buffer become stale, e.g.
i: [VLoad a (m a)]
j: [VWrite a v]

As j's write is committed to memory the (history-)entry in i's store-buffer becomes stale. But as the
virtual machine has also already executed the load this is fine.

Complicated case:
i: [VLoad a (m a)]
j: [Write a v, Release {a}]

As i sees (m a) and not v it has to come first. 
But if j has already executed and released a, there is no
safety-violation when i is poised to execute the VLoad afterwards. This would lead to the wrong value (v instead of "m a", but we cannot rule out the scheduling by a safety argument.

Currently we avoid this situation, by coupling Releases to volatile Writes. Hence the Release is
not yet executed in the virtual machine and thus VLoad of i is unsafe.

Another solution would be to go back in time, to the state before Release {a} is executed, and argue
that this scheduling is also valid and would lead to a safeness violation and thus must be disregarded.


Some thoughts.:

I think some key to free-flowing release story is the following:

    flushed      suspended
   /       \ /               \
i: […, …  , Write True,…,…]
j: […, …  , Write True,…,…]
   \       / \               /
     races     races ruled 
      ok         out by 
                 safety

Races in the flushed part are ok, since they are already simulated in the virtual machine.
The problem is how to formulate the invariants and all the stuff that this works out in the
proof. E.g. 

i: [Write a v, Release {a}]
j: poised: Rread a

So the idea doesn't seem to work. It may only work as long as no write is involved. (i.e. Release / Acq)

FIXME: 
shared: does not seem to buy anything right now, as owner has to do volatile reads and writes.
(It should be possible that the owner does non-volatile reads, but as he is the only writer reading makes no
sense)
LMS: maybe store restriction could be relayed to owned + shared, if there is no sharing anymore there are no restrictions.

What could be better is read-only memory which can be read accessed non-volatily.

a: either owned by some thread;
   either unowned by any thread and read/write;
   either unowned by any thread and read-only:a

but how does memory become read-only?

Free-Flowing releases:
Take proof as is, and augment machnes with set of released since last flush state.
xsafety: nobody is allowed to acquire / write to released things:

Second: proof that safety ==> xsafety for virtual machine (by induction on ->*, and an initial configuration).
*)

(*
FIXME, think about this:
Introduce additional control state for the direct-machine to keep track of whether
there is an outstanding volatile store, and the set of acquired variables since the
last flush. Add this protocol to the 'good'-stuff. This should make the 
enough-flushs invariant unnecessary, and it would make it more straightforward to introduce
state-dependencies to A,L,R (maybe from the temporaries and ghost state?). If state dependencies make sense depends on the integration into PIMP.   


A stricter separation of ghost and ordinary state would be nice.
shared: global ghost state
owned: local ghost state.

SB-Machine independent of ghost state.
D-Machine uses all of it.
SBH-Machine uses it only as history information 


## notes: ernie 3. July 2009
#################################################################


final theorem:
 safe-free-flowing ⟹ sb-machine is sequential consistent

we basically have (modulo dummy delay) :
 safe-delayed ⟹ sb-machine is sequential consistent


need theorem
safe-free-flowing ⟹ safe-dalayed

prf by contraposition and induction on execution:

S unsafe-dalayed ⟹ S unsafe free-flowing
induction on S:
let t be last step of T (owning the disuputed address)

by Lemma: ∃S'. |S'| < |S| and
   S' ∘ t ∼ S and S' reachable

   case: t = release (L) then S' is unsafe free flowing ⟹ S unsafe free flowing
   otherwise: S' unsafe delayed by induction S' unsafe free flowing ⟹ S unsafe free flowing



Lemma:

S is reachable and 
last step of thread T is t, which is not a volatile write
then

∃S' reachable. |S'| < |S| ∧ S' ∘ t ∼ S   (∼: states are equal modulo stuff owned by T)
proof

 * S = <>, tivial.
 * S = S1 ∘ t, trivial.
 * S = S1 ∘ u,  (u ≠ t)
    (induction): ∃S1' ∼ S1 ∘ t
      moreover: t ∘ u ∼ u ∘ t
    
      ultimately: S = S1 ∘ u ∼ S1' ∘ t ∘ u ∼ S1' ∘ u ∘ t ∼ S1 ∘ t
   
6.10.2009: 
I don't think this proof works. The problem are volatile reads. We cannot just reschedule them
to the end, since the memory content may well be changed by intermediate volatile writes of other
threads.

It might be possible to come up with another lemma that says that we can delete the
last action thread T (if its not a volatile write), since this operation should not
have an effect on the following ops of the other threads.   
*)

(* ****************** 2.10.2009 *************************** *)
(*
Some new thoughts to introduce free-flowing releases in current proof scheme:

* introduce safety notion on store-buffer machine as well: safe_sb
* safe_sb takes also store-buffer content into accout, e.g. 
  I'm not allowed to acquire what is acquired in someone elses sb.
* make safe_sb a part of invariant on sb machine.
* to show that safe_sb is preserved builds on safe_reach of virtual machine

* safe_sb makes it possible to deal with releases in prefixes:
     
aforementioned complicated case is not safe_sb, as a is still owned (and not shared) by j
i: [VLoad a (m a)]
j: [Write a v, Release {a}]

what about:
i: [VLoad a (m a)]
j: [VWrite a v, Release {a}] (a is owned by j but is also shared)

j's sb is suspended. Hence v has not yet made it to memory. Hence i sees (m a) as expected.

Examples: preservation of safe_sb
config 1
i: Acq a 
j: Acq a   

  is safe_sb (if nowbody owns a)

config 2:
i: Acq a
j: [Acq a]

  should not be safe_sb.
  
  rule out by safe_reach on config 1.
  fast-forward either sb of i or j and execute the Acq a. Then the remaining Acq a is unsafe.

Important for safe_sb is, that for the current thread, the ghost-state (on ownership) is after executing its sb!
Otherwise one has problems with quite simple things, e.g. a step 
(1) i: Acq a; Write a;    owns = {}
(2) i: Write a; [Acq a]   owns = {}

assume (1) is safe_sb; Why should (2) be safe_sb?
It is if we consider acquire [Acq a] owns = {a} instead. And we know that it has to be acquired somewhere
in the store-buffer by exploiting safe_reach.

Maybe its better not to use safe_sb on the sb-state but on the state of the virtual machine.
Better in the sense that it is easier to connect it with safe_reach.


### 5.10 ###
Can't get safe_sb inductive.

consider:
i: [Rel a]
j: someting; Acq a; []
this is still safe_sb; but after one step where something is executed I have
i: [Rel a]
j: Acq a; []
this should not be safe_sb; but I have no arguments at hand.

### next idea ###
The control flow (and the reads) of a thread do not depend on the 'flush' parts of *other* threads, but can depend
on their own flush parts.

Moreover the information of the store buffer machine is not enough to construct a virtual execution that will
violate safety. We have to go back in time even further. An virtual machine that just suspends everything is not
inductive according to the rules of virtual machine steps: e.g.
State 1)
m
i: [VRead a (m a)]
j: [VWrite a 5]

State 2)
m(a:=5)
i: [VRead a (m a)]
j: []

State 1 could still be executed on a virtual machine that has i and j suspended, but State 2 not 
(since a is already visible).

What could help is we just take an extra starting config for the virtual machine for which we have:

csb ∼ cv
csb ∼i ci; (we get virtual config ci by flushing (only) store buffer i until first volatile write, 
            all other buffers are suspended)

c →* cv
∀i. c →* ci
safe_reach c;

we show: csb → csb' ⟹ cv →* cv'

(hence we immediately have inductivity of c with respect to steps from cv : c →* cv →* cv')
What about ci → ci'

  * for thread ci it should work similar as for cv'
  * if thread i takes a step then we also have to justify cj → cj'. 
    when something enters the storebuffer i it is just stuttering (as i is just suspended)
    when something of thread i leaves the store-buffer it is also in the front of
    the instructions in i (since suspended) and may only depend on volatile writes
    in thread j, but those are also suspended.


The next question is can we instantiate the theorem with some kind of initial state?
For a state where the store-buffers are all empty we should have:
csb == cv == c == ci 

Hence we trivially have all reachability constraints, and safe_reach c == safe_reach cv

Problem with this approach: just suspending all other threads does not work.
consider justification of thread i for step ci → ci', where thread j takes a step, namely a volatile read.
volatile read can become stale in the store buffer (e.g. by a volatile write in thread i). Hence
we cannot simulate this read. (Thats why we flush all store-buffers until first volatile write in cv).
i: [VWrite a 5]
j: [VRead a 0]

i's virtual view:
i: VWrite a 5
j: VRead a 0

if i takes a step with Write a 5, we have a problem simulating j's stale read.

j's virtual view:
i: VWrite a 5
j: (Read has already happened)




Other problem:
i: [VWrite b 5] Acq a
j: [Rel a; VRead b 0]

Thread j has to execute before i to justify the read. But intuitively in ci, we wanted to delay the 'Rel a' such that
we can say 'Acq a' in thread i is unsafe.

#### 08.10.2009 ##################################
### Refined Approach for Free Flowing Releases ###
##################################################

  General setup
  -------------

Two safety predicates: 
 * safe_free_flowing: free flowing ghost releases in instruction stream
 * safe_delayed: release is delayed until next volatile write

safe_delayed is a variant of our current model, where 'Ghost' can only acquire, 
and releases are coupled with volatile (or interlocked) writes. However we don't 
want to introduce 2 different instruction streams (one with releases
and one without but with a proper annotation at the next volatile write). 
Instead we attempt to model safe-delayed with additional ghost state. 
Mainly a thread local set 'rel' of releases addresses (since the last volatile write). 
A volatile write resets this set, and a ghost release adds to this set, 
and safe_delayed may check these sets, whereas safe_free_flowing ignores these sets. 
Example: thread i poised for "Acquire A"

safe_free_flowing: 

  ∀j ≠ i. A ∩ ownsj = {}  ("A not owned by others")

safe_delayed:

  ∀j ≠ i. A ∩ (ownsj ∪ relj)= {}  ("A not owned by others, 
                                   or release not yet committed")

(sanity check: safe_delayed is more restrictive than safe_free_flowing, i.e. 
  safe_delayed c ⟹ safe_free_flowing c)

safe_reach safe c ≡ ∀c'. c ⇒* c' ⟹ safe c'
safe_reachn safe c ≡ ∀c'. c ⇒n c' ⟹ safe c'

Then we show two main theorems:

Theorem 1 (Simulation):
[|csbsb csb'; csb ∼ c; safe_reach safe_delayed c|]
⟹
∃c'. csb' ∼ c' ∧ csb* c'

Proof.
  Hopefully straightforward adoption of current proof.
Qed. 
Theorem 1 can easily be extended to many steps csbsb* csb'

Definition: init c ("initial state")
  Intuition: state where safe_delayed and safe_free_flowing are in-sync
  That is: all reli = {};

Fact:
init c ⟹ safe_free_flowing c ⟶ safe_delayed c;


Theorem 2 (Safety):
[|safe_reach safe_free_flowing c; init c|]
⟹
safe_reach safe_delayed c
Proof.

One basic ingredient is contraposition: 
from an unsafe_delayed computation we attempt to construct an unsafe_free_flowing computation.

Scenario for intuition:

  thread i:  …    Rel {a} …    |   … VWrite 
  thread j:  …       |    … Acq {a} 

thread i is somewhere in the computation between its release and the next volatile write.
thread j tries to Acq {a} inbetween.

This is safe_free_flowing (as the release has already happened), but not safe_delayed.
We want to argue that there is another scheduling of the global computation 
such that we also get a violation of safe_free_flowing.

For thread i we know that there are no volatile writes after the release until 
we hit the violation of safe_delayed. Intuitively this means that the other 
threads 'do not depend' on what is computed inside thread i after the release.
Note that the opposite is not the case. There can be volatile reads in thread i, 
which depend on (volatile) writes of 
other threads.

e.g.

  thread i:  …    Rel {a} …            … VRead b 5 …    |  … VWrite 
  thread j:  …       |    … VWrite b 6 …          …  Acq {a} 

This means that Read b 5 must come before VWrite b 6.

In general this implies that we cannot just reorder all steps of thread i 
(beginning from Rel {a}) after the Acq {a} in thread j to construct an 
unsafe_free_flowing state. 

It suggests that we instead try to *remove* all steps from thread i 
(beginning from and including Rel {a}) from the global computation, 
and argue that the other threads (especially j) can still do their 
computations until we reach a violation of safe_free_flowing 
(at latest at the Acq {a}).

There can be other violations before (but that is also fine) e.g.
  

  thread i:  …    Rel {a} …   Acq {b}; Write b 10; Rel {b}  …    |  … VWrite 
  thread j:  …       |    … VRead b 20     …              …   Acq {a}

We attempt to construct an unsafe_free_flowing state for the conflict 
with respect to a. But while we remove instructions from thread i, we 
encounter the violation with respect to b. As any violation of safe_free_flowing
is fine, this still fits into the proof.

The proof in more detail.
We do induction on the length of the computation.

[|safe_reachn safe_free_flowing c; init c|]
⟹
safe_reachn safe_delayed c

  Case n=0
  --------

From fact on "init c" we know that safe_free_flowing c ⟹ safe_delayed c;
safe_reach0 safe_free_flowing c <=> safe_free_flowing c ⟹ safe_delayed c <=> safe_reach0 safe_delayed c

  Case n → n+1
  -------------

Consider a trace: c(i) where c(0) = c; c(i) ⇒ c(i+1) for i <= n;
if there would be an k ≤ n for which ¬ safe_delayed c(k) we have ¬ safe_reachn safe_delayed c and by induction hypothesis
also ¬ safe_reachn safe_free_flowing c.

So we have:

∀k ≤ n. safe_delayed c(k)

¬ safe_delayed c(n+1)

Moreover we assume

∀k ≤ n + 1. safe_free_flowing c(k)

(if we have  ¬ safe_free_flowing c(k) we have ¬ safe_reachn safe_free_flowing c and are finished)

We do case distinction on '¬ safe_delayed c(n+1)'.
Some cases are trivially ruled out because of 'safe_free_flowing c(n+1)'.

Let us consider the case of 'safety violation due to an Acq A'

We get two racing threads i,j. Let j be the one issuing the 'Acq A'.

From ¬ safe_delayed c(n+1):

  A ∩ (ownsi ∪ reli) ≠ {}

From safe_free_flowing c(n+1):

 A ∩ (ownsi) = {}

hence there is an a ∈ A and a ∈ reli.

Let k < n+1 be the index where thread i did its last step in the transition from : c(k) ⇒ c(k+1).
So for c(k+1) … c(n+1) the thread configuration of i does not change.


                            last step of thread i
                                    |
c(0) ⇒ …            ⇒ c(k-1) ⇒ c(k) ⇒ c(k+1) ⇒ c(k+2) ⇒ … ⇒ c(n+1) 
  
c(k) ⇒ c(k+1) can't be a volatile or interlocked write of thread i. 
Otherwise reli would be {} beginning at c(k+1) and thus there would be no a ∈ reli. (in general only ops, where we cannot assert rel={} in the post state,i.e. reads, non-volatile writes and ghost ops)

We want to remove the step c(k) ⇒ c(k+1) from the computation (by undoing tread i's last step) and argue on 
replaying the rest of the computation:

c(0) ⇒ …            ⇒ c(k-1) ⇒ c'(k+1) ⇒ … ⇒ c'(n+1) 

At latest when we reach c'(n+1) we encounter a violation of safe_delayed c'(k) (from our initial race)
This requires a LEMMA!
As the length of the trace is now ≤ n (since step k is removed) we can 
use the induction hypothesis to obtain ¬ safe_reachn safe_free_flowing c 
and hence ¬ safe_reachn+1 safe_free_flowing c.

On the LEMMA:
We want to argue on the step's like c(k-1) ⇒ c'(k+1) and then 
continuing c'(k+1) ⇒ c'(k+2).
We can always do case distinction on 'safe_delayed'. If 'not safe_delayed' 
of any config we are fine with the main proof. So for the lemma we can 
assume "safe_delayed" of the initial state. 

Consider:

* (ts,S,m) ⇒ (ts',S',m')
* tsi = ts'i  (thread i does not change)
* safe_delayed (ts,S,m)
* (uts,uS,um) is obtained from ts by "undoing tread i's last step"
  - utsj = tsj (for j≠i)
  (- (uts,uS,um) ⇒ (ts,S,m) (by a step of thread i))
  - safe_delayed (uts,uS,um)
  - ∀a ∉ uownsi. um(a) = m(a)
show:
  (uts,uS,um) ⇒ (uts',uS',um')

where
* utsi = uts'i
* uts'j = ts'j  (for j≠i)
* ∀a ∈ uownsi. um'(a) = um(a) ( this may not be necessary to know (since i only care about steps of other threads )
* ∀a ∉ uownsi. um'(a) = m'(a)

We want to extend this simulation to a trace c(0)… c(k)
We either can simulate with a trace c'(0) … c'(k) and have safe_delayed c'(k) or 
we encounter an intermediate config ¬ safe_delayed c'(i)  for (i ≤ k) 
(from which simulation may no longer be possible).

LEMMA:
Assume:
* trace c(0) ⇒ … ⇒ c(k)
* ∀l,m ≤ k. ts(l)i = ts(m)i   (tread i does not change)
* undo config for thread i: uti, uS, um (initial new configuration)

Show:
∃trace c', x ≤ k. 
(c' simulate c up to step x, and all reached states are safe_delayed)
* ∀i ≤ x. safe_delayed c'(i)
* x < k ⟶ ¬ safe_delayed c'(x+1)
* ∀l ≤ x. ts'(l)i = uti       (tread i does not change)
* S'(0) = uS
* m'(0) = um

* ∀l ≤ x. ∀j≠i. ts'(l)j = ts(l)j            (tread j is simulated)
* ∀l ≤ x. ∀a ∉ uownsi. m'(l)(a) = m(l)(a) 
* ∀l ≤ x. ∀a ∉ uownsi. S'(l)(a) = S(l)(a) 
* ∀l,o ≤ x. ∀a ∈ uownsi. S'(l)(a) = S'(o)(a)  (sharing of thread i stays constant)

(In case we reach the final state k, we know enough about preservation of the 
ownership and sharing ghost state to (re)construct the initial race in our main proof).




######################################################################
### Thoughts on extensions ###########################################
######################################################################

There is a common ideom for concurrency control that we currently
cannot deal with properly within our programming discipline, typically for
acquiring like ops (like acquiring a lock) and releases:

Acquire
Barrier

Barrier
Release

The Acquire is an interlocked write (or a volatile write followed by a barrier)
but the release is an ordinary volatile write, where maybe there is a barrier in front of it
(I think the reason for this is to prohibit certain compiler optimizations) but there is no
barrier after it. That means in our model, that the release leaves the store-buffer-state dirty.
And we would have to insert a flush before the next volatile read.
This extra flush is currently not done. The intuitive reason behind this is, that it is
totally ok if other threads can only observe the release with a delay, since they don't do any
harm by waiting for the release (typically they just wait in an polling loop).
Ideally what we would like in our model is:

1. We want to allow releasing writes (like lock release or leave-turnstile) to leave the store-buffer-state clean, 
such that subsequent volatile reads do not require a flush.

2. The idea why this works is that other threads that may depend on the release but have not yet seen the releasing write do nothing bad, but are basically in a polling - loop (like lock acquire or wait-for-turnstile). So when the releasing write is not yet out of the store-buffer the store-buffer machine my take an extra polling loop (compared to an sequential consistent execution, where the write is considered immediately flushed to memory). But any final state (where all store-buffers are empty) should still be reachable by an sequential consistent execution.

For simulation 2. suggests that the waiting thread of the virtual machine should just stay in the same state as long as the 
sb-machine is captured in the polling loop. So the sb-machine takes extra turns, whereas the virtual machine does nothing. Formally, this is an issue for our model, since the sb-machine may read a lot of temporaries during the idle loop, all of which are not at all
read in the virtual machine. So expressing the simulation relation seems quite odd.
Another viewpoint is to introduce an intermediate (nondeterministic) program in the virtual machine, that always can descide and take an extra polling loop (without reading anything). This simulates the sb-program (modulo the extra reads to temporaries in the sb-program). In a second step we can argue that we can refine the intermediate program to the real program (all in the virtual machine) and still calculate the same *result*.

For wait-for-turnstile the pooling loop is even harder to justify being irrelevant compared to a spin-lock-acquire.
Wait-for-turnstile takes 2 samples and compares if they have changed. If not it waits and takes another sample, otherwise it continues.
So we really have to be able to compare the 2 values and deduce something from it. So they can't be just regarded as arbitrary values.

Spinlock-example:

Thread i: (release l in sb)
[l := 0 (release l)]

Thread j: (trying to acquire)
while (test&set l)...

If we maintain our flushing policy, and treat the l:=0 as nonvolatile (to keep the status clean), the release is already visible
in the virtual machine and thus the test&set succeeds emmediately, whereas in the sb-machine thraed j has to spin.

Idea:
1. Releasing writes are treated akin to non-volatile writes with respect to flushing policy 
(maybe we give them an extra name at some point)
2. To justify the non-volatile status when we want to write to them they have to be owned and *unshared*
   => This currently prohibits other threads to even read from them, which is bad for a memory cell that is used
   for synchronisation.
3. To compensate for that we relax the safety restriction for volatile reads.
   => Besides the 'shared' info we introduce a 'last' info (or maybe it is somehow merged with the shared info)  
       
      last:: address => value option

   With 'last' we store the last value written at the point we acquired the address in there (e.g. the 1 when acquiring a lock)
   Note that our ownership model hopefully guarantees that this is unique. There may be only one thread
   waking on the lock as long as the release is in its store-buffer.

   acquiring-write:
   l := 1; shared = old(shared)(l := None); owns = old(owns) + {l}; last = old(last)(l:=Some 1)
   releasing-write:
   l := 0; shared = old(shared)(l := Some True); owns = old(owns) - {l}; last=old(last)

   Volatile read's are now always allowed according to safety (maybe we introduce a flag to distinguish
   those liberated volatile reads from the things we allowed before (owned or shared)).

   


   The semantics of a volatile read is also changed to nondeterministically choose between either reading
   from memory or the last-info. Note that the semantics is *only* changed in the virtual machine. The sb machine
   stays the same.

   -------------------------
   read a t → tmps(t := m a)

   last a = Some v
   -------------------------
   read a t → tmps(t := v)

   (nondeterminism ensures that a program that always reads from memory is a refinement of our program)


By reading from last we attempt to be able to simulate the sb machine taking extra loops at program points where
the virtual machine could already proceed.

Spin-Lock example:
invariant i.  m l = 0 ==> lock is free
invariant ii. last l = Some v ==> v = 1
The virtual machine may always decide to take an extra loop (in case the lock was acquired at least once), alltough m l = 0;

Note:
* Values stored in last are always values that actually where (or will be) in memory at some point.


Turnstiles:
  enter-turnstile
  x := x + 1; shared = old(shared)(x:= None); owns = old(owns) + {x}; last = old(last)(x := Some (m x + 1));
  leave-turnstile
  x := x + 1; shared = old(shared)(x:= Some True); owns = old(owns) - {x}; last = old(last);

  wait-for turnstile
  takes a sample of x and when it is odd waits until x changes (by taking at least one other sample of x)

invariant i. even (m x) ==> not in turnstile
invariant ii. last x = Some v ==> odd(v)
Note that for turnstiles we know that only one thread will increment the counter;
(What do we need to know as invariants about the relation of last x and m x?)

To see if a turnstile is cautious enough we have various cases:
1. both samples are from memory: should be safe (if turnstiles make sense at all)
2. first sample from last, second sample from memory:
   * first sample from last: we think thread is in turnstile as it is always odd
   * second sample from memory: if its the same as the first sample we continue waiting (safe)
                                if not last was either outdated and we werent in the turnstile at all, 
                                or we have really left ts at least once (safe)
3. first sample from memory, second sample from last:
   * as there is only one writer to the ts, the sample from last is always at least the counter we have seen by the first sample (safe)
4. both samples from last
   * both sample are the same, we continue waiting (safe)
   * samples are different: As only one thread writes the ts, the 'last' samples also follow the program order, hence the
     thread has left the ts at least once (safe).



Think-tank

One might try to update 'last' not only on acquire, but simply every time a write actually hits memory.
Thinking of this 'last' actually just is the memory of the store-buffer machine. To correctly update 'last' in the
virtual machine is then difficult. Currently we maintain two values: the last one (being made interlocked or with a barrier)
and the (most actual) one in the virtual machine.
What about making last a set? Storing everything what happens between the acquire and the release (or flush)?
It seems to be a history.
So the general thing seems to be:
history: address ⇒ value list;
keeping track of all values (and their order) between two flushes (or interlocked things).
A volatile read has to be able to deal with any value in the histroy, starting from the one the
thread has read last.

If we would introduce a history for each thread, we would kind of explicitely model store-buffers.
This would be the 'ultimate solution' for TSO.: 

A histroy that should work:

history:: address ⇒ (lower-bound,upper-bound,(cnt,value) set)
history a = (lb,ub,V); (n,v) ∈ V ⟹ lb ≤ ub ∧ lb ≤ n ∧ n ≤ ub; 

* every write increases the upper-bound and adds the value;
* every read increases the lower-bound (and removes all lower values)
* every flush empties the set;

write a v: 
  h a = (lb,ub,V) ⟹ h'=h(a:=(lb,ub+1,V ∪ {(ub+1,v)}))

read a v: 
  h a = (lb,ub,V); (n,v) ∈ V; lb ≤ n ∧ n ≤ ub  ⟹ h'=h(a:=(lb+1,ub,V ))

flush a
 h'=h(a:=(0,0,{}))

            read a v    
(syntax: h ----------> h')
   

The single global histroy should be fine for our case, since there is only one writer, and the reads are ordered (as they are volatile) and are in sync between the sb and the virtual machine.

26.11.2009
More thoughts on history:
Currently we see 4 operations on the history:

Read (for stale reads), Write (non-volatile) , Commit (fence), WriteCommit (volatile write and LMS)

1. A read of a thread that has an outstanding write to the same address in the store buffer, always gets this value and
  should not have influence on the lb of the history for other threads.

2. In the concurrency control we thought of up to now (spinlocks release, turnstiles, versioning writer) 
   the only use of the lb for reads is to guarantee an order in which a thread sees updates when it really
   issues multiple reads. We currently do not need arguments across threads like (if thread i has already seen a
   certain value than thread j has to see a newer one).

1 and 2 suggest that it might be a good idea to make the lb's thread local. This gives us some nice properties:
* a read only effects the thread local lb (and neiter the lb of another thread nor the history)
* a write, commit, writecommit only effects the history and no lb.
* write, commite,writecommit are completely deterministic

* To easily model fences we take the thread id into the history.

history:: address => (hb,i,V :: timestamp ⇒ value option)
lb (per thread): address ⇒ lb


selectors: 
* top (hb,i,V) = hb
* thread (hb,i,V) = i
* vals (hb,i,V) = V 

invariants:
* ∀a hb x V. h a = (hb,x,V) ⟶ ∀t ∈ dom V. t ≤ hb

* ∀a hb x V. h a = (hb,x,V) ⟶ hb ∈ dom V 

* for all threads: ∀a. lb a ≤ top (h a)

* the memory value of the sb-machine and all the values in takeWhile (Not ∘ is_volatile_Write) are in the history of the virtual machine.
  * msb a ∈ range (vals (h a))
  * Write False a v ∈ set (takeWhile (Not ∘ is_volatile_Write)) sb ⟹ v ∈ range (vals (h a))

* The top-value in the history of the virtual machine is the memory value in the virtual machine:
  vals (h a) (top (h a)) = Some (m a)

* Commit preserves top value
  
 

h'=λa. let (hb,j,V) = h a 
       in if i=j then (hb,i,V|`{hb})
          else h a
---------------------------------------
        Commit i h h'


h a = (hb,j,V)      h'=h(a:= (hb+1,i,V(hb+1 ↦ v))) 
----------------------------------------------------------
        Write i a v h h'

WriteCommit is simply first Write then Commit (in this order, to make sure that there is only one value in V)
After a WriteCommit the value read is determined.

(* not a transition just a read *)
h a = (hb,j,V)      V t = Some v    lb ≤ t
----------------------------------------------
         Read a v t lb h h
The thread updates his lb: lb(a:=t);




*)

(* 

(\([^,]*\),\([^,]*\),\([^,]*\),\([^,]*\),\([^,]*\),\([^,]*\),\([^,]*\),\([^,]*?\))

(\([^()]*\),\([^()]*\),\([^()]*\),\([^()]*\),\([^()]*\),\([^()]*\),\([^()]*\),\([^()]*\))
(\1,\2,\3,\4,\5,\6,\8)
*)

subsection ‹Memory Instructions›

type_synonym addr = nat 
type_synonym val = nat
type_synonym tmp = nat


type_synonym tmps = "tmp  val option" 
type_synonym sop = "tmp set × (tmps  val)" ― ‹domain and function›

locale valid_sop =
fixes sop :: "sop"
assumes valid_sop: "D f θ. 
          sop=(D,f); D  dom θ 
           
          f θ = f (θ|`D)"

type_synonym memory = "addr  val"
type_synonym owns = "addr set"
type_synonym rels = "addr  bool option"
type_synonym shared = "addr  bool option" 
type_synonym acq = "addr set"
type_synonym rel = "addr set"
type_synonym lcl = "addr set"
type_synonym wrt = "addr set"
type_synonym cond = "tmps  bool"
type_synonym ret = "val  val  val"

datatype instr = Read bool addr tmp 
               | Write bool addr sop acq lcl rel wrt
               | RMW addr tmp sop cond ret acq lcl rel wrt  
               | Fence 
               | Ghost acq lcl rel wrt

type_synonym instrs = "instr list"


type_synonym ('p,'sb,'dirty,'owns,'rels) thread_config = 
  "'p × instrs × tmps × 'sb × 'dirty × 'owns × 'rels"
type_synonym ('p,'sb,'dirty,'owns,'rels,'shared) global_config = 
  "('p,'sb,'dirty,'owns,'rels) thread_config list × memory × 'shared "

definition "owned t = (let (p,instrs,θ,sb,𝒟,𝒪,) = t in 𝒪)" 

lemma owned_simp [simp]: "owned (p,instrs,θ,sb,𝒟,𝒪,) = (𝒪)"
  by (simp add: owned_def)

definition "𝒪_sb t = (let (p,instrs,θ,sb,𝒟,𝒪,) = t in (𝒪,sb))" 

lemma 𝒪_sb_simp [simp]: "𝒪_sb (p,instrs,θ,sb,𝒟,𝒪,) = (𝒪,sb)"
  by (simp add: 𝒪_sb_def)

definition "released t = (let (p,instrs,θ,sb,𝒟,𝒪,) = t in )" 

lemma released_simp [simp]: "released (p,instrs,θ,sb,𝒟,𝒪,) = ()"
  by (simp add: released_def)

lemma list_update_id': "v = xs ! i  xs[i := v] = xs"
  by simp

lemmas converse_rtranclp_induct5 = 
converse_rtranclp_induct [where a="(m,sb,𝒪,,𝒮)" and b="(m',sb',𝒪',ℛ',𝒮')", split_rule,consumes 1, case_names refl step]

subsection ‹Abstract Program Semantics›

locale memory_system = 
  fixes
  memop_step ::  "(instrs × tmps × 'sb × memory × 'dirty × 'owns × 'rels × 'shared)  
                  (instrs × tmps × 'sb × memory × 'dirty × 'owns × 'rels × 'shared)  bool" 
                    ("_ m _" [60,60] 100) and
  
  storebuffer_step:: "(memory × 'sb × 'owns × 'rels × 'shared)  (memory × 'sb × 'owns × 'rels × 'shared)  bool" ("_ sb _" [60,60] 100)


locale program =
  fixes
  program_step :: "tmps  'p  'p × instrs  bool" ("_ _ p _" [60,60,60] 100) 
  ― ‹A program only accesses the shared memory indirectly, it can read the temporaries
        and can output a sequence of memory instructions›

locale computation = memory_system + program +
  constrains
  ― ‹The constrains are only used to name the types @{typ "'sb"} and  @{typ "'p"}
  storebuffer_step:: "(memory × 'sb × 'owns × 'rels × 'shared)  (memory × 'sb × 'owns × 'rels × 'shared)  bool" and
  memop_step :: "
                  (instrs × tmps × 'sb × memory × 'dirty × 'owns × 'rels × 'shared)  
                  (instrs × tmps × 'sb × memory × 'dirty × 'owns × 'rels × 'shared)  bool" 
                     and
  program_step :: "tmps  'p  'p × instrs  bool" 
  fixes
  "record" :: "'p  'p  instrs  'sb  'sb"
begin

inductive concurrent_step :: 
  "('p,'sb,'dirty,'owns,'rels,'shared) global_config  ('p,'sb,'dirty,'owns,'rels,'shared) global_config  bool"
                               ("_  _" [60,60] 100)
where
  Program: 
   "i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,);
     θp p (p',is')   
     (ts,m,𝒮)  (ts[i:=(p',is@is',θ,record p p' is' sb,𝒟,𝒪,)],m,𝒮)"

| Memop:
   "i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,);
      (is,θ,sb,m,𝒟,𝒪,,𝒮) m (is',θ',sb',m',𝒟',𝒪',ℛ',𝒮')  
      
     (ts,m,𝒮)  (ts[i:=(p,is',θ',sb',𝒟',𝒪',ℛ')],m',𝒮')"


| StoreBuffer: 
   "i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,);
     (m,sb,𝒪,,𝒮) sb (m',sb',𝒪',ℛ',𝒮')   
     (ts,m,𝒮)  (ts[i:=(p,is,θ,sb',𝒟,𝒪',ℛ')],m',𝒮')"

definition final:: "('p,'sb,'dirty,'owns,'rels,'shared) global_config  bool"
where
 "final c = (¬ (c'. c  c'))"


lemma store_buffer_steps:
assumes sb_step: "storebuffer_step^** (m,sb,𝒪,,𝒮) (m',sb',𝒪',ℛ',𝒮')"
shows "ts. i < length ts  ts!i = (p,is,θ,sb,𝒟,𝒪,)  
         concurrent_step^** (ts,m,𝒮) (ts[i:=(p,is,θ,sb',𝒟,𝒪',ℛ')],m',𝒮')"
using sb_step 
proof (induct rule: converse_rtranclp_induct5)
  case refl then show ?case 
    by (simp add: list_update_id')
next
  case (step m sb 𝒪  𝒮 m'' sb'' 𝒪'' ℛ'' 𝒮'')
  note i_bound = i < length ts
  note ts_i = ts ! i = (p, is, θ, sb, 𝒟, 𝒪, )
  note step = (m, sb,𝒪,,𝒮) sb (m'', sb'',𝒪'',ℛ'',𝒮'')
  let ?ts' = "ts[i := (p, is, θ, sb'',𝒟, 𝒪'',ℛ'')]"
  from StoreBuffer [OF i_bound ts_i step] 
  have "(ts, m, 𝒮)  (?ts', m'', 𝒮'')".
  also
  from i_bound have i_bound': "i < length ?ts'" by simp
  from i_bound have ts'_i: "?ts'!i = (p,is,θ,sb'',𝒟,𝒪'',ℛ'')"
    by simp
  from step.hyps (3) [OF i_bound' ts'_i] i_bound
  have "concurrent_step** (?ts', m'', 𝒮'') (ts[i := (p, is, θ, sb',𝒟, 𝒪',ℛ')], m', 𝒮')"
    by (simp)
  finally
  show ?case .
qed

lemma step_preserves_length_ts: 
  assumes step: "(ts,m,𝒮)  (ts',m',𝒮')"
  shows "length ts' = length ts"
using step
apply (cases)
apply auto
done
end

lemmas concurrent_step_cases = computation.concurrent_step.cases 
[cases set, consumes 1, case_names Program Memop StoreBuffer]

definition augment_shared:: "shared  addr set  addr set  shared" ("_ ⊕⇘_ _" [61,1000,60] 61)
where
"𝒮 ⊕⇘WS  (λa. if a  S then Some (a  W) else 𝒮 a)"

definition restrict_shared:: "shared  addr set  addr set  shared" ("_ ⊖⇘_ _" [51,1000,50] 51)
where
"𝒮 ⊖⇘AL  (λa. if a  L then None 
                     else (case 𝒮 a of None  None
                            | Some writeable  Some (a  A  writeable)))"
                      
definition read_only :: "shared  addr set"
where
"read_only 𝒮  {a. (𝒮 a = Some False)}"

definition shared_le:: "shared  shared  bool" (infix "s" 50)
where 
"m1 s m2  m1 m m2  read_only m1  read_only m2"

lemma shared_leD: "m1 s m2  m1 m m2  read_only m1  read_only m2"
  by (simp add: shared_le_def)

lemma shared_le_map_le: "m1 s m2  m1 m m2"
  by (simp add: shared_le_def)

lemma shared_le_read_only_le: "m1 s m2  read_only m1  read_only m2"
  by (simp add: shared_le_def)

lemma dom_augment [simp]: "dom (m ⊕⇘WS) = dom m  S"
  by (auto simp add: augment_shared_def)

lemma augment_empty [simp]: "S  ⊕⇘x{} = S"
  by (simp add: augment_shared_def)


lemma inter_neg [simp]: "X  - L = X - L"
  by blast

lemma dom_restrict_shared [simp]: "dom (m ⊖⇘AL) = dom m - L"
  by (auto simp add: restrict_shared_def split: option.splits)

lemma restrict_shared_UNIV [simp]: "(m ⊖⇘AUNIV) = Map.empty"
  by (auto simp add: restrict_shared_def split: if_split_asm option.splits)

lemma restrict_shared_empty [simp]: "(Map.empty ⊖⇘AL) = Map.empty"
  apply (rule ext)
  by (auto simp add: restrict_shared_def split: if_split_asm option.splits)

lemma restrict_shared_in [simp]: "a  L  (m ⊖⇘AL) a = None"
  by (auto simp add: restrict_shared_def split: if_split_asm option.splits)

lemma restrict_shared_out: "a  L  (m ⊖⇘AL) a = 
  map_option (λwriteable. (a  A  writeable)) (m a)"
  by (auto simp add: restrict_shared_def split: if_split_asm option.splits)

lemma restrict_shared_out'[simp]: 
  "a  L  m a = Some writeable  (m ⊖⇘AL) a = Some (a  A  writeable)"
  by (simp add: restrict_shared_out)

lemma augment_mono_map': "A m B  (A ⊕⇘xC) m (B ⊕⇘xC)"
  by (auto simp add: augment_shared_def map_le_def domIff)

lemma augment_mono_map: "A s B  (A ⊕⇘xC) s (B ⊕⇘xC)"
  by (auto simp add:  augment_shared_def shared_le_def map_le_def read_only_def dom_def split: option.splits if_split_asm)

lemma restrict_mono_map: "A s B   (A ⊖⇘xC) s (B ⊖⇘xC)"
  by (auto simp add:  restrict_shared_def shared_le_def map_le_def read_only_def dom_def split: option.splits if_split_asm)

lemma augment_mono_aux: "dom A  dom B  dom (A ⊕⇘xC)  dom (B ⊕⇘xC)"
  by auto

lemma restrict_mono_aux: "dom A  dom B  dom (A ⊖⇘xC)  dom (B ⊖⇘xC)"
  by auto

lemma read_only_mono: "S m S'  a  read_only S  a  read_only S'"
    by (auto simp add:  map_le_def domIff read_only_def dest!: bspec)

lemma in_read_only_restrict_conv: 
  "a  read_only (𝒮 ⊖⇘AL) = (a  read_only 𝒮  a  L  a  A)"
  by (auto simp add: read_only_def restrict_shared_def split: option.splits if_split_asm)



lemma in_read_only_augment_conv: "a  read_only (𝒮 ⊕⇘WR) = (if a  R then a  W else a  read_only 𝒮)"
  by (auto simp add: read_only_def augment_shared_def)

lemmas in_read_only_convs = in_read_only_restrict_conv in_read_only_augment_conv

lemma read_only_dom: "read_only 𝒮  dom 𝒮"
  by (auto simp add: read_only_def dom_def)

lemma read_only_empty [simp]: "read_only Map.empty = {}"
  by (auto simp add: read_only_def)

lemma restrict_shared_fuse: "S ⊖⇘AL ⊖⇘BM = (S ⊖⇘(A  B)(L  M))"
apply (rule ext)
apply (auto simp add: restrict_shared_def split: option.splits if_split_asm)
done

lemma restrict_shared_empty_set [simp]: "S ⊖⇘{}{} = S"
  apply (rule ext)
  apply (auto simp add: restrict_shared_def split: option.splits if_split_asm)
  done

definition augment_rels:: "addr set  addr set  rels  rels"
where
"augment_rels S R  = (λa. if a  R
                             then (case  a of 
                                     None  Some (a  S)
                                   | Some s  Some (s  (a  S)))
                             else  a)"

declare domIff [iff del]

subsection ‹Memory Transitions›

locale gen_direct_memop_step = 
fixes emp::'rels and aug::"owns  rel  'rels  'rels"
begin
inductive gen_direct_memop_step :: "(instrs × tmps × unit × memory × bool × owns × 'rels × shared )  
                  (instrs × tmps × unit × memory × bool × owns × 'rels × shared )  bool" 
                    ("_  _" [60,60] 100)
where
  Read: "(Read volatile a t # is,θ, x, m,𝒟, 𝒪, , 𝒮) 
               (is, θ (tm a), x, m, 𝒟, 𝒪, , 𝒮)"

| WriteNonVolatile:
  "(Write False a (D,f) A L R W#is, θ, x, m, 𝒟, 𝒪, , 𝒮)  
          (is, θ, x, m(a := f θ), 𝒟, 𝒪, , 𝒮)"

| WriteVolatile:
  "(Write True a (D,f) A L R W# is, θ, x, m, 𝒟, 𝒪, , 𝒮) 
         (is, θ,  x, m(a:=f θ), True, 𝒪  A - R, emp, 𝒮 ⊕⇘WR ⊖⇘AL)"

| Fence:
  "(Fence # is, θ, x, m, 𝒟, 𝒪, , 𝒮)  (is, θ,x, m, False, 𝒪, emp, 𝒮)"

| RMWReadOnly:
  "¬ cond (θ(tm a))  
   (RMW a t (D,f) cond ret A L R W # is, θ, x, m, 𝒟, 𝒪, , 𝒮)  (is, θ(tm a),x,m, False, 𝒪, emp, 𝒮)"

| RMWWrite:
  "cond (θ(tm a))  
   (RMW a t (D,f) cond ret A L R W# is, θ, x, m, 𝒟, 𝒪, , 𝒮)  
         (is, θ(tret (m a) (f(θ(tm a)))),x, m(a:= f(θ(tm a))), False,𝒪  A - R, emp, 𝒮 ⊕⇘WR ⊖⇘AL)"

| Ghost:
  "(Ghost A L R W # is, θ, x, m, 𝒟, 𝒪, ,  𝒮) 
         (is, θ, x, m, 𝒟, 𝒪  A - R, aug (dom 𝒮) R  , 𝒮 ⊕⇘WR ⊖⇘AL)"
end

interpretation direct_memop_step: gen_direct_memop_step Map.empty augment_rels .

term direct_memop_step.gen_direct_memop_step
abbreviation direct_memop_step :: "(instrs × tmps × unit × memory × bool × owns × rels × shared )  
                  (instrs × tmps × unit × memory × bool × owns × rels × shared )  bool" 
                    ("_  _" [60,60] 100)
where
"direct_memop_step  direct_memop_step.gen_direct_memop_step"

term "x  Y"

abbreviation direct_memop_steps :: "
                  (instrs × tmps × unit × memory × bool × owns × rels × shared )  
                  (instrs × tmps × unit × memory × bool × owns × rels × shared )  
                    bool" 
                    ("_ * _" [60,60] 100)
where 
"direct_memop_steps == (direct_memop_step)^**"

term "x * Y"

interpretation virtual_memop_step: gen_direct_memop_step "()" "(λS R . ())" .

abbreviation virtual_memop_step :: "(instrs × tmps × unit × memory × bool × owns × unit × shared )  
                  (instrs × tmps × unit × memory × bool × owns × unit × shared )  bool" 
                    ("_ v _" [60,60] 100)
where
"virtual_memop_step  virtual_memop_step.gen_direct_memop_step"

term "x v Y"

abbreviation virtual_memop_steps :: "
                  (instrs × tmps × unit × memory × bool × owns × unit × shared )  
                  (instrs × tmps × unit × memory × bool × owns × unit × shared )  
                    bool" 
                    ("_ v* _" [60,60] 100)
where 
"virtual_memop_steps == (virtual_memop_step)^**"

term "x * Y"



lemma virtual_memop_step_simulates_direct_memop_step: 
  assumes step:
  "(is, θ, x, m, 𝒟, 𝒪, , 𝒮)  (is', θ', x', m', 𝒟', 𝒪', ℛ', 𝒮')"
  shows "(is, θ, x, m, 𝒟, 𝒪, (), 𝒮) v (is', θ', x', m', 𝒟', 𝒪', (), 𝒮')"
using step
apply (cases)
apply (auto intro: virtual_memop_step.gen_direct_memop_step.intros)
done

subsection ‹Safe Configurations of Virtual Machines›

inductive safe_direct_memop_state :: "owns list  nat   
                  (instrs × tmps × memory × bool × owns × shared)  bool " 
                    ("_,_ _ " [60,60,60] 100)
where
  Read: "a  𝒪  a  read_only 𝒮  (volatile  a  dom 𝒮);
          volatile  ¬ 𝒟 
        
        𝒪s,i(Read volatile a t # is, θ, m, 𝒟, 𝒪, 𝒮)"

| WriteNonVolatile:
  "a  𝒪; a  dom 𝒮 
   
   𝒪s,i(Write False a (D,f) A L R W#is, θ, m, 𝒟, 𝒪, 𝒮)"

| WriteVolatile:
  "j < length 𝒪s. ij  a  𝒪s!j;     
    A  dom 𝒮  𝒪; L  A; R  𝒪; A  R = {}; 
    j < length 𝒪s. ij   A   𝒪s!j = {};
   a  read_only 𝒮
    
   𝒪s,i(Write True a (D,f) A L R W# is, θ, m, 𝒟, 𝒪, 𝒮)"

| Fence:
  "𝒪s,i(Fence # is, θ, m, 𝒟, 𝒪, 𝒮)"

| Ghost:
  "A  dom 𝒮  𝒪; L  A; R  𝒪; A  R = {}; 
    j < length 𝒪s. ij   A  𝒪s!j = {}
    
   𝒪s,i(Ghost A L R W# is, θ, m, 𝒟, 𝒪, 𝒮)"

| RMWReadOnly:
  "¬ cond (θ(tm a)); a  𝒪  a  dom 𝒮  
   𝒪s,i(RMW a t (D,f) cond ret A L R W# is, θ, m, 𝒟, 𝒪, 𝒮)"

| RMWWrite:
  "cond (θ(tm a));  
    j < length 𝒪s. ij  a  𝒪s!j;
    A  dom 𝒮  𝒪; L  A; R  𝒪; A  R = {}; 
    j < length 𝒪s. ij  A  𝒪s!j  = {};
    a  read_only 𝒮 
    
   𝒪s,i(RMW a t (D,f) cond ret A L R W# is, θ, m, 𝒟, 𝒪, 𝒮)"

| Nil:   "𝒪s,i([], θ, m, 𝒟, 𝒪, 𝒮)"

inductive safe_delayed_direct_memop_state :: "owns list  rels list  nat   
                  (instrs × tmps × memory × bool × owns × shared)  bool " 
                    ("_,_,_ _ " [60,60,60,60] 100)
where
  Read: "a  𝒪  a  read_only 𝒮  (volatile  a  dom 𝒮);
          j < length 𝒪s. ij  (ℛs!j) a  Some False; ― ‹no release of unshared address›
          ¬ volatile  (j < length 𝒪s. ij  a  dom (ℛs!j));
          volatile  ¬ 𝒟 
        
        𝒪s,ℛs,i(Read volatile a t # is, θ, m, 𝒟, 𝒪, 𝒮)"

| WriteNonVolatile:
  "a  𝒪; a  dom 𝒮; j < length 𝒪s. ij  a  dom (ℛs!j) 
   
   𝒪s,ℛs,i(Write False a (D,f) A L R W#is, θ, m, 𝒟, 𝒪, 𝒮)"

| WriteVolatile:
  "j < length 𝒪s. ij  a  (𝒪s!j  dom (ℛs!j));     
    A  dom 𝒮  𝒪; L  A; R  𝒪; A  R = {}; 
    j < length 𝒪s. ij   A   (𝒪s!j  dom (ℛs!j)) = {};
   a  read_only 𝒮
    
   𝒪s,ℛs,i(Write True a (D,f) A L R W# is, θ, m, 𝒟, 𝒪, 𝒮)"

| Fence:
  "𝒪s,ℛs,i(Fence # is, θ, m, 𝒟, 𝒪, 𝒮)"

| Ghost:
  "A  dom 𝒮  𝒪; L  A; R  𝒪; A  R = {}; 
    j < length 𝒪s. ij   A  (𝒪s!j  dom (ℛs!j)) = {}
    
   𝒪s,ℛs,i(Ghost A L R W# is, θ, m, 𝒟, 𝒪, 𝒮)"

| RMWReadOnly:
  "¬ cond (θ(tm a)); a  𝒪  a  dom 𝒮; 
   j < length 𝒪s. ij  (ℛs!j) a  Some False ― ‹no release of unshared address›
    
   𝒪s,ℛs,i(RMW a t (D,f) cond ret A L R W# is, θ, m, 𝒟, 𝒪, 𝒮)"

| RMWWrite:
  "cond (θ(tm a));  a  𝒪  a  dom 𝒮; 
    j < length 𝒪s. ij  a  (𝒪s!j  dom (ℛs!j));
    A  dom 𝒮  𝒪; L  A; R  𝒪; A  R = {}; 
    j < length 𝒪s. ij  A  (𝒪s!j  dom (ℛs!j))  = {};
    a  read_only 𝒮 
    
   𝒪s,ℛs,i(RMW a t (D,f) cond ret A L R W# is, θ, m, 𝒟, 𝒪, 𝒮)"

| Nil:   "𝒪s,ℛs,i([], θ, m, 𝒟, 𝒪, 𝒮)"

lemma memop_safe_delayed_implies_safe_free_flowing: 
  assumes safe_delayed: "𝒪s,ℛs,i(is, θ, m, 𝒟, 𝒪, 𝒮)"
  shows "𝒪s,i(is, θ, m, 𝒟, 𝒪, 𝒮)"
using safe_delayed
proof (cases)
  case Read thus ?thesis
    by (fastforce intro!: safe_direct_memop_state.intros)
next
  case WriteNonVolatile thus ?thesis
    by (fastforce intro!: safe_direct_memop_state.intros)
next
  case WriteVolatile thus ?thesis
    by (fastforce intro!: safe_direct_memop_state.intros)
next
  case Fence thus ?thesis
    by (fastforce intro!: safe_direct_memop_state.intros)
next
  case Ghost thus ?thesis
  by (fastforce intro!: safe_direct_memop_state.Ghost)
next
  case RMWReadOnly thus ?thesis
    by (fastforce intro!: safe_direct_memop_state.intros)
next
  case RMWWrite thus ?thesis
    by (fastforce intro!: safe_direct_memop_state.RMWWrite)
next
  case Nil thus ?thesis
    by (fastforce intro!: safe_direct_memop_state.Nil)
qed

lemma memop_empty_rels_safe_free_flowing_implies_safe_delayed: 
  assumes safe: "𝒪s,i(is, θ, m, 𝒟, 𝒪, 𝒮)" 
  assumes empty: "  set ℛs.  = Map.empty"
  assumes leq: "length 𝒪s = length ℛs"
  assumes unowned_shared: "(a. (i < length 𝒪s. a  (𝒪s!i))  a  dom 𝒮)"
  assumes Os_i: "𝒪s!i = 𝒪"
  shows "𝒪s,ℛs,i(is, θ, m, 𝒟, 𝒪, 𝒮)"
using safe
proof (cases)
  case Read thus ?thesis
    using leq empty
    by (fastforce intro!: safe_delayed_direct_memop_state.Read dest: nth_mem)
next
  case WriteNonVolatile thus ?thesis
    using leq empty
    by (fastforce intro!: safe_delayed_direct_memop_state.intros dest: nth_mem)
next
  case WriteVolatile thus ?thesis
  using leq empty
    apply clarsimp
    apply (rule safe_delayed_direct_memop_state.WriteVolatile)
    apply (auto)
    apply  (drule nth_mem)
    apply  fastforce
    apply (drule nth_mem)
    apply fastforce
    done
next
  case Fence thus ?thesis
    by (fastforce intro!: safe_delayed_direct_memop_state.intros)
next
  case Ghost thus ?thesis
  using leq empty
    apply clarsimp
    apply (rule safe_delayed_direct_memop_state.Ghost)
    apply (auto)
    apply (drule nth_mem)
    apply fastforce
    done
next
  case RMWReadOnly thus ?thesis
  using leq empty
    by (fastforce intro!: safe_delayed_direct_memop_state.intros dest: nth_mem)
next
  case (RMWWrite cond t a A L R D f ret W) thus ?thesis
  using leq empty unowned_shared [rule_format, where a=a] Os_i
    apply clarsimp
    apply (rule safe_delayed_direct_memop_state.RMWWrite)
    apply (auto)
    apply  (drule nth_mem)
    apply  fastforce
    apply (drule nth_mem)
    apply fastforce
    done
next
  case Nil thus ?thesis
    by (fastforce intro!: safe_delayed_direct_memop_state.Nil)
qed


inductive id_storebuffer_step:: 
  "(memory × unit × owns × rels × shared)  (memory × unit × owns × rels × shared)  bool" ("_ I _" [60,60] 100)
where
  Id: "(m,x,𝒪,,𝒮) I (m,x,𝒪,,𝒮)"

definition empty_storebuffer_step:: "(memory × 'sb × 'owns × 'rels × 'shared)  (memory × 'sb × 'owns × 'rels × 'shared)  bool"
where
"empty_storebuffer_step c c' = False"

context program
begin

abbreviation direct_concurrent_step ::
  "('p,unit,bool,owns,rels,shared) global_config  ('p,unit,bool,owns,rels,shared) global_config  bool"
   ("_ d _" [100,60] 100)
where
  "direct_concurrent_step  
     computation.concurrent_step direct_memop_step.gen_direct_memop_step empty_storebuffer_step program_step
      (λp p' is sb. sb)"

abbreviation direct_concurrent_steps::  
  "('p,unit,bool,owns,rels,shared) global_config  ('p,unit,bool,owns,rels,shared) global_config  bool" 
    ("_ d* _" [60,60] 100)
where
"direct_concurrent_steps == direct_concurrent_step^**"  

abbreviation virtual_concurrent_step ::
  "('p,unit,bool,owns,unit,shared) global_config  ('p,unit,bool,owns,unit,shared) global_config  bool"
   ("_ v _" [100,60] 100)
where
  "virtual_concurrent_step  
     computation.concurrent_step virtual_memop_step.gen_direct_memop_step empty_storebuffer_step program_step
      (λp p' is sb. sb)"

abbreviation virtual_concurrent_steps::  
  "('p,unit,bool,owns,unit,shared) global_config  ('p,unit,bool,owns,unit,shared) global_config  bool" 
    ("_ v* _" [60,60] 100)
where
"virtual_concurrent_steps == virtual_concurrent_step^**"  

(* check if abbreviations work in both directions *)
term "x v Y"
term "x d Y"

term "x d* Y"
term "x v* Y"

end

definition
 "safe_reach step safe cfg  
        cfg'. step^** cfg cfg'  safe cfg'"

lemma safe_reach_safe_refl: "safe_reach step safe cfg  safe cfg"       
  apply (auto simp add: safe_reach_def)
  done

lemma safe_reach_safe_rtrancl: "safe_reach step safe cfg  step^** cfg cfg'  safe cfg'"       
  by (simp only: safe_reach_def)

lemma safe_reach_steps: "safe_reach step safe cfg  step^** cfg cfg'  safe_reach step safe  cfg'" 
  apply (auto simp add: safe_reach_def intro: rtranclp_trans)
  done
  
lemma safe_reach_step: "safe_reach step safe cfg  step cfg cfg'  safe_reach step safe cfg'"
  apply (erule safe_reach_steps)
  apply (erule r_into_rtranclp)
  done

context program
begin

abbreviation
 "safe_reach_direct  safe_reach direct_concurrent_step"

lemma safe_reac_direct_def':
 "safe_reach_direct safe cfg  
        cfg'. cfg d* cfg'  safe cfg'"
  by( simp add: safe_reach_def)

abbreviation
 "safe_reach_virtual  safe_reach virtual_concurrent_step"

lemma safe_reac_virtual_def':
 "safe_reach_virtual safe cfg  
        cfg'. cfg v* cfg'  safe cfg'"
  by( simp add: safe_reach_def)
end

definition
 "safe_free_flowing cfg  let (ts,m,𝒮) = cfg 
             in (i < length ts. let (p,is,θ,x,𝒟,𝒪,) = ts!i in 
                 map owned ts,i (is,θ,m,𝒟,𝒪,𝒮))"

lemma safeE: "safe_free_flowing (ts,m,𝒮);i<length ts; ts!i=(p,is,θ,x,𝒟,𝒪,)
               map owned ts,i (is,θ,m,𝒟,𝒪,𝒮)"
  by (auto simp add: safe_free_flowing_def)

definition
 "safe_delayed cfg  let (ts,m,𝒮) = cfg 
             in (i < length ts. let (p,is,θ,x,𝒟,𝒪,) = ts!i in 
                 map owned ts,map released ts,i (is,θ,m,𝒟,𝒪,𝒮))"

lemma safe_delayedE: "safe_delayed (ts,m,𝒮);i<length ts; ts!i=(p,is,θ,x,𝒟,𝒪,)
               map owned ts,map released ts,i (is,θ,m,𝒟,𝒪,𝒮)"
  by (auto simp add: safe_delayed_def)

definition "remove_rels   map (λ(p,is,θ,sb,𝒟,𝒪,). (p,is,θ,sb,𝒟,𝒪,()))"

theorem (in program) virtual_simulates_direct_step:
  assumes step: "(ts,m,𝒮) d (ts',m',𝒮')"
  shows "(remove_rels ts,m,𝒮) v (remove_rels ts',m',𝒮')"
using step
proof -
  interpret direct_computation:
    computation direct_memop_step empty_storebuffer_step program_step "λp p' is sb. sb" .
  interpret virtual_computation:
    computation virtual_memop_step empty_storebuffer_step program_step "λp p' is sb. sb" .
  from step show ?thesis
  proof (cases)
    case (Program j p "is" θ sb 𝒟 𝒪  p' is')
    then obtain
      ts': "ts' = ts[j:=(p',is@is',θ,sb,𝒟,𝒪,)]" and
      𝒮': "𝒮'=𝒮" and
      m': "m'=m" and
      j_bound: "j < length ts" and
      ts_j: "ts!j = (p,is,θ,sb,𝒟,𝒪,)" and
      prog_step: "θ p p (p', is')"
      by auto
    from ts_j j_bound have 
      vts_j: "remove_rels ts!j = (p,is,θ,sb,𝒟,𝒪,())" by (auto simp add: remove_rels_def)
    
    from virtual_computation.Program [OF _ vts_j prog_step, of m 𝒮] j_bound ts'
    show ?thesis
      by (clarsimp simp add: 𝒮' m' remove_rels_def map_update)
  next
    case (Memop j p "is" θ sb 𝒟 𝒪  is' θ' sb' 𝒟' 𝒪' ℛ')
    then obtain
      ts': "ts' = ts[j:=(p,is',θ',sb',𝒟',𝒪',ℛ')]" and
      j_bound: "j < length ts" and
      ts_j: "ts!j = (p,is,θ,sb,𝒟,𝒪,)"  and
      mem_step: "(is, θ, sb, m, 𝒟, 𝒪, , 𝒮)  (is', θ', sb',m', 𝒟',  𝒪', ℛ', 𝒮')"
      by auto

    from ts_j j_bound have 
      vts_j: "remove_rels ts!j = (p,is,θ,sb,𝒟,𝒪,())" by (auto simp add: remove_rels_def)

    from virtual_computation.Memop[OF _ vts_j virtual_memop_step_simulates_direct_memop_step [OF mem_step]] j_bound ts'
    show ?thesis
      by (clarsimp simp add: remove_rels_def map_update)
  next
    case (StoreBuffer _ p "is" θ sb 𝒟 𝒪  sb' 𝒪' ℛ')
    hence False 
      by (auto simp add: empty_storebuffer_step_def)
    thus ?thesis ..
  qed
qed

lemmas converse_rtranclp_induct_sbh_steps = converse_rtranclp_induct
[of _ "(ts,m,𝒮)" "(ts',m',𝒮')", split_rule,
   consumes 1, case_names refl step]


theorem (in program) virtual_simulates_direct_steps:
  assumes steps: "(ts,m,𝒮) d* (ts',m',𝒮')"
  shows "(remove_rels ts,m,𝒮) v* (remove_rels ts',m',𝒮')"
using steps
proof (induct rule: converse_rtranclp_induct_sbh_steps)
  case refl thus ?case by auto
next
  case (step ts m 𝒮 ts'' m'' 𝒮'')
  then obtain 
    first: "(ts, m, 𝒮) d (ts'', m'', 𝒮'')" and
    hyp: "(remove_rels ts'', m'', 𝒮'') v* (remove_rels ts', m', 𝒮')"
    by blast
  note virtual_simulates_direct_step [OF first] also note hyp
  finally
  show ?case by blast
qed

locale simple_ownership_distinct =
fixes ts::"('p,'sb,'dirty,owns,'rels) thread_config list"
assumes simple_ownership_distinct:
   "i j pi isi 𝒪i i 𝒟i θi sbi pj isj 𝒪j j 𝒟j θj sbj. 
      i < length ts; j < length ts; i  j; 
    ts!i = (pi,isi,θi,sbi,𝒟i,𝒪i,i); ts!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)
        𝒪i  𝒪j = {}"

lemma (in simple_ownership_distinct)
  simple_ownership_distinct_nth_update:
 "i p is θ 𝒪  𝒟 xs sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,);
    j < length ts. ij  (let (pj,isj,θj,sbj,𝒟j,𝒪j,j) = ts!j 
          in (𝒪')  (𝒪j) ={})   
     simple_ownership_distinct (ts[i := (p',is',θ',sb',𝒟',𝒪',ℛ')])"
  apply (unfold_locales)
  apply (clarsimp simp add: nth_list_update split: if_split_asm)
  apply   (force dest: simple_ownership_distinct simp add: Let_def)
  apply  (fastforce dest: simple_ownership_distinct simp add: Let_def)
  apply (fastforce dest: simple_ownership_distinct simp add: Let_def)
  done

locale read_only_unowned =
fixes 𝒮::shared and ts::"('p,'sb,'dirty,owns,'rels) thread_config list"
assumes read_only_unowned:
  "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)  
   
   𝒪  read_only 𝒮 = {}"

lemma (in read_only_unowned)
  read_only_unowned_nth_update:
 "i p is 𝒪  𝒟 acq θ sb. 
   i < length ts; 𝒪  read_only 𝒮  = {}  
     read_only_unowned 𝒮 (ts[i := (p,is,θ,sb,𝒟,𝒪,)])"
  apply (unfold_locales)
  apply   (auto dest: read_only_unowned
       simp add:  nth_list_update split: if_split_asm)
  done

locale unowned_shared =
fixes 𝒮::shared and ts::"('p,'sb,'dirty,owns,'rels) thread_config list"
assumes unowned_shared: "-  ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts)  dom 𝒮"

lemma (in unowned_shared)
  unowned_shared_nth_update:
  assumes i_bound: "i < length ts" 
  assumes ith: "ts!i=(p,is,xs,sb,𝒟,𝒪,)" 
  assumes subset: "𝒪   𝒪'"
  shows "unowned_shared 𝒮 (ts[i := (p',is',xs',sb',𝒟',𝒪',ℛ')])"
proof -
  from i_bound ith subset
  have " ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts)  
         ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set (ts[i := (p',is',xs',sb',𝒟',𝒪',ℛ')]))"

    apply (auto simp add: in_set_conv_nth nth_list_update split: if_split_asm)
    subgoal for x p'' is'' xs'' sb'' 𝒟''  𝒪'' ℛ'' j
    apply (case_tac "j=i")
    apply  (rule_tac x="(p',is',xs',sb',𝒟',𝒪',ℛ')" in bexI)
    apply   fastforce
    apply  (fastforce simp add: in_set_conv_nth)
    apply (rule_tac x="(p'',is'',xs'',sb'',𝒟'',𝒪'',ℛ'')" in bexI)
    apply  fastforce
    apply (fastforce simp add: in_set_conv_nth)
    done 
    done
  hence "-  ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set (ts[i := (p',is',xs',sb',𝒟',𝒪',ℛ')]))  
         -  ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts)"
    by blast
  also note unowned_shared
  finally
  show ?thesis
    by (unfold_locales)
qed

lemma (in unowned_shared) a_unowned_by_others_owned_or_shared:
  assumes i_bound: "i < length ts"
  assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)"
  assumes a_unowned_others:
        "j<length (map owned ts). i  j  
          (let 𝒪j = (map owned ts)!j in a  𝒪j)" 

  shows "a  𝒪  a  dom 𝒮"
proof -
  {
    fix j pj isj 𝒪j j 𝒟j xsj sbj
    assume a_unowned: "a  𝒪"
    assume j_bound: "j < length ts"
    assume jth: "ts!j = (pj,isj,xsj, sbj, 𝒟j, 𝒪j,j)"
    have "a  𝒪j"
    proof (cases "i=j")
      case True with a_unowned ts_i jth
      show ?thesis
	by auto
    next
      case False
      from a_unowned_others [rule_format, of j] j_bound jth False
      show ?thesis
	by auto
    qed
  } note lem = this
  {
    assume "a  𝒪"
    from lem [OF this]
    have "a   -  ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts)"
      by (fastforce simp add: in_set_conv_nth)
    with unowned_shared have "a  dom 𝒮"
      by auto
  } 
  then
  show ?thesis
    by auto
qed

lemma (in unowned_shared) unowned_shared': 
  assumes notin: "i < length ts. a  owned (ts!i)"
  shows "a  dom 𝒮"
proof -
  from notin  have "a  - ((λ(_, _, _, _, _, 𝒪, _). 𝒪) ` set ts)"
    by (force simp add: in_set_conv_nth)
  with unowned_shared show ?thesis by blast
qed

lemma unowned_shared_def': "unowned_shared 𝒮 ts = (a. (i < length ts. a  owned (ts!i))  a  dom 𝒮)"
apply rule
apply  clarsimp
apply (rule unowned_shared.unowned_shared')
apply   fastforce
apply  fastforce
apply (unfold unowned_shared_def)
apply clarsimp
subgoal for x
apply (drule_tac x=x in spec)
apply (erule impE)
apply  clarsimp
apply  (case_tac "(ts!i)")
apply  (drule nth_mem)
apply auto
done
done 

definition
 "initial cfg  let (ts,m,𝒮) = cfg 
             in unowned_shared 𝒮 ts 
                (i < length ts. let (p,is,θ,x,𝒟,𝒪,) = ts!i in 
                   = Map.empty )"

lemma initial_empty_rels: "initial (ts,m,𝒮)    set (map released ts).  = Map.empty"
  by (fastforce simp add: initial_def simp add: in_set_conv_nth)

lemma initial_unowned_shared: "initial (ts,m,𝒮)  unowned_shared 𝒮 ts"
  by (fastforce simp add: initial_def )

lemma initial_safe_free_flowing_implies_safe_delayed:
assumes init: "initial c" 
assumes safe: "safe_free_flowing c"
shows "safe_delayed c"
proof -
  obtain ts 𝒮 m where c: "c=(ts,m,𝒮)" by (cases c)
  from initial_empty_rels [OF init [simplified c]]
  have rels_empty: "set (map released ts).  = Map.empty".
  from initial_unowned_shared [OF init [simplified c]] have "unowned_shared 𝒮 ts"
    by auto
  hence us:"(a. (i < length (map owned ts). a  (map owned ts!i))  a  dom 𝒮)"
    by (simp add:unowned_shared_def')
  {
    fix i p "is" θ x 𝒟 𝒪 
    assume i_bound: "i < length ts"
    assume ts_i: "ts!i = (p,is,θ,x,𝒟,𝒪,)"
    have "map owned ts,map released ts,i (is,θ,m,𝒟,𝒪,𝒮)"
    proof -
      from safeE [OF safe [simplified c] i_bound ts_i] 
      have "map owned ts,i(is, θ, m, 𝒟, 𝒪, 𝒮)".
      from memop_empty_rels_safe_free_flowing_implies_safe_delayed [OF this rels_empty _ us] i_bound ts_i
      show ?thesis
        by simp
    qed
  }
  then show ?thesis 
    by (fastforce simp add: c safe_delayed_def)
qed


locale program_progress = program +
assumes progress: "θ p p (p',is')  p'  p  is'  []"
text ‹The assumption `progress' could be avoided if we introduce stuttering steps in lemma undo_local_step›
or make the scheduling of threads explicit, such that we can directly express that `thread i does not make a step'.
›

lemma (in program_progress) undo_local_step:
assumes step: "(ts,m,𝒮) d (ts',m',𝒮')"
assumes i_bound: "i < length ts"
assumes unchanged: "ts!i = ts'!i"
assumes safe_delayed_undo: "safe_delayed (u_ts,u_m,u_shared)" ― ‹proof should also work with weaker @{const safe_free_flowing}
assumes leq: "length u_ts = length ts"
assumes others_same: "j < length ts. ji  u_ts!j = ts!j"
assumes u_ts_i: "u_ts!i=(u_p,u_is,u_tmps,u_x,u_dirty,u_owns,u_rels)"
assumes u_m_other: "a. a  u_owns  u_m a = m a"
assumes u_m_shared: "a. a  u_owns  a  dom u_shared  u_m a = m a"
assumes u_shared: "a. a  u_owns  a  owned (ts!i)  u_shared a = 𝒮 a"
assumes dist: "simple_ownership_distinct u_ts"
assumes dist_ts: "simple_ownership_distinct ts"
shows "u_ts' u_shared' u_m'. (u_ts,u_m,u_shared) d (u_ts',u_m',u_shared') 
         ― ‹thread i is unchanged›
         u_ts'!i = u_ts!i 
         (a  u_owns. u_shared' a = u_shared a) 
         (a  u_owns. 𝒮' a = 𝒮 a) 
         (a  u_owns. u_m' a = u_m a) 
         (a  u_owns. m' a = m a) 

         ― ‹other threads are simulated›
         (j < length ts. ji  u_ts'!j = ts'!j) 
         (a. a  u_owns  a  owned (ts!i)  u_shared' a = 𝒮' a) 
         (a. a  u_owns  u_m' a = m' a)"
proof -
  interpret direct_computation:
    computation direct_memop_step empty_storebuffer_step program_step "λp p' is sb. sb" .
  from dist interpret simple_ownership_distinct u_ts .
  from step
  show ?thesis
  proof (cases)
    case (Program j p "is" θ sb 𝒟 𝒪  p' is')
    then obtain
      ts': "ts' = ts[j:=(p',is@is',θ,sb,𝒟,𝒪,)]" and
      𝒮': "𝒮'=𝒮" and
      m': "m'=m" and
      j_bound: "j < length ts" and
      ts_j: "ts!j = (p,is,θ,sb,𝒟,𝒪,)" and
      prog_step: "θ p p (p', is')"
      by auto

    from progress [OF prog_step] i_bound unchanged ts_j ts'
    have neq_j_i: "ji"
      by auto


    from others_same [rule_format, OF j_bound neq_j_i] ts_j
    have u_ts_j: "u_ts!j = (p,is,θ,sb,𝒟,𝒪,)"
      by simp
    from leq j_bound have j_bound': "j < length u_ts"
      by simp
    from leq i_bound have i_bound': "i < length u_ts"
      by simp

    from direct_computation.Program [OF j_bound' u_ts_j prog_step]
    have ustep:" (u_ts,u_m, u_shared) d (u_ts[j := (p', is @ is', θ, sb, 𝒟, 𝒪, )], u_m, u_shared)".  show ?thesis
      apply -
      apply (rule exI)
      apply (rule exI)
      apply (rule exI)
      apply (rule conjI)
      apply (rule ustep)
      using neq_j_i others_same u_m_other u_shared j_bound leq ts_j 
      apply (auto simp add: nth_list_update ts' 𝒮' m')
      done
  next
    case (Memop j p "is" θ sb 𝒟 𝒪  is' θ' sb' 𝒟' 𝒪' ℛ')
    then obtain
      ts': "ts' = ts[j:=(p,is',θ',sb',𝒟',𝒪',ℛ')]" and
      j_bound: "j < length ts" and
      ts_j: "ts!j = (p,is,θ,sb,𝒟,𝒪,)"  and
      mem_step: "(is, θ, sb, m, 𝒟, 𝒪, , 𝒮)  (is', θ', sb',m', 𝒟',  𝒪', ℛ', 𝒮')"
      by auto

    from mem_step i_bound unchanged ts_j
    have neq_j_i: "ji"
      by cases (auto simp add: ts')

    from others_same [rule_format, OF j_bound neq_j_i] ts_j
    have u_ts_j: "u_ts!j = (p,is,θ,sb,𝒟,𝒪,)"
      by simp
    from leq j_bound have j_bound': "j < length u_ts"
      by simp
    from leq i_bound have i_bound': "i < length u_ts"
      by simp
    from safe_delayedE [OF safe_delayed_undo j_bound' u_ts_j]
    have safe_j: "map owned u_ts,map released u_ts,j(is, θ, u_m, 𝒟, 𝒪, u_shared)".
    from simple_ownership_distinct [OF j_bound' i_bound' neq_j_i u_ts_j u_ts_i]
    have owns_u_owns: "𝒪  u_owns = {}" .
    from mem_step
    show ?thesis
    proof (cases)
      case (Read volatile a t)
      then obtain
        "is": "is = Read volatile a t # is'" and
        θ': "θ' = θ(t  m a)" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from safe_j [simplified "is"]
      obtain 
        access_cond: "a  𝒪  a  read_only u_shared  
                   (volatile  a  dom u_shared)" 
      
        and
        clean: "volatile  ¬ 𝒟"
        by cases auto
      have u_m_a_eq: "u_m a = m a"
      proof (cases "a  u_owns")
        case True
        with simple_ownership_distinct [OF j_bound' i_bound' neq_j_i u_ts_j u_ts_i]
        have "a  𝒪" by auto
        with access_cond read_only_dom [of u_shared] have "a  dom u_shared"
          by auto
        from u_m_shared [rule_format, OF True this] 
        show ?thesis .
      next
        case False
        from u_m_other [rule_format, OF this]
        show ?thesis .
      qed
      note Read' = direct_memop_step.Read [of volatile a t "is'" θ sb u_m 𝒟 𝒪  u_shared]
      from direct_computation.Memop [OF j_bound' u_ts_j, simplified "is", OF Read' ] 
      have ustep: "(u_ts, u_m, u_shared) d (u_ts[j := (p, is', θ(t  u_m a), sb, 𝒟, 𝒪, )], u_m, u_shared)".
      show ?thesis
        apply -
        apply (rule exI)
        apply (rule exI)
        apply (rule exI)
        apply (rule conjI)
        apply (rule ustep)
        using neq_j_i others_same u_m_other u_shared j_bound leq ts_j 
        by (auto simp add: nth_list_update ts' eqs' u_m_a_eq)
    next
      case (WriteNonVolatile a D f A L R W)
      then obtain
        "is": "is = Write False a (D, f) A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m(a:=f θ)" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from safe_j [simplified "is"]
      obtain
        owned: "a  𝒪" and unshared: "a  dom u_shared"
        by cases auto

      from simple_ownership_distinct [OF j_bound' i_bound' neq_j_i u_ts_j u_ts_i] owned
      have a_unowned_i: "a  u_owns"
        by auto
      note Write' = direct_memop_step.WriteNonVolatile [of a D f A L R W is' θ sb u_m 𝒟 𝒪  u_shared]
      from direct_computation.Memop [OF j_bound' u_ts_j, simplified "is", OF Write' ] 
      have ustep: "(u_ts, u_m, u_shared) d (u_ts[j := (p, is', θ, sb, 𝒟, 𝒪, )], u_m (a := f θ), u_shared)".
      show ?thesis
        apply -
        apply (rule exI)
        apply (rule exI)
        apply (rule exI)
        apply (rule conjI)
        apply (rule ustep)
        using neq_j_i others_same u_m_other u_shared j_bound leq ts_j a_unowned_i
        apply (auto simp add: nth_list_update ts' eqs')
        done
    next
      case (WriteVolatile a D f A L R W)
      then obtain
        "is": "is = Write True a (D, f) A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m(a:=f θ)" and
        𝒟': "𝒟'=True" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from safe_j [simplified "is"]
      obtain
        a_unowned_others: "k < length u_ts. jk  a  (map owned u_ts!k  dom (map released u_ts!k))" and
        A: "A  dom u_shared  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length u_ts. jk   A   (map owned u_ts!k  dom (map released u_ts!k)) = {}" and
        a_not_ro: "a  read_only u_shared"
        by cases auto

      note Write' = direct_memop_step.WriteVolatile [of a D f A L R W is' θ sb u_m 𝒟 𝒪  u_shared]
      from direct_computation.Memop [OF j_bound' u_ts_j, simplified "is", OF Write' ] 
      have ustep: "(u_ts, u_m, u_shared) d 
                   (u_ts[j := (p, is', θ, sb, True, 𝒪  A - R, Map.empty)], u_m (a := f θ), u_shared ⊕⇘WR ⊖⇘AL)".

      from A_unowned_others [rule_format, OF i_bound' neq_j_i] u_ts_i i_bound' 
      have A_u_owns: "A  u_owns = {}" by auto
      {
        fix a
        assume a_u_owns: "a  u_owns"
        have "(u_shared ⊕⇘WR ⊖⇘AL) a = u_shared a"
        using R_owns A_R L_A A_u_owns owns_u_owns a_u_owns
          by (auto simp add: restrict_shared_def augment_shared_def split: option.splits)
      }
      note u_owned_shared = this
      from a_unowned_others [rule_format, OF i_bound' neq_j_i] u_ts_i i_bound'  have a_u_owns: "a  u_owns" by auto
      {
        fix a
        assume a_u_owns: "a  u_owns"
        assume a_u_owns_orig: "a  owned (ts!i)"
        from u_shared [rule_format, OF a_u_owns a_u_owns_orig]
        have "(u_shared ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
        using R_owns A_R L_A A_u_owns owns_u_owns 
          by (auto simp add: restrict_shared_def augment_shared_def split: option.splits)
      }
      note u_unowned_shared = this
      {
        fix a
        assume a_u_owns: "a  u_owns"

        have "(𝒮 ⊕⇘WR ⊖⇘AL) a = 𝒮 a"
        using R_owns A_R L_A A_u_owns owns_u_owns a_u_owns
          by (auto simp add: restrict_shared_def augment_shared_def split: option.splits)
      }
      note 𝒮'_shared = this

      show ?thesis
        apply -
        apply (rule exI)
        apply (rule exI)
        apply (rule exI)
        apply (rule conjI)
        apply (rule ustep)
        using neq_j_i others_same u_m_other u_shared j_bound leq ts_j u_owned_shared a_u_owns u_unowned_shared 𝒮'_shared
        apply (auto simp add: nth_list_update ts' eqs')
        done
    next      
      case Fence
      then obtain
        "is": "is = Fence # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      note Fence' = direct_memop_step.Fence [of is' θ sb u_m 𝒟 𝒪  u_shared]
      from direct_computation.Memop [OF j_bound' u_ts_j, simplified "is", OF Fence' ] 
      have ustep: "(u_ts, u_m, u_shared) d (u_ts[j := (p, is', θ, sb, False, 𝒪, Map.empty)], u_m, u_shared)".
      show ?thesis
        apply -
        apply (rule exI)
        apply (rule exI)
        apply (rule exI)
        apply (rule conjI)
        apply (rule ustep)
        using neq_j_i others_same u_m_other u_shared j_bound leq ts_j 
        by (auto simp add: nth_list_update ts' eqs' )
    next
      case (RMWReadOnly 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
        θ': "θ' = θ(t m a)" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮" and
        cond: "¬ cond (θ(t  m a))"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      from safe_j [simplified "is"] owns_u_owns u_ts_i i_bound' neq_j_i
      obtain 
        access_cond: "a  u_owns  (a  dom u_shared  a  u_owns)"
        by cases auto
        
      from u_m_other u_m_shared access_cond
      have u_m_a_eq: "u_m a = m a"
        by auto
      from cond u_m_a_eq have cond': "¬ cond (θ(t  u_m a))"
        by auto
      note RMWReadOnly' = direct_memop_step.RMWReadOnly [of cond θ t u_m a D f ret A L R W is' sb 𝒟 𝒪  u_shared, 
        OF cond']
      from direct_computation.Memop [OF j_bound' u_ts_j, simplified "is", OF RMWReadOnly' ] 
      have ustep: "(u_ts, u_m, u_shared) d (u_ts[j := (p, is', θ(t  u_m a), sb, False, 𝒪, Map.empty)], u_m, u_shared)".
      show ?thesis
        apply -
        apply (rule exI)
        apply (rule exI)
        apply (rule exI)
        apply (rule conjI)
        apply (rule ustep)
        using neq_j_i others_same u_m_other u_shared j_bound leq ts_j
        by (auto simp add: nth_list_update ts' eqs' u_m_a_eq)
    next
      case (RMWWrite 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
        θ': "θ' = θ(t  ret (m a) (f (θ(t  m a))))" and
        sb': "sb'=sb" and
        m': "m'=m(a := f (θ(t  m a)))" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL" and
        cond: "cond (θ(t  m a))"
        by auto

      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      from safe_j [simplified "is"] owns_u_owns u_ts_i i_bound' neq_j_i
      obtain 
        access_cond: "a  u_owns  (a  dom u_shared  a  u_owns)"
        by cases auto
        
      from u_m_other u_m_shared access_cond
      have u_m_a_eq: "u_m a = m a"
        by auto
      from cond u_m_a_eq have cond': "cond (θ(t  u_m a))"
        by auto
      from safe_j [simplified "is"] cond'
      obtain
        a_unowned_others: "k < length u_ts. jk  a  (map owned u_ts!k  dom (map released u_ts!k))" and
        A: "A  dom u_shared  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length u_ts. jk   A   (map owned u_ts!k  dom (map released u_ts!k)) = {}" and
        a_not_ro: "a  read_only u_shared"
        by cases auto

      note Write' = direct_memop_step.RMWWrite [of cond θ t u_m a D f ret A L R W is' sb 𝒟 𝒪  u_shared, 
        OF cond']
      from direct_computation.Memop [OF j_bound' u_ts_j, simplified "is", OF Write' ] 
      have ustep: "(u_ts, u_m, u_shared) d 
                   (u_ts[j := (p, is', θ(t  ret (u_m a) (f (θ(t  u_m a)))), sb, False, 𝒪  A - R, Map.empty)], u_m(a := f (θ(t  u_m a))), 
                    u_shared ⊕⇘WR ⊖⇘AL)".

      from A_unowned_others [rule_format, OF i_bound' neq_j_i] u_ts_i i_bound' 
      have A_u_owns: "A  u_owns = {}" by auto
      {
        fix a
        assume a_u_owns: "a  u_owns"
        have "(u_shared ⊕⇘WR ⊖⇘AL) a = u_shared a"
        using R_owns A_R L_A A_u_owns owns_u_owns a_u_owns
          by (auto simp add: restrict_shared_def augment_shared_def split: option.splits)
      }
      note u_owned_shared = this
      from a_unowned_others [rule_format, OF i_bound' neq_j_i] u_ts_i i_bound'  have a_u_owns: "a  u_owns" by auto
      {
        fix a
        assume a_u_owns: "a  u_owns"
        assume a_u_owns_orig: "a  owned (ts!i)"
        from u_shared [rule_format, OF a_u_owns this]
        have "(u_shared ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
        using R_owns A_R L_A A_u_owns owns_u_owns 
          by (auto simp add: restrict_shared_def augment_shared_def split: option.splits)
      }
      note u_unowned_shared = this
      {
        fix a
        assume a_u_owns: "a  u_owns"

        have "(𝒮 ⊕⇘WR ⊖⇘AL) a = 𝒮 a"
        using R_owns A_R L_A A_u_owns owns_u_owns a_u_owns
          by (auto simp add: restrict_shared_def augment_shared_def split: option.splits)
      }
      note 𝒮'_shared = this
      show ?thesis
        apply -
        apply (rule exI)
        apply (rule exI)
        apply (rule exI)
        apply (rule conjI)
        apply (rule ustep)
        using neq_j_i others_same u_m_other u_shared j_bound leq ts_j u_owned_shared a_u_owns u_unowned_shared 𝒮'_shared
        apply (auto simp add: nth_list_update ts' eqs')
        done
    next
      case (Ghost A L R W) 
      then obtain
        "is": "is = Ghost A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=augment_rels (dom 𝒮) R " and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from safe_j [simplified "is"]
      obtain
        A: "A  dom u_shared  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length u_ts. jk   A   (map owned u_ts!k  dom (map released u_ts!k)) = {}" 
        by cases auto

      note Ghost' = direct_memop_step.Ghost [of A L R W is' θ sb u_m 𝒟 𝒪  u_shared]
      from direct_computation.Memop [OF j_bound' u_ts_j, simplified "is", OF Ghost' ] 
      have ustep: "(u_ts, u_m, u_shared) d 
                   (u_ts[j := (p, is', θ, sb, 𝒟, 𝒪  A - R, augment_rels (dom u_shared) R  )], u_m, 
                          u_shared ⊕⇘WR ⊖⇘AL)".

      from A_unowned_others [rule_format, OF i_bound' neq_j_i] u_ts_i i_bound' 
      have A_u_owns: "A  u_owns = {}" by auto
      {
        fix a
        assume a_u_owns: "a  u_owns"
        have "(u_shared ⊕⇘WR ⊖⇘AL) a = u_shared a"
        using R_owns A_R L_A A_u_owns owns_u_owns a_u_owns
          by (auto simp add: restrict_shared_def augment_shared_def split: option.splits)
      }
      note u_owned_shared = this
      {
        fix a
        assume a_u_owns: "a  u_owns"
        assume "a  owned (ts!i)"
        from u_shared [rule_format, OF a_u_owns this]
        have "(u_shared ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
        using R_owns A_R L_A A_u_owns owns_u_owns 
          by (auto simp add: restrict_shared_def augment_shared_def split: option.splits)
      }
      note u_unowned_shared = this

      {
        fix a
        assume a_u_owns: "a  u_owns"

        have "(𝒮 ⊕⇘WR ⊖⇘AL) a = 𝒮 a"
        using R_owns A_R L_A A_u_owns owns_u_owns a_u_owns
          by (auto simp add: restrict_shared_def augment_shared_def split: option.splits)
      }
      note 𝒮'_shared = this

      from dist_ts
      interpret dist_ts_inter: simple_ownership_distinct ts .
      from dist_ts_inter.simple_ownership_distinct [OF j_bound i_bound neq_j_i ts_j] 
      have "𝒪  owned (ts!i) = {}"
        apply (cases "ts!i")
        apply fastforce+
        done
      
      with simple_ownership_distinct [OF j_bound' i_bound' neq_j_i u_ts_j u_ts_i] R_owns u_shared

      have augment_eq: "augment_rels (dom u_shared) R  = augment_rels (dom 𝒮) R "
        apply -
        apply (rule ext)
        apply (fastforce simp add: augment_rels_def split: option.splits simp add: domIff)        
        done
        
        
      show ?thesis
        apply -
        apply (rule exI)
        apply (rule exI)
        apply (rule exI)
        apply (rule conjI)
        apply (rule ustep)
        using neq_j_i others_same u_m_other u_shared j_bound leq ts_j u_owned_shared u_unowned_shared 𝒮'_shared
        apply (auto simp add: nth_list_update ts' eqs' augment_eq)
        done
    qed
  next
    case (StoreBuffer _ p "is" θ sb 𝒟 𝒪  sb' 𝒪' ℛ')
    hence False 
      by (auto simp add: empty_storebuffer_step_def)
    thus ?thesis ..
  qed
qed


theorem (in program) safe_step_preserves_simple_ownership_distinct:
  assumes step: "(ts,m,𝒮) d (ts',m',𝒮')"
  assumes safe: "safe_delayed (ts,m,𝒮)"
  assumes dist: "simple_ownership_distinct ts"
  shows "simple_ownership_distinct ts'"
proof -
  interpret direct_computation:
    computation direct_memop_step empty_storebuffer_step program_step "λp p' is sb. sb" .
  from dist interpret simple_ownership_distinct ts .
  from step
  show ?thesis
  proof (cases)
    case (Program j p "is" θ sb 𝒟 𝒪  p' is')
    then obtain
      ts': "ts' = ts[j:=(p',is@is',θ,sb,𝒟,𝒪,)]" and
      𝒮': "𝒮'=𝒮" and
      m': "m'=m" and
      j_bound: "j < length ts" and
      ts_j: "ts!j = (p,is,θ,sb,𝒟,𝒪,)" and
      prog_step: "θ p p (p', is')"
      by auto

    from simple_ownership_distinct [OF j_bound _ _ ts_j]
    show "simple_ownership_distinct ts'"
      apply (simp only: ts')
      apply (rule simple_ownership_distinct_nth_update [OF j_bound ts_j])
      apply force
      done
  next
    case (Memop j p "is" θ sb 𝒟 𝒪  is' θ' sb' 𝒟' 𝒪' ℛ')
    then obtain
      ts': "ts' = ts[j:=(p,is',θ',sb',𝒟',𝒪',ℛ')]" and
      j_bound: "j < length ts" and
      ts_j: "ts!j = (p,is,θ,sb,𝒟,𝒪,)"  and
      mem_step: "(is, θ, sb, m, 𝒟, 𝒪, , 𝒮)  (is', θ', sb',m', 𝒟',  𝒪', ℛ', 𝒮')"
      by auto

    from safe_delayedE [OF safe j_bound ts_j]
    have safe_j: "map owned ts,map released ts,j(is, θ, m, 𝒟, 𝒪, 𝒮)".
    from mem_step
    show ?thesis
    proof (cases)
      case (Read volatile a t)
      then obtain
        "is": "is = Read volatile a t # is'" and
        θ': "θ' = θ(t  m a)" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from simple_ownership_distinct [OF j_bound _ _ ts_j]
      show "simple_ownership_distinct ts'"
        apply (simp only: ts' eqs')
        apply (rule simple_ownership_distinct_nth_update [OF j_bound ts_j])
        apply force
        done

    next
      case (WriteNonVolatile a D f A L R W)
      then obtain
        "is": "is = Write False a (D, f) A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m(a:=f θ)" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      from simple_ownership_distinct [OF j_bound _ _ ts_j]
      show "simple_ownership_distinct ts'"
        apply (simp only: ts' eqs')
        apply (rule simple_ownership_distinct_nth_update [OF j_bound ts_j])
        apply force
        done

    next
      case (WriteVolatile a D f A L R W)
      then obtain
        "is": "is = Write True a (D, f) A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m(a:=f θ)" and
        𝒟': "𝒟'=True" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from safe_j [simplified "is"]
      obtain
        a_unowned_others: "k < length ts. jk  a  (map owned ts!k  dom (map released ts!k))" and
        A: "A  dom 𝒮  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length ts. jk   A   (map owned ts!k  dom (map released ts!k)) = {}" and
        a_not_ro: "a  read_only 𝒮"
        by cases auto
      from simple_ownership_distinct [OF j_bound _ _ ts_j] R_owns A_R A_unowned_others
      show "simple_ownership_distinct ts'"
        apply (simp only: ts' eqs')
        apply (rule simple_ownership_distinct_nth_update [OF j_bound ts_j])
        apply force
        done
    next      
      case Fence
      then obtain
        "is": "is = Fence # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      from simple_ownership_distinct [OF j_bound _ _ ts_j]
      show "simple_ownership_distinct ts'"
        apply (simp only: ts' eqs')
        apply (rule simple_ownership_distinct_nth_update [OF j_bound ts_j])
        apply force
        done
    next
      case (RMWReadOnly 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
        θ': "θ' = θ(t m a)" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮" and
        cond: "¬ cond (θ(t  m a))"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      from simple_ownership_distinct [OF j_bound _ _ ts_j]
      show "simple_ownership_distinct ts'"
        apply (simp only: ts' eqs')
        apply (rule simple_ownership_distinct_nth_update [OF j_bound ts_j])
        apply force
        done
    next
      case (RMWWrite 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
        θ': "θ' = θ(t  ret (m a) (f (θ(t  m a))))" and
        sb': "sb'=sb" and
        m': "m'=m(a := f (θ(t  m a)))" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL" and
        cond: "cond (θ(t  m a))"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      from safe_j [simplified "is"] cond
      obtain
        a_unowned_others: "k < length ts. jk  a  (map owned ts!k  dom (map released ts!k))" and
        A: "A  dom 𝒮  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length ts. jk   A   (map owned ts!k  dom (map released ts!k)) = {}" and
        a_not_ro: "a  read_only 𝒮"
        by cases auto

      from simple_ownership_distinct [OF j_bound _ _ ts_j] R_owns A_R A_unowned_others
      show "simple_ownership_distinct ts'"
        apply (simp only: ts' eqs')
        apply (rule simple_ownership_distinct_nth_update [OF j_bound ts_j])
        apply force
        done
    next
      case (Ghost A L R W) 
      then obtain
        "is": "is = Ghost A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=augment_rels (dom 𝒮) R " and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from safe_j [simplified "is"]
      obtain
        A: "A  dom 𝒮  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length ts. jk   A   (map owned ts!k  dom (map released ts!k)) = {}" 
        by cases auto

      from simple_ownership_distinct [OF j_bound _ _ ts_j] R_owns A_R A_unowned_others
      show "simple_ownership_distinct ts'"
        apply (simp only: ts' eqs')
        apply (rule simple_ownership_distinct_nth_update [OF j_bound ts_j])
        apply force
        done
    qed
  next
    case (StoreBuffer _ p "is" θ sb 𝒟 𝒪  sb' 𝒪' ℛ')
    hence False 
      by (auto simp add: empty_storebuffer_step_def)
    thus ?thesis ..
  qed
qed

theorem (in program) safe_step_preserves_read_only_unowned:
  assumes step: "(ts,m,𝒮) d (ts',m',𝒮')"
  assumes safe: "safe_delayed (ts,m,𝒮)"
  assumes dist: "simple_ownership_distinct ts"
  assumes ro_unowned: "read_only_unowned 𝒮 ts"
  shows "read_only_unowned 𝒮' ts'"
proof -
  interpret direct_computation:
    computation direct_memop_step empty_storebuffer_step program_step "λp p' is sb. sb" .
  from dist interpret simple_ownership_distinct ts .
  from ro_unowned interpret read_only_unowned 𝒮 ts .
  from step
  show ?thesis
  proof (cases)
    case (Program j p "is" θ sb 𝒟 𝒪  p' is')
    then obtain
      ts': "ts' = ts[j:=(p',is@is',θ,sb,𝒟,𝒪,)]" and
      𝒮': "𝒮'=𝒮" and
      m': "m'=m" and
      j_bound: "j < length ts" and
      ts_j: "ts!j = (p,is,θ,sb,𝒟,𝒪,)" and
      prog_step: "θ p p (p', is')"
      by auto

    from read_only_unowned [OF j_bound ts_j]
    show "read_only_unowned 𝒮' ts'"
      apply (simp only: ts' 𝒮')
      apply (rule read_only_unowned_nth_update [OF j_bound])
      apply force
      done
  next
    case (Memop j p "is" θ sb 𝒟 𝒪  is' θ' sb' 𝒟' 𝒪' ℛ')
    then obtain
      ts': "ts' = ts[j:=(p,is',θ',sb',𝒟',𝒪',ℛ')]" and
      j_bound: "j < length ts" and
      ts_j: "ts!j = (p,is,θ,sb,𝒟,𝒪,)"  and
      mem_step: "(is, θ, sb, m, 𝒟, 𝒪, , 𝒮)  (is', θ', sb',m', 𝒟',  𝒪', ℛ', 𝒮')"
      by auto

    from safe_delayedE [OF safe j_bound ts_j]
    have safe_j: "map owned ts,map released ts,j(is, θ, m, 𝒟, 𝒪, 𝒮)".
    from mem_step
    show ?thesis
    proof (cases)
      case (Read volatile a t)
      then obtain
        "is": "is = Read volatile a t # is'" and
        θ': "θ' = θ(t  m a)" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from read_only_unowned [OF j_bound ts_j]
      show "read_only_unowned 𝒮' ts'"
        apply (simp only: ts' eqs')
        apply (rule read_only_unowned_nth_update [OF j_bound])
        apply force
        done

    next
      case (WriteNonVolatile a D f A L R W)
      then obtain
        "is": "is = Write False a (D, f) A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m(a:=f θ)" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      from read_only_unowned [OF j_bound ts_j]
      show "read_only_unowned 𝒮' ts'"
        apply (simp only: ts' eqs')
        apply (rule read_only_unowned_nth_update [OF j_bound])
        apply force
        done

    next
      case (WriteVolatile a D f A L R W)
      then obtain
        "is": "is = Write True a (D, f) A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m(a:=f θ)" and
        𝒟': "𝒟'=True" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from safe_j [simplified "is"]
      obtain
        a_unowned_others: "k < length ts. jk  a  (map owned ts!k  dom (map released ts!k))" and
        A: "A  dom 𝒮  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length ts. jk   A   (map owned ts!k  dom (map released ts!k)) = {}" and
        a_not_ro: "a  read_only 𝒮"
        by cases auto

      show "read_only_unowned 𝒮' ts'"
      proof (unfold_locales)
        fix i pi "isi" 𝒪i i 𝒟i θi sbi
        assume i_bound: "i < length ts'"
        assume ts'_i: "ts'!i = (pi,isi,θi, sbi, 𝒟i, 𝒪i,i)"
        show "𝒪i  read_only 𝒮' = {}"
        proof (cases "i=j")
          case True
          with read_only_unowned [OF j_bound ts_j] ts'_i A L_A R_owns A_R j_bound 
          show ?thesis
            by (auto simp add: eqs' ts' read_only_def augment_shared_def restrict_shared_def domIff split: option.splits)
        next
          case False
          from simple_ownership_distinct [OF j_bound _ False [symmetric] ts_j] ts'_i i_bound j_bound False
          have "𝒪  𝒪i = {}"
            by (fastforce simp add: ts')
          with A L_A R_owns A_R j_bound A_unowned_others [rule_format, of i] 
          read_only_unowned [of i pi isi θi sbi 𝒟i 𝒪i i]
            False i_bound ts'_i False
          show ?thesis
            by (force simp add: eqs' ts' read_only_def augment_shared_def restrict_shared_def domIff split: option.splits)
        qed
      qed
    next      
      case Fence
      then obtain
        "is": "is = Fence # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      from read_only_unowned [OF j_bound ts_j]
      show "read_only_unowned 𝒮' ts'"
        apply (simp only: ts' eqs')
        apply (rule read_only_unowned_nth_update [OF j_bound])
        apply force
        done
    next
      case (RMWReadOnly 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
        θ': "θ' = θ(t m a)" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮" and
        cond: "¬ cond (θ(t  m a))"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      from read_only_unowned [OF j_bound ts_j]
      show "read_only_unowned 𝒮' ts'"
        apply (simp only: ts' eqs')
        apply (rule read_only_unowned_nth_update [OF j_bound])
        apply force
        done
    next
      case (RMWWrite 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
        θ': "θ' = θ(t  ret (m a) (f (θ(t  m a))))" and
        sb': "sb'=sb" and
        m': "m'=m(a := f (θ(t  m a)))" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL" and
        cond: "cond (θ(t  m a))"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      from safe_j [simplified "is"] cond
      obtain
        a_unowned_others: "k < length ts. jk  a  (map owned ts!k  dom (map released ts!k))" and
        A: "A  dom 𝒮  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length ts. jk   A   (map owned ts!k  dom (map released ts!k)) = {}" and
        a_not_ro: "a  read_only 𝒮"
        by cases auto

      show "read_only_unowned 𝒮' ts'"
      proof (unfold_locales)
        fix i pi "isi" 𝒪i i 𝒟i θi sbi
        assume i_bound: "i < length ts'"
        assume ts'_i: "ts'!i = (pi,isi,θi, sbi, 𝒟i, 𝒪i,i)"
        show "𝒪i  read_only 𝒮' = {}"
        proof (cases "i=j")
          case True
          with read_only_unowned [OF j_bound ts_j] ts'_i A L_A R_owns A_R j_bound 
          show ?thesis
            by (auto simp add: eqs' ts' read_only_def augment_shared_def restrict_shared_def domIff split: option.splits)
        next
          case False
          from simple_ownership_distinct [OF j_bound _ False [symmetric] ts_j] ts'_i i_bound j_bound False
          have "𝒪  𝒪i = {}"
            by (fastforce simp add: ts')
          with A L_A R_owns A_R j_bound A_unowned_others [rule_format, of i] 
          read_only_unowned [of i pi isi θi sbi 𝒟i 𝒪i i]
            False i_bound ts'_i False
          show ?thesis
            by (force simp add: eqs' ts' read_only_def augment_shared_def restrict_shared_def domIff split: option.splits)
        qed
      qed
    next
      case (Ghost A L R W) 
      then obtain
        "is": "is = Ghost A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=augment_rels (dom 𝒮) R " and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from safe_j [simplified "is"]
      obtain
        A: "A  dom 𝒮  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length ts. jk   A   (map owned ts!k  dom (map released ts!k)) = {}" 
        by cases auto

      show "read_only_unowned 𝒮' ts'"
      proof (unfold_locales)
        fix i pi "isi" 𝒪i i 𝒟i θi sbi
        assume i_bound: "i < length ts'"
        assume ts'_i: "ts'!i = (pi,isi,θi, sbi, 𝒟i, 𝒪i,i)"
        show "𝒪i  read_only 𝒮' = {}"
        proof (cases "i=j")
          case True
          with read_only_unowned [OF j_bound ts_j] ts'_i A L_A R_owns A_R j_bound 
          show ?thesis
            by (auto simp add: eqs' ts' read_only_def augment_shared_def restrict_shared_def domIff split: option.splits)
        next
          case False
          from simple_ownership_distinct [OF j_bound _ False [symmetric] ts_j] ts'_i i_bound j_bound False
          have "𝒪  𝒪i = {}"
            by (fastforce simp add: ts')
          with A L_A R_owns A_R j_bound A_unowned_others [rule_format, of i] 
          read_only_unowned [of i pi isi θi sbi 𝒟i 𝒪i i]
            False i_bound ts'_i False
          show ?thesis
            by (force simp add: eqs' ts' read_only_def augment_shared_def restrict_shared_def domIff split: option.splits)
        qed
      qed
    qed
  next
    case (StoreBuffer _ p "is" θ sb 𝒟 𝒪  sb' 𝒪' ℛ')
    hence False 
      by (auto simp add: empty_storebuffer_step_def)
    thus ?thesis ..
  qed
qed


theorem (in program) safe_step_preserves_unowned_shared:
  assumes step: "(ts,m,𝒮) d (ts',m',𝒮')"
  assumes safe: "safe_delayed (ts,m,𝒮)"
  assumes dist: "simple_ownership_distinct ts"
  assumes unowned_shared: "unowned_shared 𝒮 ts"
  shows "unowned_shared 𝒮' ts'"
proof -
  interpret direct_computation:
    computation direct_memop_step empty_storebuffer_step program_step "λp p' is sb. sb" .
  from dist interpret simple_ownership_distinct ts .
  from unowned_shared interpret unowned_shared 𝒮 ts .
  from step
  show ?thesis
  proof (cases)
    case (Program j p "is" θ sb 𝒟 𝒪  p' is')
    then obtain
      ts': "ts' = ts[j:=(p',is@is',θ,sb,𝒟,𝒪,)]" and
      𝒮': "𝒮'=𝒮" and
      m': "m'=m" and
      j_bound: "j < length ts" and
      ts_j: "ts!j = (p,is,θ,sb,𝒟,𝒪,)" and
      prog_step: "θ p p (p', is')"
      by auto

    show "unowned_shared 𝒮' ts'"
      apply (simp only: ts' 𝒮')
      apply (rule unowned_shared_nth_update [OF j_bound ts_j] )
      apply force
      done
  next
    case (Memop j p "is" θ sb 𝒟 𝒪  is' θ' sb' 𝒟' 𝒪' ℛ')
    then obtain
      ts': "ts' = ts[j:=(p,is',θ',sb',𝒟',𝒪',ℛ')]" and
      j_bound: "j < length ts" and
      ts_j: "ts!j = (p,is,θ,sb,𝒟,𝒪,)"  and
      mem_step: "(is, θ, sb, m, 𝒟, 𝒪, , 𝒮)  (is', θ', sb',m', 𝒟',  𝒪', ℛ', 𝒮')"
      by auto

    from safe_delayedE [OF safe j_bound ts_j]
    have safe_j: "map owned ts,map released ts,j(is, θ, m, 𝒟, 𝒪, 𝒮)".
    from mem_step
    show ?thesis
    proof (cases)
      case (Read volatile a t)
      then obtain
        "is": "is = Read volatile a t # is'" and
        θ': "θ' = θ(t  m a)" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      show "unowned_shared 𝒮' ts'"
        apply (simp only: ts' eqs')
        apply (rule unowned_shared_nth_update [OF j_bound ts_j])
        apply force
        done

    next
      case (WriteNonVolatile a D f A L R W)
      then obtain
        "is": "is = Write False a (D, f) A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m(a:=f θ)" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      
      show "unowned_shared 𝒮' ts'"
        apply (simp only: ts' eqs')
        apply (rule unowned_shared_nth_update [OF j_bound ts_j])
        apply force
        done

    next
      case (WriteVolatile a D f A L R W)
      then obtain
        "is": "is = Write True a (D, f) A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m(a:=f θ)" and
        𝒟': "𝒟'=True" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from safe_j [simplified "is"]
      obtain
        a_unowned_others: "k < length ts. jk  a  (map owned ts!k  dom (map released ts!k))" and
        A: "A  dom 𝒮  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length ts. jk   A   (map owned ts!k  dom (map released ts!k)) = {}" and
        a_not_ro: "a  read_only 𝒮"
        by cases auto


      show "unowned_shared 𝒮' ts'"
      apply (clarsimp simp add: unowned_shared_def')
      using A R_owns L_A A_R A_unowned_others ts_j j_bound
      apply (auto simp add: 𝒮' ts' 𝒪')
      apply  (rule  unowned_shared')
      apply  clarsimp
      apply  (drule_tac x=i in spec)
      apply  (case_tac "i=j")
      apply   clarsimp
      apply  clarsimp
      apply (drule_tac x=j in spec)
      apply auto
      done
    next      
      case Fence
      then obtain
        "is": "is = Fence # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      
      show "unowned_shared 𝒮' ts'"
        apply (simp only: ts' eqs')
        apply (rule unowned_shared_nth_update [OF j_bound ts_j])
        apply force
        done
    next
      case (RMWReadOnly 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
        θ': "θ' = θ(t m a)" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮" and
        cond: "¬ cond (θ(t  m a))"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      show "unowned_shared 𝒮' ts'"
        apply (simp only: ts' eqs')
        apply (rule unowned_shared_nth_update [OF j_bound ts_j])
        apply force
        done
    next
      case (RMWWrite 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
        θ': "θ' = θ(t  ret (m a) (f (θ(t  m a))))" and
        sb': "sb'=sb" and
        m': "m'=m(a := f (θ(t  m a)))" and
        𝒟': "𝒟'=False" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=Map.empty" and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL" and
        cond: "cond (θ(t  m a))"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'
      from safe_j [simplified "is"] cond
      obtain
        a_unowned_others: "k < length ts. jk  a  (map owned ts!k  dom (map released ts!k))" and
        A: "A  dom 𝒮  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length ts. jk   A   (map owned ts!k  dom (map released ts!k)) = {}" and
        a_not_ro: "a  read_only 𝒮"
        by cases auto

      show "unowned_shared 𝒮' ts'"
      apply (clarsimp simp add: unowned_shared_def')
      using A R_owns L_A A_R A_unowned_others ts_j j_bound
      apply (auto simp add: 𝒮' ts' 𝒪')
      apply  (rule  unowned_shared')
      apply  clarsimp
      apply  (drule_tac x=i in spec)
      apply  (case_tac "i=j")
      apply   clarsimp
      apply  clarsimp
      apply (drule_tac x=j in spec)
      apply auto
      done
    next
      case (Ghost A L R W) 
      then obtain
        "is": "is = Ghost A L R W # is'" and
        θ': "θ' = θ" and
        sb': "sb'=sb" and
        m': "m'=m" and
        𝒟': "𝒟'=𝒟" and
        𝒪': "𝒪'=𝒪  A - R" and
        ℛ': "ℛ'=augment_rels (dom 𝒮) R " and
        𝒮': "𝒮'=𝒮 ⊕⇘WR ⊖⇘AL"
        by auto
      note eqs' = θ' sb' m' 𝒟' 𝒪' ℛ' 𝒮'

      from safe_j [simplified "is"]
      obtain
        A: "A  dom 𝒮  𝒪" and L_A: "L  A" and R_owns: "R  𝒪" and A_R: "A  R = {}" and
        A_unowned_others: "k < length ts. jk   A   (map owned ts!k  dom (map released ts!k)) = {}" 
        by cases auto
      show "unowned_shared 𝒮' ts'"
      apply (clarsimp simp add: unowned_shared_def')
      using A R_owns L_A A_R A_unowned_others ts_j j_bound
      apply (auto simp add: 𝒮' ts' 𝒪')
      apply  (rule  unowned_shared')
      apply  clarsimp
      apply  (drule_tac x=i in spec)
      apply  (case_tac "i=j")
      apply   clarsimp
      apply  clarsimp
      apply (drule_tac x=j in spec)
      apply auto
      done
    qed
  next
    case (StoreBuffer _ p "is" θ sb 𝒟 𝒪  sb' 𝒪' ℛ')
    hence False 
      by (auto simp add: empty_storebuffer_step_def)
    thus ?thesis ..
  qed
qed

locale program_trace = program +
fixes c       ― ‹enumeration of configurations: @{text "c n  ⇒d c (n + 1) ... ⇒d c (n + k)"} 
fixes n::nat  ― ‹starting index›
fixes k::nat  ― ‹steps›

assumes step: "l. l < k  c (n+l) d c (n + (Suc l))"

abbreviation (in program)
"trace  program_trace program_step"

lemma (in program) trace_0 [simp]: "trace c n 0"
apply (unfold_locales)
apply auto
done

lemma split_less_Suc: "(x<Suc k. P x) =  (P k  (x<k. P x))"
  apply rule
  apply  clarsimp
  apply clarsimp
  apply (case_tac "x = k") 
  apply auto
  done
  
lemma split_le_Suc: "(xSuc k. P x) =  (P (Suc k)  (xk. P x))"
  apply rule
  apply  clarsimp
  apply clarsimp
  apply (case_tac "x = Suc k") 
  apply auto
  done

lemma (in program) steps_to_trace: 
assumes steps: "x d* y"
shows "c k. trace c 0 k  c 0 = x  c k = y" 
using steps
proof (induct)
  case base
  thus ?case
    apply (rule_tac x="λk. x" in exI)
    apply (rule_tac x=0 in exI)
    by (auto simp add: program_trace_def)
next
  case (step y z)
  have first: "x d* y" by fact
  have last: "y d z" by fact
  from step.hyps obtain c k where
    trace: "trace c 0 k" and c_0: "c 0 = x" and c_k: "c k = y"
    by auto
  define c' where "c' == λi. (if i  k then c i else z)"
  from trace last c_k have "trace c' 0 (k + 1)"
    apply (clarsimp simp add: c'_def program_trace_def)
    apply (subgoal_tac "l=k")
    apply  (simp)
    apply (simp)
    done
  with c_0 
  show ?case
    apply -
    apply (rule_tac x="c'" in exI)
    apply (rule_tac x="k + 1" in exI)
    apply (auto simp add: c'_def)
    done
qed


lemma (in program) trace_preserves_length_ts: 
  "l x. trace c n k  l  k  x  k  length (fst (c (n + l))) = length (fst (c (n + x)))"
proof (induct k)
  case 0
  thus ?case by auto
next
  case (Suc k)
  then obtain trace_suc: "trace c n (Suc k)" and
    l_suc: "l  Suc k" and
    x_suc: "x  Suc k"
    by simp
  interpret direct_computation:
    computation direct_memop_step empty_storebuffer_step program_step "λp p' is sb. sb" .

  from trace_suc obtain
    trace_k: "trace c n k" and
    last_step: "c (n + k) d c (n + (Suc k)) "
    by (clarsimp simp add: program_trace_def)
  obtain ts 𝒮 m where c_k: "c (n + k) = (ts, m, 𝒮)" by (cases "c (n + k)") 
  obtain ts' 𝒮' m' where c_suc_k: "c (n + (Suc k)) = (ts', m', 𝒮')" by (cases "c (n + (Suc k))") 
  from direct_computation.step_preserves_length_ts [OF last_step [simplified c_k c_suc_k]] c_k c_suc_k
  have leq: "length (fst (c (n + Suc k))) = length (fst (c (n + k)))"
    by simp

  show ?case
  proof (cases "l = Suc k") 
    case True
    note l_suc = this
    show ?thesis
    proof (cases "x = Suc k") 
      case True with l_suc show ?thesis by simp
    next
      case False
      with x_suc have "x  k" by simp
      from Suc.hyps [OF trace_k this, of k]
      have "length (fst (c (n + x))) = length (fst (c (n + k)))"
        by simp
      with leq show ?thesis using l_suc by simp
    qed
  next
    case False
    with l_suc have l_k: "l  k"
      by auto
    show ?thesis
    proof (cases "x = Suc k") 
      case True
      from Suc.hyps [OF trace_k l_k, of k]
      have "length (fst (c (n + l))) = length (fst (c (n + k)))" by simp
      with leq True show ?thesis by simp
    next
      case False
      with x_suc have "x  k" by simp
      from Suc.hyps [OF trace_k l_k this]
      show ?thesis by simp
    qed
  qed
qed

lemma (in program) trace_preserves_simple_ownership_distinct: 
  assumes dist: "simple_ownership_distinct (fst (c n))"
  shows "l. trace c n k  (x < k. safe_delayed (c (n + x)))   
           l  k    simple_ownership_distinct (fst (c (n + l)))"
proof (induct k)
  case 0 thus ?case using dist by auto
next
  case (Suc k)
  then obtain 
    trace_suc: "trace c n (Suc k)" and
    safe_suc: "x<Suc k. safe_delayed (c (n + x))" and
    l_suc: "l  Suc k"
    by simp

  from trace_suc obtain
    trace_k: "trace c n k" and
    last_step: "c (n + k) d c (n + (Suc k)) "
    by (clarsimp simp add: program_trace_def)

  obtain ts 𝒮 m where c_k: "c (n + k) = (ts, m, 𝒮)" by (cases "c (n + k)") 
  obtain ts' 𝒮' m' where c_suc_k: "c (n + (Suc k)) = (ts', m', 𝒮')" by (cases "c (n + (Suc k))") 

  from safe_suc c_suc_k c_k
  obtain 
    safe_up_k: "x<k. safe_delayed (c (n + x))" and
    safe_k: "safe_delayed (ts,m,𝒮)" 
    by (auto simp add: split_le_Suc) 
  from Suc.hyps [OF trace_k safe_up_k] 
  have hyp: "l  k. simple_ownership_distinct (fst (c (n + l)))"
    by simp

  from Suc.hyps [OF trace_k safe_up_k, of k] c_k
  have "simple_ownership_distinct ts"
    by simp

  from safe_step_preserves_simple_ownership_distinct [OF last_step[simplified c_k c_suc_k] safe_k this]
  have "simple_ownership_distinct ts'".
  then show ?case
  using c_suc_k hyp l_suc
    apply (cases "l=Suc k")
    apply (auto simp add: split_less_Suc)
    done
qed

lemma (in program) trace_preserves_read_only_unowned: 
  assumes dist: "simple_ownership_distinct (fst (c n))"
  assumes ro: "read_only_unowned (snd (snd (c n))) (fst (c n))"
  shows "l. trace c n k  (x < k. safe_delayed (c (n + x)))   
           l  k    read_only_unowned (snd (snd (c (n + l)))) (fst (c (n + l)))"
proof (induct k)
  case 0 thus ?case using ro by auto
next
  case (Suc k)
  then obtain 
    trace_suc: "trace c n (Suc k)" and
    safe_suc: "x<Suc k. safe_delayed (c (n + x))" and
    l_suc: "l  Suc k"
    by simp

  from trace_suc obtain
    trace_k: "trace c n k" and
    last_step: "c (n + k) d c (n + (Suc k)) "
    by (clarsimp simp add: program_trace_def)

  obtain ts 𝒮 m where c_k: "c (n + k) = (ts, m, 𝒮)" by (cases "c (n + k)") 
  obtain ts' 𝒮' m' where c_suc_k: "c (n + (Suc k)) = (ts', m', 𝒮')" by (cases "c (n + (Suc k))") 

  from safe_suc c_suc_k c_k
  obtain 
    safe_up_k: "x<k. safe_delayed (c (n + x))" and
    safe_k: "safe_delayed (ts,m,𝒮)" 
    by (auto simp add: split_le_Suc) 
  from Suc.hyps [OF trace_k safe_up_k] 
  have hyp: "l  k. read_only_unowned (snd (snd (c (n + l)))) (fst (c (n + l)))"
    by simp

  from Suc.hyps [OF trace_k safe_up_k, of k] c_k
  have ro': "read_only_unowned 𝒮 ts"
    by simp

  from trace_preserves_simple_ownership_distinct [where c=c and n=n, OF dist trace_k safe_up_k, of k] c_k
  have dist': "simple_ownership_distinct ts" by simp

  from safe_step_preserves_read_only_unowned [OF last_step[simplified c_k c_suc_k] safe_k dist' ro']
  have "read_only_unowned 𝒮' ts'".
  then show ?case
  using c_suc_k hyp l_suc
    apply (cases "l=Suc k")
    apply (auto simp add: split_less_Suc)
    done
qed

lemma (in program) trace_preserves_unowned_shared: 
  assumes dist: "simple_ownership_distinct (fst (c n))"
  assumes ro: "unowned_shared (snd (snd (c n))) (fst (c n))"
  shows "l. trace c n k  (x < k. safe_delayed (c (n + x)))   
           l  k    unowned_shared (snd (snd (c (n + l)))) (fst (c (n + l)))"
proof (induct k)
  case 0 thus ?case using ro by auto
next
  case (Suc k)
  then obtain 
    trace_suc: "trace c n (Suc k)" and
    safe_suc: "x<Suc k. safe_delayed (c (n + x))" and
    l_suc: "l  Suc k"
    by simp

  from trace_suc obtain
    trace_k: "trace c n k" and
    last_step: "c (n + k) d c (n + (Suc k)) "
    by (clarsimp simp add: program_trace_def)

  obtain ts 𝒮 m where c_k: "c (n + k) = (ts, m, 𝒮)" by (cases "c (n + k)") 
  obtain ts' 𝒮' m' where c_suc_k: "c (n + (Suc k)) = (ts', m', 𝒮')" by (cases "c (n + (Suc k))") 

  from safe_suc c_suc_k c_k
  obtain 
    safe_up_k: "x<k. safe_delayed (c (n + x))" and
    safe_k: "safe_delayed (ts,m,𝒮)" 
    by (auto simp add: split_le_Suc) 
  from Suc.hyps [OF trace_k safe_up_k] 
  have hyp: "l  k. unowned_shared (snd (snd (c (n + l)))) (fst (c (n + l)))"
    by simp

  from Suc.hyps [OF trace_k safe_up_k, of k] c_k
  have ro': "unowned_shared 𝒮 ts"
    by simp

  from trace_preserves_simple_ownership_distinct [where c=c and n=n, OF dist trace_k safe_up_k, of k] c_k
  have dist': "simple_ownership_distinct ts" by simp

  from safe_step_preserves_unowned_shared [OF last_step[simplified c_k c_suc_k] safe_k dist' ro']
  have "unowned_shared 𝒮' ts'".
  then show ?case
  using c_suc_k hyp l_suc
    apply (cases "l=Suc k")
    apply (auto simp add: split_less_Suc)
    done
qed


theorem (in program_progress) undo_local_steps:
assumes steps: "trace c n k"
assumes c_n: "c n = (ts,m,𝒮)"
assumes unchanged: "l  k. (tsl 𝒮l ml . c (n + l) = (tsl,ml,𝒮l)  tsl!i=ts!i)"
assumes safe: "safe_delayed (u_ts, u_m, u_shared)"
assumes leq: "length u_ts = length ts"
assumes i_bound: "i < length ts"
assumes others_same: "j < length ts. ji  u_ts!j = ts!j"
assumes u_ts_i: "u_ts!i=(u_p,u_is,u_tmps,u_sb,u_dirty,u_owns,u_rels)"
assumes u_m_other: "a. a  u_owns  u_m a = m a"
assumes u_m_shared: "a. a  u_owns  a  dom u_shared  u_m a = m a"
assumes u_shared: "a. a  u_owns  a  owned (ts!i)  u_shared a = 𝒮 a"
assumes dist: "simple_ownership_distinct u_ts"
assumes dist_ts: "simple_ownership_distinct ts"
assumes safe_orig: "x.  x < k  safe_delayed (c (n + x))" 
shows "c' l. l  k  trace c' n l 

        c' n = (u_ts, u_m, u_shared) 
        (x  l. length (fst (c' (n + x))) = length (fst (c (n + x)))) 

        (x < l. safe_delayed (c' (n + x))) 
        (l < k  ¬ safe_delayed (c' (n + l)))  
       
        (x  l. tsx 𝒮x mx tsx' 𝒮x' mx' . c (n + x) = (tsx,mx,𝒮x)  c' (n+ x) = (tsx',mx',𝒮x')  
          tsx'!i=u_ts!i 
          (a  u_owns. 𝒮x' a = u_shared a) 
          (a  u_owns. 𝒮x a = 𝒮 a) 
          (a  u_owns. mx' a = u_m a) 
          (a  u_owns. mx a = m a)) 
        
        (x  l. tsx 𝒮x mx tsx' 𝒮x' mx'. c (n + x) = (tsx,mx,𝒮x)  c' (n + x) = (tsx',mx',𝒮x')  
         (j < length tsx. ji  tsx'!j = tsx!j) 
         (a. a  u_owns  a  owned (ts!i)  𝒮x' a = 𝒮x a) 
         (a. a  u_owns  mx' a = mx a))
"
using steps unchanged safe_orig
proof (induct k)
  case 0
  show ?case
    apply (rule_tac x="λ l. (u_ts, u_m, u_shared)" in exI)
    apply (rule_tac x=0 in exI)
    thm c_n
    apply (simp add: c_n)
    apply (clarsimp simp add: 0 leq others_same u_m_other u_shared)
    done
next
  case (Suc k)
  then obtain  
    trace_suc: "trace c n (Suc k)" and
    unchanged_suc: "lSuc k. tsl 𝒮l ml. c (n + l) = (tsl, ml, 𝒮l)  tsl ! i = ts ! i" and
    safe_orig: "x<k. safe_delayed (c (n + x))"
    by simp

  interpret direct_computation:
    computation direct_memop_step empty_storebuffer_step program_step "λp p' is sb. sb" .
  from trace_suc obtain
    trace_k: "trace c n k" and
    last_step: "c (n + k) d c (n + (Suc k)) "
    by (clarsimp simp add: program_trace_def)

  from unchanged_suc obtain
    unchanged_k: "lk. tsl 𝒮l ml. c (n + l) = (tsl, ml, 𝒮l)  tsl ! i = ts ! i" and
    unchanged_suc_k: "tsl 𝒮l ml. c (n + (Suc k)) = (tsl, ml, 𝒮l)  tsl ! i = ts ! i"
    apply -
    apply (rule that)
    apply auto
    apply (drule_tac x=l in spec) (* Why the hack doesn't Isabelle try this *)
    apply simp
    done
  

  from Suc.hyps [OF trace_k unchanged_k safe_orig] obtain c' l where
    l_k: "l  k" and
    trace_c'_l: "trace c' n l" and
    safe_l: "(x<l. safe_delayed (c' (n + x)))" and
    unsafe_l: "(l < k  ¬ safe_delayed (c' (n + l)))" and
    c'_n: "c' n = (u_ts, u_m, u_shared)" and
    leq_l: "(xl. length (fst (c' (n + x))) = length (fst (c (n + x))))" and
    unchanged_i: "(xl. tsx 𝒮x mx tsx' 𝒮x' mx'.
                   c (n + x) = (tsx, mx, 𝒮x) 
                   c' (n + x) = (tsx', mx', 𝒮x') 
                   tsx' ! i = u_ts ! i 
                   (au_owns. 𝒮x' a = u_shared a) 
                   (au_owns. 𝒮x a = 𝒮 a) 
                   (au_owns. mx' a = u_m a) 
                   (au_owns. mx a = m a))" and
     sim: "(xl. tsx 𝒮x mx tsx' 𝒮x' mx'.
                   c (n + x) = (tsx, mx, 𝒮x) 
                   c' (n + x) = (tsx', mx', 𝒮x') 
                   (j<length tsx. j  i  tsx' ! j = tsx ! j) 
                   (a. a  u_owns  a  owned (ts!i)  𝒮x' a = 𝒮x a) 
                   (a. a  u_owns  mx' a = mx a))"
      by auto
    show ?case
    proof (cases "l < k")
      case True
      with True trace_c'_l safe_l unsafe_l unchanged_i sim leq_l c'_n
      show ?thesis
        apply -
        apply (rule_tac x="c'" in exI)
        apply (rule_tac x="l" in exI)
        apply auto
        done
    next
      case False
      with l_k have l_k: "l=k" by auto
      show ?thesis
      proof (cases "safe_delayed (c' (n + k))")
        case False
        with False l_k trace_c'_l safe_l unsafe_l unchanged_i sim leq_l c'_n
        show ?thesis
          apply -
          apply (rule_tac x="c'" in exI)
          apply (rule_tac x="k" in exI)
          apply auto
          done
      next
        case True
        note safe_k = this

        obtain tsk 𝒮k mk where c_k: "c (n + k) = (tsk,mk,𝒮k)"
          by (cases "c (n + k)")

        obtain tsk' 𝒮k' mk' where c_suc_k: "c (n + (Suc k)) = (tsk',mk',𝒮k')"
          by (cases "c (n + (Suc k))")

        obtain u_tsk u_sharedk u_mk where c'_k: "c' (n + k) = (u_tsk, u_mk, u_sharedk)"
          by (cases "c' (n + k)")

        from trace_preserves_length_ts [OF trace_k, of k 0] c_n c_k i_bound
        have i_bound_k: "i < length tsk"
          by simp

        from leq_l [rule_format, simplified l_k, of k] c_k c'_k
        have leq: "length u_tsk = length tsk"
          by simp
        
        note last_step = last_step [simplified c_k c_suc_k]
        from unchanged_suc_k c_suc_k
        have "tsk'!i = ts!i"
          by auto
        moreover from unchanged_k [rule_format, of k] c_k
        have unch_k_i: "tsk!i=ts!i"
          by auto
        ultimately have ts_eq: "tsk!i=tsk'!i"
          by simp

        from unchanged_i [simplified l_k, rule_format, OF _ c_k c'_k]
        obtain 
          u_ts_eq: "u_tsk ! i = u_ts ! i" and
          unchanged_shared: "au_owns. u_sharedk a = u_shared a"  and
          unchanged_shared_orig: "au_owns. 𝒮k a = 𝒮 a" and
          unchanged_owns: "au_owns. u_mk a = u_m a" and
          unchanged_owns_orig: "au_owns. mk a = m a"
          by fastforce

        from u_ts_eq u_ts_i
        have u_tsk_i: "u_tsk!i=(u_p,u_is,u_tmps,u_sb,u_dirty,u_owns,u_rels)"
          by auto
        from sim [simplified l_k, rule_format, of k, OF _ c_k c'_k]  
        obtain
          ts_sim: "(j<length tsk. j  i  u_tsk ! j = tsk ! j)" and
          shared_sim: "(a. a  u_owns  a  owned (tsk!i)  u_sharedk a = 𝒮k a)" and
          mem_sim: "(a. a  u_owns  u_mk a = mk a)"
          by (auto simp add: unch_k_i)

        

        from unchanged_owns_orig unchanged_owns u_m_shared unchanged_shared
        have unchanged_owns_shared: "a. a  u_owns  a  dom u_sharedk  u_mk a = mk a"
          by (auto simp add: simp add: domIff)

        from safe_l l_k safe_k
        have safe_up_k: "x<k. safe_delayed (c' (n + x))"
          apply clarsimp
          done
        from trace_preserves_simple_ownership_distinct [OF _ trace_c'_l [simplified l_k] safe_up_k, 
          simplified c'_n, simplified, OF dist, of k] c'_k
        have dist': "simple_ownership_distinct u_tsk"
          by simp
        

        from trace_preserves_simple_ownership_distinct [OF _ trace_k, simplified c_n, simplified, OF dist_ts safe_orig, of k] 
        c_k
        have dist_orig': "simple_ownership_distinct tsk"
          by simp

        from undo_local_step [OF last_step i_bound_k ts_eq safe_k [simplified c'_k] leq ts_sim u_tsk_i mem_sim 
          unchanged_owns_shared shared_sim dist' dist_orig']
        obtain u_ts' u_shared' u_m' where
           step': "(u_tsk, u_mk, u_sharedk) d (u_ts', u_m', u_shared')" and
           ts_eq': "u_ts' ! i = u_tsk ! i" and
           unchanged_shared': "(au_owns. u_shared' a = u_sharedk a)" and
           unchanged_shared_orig': "(au_owns. 𝒮k' a = 𝒮k a)" and
           unchanged_owns': "(au_owns. u_m' a = u_mk a)" and
           unchanged_owns_orig': "(au_owns. mk' a = mk a)" and
           sim_ts': "(j<length tsk. j  i  u_ts' ! j = tsk' ! j)" and
           sim_shared': "(a. a  u_owns  a  owned (tsk ! i)  u_shared' a = 𝒮k' a)" and
           sim_m': "(a. a  u_owns  u_m' a = mk' a)"
          by auto

        define c'' where "c'' == λl. if l  n + k then c' l else (u_ts', u_m', u_shared')"
        have [simp]: "x  n + k. c'' x = c' x"
          by (auto simp add: c''_def)
        have [simp]: "c'' (Suc (n + k)) = (u_ts', u_m', u_shared')"
          by (auto simp add: c''_def)

        from trace_c'_l l_k step' c'_k  have trace': "trace c'' n (Suc k)"
         apply (simp add: program_trace_def)
         apply (clarsimp simp add: split_less_Suc)
         done

        from direct_computation.step_preserves_length_ts [OF last_step]
        have leq_tsk': "length tsk' = length tsk".

        with direct_computation.step_preserves_length_ts [OF step'] leq
        have leq': "length u_ts' = length tsk"
          by simp
        show ?thesis
          apply (rule_tac x=c'' in exI)
          apply (rule_tac x="Suc k" in exI)
          using safe_l l_k unchanged_i sim c_suc_k leq_l c'_n leq'
          apply (clarsimp simp add: split_less_Suc split_le_Suc safe_k trace' leq_tsk' sim_ts' sim_shared' sim_m' unch_k_i
            
            ts_eq' u_ts_eq
            unchanged_shared' unchanged_shared unchanged_shared_orig unchanged_shared_orig'
            unchanged_owns' unchanged_owns
            unchanged_owns_orig' unchanged_owns_orig )
          done
    qed (* FIXME: indentation *)
  qed         
qed


locale program_safe_reach_upto = program +
  fixes n fixes safe fixes c0
  assumes safe_config: "k  n; trace c 0 k; c 0 = c0; l  k   safe (c l)"

abbreviation (in program)
  "safe_reach_upto  program_safe_reach_upto program_step"

lemma (in program) safe_reach_upto_le:
  assumes safe: "safe_reach_upto n safe c0"
  assumes m_n: "m  n"
  shows "safe_reach_upto m safe c0"
using safe m_n
apply (clarsimp simp add: program_safe_reach_upto_def)
  subgoal for k c
    apply (subgoal_tac "k  n")
     apply blast
    apply simp
  done
done


lemma (in program) last_action_of_thread:
assumes trace: "trace c 0 k"
shows 
  "― ‹thread i never executes›  
  (l  k. fst (c l)!i = fst (c k)!i)  
  ― ‹thread i has a last step in the trace›
  (last <  k.  
    fst (c last)!i  fst (c (Suc last))!i 
    (l. last < l  l  k  fst (c l)!i = fst (c k)!i)) "
using trace
proof (induct k)    
  case 0 thus ?case
    by auto
next
  case (Suc k)
  hence "trace c 0 (Suc k)" by simp
  then
  obtain
    trace_k: "trace c 0 k" and
    last_step: "c k d c (Suc k) "
    by (clarsimp simp add: program_trace_def)  
  
  show ?case
  proof (cases "fst (c k)!i = fst (c (Suc k))!i")
    case False
    then show ?thesis
      apply -
      apply (rule disjI2)
      apply (rule_tac x=k in exI)
      apply clarsimp
      apply (subgoal_tac "l=Suc k")
      apply auto
      done
  next
    case True
    note idle_i = this


    {
      assume same: "(lk. fst (c l) ! i = fst (c k) ! i)"
      have ?thesis
        apply -
        apply (rule disjI1)
        apply clarsimp
        apply (case_tac "l=Suc k")
        apply  (simp add: idle_i)
        apply (rule same [simplified idle_i, rule_format])
        apply simp
        done
    }
    moreover
    {
      fix last
      assume last_k: "last < k"
      assume last_step: "fst (c last) ! i  fst (c (Suc last)) ! i"
      assume idle: "(l>last. l  k  fst (c l) ! i = fst (c k) ! i)"
      have ?thesis
        apply -
        apply (rule disjI2)
        apply (rule_tac x=last in exI)
        using last_k 
        apply (simp add:  last_step)
        using idle [simplified idle_i] 
        apply clarsimp
        apply (case_tac "l=Suc k")
        apply  clarsimp
        apply clarsimp
        done
    }
    moreover note Suc.hyps [OF trace_k]
    ultimately
    show ?thesis
      by blast
  qed
qed

lemma (in program) sequence_traces:
assumes trace1: "trace c1 0 k"
assumes trace2: "trace c2 m l"
assumes seq: "c2 m = c1 k"
assumes c_def: "c = (λx. if x   k then c1 x else (c2 (m + x -k)))"
shows "trace c 0 (k + l)"
proof -
  from trace1
  interpret trace1: program_trace program_step c1 0 k .
  from trace2
  interpret trace2: program_trace program_step c2 m l .
  {
    fix x
    assume x_bound: "x < (k + l)" 
    have "c x d c (Suc x)"
    proof (cases "x < k")
      case True
      from trace1.step [OF True] True
      show ?thesis
        by (simp add: c_def)
    next
      case False
      hence k_x: "k  x"
        by auto
      with x_bound have bound: "x - k < l"
        by auto
      from k_x have eq: "(Suc (m + x) - k) = Suc (m + x - k)"
        by simp
      from trace2.step [OF bound] k_x seq
      show ?thesis
        by (auto simp add: c_def eq)
    qed
  }
  thus ?thesis
    by (auto simp add: program_trace_def)
qed
      
theorem (in program_progress) safe_free_flowing_implies_safe_delayed:
  assumes init: "initial c0"
  assumes dist: "simple_ownership_distinct (fst c0)"
  assumes read_only_unowned: "read_only_unowned (snd (snd c0)) (fst c0)"
  assumes unowned_shared: "unowned_shared (snd (snd c0)) (fst c0)"
  assumes safe_reach_ff: "safe_reach_upto n safe_free_flowing c0"
  shows "safe_reach_upto n safe_delayed c0"
using safe_reach_ff 
proof (induct n)
  case 0
  hence "safe_reach_upto 0 safe_free_flowing c0" by simp
  hence "safe_free_flowing c0"
    by (auto simp add: program_safe_reach_upto_def)
  from initial_safe_free_flowing_implies_safe_delayed [OF init this]
  have "safe_delayed c0".
  then show ?case
    by (simp add: program_safe_reach_upto_def)
next
  case (Suc n)
  hence safe_reach_suc: "safe_reach_upto (Suc n) safe_free_flowing c0" by simp
  then interpret safe_reach_suc_inter: program_safe_reach_upto program_step "(Suc n)" safe_free_flowing c0 .
  from safe_reach_upto_le [OF safe_reach_suc ]  
  have safe_reach_n: "safe_reach_upto n safe_free_flowing c0" by simp
  from Suc.hyps [OF this]
  have safe_delayed_reach_n: "safe_reach_upto n safe_delayed c0".
  then interpret safe_delayed_reach_inter: program_safe_reach_upto program_step "n" safe_delayed c0 .
  interpret direct_computation:
    computation direct_memop_step empty_storebuffer_step program_step "λp p' is sb. sb" .
  show ?case
  proof (cases "safe_reach_upto (Suc n) safe_delayed c0")
    case True thus ?thesis .
  next
    case False
    from safe_delayed_reach_n False
    obtain c where
      trace: "trace c 0 (Suc n)" and
      c_0: "c 0 = c0" and
      safe_delayed_upto_n: "kn. safe_delayed (c k)" and
      violation_delayed_suc: "¬ safe_delayed (c (Suc n))"
    proof -
      from False
      obtain c k l where 
        k_suc: "k  Suc n" and
        trace_k: "trace c 0 k" and 
        l_k: "l  k" and
        violation: "¬ safe_delayed (c l)" and
        start: "c 0 = c0"
        by (clarsimp simp add: program_safe_reach_upto_def)
    
      show ?thesis
      proof (cases "k = Suc n")
        case False
        with k_suc have "k  n"
          by auto
        from safe_delayed_reach_inter.safe_config [where c=c, OF this trace_k start l_k]
        have "safe_delayed (c l)".
        with violation have False by simp
        thus ?thesis ..
      next
        case True
        note k_suc_n = this
        from trace_k True have trace_n: "trace c 0 n"
          by (auto simp add: program_trace_def)
        show ?thesis
        proof (cases "l=Suc n") 
          case False
          with k_suc_n l_k have "l  n" by simp
          from safe_delayed_reach_inter.safe_config [where c=c, OF _ trace_n start this ]
          have "safe_delayed (c l)" by simp
          with violation have False by simp
          thus ?thesis ..
        next
          case True
          from safe_delayed_reach_inter.safe_config [where c=c, OF _ trace_n start]
          have "kn. safe_delayed (c k)" by simp
          with True k_suc_n trace_k start violation
          show ?thesis
            apply -
            apply (rule that)
            apply auto
            done
        qed    
      qed
    qed

    from trace
    interpret trace_inter: program_trace program_step c 0 "Suc n" .

    from safe_reach_suc_inter.safe_config [where c=c, OF _ trace c_0]
    have safe_suc: "safe_free_flowing (c (Suc n))"
      by auto
    
    obtain ts 𝒮 m where c_suc: "c (Suc n) = (ts,m,𝒮)" by (cases "c (Suc n)")
    from violation_delayed_suc c_suc
    obtain i p "is" θ sb 𝒟 𝒪  where
      i_bound: "i < length ts" and
      ts_i: "ts ! i = (p,is,θ,sb,𝒟,𝒪,)" and
      violation_i: "¬ map owned ts,map released ts,i (is,θ,m,𝒟,𝒪,𝒮) "
      by (fastforce simp add: safe_free_flowing_def safe_delayed_def)

    from trace_preserves_unowned_shared [where c=c and n=0 and l="Suc n", 
          simplified c_0, OF dist unowned_shared trace] safe_delayed_upto_n  c_suc
    have "unowned_shared 𝒮 ts" by auto
    then interpret unowned_shared 𝒮 ts .

    
    from violation_i obtain ins is' where "is": "is = ins#is'"
      by (cases "is") (auto simp add: safe_delayed_direct_memop_state.Nil)
    from safeE [OF safe_suc [simplified c_suc] i_bound ts_i]
    have safe_i: "map owned ts,i(is, θ, m, 𝒟, 𝒪, 𝒮)".

    define races where "races == λ. (case ins of
         Read volatile a t  ( a = Some False)  (¬ volatile  a  dom )
       | Write volatile a sop A L R W  (a  dom   (volatile  A  dom   {}))
       | Ghost A L R W  (A  dom   {})
       | RMW a t (D,f) cond ret A L R W  (if cond (θ(t  m a)) 
                                           then a  dom   A  dom   {}
                                           else  a = Some False)
       | _  False)"



    {
      assume no_race: 
        " j. j < length ts  ji  ¬ races (released (ts!j))"
      from safe_i 
      have "map owned ts,map released ts,i (is,θ,m,𝒟,𝒪,𝒮) "
      proof cases
        case Read
        thus ?thesis
          using "is" no_race
          by (auto simp add: races_def intro: safe_delayed_direct_memop_state.intros)
      next
        case WriteNonVolatile
        thus ?thesis
          using "is" no_race
          by (auto simp add: races_def intro: safe_delayed_direct_memop_state.intros)
      next
        case WriteVolatile
        thus ?thesis
          using "is" no_race
          apply (clarsimp simp add: races_def) 
          apply (rule safe_delayed_direct_memop_state.intros)
          apply auto
          done
      next
        case Fence
        thus ?thesis
          using "is" no_race
          by (auto simp add: races_def intro: safe_delayed_direct_memop_state.intros)
      next
        case Ghost
        thus ?thesis
          using "is" no_race
          apply (clarsimp simp add: races_def) 
          apply (rule safe_delayed_direct_memop_state.intros)
          apply auto
          done
      next
        case RMWReadOnly
        thus ?thesis
          using "is" no_race
          by (auto simp add: races_def intro: safe_delayed_direct_memop_state.intros)
      next
        case (RMWWrite cond t a _ _ A _ 𝒪)
        thus ?thesis
          using "is" no_race unowned_shared' [rule_format, of a] ts_i
          apply (clarsimp simp add: races_def) 
          apply (rule safe_delayed_direct_memop_state.RMWWrite)
          apply auto
          apply force
          done
      next
        case Nil with "is" show ?thesis by auto
      qed
    }
    with violation_i
    obtain j where
      j_bound: "j < length ts" and
      neq_j_i: "j  i" and
      race: "races (released (ts!j))"
      by auto
    
    obtain pj "isj" θj sbj 𝒟j 𝒪j j where
      ts_j: "ts!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
      apply (cases "ts!j")
      apply force
      done

    from race 
    have j_non_empty: "j  Map.empty"
      by (auto simp add: ts_j races_def split: instr.splits if_split_asm)

    {
      assume idle_j: "lSuc n. fst (c l) ! j = fst (c (Suc n)) ! j"
      have ?thesis
      proof -
        from idle_j [rule_format, of 0] c_suc  c_0 ts_j
        have c0_j: "fst c0 ! j = ts!j"
          by clarsimp
        from trace_preserves_length_ts [OF trace, of 0 "Suc n"] c_0 c_suc
        have "length (fst c0) = length ts"
          by clarsimp
        with j_bound have "j < length (fst c0)"
          by simp
        with nth_mem [OF this] init c0_j ts_j 
        have "j = Map.empty"
          by (auto simp add: initial_def)
        with j_non_empty have False
          by simp
        thus ?thesis ..
      qed
    }
    moreover
    {
      fix last
      assume last_bound: "last<Suc n" 
      assume last_step_changed_j: "fst (c last) ! j  fst (c (Suc last)) ! j"
      assume idle_rest: "l>last. l  Suc n  fst (c l) ! j = fst (c (Suc n)) ! j"
      have ?thesis
      proof -
        obtain tsl 𝒮l ml where
           c_last: "c last = (tsl,ml,𝒮l)"
          by (cases "c last")
        obtain tsl' 𝒮l' ml' where
           c_last': "c (Suc last) = (tsl',ml',𝒮l')"
          by (cases "c (Suc last)")
        from idle_rest [rule_format, of "Suc last" ] c_suc c_last' last_bound
        have tsl'_j: "tsl'!j = ts!j"
          by auto

        from last_step_changed_j c_last c_last'
        have j_changed: "tsl!j  tsl'!j"
           by auto

        from trace_inter.step [OF last_bound] c_last c_last'
        have last_step: "(tsl,ml,𝒮l) d (tsl',ml',𝒮l')"
          by simp

        obtain pl "isl" θl sbl 𝒟l 𝒪l l where
          tsl_j: "tsl!j = (pl,isl,θl,sbl,𝒟l,𝒪l,l)"
          apply (cases "tsl!j")
          apply force
          done
(*
        obtain pl' "isl'" θl' sbl' 𝒟l' 𝒪l' ℛl' where
          tsl'_j: "tsl'!j = (pl',isl',θl',sbl',𝒟l',𝒪l',ℛl')"
          apply (cases "tsl'!j")
          apply force
          done
*)      
        from trace_preserves_length_ts [OF trace, of last "Suc n"] c_last c_suc last_bound
        have leql: "length tsl = length ts"
          by simp
        with j_bound have j_boundl: "j < length tsl"
          by simp

        from trace have trace_n: "trace c 0 n"
          by (auto simp add: program_trace_def)
        
        from safe_delayed_reach_inter.safe_config [where k=n and c=c and l=last, OF _ trace_n c_0] last_bound c_last
        have safe_delayed_last: "safe_delayed (tsl,ml,𝒮l)"
          by auto
        
        from safe_delayed_reach_inter.safe_config [where c=c, OF _ trace_n c_0] 
        have safe_delayed_upto_n: "x<n. safe_delayed (c (0 + x))"
          by auto
        from trace_preserves_simple_ownership_distinct [where c=c and n=0 and l=last, 
          simplified c_0, OF dist trace_n safe_delayed_upto_n]
          last_bound c_last
        have dist_last: "simple_ownership_distinct tsl" 
          by auto

        from trace_preserves_read_only_unowned [where c=c and n=0 and l=last, 
          simplified c_0, OF dist read_only_unowned trace_n safe_delayed_upto_n]
          last_bound c_last
        have ro_last_last: "read_only_unowned 𝒮l tsl" 
          by auto


        
        from safe_delayed_reach_inter.safe_config [where c=c, OF _ trace_n c_0] 
        have safe_delayed_upto_suc_n: "x<Suc n. safe_delayed (c (0 + x))"
          by auto

        from trace_preserves_simple_ownership_distinct [where c=c and n=0 and l="Suc last", 
          simplified c_0, OF dist trace safe_delayed_upto_suc_n]
          last_bound c_last'
        have dist_last': "simple_ownership_distinct tsl'" 
          by auto
        from trace last_bound have trace_last: "trace c 0 last"
          by (auto simp add: program_trace_def)

        from trace last_bound have trace_rest: "trace c (Suc last) (n - last)"
          by (auto simp add: program_trace_def)

        from idle_rest last_bound
        have idle_rest':
            "ln - last.
                      tsl 𝒮l ml. c (Suc last + l) = (tsl, ml, 𝒮l)  tsl ! j = tsl' ! j"
          apply clarsimp
          apply (drule_tac x="Suc (last + l)" in spec)
          apply (auto simp add: c_last' c_suc tsl'_j)
          done

        from safe_delayed_upto_suc_n [rule_format, of last] last_bound 
        have safe_delayed_last: "safe_delayed (tsl, ml, 𝒮l)"
          by (auto simp add: c_last)
        from safe_delayedE [OF this j_boundl tsl_j] 
        have safel: "map owned tsl,map released tsl,j(isl, θl, ml, 𝒟l, 𝒪l, 𝒮l)".
        
        from safe_delayed_reach_inter.safe_config [where c=c, OF _ trace_n c_0] 
        have safe_delayed_upto_last: "x<n - last. safe_delayed (c (Suc (last + x)))"
          by auto

        from last_step
        show ?thesis
        proof (cases)
          case (Program i' _ _ _ _ _ _ _ p' is')
          with j_changed j_boundl tsl_j 
          obtain
            tsl': "tsl' = tsl[j:=(p',isl@is',θl,sbl,𝒟l,𝒪l,l)]" and
            𝒮l': "𝒮l'=𝒮l" and
            ml': "ml'=ml" and
            prog_step: "θl pl p (p', is')"
            by (cases "i'=j") auto
          from tsl'_j tsl' ts_j j_boundl
          obtain eqs: "p'=pj" "isl@is'=isj" "θl=θj" "𝒟l=𝒟j" "𝒪l=𝒪j" "l=j"
            by auto


          from undo_local_steps [where c=c, OF trace_rest c_last' idle_rest' safe_delayed_last, simplified tsl', 
            simplified,
            OF j_boundl tsl_j [simplified], simplified ml' 𝒮l', simplified, OF dist_last  
            dist_last' [simplified tsl',simplified] safe_delayed_upto_last]
          obtain c' k where
            k_bound: "k  n - last" and
            trace_c': "trace c' (Suc last) k" and
            c'_first: "c' (Suc last) = (tsl, ml, 𝒮l)" and
            c'_leq: "(xk. length (fst (c' (Suc (last + x)))) = length (fst (c (Suc (last + x)))))" and
            c'_safe: "(x<k. safe_delayed (c' (Suc (last + x))))" and
            c'_unsafe: "(k < n - last  ¬ safe_delayed (c' (Suc (last + k))))" and
            c'_unch: 
               "(xk. tsx 𝒮x mx.
                   c (Suc (last + x)) = (tsx, mx, 𝒮x) 
                   (tsx' 𝒮x' mx'.
                       c' (Suc (last + x)) = (tsx', mx', 𝒮x') 
                       tsx' ! j = tsl ! j 
                       (a𝒪l. 𝒮x' a = 𝒮l a) 
                       (a𝒪l. 𝒮x a = 𝒮l a) 
                       (a𝒪l. mx' a = ml a)  (a𝒪l. mx a = ml a)))" and
            c'_sim:
               "(xk. tsx 𝒮x mx.
                   c (Suc (last + x)) = (tsx, mx, 𝒮x) 
                   (tsx' 𝒮x' mx'.
                       c' (Suc (last + x)) = (tsx', mx', 𝒮x') 
                       (ja<length tsx. ja  j  tsx' ! ja = tsx ! ja) 
                       (a. a  𝒪l  𝒮x' a = 𝒮x a) 
                       (a. a  𝒪l  mx' a = mx a)))"
            by auto

          obtain c_undo where c_undo:  "c_undo = (λx. if x  last then c x else c' (Suc last + x - last))"
            by blast
          have c_undo_0: "c_undo 0 = c0"
            by (auto simp add: c_undo c_0)
          from sequence_traces [OF trace_last trace_c', simplified c_last, OF c'_first c_undo]
          have trace_undo: "trace c_undo 0 (last + k)" .
          obtain u_ts u_shared u_m where 
            c_undo_n: "c_undo n = (u_ts,u_m, u_shared)"
            by (cases "c_undo n") 
          with last_bound c'_first c_last
          have c'_suc: "c' (Suc n) = (u_ts,u_m, u_shared)"
            apply (auto simp add: c_undo split: if_split_asm)
            apply (subgoal_tac "n=last")
            apply auto
            done

          
          show ?thesis  
          proof (cases "k < n - last")
            case True
            with c'_unsafe have unsafe: "¬ safe_delayed (c_undo (last + k))"
              by (auto simp add: c_undo c_last c'_first)
            from True have "last + k  n"
              by auto
            from safe_delayed_reach_inter.safe_config [OF this trace_undo, of "last + k"]
            have "safe_delayed (c_undo (last + k))"
              by (auto simp add: c_undo c_0)
            with unsafe have False by simp
            thus ?thesis ..
          next
            case False
            with k_bound have k: "k = n - last"
              by auto
            have eq': "Suc (last + (n - last)) = Suc n"
              using last_bound
              by simp
            from c'_unch [rule_format, of k, simplified k eq', OF _ c_suc c'_suc]
            obtain u_ts_j: "u_ts!j = tsl!j" and
              shared_unch: "a𝒪l. u_shared a = 𝒮l a" and
              shared_orig_unch: "a𝒪l. 𝒮 a = 𝒮l a" and
              mem_unch: "a𝒪l. u_m a = ml a" and 
              mem_unch_orig: "a𝒪l. m a = ml a"
              by auto

            from c'_sim [rule_format, of k, simplified k eq', OF _ c_suc c'_suc] i_bound neq_j_i
            obtain u_ts_i: "u_ts!i = ts!i" and
               shared_sim: "a. a  𝒪l  u_shared a = 𝒮 a" and
               mem_sim: "a. a  𝒪l  u_m a = m a"
              by auto
          
            from c'_leq [rule_format, of k] c'_suc c_suc
            have leq_u_ts: "length u_ts = length ts"
              by (auto simp add: eq' k)

            from j_bound leq_u_ts
            have j_bound_u: "j < length u_ts"
              by simp
            from i_bound leq_u_ts
            have i_bound_u: "i < length u_ts"
              by simp
            from k last_bound have l_k_eq: "last + k = n"
              by auto
            from safe_delayed_reach_inter.safe_config [OF _ trace_undo, simplified l_k_eq] 
              k c_0 last_bound
            have safe_delayed_c_undo': "x n. safe_delayed (c_undo x)"
              by (auto simp add: c_undo split: if_split_asm)
            hence safe_delayed_c_undo: "x<n. safe_delayed (c_undo x)"
              by (auto)
            from trace_preserves_simple_ownership_distinct [OF _ trace_undo, 
              simplified l_k_eq c_undo_0, simplified, OF dist this, of n] dist c_undo_n
            have dist_u_ts: "simple_ownership_distinct u_ts"
              by auto
            then interpret dist_u_ts_inter: simple_ownership_distinct u_ts .
          
            {
              fix a
              have "u_m a = m a"
              proof (cases "a  𝒪l") 
                case True with mem_unch
                have "u_m a = ml a"
                  by auto
                moreover
                from True mem_unch_orig
                have "m a = ml a" 
                  by auto
                ultimately show ?thesis by simp
              next
                case False
                with mem_sim
                show ?thesis
                  by auto
              qed
            } hence u_m_eq: "u_m = m" by - (rule ext, auto) 

            {
              fix a
              have "u_shared a = 𝒮 a"
              proof (cases "a  𝒪l") 
                case True with shared_unch
                have "u_shared a = 𝒮l a"
                  by auto
                moreover
                from True shared_orig_unch
                have "𝒮 a = 𝒮l a" 
                  by auto
                ultimately show ?thesis by simp
              next
                case False
                with shared_sim
                show ?thesis
                  by auto
              qed
            } hence u_shared_eq: "u_shared = 𝒮" by - (rule ext, auto) 

            {
              assume safe: "map owned u_ts,map released u_ts,i (is,θ,u_m,𝒟,𝒪,u_shared) "
              then have False
              proof cases
                case Read
                then show ?thesis
                using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                  by (auto simp add:eqs races_def split: if_split_asm)
              next
                case WriteNonVolatile
                then show ?thesis
                using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                  by (auto simp add:eqs races_def split: if_split_asm)
              next
                case WriteVolatile
                then show ?thesis
                using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                  apply (auto simp add:eqs races_def split: if_split_asm)
                  apply fastforce
                  done
              next
                case Fence
                then show ?thesis
                using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                  by (auto simp add:eqs races_def split: if_split_asm)
              next
                case Ghost
                then show ?thesis
                using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                  apply (auto simp add:eqs races_def split: if_split_asm)
                  apply fastforce
                  done
              next
                case (RMWReadOnly cond t a D f ret A L R W)
                then show ?thesis
                using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                  by (auto simp add:eqs races_def u_shared_eq u_m_eq split: if_split_asm)
              next
                case RMWWrite
                then show ?thesis
                using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                  apply (auto simp add:eqs races_def u_shared_eq u_m_eq split: if_split_asm)
                  apply fastforce+
                  done
              next
                case Nil
                then show ?thesis
                using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                  by (auto simp add:eqs races_def split: if_split_asm)
              qed
            }
            hence "¬ safe_delayed (u_ts, u_m, u_shared)"
              apply (clarsimp simp add: safe_delayed_def)
              apply (rule_tac x=i in exI)
              using u_ts_i ts_i i_bound_u
              apply auto
              done
            moreover
            from safe_delayed_c_undo' [rule_format, of n] c_undo_n
            have "safe_delayed (u_ts, u_m, u_shared)"
              by simp
            ultimately have False
              by simp
            thus ?thesis 
              by simp
          qed
        next
          case (Memop i' _ _ _ _ _ _ _ "isl'" θl' sbl' 𝒟l' 𝒪l' l')
          with j_changed j_boundl tsl_j 
          obtain             
            tsl': "tsl' = tsl[j:=(pl,isl',θl',sbl',𝒟l',𝒪l',l')]" and
            mem_step: "(isl, θl, sbl, ml, 𝒟l, 𝒪l, l,𝒮l)  
              (isl', θl', sbl', ml', 𝒟l', 𝒪l', l', 𝒮l')"
              by (cases "i'=j") auto

          from mem_step  
          show ?thesis
          proof (cases)
            case (Read volatile a t)
            then obtain
              "isl": "isl = Read volatile a t # isl'" and
              θl': "θl' = θl(t  ml a)"  and
              sbl': "sbl'=sbl" and
              𝒟l': "𝒟l'=𝒟l" and
              𝒪l': "𝒪l' = 𝒪l" and
              l': "l'= l" and
              𝒮l': "𝒮l'=𝒮l" and
              ml': "ml' = ml"
              by auto
            note eqs' = θl' sbl' 𝒟l' 𝒪l' l' 𝒮l' ml'
            from tsl'_j tsl' ts_j j_boundl eqs'
            obtain eqs: "pl=pj" "isl'=isj" "θl(t  ml a)=θj" "𝒟l=𝒟j" "𝒪l=𝒪j" "l=j"
              by auto

            (* FIXME: now the same proof as for Program step *)
            from undo_local_steps [where c=c, OF trace_rest c_last' idle_rest' safe_delayed_last, simplified tsl', 
            simplified,
            OF j_boundl tsl_j [simplified], simplified ml' 𝒮l', simplified, OF dist_last 
              dist_last' [simplified tsl',simplified] safe_delayed_upto_last]
            obtain c' k where
              k_bound: "k  n - last" and
              trace_c': "trace c' (Suc last) k" and
              c'_first: "c' (Suc last) = (tsl, ml, 𝒮l)" and
              c'_leq: "(xk. length (fst (c' (Suc (last + x)))) = length (fst (c (Suc (last + x)))))" and
              c'_safe: "(x<k. safe_delayed (c' (Suc (last + x))))" and
              c'_unsafe: "(k < n - last  ¬ safe_delayed (c' (Suc (last + k))))" and
              c'_unch: 
               "(xk. tsx 𝒮x mx.
                   c (Suc (last + x)) = (tsx, mx, 𝒮x) 
                   (tsx' 𝒮x' mx'.
                       c' (Suc (last + x)) = (tsx', mx', 𝒮x') 
                       tsx' ! j = tsl ! j 
                       (a𝒪l. 𝒮x' a = 𝒮l a) 
                       (a𝒪l. 𝒮x a = 𝒮l a) 
                       (a𝒪l. mx' a = ml a)  (a𝒪l. mx a = ml a)))" and
              c'_sim:
               "(xk. tsx 𝒮x mx.
                   c (Suc (last + x)) = (tsx, mx, 𝒮x) 
                   (tsx' 𝒮x' mx'.
                       c' (Suc (last + x)) = (tsx', mx', 𝒮x') 
                       (ja<length tsx. ja  j  tsx' ! ja = tsx ! ja) 
                       (a. a  𝒪l  𝒮x' a = 𝒮x a) 
                       (a. a  𝒪l  mx' a = mx a)))"
              by (clarsimp simp add: 𝒪l')
            obtain c_undo where c_undo:  "c_undo = (λx. if x  last then c x else c' (Suc last + x - last))"
              by blast
            have c_undo_0: "c_undo 0 = c0"
              by (auto simp add: c_undo c_0)
            from sequence_traces [OF trace_last trace_c', simplified c_last, OF c'_first c_undo]
            have trace_undo: "trace c_undo 0 (last + k)" .
            obtain u_ts u_shared u_m where 
              c_undo_n: "c_undo n = (u_ts,u_m, u_shared)"
              by (cases "c_undo n") 
            with last_bound c'_first c_last
            have c'_suc: "c' (Suc n) = (u_ts,u_m, u_shared)"
              apply (auto simp add: c_undo split: if_split_asm)
              apply (subgoal_tac "n=last")
              apply auto
              done

          
            show ?thesis  
            proof (cases "k < n - last")
              case True
              with c'_unsafe have unsafe: "¬ safe_delayed (c_undo (last + k))"
                by (auto simp add: c_undo c_last c'_first)
              from True have "last + k  n"
                by auto
              from safe_delayed_reach_inter.safe_config [OF this trace_undo, of "last + k"]
              have "safe_delayed (c_undo (last + k))"
                by (auto simp add: c_undo c_0)
              with unsafe have False by simp
              thus ?thesis ..
            next
              case False
              with k_bound have k: "k = n - last"
                by auto
              have eq': "Suc (last + (n - last)) = Suc n"
                using last_bound
                by simp
              from c'_unch [rule_format, of k, simplified k eq', OF _ c_suc c'_suc]
              obtain u_ts_j: "u_ts!j = tsl!j" and
                shared_unch: "a𝒪l. u_shared a = 𝒮l a" and
                shared_orig_unch: "a𝒪l. 𝒮 a = 𝒮l a" and
                mem_unch: "a𝒪l. u_m a = ml a" and 
                mem_unch_orig: "a𝒪l. m a = ml a"
                by auto

              from c'_sim [rule_format, of k, simplified k eq', OF _ c_suc c'_suc] i_bound neq_j_i
              obtain u_ts_i: "u_ts!i = ts!i" and
                 shared_sim: "a. a  𝒪l  u_shared a = 𝒮 a" and
                 mem_sim: "a. a  𝒪l  u_m a = m a"
                by auto
          
              from c'_leq [rule_format, of k] c'_suc c_suc
              have leq_u_ts: "length u_ts = length ts"
                by (auto simp add: eq' k)

              from j_bound leq_u_ts
              have j_bound_u: "j < length u_ts"
                by simp
              from i_bound leq_u_ts
              have i_bound_u: "i < length u_ts"
                by simp
              from k last_bound have l_k_eq: "last + k = n"
                by auto
              from safe_delayed_reach_inter.safe_config [OF _ trace_undo, simplified l_k_eq] 
                k c_0 last_bound
              have safe_delayed_c_undo': "xn. safe_delayed (c_undo x)"
                by (auto simp add: c_undo split: if_split_asm)
              hence safe_delayed_c_undo: "x<n. safe_delayed (c_undo x)"
                by (auto)
              from trace_preserves_simple_ownership_distinct [OF _ trace_undo, 
                simplified l_k_eq c_undo_0, simplified, OF dist this, of n] dist c_undo_n
              have dist_u_ts: "simple_ownership_distinct u_ts"
                by auto
              then interpret dist_u_ts_inter: simple_ownership_distinct u_ts .

              
              {
                fix a
                have "u_m a = m a"
                proof (cases "a  𝒪l") 
                  case True with mem_unch
                  have "u_m a = ml a"
                    by auto
                  moreover
                  from True mem_unch_orig
                  have "m a = ml a" 
                    by auto
                  ultimately show ?thesis by simp
                next
                  case False
                  with mem_sim
                  show ?thesis
                    by auto
                qed
              } hence u_m_eq: "u_m = m" by - (rule ext, auto) 

              {
                fix a
                have "u_shared a = 𝒮 a"
                proof (cases "a  𝒪l") 
                  case True with shared_unch
                  have "u_shared a = 𝒮l a"
                    by auto
                  moreover
                  from True shared_orig_unch
                  have "𝒮 a = 𝒮l a" 
                    by auto
                  ultimately show ?thesis by simp
                next
                  case False
                  with shared_sim
                  show ?thesis
                    by auto
                qed
              } hence u_shared_eq: "u_shared = 𝒮" by - (rule ext, auto) 

              {
                assume safe: "map owned u_ts,map released u_ts,i (is,θ,u_m,𝒟,𝒪,u_shared) "
                then have False
                proof cases
                  case Read
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def split: if_split_asm)
                next
                  case WriteNonVolatile
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def split: if_split_asm)
                next
                  case WriteVolatile
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    apply (auto simp add:eqs races_def split: if_split_asm)
                    apply fastforce
                    done
                next
                  case Fence
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def split: if_split_asm)
                next
                  case Ghost
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    apply (auto simp add:eqs races_def split: if_split_asm)
                    apply fastforce
                    done
                next
                  case (RMWReadOnly cond t a D f ret A L R W)
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def u_shared_eq u_m_eq split: if_split_asm)
                next
                  case RMWWrite
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    apply (auto simp add:eqs races_def u_shared_eq u_m_eq split: if_split_asm)
                    apply fastforce+
                    done
                next
                  case Nil
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def split: if_split_asm)
                qed
              }
              hence "¬ safe_delayed (u_ts, u_m, u_shared)"
                apply (clarsimp simp add: safe_delayed_def)
                apply (rule_tac x=i in exI)
                using u_ts_i ts_i i_bound_u
                apply auto
                done
              moreover
              from safe_delayed_c_undo' [rule_format, of n] c_undo_n
              have "safe_delayed (u_ts, u_m, u_shared)"
                by simp
              ultimately have False
                by simp
              thus ?thesis 
                by simp
            qed
          next
            case (WriteNonVolatile  a D f A L R W)
            then obtain
              "isl": "isl = Write False a (D, f) A L R W # isl'" and
              θl': "θl' = θl"  and
              sbl': "sbl'=sbl" and
              𝒟l': "𝒟l'=𝒟l" and
              𝒪l': "𝒪l' = 𝒪l" and
              l': "l'= l" and
              𝒮l': "𝒮l'=𝒮l" and
              ml': "ml' = ml(a:=f θl)"
              by auto
            note eqs' = θl' sbl' 𝒟l' 𝒪l' l' 𝒮l' ml'
            from tsl'_j tsl' ts_j j_boundl eqs'
            obtain eqs: "pl=pj" "isl'=isj" "θl=θj" "𝒟l=𝒟j" "𝒪l=𝒪j" 
              "l=j"
              by auto

            from safel [simplified "isl"]
            obtain a_owned: "a  𝒪l" and a_unshared: "a  dom 𝒮l"
              by cases auto
            have ml_unch_unowned: "a'. a'  𝒪l  ml a' = (ml(a := f θl)) a'" 
            using a_owned by auto

            have ml_unch_unshared: "a'. a'  𝒪l  a'  dom 𝒮l  ml a' = (ml(a := f θl)) a'" 
            using a_unshared by auto
            
            from undo_local_steps [where c=c, OF trace_rest c_last' idle_rest' safe_delayed_last, simplified tsl', 
              simplified,
              OF j_boundl tsl_j [simplified], simplified ml' 𝒮l',OF  ml_unch_unowned ml_unch_unshared, simplified, 
              OF dist_last dist_last' [simplified tsl',simplified] safe_delayed_upto_last]

            obtain c' k where
              k_bound: "k  n - last" and
              trace_c': "trace c' (Suc last) k" and
              c'_first: "c' (Suc last) = (tsl, ml, 𝒮l)" and
              c'_leq: "(xk. length (fst (c' (Suc (last + x)))) = length (fst (c (Suc (last + x)))))" and
              c'_safe: "(x<k. safe_delayed (c' (Suc (last + x))))" and
              c'_unsafe: "(k < n - last  ¬ safe_delayed (c' (Suc (last + k))))" and
              c'_unch: 
               "(xk. tsx 𝒮x mx.
                   c (Suc (last + x)) = (tsx, mx, 𝒮x) 
                   (tsx' 𝒮x' mx'.
                       c' (Suc (last + x)) = (tsx', mx', 𝒮x') 
                       tsx' ! j = tsl ! j 
                       (a𝒪l. 𝒮x' a = 𝒮l a) 
                       (a𝒪l. 𝒮x a = 𝒮l a) 
                       (a𝒪l. mx' a = ml a)  (a'𝒪l. mx a' = (ml(a := f θl)) a')))" and
              c'_sim:
               "(xk. tsx 𝒮x mx.
                   c (Suc (last + x)) = (tsx, mx, 𝒮x) 
                   (tsx' 𝒮x' mx'.
                       c' (Suc (last + x)) = (tsx', mx', 𝒮x') 
                       (ja<length tsx. ja  j  tsx' ! ja = tsx ! ja) 
                       (a. a  𝒪l  𝒮x' a = 𝒮x a) 
                       (a. a  𝒪l  mx' a = mx a)))"
              by (clarsimp simp add: 𝒪l')

            obtain c_undo where c_undo:  "c_undo = (λx. if x  last then c x else c' (Suc last + x - last))"
              by blast
            have c_undo_0: "c_undo 0 = c0"
              by (auto simp add: c_undo c_0)
            from sequence_traces [OF trace_last trace_c', simplified c_last, OF c'_first c_undo]
            have trace_undo: "trace c_undo 0 (last + k)" .
            obtain u_ts u_shared u_m where 
              c_undo_n: "c_undo n = (u_ts,u_m, u_shared)"
              by (cases "c_undo n") 
            with last_bound c'_first c_last
            have c'_suc: "c' (Suc n) = (u_ts,u_m, u_shared)"
              apply (auto simp add: c_undo split: if_split_asm)
              apply (subgoal_tac "n=last")
              apply auto
              done

          
            show ?thesis  
            proof (cases "k < n - last")
              case True
              with c'_unsafe have unsafe: "¬ safe_delayed (c_undo (last + k))"
                by (auto simp add: c_undo c_last c'_first)
              from True have "last + k  n"
                by auto
              from safe_delayed_reach_inter.safe_config [OF this trace_undo, of "last + k"]
              have "safe_delayed (c_undo (last + k))"
                by (auto simp add: c_undo c_0)
              with unsafe have False by simp
              thus ?thesis ..
            next
              case False
              with k_bound have k: "k = n - last"
                by auto
              have eq': "Suc (last + (n - last)) = Suc n"
                using last_bound
                by simp
              from c'_unch [rule_format, of k, simplified k eq', OF _ c_suc c'_suc]
              obtain u_ts_j: "u_ts!j = tsl!j" and
                shared_unch: "a𝒪l. u_shared a = 𝒮l a" and
                shared_orig_unch: "a𝒪l. 𝒮 a = 𝒮l a" and
                mem_unch: "a𝒪l. u_m a = ml a" and 
                mem_unch_orig: "a'𝒪l. m a' = (ml(a := f θl)) a'"
                by auto

              from c'_sim [rule_format, of k, simplified k eq', OF _ c_suc c'_suc] i_bound neq_j_i
              obtain u_ts_i: "u_ts!i = ts!i" and
                 shared_sim: "a. a  𝒪l  u_shared a = 𝒮 a" and
                 mem_sim: "a. a  𝒪l  u_m a = m a"
                by auto
          
              from c'_leq [rule_format, of k] c'_suc c_suc
              have leq_u_ts: "length u_ts = length ts"
                by (auto simp add: eq' k)

              from j_bound leq_u_ts
              have j_bound_u: "j < length u_ts"
                by simp
              from i_bound leq_u_ts
              have i_bound_u: "i < length u_ts"
                by simp
              from k last_bound have l_k_eq: "last + k = n"
                by auto
              from safe_delayed_reach_inter.safe_config [OF _ trace_undo, simplified l_k_eq] 
                k c_0 last_bound
              have safe_delayed_c_undo': "xn. safe_delayed (c_undo x)"
                by (auto simp add: c_undo split: if_split_asm)              
              hence safe_delayed_c_undo: "x<n. safe_delayed (c_undo x)"
                by auto
              from trace_preserves_simple_ownership_distinct [OF _ trace_undo, 
                simplified l_k_eq c_undo_0, simplified, OF dist this, of n] dist c_undo_n
              have dist_u_ts: "simple_ownership_distinct u_ts"
                by auto
              then interpret dist_u_ts_inter: simple_ownership_distinct u_ts .
          (* FIXME delete
              {
                fix a
                have "u_m a = m a"
                proof (cases "a ∈ 𝒪l") 
                  case True with mem_unch
                  have "u_m a = ml a"
                    by auto
                  moreover
                  from True mem_unch_orig
                  have "m a = ml a" 
                    by auto
                  ultimately show ?thesis by simp
                next
                  case False
                  with mem_sim
                  show ?thesis
                    by auto
                qed
              } hence u_m_eq: "u_m = m" by - (rule ext, auto) 
*)
              {
                fix a
                have "u_shared a = 𝒮 a"
                proof (cases "a  𝒪l") 
                  case True with shared_unch
                  have "u_shared a = 𝒮l a"
                    by auto
                  moreover
                  from True shared_orig_unch
                  have "𝒮 a = 𝒮l a" 
                    by auto
                  ultimately show ?thesis by simp
                next
                  case False
                  with shared_sim
                  show ?thesis
                    by auto
                qed
              } hence u_shared_eq: "u_shared = 𝒮" by - (rule ext, auto) 

              {
                assume safe: "map owned u_ts,map released u_ts,i (is,θ,u_m,𝒟,𝒪,u_shared) "
                then have False
                proof cases
                  case Read
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def split: if_split_asm)
                next
                  case WriteNonVolatile
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def split: if_split_asm)
                next
                  case WriteVolatile
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    apply (auto simp add:eqs races_def split: if_split_asm)
                    apply fastforce
                    done
                next
                  case Fence
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def split: if_split_asm)
                next
                  case Ghost
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    apply (auto simp add:eqs races_def split: if_split_asm)
                    apply fastforce
                    done
                next
                  case (RMWReadOnly cond t a' D f ret A L R W)
                  with ts_i "is" obtain
                    ins: "ins = RMW a' t (D, f) cond ret A L R W" and
                    owned_or_shared: "a'  𝒪  a'  dom u_shared" and
                    cond: "¬ cond (θ(t  u_m a'))" and
                    rels_race: "j<length (map owned u_ts). i  j  ((map released u_ts) ! j) a'  Some False"
                    by auto
                  from dist_u_ts_inter.simple_ownership_distinct [OF j_bound_u i_bound_u neq_j_i u_ts_j [simplified tsl_j] 
                    u_ts_i [simplified ts_i]] 
                  have dist: "𝒪l  𝒪 = {}"
                    by auto
                  from owned_or_shared dist a_owned a_unshared shared_orig_unch
                  have a'_a: "a'a"
                    by (auto simp add: u_shared_eq domIff)
                  have u_m_eq: "u_m a' = m a'"
                  proof (cases "a'  𝒪l") 
                    case True with mem_unch
                    have "u_m a' = ml a'"
                      by auto
                    moreover
                    from True mem_unch_orig a'_a
                    have "m a' = ml a'" 
                      by auto
                    ultimately show ?thesis by simp
                  next
                    case False
                    with mem_sim
                    show ?thesis
                      by auto
                  qed
                  with ins cond rels_race show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def u_shared_eq u_m_eq split: if_split_asm)
                next
                  case (RMWWrite cond t a' A L R D f ret W)
                  with ts_i "is" obtain
                    ins: "ins = RMW a' t (D, f) cond ret A L R W" and
                    cond: "cond (θ(t  u_m a'))" and
                    a': "j<length (map owned u_ts). i  j  a'  (map owned u_ts) ! j  dom ((map released u_ts) ! j)" and
                    safety:
                      "A  dom u_shared  𝒪" "L  A" "R  𝒪" "A  R = {}"
                      "j<length (map owned u_ts). i  j  A  ((map owned u_ts) ! j  dom ((map released u_ts) ! j)) = {}"
                      "a'  read_only u_shared"
                    by auto
                  from a'[rule_format, of j] j_bound_u u_ts_j tsl_j neq_j_i
                  have "a'  𝒪l"
                    by auto
                  from mem_sim [rule_format, OF this]
                  have u_m_eq: "u_m a' = m a'"
                    by auto
                  
                  with ins cond safety a' show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    apply (auto simp add:eqs races_def u_shared_eq u_m_eq split: if_split_asm)
                    apply fastforce
                    done
                next
                  case Nil
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def split: if_split_asm)
                qed
              }
              hence "¬ safe_delayed (u_ts, u_m, u_shared)"
                apply (clarsimp simp add: safe_delayed_def)
                apply (rule_tac x=i in exI)
                using u_ts_i ts_i i_bound_u
                apply auto
                done
              moreover
              from safe_delayed_c_undo' [rule_format, of n] c_undo_n
              have "safe_delayed (u_ts, u_m, u_shared)"
                by simp
              ultimately have False
                by simp
              thus ?thesis 
                by simp
            qed
          next
            case WriteVolatile
            with tsl'_j tsl' ts_j j_boundl have "j = Map.empty"
              by auto
            with j_non_empty have False by auto
            thus ?thesis ..
          next
            case Fence
            with tsl'_j tsl' ts_j j_boundl have "j = Map.empty"
              by auto
            with j_non_empty have False by auto
            thus ?thesis ..
          next
            case RMWReadOnly
            with tsl'_j tsl' ts_j j_boundl have "j = Map.empty"
              by auto
            with j_non_empty have False by auto
            thus ?thesis ..
          next
            case RMWWrite
            with tsl'_j tsl' ts_j j_boundl have "j = Map.empty"
              by auto
            with j_non_empty have False by auto
            thus ?thesis ..
          next
            case (Ghost A L R W)
            then obtain
              "isl": "isl = Ghost A L R W # isl'" and
              θl': "θl' = θl"  and
              sbl': "sbl'=sbl" and
              𝒟l': "𝒟l'=𝒟l" and
              𝒪l': "𝒪l' = 𝒪l  A - R" and
              l': "l'= augment_rels (dom 𝒮l) R l" and
              𝒮l': "𝒮l'=𝒮l ⊕⇘WR ⊖⇘AL" and
              ml': "ml' = ml"
              by auto
            note eqs' = θl' sbl' 𝒟l' 𝒪l' l' 𝒮l' ml'
            from tsl'_j tsl' ts_j j_boundl eqs'
            obtain eqs: "pl=pj" "isl'=isj" "θl=θj" "𝒟l=𝒟j" "𝒪l  A - R = 𝒪j" 
              "augment_rels (dom 𝒮l) R l=j"
              by auto

            from safel [simplified "isl"]
            obtain
            A_shared_owned: "A  dom 𝒮l  𝒪l" and L_A: "L  A" and R_owns: "R  𝒪l" and A_R: "A  R = {}" and
            "j' < length (map owned tsl). jj'   A  ((map owned tsl)!j'  dom ((map released tsl)!j')) = {}"
              by cases auto


            from A_shared_owned L_A R_owns A_R
            have shared_eq: "a. a  𝒪l  a  𝒪l'  𝒮l a = (𝒮l ⊕⇘WR ⊖⇘AL) a"
              by (auto simp add: restrict_shared_def augment_shared_def 𝒪l' split: option.splits)
            
            from undo_local_steps [where c=c, OF trace_rest c_last' idle_rest' safe_delayed_last, simplified tsl', 
              simplified,
              OF j_boundl tsl_j [simplified], simplified ml' 𝒮l', simplified, 
              OF  shared_eq dist_last dist_last' [simplified tsl',simplified] safe_delayed_upto_last]

            obtain c' k where
              k_bound: "k  n - last" and
              trace_c': "trace c' (Suc last) k" and
              c'_first: "c' (Suc last) = (tsl, ml, 𝒮l)" and
              c'_leq: "(xk. length (fst (c' (Suc (last + x)))) = length (fst (c (Suc (last + x)))))" and
              c'_safe: "(x<k. safe_delayed (c' (Suc (last + x))))" and
              c'_unsafe: "(k < n - last  ¬ safe_delayed (c' (Suc (last + k))))" and
              c'_unch: 
               "(xk. tsx 𝒮x mx.
                   c (Suc (last + x)) = (tsx, mx, 𝒮x) 
                   (tsx' 𝒮x' mx'.
                       c' (Suc (last + x)) = (tsx', mx', 𝒮x') 
                       tsx' ! j = tsl ! j 
                       (a𝒪l. 𝒮x' a = 𝒮l a) 
                       (a𝒪l. 𝒮x a = (𝒮l ⊕⇘WR ⊖⇘AL) a) 
                       (a𝒪l. mx' a = ml a)  (a'𝒪l. mx a' = (ml) a')))" and
              c'_sim:
               "(xk. tsx 𝒮x mx.
                   c (Suc (last + x)) = (tsx, mx, 𝒮x) 
                   (tsx' 𝒮x' mx'.
                       c' (Suc (last + x)) = (tsx', mx', 𝒮x') 
                       (ja<length tsx. ja  j  tsx' ! ja = tsx ! ja) 
                       (a. a  𝒪l  a  𝒪l'  𝒮x' a = 𝒮x a) 
                       (a. a  𝒪l  mx' a = mx a)))"

              by (clarsimp )
            obtain c_undo where c_undo:  "c_undo = (λx. if x  last then c x else c' (Suc last + x - last))"
              by blast
            have c_undo_0: "c_undo 0 = c0"
              by (auto simp add: c_undo c_0)
            from sequence_traces [OF trace_last trace_c', simplified c_last, OF c'_first c_undo]
            have trace_undo: "trace c_undo 0 (last + k)" .
            obtain u_ts u_shared u_m where 
              c_undo_n: "c_undo n = (u_ts,u_m, u_shared)"
              by (cases "c_undo n") 
            with last_bound c'_first c_last
            have c'_suc: "c' (Suc n) = (u_ts,u_m, u_shared)"
              apply (auto simp add: c_undo split: if_split_asm)
              apply (subgoal_tac "n=last")
              apply auto
              done

          
            show ?thesis  
            proof (cases "k < n - last")
              case True
              with c'_unsafe have unsafe: "¬ safe_delayed (c_undo (last + k))"
                by (auto simp add: c_undo c_last c'_first)
              from True have "last + k  n"
                by auto
              from safe_delayed_reach_inter.safe_config [OF this trace_undo, of "last + k"]
              have "safe_delayed (c_undo (last + k))"
                by (auto simp add: c_undo c_0)
              with unsafe have False by simp
              thus ?thesis ..
            next
              case False
              with k_bound have k: "k = n - last"
                by auto
              have eq': "Suc (last + (n - last)) = Suc n"
                using last_bound
                by simp
              from c'_unch [rule_format, of k, simplified k eq', OF _ c_suc c'_suc]
              obtain u_ts_j: "u_ts!j = tsl!j" and
                shared_unch: "a𝒪l. u_shared a = 𝒮l a" and
                shared_orig_unch: "a𝒪l. 𝒮 a = (𝒮l ⊕⇘WR ⊖⇘AL) a" and
                mem_unch: "a𝒪l. u_m a = ml a" and 
                mem_unch_orig: "a'𝒪l. m a' = ml a'"
                by auto

              from c'_sim [rule_format, of k, simplified k eq', OF _ c_suc c'_suc] i_bound neq_j_i
              obtain u_ts_i: "u_ts!i = ts!i" and
                 shared_sim: "a. a  𝒪l  a  𝒪l'  u_shared a = 𝒮 a" and
                 mem_sim: "a. a  𝒪l  u_m a = m a"
                by auto
          
              from c'_leq [rule_format, of k] c'_suc c_suc
              have leq_u_ts: "length u_ts = length ts"
                by (auto simp add: eq' k)

              from j_bound leq_u_ts
              have j_bound_u: "j < length u_ts"
                by simp
              from i_bound leq_u_ts
              have i_bound_u: "i < length u_ts"
                by simp
              from k last_bound have l_k_eq: "last + k = n"
                by auto
              from safe_delayed_reach_inter.safe_config [OF _ trace_undo, simplified l_k_eq] 
                k c_0 last_bound
              have safe_delayed_c_undo': "xn. safe_delayed (c_undo x)"
                by (auto simp add: c_undo split: if_split_asm)              
              hence safe_delayed_c_undo: "x<n. safe_delayed (c_undo x)"
                by auto
              from trace_preserves_simple_ownership_distinct [OF _ trace_undo, 
                simplified l_k_eq c_undo_0, simplified, OF dist this, of n] dist c_undo_n
              have dist_u_ts: "simple_ownership_distinct u_ts"
                by auto
              then interpret dist_u_ts_inter: simple_ownership_distinct u_ts .
              {
                fix a
                have "u_m a = m a"
                proof (cases "a  𝒪l") 
                  case True with mem_unch
                  have "u_m a = ml a"
                    by auto
                  moreover
                  from True mem_unch_orig
                  have "m a = ml a" 
                    by auto
                  ultimately show ?thesis by simp
                next
                  case False
                  with mem_sim
                  show ?thesis
                    by auto
                qed
              } hence u_m_eq: "u_m = m" by - (rule ext, auto) 
              {
                assume safe: "map owned u_ts,map released u_ts,i (is,θ,u_m,𝒟,𝒪,u_shared) "
                then have False
                proof cases
                  case (Read a volatile t)
                  with ts_i "is" obtain
                    ins: "ins = Read volatile a t"  and
                    access_cond: "a  𝒪  a  read_only u_shared  volatile  a  dom u_shared" and
                    rels_cond: "j<length u_ts. i  j  ((map released u_ts) ! j) a  Some False" and
                    rels_non_volatile_cond: "¬ volatile  (j<length u_ts. i  j  a  dom ((map released u_ts) ! j) )" and
                    clean: "volatile  ¬ 𝒟"
                    by auto

                  from race ts_j
                  have rc: "augment_rels (dom 𝒮l) R l a = Some False  
                            (¬ volatile  a  dom (augment_rels (dom 𝒮l) R l))"
                    by (auto simp add: races_def ins eqs)
                  from rels_cond [rule_format, simplified, OF j_bound_u neq_j_i [symmetric]] u_ts_j tsl_j j_bound_u
                  have l_a: "l a  Some False"
                    by auto
                  from dist_u_ts_inter.simple_ownership_distinct [OF j_bound_u i_bound_u neq_j_i u_ts_j [simplified tsl_j] 
                    u_ts_i [simplified ts_i]] 
                  have dist: "𝒪l  𝒪 = {}"
                    by auto

                  show ?thesis 
                  proof (cases volatile)
                    case True
                    note volatile=this
                    show ?thesis
                    proof (cases "a  R")
                      case False 
                      with rc l_a show False
                        by (auto simp add: augment_rels_def volatile)
                    next
                      case True
                      with R_owns
                      have a_ownsl: "a  𝒪l"
                        by auto
                      from shared_unch [rule_format, OF a_ownsl]
                      have u_shared_eq: "u_shared a = 𝒮l a"
                        by auto
                      from a_ownsl dist have "a  𝒪"
                        by auto
                      moreover
                      {
                        assume "a  read_only u_shared"
                        with u_shared_eq have "𝒮l a = Some False"
                          by (auto simp add: read_only_def)
                        with rc True l_a have False
                          by (auto simp add: augment_rels_def split: option.splits simp add: domIff volatile)
                      }
                      moreover
                      {
                        assume "a  dom u_shared"
                        with u_shared_eq rc True l_a have False
                          by (auto simp add: augment_rels_def split: option.splits simp add: domIff volatile)
                      }
                      ultimately show False
                      using access_cond 
                        by auto
                    qed
                  next
                    case False
                    note non_volatile = this
                    from rels_non_volatile_cond [rule_format, OF False j_bound_u neq_j_i [symmetric]] u_ts_j tsl_j j_bound_u
                    have l_a: "l a = None"
                      by (auto simp add: domIff)
                    show ?thesis
                    proof (cases "a  R")
                      case False 
                      with rc l_a show False
                        by (auto simp add: augment_rels_def non_volatile domIff)
                    next
                      case True
                      with R_owns
                      have a_ownsl: "a  𝒪l"
                        by auto
                      from shared_unch [rule_format, OF a_ownsl]
                      have u_shared_eq: "u_shared a = 𝒮l a"
                        by auto
                      from a_ownsl dist have a_unowned: "a  𝒪"
                        by auto
                      moreover
                      from ro_last_last interpret
                      read_only_unowned 𝒮l tsl .
                      from read_only_unowned [OF j_boundl tsl_j] a_ownsl have a_unsh: "a   read_only 𝒮l" by auto
                      {
                        assume "a  read_only u_shared"
                        with u_shared_eq have sh: "𝒮l a = Some False"
                          by (auto simp add: read_only_def)
                        
                        with rc True l_a access_cond u_shared_eq a_unowned sh a_ownsl a_unsh have False
                          by (auto simp add: augment_rels_def split: option.splits simp add: domIff non_volatile read_only_def)
                      }
                      moreover
                      {
                        assume "a  dom u_shared"
                        with u_shared_eq rc True l_a a_ownsl a_unsh access_cond dist have False
                          by (auto simp add: augment_rels_def split: option.splits simp add: domIff non_volatile read_only_def)
                      }
                      ultimately show False
                      using access_cond 
                        by (auto)
                    qed
                  qed
                next
                  case (WriteNonVolatile a D f A' L' R' W')
                  with ts_i "is" obtain
                    ins: "ins = Write False a (D, f) A' L' R' W'"  and
                    a_owned: "a  𝒪" and a_unshared: "a  dom u_shared" and
                    a_unreleased: "j<length u_ts. i  j  a  dom ((map released u_ts) ! j)"
                    by auto
                  from dist_u_ts_inter.simple_ownership_distinct [OF j_bound_u i_bound_u neq_j_i u_ts_j [simplified tsl_j] 
                    u_ts_i [simplified ts_i]] 
                  have dist: "𝒪l  𝒪 = {}"
                    by auto
                  from race ts_j
                  have rc: "a  dom (augment_rels (dom 𝒮l) R l)"
                    by (auto simp add: races_def ins eqs)
                  from a_unreleased [rule_format, simplified, OF j_bound_u neq_j_i [symmetric]] u_ts_j tsl_j j_bound_u
                  have l_a: "a  dom l"
                    by auto
                  show False
                  proof (cases "a  R")
                    case False 
                    with rc l_a show False
                      by (auto simp add: augment_rels_def domIff)
                  next
                    case True
                    with R_owns
                    have a_ownsl: "a  𝒪l"
                      by auto
                    with a_owned dist show False
                      by auto
                  qed
                next
                  case (WriteVolatile a A' L' R' D f W')
                  with ts_i "is" obtain
                    ins: "ins = Write True a (D, f) A' L' R' W'"  and
                    a_un_owned_released: "j<length u_ts. i  j  
                      a  ((map owned u_ts) ! j)  a  dom ((map released u_ts) ! j)" and
                    A'_owns_shared: "A'  dom u_shared  𝒪" and
                    L'_A': "L'  A'" and
                    R'_owned: "R'  𝒪" and
                    A'_R': "A'  R' = {}" and
                    acq_ok: "j<length u_ts. i  j  A'  ((map owned u_ts) ! j  dom ((map released u_ts) ! j)) = {}" and
                   
                    writeable: "a  read_only u_shared"
                    by auto
                  from a_un_owned_released [rule_format, simplified, OF j_bound_u neq_j_i [symmetric]] u_ts_j tsl_j j_bound_u
                  obtain 𝒪l_a: "a  𝒪l" and l_a: "a  dom (l)"
                    by auto
                  from acq_ok [rule_format, simplified, OF j_bound_u neq_j_i [symmetric]] u_ts_j tsl_j j_bound_u
                  obtain 𝒪l_A': "A'  𝒪l = {}" and l_A': "A'  dom (l) = {}"
                    by auto
                  {
                    assume rc: "a  dom (augment_rels (dom 𝒮l) R l)"
                    have False
                    proof (cases "a  R")
                      case False 
                      with rc l_a show False
                        by (auto simp add: augment_rels_def domIff)
                    next
                      case True
                      with R_owns
                      have a_ownsl: "a  𝒪l"
                        by auto
                      with 𝒪l_a show False
                        by auto
                    qed
                  }
                  moreover 
                  {
                    assume rc: "A'  dom (augment_rels (dom 𝒮l) R l)  {}"
                    then obtain a' where a'_A': "a'  A'" and a'_aug: "a'  dom (augment_rels (dom 𝒮l) R l)"
                      by auto
                    have False
                    proof (cases "a'  R")
                      case False 
                      with a'_aug a'_A' l_A' show False
                        by (auto simp add: augment_rels_def domIff)
                    next
                      case True
                      with R_owns have a'_ownsl: "a'  𝒪l"
                        by auto
                      with 𝒪l_A' a'_A' show False
                        by auto
                    qed
                  }
                  ultimately show False
                  using race ts_j 
                    by (auto simp add: races_def ins eqs)
                next
                  case Fence
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def split: if_split_asm)
                next
                  case (Ghost A' L' R' W')
                  with ts_i "is" obtain
                    ins: "ins = Ghost A' L' R' W'"  and
                    A'_owns_shared: "A'  dom u_shared  𝒪" and
                    L'_A': "L'  A'" and
                    R'_owned: "R'  𝒪" and
                    A'_R': "A'  R' = {}" and
                    acq_ok: "j<length u_ts. i  j  A'  ((map owned u_ts) ! j  dom ((map released u_ts) ! j)) = {}" 
                    by auto
                  from acq_ok [rule_format, simplified, OF j_bound_u neq_j_i [symmetric]] u_ts_j tsl_j j_bound_u
                  obtain 𝒪l_A': "A'  𝒪l = {}" and l_A': "A'  dom (l) = {}"
                    by auto
                  
                  from race ts_j 
                  obtain a'  where a'_A': "a'  A'" and 
                    a'_aug: "a'  dom (augment_rels (dom 𝒮l) R l)"
                    by (auto simp add: races_def ins eqs)
                  show False
                  proof (cases "a'  R")
                    case False 
                    with a'_aug a'_A' l_A' show False
                      by (auto simp add: augment_rels_def domIff)
                  next
                    case True
                    with R_owns have a'_ownsl: "a'  𝒪l"
                      by auto
                    with 𝒪l_A' a'_A' show False
                      by auto
                  qed
                next
                  case (RMWReadOnly cond t a D f ret A' L' R' W')
                  with ts_i "is" obtain
                    ins: "ins = RMW a t (D, f) cond ret A' L' R' W'" and
                    owned_or_shared: "a  𝒪  a  dom u_shared" and
                    cond: "¬ cond (θ(t  u_m a))" and
                    rels_race: "j<length (map owned u_ts). i  j  ((map released u_ts) ! j) a  Some False"
                    by auto
                  from dist_u_ts_inter.simple_ownership_distinct [OF j_bound_u i_bound_u neq_j_i u_ts_j [simplified tsl_j] 
                    u_ts_i [simplified ts_i]] 
                  have dist: "𝒪l  𝒪 = {}"
                    by auto
                  from race ts_j cond
                  have rc: "augment_rels (dom 𝒮l) R l a = Some False"
                    by (auto simp add: races_def ins eqs u_m_eq)

                  from rels_race [rule_format, simplified, OF j_bound_u neq_j_i [symmetric]] 
                    u_ts_j tsl_j j_bound_u
                  have l_a: "l a  Some False"
                    by auto

                  show ?thesis
                  proof (cases "a  R")
                    case False 
                    with rc l_a show False
                      by (auto simp add: augment_rels_def)
                  next
                    case True
                    with R_owns
                    have a_ownsl: "a  𝒪l"
                      by auto
                    from shared_unch [rule_format, OF a_ownsl]
                    have u_shared_eq: "u_shared a = 𝒮l a"
                      by auto
                    from a_ownsl dist have "a  𝒪"
                      by auto
                    with u_shared_eq rc True l_a owned_or_shared show False
                      by (auto simp add: augment_rels_def split: option.splits simp add: domIff)
                  qed
                next
                  case (RMWWrite cond t a A' L' R' D f ret W')
                  with ts_i "is" obtain
                    ins: "ins = RMW a t (D, f) cond ret A' L' R' W'" and
                    cond: "cond (θ(t  u_m a))" and
                    a_un_owned_released: "j<length (map owned u_ts). i  j  a  (map owned u_ts) ! j  dom ((map released u_ts) ! j)" and
                    A'_owns_shared:"A'  dom u_shared  𝒪" and
                    L'_A': "L'  A'" and
                    R'_owned: "R'  𝒪" and
                    A'_R': "A'  R' = {}" and
                    acq_ok: "j<length (map owned u_ts). i  j  A'  ((map owned u_ts) ! j  dom ((map released u_ts) ! j)) = {}" and 
                    writeable: "a  read_only u_shared"
                    by auto


                  from a_un_owned_released [rule_format, simplified, OF j_bound_u neq_j_i [symmetric]] u_ts_j tsl_j j_bound_u
                  obtain 𝒪l_a: "a  𝒪l" and l_a: "a  dom (l)"
                    by auto
                  from acq_ok [rule_format, simplified, OF j_bound_u neq_j_i [symmetric]] u_ts_j tsl_j j_bound_u
                  obtain 𝒪l_A': "A'  𝒪l = {}" and l_A': "A'  dom (l) = {}"
                    by auto
                  {
                    assume rc: "a  dom (augment_rels (dom 𝒮l) R l)"
                    have False
                    proof (cases "a  R")
                      case False 
                      with rc l_a show False
                        by (auto simp add: augment_rels_def domIff)
                    next
                      case True
                      with R_owns
                      have a_ownsl: "a  𝒪l"
                        by auto
                      with 𝒪l_a show False
                        by auto
                    qed
                  }
                  moreover 
                  {
                    assume rc: "A'  dom (augment_rels (dom 𝒮l) R l)  {}"
                    then obtain a' where a'_A': "a'  A'" and a'_aug: "a'  dom (augment_rels (dom 𝒮l) R l)"
                      by auto
                    have False
                    proof (cases "a'  R")
                      case False 
                      with a'_aug a'_A' l_A' show False
                        by (auto simp add: augment_rels_def domIff)
                    next
                      case True
                      with R_owns have a'_ownsl: "a'  𝒪l"
                        by auto
                      with 𝒪l_A' a'_A' show False
                        by auto
                    qed
                  }
                  ultimately show False
                  using race ts_j cond
                    by (auto simp add: races_def ins eqs u_m_eq)
                next
                next
                  case Nil
                  then show ?thesis
                  using  ts_i tsl_j race "is" j_bound i_bound u_ts_i u_ts_j leq_u_ts neq_j_i ts_j
                    by (auto simp add:eqs races_def split: if_split_asm)
                qed
              }
              hence "¬ safe_delayed (u_ts, u_m, u_shared)"
                apply (clarsimp simp add: safe_delayed_def)
                apply (rule_tac x=i in exI)
                using u_ts_i ts_i i_bound_u
                apply auto
                done
              moreover
              from safe_delayed_c_undo' [rule_format, of n] c_undo_n
              have "safe_delayed (u_ts, u_m, u_shared)"
                by simp
              ultimately have False
                by simp
              thus ?thesis 
                by simp
            qed
          qed
        next
          case (StoreBuffer _ p "is" θ sb 𝒟 𝒪  sb' 𝒪' ℛ')
          hence False 
            by (auto simp add: empty_storebuffer_step_def)
          thus ?thesis ..
        qed
      qed
      } (* FIXME indentation *)
      ultimately show ?thesis
      using last_action_of_thread [where i=j, OF trace]
         by blast
   qed
 qed(* FIXME indentation *)

datatype 'p memref = 
   Writesb bool addr sop val acq lcl rel wrt
 | Readsb bool addr tmp val 
 | Progsb 'p 'p "instrs"
 | Ghostsb acq lcl rel wrt

type_synonym 'p store_buffer = "'p memref list"
inductive flush_step:: "memory × 'p store_buffer × owns × rels × shared  memory × 'p store_buffer × owns × rels × shared  bool" 
  ("_ f _" [60,60] 100)
where
  Writesb: "𝒪' = (if volatile then 𝒪  A - R else 𝒪);
           𝒮' = (if volatile then 𝒮 ⊕⇘WR ⊖⇘AL else 𝒮);
          ℛ'=(if volatile then Map.empty else )
          
          (m, Writesb volatile a sop v A L R W# rs,𝒪,,𝒮) f (m(a := v), rs,𝒪',ℛ',𝒮')"
| Readsb: "(m, Readsb volatile a t v#rs,𝒪,,𝒮) f (m, rs,𝒪,, 𝒮)"
| Progsb: "(m, Progsb p p' is#rs,𝒪,, 𝒮) f (m, rs,𝒪,, 𝒮)"
| Ghost: "(m, Ghostsb A L R W# rs,𝒪,,𝒮) f (m, rs,𝒪  A - R, augment_rels (dom 𝒮) R , 𝒮 ⊕⇘WR ⊖⇘AL )"

abbreviation flush_steps::"memory × 'p store_buffer × owns × rels × shared  memory × 'p store_buffer × owns × rels × shared bool" 
  ("_ f* _" [60,60] 100)
where
"flush_steps == flush_step^**"

term "x f* Y"

lemmas flush_step_induct =  
  flush_step.induct [split_format (complete),
  consumes 1, case_names Writesb Readsb Progsb Ghost]

inductive store_buffer_step:: "memory × 'p store_buffer × 'owns × 'rels × 'shared  memory × 'p memref list × 'owns × 'rels × 'shared  bool" 
  ("_ w _" [60,60] 100)
where
  SBWritesb: "
          (m, Writesb volatile a sop v A L R W# rs,𝒪,,𝒮) w (m(a := v), rs,𝒪,,𝒮)"

abbreviation store_buffer_steps::"memory × 'p store_buffer × 'owns × 'rels × 'shared  memory × 'p store_buffer × 'owns × 'rels × 'shared bool" 
  ("_ →w* _" [60,60] 100)
where
"store_buffer_steps == store_buffer_step^**"

term "x →w* Y"

fun buffered_val :: "'p memref list  addr  val option"
where
  "buffered_val [] a = None"
| "buffered_val (r # rs) a' = 
   (case r of
      Writesb volatile a _ v _ _ _ _  (case buffered_val rs a' of 
                               None  (if a'=a then Some v else None)
                             | Some v'  Some v')
     | _  buffered_val rs a')"

definition address_of :: "'p memref  addr set"
where
"address_of r = (case r of Writesb volatile a _ v _ _ _ _  {a} | Readsb volatile a t v  {a} |
                  _  {})"

lemma address_of_simps [simp]: 
"address_of (Writesb volatile a sop v A L R W) = {a}"
"address_of (Readsb volatile a t v) = {a}"
"address_of (Progsb p p' is) = {}"
"address_of (Ghostsb A L R W) = {}"
  by (auto simp add: address_of_def)

definition is_volatile :: "'p memref  bool"
where
"is_volatile r = (case r of Writesb volatile a _ v _ _ _ _ volatile | Readsb volatile a t v  volatile
 | _  False)"

lemma is_volatile_simps [simp]: 
"is_volatile (Writesb volatile a sop v A L R W) = volatile"
"is_volatile (Readsb volatile a t v) = volatile"
"is_volatile (Progsb p p' is) = False"
"is_volatile (Ghostsb A L R W) = False"
  by (auto simp add: is_volatile_def)

definition is_Writesb:: "'p memref  bool"
where
"is_Writesb r = (case r of Writesb volatile a _ v _ _ _ _ True | _  False)"

definition is_Readsb:: "'p memref  bool"
where
"is_Readsb r = (case r of Readsb volatile a t v  True | _  False)"

definition is_Progsb:: "'p memref  bool"
where
"is_Progsb r = (case r of Progsb _ _ _  True | _  False)"

definition is_Ghostsb:: "'p memref  bool"
where
"is_Ghostsb r = (case r of Ghostsb _ _ _ _  True | _  False)"

lemma is_Writesb_simps [simp]: 
"is_Writesb (Writesb volatile a sop v A L R W) = True"
"is_Writesb (Readsb volatile a t v) = False"
"is_Writesb (Progsb p p' is) = False"
"is_Writesb (Ghostsb A L R W) = False"
  by (auto simp add: is_Writesb_def)

lemma is_Readsb_simps [simp]: 
"is_Readsb (Readsb volatile a t v) = True"
"is_Readsb (Writesb volatile a sop v A L R W) = False"
"is_Readsb (Progsb p p' is) = False"
"is_Readsb (Ghostsb A L R W) = False"
  by (auto simp add: is_Readsb_def)

lemma is_Progsb_simps [simp]: 
"is_Progsb (Readsb volatile a t v) = False"
"is_Progsb (Writesb volatile a sop v A L R W) = False"
"is_Progsb (Progsb p p' is) = True"
"is_Progsb (Ghostsb A L R W) = False"
  by (auto simp add: is_Progsb_def)

lemma is_Ghostsb_simps [simp]: 
"is_Ghostsb (Readsb volatile a t v) = False"
"is_Ghostsb (Writesb volatile a sop v A L R W) = False"
"is_Ghostsb (Progsb p p' is) = False"
"is_Ghostsb (Ghostsb A L R W) = True"
  by (auto simp add: is_Ghostsb_def)

definition is_volatile_Writesb:: "'p memref  bool"
where
"is_volatile_Writesb r = (case r of Writesb volatile a _ v _ _ _ _ volatile | _  False)"

lemma is_volatile_Writesb_simps [simp]: 
"is_volatile_Writesb (Writesb volatile a sop v A L R W) = volatile"
"is_volatile_Writesb (Readsb volatile a t v) = False"
"is_volatile_Writesb (Progsb p p' is) = False"
"is_volatile_Writesb (Ghostsb A L R W) = False"
  by (auto simp add: is_volatile_Writesb_def)

lemma is_volatile_Writesb_address_of [simp]: "is_volatile_Writesb x  address_of x  {}"
  by (cases x) auto

definition is_volatile_Readsb:: "'p memref  bool"
where
"is_volatile_Readsb r = (case r of Readsb volatile a t v  volatile | _  False)"

lemma is_volatile_Readsb_simps [simp]: 
"is_volatile_Readsb (Readsb volatile a t v) = volatile"
"is_volatile_Readsb (Writesb volatile a sop v A L R W) = False"
"is_volatile_Readsb (Progsb p p' is) = False"
"is_volatile_Readsb (Ghostsb A L R W) = False"
  by (auto simp add: is_volatile_Readsb_def)

definition is_non_volatile_Writesb:: "'p memref  bool"
where
"is_non_volatile_Writesb r = (case r of Writesb volatile a _ v _ _ _ _ ¬ volatile | _  False)"

lemma is_non_volatile_Writesb_simps [simp]: 
"is_non_volatile_Writesb (Writesb volatile a sop v A L R W) = (¬ volatile)"
"is_non_volatile_Writesb (Readsb volatile a t v) = False"
"is_non_volatile_Writesb (Progsb p p' is) = False"
"is_non_volatile_Writesb (Ghostsb A L R W) = False"
  by (auto simp add: is_non_volatile_Writesb_def)

definition is_non_volatile_Readsb:: "'p memref  bool"
where
"is_non_volatile_Readsb r = (case r of Readsb volatile a t v  ¬ volatile | _  False)"

lemma is_non_volatile_Readsb_simps [simp]: 
"is_non_volatile_Readsb (Readsb volatile a t v) = (¬ volatile)"
"is_non_volatile_Readsb (Writesb volatile a sop v A L R W) = False"
"is_non_volatile_Readsb (Progsb p p' is) = False"
"is_non_volatile_Readsb (Ghostsb A L R W) = False"
  by (auto simp add: is_non_volatile_Readsb_def)

lemma is_volatile_split: "is_volatile r = 
  (is_volatile_Readsb r  is_volatile_Writesb r)"
  by (cases r) auto

lemma is_non_volatile_split: 
  "¬ is_volatile r = (is_non_volatile_Readsb r  is_non_volatile_Writesb r  is_Progsb r  is_Ghostsb r)"
  by (cases r) auto

fun outstanding_refs:: "('p memref  bool)  'p memref list  addr set"
where
  "outstanding_refs P [] = {}"
| "outstanding_refs P (r#rs) = (if P r then (address_of r)  (outstanding_refs P rs)
                                else outstanding_refs P rs)"

lemma outstanding_refs_conv: "outstanding_refs P sb = (address_of ` {r. r  set sb  P r})"
  by (induct sb) auto

lemma outstanding_refs_append: 
  "ys. outstanding_refs vol (xs@ys) = outstanding_refs vol xs  outstanding_refs vol ys"
  by (auto simp add: outstanding_refs_conv)

(*
lemma outstanding_refs_empty_conv:
"(outstanding_refs P sb = {}) = (∀r ∈ set sb. ¬ (P r))"
  by (auto simp add: outstanding_refs_conv)
*)

lemma outstanding_refs_empty_negate: "(outstanding_refs P sb = {})  
       (outstanding_refs (Not  P) sb = (address_of ` set sb))"
  by (auto simp add: outstanding_refs_conv)

lemma outstanding_refs_mono_pred:
  "sb sb'. 
     r. P r  P' r  outstanding_refs P sb  outstanding_refs P' sb"
  by (auto simp add: outstanding_refs_conv)

lemma outstanding_refs_mono_set:
  "sb sb'. 
     set sb  set sb'  outstanding_refs P sb  outstanding_refs P sb'"
  by (auto simp add: outstanding_refs_conv)

lemma outstanding_refs_takeWhile:
"outstanding_refs P (takeWhile P' sb)  outstanding_refs P sb"
apply (rule outstanding_refs_mono_set)
apply (auto dest: set_takeWhileD)
done

lemma outstanding_refs_subsets:
  "outstanding_refs is_volatile_Writesb sb  outstanding_refs is_Writesb sb"
  "outstanding_refs is_non_volatile_Writesb sb  outstanding_refs is_Writesb sb"

  "outstanding_refs is_volatile_Readsb sb  outstanding_refs is_Readsb sb"
  "outstanding_refs is_non_volatile_Readsb sb  outstanding_refs is_Readsb sb"

  "outstanding_refs is_non_volatile_Writesb sb  outstanding_refs (Not  is_volatile) sb"
  "outstanding_refs is_non_volatile_Readsb sb  outstanding_refs (Not  is_volatile) sb"

  "outstanding_refs is_volatile_Writesb sb  outstanding_refs (is_volatile) sb"
  "outstanding_refs is_volatile_Readsb sb  outstanding_refs (is_volatile) sb"

  "outstanding_refs is_non_volatile_Writesb sb  outstanding_refs (Not  is_volatile_Writesb) sb"
  "outstanding_refs is_non_volatile_Readsb sb  outstanding_refs (Not  is_volatile_Writesb) sb"
  "outstanding_refs is_volatile_Readsb sb  outstanding_refs (Not  is_volatile_Writesb) sb"
  "outstanding_refs is_Readsb sb  outstanding_refs (Not  is_volatile_Writesb) sb"
by (auto intro!:outstanding_refs_mono_pred simp add: is_volatile_Writesb_def is_non_volatile_Writesb_def 
  is_volatile_Readsb_def is_non_volatile_Readsb_def is_Readsb_def split: memref.splits)


lemma outstanding_non_volatile_refs_conv: 
  "outstanding_refs (Not  is_volatile) sb =
   outstanding_refs is_non_volatile_Writesb sb  outstanding_refs is_non_volatile_Readsb sb"
apply (induct sb) 
apply simp
  subgoal for a sb
    by (case_tac a, auto)
done


lemma outstanding_volatile_refs_conv: 
  "outstanding_refs is_volatile sb =
   outstanding_refs is_volatile_Writesb sb  outstanding_refs is_volatile_Readsb sb"
apply (induct sb) 
apply simp
  subgoal for a sb
    by (case_tac a, auto)
done

lemma outstanding_is_Writesb_refs_conv: 
  "outstanding_refs is_Writesb sb =
   outstanding_refs is_non_volatile_Writesb sb  outstanding_refs is_volatile_Writesb sb"
apply (induct sb) 
apply simp
  subgoal for a sb
    by (case_tac a, auto)
done

lemma outstanding_is_Readsb_refs_conv: 
  "outstanding_refs is_Readsb sb =
   outstanding_refs is_non_volatile_Readsb sb  outstanding_refs is_volatile_Readsb sb"
apply (induct sb) 
apply simp
  subgoal for a sb
    by (case_tac a, auto)
done

lemma outstanding_not_volatile_Readsb_refs_conv: "outstanding_refs (Not  is_volatile_Readsb) sb =
       outstanding_refs is_Writesb sb  outstanding_refs is_non_volatile_Readsb sb"
apply (induct sb)
apply (clarsimp)
  subgoal for a sb
    by (case_tac a, auto)
done


lemmas misc_outstanding_refs_convs = outstanding_non_volatile_refs_conv outstanding_volatile_refs_conv
outstanding_is_Writesb_refs_conv outstanding_is_Readsb_refs_conv outstanding_not_volatile_Readsb_refs_conv

lemma no_outstanding_vol_write_takeWhile_append: "outstanding_refs is_volatile_Writesb sb = {}  
  takeWhile (Not  is_volatile_Writesb) (sb@xs) = sb@(takeWhile (Not  is_volatile_Writesb) xs)"
apply (induct sb)
apply (auto split: if_split_asm)
done

lemma outstanding_vol_write_takeWhile_append: "outstanding_refs is_volatile_Writesb sb  {}  
  takeWhile (Not  is_volatile_Writesb) (sb@xs) = (takeWhile (Not  is_volatile_Writesb) sb)"
apply (induct sb)
apply (auto split: if_split_asm)
done


lemma no_outstanding_vol_write_dropWhile_append: "outstanding_refs is_volatile_Writesb sb = {}  
  dropWhile (Not  is_volatile_Writesb) (sb@xs) = (dropWhile (Not  is_volatile_Writesb) xs)"
apply (induct sb)
apply (auto split: if_split_asm)
done

lemma outstanding_vol_write_dropWhile_append: "outstanding_refs is_volatile_Writesb sb  {}  
  dropWhile (Not  is_volatile_Writesb) (sb@xs) = (dropWhile (Not  is_volatile_Writesb) sb)@xs"
apply (induct sb)
apply (auto split: if_split_asm)
done

lemmas outstanding_vol_write_take_drop_appends =
no_outstanding_vol_write_takeWhile_append
outstanding_vol_write_takeWhile_append
no_outstanding_vol_write_dropWhile_append
outstanding_vol_write_dropWhile_append

lemma outstanding_refs_is_non_volatile_Writesb_takeWhile_conv:
  "outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sb) =
       outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sb)"
apply (induct sb)
apply  clarsimp
  subgoal for a sb
    by (case_tac a, auto)
done


lemma dropWhile_not_vol_write_empty:
  "outstanding_refs is_volatile_Writesb sb = {}  (dropWhile (Not  is_volatile_Writesb) sb) = []"
apply (induct sb)
apply (auto split: if_split_asm)
done

lemma takeWhile_not_vol_write_outstanding_refs:
  "outstanding_refs is_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sb) = {}"
apply (induct sb)
apply (auto split: if_split_asm)
done

lemma no_volatile_Writesbs_conv: "(outstanding_refs is_volatile_Writesb sb = {}) = 
       (r  set sb. (v' sop' a' A L R W. r  Writesb True a' sop' v' A L R W))"
  by (force simp add: outstanding_refs_conv is_volatile_Writesb_def split: memref.splits)

lemma no_volatile_Readsbs_conv: "(outstanding_refs is_volatile_Readsb sb = {}) = 
       (r  set sb. (v' t' a'. r  Readsb True a' t' v'))"
  by (force simp add: outstanding_refs_conv is_volatile_Readsb_def split: memref.splits)



inductive sb_memop_step :: "(instrs × tmps × 'p store_buffer × memory × 'dirty × 'owns × 'rels × 'shared )  
                  (instrs × tmps × 'p store_buffer × memory × 'dirty × 'owns × 'rels × 'shared )  bool" 
                    ("_ sb _" [60,60] 100)
where
  SBReadBuffered: 
  "buffered_val sb a = Some v
   
   (Read volatile a t # is,θ, sb, m,𝒟, 𝒪, , 𝒮) sb
          (is, θ (tv), sb, m,𝒟, 𝒪,, 𝒮)"

| SBReadUnbuffered: 
  "buffered_val sb a = None 
   
   (Read volatile a t # is, θ, sb, m,𝒟, 𝒪, , 𝒮) sb
          (is, θ (tm a), sb, m,𝒟, 𝒪, , 𝒮)"

| SBWriteNonVolatile:
  "(Write False a (D,f) A L R W#is, θ, sb, m,𝒟,𝒪, , 𝒮) sb
          (is, θ, sb@ [Writesb False a (D,f) (f θ) A L R W], m,𝒟, 𝒪, , 𝒮)"

| SBWriteVolatile:
   
   "(Write True a (D,f) A L R W# is, θ, sb, m,𝒟, 𝒪, , 𝒮) sb
         (is, θ, sb@[Writesb True a (D,f) (f θ) A L R W], m,𝒟, 𝒪, , 𝒮)"

| SBFence:
  "(Fence # is, θ, [], m,𝒟, 𝒪, , 𝒮) sb (is, θ, [], m,𝒟, 𝒪, , 𝒮)"

| SBRMWReadOnly:
  "¬ cond (θ(tm a))  
   (RMW a t (D,f) cond ret A L R W# is, θ, [], m,𝒟, 𝒪, , 𝒮) sb (is, θ(tm a),[], m,𝒟, 𝒪, , 𝒮)"

| SBRMWWrite:
  "cond (θ(tm a))  
   (RMW a t (D,f) cond ret A L R W# is, θ, [], m,𝒟, 𝒪, , 𝒮) sb
         (is, θ(tret (m a) (f(θ(tm a)))),[], m(a:= f(θ(tm a))),𝒟, 𝒪, , 𝒮)"

| SBGhost:
  "(Ghost A  L R W# is, θ, sb, m,𝒟, 𝒪, , 𝒮) sb
         (is, θ, sb, m,𝒟, 𝒪, , 𝒮)"


inductive sbh_memop_step :: "
                  (instrs × tmps × 'p store_buffer × memory × bool × owns × rels × shared ) 
                  (instrs × tmps × 'p store_buffer × memory × bool × owns × rels × shared )  bool" 
                    ("_ sbh _" [60,60] 100)
where
  SBHReadBuffered: 
  "buffered_val sb a = Some v
   
   (Read volatile a t # is, θ, sb, m, 𝒟, 𝒪, , 𝒮) sbh
          (is, θ (tv), sb@[Readsb volatile a t v], m, 𝒟, 𝒪, , 𝒮)"

| SBHReadUnbuffered: 
  "buffered_val sb a = None 
   
   (Read volatile a t # is, θ, sb, m, 𝒟, 𝒪, , 𝒮) sbh
          (is, θ (tm a), sb@[Readsb volatile a t (m a)], m, 𝒟, 𝒪, , 𝒮)"

| SBHWriteNonVolatile:
  "(Write False a (D,f) A L R W#is, θ, sb, m, 𝒟, 𝒪, , 𝒮) sbh
          (is, θ, sb@ [Writesb False a (D,f) (f θ) A L R W], m, 𝒟, 𝒪, , 𝒮)"

| SBHWriteVolatile:
  "(Write True a (D,f) A L R W# is, θ, sb, m, 𝒟, 𝒪, , 𝒮) sbh
         (is, θ, sb@[Writesb True a (D,f) (f θ) A L R W], m, True, 𝒪, , 𝒮)"

| SBHFence:
  "(Fence # is, θ, [], m, 𝒟, 𝒪, , 𝒮) sbh (is, θ, [], m, False, 𝒪, Map.empty, 𝒮)"

| SBHRMWReadOnly:
  "¬ cond (θ(tm a))  
   (RMW a t (D,f) cond ret A L R W# is, θ, [], m, 𝒟, 𝒪, , 𝒮) sbh (is, θ(tm a),[], m, False, 𝒪, Map.empty, 𝒮)"

| SBHRMWWrite:
  "cond (θ(tm a))  
   (RMW a t (D,f) cond ret A L R W# is, θ, [], m, 𝒟, 𝒪, , 𝒮) sbh
         (is, θ(tret (m a) (f(θ(tm a)))),[], m(a:= f(θ(tm a))), False, 𝒪  A - R,Map.empty, 𝒮 ⊕⇘WR ⊖⇘AL)"

| SBHGhost:
  "(Ghost A L R W# is, θ, sb, m, 𝒟, 𝒪, , 𝒮) sbh
         (is, θ, sb@[Ghostsb A L R W], m, 𝒟, 𝒪, , 𝒮)"


interpretation direct:  memory_system direct_memop_step id_storebuffer_step .
interpretation sb: memory_system sb_memop_step store_buffer_step .
interpretation sbh: memory_system sbh_memop_step flush_step .

primrec non_volatile_owned_or_read_only:: "bool  shared  owns  'a memref list  bool"
where
"non_volatile_owned_or_read_only pending_write 𝒮 𝒪 [] = True"
| "non_volatile_owned_or_read_only pending_write 𝒮 𝒪 (x#xs) =
  (case x of
    Readsb volatile a t v  
     (¬volatile  pending_write  (a  𝒪  a  read_only 𝒮))  
      non_volatile_owned_or_read_only pending_write 𝒮 𝒪 xs
  | Writesb volatile a sop v A L R W  
     (if volatile then non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) xs
      else a  𝒪  non_volatile_owned_or_read_only pending_write 𝒮 𝒪 xs)
  | Ghostsb A L R W  non_volatile_owned_or_read_only pending_write (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) xs
  | _  non_volatile_owned_or_read_only pending_write 𝒮 𝒪 xs)"

primrec acquired :: "bool  'a memref list  addr set  addr set"
where
"acquired pending_write [] A = (if pending_write then A else {})"
| "acquired pending_write (x#xs) A =
  (case x of
     Writesb volatile _ _ _ A' L R W 
        (if volatile then acquired True xs (if pending_write then (A  A' - R) else (A' - R)) 
         else acquired pending_write xs A)
   | Ghostsb A' L R W  acquired pending_write xs (if pending_write then (A  A' - R) else A) 
   | _  acquired pending_write xs A)"

primrec share :: "'a memref list  shared  shared"
where
"share [] S = S"
| "share (x#xs) S =
  (case x of
     Writesb volatile _ _ _ A L R W  (if volatile then (share xs (S ⊕⇘WR ⊖⇘AL)) else share xs S)
   | Ghostsb A L R W  share xs (S ⊕⇘WR ⊖⇘AL)
   | _  share xs S)"

primrec acquired_reads :: "bool  'a memref list  addr set  addr set"
where
"acquired_reads pending_write [] A = {}"
| "acquired_reads pending_write (x#xs) A =
  (case x of
     Readsb volatile a t v  (if pending_write  ¬ volatile  a  A 
                             then insert a (acquired_reads pending_write xs A)
                             else acquired_reads pending_write xs A)
   | Writesb volatile _ _ _ A' L R W  
         (if volatile then acquired_reads True xs (if pending_write then (A  A' - R) else (A' - R)) 
          else acquired_reads pending_write xs A)
   | Ghostsb A' L R W  acquired_reads pending_write xs (A  A' - R)
   | _  acquired_reads pending_write xs A)"

lemma union_mono_aux: "A  B  A  C  B  C"
  by blast

lemma set_minus_mono_aux: "A  B  A - C  B - C"
  by blast

lemma acquired_mono: "A B pending_write. A  B  acquired pending_write xs A  acquired pending_write xs B"
apply (induct xs)
apply  simp
subgoal for a xs A B pending_write
apply (case_tac a ) 
apply    clarsimp 
         subgoal for volatile a1 D f v A' L R W x
           apply (drule_tac C=A' in union_mono_aux)
           apply (drule_tac C="R" in set_minus_mono_aux)
           apply blast
           done 
apply   clarsimp
apply  clarsimp
apply clarsimp
subgoal for A' L R W x
  apply (drule_tac C=A' in union_mono_aux)
  apply (drule_tac C="R" in set_minus_mono_aux)
  apply blast
  done
done
done


lemma acquired_mono_in: 
  assumes x_in: "x  acquired pending_write xs A" 
  assumes sub: "A  B" 
  shows "x  acquired pending_write xs B"
using acquired_mono [OF sub, of pending_write xs] x_in
by blast

lemma acquired_no_pending_write:"A B. acquired False xs A = acquired False xs B"
  by (induct xs) (auto split: memref.splits)

lemma acquired_no_pending_write_in:
  "x  acquired False xs A  x  acquired False xs B"
  apply (subst acquired_no_pending_write)
  apply auto
  done

lemma acquired_pending_write_mono_in: "A B. x  acquired False xs A  x  acquired True xs B"
apply (induct xs)
apply (auto split: memref.splits if_split_asm intro: acquired_mono_in)
done

lemma acquired_pending_write_mono: "acquired False xs A  acquired True xs B"
  by (auto intro: acquired_pending_write_mono_in)

lemma acquired_append: "A pending_write. acquired pending_write (xs@ys) A = 
 acquired (pending_write  outstanding_refs is_volatile_Writesb xs  {}) ys (acquired pending_write xs A)"
  apply (induct xs)
  apply (auto split: memref.splits intro: acquired_no_pending_write_in)
  done

lemma acquired_take_drop: 
  "acquired (pending_write  outstanding_refs is_volatile_Writesb (takeWhile P xs)  {}) 
      (dropWhile P xs) (acquired pending_write (takeWhile P xs) A) = 
   acquired pending_write xs A"
proof -
  have "acquired pending_write xs A = acquired pending_write ((takeWhile P xs)@(dropWhile P xs)) A"
    by simp
  also
  from acquired_append [where xs="(takeWhile P xs)" and ys="(dropWhile P xs)"]
  have " = acquired (pending_write  outstanding_refs is_volatile_Writesb (takeWhile P xs)  {}) 
      (dropWhile P xs) (acquired pending_write (takeWhile P xs) A)"
    by simp
  finally show ?thesis
    by simp
qed

lemma share_mono: "A B. dom A  dom B  dom (share xs A)  dom (share xs B)"
apply (induct xs)
apply  simp
subgoal for a xs A B
apply (case_tac a)
apply    (clarsimp iff del: domIff)
         subgoal for volatile a1 D f v A' L R W x
         apply (drule_tac C="R" and x="W" in augment_mono_aux)
         apply (drule_tac C="L" in restrict_mono_aux)
         apply blast
         done
apply   clarsimp
apply  clarsimp
apply (clarsimp iff del: domIff)
subgoal for A' L R W x
apply (drule_tac C="R" and x="W" in augment_mono_aux)
apply (drule_tac C="L" in restrict_mono_aux)
apply blast
done
done
done

lemma share_mono_in: 
  assumes x_in: "x  dom (share xs A)" 
  assumes sub: "dom A  dom B" 
  shows "x  dom (share xs B)"
using share_mono [OF sub, of xs] x_in
by blast

lemma acquired_reads_mono: 
  "A B pending_write. A  B  acquired_reads pending_write xs A  acquired_reads pending_write xs B"
apply (induct xs)
apply  simp
subgoal for a xs A B pending_write
apply (case_tac a)
apply    clarsimp
         subgoal for volatile a1 D f v A' L R W x
         apply (drule_tac C="A'" in union_mono_aux)
         apply (drule_tac C="R" in set_minus_mono_aux)
         apply blast
         done 
apply   clarsimp
apply   blast
apply  clarsimp
apply clarsimp
subgoal for A' L R W x
apply (drule_tac C="A'" in union_mono_aux)
apply (drule_tac C="R" in set_minus_mono_aux)
apply blast
done
done
done

lemma acquired_reads_mono_in:
  assumes x_in: "x  acquired_reads pending_write xs A" 
  assumes sub: "A  B" 
  shows "x  acquired_reads pending_write xs B"
using acquired_reads_mono [OF sub, of pending_write xs] x_in
by blast

lemma acquired_reads_no_pending_write: "A B. acquired_reads False xs A = acquired_reads False xs B"
  by (induct xs) (auto split: memref.splits)

lemma acquired_reads_no_pending_write_in:
"x  acquired_reads False xs A  x  acquired_reads False xs B"
  apply (subst acquired_reads_no_pending_write)
  apply blast
  done

lemma acquired_reads_pending_write_mono:
  "A. acquired_reads False xs A  acquired_reads True xs A"
  by (induct xs) (auto split: memref.splits intro: acquired_reads_mono_in )

lemma acquired_reads_pending_write_mono_in: 
  assumes x_in: "x  acquired_reads False xs A" 
  shows "x  acquired_reads True xs A"
using acquired_reads_pending_write_mono [of xs A] x_in
by blast

lemma acquired_reads_append: "pending_write A. acquired_reads pending_write (xs@ys) A = 
  acquired_reads pending_write xs A  
  acquired_reads (pending_write  (outstanding_refs is_volatile_Writesb xs  {})) ys 
   (acquired pending_write xs A)"
proof (induct xs)
  case Nil thus ?case by (auto dest: acquired_reads_no_pending_write_in)
next
  case (Cons x xs)
  show ?case
  proof (cases x)
    case (Writesb volatile a sop v A L R W)
    show ?thesis
    proof (cases volatile)
      case False
      show ?thesis
	using Cons.hyps
	by (auto simp add: Writesb False)
    next
      case True
      show ?thesis
	using Cons.hyps
	by (auto simp add: Writesb True)
    qed
  next
    case (Readsb volatile a t v)
    show ?thesis
    proof (cases volatile)
      case False
      show ?thesis
	using Cons.hyps
	by (auto simp add: Readsb False)
    next
      case True
      show ?thesis
	using Cons.hyps
	by (auto simp add: Readsb True)
    qed
  next
    case Progsb
    with Cons.hyps show ?thesis by auto
  next
    case (Ghostsb A' L R W)
    have "(acquired False xs (A  A' -R )) = (acquired False xs A)"
      by (simp add: acquired_no_pending_write)
    with Cons.hyps show ?thesis by (auto simp add: Ghostsb)
  qed
qed

lemma in_acquired_reads_no_pending_write_outstanding_write: 
"A. a  acquired_reads False xs A  outstanding_refs (is_volatile_Writesb) xs  {}"
  apply (induct xs)
  apply simp
  apply (auto split: memref.splits)
  apply auto 
  done

lemma augment_read_only_mono: "read_only 𝒮  read_only 𝒮'  
  read_only (𝒮 ⊕⇘WR)  read_only (𝒮' ⊕⇘WR)"
  by (auto simp add: augment_shared_def read_only_def)

lemma restrict_read_only_mono: "read_only 𝒮  read_only 𝒮'  
  read_only (𝒮 ⊖⇘AL)  read_only (𝒮' ⊖⇘AL)"
  apply (clarsimp simp add: restrict_shared_def read_only_def split: option.splits if_split_asm)
  apply (rule conjI)
  apply  blast
  apply fastforce
  done


lemma share_read_only_mono: "𝒮 𝒮'. read_only 𝒮  read_only 𝒮' 
        read_only (share sb 𝒮)  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
      note read_only 𝒮  read_only 𝒮'
      from augment_read_only_mono [OF this]
      have "read_only (𝒮 ⊕⇘WR)  read_only (𝒮' ⊕⇘WR)".
      from restrict_read_only_mono [OF this, of A L]
      have "read_only (𝒮 ⊕⇘WR ⊖⇘AL)  read_only (𝒮' ⊕⇘WR ⊖⇘AL)".
      from Cons.hyps [OF this]
      show ?thesis
	by (clarsimp 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)
    note read_only 𝒮  read_only 𝒮'
    from augment_read_only_mono [OF this]
    have "read_only (𝒮 ⊕⇘WR)  read_only (𝒮' ⊕⇘WR)".
    from restrict_read_only_mono [OF this, of A L]
    have "read_only (𝒮 ⊕⇘WR ⊖⇘AL)  read_only (𝒮' ⊕⇘WR ⊖⇘AL)".
   from Cons.hyps [OF this]
   show ?thesis
     by (clarsimp simp add: Ghostsb)
  qed
qed


(*
lemma read_only_share_takeWhile:
  "⋀𝒮. read_only (share (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒮)
       ⊆ read_only 𝒮"
proof (induct sb)
  case Nil thus ?case by simp
next
  case (Cons x sb)
  show ?case
  proof (cases x)
    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)
  next
    case (Ghostsb A L R W)
    note Cons.hyps [of "𝒮 ⊖A L"] 
    moreover
    have "read_only (𝒮 ⊖A L) ⊆ read_only 𝒮"
      by (auto simp add: in_read_only_restrict_conv)
    ultimately
    show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed

lemma read_only_share_takeWhile_in:
"a ∈ read_only (share (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒮) ⟹ a ∈  read_only 𝒮"
using read_only_share_takeWhile
by blast
*)
lemma non_volatile_owned_or_read_only_append: 
"𝒪 𝒮 pending_write. non_volatile_owned_or_read_only pending_write 𝒮 𝒪 (xs@ys)
         = (non_volatile_owned_or_read_only pending_write 𝒮 𝒪 xs  
            non_volatile_owned_or_read_only (pending_write  outstanding_refs is_volatile_Writesb xs  {}) 
             (share xs 𝒮) (acquired True xs 𝒪) ys)"
apply (induct xs)
apply (auto split: memref.splits)
done

lemma non_volatile_owned_or_read_only_mono:
"𝒪 𝒪' 𝒮 pending_write. 𝒪  𝒪'  non_volatile_owned_or_read_only pending_write 𝒮 𝒪 xs 
   non_volatile_owned_or_read_only pending_write 𝒮 𝒪' xs"
  apply (induct xs)
  apply  simp
  subgoal for a xs 𝒪 𝒪' 𝒮 pending_write
  apply (case_tac a)
  apply    (clarsimp split: if_split_asm)
           subgoal for volatile a1 D f v A L R W
           apply (drule_tac C="A" in union_mono_aux)
           apply (drule_tac C="R" in set_minus_mono_aux)
           apply blast
           done
  apply    fastforce
  apply   fastforce
  apply  fastforce
  apply clarsimp
  subgoal for A L R W
  apply (drule_tac C="A" in union_mono_aux)
  apply (drule_tac C="R" in set_minus_mono_aux)
  apply blast
  done
  done
  done

lemma non_volatile_owned_or_read_only_shared_mono:
"𝒮 𝒮' 𝒪 pending_write. 𝒮 s 𝒮'  non_volatile_owned_or_read_only pending_write 𝒮 𝒪 xs 
   non_volatile_owned_or_read_only pending_write 𝒮' 𝒪 xs"
  apply (induct xs)
  apply  simp
  subgoal for a xs 𝒮 𝒮' 𝒪 pending_write
  apply (case_tac a)
  apply    (clarsimp split: if_split_asm) 
           subgoal for volatile a1 D f v A L R W
           apply (frule_tac C="R" and x="W" in augment_mono_map)
           apply (drule_tac A="𝒮 ⊕⇘WR" and C="L" in restrict_mono_map)
           apply (fastforce dest: read_only_mono)
           done
  apply   (fastforce dest: read_only_mono shared_leD)
  apply  fastforce
  subgoal for A L R W
  apply (frule_tac C="R" and x="W" in augment_mono_map)
  apply (drule_tac A="𝒮 ⊕⇘WR" and C="L" in restrict_mono_map)
  apply (fastforce dest: read_only_mono)
  done
  done
  done

lemma non_volatile_owned_or_read_only_pending_write_antimono:
"𝒪 𝒮. non_volatile_owned_or_read_only True 𝒮 𝒪 xs 
   non_volatile_owned_or_read_only False 𝒮 𝒪 xs"
  by (induct xs) (auto split: memref.splits)

primrec all_acquired :: "'a memref list  addr set"
where 
  "all_acquired [] = {}"
|  "all_acquired (i#is) =
    (case i of
       Writesb volatile _ _ _ A L R W  (if volatile then A  all_acquired is else all_acquired is)
     | Ghostsb A L R W  A  all_acquired is
     | _  all_acquired is)"

lemma all_acquired_append: "all_acquired (xs@ys) = all_acquired xs  all_acquired ys"
  apply (induct xs)
  apply (auto split: memref.splits)
  done

lemma acquired_reads_all_acquired: "𝒪 pending_write.
  acquired_reads pending_write sb 𝒪  𝒪  all_acquired sb"
apply (induct sb)
apply  clarsimp
apply (auto split: memref.splits)
done

(*
lemma acquired_takeWhile_non_volatile_Writesb: 
  "⋀A. (acquired True (takeWhile (Not ∘ is_volatile_Writesb) sb) A) = 
         A ∪ all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sb)"
*)
lemma acquired_takeWhile_non_volatile_Writesb: 
  "A. (acquired True (takeWhile (Not  is_volatile_Writesb) sb) A)  
         A  all_acquired (takeWhile (Not  is_volatile_Writesb) sb)"
apply (induct sb)
apply  clarsimp 
subgoal for a sb A
apply (case_tac a)
apply auto
done
done

lemma acquired_False_takeWhile_non_volatile_Writesb:
  "acquired False (takeWhile (Not  is_volatile_Writesb) sb) A = {}"
  apply (induct sb)
   apply simp
  subgoal for a sb
    by (case_tac a) auto
  done  

lemma outstanding_refs_takeWhile_opposite: "outstanding_refs P (takeWhile (Not  P) xs) = {}"   
apply (induct xs)
apply auto
done


lemma no_outstanding_volatile_Writesb_acquired:
  "outstanding_refs is_volatile_Writesb sb = {}  acquired False sb A = {}"
  apply (induct sb)
   apply simp
  subgoal for a sb
    by (case_tac a) auto
  done

lemma acquired_all_acquired:"pending_write A. acquired pending_write xs A  A  all_acquired xs"
  apply (induct xs)
  apply (auto split: memref.splits)
  done

lemma acquired_all_acquired_in: "x  acquired pending_write xs A  x  A  all_acquired xs"
  using acquired_all_acquired
  by blast



primrec sharing_consistent:: "shared  owns  'a memref list   bool"
where
"sharing_consistent 𝒮 𝒪 [] = True"
| "sharing_consistent 𝒮 𝒪 (r#rs) =
   (case r of
     Writesb volatile _ _ _ A L R W  
      (if volatile then A  dom 𝒮  𝒪  L  A  A  R = {}  R  𝒪  
                       sharing_consistent (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) rs
      else sharing_consistent 𝒮 𝒪 rs)  
   | Ghostsb A L R W   A  dom 𝒮  𝒪  L  A  A  R = {}  R  𝒪  
        sharing_consistent (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) rs
   | _  sharing_consistent 𝒮 𝒪 rs)"

lemma sharing_consistent_all_acquired:
  "𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb  all_acquired sb  dom 𝒮  𝒪"
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.hyps [where 𝒮="(𝒮 ⊕⇘WR ⊖⇘AL)" and 𝒪="(𝒪  A - R)"] Cons.prems
      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)       
    with Cons.hyps [where 𝒮="(𝒮 ⊕⇘WR ⊖⇘AL)" and 𝒪="(𝒪  A - R)"] Cons.prems show ?thesis by auto
  qed
qed

lemma sharing_consistent_append:
"𝒮 𝒪. sharing_consistent 𝒮 𝒪 (xs@ys) =
     (sharing_consistent 𝒮 𝒪 xs  
      sharing_consistent (share xs 𝒮) (acquired True xs 𝒪) ys)"
apply (induct xs)
apply (auto split: memref.splits)
done

primrec read_only_reads :: "owns  'a memref list  addr set"
where
"read_only_reads 𝒪 [] = {}"
| "read_only_reads 𝒪 (x#xs) =
  (case x of
     Readsb volatile a t v  (if ¬ volatile  a  𝒪 
                             then insert a (read_only_reads 𝒪 xs)
                             else read_only_reads 𝒪 xs)
   | Writesb volatile _ _ _ A L R W  
         (if volatile then read_only_reads (𝒪  A - R) xs  
          else read_only_reads 𝒪 xs )
   | Ghostsb A L R W  read_only_reads (𝒪  A - R) xs 
   | _  read_only_reads 𝒪 xs)"

lemma read_only_reads_append:
"𝒪. read_only_reads 𝒪 (xs@ys) = 
  read_only_reads 𝒪 xs  read_only_reads (acquired True xs 𝒪) ys"
  apply (induct xs)
   apply simp
  subgoal for a xs 𝒪
    by (case_tac a) auto
  done

lemma read_only_reads_antimono:
  "𝒪 𝒪'. 
  𝒪  𝒪'  read_only_reads 𝒪' sb  read_only_reads 𝒪 sb"
  apply (induct sb)
  apply  simp
  subgoal for a sb 𝒪 𝒪'
  apply (case_tac a)
  apply    (clarsimp split: if_split_asm)
           subgoal for volatile a1 D f v A L R W
           apply (drule_tac C="A" in union_mono_aux)
           apply (drule_tac C="R" in set_minus_mono_aux)
           apply blast
           done 
  apply   auto
  subgoal for A L R W x
  apply (drule_tac C="A" in union_mono_aux)
  apply (drule_tac C="R" in set_minus_mono_aux)
  apply blast
  done
  done
  done

primrec non_volatile_writes_unshared:: "shared  'a memref list  bool"
where
"non_volatile_writes_unshared 𝒮 [] = True"
| "non_volatile_writes_unshared 𝒮 (x#xs) =
  (case x of
    Writesb volatile a sop v A L R W  (if volatile then non_volatile_writes_unshared (𝒮 ⊕⇘WR ⊖⇘AL) xs
                                     else a  dom 𝒮  non_volatile_writes_unshared 𝒮 xs)
  | Ghostsb A L R W   non_volatile_writes_unshared (𝒮 ⊕⇘WR ⊖⇘AL) xs
  | _  non_volatile_writes_unshared 𝒮 xs)"


lemma non_volatile_writes_unshared_append: 
"𝒮. non_volatile_writes_unshared 𝒮 (xs@ys)
         = (non_volatile_writes_unshared 𝒮 xs  non_volatile_writes_unshared (share xs 𝒮) ys)"
apply (induct xs)
apply (auto split: memref.splits)
done

lemma non_volatile_writes_unshared_antimono:
"𝒮 𝒮'. dom 𝒮  dom 𝒮'  non_volatile_writes_unshared 𝒮' xs 
   non_volatile_writes_unshared 𝒮 xs"
  apply (induct xs)
  apply  simp
  subgoal for a xs 𝒮 𝒮'
  apply (case_tac a)
  apply     (clarsimp split: if_split_asm)
            subgoal for volatile a1 D f v A L R W
            apply (drule_tac C="R" in augment_mono_aux)
            apply (drule_tac C="L" in restrict_mono_aux)
            apply blast
            done
  apply    fastforce
  apply   fastforce
  apply  fastforce
  apply (clarsimp split: if_split_asm)
  subgoal for A L R W
  apply (drule_tac C="R" in augment_mono_aux)
  apply (drule_tac C="L" in restrict_mono_aux)
  apply blast
  done
  done
  done 

primrec  no_write_to_read_only_memory:: "shared  'a memref list  bool"
where
"no_write_to_read_only_memory 𝒮 [] = True"
| "no_write_to_read_only_memory 𝒮 (x#xs) =
  (case x of
    Writesb volatile a sop v A L R W  a  read_only 𝒮 
                                      (if volatile then no_write_to_read_only_memory (𝒮 ⊕⇘WR ⊖⇘AL) xs
                                       else no_write_to_read_only_memory 𝒮 xs)
  | Ghostsb A L R W   no_write_to_read_only_memory (𝒮 ⊕⇘WR ⊖⇘AL) xs
  | _  no_write_to_read_only_memory 𝒮 xs)"

lemma no_write_to_read_only_memory_append: 
"𝒮. no_write_to_read_only_memory 𝒮 (xs@ys)
         = (no_write_to_read_only_memory 𝒮 xs  no_write_to_read_only_memory (share xs 𝒮) ys)"
apply (induct xs)
apply  simp
subgoal for a xs 𝒮
  by (case_tac a) auto
done

lemma no_write_to_read_only_memory_antimono:
"𝒮 𝒮'. 𝒮 s 𝒮'  no_write_to_read_only_memory 𝒮' xs 
   no_write_to_read_only_memory 𝒮 xs"
  apply (induct xs)
  apply  simp
  subgoal for a xs 𝒮 𝒮' 
  apply (case_tac a)
  apply    (clarsimp split: if_split_asm) 
             subgoal for volatile a1 D f v A L R W 
             apply (frule_tac C="R" and x="W" in augment_mono_map)
             apply (drule_tac A="𝒮 ⊕⇘WR" and C="L" and x="A" in restrict_mono_map)
             apply (fastforce dest: read_only_mono shared_leD)
             done
  apply    (fastforce dest: read_only_mono shared_leD)
  apply   fastforce
  apply  fastforce
  apply (clarsimp)
  subgoal for A L R W
  apply (frule_tac C="R" and x="W" in augment_mono_map)
  apply (drule_tac A="𝒮 ⊕⇘WR" and C="L" and x="A" in restrict_mono_map)
  apply (fastforce dest: read_only_mono shared_leD)
  done
  done
  done

locale outstanding_non_volatile_refs_owned_or_read_only =
fixes 𝒮::shared
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes outstanding_non_volatile_refs_owned_or_read_only:
  "i is 𝒪  𝒟 θ sb p. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)  
   
   non_volatile_owned_or_read_only False 𝒮 𝒪 sb"

locale outstanding_volatile_writes_unowned_by_others =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes outstanding_volatile_writes_unowned_by_others: 
  "i pi isi 𝒪i i 𝒟i θi sbi j pj isj 𝒪j j 𝒟j θj sbj. 
   i < length ts; j < length ts; ij;
    ts!i = (pi,isi,θi,sbi,𝒟i,𝒪i,i); ts!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)
   
   
    (𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sbi = {}"


locale read_only_reads_unowned =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes read_only_reads_unowned: 
  "i pi isi 𝒪i i 𝒟i θi sbi j pj isj 𝒪j j 𝒟j θj sbj. 
   i < length ts; j < length ts; ij;
    ts!i = (pi,isi,θi,sbi,𝒟i,𝒪i,i); ts!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)
   
   
    (𝒪j  all_acquired sbj)  
     read_only_reads (acquired True 
                          (takeWhile (Not  is_volatile_Writesb) sbi) 𝒪i) 
                          (dropWhile (Not  is_volatile_Writesb) sbi) = {}"

locale ownership_distinct =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes ownership_distinct:
   "i j pi isi 𝒪i i 𝒟i θi sbi pj isj 𝒪j j 𝒟j θj sbj. 
      i < length ts; j < length ts; i  j; 
    ts!i = (pi,isi,θi,sbi,𝒟i,𝒪i,i); ts!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)
        (𝒪i  all_acquired sbi)  (𝒪j  all_acquired sbj) = {}"


locale valid_ownership = 
  outstanding_non_volatile_refs_owned_or_read_only + 
  outstanding_volatile_writes_unowned_by_others + 
  read_only_reads_unowned +
  ownership_distinct

locale outstanding_non_volatile_writes_unshared =
fixes 𝒮::shared and ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes outstanding_non_volatile_writes_unshared:
  "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)  
   
   non_volatile_writes_unshared 𝒮 sb"


locale sharing_consis =
fixes 𝒮::shared and ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes sharing_consis:
  "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)  
   
   sharing_consistent 𝒮 𝒪 sb"




locale no_outstanding_write_to_read_only_memory =
fixes 𝒮::shared and ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes no_outstanding_write_to_read_only_memory:
  "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)  
   
   no_write_to_read_only_memory 𝒮 sb"


locale valid_sharing = 
  outstanding_non_volatile_writes_unshared +
  sharing_consis +
  read_only_unowned +
  unowned_shared +
  no_outstanding_write_to_read_only_memory

locale valid_ownership_and_sharing = valid_ownership +
  outstanding_non_volatile_writes_unshared +
  sharing_consis +
  no_outstanding_write_to_read_only_memory


lemma (in read_only_reads_unowned)
  read_only_reads_unowned_nth_update:
 "i p is 𝒪  𝒟  θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,); 
     read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb') 𝒪') 
       (dropWhile (Not  is_volatile_Writesb) sb')  read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪) 
       (dropWhile (Not  is_volatile_Writesb) sb);
     𝒪'  all_acquired sb'  𝒪  all_acquired sb  
     read_only_reads_unowned (ts[i := (p',is',θ',sb',𝒟',𝒪',ℛ')])"
  apply (unfold_locales)
  apply (clarsimp simp add: nth_list_update split: if_split_asm)
  apply   (fastforce dest: read_only_reads_unowned)+
  done

 
lemma outstanding_non_volatile_refs_owned_or_read_only_tl: 
  "outstanding_non_volatile_refs_owned_or_read_only 𝒮 (t#ts)  outstanding_non_volatile_refs_owned_or_read_only 𝒮 ts"
  by (force simp add: outstanding_non_volatile_refs_owned_or_read_only_def)

lemma outstanding_volatile_writes_unowned_by_others_tl: 
  "outstanding_volatile_writes_unowned_by_others (t#ts)  outstanding_volatile_writes_unowned_by_others ts"
  apply (clarsimp simp add: outstanding_volatile_writes_unowned_by_others_def)
  apply fastforce
  done


lemma read_only_reads_unowned_tl: 
  "read_only_reads_unowned  (t # ts) 
    read_only_reads_unowned  (ts)"
  apply (clarsimp simp add: read_only_reads_unowned_def)
  apply fastforce
  done



lemma ownership_distinct_tl:
  assumes dist: "ownership_distinct (t#ts)" 
  shows "ownership_distinct ts"
proof -
  from dist
  interpret ownership_distinct "t#ts" .
  
  show ?thesis
  proof (rule ownership_distinct.intro)
    fix i j p "is" 𝒪  𝒟 xs sb p' is' 𝒪' ℛ' 𝒟' xs' sb'
    assume i_bound: "i < length ts" 
      and j_bound: "j < length ts" 
      and neq: "i  j" 
      and ith: "ts ! i = (p,is,xs,sb,𝒟,𝒪,)"
      and jth: "ts ! j = (p',is', xs', sb',𝒟', 𝒪',ℛ')"
    from i_bound j_bound neq ith jth
    show "(𝒪  all_acquired sb)  (𝒪'  all_acquired sb') = {}"
      by - (rule ownership_distinct [of "Suc i" "Suc j"],auto)
  qed
qed

lemma valid_ownership_tl: "valid_ownership 𝒮 (t#ts)  valid_ownership 𝒮 ts"
  by (auto simp add: valid_ownership_def 
    intro: outstanding_volatile_writes_unowned_by_others_tl 
    outstanding_non_volatile_refs_owned_or_read_only_tl ownership_distinct_tl
    read_only_reads_unowned_tl)


lemma sharing_consistent_takeWhile: 
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  shows "sharing_consistent 𝒮 𝒪 (takeWhile P sb)"
proof -
  from consis have "sharing_consistent 𝒮 𝒪 (takeWhile P sb @ dropWhile P sb)"
    by simp
  with sharing_consistent_append [of _ _ "takeWhile P sb" "dropWhile P sb"] 
  show ?thesis
    by simp
qed
    
lemma sharing_consis_tl: "sharing_consis 𝒮 (t#ts)  sharing_consis 𝒮 ts"
  by (auto simp add: sharing_consis_def)

lemma sharing_consis_Cons: 
  "sharing_consis 𝒮 ts; sharing_consistent 𝒮 𝒪 sb
    sharing_consis 𝒮 ((p,is,θ,sb,𝒟,𝒪,)#ts)"
  apply (clarsimp simp add: sharing_consis_def)
  subgoal for i pa isa 𝒪' ℛ' 𝒟' θ' sba
    by (case_tac i) auto
  done

lemma outstanding_non_volatile_writes_unshared_tl:
  "outstanding_non_volatile_writes_unshared 𝒮 (t#ts)  
  outstanding_non_volatile_writes_unshared 𝒮 ts"
  by (auto simp add: outstanding_non_volatile_writes_unshared_def)

lemma no_outstanding_write_to_read_only_memory_tl:
  "no_outstanding_write_to_read_only_memory 𝒮 (t#ts) 
  no_outstanding_write_to_read_only_memory 𝒮 ts"
  by (auto simp add: no_outstanding_write_to_read_only_memory_def)

lemma valid_ownership_and_sharing_tl: 
  "valid_ownership_and_sharing 𝒮 (t#ts)  valid_ownership_and_sharing 𝒮 ts"
  apply (clarsimp simp add: valid_ownership_and_sharing_def)
  apply (auto intro: valid_ownership_tl
    outstanding_non_volatile_writes_unshared_tl
    no_outstanding_write_to_read_only_memory_tl
    sharing_consis_tl)
  done


lemma non_volatile_owned_or_read_only_outstanding_non_volatile_writes: 
  "𝒪 𝒮 pending_write. non_volatile_owned_or_read_only pending_write 𝒮 𝒪 sb
   
  outstanding_refs is_non_volatile_Writesb sb  𝒪  all_acquired 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
      from Cons.hyps [of True "(𝒮 ⊕⇘WR ⊖⇘AL)" "(𝒪  A - R)"] Cons.prems
      show ?thesis
	by (auto 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.hyps [of pending_write "(𝒮 ⊕⇘WR ⊖⇘AL)" "(𝒪  A - R)"] Cons.prems
    show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed

lemma (in outstanding_non_volatile_refs_owned_or_read_only) outstanding_non_volatile_writes_owned:
  assumes i_bound: "i < length ts" 
  assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)"
  shows "outstanding_refs is_non_volatile_Writesb sb  𝒪  all_acquired sb"
using non_volatile_owned_or_read_only_outstanding_non_volatile_writes [OF outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ts_i]]
by blast







lemma non_volatile_reads_acquired_or_read_only: 
  "𝒪 𝒮. non_volatile_owned_or_read_only True 𝒮 𝒪 sb; sharing_consistent 𝒮 𝒪 sb
   
  outstanding_refs is_non_volatile_Readsb sb  𝒪  all_acquired sb  read_only 𝒮"
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 non_vol: "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R)  sb" and 
	A_shared_onws: "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"
	by (clarsimp simp add: Writesb True )

      from Cons.hyps [OF non_vol consis']
      have hyp: "outstanding_refs is_non_volatile_Readsb sb
                  𝒪  A - R  all_acquired sb  read_only (𝒮 ⊕⇘WR ⊖⇘AL)".
      with R_owns A_R L_A
      show ?thesis
	apply (clarsimp simp add: Writesb True )
	apply (drule (1) rev_subsetD)
	apply (auto simp add: in_read_only_convs split: if_split_asm)
	done
    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 non_vol: "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R)  sb" and 
      A_shared_onws: "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"
      by (clarsimp simp add: Ghostsb )

    from Cons.hyps [OF non_vol consis']
    have hyp: "outstanding_refs is_non_volatile_Readsb sb
       𝒪  A - R  all_acquired sb  read_only (𝒮 ⊕⇘WR ⊖⇘AL)".
    with R_owns A_R L_A
    show ?thesis
      apply (clarsimp simp add: Ghostsb )
      apply (drule (1) rev_subsetD)
      apply (auto simp add: in_read_only_convs split: if_split_asm)
      done
  qed
qed


lemma non_volatile_reads_acquired_or_read_only_reads: 
  "𝒪 𝒮 pending_write. non_volatile_owned_or_read_only pending_write 𝒮 𝒪 sb
   
  outstanding_refs is_non_volatile_Readsb sb  𝒪  all_acquired sb  read_only_reads 𝒪 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

      from Cons.prems obtain non_vol: "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R)  sb" 
	by (clarsimp simp add: Writesb True )

      from Cons.hyps [OF non_vol ]
      have hyp: "outstanding_refs is_non_volatile_Readsb sb
                  𝒪  A - R  all_acquired sb  read_only_reads (𝒪  A - R) sb".
      then 
      show ?thesis
	by (auto 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 non_vol: "non_volatile_owned_or_read_only pending_write (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R)  sb" 
      by (clarsimp simp add: Ghostsb )

    from Cons.hyps [OF non_vol ]
    have hyp: "outstanding_refs is_non_volatile_Readsb sb
                  𝒪  A - R  all_acquired sb  read_only_reads (𝒪  A - R) sb".
    then 
    show ?thesis
      by (auto simp add: Ghostsb )
  qed
qed




lemma non_volatile_owned_or_read_only_outstanding_refs: 
  "𝒪 𝒮 pending_write. non_volatile_owned_or_read_only pending_write 𝒮 𝒪 sb
   
  outstanding_refs (Not  is_volatile) sb  𝒪  all_acquired sb  read_only_reads 𝒪 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
      from Cons.hyps [of True "(𝒮 ⊕⇘WR ⊖⇘AL)" "(𝒪  A - R)"] Cons.prems
      show ?thesis
	by (auto 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.hyps [of pending_write "(𝒮 ⊕⇘WR ⊖⇘AL)" "(𝒪  A - R)"] Cons.prems
    show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed



lemma no_unacquired_write_to_read_only:
"𝒮 𝒪. no_write_to_read_only_memory 𝒮 sb;sharing_consistent 𝒮 𝒪 sb;
 a  read_only 𝒮; a  (𝒪  all_acquired sb) 
        a  outstanding_refs is_Writesb 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

      from Cons.prems obtain no_wrt: "no_write_to_read_only_memory (𝒮 ⊕⇘WR ⊖⇘AL) sb" and 
	A_shared_onws: "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_ro: "a  read_only 𝒮" and
	a_A: "a  A" and a_all_acq: "a  all_acquired sb" and a_owns: "a  𝒪" and 
	a'_notin: "a'  read_only 𝒮" 
	by ( simp add: Writesb True )
      
      from a'_notin a_ro have neq_a_a': "aa'"
	by blast

      from a_A a_all_acq a_owns
      have a_notin': "a  𝒪  A - R  all_acquired sb"
	by auto
      from a_ro L_A a_A R_owns a_owns
      have "a  read_only (𝒮 ⊕⇘WR ⊖⇘AL)"
	by (auto simp add: in_read_only_convs split: if_split_asm)

      from Cons.hyps [OF no_wrt consis' this a_notin']
      have "a  outstanding_refs is_Writesb sb".
      with neq_a_a'
      show ?thesis
	by (clarsimp simp add: Writesb True)
    next
      case False with Cons
      show ?thesis
	by (auto simp add: Writesb False)
    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 no_wrt: "no_write_to_read_only_memory (𝒮 ⊕⇘WR ⊖⇘AL) sb" and 
	A_shared_onws: "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_ro: "a  read_only 𝒮" and
	a_A: "a  A" and a_all_acq: "a  all_acquired sb" and a_owns: "a  𝒪" 
	by ( simp add: Ghostsb )
      
      
      from a_A a_all_acq a_owns
      have a_notin': "a  𝒪  A - R  all_acquired sb"
	by auto
      from a_ro L_A a_A R_owns a_owns
      have "a  read_only (𝒮 ⊕⇘WR ⊖⇘AL)"
	by (auto simp add: in_read_only_convs split: if_split_asm)

      from Cons.hyps [OF no_wrt consis' this a_notin']
      have "a  outstanding_refs is_Writesb sb".
      then
      show ?thesis
	by (clarsimp simp add: Ghostsb)
  qed
qed

lemma read_only_reads_read_only: 
  "𝒮 𝒪. non_volatile_owned_or_read_only True 𝒮 𝒪 sb;
  sharing_consistent 𝒮 𝒪 sb
  
  read_only_reads 𝒪 sb  𝒪  all_acquired sb  read_only 𝒮"
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 non_vol: "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R)  sb" and 
	A_shared_onws: "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"
	by (clarsimp simp add: Writesb True )

      from Cons.hyps [OF non_vol consis']
      have hyp: "read_only_reads (𝒪  A - R) sb
                  𝒪  A - R  all_acquired sb  read_only (𝒮 ⊕⇘WR ⊖⇘AL)".

      {
	fix a'
	assume a'_in: "a'  read_only_reads (𝒪  A - R) sb"
	assume a'_unowned: "a'  𝒪"
	assume a'_unacq: "a'  all_acquired sb"
	assume a'_A: "a'  A"
	have "a'  read_only 𝒮"
	proof -
	  from a'_in hyp a'_unowned a'_unacq a'_A 
	  have "a'  read_only (𝒮 ⊕⇘WR ⊖⇘AL)"
	    by auto
	  
	  with L_A R_owns a'_unowned
	  show ?thesis
	    by (auto simp add: in_read_only_convs split:if_split_asm)
	qed
      }
	
      then
	  
      show ?thesis
	apply (clarsimp simp add: Writesb True simp del: o_apply)
	apply force
	done
    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 non_vol: "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R)  sb" and 
      A_shared_onws: "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"
      by (clarsimp simp add: Ghostsb )

    from Cons.hyps [OF non_vol consis']
    have hyp: "read_only_reads (𝒪  A - R) sb
                  𝒪  A - R  all_acquired sb  read_only (𝒮 ⊕⇘WR ⊖⇘AL)".

    {
      fix a'
      assume a'_in: "a'  read_only_reads (𝒪  A - R) sb"
      assume a'_unowned: "a'  𝒪"
      assume a'_unacq: "a'  all_acquired sb"
      assume a'_A: "a'  A"
      have "a'  read_only 𝒮"
      proof -
	from a'_in hyp a'_unowned a'_unacq a'_A 
	have "a'  read_only (𝒮 ⊕⇘WR ⊖⇘AL)"
	  by auto
	  
	with L_A R_owns a'_unowned
	show ?thesis
	  by (auto simp add: in_read_only_convs split:if_split_asm)
        qed
    }
	
    then
	  
    show ?thesis
      apply (clarsimp simp add: Ghostsb simp del: o_apply)
      apply force
      done
    
  qed
qed

lemma no_unacquired_write_to_read_only_reads:
"𝒮 𝒪 . no_write_to_read_only_memory 𝒮 sb;
non_volatile_owned_or_read_only True 𝒮 𝒪 sb; sharing_consistent 𝒮 𝒪 sb;
 a  read_only_reads 𝒪 sb; a  (𝒪  all_acquired sb) 
        a  outstanding_refs is_Writesb 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

      from Cons.prems obtain no_wrt: "no_write_to_read_only_memory (𝒮 ⊕⇘WR ⊖⇘AL) sb" and 
	non_vol: "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R)  sb" and 
	A_shared_onws: "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_ro: "a  read_only_reads (𝒪  A - R) sb" and
	a_A: "a  A" and a_all_acq: "a  all_acquired sb" and a_owns: "a  𝒪" and 
	a'_notin: "a'  read_only 𝒮" 
	by ( simp add: Writesb True )

      from read_only_reads_read_only [OF non_vol consis' ] a_ro a_owns a_all_acq a_A
      have "a  read_only (𝒮 ⊕⇘WR ⊖⇘AL)" 
	by auto
      with a'_notin R_owns a_owns have neq_a_a': "aa'"
	by (auto simp add:  in_read_only_convs split: if_split_asm)
      

      from a_A a_all_acq a_owns
      have a_notin': "a  𝒪  A - R  all_acquired sb"
	by auto

      from Cons.hyps [OF no_wrt non_vol consis' a_ro a_notin'] 
      have "a  outstanding_refs is_Writesb sb".
      then 
      show ?thesis
	using neq_a_a'
	by (auto simp add: Writesb True)
    next
      case False with Cons
      show ?thesis
	by (auto simp add: Writesb False)
    qed
  next
    case (Readsb volatile a' t v) 
    show ?thesis
    proof (cases volatile)
      case True
      with Cons show ?thesis
	by  (auto simp add: Readsb)
    next
      case False
      note non_volatile = this
      from Cons.prems obtain no_wrt': "no_write_to_read_only_memory 𝒮 sb" and 
	consis':"sharing_consistent 𝒮 𝒪 sb" and
	a_in: "a  (if a'  𝒪 then insert a' (read_only_reads 𝒪 sb)
                 else read_only_reads 𝒪 sb)" and
	a'_owns_shared: "a'  𝒪  a'  read_only 𝒮" and 
	non_vol': "non_volatile_owned_or_read_only True 𝒮 𝒪 sb" and
        a_owns: "a  𝒪  all_acquired sb"
	by (clarsimp simp add: Readsb False)

      show ?thesis
      proof (cases "a'  𝒪")
	case True
	with a_in have "a  read_only_reads 𝒪 sb"
	  by auto
	from Cons.hyps [OF no_wrt' non_vol' consis' this a_owns]
	show ?thesis
	  by (clarsimp simp add: Readsb)
      next
	case False
	note a'_unowned = this
	with a_in have a_in': "a  insert a' (read_only_reads 𝒪 sb)" by auto
	from a'_owns_shared False have a'_read_only: "a'  read_only 𝒮" by auto
	show ?thesis
	proof (cases "a=a'")
	  case False
	  with a_in' have "a  (read_only_reads 𝒪 sb)" by auto
	  from Cons.hyps [OF no_wrt' non_vol' consis' this a_owns]
	  show ?thesis
	    by (simp add: Readsb)
	next
	  case True
	  from no_unacquired_write_to_read_only [OF no_wrt' consis' a'_read_only] a_owns True
	  
	  have "a'  outstanding_refs is_Writesb sb"
	    by auto
	  then show ?thesis
	    by (simp add: Readsb True)
	qed
      qed
    qed
  next
    case Progsb with Cons
    show ?thesis
      by (auto)
  next
    case (Ghostsb A L R W)
    from Cons.prems obtain no_wrt: "no_write_to_read_only_memory (𝒮 ⊕⇘WR ⊖⇘AL) sb" and 
      non_vol: "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R)  sb" and 
      A_shared_onws: "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_ro: "a  read_only_reads (𝒪  A - R) sb" and
      a_A: "a  A" and a_all_acq: "a  all_acquired sb" and a_owns: "a  𝒪" 
      by ( simp add: Ghostsb )

    from read_only_reads_read_only [OF non_vol consis' ] a_ro a_owns a_all_acq a_A
    have "a  read_only (𝒮 ⊕⇘WR ⊖⇘AL)" 
      by auto
      

    from a_A a_all_acq a_owns
    have a_notin': "a  𝒪  A - R  all_acquired sb"
      by auto

    from Cons.hyps [OF no_wrt non_vol consis' a_ro a_notin'] 
    have "a  outstanding_refs is_Writesb sb".
    then 
    show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed


lemma no_unacquired_write_to_read_only'':
  assumes no_wrt: "no_write_to_read_only_memory 𝒮 sb"
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  shows "read_only 𝒮  outstanding_refs is_Writesb sb  𝒪  all_acquired sb"
using no_unacquired_write_to_read_only [OF no_wrt consis]
by auto

lemma no_unacquired_volatile_write_to_read_only:
  assumes no_wrt: "no_write_to_read_only_memory 𝒮 sb"
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  shows "read_only 𝒮  outstanding_refs is_volatile_Writesb sb  𝒪  all_acquired sb"
proof -
  have "outstanding_refs is_volatile_Writesb sb  outstanding_refs is_Writesb sb"
    apply (rule outstanding_refs_mono_pred)
    apply (auto simp add: is_volatile_Writesb_def split: memref.splits)
    done
  with no_unacquired_write_to_read_only'' [OF no_wrt consis]
  show ?thesis by blast
qed

lemma no_unacquired_non_volatile_write_to_read_only_reads:
  assumes no_wrt: "no_write_to_read_only_memory 𝒮 sb"
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  shows "read_only 𝒮  outstanding_refs is_non_volatile_Writesb sb  𝒪  all_acquired sb"
proof -
  from outstanding_refs_subsets 
  have "outstanding_refs is_non_volatile_Writesb sb  outstanding_refs is_Writesb sb" by - assumption
  with no_unacquired_write_to_read_only'' [OF no_wrt consis]
  show ?thesis by blast
qed


lemma no_unacquired_write_to_read_only_reads':
  assumes no_wrt: "no_write_to_read_only_memory 𝒮 sb"
  assumes non_vol: "non_volatile_owned_or_read_only True 𝒮 𝒪 sb"
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  shows "read_only_reads 𝒪 sb  outstanding_refs is_Writesb sb  𝒪  all_acquired sb"
using no_unacquired_write_to_read_only_reads [OF no_wrt non_vol consis]
by auto

lemma no_unacquired_volatile_write_to_read_only_reads:
  assumes no_wrt: "no_write_to_read_only_memory 𝒮 sb"
  assumes non_vol: "non_volatile_owned_or_read_only True 𝒮 𝒪 sb"
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  shows "read_only_reads 𝒪 sb  outstanding_refs is_volatile_Writesb sb  𝒪  all_acquired sb"
proof -
  have "outstanding_refs is_volatile_Writesb sb  outstanding_refs is_Writesb sb"
    apply (rule outstanding_refs_mono_pred)
    apply (auto simp add: is_volatile_Writesb_def split: memref.splits)
    done
  with no_unacquired_write_to_read_only_reads [OF no_wrt non_vol consis]
  show ?thesis by blast
qed

lemma no_unacquired_non_volatile_write_to_read_only:
  assumes no_wrt: "no_write_to_read_only_memory 𝒮 sb"
  assumes non_vol: "non_volatile_owned_or_read_only True 𝒮 𝒪 sb"
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  shows "read_only_reads 𝒪 sb  outstanding_refs is_non_volatile_Writesb sb  𝒪  all_acquired sb"
proof -
  from outstanding_refs_subsets 
  have "outstanding_refs is_non_volatile_Writesb sb  outstanding_refs is_Writesb sb" by - assumption
  with no_unacquired_write_to_read_only_reads [OF no_wrt non_vol consis]
  show ?thesis by blast
qed



lemma set_dropWhileD: "x  set (dropWhile P xs)  x  set xs"
  by (induct xs) (auto split: if_split_asm)

lemma outstanding_refs_takeWhileD:
  "x  outstanding_refs P (takeWhile P' sb)  x  outstanding_refs P sb"
  using outstanding_refs_takeWhile
  by blast

lemma outstanding_refs_dropWhileD:
  "x  outstanding_refs P (dropWhile P' sb)  x  outstanding_refs P sb"
  by (auto dest: set_dropWhileD simp add: outstanding_refs_conv)



lemma dropWhile_ConsD: "dropWhile P xs = y#ys  ¬ P y"
  by (simp add: dropWhile_eq_Cons_conv)


lemma non_volatile_owned_or_read_only_drop:
  "non_volatile_owned_or_read_only False 𝒮 𝒪 sb
   non_volatile_owned_or_read_only True 
      (share (takeWhile (Not  is_volatile_Writesb) sb) 𝒮) 
      (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪) 
      (dropWhile (Not  is_volatile_Writesb) sb)"
using non_volatile_owned_or_read_only_append [of False 𝒮 𝒪 "(takeWhile (Not  is_volatile_Writesb) sb)" 
  "(dropWhile (Not  is_volatile_Writesb) sb)"]
apply (cases "outstanding_refs is_volatile_Writesb sb = {}")
apply  (clarsimp simp add: outstanding_vol_write_take_drop_appends 
  takeWhile_not_vol_write_outstanding_refs dropWhile_not_vol_write_empty)
apply(clarsimp simp add: outstanding_vol_write_take_drop_appends 
  takeWhile_not_vol_write_outstanding_refs dropWhile_not_vol_write_empty )
apply (case_tac "(dropWhile (Not  is_volatile_Writesb) sb)")
apply  (fastforce simp add: outstanding_refs_conv)
apply (frule dropWhile_ConsD)
apply (clarsimp split: memref.splits)
done


lemma  read_only_share: "𝒮 𝒪. 
  sharing_consistent 𝒮 𝒪 sb  
        read_only (share sb 𝒮)  read_only 𝒮  𝒪  all_acquired 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
      from Cons.prems obtain 
	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"
	by (clarsimp simp add: Writesb True )
      from Cons.hyps [OF  consis']
      have "read_only (share sb (𝒮 ⊕⇘WR ⊖⇘AL))
               read_only (𝒮 ⊕⇘WR ⊖⇘AL)  (𝒪  A - R)  all_acquired sb"
        by auto
      also from A_shared_owns L_A R_owns A_R
      have "read_only (𝒮 ⊕⇘WR ⊖⇘AL)  (𝒪  A - R)  all_acquired sb 
        read_only 𝒮  𝒪  (A  all_acquired sb)"
        by (auto simp add: read_only_def augment_shared_def restrict_shared_def split: option.splits)
      finally
      show ?thesis
        by (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_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"
      by (clarsimp simp add: Ghostsb )
    from Cons.hyps [OF  consis']
    have "read_only (share sb (𝒮 ⊕⇘WR ⊖⇘AL))
               read_only (𝒮 ⊕⇘WR ⊖⇘AL)  (𝒪  A - R)  all_acquired sb"
      by auto
    also from A_shared_owns L_A R_owns A_R
    have "read_only (𝒮 ⊕⇘WR ⊖⇘AL)  (𝒪  A - R)  all_acquired sb 
        read_only 𝒮  𝒪  (A  all_acquired sb)"
      by (auto simp add: read_only_def augment_shared_def restrict_shared_def split: option.splits)
    finally
    show ?thesis
      by (simp add: Ghostsb)
  qed
qed


      
lemma (in valid_ownership_and_sharing) outstanding_non_write_non_vol_reads_drop_disj:
assumes i_bound: "i < length ts"
assumes j_bound: "j < length ts"
assumes neq_i_j: "i  j"
assumes ith: "ts!i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
assumes jth: "ts!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
shows "outstanding_refs is_Writesb (dropWhile (Not  is_volatile_Writesb) sbi)  
         outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) sbj)
         = {}" 
proof -

  let ?take_j = "(takeWhile (Not  is_volatile_Writesb) sbj)"
  let ?drop_j = "(dropWhile (Not  is_volatile_Writesb) sbj)"

  let ?take_i = "(takeWhile (Not  is_volatile_Writesb) sbi)"
  let ?drop_i = "(dropWhile (Not  is_volatile_Writesb) sbi)"


  note nvo_i = outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ith]
  note nvo_j = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound jth]
  note nro_i = no_outstanding_write_to_read_only_memory [OF i_bound ith]
  with no_write_to_read_only_memory_append [of 𝒮 ?take_i ?drop_i]
  have nro_drop_i: "no_write_to_read_only_memory (share ?take_i 𝒮) ?drop_i"
    by simp
  note nro_j = no_outstanding_write_to_read_only_memory [OF j_bound jth]
  with no_write_to_read_only_memory_append [of 𝒮 ?take_j ?drop_j]
  have nro_drop_j: "no_write_to_read_only_memory (share ?take_j 𝒮) ?drop_j"
    by simp
  from outstanding_volatile_writes_unowned_by_others [OF i_bound j_bound neq_i_j ith jth]
  have dist: "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sbi = {}".
  note own_dist = ownership_distinct [OF i_bound j_bound neq_i_j ith jth]




  from sharing_consis [OF j_bound jth]
  have consis_j: "sharing_consistent 𝒮 𝒪j sbj".
  with sharing_consistent_append [of 𝒮 𝒪j ?take_j ?drop_j]
  obtain
    consis_take_j: "sharing_consistent 𝒮 𝒪j ?take_j" and
    consis_drop_j: "sharing_consistent (share ?take_j 𝒮) (acquired True ?take_j 𝒪j) ?drop_j"
    by simp
    
  from sharing_consis [OF i_bound ith]
  have consis_i: "sharing_consistent 𝒮 𝒪i sbi".
  with sharing_consistent_append [of 𝒮 𝒪i ?take_i ?drop_i]
  have consis_drop_i: "sharing_consistent (share ?take_i 𝒮) (acquired True ?take_i 𝒪i) ?drop_i"
    by simp


  {
    fix x
    assume x_in_drop_i: "x  outstanding_refs is_Writesb ?drop_i"
    assume x_in_drop_j: "x  outstanding_refs is_non_volatile_Readsb ?drop_j"
    have False
    proof -
      from x_in_drop_i have x_in_i: "x  outstanding_refs is_Writesb sbi"
	using outstanding_refs_append [of is_Writesb ?take_i ?drop_i] by auto

      from x_in_drop_j have x_in_j: "x  outstanding_refs is_non_volatile_Readsb sbj"
	using outstanding_refs_append [of is_non_volatile_Readsb ?take_j ?drop_j]
	by auto
      from non_volatile_owned_or_read_only_drop [OF nvo_j]
      have nvo_drop_j: "non_volatile_owned_or_read_only True (share ?take_j 𝒮) (acquired True ?take_j 𝒪j) ?drop_j".

      from non_volatile_reads_acquired_or_read_only_reads [OF nvo_drop_j ] x_in_drop_j 
        acquired_takeWhile_non_volatile_Writesb [of sbj 𝒪j]
      have x_j: "x  𝒪j  all_acquired sbj  read_only_reads (acquired True ?take_j 𝒪j) ?drop_j" 
	using all_acquired_append [of ?take_j ?drop_j]
	by ( auto )

      {
	assume x_in_vol_drop_i: "x  outstanding_refs is_volatile_Writesb ?drop_i"
	hence x_in_vol_i: "x  outstanding_refs is_volatile_Writesb sbi"
	  using outstanding_refs_append [of is_volatile_Writesb ?take_i ?drop_i]
	  by auto

	from outstanding_volatile_writes_unowned_by_others [OF i_bound j_bound neq_i_j ith jth]
	have "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sbi = {}". 
	with x_in_vol_i x_j obtain  
	  x_unacq_j: "x   𝒪j  all_acquired sbj" and
	  x_ror_j: "x  read_only_reads (acquired True ?take_j 𝒪j) ?drop_j"
	  by auto
	from read_only_reads_unowned [OF j_bound i_bound neq_i_j [symmetric] jth ith] x_ror_j
	have "x  𝒪i  all_acquired sbi"
	  by auto

	moreover 


	from read_only_reads_read_only [OF nvo_drop_j  consis_drop_j] x_ror_j x_unacq_j 
	  all_acquired_append [of ?take_j ?drop_j] acquired_takeWhile_non_volatile_Writesb [of sbj 𝒪j]
	have "x  read_only (share ?take_j 𝒮)"
	  by (auto)
        
        from read_only_share [OF consis_take_j] this x_unacq_j all_acquired_append [of ?take_j ?drop_j]
        have "x  read_only 𝒮"
	  by auto

	with no_unacquired_write_to_read_only'' [OF nro_i consis_i]  x_in_i
	have "x  𝒪i  all_acquired sbi"
	  by auto

	ultimately have False by auto
      }
      moreover
      {
	assume x_in_non_vol_drop_i: "x  outstanding_refs is_non_volatile_Writesb ?drop_i"
	hence "x  outstanding_refs is_non_volatile_Writesb sbi"
	  using outstanding_refs_append [of is_non_volatile_Writesb ?take_i ?drop_i]
	  by auto
	with non_volatile_owned_or_read_only_outstanding_non_volatile_writes [OF nvo_i]
	have "x  𝒪i  all_acquired sbi" by auto

	moreover

	with x_j own_dist obtain
	  x_unacq_j: "x   𝒪j  all_acquired sbj" and
	  x_ror_j: "x  read_only_reads (acquired True ?take_j 𝒪j) ?drop_j"
	  by auto
	from read_only_reads_unowned [OF j_bound i_bound neq_i_j [symmetric] jth ith] x_ror_j
	have "x  𝒪i  all_acquired sbi"
	  by auto

	ultimately have False
	  by auto
      }
      ultimately
      
      show ?thesis
	using x_in_drop_i x_in_drop_j
	by (auto simp add: misc_outstanding_refs_convs)
    qed
  }
  thus ?thesis
    by auto
qed
  
lemma (in valid_ownership_and_sharing) outstanding_non_volatile_write_disj:
assumes i_bound: "i < length ts"
assumes j_bound: "j < length ts"
assumes neq_i_j: "i  j"
assumes ith: "ts!i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
assumes jth: "ts!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
shows "outstanding_refs (is_non_volatile_Writesb) (takeWhile (Not  is_volatile_Writesb) sbi)  
        (outstanding_refs is_volatile_Writesb sbj   
         outstanding_refs is_non_volatile_Writesb sbj  
         outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) sbj)  
         (outstanding_refs is_non_volatile_Readsb (takeWhile (Not  is_volatile_Writesb) sbj) - 
          read_only_reads 𝒪j (takeWhile (Not  is_volatile_Writesb) sbj))  
         (𝒪j  all_acquired (takeWhile (Not  is_volatile_Writesb) sbj))
        ) = {}" (is "?non_vol_writes_i  ?not_volatile_j = {}")
proof -
  note nro_i = no_outstanding_write_to_read_only_memory [OF i_bound ith]
  note nro_j = no_outstanding_write_to_read_only_memory [OF j_bound jth]
  note nvo_j = outstanding_non_volatile_refs_owned_or_read_only [OF j_bound jth]
  note nvo_i = outstanding_non_volatile_refs_owned_or_read_only [OF i_bound ith]
    
  from outstanding_volatile_writes_unowned_by_others [OF i_bound j_bound neq_i_j ith jth]
  have dist: "(𝒪j  all_acquired sbj)  outstanding_refs is_volatile_Writesb sbi = {}".
  
  from outstanding_volatile_writes_unowned_by_others [OF j_bound i_bound neq_i_j [symmetric] jth ith]
  have dist_j: "(𝒪i  all_acquired sbi)  outstanding_refs is_volatile_Writesb sbj = {}".
  
  note own_dist = ownership_distinct [OF i_bound j_bound neq_i_j ith jth]
  
  from sharing_consis [OF j_bound jth]
  have consis_j: "sharing_consistent 𝒮 𝒪j sbj".
  
  from sharing_consis [OF i_bound ith]
  have consis_i: "sharing_consistent 𝒮 𝒪i sbi".
  
  let ?take_j = "(takeWhile (Not  is_volatile_Writesb) sbj)"
  let ?drop_j = "(dropWhile (Not  is_volatile_Writesb) sbj)"


  { 
    fix x
    assume x_in_take_i: "x  ?non_vol_writes_i"
    assume x_in_j: "x  ?not_volatile_j"
    from x_in_take_i have x_in_i: "x  outstanding_refs (is_non_volatile_Writesb) sbi"
      by (auto dest: outstanding_refs_takeWhileD)
    from non_volatile_owned_or_read_only_outstanding_non_volatile_writes [OF nvo_i] x_in_i
    have x_in_owns_acq_i: "x  𝒪i  all_acquired sbi"
      by auto
    have False
    proof -
      {
	assume x_in_j: "x  outstanding_refs is_volatile_Writesb sbj" 
	with dist_j have x_notin: "x  (𝒪i  all_acquired sbi)"
	  by auto
	with x_in_owns_acq_i have False
	  by auto
      }
      moreover
      {
	assume x_in_j: "x  outstanding_refs is_non_volatile_Writesb sbj"
	from non_volatile_owned_or_read_only_outstanding_non_volatile_writes [OF nvo_j] x_in_j
	have "x  𝒪j  all_acquired sbj"
	  by auto
	with x_in_owns_acq_i own_dist
	have False
	  by auto
      }
      moreover
      {
	assume x_in_j: "x  outstanding_refs is_non_volatile_Readsb ?drop_j"
	  
	from non_volatile_owned_or_read_only_drop [OF nvo_j]
	have nvo': "non_volatile_owned_or_read_only True (share ?take_j 𝒮) (acquired True ?take_j 𝒪j) ?drop_j".

	from non_volatile_owned_or_read_only_outstanding_refs [OF nvo'] x_in_j
	have "x  acquired True ?take_j 𝒪j  all_acquired ?drop_j   
	  read_only_reads (acquired True ?take_j 𝒪j) ?drop_j"
	  by (auto simp add: misc_outstanding_refs_convs)
	
	moreover 
	from acquired_append [of True ?take_j ?drop_j 𝒪j] acquired_all_acquired [of True ?take_j 𝒪j]
	  all_acquired_append [of ?take_j ?drop_j]
	have "acquired True ?take_j 𝒪j  all_acquired ?drop_j  𝒪j  all_acquired sbj"
	  by auto
	ultimately 
	have "x  read_only_reads (acquired True ?take_j 𝒪j) ?drop_j"
	  using x_in_owns_acq_i own_dist
	  by auto
	
	with read_only_reads_unowned [OF j_bound i_bound neq_i_j [symmetric] jth ith] x_in_owns_acq_i
	have False
	  by auto
      }
      moreover
      {
	assume x_in_j: "x  outstanding_refs is_non_volatile_Readsb ?take_j" 
	assume x_notin: "x  read_only_reads 𝒪j ?take_j"
	from non_volatile_owned_or_read_only_append [where xs="?take_j" and ys="?drop_j"] nvo_j
	have "non_volatile_owned_or_read_only False 𝒮 𝒪j ?take_j"
	  by auto
	
	from non_volatile_owned_or_read_only_outstanding_refs [OF this]  x_in_j x_notin
	have "x  𝒪j  all_acquired ?take_j" 
	  by (auto simp add: misc_outstanding_refs_convs )
	with all_acquired_append [of ?take_j ?drop_j] x_in_owns_acq_i own_dist
	have False
	  by auto
      }
      moreover
      {
	assume x_in_j: "x  𝒪j  all_acquired ?take_j"
	moreover
	from all_acquired_append [of ?take_j ?drop_j]
	have "all_acquired ?take_j  all_acquired sbj"
	  by auto
	ultimately have False
	  using x_in_owns_acq_i own_dist
	  by auto
      }
      ultimately show ?thesis
	using x_in_take_i x_in_j
	by (auto simp add: misc_outstanding_refs_convs)
    qed
  }
  then show ?thesis
    by auto
qed

lemma (in valid_ownership_and_sharing) outstanding_non_volatile_write_not_volatile_read_disj:
assumes i_bound: "i < length ts"
assumes j_bound: "j < length ts"
assumes neq_i_j: "i  j"
assumes ith: "ts!i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
assumes jth: "ts!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
shows "outstanding_refs (is_non_volatile_Writesb) (takeWhile (Not  is_volatile_Writesb) sbi)  
        outstanding_refs (Not  is_volatile_Readsb) (dropWhile (Not  is_volatile_Writesb) sbj) = {}" 
  (is "?non_vol_writes_i  ?not_volatile_j = {}")
proof -
  have "outstanding_refs (Not  is_volatile_Readsb) (dropWhile (Not  is_volatile_Writesb) sbj)  
    outstanding_refs is_volatile_Writesb sbj   
    outstanding_refs is_non_volatile_Writesb sbj  
    outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) sbj)"
    by (auto simp add: misc_outstanding_refs_convs dest: outstanding_refs_dropWhileD)
  with outstanding_non_volatile_write_disj [OF i_bound j_bound neq_i_j ith jth]
  show ?thesis
    by blast
qed



lemma (in valid_ownership_and_sharing) outstanding_refs_is_Writesb_takeWhile_disj:
        "i < length ts. (j < length ts. i  j 
                  (let (_,_,_,sbi,_,_,_) = ts!i;
                      (_,_,_,sbj,_,_,_) = ts!j
                   in outstanding_refs is_Writesb sbi  
                      outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj) = {}))"
proof -
  {
    fix i j pi isi 𝒪i i 𝒟i θi sbi pj isj 𝒪j j 𝒟j θj sbj
    assume i_bound: "i < length ts"
    assume j_bound: "j < length ts"
    assume neq_i_j: "i  j"
    assume ith: "ts!i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
    assume jth: "ts!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
    from outstanding_non_volatile_write_disj [OF j_bound i_bound neq_i_j[symmetric] jth ith]
    have "outstanding_refs is_Writesb sbi  
                      outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj) = {}"
      apply (clarsimp simp add: outstanding_refs_is_non_volatile_Writesb_takeWhile_conv)
      apply (auto simp add: misc_outstanding_refs_convs )
      done
  }
  thus ?thesis
    by (fastforce simp add: Let_def)
qed



fun read_tmps:: "'p store_buffer  tmp set"
where
  "read_tmps [] = {}"
| "read_tmps (r#rs) =
     (case r of 
       Readsb volatile a t v  insert t (read_tmps rs)
      | _  read_tmps rs)"


lemma in_read_tmps_conv:
  "(t  read_tmps xs) = (volatile a v. Readsb volatile a t v  set xs)"
  by (induct xs) (auto split: memref.splits)


lemma read_tmps_mono: "ys. set xs  set ys  read_tmps xs  read_tmps ys"
  by (fastforce simp add: in_read_tmps_conv)



fun distinct_read_tmps:: "'p store_buffer  bool"
where
  "distinct_read_tmps [] = True"
| "distinct_read_tmps (r#rs) =
     (case r of 
         Readsb volatile a t v  t  (read_tmps rs)  distinct_read_tmps rs
       | _  distinct_read_tmps rs)"
 
lemma distinct_read_tmps_conv:
 "distinct_read_tmps xs = (i < length xs. j < length xs. i  j 
      (case xs!i of 
         Readsb _ _ ti _  case xs!j of Readsb _ _ tj _  ti  tj | _  True
       | _  True))"
― ‹Nice lemma, ugly proof.›
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x xs)
  show ?case
  proof (cases x)
    case (Writesb volatile a sop v)
    with Cons.hyps show ?thesis
      apply -
      apply (rule iffI [rule_format]) 
      apply  clarsimp
             subgoal for i j
             apply (case_tac i)
             apply  fastforce
             apply (case_tac j)
             apply  (fastforce split: memref.splits) 
             apply (clarsimp cong: memref.case_cong)
             done
      apply clarsimp
      subgoal for i j
      apply (erule_tac x="Suc i" in allE)
      apply  clarsimp
      apply (erule_tac x="Suc j" in allE)
      apply  (clarsimp cong: memref.case_cong)
      done
      done
  next
    case (Readsb volatile a t v)
    with Cons.hyps show ?thesis
      apply -
      apply (rule iffI [rule_format]) 
      apply  clarsimp
             subgoal for i j 
             apply (case_tac i)
             apply  clarsimp
             apply  (case_tac j)
             apply   clarsimp
             apply  (fastforce split: memref.splits simp add: in_read_tmps_conv dest: nth_mem)
             apply (clarsimp)
             apply (case_tac j)
             apply  (fastforce split: memref.splits simp add: in_read_tmps_conv dest: nth_mem)
             apply (clarsimp cong: memref.case_cong)
             done
      apply clarsimp
      apply (rule conjI)
      apply (clarsimp simp add: in_read_tmps_conv)
      apply  (erule_tac x="0" in allE)
      apply  (clarsimp simp add: in_set_conv_nth)
             subgoal for volatile' a' v' i
             apply  (erule_tac x="Suc i" in allE)
             apply  clarsimp
             done
      apply clarsimp
      subgoal for i j
      apply (erule_tac x="Suc i" in allE)
      apply clarsimp
      apply (erule_tac x="Suc j" in allE)
      apply (clarsimp cong: memref.case_cong)
      done
      done
  next
    case Progsb
    with Cons.hyps show ?thesis
      apply -
      apply (rule iffI [rule_format]) 
      apply  clarsimp
             subgoal for i j
             apply (case_tac i)
             apply  fastforce
             apply (case_tac j)
             apply  (fastforce split: memref.splits) 
             apply (clarsimp cong: memref.case_cong)
             done
      apply clarsimp
      subgoal for i j
      apply (erule_tac x="Suc i" in allE)
      apply  clarsimp
      apply (erule_tac x="Suc j" in allE)
      apply  (clarsimp cong: memref.case_cong)
      done
      done
  next
    case Ghostsb
    with Cons.hyps show ?thesis
      apply -
      apply (rule iffI [rule_format]) 
      apply  clarsimp
             subgoal for i j
             apply (case_tac i)
             apply  fastforce
             apply (case_tac j)
             apply  (fastforce split: memref.splits) 
             apply (clarsimp cong: memref.case_cong)
             done
      apply clarsimp
      subgoal for i j
      apply (erule_tac x="Suc i" in allE)
      apply  clarsimp
      apply (erule_tac x="Suc j" in allE)
      apply (clarsimp cong: memref.case_cong)
      done
      done
  qed
qed

fun load_tmps:: "instrs  tmp set"
where
  "load_tmps [] = {}"
| "load_tmps (i#is) =
     (case i of
        Read volatile a t  insert t (load_tmps is)
      | RMW _ t _ _ _ _ _ _ _   insert t (load_tmps is)
      | _  load_tmps is)"

lemma in_load_tmps_conv:
  "(t  load_tmps xs) = ((volatile a. Read volatile a t  set xs) 
                         (a sop cond ret A L R W. RMW a t sop cond ret A L R W  set xs))"
  by (induct xs) (auto split: instr.splits)

lemma load_tmps_mono: "ys. set xs  set ys  load_tmps xs  load_tmps ys"
  by (fastforce simp add: in_load_tmps_conv)

fun distinct_load_tmps:: "instrs  bool"
where
  "distinct_load_tmps [] = True"
| "distinct_load_tmps (r#rs) =
     (case r of 
         Read volatile a t  t  (load_tmps rs)  distinct_load_tmps rs
       | RMW a t sop cond ret A L R W  t  (load_tmps rs)  distinct_load_tmps rs
       | _  distinct_load_tmps rs)"


locale load_tmps_distinct =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes load_tmps_distinct:
  "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)  
   
   distinct_load_tmps is"

locale read_tmps_distinct =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes read_tmps_distinct:
  "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)  
   
   distinct_read_tmps sb"

locale load_tmps_read_tmps_distinct =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes load_tmps_read_tmps_distinct:
  "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)  
   
   load_tmps is  read_tmps sb = {}"

locale tmps_distinct = 
  load_tmps_distinct + 
  read_tmps_distinct + 
  load_tmps_read_tmps_distinct

lemma rev_read_tmps: "read_tmps (rev xs) = read_tmps xs"
  by (auto simp add: in_read_tmps_conv)

lemma rev_load_tmps: "load_tmps (rev xs) = load_tmps xs"
  by (auto simp add: in_load_tmps_conv)

lemma distinct_read_tmps_append: "ys. distinct_read_tmps (xs @ ys) = 
        (distinct_read_tmps xs  distinct_read_tmps ys 
        read_tmps xs  read_tmps ys = {})" 
by (induct xs) (auto split: memref.splits simp add: in_read_tmps_conv)

lemma distinct_load_tmps_append: "ys. distinct_load_tmps (xs @ ys) = 
        (distinct_load_tmps xs  distinct_load_tmps ys 
        load_tmps xs  load_tmps ys = {})" 
apply (induct xs) 
apply (auto split: instr.splits simp add: in_load_tmps_conv)
done

lemma read_tmps_append: "read_tmps (xs@ys) = (read_tmps xs  read_tmps ys)"
  by (fastforce simp add: in_read_tmps_conv)

lemma load_tmps_append: "load_tmps (xs@ys) = (load_tmps xs  load_tmps ys)"
  by (fastforce simp add: in_load_tmps_conv)

fun write_sops:: "'p store_buffer  sop set"
where
  "write_sops [] = {}"
| "write_sops (r#rs) =
     (case r of 
       Writesb volatile a sop v _ _ _ _ insert sop (write_sops rs)
      | _  write_sops rs)"

lemma in_write_sops_conv:
  "(sop  write_sops xs) = (volatile a v A L R W. Writesb volatile a sop v A L R W  set xs)"
  apply (induct xs) 
  apply  simp
  apply (auto split: memref.splits) 
  apply  force
  apply force
  done

lemma write_sops_mono: "ys. set xs  set ys  write_sops xs  write_sops ys"
  by (fastforce simp add: in_write_sops_conv)

lemma write_sops_append: "write_sops (xs@ys) = write_sops xs  write_sops ys"
  by (force simp add: in_write_sops_conv)

  
fun store_sops:: "instrs  sop set"
where
  "store_sops [] = {}"
| "store_sops (i#is) =
     (case i of
        Write volatile a sop _ _ _ _  insert sop (store_sops is)
      | RMW a t sop cond ret A L R W  insert sop (store_sops is) 
      | _  store_sops is)"

lemma in_store_sops_conv:
  "(sop  store_sops xs) = ((volatile a A L R W. Write volatile a sop A L R W  set xs) 
                            (a t cond ret A L R W. RMW a t sop cond ret A L R W  set xs))"
  by (induct xs) (auto split: instr.splits)

lemma store_sops_mono: "ys. set xs  set ys  store_sops xs  store_sops ys"
  by (fastforce simp add: in_store_sops_conv)

lemma store_sops_append: "store_sops (xs@ys) = store_sops xs  store_sops ys"
  by (force simp add: in_store_sops_conv)

locale valid_write_sops =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes valid_write_sops:
  "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,) 
   
   sop  write_sops sb. valid_sop sop"

locale valid_store_sops =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes valid_store_sops:
  "i is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,) 
   
   sop  store_sops is. valid_sop sop"

locale valid_sops = valid_write_sops + valid_store_sops

text ‹The value stored in a non-volatile @{const "Readsb"} in the store-buffer has to match the
 last value written to the same address in the store buffer 
 or the memory content if there is no corresponding write in the store buffer.
 No volatile read may follow a volatile write.
 Volatile reads in the store buffer may refer to a stale value:
  e.g. imagine one writer and multiple readers
›

(* Readsb-only reads in the takeWhile part of the store buffer may become stale*)
(* FIXME: The flushing stuff: outstanding_refs is_volatile_Readsb rs = {} ∧ acquired_reads True rs (A - R) = {}
   does not fit well in this definition (it is not memory dependent).
   Maybe it would fit better to sharing_consistent.
*)
fun reads_consistent:: "bool  owns  memory  'p store_buffer  bool"
where
  "reads_consistent pending_write 𝒪 m [] = True"
| "reads_consistent pending_write 𝒪 m (r#rs) = 
   (case r of 
      Readsb volatile a t v  (¬ volatile  (pending_write  a  𝒪)  v = m a)  
                             reads_consistent pending_write 𝒪 m rs
    | Writesb volatile a sop v A L R W  
        (if volatile then
             outstanding_refs is_volatile_Readsb rs = {}  
             reads_consistent True (𝒪  A - R) (m(a := v)) rs
         else reads_consistent pending_write 𝒪 (m(a := v)) rs)
    | Ghostsb A L R W  reads_consistent pending_write (𝒪  A - R) m rs
    | _  reads_consistent pending_write 𝒪 m rs
   )"

fun volatile_reads_consistent:: "memory  'p store_buffer  bool"
where
  "volatile_reads_consistent m [] = True"
| "volatile_reads_consistent m (r#rs) = 
   (case r of 
      Readsb volatile a t v  (volatile  v = m a)  volatile_reads_consistent m rs
    | Writesb volatile a sop v A L R W  volatile_reads_consistent (m(a := v)) rs 
    | _  volatile_reads_consistent m rs
   )"

fun flush:: "'p store_buffer  memory  memory"
where
  "flush [] m = m"
| "flush (r#rs) m =
     (case r of 
        Writesb volatile a _ v _ _ _ _  flush rs (m(a:=v))
      | _  flush rs m)"

lemma reads_consistent_pending_write_antimono:
  "𝒪 m. reads_consistent True 𝒪 m sb  reads_consistent False 𝒪 m sb"
apply (induct sb)
apply  simp
subgoal for a sb 𝒪 m
  by (case_tac a) auto
done

lemma reads_consistent_owns_antimono:
  "𝒪 𝒪' pending_write m.
       𝒪  𝒪'  reads_consistent pending_write 𝒪' m sb  reads_consistent pending_write 𝒪 m sb"
apply (induct sb)
apply  simp
subgoal for a sb 𝒪 𝒪' pending_write m
apply (case_tac a)
apply    (clarsimp split: if_split_asm)
         subgoal for volatile a D f v A L R W
         apply (drule_tac C="A" in union_mono_aux)
         apply (drule_tac C="R" in set_minus_mono_aux)
         apply blast
         done
apply   fastforce
apply  fastforce
apply clarsimp
subgoal for A L R W 
apply (drule_tac C="A" in union_mono_aux)
apply (drule_tac C="R" in set_minus_mono_aux)
apply blast
done
done
done

lemma acquired_reads_mono': "x  acquired_reads b xs A  acquired_reads b xs B = {}  A  B  False"
apply (drule acquired_reads_mono_in [where B=B])
apply auto
done


lemma reads_consistent_append: 
  "m pending_write 𝒪. reads_consistent pending_write 𝒪 m (xs@ys) = 
      (reads_consistent pending_write 𝒪 m xs 
       reads_consistent (pending_write  outstanding_refs is_volatile_Writesb xs  {}) 
          (acquired True xs 𝒪) (flush xs m) ys 
       (outstanding_refs is_volatile_Writesb xs  {}
         outstanding_refs is_volatile_Readsb ys = {} ))"
apply (induct xs)
apply  clarsimp
subgoal for a xs m pending_write 𝒪
apply (case_tac a)
apply (auto simp add: outstanding_refs_append acquired_reads_append
dest: acquired_reads_mono_in acquired_pending_write_mono_in acquired_reads_mono' acquired_mono_in)
done
done


lemma reads_consistent_mem_eq_on_non_volatile_reads:
  assumes mem_eq: "a  A. m' a = m a"
  assumes subset: "outstanding_refs (is_non_volatile_Readsb) sb  A"
  ― ‹We could be even more restrictive here, only the non volatile reads that are
        not buffered in @{term "sb"} have to be the same.›
  assumes consis_m: "reads_consistent pending_write 𝒪 m sb"
  shows "reads_consistent pending_write 𝒪 m' sb"
using mem_eq subset consis_m 
proof (induct sb arbitrary: m' m pending_write 𝒪)
  case Nil thus ?case by simp
next
  case (Cons r sb)
  note mem_eq = a  A. m' a = m a
  note subset = outstanding_refs (is_non_volatile_Readsb) (r#sb)  A
  note consis_m = reads_consistent pending_write 𝒪 m (r#sb)

  from subset have subset': "outstanding_refs is_non_volatile_Readsb sb  A"
    by (auto simp add: Writesb)
  show ?case
  proof (cases r)
    case (Writesb volatile a sop v A' L R W)
    from mem_eq
    have mem_eq': 
      "a'  A. (m'(a:=v)) a' = (m(a:=v)) a'"
      by (auto)
    show ?thesis
    proof (cases volatile)
      case True
      from consis_m obtain
	consis': "reads_consistent True (𝒪  A' - R) (m(a := v)) sb" and
        no_volatile_Readsb: "outstanding_refs is_volatile_Readsb sb = {}" 
	by (simp add: Writesb True)

      from Cons.hyps [OF mem_eq' subset' consis']
      have "reads_consistent True (𝒪  A' - R) (m'(a := v)) sb".
      with no_volatile_Readsb
      show ?thesis
	by (simp add: Writesb True)
    next
      case False
      from consis_m obtain consis': "reads_consistent pending_write 𝒪 (m(a := v)) sb" 
	by (simp add: Writesb False)
      from Cons.hyps [OF mem_eq' subset' consis']
      have "reads_consistent pending_write 𝒪 (m'(a := v)) sb".
      then
      show ?thesis
	by (simp add: Writesb False)
    qed
  next
    case (Readsb volatile a t v)
    from mem_eq
    have mem_eq': 
      "a'  A. m' a' = m a'"
      by (auto)
    show ?thesis
    proof (cases volatile)
      case True
      from consis_m obtain	
	consis': "reads_consistent pending_write 𝒪 m sb"  
	by (simp add: Readsb True)
      from Cons.hyps [OF mem_eq' subset' consis']
      show ?thesis
	by (simp add: Readsb True)
    next
      case False
      from consis_m obtain	
	consis': "reads_consistent pending_write 𝒪 m sb"  and v: "(pending_write  a  𝒪)  v=m a" 
	by (simp add: Readsb False)
      from mem_eq subset Readsb have "m' a = m a"
	by (auto simp add: False)
      with Cons.hyps [OF mem_eq' subset' consis'] v
      show ?thesis
	by (simp add: Readsb False)
    qed
  next
    case Progsb with Cons show ?thesis by auto
  next
    case Ghostsb with Cons show ?thesis by auto
  qed
qed



lemma volatile_reads_consistent_mem_eq_on_volatile_reads:
  assumes mem_eq: "a  A. m' a = m a"
  assumes subset: "outstanding_refs (is_volatile_Readsb) sb  A"
  ― ‹We could be even more restrictive here, only the non volatile reads that are
        not buffered in @{term "sb"} have to be the same.›
  assumes consis_m: "volatile_reads_consistent m sb"
  shows "volatile_reads_consistent m' sb"
using mem_eq subset consis_m 
proof (induct sb arbitrary: m' m)
  case Nil thus ?case by simp
next
  case (Cons r sb)
  note mem_eq = a  A. m' a = m a
  note subset = outstanding_refs (is_volatile_Readsb) (r#sb)  A
  note consis_m = volatile_reads_consistent m (r#sb)

  from subset have subset': "outstanding_refs is_volatile_Readsb sb  A"
    by (auto simp add: Writesb)
  show ?case
  proof (cases r)
    case (Writesb volatile a sop v A' L R W)
    from mem_eq
    have mem_eq': 
      "a'  A. (m'(a:=v)) a' = (m(a:=v)) a'"
      by (auto)
    show ?thesis
    proof (cases volatile)
      case True
      from consis_m obtain
	consis': "volatile_reads_consistent (m(a := v)) sb"
	by (simp add: Writesb True)

      from Cons.hyps [OF mem_eq' subset' consis']
      have "volatile_reads_consistent (m'(a := v)) sb".
      then
      show ?thesis
	by (simp add: Writesb True)
    next
      case False
      from consis_m obtain consis': "volatile_reads_consistent (m(a := v)) sb" 
	by (simp add: Writesb False)
      from Cons.hyps [OF mem_eq' subset' consis']
      have "volatile_reads_consistent (m'(a := v)) sb".
      then
      show ?thesis
	by (simp add: Writesb False)
    qed
  next
    case (Readsb volatile a t v)
    from mem_eq
    have mem_eq': 
      "a'  A. m' a' = m a'"
      by (auto)
    show ?thesis
    proof (cases volatile)
      case False
      from consis_m obtain	
	consis': "volatile_reads_consistent m sb"  
	by (simp add: Readsb False)
      from Cons.hyps [OF mem_eq' subset' consis']
      show ?thesis
	by (simp add: Readsb False)
    next
      case True
      from consis_m obtain	
	consis': "volatile_reads_consistent m sb"  and v: "v=m a" 
	by (simp add: Readsb True)
      from mem_eq subset Readsb v have "v = m' a"
	by (auto simp add: True)
      with Cons.hyps [OF mem_eq' subset' consis'] 
      show ?thesis
	by (simp add: Readsb True)
    qed
  next
    case Progsb with Cons show ?thesis by auto
  next
    case Ghostsb with Cons show ?thesis by auto
  qed
qed

locale valid_reads =
fixes m::"memory"  and ts::"('p, 'p store_buffer,bool,owns,rels) thread_config list" 
assumes valid_reads: "i p is 𝒪  𝒟 θ sb. 
          i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)   
            reads_consistent False 𝒪 m sb"

lemma valid_reads_Cons: "valid_reads m (t#ts) = 
  (let (_,_,_,sb,_,𝒪,_) = t in reads_consistent False 𝒪 m sb  valid_reads m ts)"
apply (auto simp add: valid_reads_def)
subgoal for p' is' θ' sb' 𝒟' 𝒪' ℛ' i p "is" θ sb 𝒟 𝒪 ℛ
apply (case_tac i)
apply auto
done
done



text Readsbs› and writes have in the store-buffer have to conform to the 
  valuation of temporaries.›
context program
begin
fun history_consistent:: "tmps  'p  'p store_buffer  bool"
where
  "history_consistent θ p [] = True"
| "history_consistent θ p (r#rs) =
    (case r of
       Readsb vol a t v  
         (case θ t of Some v'  v=v'  history_consistent θ p rs | _  False) 
     | Writesb vol a (D,f) v _ _ _ _  
           D  dom θ  f θ = v  D  read_tmps rs = {}  history_consistent θ p rs
     | Progsb p1 p2 is  p1=p  
                           θ|`(dom θ - read_tmps rs) p1 p (p2,is)  
                           history_consistent θ p2 rs
     | _  history_consistent θ p rs)"
end

fun hd_prog:: "'p  'p store_buffer  'p"
where
  "hd_prog p [] = p"
| "hd_prog p (i#is) = (case i of
      Progsb p' _ _  p'
    | _  hd_prog p is)" 

fun last_prog:: "'p  'p store_buffer  'p"
where
  "last_prog p [] = p"
| "last_prog p (i#is) = (case i of
      Progsb _ p' _  last_prog p' is
    | _  last_prog p is)" 

locale valid_history = program +
constrains 
  program_step :: "tmps  'p  'p × instrs  bool"  
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list" 
assumes valid_history: "i p is 𝒪  𝒟 θ sb. 
          i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)   
            program.history_consistent program_step θ (hd_prog p sb) sb"

fun data_dependency_consistent_instrs:: "addr set  instrs  bool"
where
  "data_dependency_consistent_instrs T [] = True"
| "data_dependency_consistent_instrs T (i#is) =
     (case i of
        Write volatile a (D,f) _ _ _ _  D  T  D  load_tmps is = {}  data_dependency_consistent_instrs T is
      | RMW a t (D,f) cond ret _ _ _ _  D  T  D  load_tmps is = {}  data_dependency_consistent_instrs (insert t T) is
      | Read _ _ t  data_dependency_consistent_instrs (insert t T) is
      | _  data_dependency_consistent_instrs T is)"

lemma data_dependency_consistent_mono:
" T T'. data_dependency_consistent_instrs T is; T  T'  data_dependency_consistent_instrs T' is"
apply (induct "is")
apply clarsimp
subgoal for a "is" T T'
apply (case_tac a)
apply     clarsimp
          subgoal for volatile a' t
          apply (drule_tac a=t in insert_mono)
          apply clarsimp
          done
apply    fastforce
apply   clarsimp
        subgoal for a' t D f cond ret A L R W 
        apply (frule_tac a=t in insert_mono)
        apply fastforce
        done
apply  fastforce
apply fastforce
done
done


lemma data_dependency_consistent_instrs_append:
  "ys T . data_dependency_consistent_instrs T (xs@ys) =
       (data_dependency_consistent_instrs T xs  
        data_dependency_consistent_instrs (T  load_tmps xs) ys 
        load_tmps ys  (fst ` store_sops xs) = {})"
apply (induct xs)
apply (auto split: instr.splits simp add: load_tmps_append intro: data_dependency_consistent_mono)
done

locale valid_data_dependency =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list" 
assumes data_dependency_consistent_instrs: 
  "i p is 𝒪 𝒟 θ sb. 
          i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)   
            data_dependency_consistent_instrs (dom θ) is"
assumes load_tmps_write_tmps_distinct: 
  "i p is 𝒪 𝒟 θ sb. 
          i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)   
            load_tmps is  (fst ` write_sops sb) = {}"

locale load_tmps_fresh =
fixes ts::"('p, 'p store_buffer,bool,owns,rels) thread_config list" 
assumes load_tmps_fresh: 
  "i p is 𝒪 𝒟 θ sb. 
          i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)   
            load_tmps is  dom θ = {}"

fun acquired_by_instrs :: "instrs  addr set  addr set"
where
  "acquired_by_instrs [] A = A"
| "acquired_by_instrs (i#is) A = 
   (case i of
      Read _ _ _  acquired_by_instrs is A
    | Write volatile _ _ A' L R W  acquired_by_instrs is (if volatile then (A  A' - R) else A)
    | RMW a t sop cond ret A' L R W   acquired_by_instrs is {}
    | Fence    acquired_by_instrs is {}
    | Ghost A' L R W  acquired_by_instrs is (A  A' - R))"

fun acquired_loads :: "bool  instrs  addr set  addr set"
where
  "acquired_loads pending_write [] A = {}"
| "acquired_loads pending_write (i#is) A =
   (case i of
      Read volatile a _  (if pending_write  ¬ volatile  a  A 
                            then insert a (acquired_loads pending_write is A)
                            else acquired_loads pending_write is A)
    | Write volatile _ _ A' L R W  (if volatile then acquired_loads True is (if pending_write then (A  A' - R) else {})
                             else acquired_loads pending_write is A)
    | RMW a t sop cond ret A' L R W   acquired_loads pending_write is {}
    | Fence    acquired_loads pending_write is {}
    | Ghost A' L R W  acquired_loads pending_write is (A  A' - R))"

lemma acquired_by_instrs_mono: 
  " A B. A  B  acquired_by_instrs is A  acquired_by_instrs is B"
apply (induct "is")
apply  simp
subgoal for a "is" A B
apply (case_tac a)
apply      clarsimp
apply     clarsimp
          subgoal for volatile a' D f A' L R W x 
          apply (drule_tac C=A' in union_mono_aux)
          apply (drule_tac C="R" in set_minus_mono_aux)
          apply blast
          done
apply   clarsimp
apply  clarsimp
apply clarsimp
subgoal for A' L R W x
apply (drule_tac C=A' in union_mono_aux)
apply (drule_tac C="R" in set_minus_mono_aux)
apply blast
done
done
done

lemma acquired_by_instrs_mono_in:
  assumes x_in: "x  acquired_by_instrs is A" 
  assumes sub: "A  B" 
  shows "x  acquired_by_instrs is B"
using acquired_by_instrs_mono [OF sub, of "is"] x_in
by blast


locale enough_flushs =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list" 
assumes clean_no_outstanding_volatile_Writesb: 
  "i p is 𝒪  𝒟 θ sb. 
     i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,);¬ 𝒟  
       (outstanding_refs is_volatile_Writesb sb = {})"

fun prog_instrs:: "'p store_buffer  instrs"
where 
  "prog_instrs [] = []"
 |"prog_instrs (i#is) = (case i of
    Progsb _ _ is'  is' @ prog_instrs is
   | _  prog_instrs is)"

fun instrs:: "'p store_buffer  instrs"
where
  "instrs [] = []"
| "instrs (i#is) = (case i of
     Writesb volatile a sop v A L R W  Write volatile a sop A L R W# instrs is
   | Readsb volatile a t v  Read volatile a t # instrs is
   | Ghostsb A L R W  Ghost A L R W# instrs is
   | _  instrs is)"

locale causal_program_history =
fixes "issb" and sb
assumes causal_program_history: 
  "sb1 sb2. sb=sb1@sb2  is. instrs sb2 @ issb = is @ prog_instrs sb2"

lemma causal_program_history_empty [simp]: "causal_program_history is []"
  by (rule causal_program_history.intro) simp

lemma causal_program_history_suffix:
  "causal_program_history issb (sb@sb')  causal_program_history issb sb'"
  by (auto simp add: causal_program_history_def)

locale valid_program_history = 
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list" 
assumes valid_program_history: 
  "i p is 𝒪  𝒟 θ sb. 
     i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)   
      causal_program_history is sb"

assumes valid_last_prog:
  "i p is 𝒪  𝒟 θ sb. 
     i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)   
      last_prog p sb = p"

lemma (in valid_program_history) valid_program_history_nth_update:
 "i < length ts; causal_program_history is sb; last_prog p sb = p 
   
   valid_program_history (ts [i:=(p,is,θ,sb,𝒟,𝒪,)])"
  by (rule valid_program_history.intro)
     (auto dest: valid_program_history valid_last_prog
    simp add: nth_list_update split: if_split_asm)

lemma (in outstanding_non_volatile_refs_owned_or_read_only)
  outstanding_non_volatile_refs_owned_instructions_read_value_independent:
 "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)   
     outstanding_non_volatile_refs_owned_or_read_only 𝒮 (ts[i := (p',is',θ',sb,𝒟',𝒪,ℛ')])"
  by (unfold_locales)
     (auto dest: outstanding_non_volatile_refs_owned_or_read_only 
       simp add:  nth_list_update split: if_split_asm)

lemma (in outstanding_non_volatile_refs_owned_or_read_only)
  outstanding_non_volatile_refs_owned_or_read_only_nth_update:
 "i is 𝒪 𝒟   θ sb. 
   i < length ts; non_volatile_owned_or_read_only False 𝒮 𝒪 sb  
     outstanding_non_volatile_refs_owned_or_read_only 𝒮 (ts[i := (p,is,θ,sb,𝒟,𝒪,)])"
by (unfold_locales)
   (auto dest: outstanding_non_volatile_refs_owned_or_read_only 
       simp add:  nth_list_update split: if_split_asm)

lemma (in outstanding_volatile_writes_unowned_by_others)
  outstanding_volatile_writes_unowned_by_others_instructions_read_value_independent:
 "i p is 𝒪  𝒟  θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)   
     outstanding_volatile_writes_unowned_by_others (ts[i := (p',is',θ',sb,𝒟',𝒪,ℛ')])"
  by (unfold_locales)
     (auto dest: outstanding_volatile_writes_unowned_by_others 
       simp add:  nth_list_update split: if_split_asm)

lemma (in read_only_reads_unowned)
  read_only_unowned_instructions_read_value_independent:
 "i p is 𝒪  𝒟  θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)   
     read_only_reads_unowned (ts[i := (p',is',θ',sb,𝒟',𝒪,ℛ')])"
  by (unfold_locales)
     (auto dest: read_only_reads_unowned
       simp add:  nth_list_update split: if_split_asm)



lemma Writesb_in_outstanding_refs:
  "Writesb True a sop v A L R W  set xs  a  outstanding_refs is_volatile_Writesb xs"
  by (induct xs) (auto split:memref.splits)


lemma (in outstanding_volatile_writes_unowned_by_others)
  outstanding_volatile_writes_unowned_by_others_store_buffer:
 "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,); 
    outstanding_refs is_volatile_Writesb sb'  outstanding_refs is_volatile_Writesb sb;
    all_acquired sb'  all_acquired sb  
     outstanding_volatile_writes_unowned_by_others (ts[i := (p',is',θ',sb',𝒟',𝒪,ℛ')])"
  apply (unfold_locales)
  apply (fastforce dest: outstanding_volatile_writes_unowned_by_others 
         simp add:  nth_list_update split: if_split_asm)
  done


lemma (in ownership_distinct)
  ownership_distinct_instructions_read_value_store_buffer_independent:
 "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,);
     all_acquired sb'  all_acquired sb  
     ownership_distinct (ts[i := (p',is',θ',sb',𝒟',𝒪,ℛ')])"
  by (unfold_locales)
     (auto dest: ownership_distinct 
       simp add:  nth_list_update split: if_split_asm)


lemma (in ownership_distinct)
  ownership_distinct_nth_update:
 "i p is 𝒪  𝒟 xs sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,);
    j < length ts. ij  (let (pj,isj,θj,sbj,𝒟j,𝒪j,j) = ts!j 
          in (𝒪'  all_acquired sb')  (𝒪j  all_acquired sbj) ={})   
     ownership_distinct (ts[i := (p',is',θ',sb',𝒟',𝒪',ℛ')])"
  apply (unfold_locales)
  apply (clarsimp simp add: nth_list_update split: if_split_asm)
  apply   (force dest: ownership_distinct simp add: Let_def)
  apply  (fastforce dest: ownership_distinct simp add: Let_def)
  apply (fastforce dest: ownership_distinct simp add: Let_def)
  done


lemma (in valid_write_sops) valid_write_sops_nth_update:
          "i < length ts; sop  write_sops sb. valid_sop sop  
            valid_write_sops (ts[i := (p,is,xs,sb,𝒟,𝒪,)])"
  by (unfold valid_write_sops_def)
     (auto dest: valid_write_sops simp add: nth_list_update split: if_split_asm)

lemma (in valid_store_sops) valid_store_sops_nth_update:
          "i < length ts; sop  store_sops is. valid_sop sop  
            valid_store_sops (ts[i := (p,is,xs,sb,𝒟,𝒪,)])"
  by (unfold valid_store_sops_def)
     (auto dest: valid_store_sops simp add: nth_list_update split: if_split_asm)

lemma (in valid_sops) valid_sops_nth_update:
          "i < length ts; sop  write_sops sb. valid_sop sop;
            sop  store_sops is. valid_sop sop  
            valid_sops (ts[i := (p,is,xs,sb,𝒟,𝒪,)])"
  by (unfold valid_sops_def valid_write_sops_def valid_store_sops_def)
     (auto dest: valid_write_sops valid_store_sops 
       simp add: nth_list_update split: if_split_asm)


lemma (in valid_data_dependency) valid_data_dependency_nth_update:
          "i < length ts; data_dependency_consistent_instrs (dom θ) is; 
            load_tmps is  (fst ` write_sops sb) = {}  
            valid_data_dependency (ts[i := (p,is,θ,sb,𝒟,𝒪,)])"
  by (unfold valid_data_dependency_def)
     (force dest: data_dependency_consistent_instrs load_tmps_write_tmps_distinct 
         simp add: nth_list_update split: if_split_asm)

lemma (in enough_flushs) enough_flushs_nth_update:
 "i < length ts; 
   ¬ 𝒟  (outstanding_refs is_volatile_Writesb sb = {})
    
     enough_flushs (ts[i := (p,is,θ,sb,𝒟,𝒪,)])"

  apply (unfold_locales)
  apply  (force simp add: nth_list_update split: if_split_asm dest: clean_no_outstanding_volatile_Writesb)
  done

lemma (in outstanding_non_volatile_writes_unshared) 
  outstanding_non_volatile_writes_unshared_nth_update:
          "i < length ts; non_volatile_writes_unshared 𝒮 sb  
            outstanding_non_volatile_writes_unshared 𝒮 (ts[i := (p,is,xs,sb,𝒟,𝒪,)])"
  by (unfold_locales)
     (auto dest: outstanding_non_volatile_writes_unshared 
       simp add: nth_list_update split: if_split_asm)

lemma (in sharing_consis) 
  sharing_consis_nth_update:
          "i < length ts; sharing_consistent 𝒮 𝒪 sb  
            sharing_consis 𝒮 (ts[i := (p,is,xs,sb,𝒟,𝒪,)])"
  by (unfold_locales)
     (auto dest: sharing_consis 
       simp add: nth_list_update split: if_split_asm)



lemma (in no_outstanding_write_to_read_only_memory) 
  no_outstanding_write_to_read_only_memory_nth_update:
          "i < length ts; no_write_to_read_only_memory 𝒮 sb  
            no_outstanding_write_to_read_only_memory 𝒮 (ts[i := (p,is,xs,sb,𝒟,𝒪,)])"
  by (unfold_locales)
     (auto dest:  no_outstanding_write_to_read_only_memory
       simp add: nth_list_update split: if_split_asm)

lemma in_Union_image_nth_conv: "a   (f ` set xs)  i. i < length xs  a  f (xs!i)"    
  by (auto simp add: in_set_conv_nth)

lemma in_Inter_image_nth_conv: "a   (f ` set xs) = (i < length xs. a  f (xs!i))"    
  by (force simp add:  in_set_conv_nth)



lemma release_ownership_nth_update:
  assumes R_subset: "R  𝒪"
  shows "i. i < length ts; ts!i = (p,is,xs,sb,𝒟,𝒪,);
          ownership_distinct ts
      ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set (ts[i:=(p',is',xs',sb',𝒟',𝒪 - R,ℛ')]))
        = (( ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts))  - R )"
proof (induct ts)
  case Nil thus ?case by simp
next
  case (Cons t ts)
  note i_bound = i < length (t # ts)
  note ith = (t # ts) ! i = (p,is,xs, sb, 𝒟, 𝒪,)
  note dist = ownership_distinct (t#ts)
  then interpret ownership_distinct "t#ts".
  from dist
  have dist': "ownership_distinct ts"
    by (rule ownership_distinct_tl)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have t: "t = (p,is,xs,sb,𝒟,𝒪,)"
      by simp
    have "R  ( ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts)) = {}"
    proof -
      {
	fix x
	assume x_R: "x  R"
	assume x_ls: "x  ( ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts))"
	then obtain j pj "isj" 𝒪j j 𝒟j xsj sbj where
	  j_bound: "j < length ts" and
	  jth: "ts!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)" and
	  x_in: "x  𝒪j"
	  by (fastforce simp add: in_set_conv_nth )
	from  j_bound jth 0 
	have "(𝒪  all_acquired sb)  (𝒪j  all_acquired sbj)= {}"
	  apply -
	  apply (rule ownership_distinct [OF i_bound _ _ ith, of "Suc j"])
	  apply clarsimp+
	  apply blast
	  done
	
	with x_R R_subset x_in have False
	  by auto
      }
      thus ?thesis
	by blast
    qed
    then
    show ?thesis
      by (auto simp add: 0 t)  
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l xsl sbl where t: "t = (pl,isl,xsl,sbl,𝒟l,𝒪l,l)"
      by (cases t)

    have n_bound: "n < length ts"
      using i_bound by (simp add: Suc)
    have nth: "ts!n = (p,is,xs,sb,𝒟,𝒪,)"
      using ith by (simp add: Suc)
    
    have "R  (𝒪l  all_acquired sbl) = {}"
    proof -
      {
	fix x
	assume x_R: "x  R"
	assume x_ownsl: "x  (𝒪l  all_acquired sbl)"
	from t 
	have "(𝒪  all_acquired sb)  (𝒪l  all_acquired sbl)= {}"
	  apply -
	  apply (rule ownership_distinct [OF i_bound _ _ ith, of "0"])
	  apply (auto simp add: Suc)
	  done
	with x_ownsl x_R R_subset have False
	  by auto
      }
      thus ?thesis
	by blast
    qed
    with Cons.hyps [OF n_bound nth dist']
    show ?thesis
      by (auto simp add: Suc t)
  qed
qed
   
lemma acquire_ownership_nth_update:
  shows "i. i < length ts; ts!i = (p,is,xs,sb,𝒟,𝒪,)
      ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set (ts[i:=(p',is',xs',sb',𝒟',𝒪  A,ℛ')]))
        = (( ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts))   A )"
proof (induct ts)
  case Nil thus ?case by simp
next
  case (Cons t ts)
  note i_bound = i < length (t # ts)
  note ith = (t # ts) ! i = (p,is, xs, sb, 𝒟, 𝒪, )
  show ?case
  proof (cases i)
    case 0
    from ith 0 have t: "t = (p,is,xs,sb,𝒟,𝒪,)"
      by simp
    show ?thesis
      by (auto simp add: 0 t)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l xsl sbl where t: "t = (pl,isl,xsl,sbl,𝒟l,𝒪l,l)"
      by (cases t)

    have n_bound: "n < length ts"
      using i_bound by (simp add: Suc)
    have nth: "ts!n = (p,is,xs,sb,𝒟,𝒪,)"
      using ith by (simp add: Suc)
    from Cons.hyps [OF n_bound nth]
    show ?thesis
      by (auto simp add: Suc t)
  qed
qed

lemma acquire_release_ownership_nth_update:
  assumes R_subset: "R  𝒪"
  shows "i. i < length ts; ts!i = (p,is,xs,sb,𝒟,𝒪,);
          ownership_distinct ts
      ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set (ts[i:=(p',is',xs',sb',𝒟',𝒪  A - R,ℛ')]))
        = (( ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts))   A - R )"
proof (induct ts)
  case Nil thus ?case by simp
next
  case (Cons t ts)
  note i_bound = i < length (t # ts)
  note ith = (t # ts) ! i = (p,is, xs, sb,𝒟, 𝒪,)
  note dist = ownership_distinct (t#ts)
  then interpret ownership_distinct "t#ts".
  from dist
  have dist': "ownership_distinct ts"
    by (rule ownership_distinct_tl)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have t: "t = (p,is,xs,sb,𝒟,𝒪,)"
      by simp
    have "R  ( ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts)) = {}"
    proof -
      {
	fix x
	assume x_R: "x  R"
	assume x_ls: "x  ( ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts))"
	then obtain j pj "isj" 𝒪j j 𝒟j xsj sbj where
	  j_bound: "j < length ts" and
	  jth: "ts!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)" and
	  x_in: "x  𝒪j"
	  by (fastforce simp add: in_set_conv_nth )
	from  j_bound jth 0 
	have "(𝒪  all_acquired sb)  (𝒪j  all_acquired sbj)= {}"
	  apply -
	  apply (rule ownership_distinct [OF i_bound _ _ ith, of "Suc j"])
	  apply clarsimp+
	  apply blast
	  done
	
	with x_R R_subset x_in have False
	  by auto
      }
      thus ?thesis
	by blast
    qed
    then
    show ?thesis
      by (auto simp add: 0 t)  
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l xsl sbl where t: "t = (pl,isl,xsl,sbl,𝒟l,𝒪l,l)"
      by (cases t)

    have n_bound: "n < length ts"
      using i_bound by (simp add: Suc)
    have nth: "ts!n = (p,is,xs,sb,𝒟,𝒪,)"
      using ith by (simp add: Suc)
    
    have "R  (𝒪l  all_acquired sbl) = {}"
    proof -
      {
	fix x
	assume x_R: "x  R"
	assume x_ownsl: "x  (𝒪l  all_acquired sbl)"
	from t 
	have "(𝒪  all_acquired sb)  (𝒪l  all_acquired sbl)= {}"
	  apply -
	  apply (rule ownership_distinct [OF i_bound _ _ ith, of "0"])
	  apply (auto simp add: Suc)
	  done
	with x_ownsl x_R R_subset have False
	  by auto
      }
      thus ?thesis
	by blast
    qed
    with Cons.hyps [OF n_bound nth dist']
    show ?thesis
      by (auto simp add: Suc t)
  qed
qed



lemma (in valid_history) valid_history_nth_update:
          "i < length ts; history_consistent θ (hd_prog p sb) sb   
            valid_history program_step (ts[i := (p,is,θ,sb,𝒟,𝒪,)])"
  by (unfold_locales)
     (auto dest: valid_history simp add: nth_list_update split: if_split_asm)

lemma (in valid_reads) valid_reads_nth_update:
          "i < length ts; reads_consistent False 𝒪 m sb   
            valid_reads m (ts[i := (p,is,xs,sb,𝒟,𝒪,)])"
  by (unfold_locales)
     (auto dest: valid_reads simp add: nth_list_update split: if_split_asm)

lemma (in load_tmps_distinct) load_tmps_distinct_nth_update:
          "i < length ts; distinct_load_tmps is  
            load_tmps_distinct (ts[i := (p,is,xs,sb,𝒟,𝒪,)])"
  by (unfold_locales)
     (auto dest: load_tmps_distinct simp add: nth_list_update split: if_split_asm)

lemma (in read_tmps_distinct) read_tmps_distinct_nth_update:
          "i < length ts; distinct_read_tmps sb  
            read_tmps_distinct (ts[i := (p,is,xs,sb,𝒟,𝒪,)])"
  by (unfold_locales)
     (auto dest: read_tmps_distinct simp add: nth_list_update split: if_split_asm)

lemma (in load_tmps_read_tmps_distinct) load_tmps_read_tmps_distinct_nth_update:
          "i < length ts; load_tmps is  read_tmps sb = {}  
            load_tmps_read_tmps_distinct (ts[i := (p,is,xs,sb,𝒟,𝒪,)])"
  by (unfold_locales)
     (auto dest: load_tmps_read_tmps_distinct simp add: nth_list_update split: if_split_asm)

lemma (in load_tmps_fresh) load_tmps_fresh_nth_update:
          "i < length ts; 
            load_tmps is  dom θ = {}  
            load_tmps_fresh (ts[i := (p,is,θ,sb,𝒟,𝒪,)])"
  by (unfold_locales)
     (fastforce dest: load_tmps_fresh
         simp add: nth_list_update split: if_split_asm)




fun flush_all_until_volatile_write:: 
  "('p,'p store_buffer,'dirty,'owns,'rels) thread_config list  memory  memory"
where
  "flush_all_until_volatile_write [] m = m"
| "flush_all_until_volatile_write ((_, _, _, sb,_, _)#ts) m =
     flush_all_until_volatile_write ts (flush (takeWhile (Not  is_volatile_Writesb) sb) m)"

fun share_all_until_volatile_write:: 
  "('p,'p store_buffer,'dirty,'owns,'rels) thread_config list  shared  shared"
where
  "share_all_until_volatile_write [] S = S"
| "share_all_until_volatile_write ((_, _, _, sb,_,_)#ts) S =
     share_all_until_volatile_write ts (share (takeWhile (Not  is_volatile_Writesb) sb) S)"





lemma takeWhile_dropWhile_real_prefix: 
  "x  set xs; ¬ P x  y ys. xs=takeWhile P xs @ y#ys  ¬ P y  dropWhile P xs = y#ys"   
  by (induct xs) auto

lemma buffered_val_witness: "buffered_val sb a = Some v  
  volatile sop A L R W.  Writesb volatile a sop v A L R W  set sb"
  apply (induct sb) 
  apply  simp
  apply (clarsimp split: memref.splits option.splits if_split_asm)
  apply  blast
  apply blast
  done


lemma flush_append_Readsb:
  "m. (flush (takeWhile (Not  is_volatile_Writesb) (sb @ [Readsb volatile a t v])) m)
       = flush (takeWhile (Not  is_volatile_Writesb) sb) m"
by (induct sb) (auto split: memref.splits)

lemma flush_append_write:
   "m. (flush (sb @ [Writesb volatile a sop v A L R W]) m) = (flush sb m) (a:=v)"
by (induct sb) (auto split: memref.splits)

lemma flush_append_Progsb:
"m. (flush (takeWhile (Not  is_volatile_Writesb) (sb @ [Progsb p1 p2 mis])) m) = 
       (flush (takeWhile (Not  is_volatile_Writesb) sb) m) "
  by (induct sb) (auto split: memref.splits)

lemma flush_append_Ghostsb:
"m. (flush (takeWhile (Not  is_volatile_Writesb) (sb @ [Ghostsb A L R W])) m) = 
       (flush (takeWhile (Not  is_volatile_Writesb) sb) m) "
  by (induct sb) (auto split: memref.splits)

lemma share_append: "S. share (xs@ys) S = share ys (share xs S)"
  by (induct xs) (auto split: memref.splits)

lemma share_append_Readsb:
  "S. (share (takeWhile (Not  is_volatile_Writesb) (sb @ [Readsb volatile a t v])) S)
       = share (takeWhile (Not  is_volatile_Writesb) sb) S"
 by (induct sb) (auto split: memref.splits)

lemma share_append_Writesb:
  "S. (share (takeWhile (Not  is_volatile_Writesb) (sb @ [Writesb volatile a sop v A L R W])) S)
       = share (takeWhile (Not  is_volatile_Writesb) sb) S"
by (induct sb) (auto split: memref.splits)

lemma share_append_Progsb:
"S. (share (takeWhile (Not  is_volatile_Writesb) (sb @ [Progsb p1 p2 mis])) S) = 
       (share (takeWhile (Not  is_volatile_Writesb) sb) S) "
  by (induct sb) (auto split: memref.splits)

lemma in_acquired_no_pending_write_outstanding_write: 
  "a  acquired False sb A  outstanding_refs is_volatile_Writesb sb  {}"
apply (induct sb)
apply (auto split: memref.splits)
done

lemma flush_buffered_val_conv:
  "m. flush sb m a = (case buffered_val sb a of None  m a | Some v  v)"
  by (induct sb) (auto split: memref.splits option.splits)

(*
lemma reads_consistent_unbuffered_snoc: 
  "⋀m. buffered_val sb a = None ⟹ m a = v ⟹ reads_consistent m sb ⟹
        volatile ⟶ 
          outstanding_refs is_volatile_Writesb sb = {} ⟹ 
        ¬ volatile ⟶ a ∉ acquired False sb {}
  ⟹ reads_consistent m (sb @ [Readsb volatile a t v])"
  by (simp add: reads_consistent_append flush_buffered_val_conv)
*)

lemma reads_consistent_unbuffered_snoc: 
  "m. buffered_val sb a = None  m a = v  reads_consistent pending_write 𝒪 m sb 
        volatile  
          outstanding_refs is_volatile_Writesb sb = {} 
   reads_consistent pending_write 𝒪 m (sb @ [Readsb volatile a t v])"
  by (simp add: reads_consistent_append flush_buffered_val_conv)

lemma reads_consistent_buffered_snoc: 
  "m. buffered_val sb a = Some v   reads_consistent pending_write 𝒪 m sb 
        volatile  outstanding_refs is_volatile_Writesb sb = {} 
   reads_consistent pending_write 𝒪 m (sb @ [Readsb volatile a t v])"
  by (simp add: reads_consistent_append flush_buffered_val_conv)

lemma reads_consistent_snoc_Writesb:
  "m. reads_consistent pending_write 𝒪 m sb  
  reads_consistent pending_write 𝒪 m (sb @ [Writesb volatile a sop v A L R W])"
  by (simp add: reads_consistent_append)

lemma reads_consistent_snoc_Progsb:
  "m. reads_consistent pending_write 𝒪 m sb  reads_consistent pending_write 𝒪 m (sb @ [Progsb p1 p2 mis])"
  by (simp add: reads_consistent_append)

lemma reads_consistent_snoc_Ghostsb:
  "m. reads_consistent pending_write 𝒪 m sb  reads_consistent pending_write 𝒪 m (sb @ [Ghostsb A L R W])"
  by (simp add: reads_consistent_append)

(* FIXME: move to map.thy *)
lemma restrict_map_id [simp]:"m |` dom m = m"
  apply (rule ext)
  subgoal for x
  apply (case_tac "m x")
  apply (auto simp add: restrict_map_def domIff)
  done
  done

lemma flush_all_until_volatile_write_Read_commute:
  shows "m i.  i < length ls; ls!i=(p,Read volatile a t#is,θ,sb,𝒟,𝒪,)
      
    
    flush_all_until_volatile_write 
       (ls[i := (p,is , θ(tv), sb @ [Readsb volatile a t v],𝒟',𝒪',ℛ')]) m =
    flush_all_until_volatile_write ls m"
proof (induct ls)
  case Nil thus ?case
    by simp
next
  case (Cons l ls)
  note i_bound =  i < length (l#ls)
  note ith = (l#ls)!i = (p,Read volatile a t#is,θ,sb,𝒟,𝒪,)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,Read volatile a t#is,θ,sb,𝒟,𝒪,)"
      by simp
    thus ?thesis 
      by (simp add: 0 flush_append_Readsb del: fun_upd_apply )
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l,l)"
      by (cases l)
    from i_bound ith
    have "flush_all_until_volatile_write
      (ls[n := (p,is , θ(tv), sb @ [Readsb volatile a t v],𝒟',𝒪',ℛ') ]) 
      (flush (takeWhile (Not  is_volatile_Writesb) sbl) m) =
      flush_all_until_volatile_write ls (flush (takeWhile (Not  is_volatile_Writesb) sbl) m)"
      apply -
      apply (rule Cons.hyps)
      apply (auto simp add: Suc l)
      done
    
    then
    show ?thesis
      by (simp add: Suc l del: fun_upd_apply)
  qed
qed

lemma flush_all_until_volatile_write_append_Ghost_commute:
  "i m. i < length ts; ts!i=(p,is,θ,sb,𝒟,𝒪,)
        flush_all_until_volatile_write (ts[i := (p',is',θ', sb@[Ghostsb A L R W], 𝒟', 𝒪',ℛ')]) m
       = flush_all_until_volatile_write ts m"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons l ts)
  note i_bound =  i < length (l#ts)
  note ith = (l#ts)!i = (p,is,θ,sb,𝒟,𝒪,)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,is,θ,sb,𝒟,𝒪,)"
      by simp
    thus ?thesis 
      by (simp add: 0 flush_append_Ghostsb del: fun_upd_apply)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l,l)"
      by (cases l)

    from i_bound ith
    have "flush_all_until_volatile_write 
              (ts[n := (p',is',θ', sb@[Ghostsb A L R W], 𝒟', 𝒪',ℛ')])
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m) =
         flush_all_until_volatile_write ts
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m)"
      apply -
      apply (rule Cons.hyps)
      apply (auto simp add: Suc l)
      done

    then show ?thesis
      by (simp add: Suc l)
  qed
qed


lemma update_commute:
assumes g_unchanged: "a m. a  G  g m a = m a"
assumes g_independent: "a m. a  G  g (f m) a = g m a"
assumes f_unchanged: "a m. a  F  f m a = m a"
assumes f_independent: "a m. a  F  f (g m) a = f m a"
assumes disj: "G  F = {}"
shows "f (g m) = g (f m)"
proof 
  fix a
  show "f (g m) a = g (f m) a"
  proof (cases "a  G")
    case True
    with disj have a_notin_F: "a  F"
      by blast
    from f_unchanged [rule_format, OF a_notin_F, of "g m"]
    have "f (g m) a = g m a" .
    also
    from g_independent [rule_format, OF True]
    have " = g (f m) a" by simp
    finally show ?thesis .
  next
    case False
    note a_notin_G = this
    show ?thesis
    proof (cases "a  F")
      case True
      from f_independent [rule_format, OF True]
      have "f (g m) a = f m a" by simp
      also
      from g_unchanged [rule_format, OF a_notin_G]
      have " = g (f m) a"
	by simp
      finally show ?thesis .
    next
      case False
      from f_unchanged [rule_format, OF False]
      have "f (g m) a = g m a".
      also
      from g_unchanged [rule_format, OF a_notin_G]
      have " = m a" .
      also       
      from f_unchanged [rule_format, OF False]
      have " = f m a" by simp
      also
      from g_unchanged [rule_format, OF a_notin_G]
      have " = g (f m) a"
	by simp
      finally show ?thesis .
    qed
  qed
qed
      
      
lemma update_commute':
assumes g_unchanged: "a m. a  G  g m a = m a"
assumes g_independent: "a m1 m2. a  G  g m1 a = g m2 a"
assumes f_unchanged: "a m. a  F  f m a = m a"
assumes f_independent: "a m1 m2. a  F  f m1 a = f m2 a"
assumes disj: "G  F = {}"
shows "f (g m) = g (f m)"
proof -
  from g_independent have g_ind': "a m. a  G  g (f m) a = g m a" by blast
  from f_independent have f_ind': "a m. a  F  f (g m) a = f m a" by blast
  from update_commute [OF g_unchanged g_ind' f_unchanged f_ind' disj]
  show ?thesis .
qed

lemma flush_unchanged_addresses: "m. a  outstanding_refs is_Writesb sb  flush sb m a = m a"
proof (induct sb)
  case Nil thus ?case by simp
next
  case (Cons r sb)
  note a_notin = a  outstanding_refs is_Writesb (r#sb)
  show ?case
  proof (cases r)
    case (Writesb volatile a' sop v)
    from a_notin obtain neq_a_a': "aa'" and a_notin': "a  outstanding_refs is_Writesb sb"
      by (simp add: Writesb)
    from Cons.hyps [OF a_notin', of "m(a':=v)"] neq_a_a'
    show ?thesis
      apply (simp add: Writesb del: fun_upd_apply)
      apply simp
      done
  next
    case (Readsb volatile a' t v)
    from a_notin obtain a_notin': "a  outstanding_refs is_Writesb sb"
      by (simp add: Readsb)
    from Cons.hyps [OF a_notin', of "m"] 
    show ?thesis
      by (simp add: Readsb)
  next
    case Progsb with Cons show ?thesis by simp
  next
    case Ghostsb with Cons show ?thesis by simp
  qed
qed

lemma flushed_values_mem_independent:
  "m m' a. a  outstanding_refs is_Writesb sb   flush sb m' a = flush sb m a"
proof (induct sb)
  case Nil thus ?case by simp
next
  case (Cons r sb)
  show ?case
  proof (cases r)
    case (Writesb volatile a' sop' v')
    have "flush sb (m'(a' := v')) a' = flush sb (m(a' := v')) a'"
    proof (cases "a'  outstanding_refs is_Writesb sb")
      case True
      from Cons.hyps [OF this]
      show ?thesis .
    next
      case False
      from flush_unchanged_addresses [OF False]
      show ?thesis
	by simp
    qed
    with Cons.hyps Cons.prems
    show ?thesis
      by (auto simp add: Writesb)
  next
    case Readsb thus ?thesis using Cons
      by auto
  next
    case Progsb thus ?thesis using Cons
      by auto
  next
    case Ghostsb thus ?thesis using Cons
      by auto
  qed
qed

lemma flush_all_until_volatile_write_unchanged_addresses:
  "m. a  ((λ(_,_,_,sb,_,_,_). outstanding_refs is_Writesb 
                            (takeWhile (Not  is_volatile_Writesb) sb)) ` set ls) 
       flush_all_until_volatile_write ls m a = m a"
proof (induct ls)
  case Nil thus ?case by simp
next
  case (Cons l ls)
  obtain p "is" 𝒪  𝒟 xs sb where l: "l=(p,is,xs,sb,𝒟,𝒪,)"
    by (cases l)
  note a   ((λ(_,_,_,sb,_,_,_). outstanding_refs is_Writesb 
                            (takeWhile (Not  is_volatile_Writesb) sb)) ` set (l#ls))
  then obtain
    a_notin_sb: "a  outstanding_refs is_Writesb 
                            (takeWhile (Not  is_volatile_Writesb) sb)" and
    a_notin_ls: "a   ((λ(_,_,_,sb,_,_,_). outstanding_refs is_Writesb 
                            (takeWhile (Not  is_volatile_Writesb) sb)) ` set ls)"
    by (auto simp add: l)

  
  from Cons.hyps [OF a_notin_ls]
  have "flush_all_until_volatile_write ls (flush (takeWhile (Not  is_volatile_Writesb) sb) m) a 
        = 
        (flush (takeWhile (Not  is_volatile_Writesb) sb) m) a".

  also

  from flush_unchanged_addresses [OF a_notin_sb]
  have "(flush (takeWhile (Not  is_volatile_Writesb) sb) m) a = m a".
  finally
  show ?case
    by (simp add: l)
qed

lemma  notin_outstanding_non_volatile_takeWhile_lem:
  "a  outstanding_refs (Not  is_volatile) sb
        
        a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sb)"
apply (induct sb)
apply (auto simp add: is_Writesb_def split: if_split_asm memref.splits)
done

lemma  notin_outstanding_non_volatile_takeWhile_lem':
  "a  outstanding_refs is_non_volatile_Writesb sb
        
        a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sb)"
apply (induct sb)
apply (auto simp add: is_Writesb_def split: if_split_asm memref.splits)
done

lemma notin_outstanding_non_volatile_takeWhile_Un_lem':
"a   ((λ(_,_,_,sb,_,_,_). outstanding_refs (Not  is_volatile) sb) ` set ls)
  a   ((λ(_,_,_,sb,_,_,_). outstanding_refs is_Writesb 
                            (takeWhile (Not  is_volatile_Writesb) sb)) ` set ls)"
proof (induct ls)
  case Nil thus ?case by simp
next
  case (Cons l ls)
  obtain p "is" 𝒪  𝒟 xs sb where l: "l=(p,is,xs,sb,𝒟,𝒪,)"
    by (cases l)

  from Cons.prems
  obtain 
    a_notin_sb: "a  outstanding_refs (Not  is_volatile) sb" and
    a_notin_ls: "a   ((λ(_,_,_,sb,_,_,_). outstanding_refs (Not  is_volatile) sb) ` set ls)"
    by (force simp add: l simp del: o_apply) 
  from notin_outstanding_non_volatile_takeWhile_lem [OF a_notin_sb]
   Cons.hyps [OF a_notin_ls]
  show ?case
    by (auto simp add: l simp del: o_apply)
qed

lemma flush_all_until_volatile_write_unchanged_addresses':
  assumes  notin: "a   ((λ(_,_,_,sb,_,_,_). outstanding_refs (Not  is_volatile) sb) ` set ls)"
  shows "flush_all_until_volatile_write ls m a = m a"
using notin_outstanding_non_volatile_takeWhile_Un_lem' [OF notin]
by (auto intro: flush_all_until_volatile_write_unchanged_addresses)

lemma flush_all_until_volatile_wirte_mem_independent:
  "m m'. a   ((λ(_,_,_,sb,_,_,_). outstanding_refs is_Writesb 
                     (takeWhile (Not  is_volatile_Writesb) sb)) ` set ls) 
          flush_all_until_volatile_write ls m' a = flush_all_until_volatile_write ls m a"
proof (induct ls)
  case Nil thus ?case by simp
next
  case (Cons l ls)
  obtain p "is" 𝒪  𝒟 xs sb where l: "l=(p,is,xs,sb,𝒟,𝒪,)"
    by (cases l)
  note a_in = a   ((λ(_,_,_,sb,_,_,_). outstanding_refs is_Writesb 
                     (takeWhile (Not  is_volatile_Writesb) sb)) ` set (l#ls))
  show ?case 
  proof (cases "a   ((λ(_,_,_,sb,_,_,_). outstanding_refs is_Writesb 
                     (takeWhile (Not  is_volatile_Writesb) sb)) ` set ls)") 
    case True
    from Cons.hyps [OF this]
    show ?thesis
      by (simp add: l)
  next
    case False
    with a_in
    have "a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sb)"
      by (auto simp add: l)
    from flushed_values_mem_independent [rule_format, OF this]
    have "flush (takeWhile (Not  is_volatile_Writesb) sb) m' a =
          flush (takeWhile (Not  is_volatile_Writesb) sb) m a".
    with flush_all_until_volatile_write_unchanged_addresses [OF False]
    show ?thesis
      by (auto simp add: l)
  qed
qed

lemma flush_all_until_volatile_write_buffered_val_conv: 
  assumes no_volatile_Writesb: "outstanding_refs is_volatile_Writesb sb = {}"
  shows"m i. i < length ls; ls!i = (p,is,xs,sb,𝒟,𝒪,);
               
        j < length ls. i  j 
                (let (_,_,_,sbj,_,_,_) = ls!j 
                 in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))  
      flush_all_until_volatile_write ls m a =
        (case buffered_val sb a of None  m a | Some v  v)"
proof (induct ls)
  case Nil thus ?case
    by simp
next
  case (Cons l ls)
  note i_bound =  i < length (l#ls)
  note ith = (l#ls)!i = (p,is,xs,sb,𝒟,𝒪,)
  note notin = j < length (l#ls). i  j 
                (let (_,_,_,sbj,_,_,_) = (l#ls)!j 
                 in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,is,xs,sb,𝒟,𝒪,)"
      by simp
    from no_volatile_Writesb have take_all: "takeWhile (Not  is_volatile_Writesb) sb = sb"
      by (auto simp add: outstanding_refs_conv)

    have "a  ((λ(_,_, _, sb, _,_,_).
            outstanding_refs is_Writesb
             (takeWhile (Not  is_volatile_Writesb) sb)) ` set ls)" (is "a  ?LS")
    proof 
      assume "a  ?LS" 
      from in_Union_image_nth_conv [OF this]
      obtain j pj "isj" "𝒪j" j "𝒟j" "xsj" "sbj" where 
	j_bound: "j < length ls" and
	jth: "ls!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)" and
	a_in_j: "a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	by fastforce
      from a_in_j obtain v' sop' A L R W where "Writesb False a sop' v' A L R W set (takeWhile (Not  is_volatile_Writesb) sbj)" 
  apply (clarsimp simp add: outstanding_refs_conv )
  subgoal for x
	apply (case_tac x)
	apply    clarsimp
	apply    (frule set_takeWhileD) 
	apply auto
	done
  done
      with notin [rule_format, of "Suc j"] j_bound jth
      show False
	by (force simp add: 0  outstanding_refs_conv is_non_volatile_Writesb_def 
	  split: memref.splits)
    qed
    from flush_all_until_volatile_write_unchanged_addresses [OF this]
    have "flush_all_until_volatile_write ls (flush sb m) a = (flush sb m) a"
      by (simp add: take_all)
    then 
    show ?thesis 
      by (simp add: 0 l take_all flush_buffered_val_conv)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l xsl sbl where l: "l = (pl,isl,xsl,sbl,𝒟l,𝒪l,l )"
      by (cases l)

    from i_bound ith notin
    have "flush_all_until_volatile_write ls
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m) a
          = (case buffered_val sb a of None  
               (flush (takeWhile (Not  is_volatile_Writesb) sbl) m) a | Some v  v)"
      apply -
      apply (rule Cons.hyps)
      apply (force simp add: Suc Let_def simp del: o_apply)+
      done

    moreover
    from notin [rule_format, of 0] l
    have "a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbl)"
      by (auto simp add: Let_def outstanding_refs_conv Suc )
    then
    have "a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbl)" 
      apply (clarsimp simp add: outstanding_refs_conv is_Writesb_def split: memref.splits dest: set_takeWhileD)
      apply (frule set_takeWhileD)
      apply force
      done

    from flush_unchanged_addresses [OF this]
    have "(flush (takeWhile (Not  is_volatile_Writesb) sbl) m) a = m a" .

    ultimately
    show ?thesis
      by (simp add: Suc l split: option.splits)
  qed
qed


context program
begin

abbreviation sb_concurrent_step ::
  "('p,'p store_buffer,'dirty,'owns,'rels,'shared) global_config  ('p,'p store_buffer,'dirty,'owns,'rels,'shared) global_config  bool"
    ("_ sb _" [60,60] 100)
where
  "sb_concurrent_step  
     computation.concurrent_step sb_memop_step store_buffer_step program_step (λp p' is sb. sb)"

term "x sb Y"

abbreviation (in program) sb_concurrent_steps::
  "('p,'p store_buffer,'dirty,'owns,'rels,'shared) global_config  ('p,'p store_buffer,'dirty,'owns,'rels,'shared) global_config  bool"
    ("_ sb* _" [60,60] 100)
where
"sb_concurrent_steps  sb_concurrent_step^**"

term "x sb* Y"

abbreviation sbh_concurrent_step ::
  "('p,'p store_buffer,bool,owns,rels,shared) global_config  ('p,'p store_buffer,bool,owns,rels,shared) global_config  bool"
    ("_ sbh _" [60,60] 100)
where
  "sbh_concurrent_step  
     computation.concurrent_step sbh_memop_step flush_step program_step 
      (λp p' is sb. sb @ [Progsb p p' is] )"

term "x sbh Y"

abbreviation sbh_concurrent_steps::
  "('p,'p store_buffer,bool,owns,rels,shared) global_config  ('p,'p store_buffer,bool,owns,rels,shared) global_config  bool"
    ("_ sbh* _" [60,60] 100)
where
"sbh_concurrent_steps  sbh_concurrent_step^**"

term "x sbh* Y"
end 

lemma instrs_append_Readsb:
  "instrs (sb@[Readsb volatile a t v]) = instrs sb @ [Read volatile a t]"
  by (induct sb) (auto split: memref.splits)

lemma instrs_append_Writesb:
  "instrs (sb@[Writesb volatile a sop v A L R W]) = instrs sb @ [Write volatile a sop A L R W]"
  by (induct sb) (auto split: memref.splits)

lemma instrs_append_Ghostsb:
  "instrs (sb@[Ghostsb A L R W]) = instrs sb @ [Ghost A L R W]"
  by (induct sb) (auto split: memref.splits)


lemma prog_instrs_append_Ghostsb:
  "prog_instrs (sb@[Ghostsb A L R W]) = prog_instrs sb"
  by (induct sb) (auto split: memref.splits)

lemma prog_instrs_append_Readsb:
  "prog_instrs (sb@[Readsb volatile a t v]) = prog_instrs sb "
  by (induct sb) (auto split: memref.splits)

lemma prog_instrs_append_Writesb:
  "prog_instrs (sb@[Writesb volatile a sop v A L R W]) = prog_instrs sb"
  by (induct sb) (auto split: memref.splits)

lemma hd_prog_append_Readsb:
  "hd_prog p (sb@[Readsb volatile a t v]) = hd_prog p sb"
  by (induct sb) (auto split: memref.splits)

lemma hd_prog_append_Writesb:
  "hd_prog p (sb@[Writesb volatile a sop v A L R W]) = hd_prog p sb"
  by (induct sb) (auto split: memref.splits)

lemma flush_update_other: "m. a  outstanding_refs (Not  is_volatile) sb 
        outstanding_refs (is_volatile_Writesb) sb = {} 
       flush sb (m(a:=v)) = (flush sb m)(a := v)"
  by (induct sb)
     (auto split: memref.splits if_split_asm simp add: fun_upd_twist)

lemma flush_update_other': "m. a  outstanding_refs (is_non_volatile_Writesb) sb 
        outstanding_refs (is_volatile_Writesb) sb = {} 
       flush sb (m(a:=v)) = (flush sb m)(a := v)"
  by (induct sb)
     (auto split: memref.splits if_split_asm simp add: fun_upd_twist)

lemma flush_update_other'': "m. a  outstanding_refs (is_non_volatile_Writesb) sb 
        a  outstanding_refs (is_volatile_Writesb) sb 
       flush sb (m(a:=v)) = (flush sb m)(a := v)"
  by (induct sb)
     (auto split: memref.splits if_split_asm simp add: fun_upd_twist)

lemma flush_all_until_volatile_write_update_other:     
"m. j < length ts. 
                (let (_,_,_,sbj,_,_,_) = ts!j 
                 in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)) 
   
  flush_all_until_volatile_write ts (m(a := v)) =
    (flush_all_until_volatile_write ts m)(a := v)"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons t ts)
  note notin = j < length (t#ts). 
                (let (_,_,_,sbj,_,_,_) = (t#ts)!j 
                 in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))
  hence notin': "j < length ts. 
                (let (_,_,_,sbj,_,_,_) = ts!j 
                 in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))"
    by auto

  obtain pl "isl" 𝒪l l 𝒟l xsl sbl where t: "t = (pl,isl,xsl,sbl,𝒟l,𝒪l,l)"
    by (cases t)

  have no_write: 
    "outstanding_refs (is_volatile_Writesb) (takeWhile (Not  is_volatile_Writesb) sbl) = {}"
    by (auto simp add: outstanding_refs_conv dest: set_takeWhileD)

  from notin [rule_format, of 0] t
  have a_notin: 
    "a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbl)"
    by (auto )

  from flush_update_other' [OF a_notin no_write]
  have "(flush (takeWhile (Not  is_volatile_Writesb) sbl) (m(a := v))) =
          (flush (takeWhile (Not  is_volatile_Writesb) sbl) m)(a := v)".
  with Cons.hyps [OF notin', of "(flush (takeWhile (Not  is_volatile_Writesb) sbl) m)"]
  show ?case
    by (simp add: t del: fun_upd_apply)
qed

lemma flush_all_until_volatile_write_append_non_volatile_write_commute: 
  assumes no_volatile_Writesb: "outstanding_refs is_volatile_Writesb sb = {}" 
  shows "m i. i < length ts; ts!i = (p,is,xs,sb,𝒟,𝒪,);
      j < length ts. i  j 
                (let (_,_,_,sbj,_,_,_) = ts!j 
                 in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)) 
    flush_all_until_volatile_write (ts[i := (p',is', xs, sb @ [Writesb False a sop v A L R W],𝒟', 𝒪,ℛ')]) m =
    (flush_all_until_volatile_write ts m)(a := v)"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons t ts)
  note i_bound =  i < length (t#ts)
  note ith = (t#ts)!i = (p,is,xs,sb,𝒟,𝒪,)
  note notin = j < length (t#ts). i  j 
                (let (_,_,_,sbj,_,_,_) = (t#ts)!j 
                 in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))
  show ?case
  proof (cases i)
    case 0
    from ith 0 have t: "t = (p,is,xs,sb,𝒟,𝒪,)"
      by simp
    from no_volatile_Writesb have take_all: "takeWhile (Not  is_volatile_Writesb) sb = sb"
      by (auto simp add: outstanding_refs_conv)

    from no_volatile_Writesb 
    have take_all': "takeWhile (Not  is_volatile_Writesb) (sb @ [Writesb False a sop v A L R W]) = 
            (sb @ [Writesb False a sop v A L R W])"
      by (auto simp add: outstanding_refs_conv)
    from notin 
    have "j < length ts. 
                (let (_,_,_,sbj,_,_,_) = ts!j 
                 in a  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))"
      by (auto simp add: 0)

    from flush_all_until_volatile_write_update_other [OF this]
    show ?thesis 
      by (simp add: 0 t take_all' take_all flush_append_write del: fun_upd_apply)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l xsl sbl where t: "t = (pl,isl,xsl,sbl,𝒟l,𝒪l,l)"
      by (cases t)
    from i_bound ith notin 
    have "flush_all_until_volatile_write
            (ts[n := (p',is',xs, sb @ [Writesb False a sop v A L R W], 𝒟', 𝒪,ℛ')])
            (flush (takeWhile (Not  is_volatile_Writesb) sbl) m) =
          (flush_all_until_volatile_write ts
              (flush (takeWhile (Not  is_volatile_Writesb) sbl) m))
              (a := v)"
      apply -
      apply (rule Cons.hyps) 
      apply (auto simp add: Suc simp del: o_apply)
      done

    then
    show ?thesis
      by (simp add: t Suc del: fun_upd_apply)
  qed
qed

lemma flush_all_until_volatile_write_append_unflushed: 
  assumes volatile_Writesb: "¬ outstanding_refs is_volatile_Writesb sb = {}" 
  shows "m i. i < length ts; ts!i = (p,is,xs,sb,𝒟,𝒪,) 
    flush_all_until_volatile_write (ts[i := (p',is', xs, sb @ sbx,𝒟', 𝒪,ℛ')]) m =
    (flush_all_until_volatile_write ts m)"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons l ts)
  note i_bound =  i < length (l#ts)
  note ith = (l#ts)!i = (p,is,xs,sb,𝒟,𝒪,)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,is,xs,sb,𝒟,𝒪,)"
      by simp
    from volatile_Writesb
    obtain r where r_in: "r  set sb" and volatile_r: "is_volatile_Writesb r"
      by (auto simp add: outstanding_refs_conv)
    from takeWhile_append1 [OF r_in, of "(Not  is_volatile_Writesb)" ] volatile_r

    have "(flush (takeWhile (Not  is_volatile_Writesb) (sb @ sbx)) m) =
          (flush (takeWhile (Not  is_volatile_Writesb) sb ) m)"
      by auto
    then
    show ?thesis
      by (simp add: 0 l)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l xsl sbl where l: "l = (pl,isl,xsl,sbl,𝒟l,𝒪l,l)"
      by (cases l)

    from Cons.hyps [of n] i_bound ith
    show ?thesis
      by (simp add: l Suc)
  qed
qed

lemma flush_all_until_volatile_nth_update_unused: 
  shows "m i. i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,) 
    flush_all_until_volatile_write (ts[i := (p',is',θ', sb, 𝒟', 𝒪',ℛ')]) m =
    (flush_all_until_volatile_write ts m)"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons l ts)
  note i_bound =  i < length (l#ts)
  note ith = (l#ts)!i = (p,is,θ,sb,𝒟,𝒪,)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,is,θ,sb,𝒟,𝒪,)"
      by simp
    show ?thesis
      by (simp add: 0 l)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l,l)"
      by (cases l)

    from Cons.hyps [of n] i_bound ith
    show ?thesis
      by (simp add: l Suc)
  qed
qed

lemma flush_all_until_volatile_write_append_volatile_write_commute:  
  assumes no_volatile_Writesb: "outstanding_refs is_volatile_Writesb sb = {}" 
  shows "m i. i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,) 
    flush_all_until_volatile_write
     (ts[i := (p',is', θ, sb @ [Writesb True a sop v A L R W],𝒟', 𝒪,ℛ')]) m
   = flush_all_until_volatile_write ts m"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons l ts)
  note i_bound =  i < length (l#ts)
  note ith = (l#ts)!i = (p,is,θ,sb,𝒟,𝒪,)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,is,θ,sb,𝒟,𝒪,)"
      by simp
    from no_volatile_Writesb
    have s1: "takeWhile (Not  is_volatile_Writesb) sb  = sb"
      by (auto simp add: outstanding_refs_conv)

    from no_volatile_Writesb
    have s2: "(takeWhile (Not  is_volatile_Writesb) (sb @ [Writesb True a sop v A L R W])) = sb"
      by (auto simp add: outstanding_refs_conv)

    show ?thesis
      by (simp add: 0 l s1 s2)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l,l)"
      by (cases l)

    from Cons.hyps [of n] i_bound ith
    show ?thesis
      by (simp add: l Suc)
  qed
qed

lemma reads_consistent_update: 
  "pending_write 𝒪 m. reads_consistent pending_write 𝒪 m sb  
       a  outstanding_refs (Not  is_volatile) sb       
       reads_consistent pending_write 𝒪 (m(a := v)) sb"
apply (induct sb)
apply simp
apply  (clarsimp split: memref.splits if_split_asm
         simp add: fun_upd_twist)
subgoal for sb 𝒪 m x11 addr val A R pending_write
apply (case_tac "a=addr")
apply simp
apply (fastforce simp add: fun_upd_twist)
done
done

lemma (in program) history_consistent_hd_prog: "p. history_consistent θ p' xs
 history_consistent θ (hd_prog p xs) xs"
apply (induct xs)
apply  simp
apply (auto split: memref.splits option.splits)
done

locale valid_program = program +
  fixes valid_prog
  assumes valid_prog_inv: "θp p (p',is'); valid_prog p  valid_prog p'"

lemma (in valid_program) history_consistent_appendD: 
  "θ ys p. sop  write_sops xs. valid_sop sop 
                read_tmps xs  read_tmps ys = {}  
          history_consistent θ p (xs@ys)  
           (history_consistent (θ|` (dom θ - read_tmps ys)) p xs  
            history_consistent θ (last_prog p xs) ys 
            read_tmps ys  (fst ` write_sops xs) = {})" 
proof (induct xs)
  case Nil thus ?case
    by auto
next
  case (Cons x xs)
  note valid_sops = sopwrite_sops (x # xs). valid_sop sop
  note read_tmps_dist = read_tmps (x#xs)  read_tmps ys = {}
  note consis = history_consistent θ p ((x#xs)@ys)
  show ?case
  proof (cases x)
    case (Writesb volatile a sop v)
    obtain D f where sop: "sop=(D,f)"
      by (cases sop) 
    from consis obtain 
      D_tmps: "D  dom θ" and
      f_v: "f θ = v" and
      D_read_tmps: "D  read_tmps (xs @ ys) = {}" and
      consis': "history_consistent θ p (xs @ ys)" 
      by (simp add: Writesb sop)
    from valid_sops obtain 
      valid_Df: "valid_sop (D,f)" and
      valid_sops': "sopwrite_sops xs. valid_sop sop"
      by (auto simp add: Writesb sop)
    from valid_Df
    interpret valid_sop "(D,f)" .
    from read_tmps_dist have read_tmps_dist': "read_tmps xs  read_tmps ys = {}"
      by (simp add: Writesb)


    from D_read_tmps have D_ys: "D  read_tmps ys = {}"
      by (auto simp add: read_tmps_append)
    with D_tmps have D_subset: "D  dom θ - read_tmps ys"
      by auto
    moreover
    
    from valid_sop [OF refl D_tmps]
    have "f θ = f (θ |` D)".
    moreover
    let ?θ' = "θ |` (dom θ - read_tmps ys)"
    from D_subset
    have "?θ' |` D = θ |` D" 
      apply -
      apply (rule ext)
      by (auto simp add: restrict_map_def)
    moreover
    from D_subset
    have D_tmps': "D  dom ?θ'"
      by auto
    ultimately 
    have f_v': "f ?θ' = v"
      using valid_sop [OF refl D_tmps'] f_v
      by simp
    from D_read_tmps
    have "D  read_tmps xs = {}"
      by (auto simp add: read_tmps_append)
    with Cons.hyps [OF valid_sops' read_tmps_dist' consis'] D_tmps D_subset f_v' D_ys
    show ?thesis
      by (auto simp add: Writesb sop)
  next
    case (Readsb volatile a t v)
    from consis obtain 
      tmps_t: "θ t = Some v" and
      consis': "history_consistent θ p (xs @ ys)"
      by (simp add: Readsb split: option.splits)

    from read_tmps_dist
    obtain t_ys: "t  read_tmps ys" and read_tmps_dist': "read_tmps xs  read_tmps ys = {}"
      by (auto simp add: Readsb)
    from valid_sops have valid_sops': "sopwrite_sops xs. valid_sop sop"
      by (auto simp add: Readsb)
    from t_ys tmps_t
    have "(θ |` (dom θ - read_tmps ys)) t = Some v"
      by (auto simp add: restrict_map_def domIff)
    with Cons.hyps [OF valid_sops' read_tmps_dist' consis']

    show ?thesis
      by (auto simp add: Readsb)
  next
    case (Progsb p1 p2 mis)
    from consis obtain p1_p: "p1 = p" and
     prog_step: "θ |` (dom θ - read_tmps (xs @ ys)) p1 p (p2, mis)" and
     consis': "history_consistent θ p2 (xs @ ys)"
      by (auto simp add: Progsb)

    let ?θ' = "θ |` (dom θ - read_tmps ys)"
    have eq: "?θ' |` (dom ?θ' - read_tmps xs) = θ |` (dom θ - read_tmps (xs @ ys))"
      apply (rule ext)
      apply (auto simp add: read_tmps_append restrict_map_def domIff split: if_split_asm)
      done

    from valid_sops have valid_sops': "sopwrite_sops xs. valid_sop sop"
      by (auto simp add: Progsb)
    from read_tmps_dist
    obtain read_tmps_dist': "read_tmps xs  read_tmps ys = {}"
      by (auto simp add: Progsb)
    from Cons.hyps [OF valid_sops' read_tmps_dist' consis'] p1_p prog_step eq
    show ?thesis
      by (simp add: Progsb)
  next
    case Ghostsb
    with Cons show ?thesis
      by auto
  qed
qed

lemma (in valid_program) history_consistent_appendI: 
  "θ ys p. sop  write_sops xs. valid_sop sop 
  history_consistent (θ|` (dom θ - read_tmps ys)) p xs 
 history_consistent θ (last_prog p xs) ys 
 read_tmps ys  (fst ` write_sops xs) = {}  valid_prog p 
           history_consistent θ p (xs@ys)" 
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x xs)
  note valid_sops = sopwrite_sops (x # xs). valid_sop sop
  note consis_xs = history_consistent (θ |` (dom θ - read_tmps ys)) p (x # xs)
  note consis_ys = history_consistent θ (last_prog p (x # xs)) ys
  note dist = read_tmps ys  (fst ` write_sops (x # xs)) = {}
  note valid_p = valid_prog p
  show ?case
  proof (cases x)
    case (Writesb volatile a sop v)
    obtain D f where sop: "sop=(D,f)"
      by (cases sop) 
    from consis_xs obtain 
      D_tmps: "D  dom θ - read_tmps ys" and
      f_v: "f (θ |` (dom θ - read_tmps ys)) = v" (is "f  = v") and
      D_read_tmps: "D  read_tmps xs = {}" and
      consis': "history_consistent (θ |` (dom θ - read_tmps ys)) p xs" 
      by (simp add: Writesb sop)

    from D_tmps D_read_tmps 
    have "D  read_tmps (xs @ ys) = {}"
      by (auto simp add: read_tmps_append)
    moreover
    from D_tmps have D_tmps': "D  dom θ"
      by auto
    moreover 
    from valid_sops obtain 
      valid_Df: "valid_sop (D,f)" and
      valid_sops': "sopwrite_sops xs. valid_sop sop"
      by (auto simp add: Writesb sop)
    from valid_Df
    interpret valid_sop "(D,f)" .

    from D_tmps
    have tmps_eq: "θ |` ((dom θ - read_tmps ys)  D) = θ |` D"
      apply -
      apply (rule ext) 
      apply (auto simp add: restrict_map_def)
      done
    from D_tmps
    have "f  = f ( |` D)"
      apply -
      apply (rule valid_sop [OF refl ])
      apply auto
      done
    with valid_sop [OF refl D_tmps'] f_v D_tmps

    have "f θ = v"
      by (clarsimp simp add: tmps_eq)
    moreover
    from consis_ys have consis_ys': "history_consistent θ (last_prog p xs) ys"
      by (auto simp add: Writesb)

    from dist have dist': "read_tmps ys  (fst ` write_sops xs) = {}"
      by (auto simp add: Writesb)

    moreover note Cons.hyps [OF valid_sops' consis' consis_ys' dist' valid_p]

    ultimately show ?thesis
      by (simp add: Writesb sop)
  next
    case (Readsb volatile a t v)
    from consis_xs obtain
      t_v: "(θ |` (dom θ - read_tmps ys)) t = Some v" and
      consis_xs': "history_consistent (θ |` (dom θ - read_tmps ys)) p xs"
      by (clarsimp simp add: Readsb split: option.splits)
    from t_v have "θ t = Some v"
      by (auto simp add: restrict_map_def split: if_split_asm)
    moreover
    from valid_sops obtain 
      valid_sops': "sopwrite_sops xs. valid_sop sop"
      by (auto simp add: Readsb)
    from consis_ys have consis_ys': "history_consistent θ (last_prog p xs) ys"
      by (auto simp add: Readsb)
    from dist have dist': "read_tmps ys  (fst ` write_sops xs) = {}"
      by (auto simp add: Readsb)

    note Cons.hyps [OF valid_sops' consis_xs' consis_ys' dist' valid_p]
    ultimately
    show ?thesis
      by (simp add: Readsb)
  next
    case (Progsb p1 p2 mis)
    let  = "θ |` (dom θ - read_tmps ys)"
    from consis_xs  obtain 
      p1_p: "p1 = p" and
      prog_step: " |` (dom  - read_tmps xs) p1 p (p2, mis)" and
      consis': "history_consistent  p2 xs"
      by (auto simp add: Progsb)
    
    (*let ?θ' = "θ |` (dom θ - read_tmps ys)"*)
    have eq: " |` (dom  - read_tmps xs) = θ |` (dom θ - read_tmps (xs @ ys))"
      apply (rule ext)
      apply (auto simp add: read_tmps_append restrict_map_def domIff split: if_split_asm)
      done

    from prog_step eq
    have "θ |` (dom θ - read_tmps (xs @ ys)) p1 p (p2, mis)" by simp
    moreover
    from valid_sops obtain 
      valid_sops': "sopwrite_sops xs. valid_sop sop"
      by (auto simp add: Progsb)
    from consis_ys have consis_ys': "history_consistent θ (last_prog p2 xs) ys"
      by (auto simp add: Progsb)
    from dist have dist': "read_tmps ys  (fst ` write_sops xs) = {}"
      by (auto simp add: Progsb)

    note Cons.hyps [OF valid_sops' consis' consis_ys' dist' valid_prog_inv [OF prog_step valid_p [simplified p1_p [symmetric]]]]
    ultimately
    show ?thesis
      by (simp add: Progsb p1_p)
  next
    case Ghostsb
    with Cons show ?thesis
      by auto
  qed
qed

lemma (in valid_program) history_consistent_append_conv: 
  "θ ys p. sop  write_sops xs. valid_sop sop 
                read_tmps xs  read_tmps ys = {}  valid_prog p 
          history_consistent θ p (xs@ys) = 
           (history_consistent (θ|` (dom θ - read_tmps ys)) p xs  
            history_consistent θ (last_prog p xs) ys 
            read_tmps ys  (fst ` write_sops xs) = {})"
apply rule
apply  (rule history_consistent_appendD,assumption+)
apply (rule history_consistent_appendI)
apply auto
done

lemma instrs_takeWhile_dropWhile_conv:
  "instrs xs = instrs (takeWhile P xs) @ instrs (dropWhile P xs)"
by (induct xs) (auto split: memref.splits)



lemma (in program) history_consistent_hd_prog_p: 
  "p. history_consistent θ p xs  p = hd_prog p xs"
  by (induct xs) (auto split: memref.splits option.splits)

lemma instrs_append: "ys. instrs (xs@ys) = instrs xs @ instrs ys"
  by (induct xs) (auto split: memref.splits)

lemma prog_instrs_append: "ys. prog_instrs (xs@ys) = prog_instrs xs @ prog_instrs ys"
  by (induct xs) (auto split: memref.splits)

lemma prog_instrs_empty: "r  set xs. ¬ is_Progsb r  prog_instrs xs = []"
  by (induct xs) (auto split: memref.splits)

lemma length_dropWhile [termination_simp]: "length (dropWhile P xs)  length xs"
  by (induct xs) auto

lemma prog_instrs_filter_is_Progsb: "prog_instrs (filter (is_Progsb) xs) = prog_instrs xs"
  by (induct xs) (auto split: memref.splits)


lemma Cons_to_snoc: "x. ys y. (x#xs) = (ys@[y])"
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x1 xs)
  from Cons [of x1] obtain ys y where "x1#xs = ys @ [y]"
    by auto
  then
  show ?case
    by simp
qed

lemma causal_program_history_Read:
  assumes causal_Read: "causal_program_history (Read volatile a t # issb) sb" 
  shows "causal_program_history issb (sb @ [Readsb volatile a t v])"
proof
  fix sb1 sb2
  assume sb: "sb @ [Readsb volatile a t v] = sb1 @ sb2"
  from causal_Read
  interpret causal_program_history "Read volatile a t # issb" "sb" .
  show "is. instrs sb2 @ issb = is @ prog_instrs sb2"
  proof (cases sb2)
    case Nil
    thus ?thesis
      by simp
  next
    case (Cons r sb')
    from Cons_to_snoc [of r sb'] Cons obtain ys y where sb2_snoc: "sb2=ys@[y]"
      by auto
    with sb obtain y: "y = Readsb volatile a t v" and sb: "sb = sb1@ys"
      by simp

    from causal_program_history [OF sb] obtain "is" where
      "instrs ys @ Read volatile a t # issb = is @ prog_instrs ys"
      by auto
    then show ?thesis
      by (simp add: sb2_snoc y instrs_append prog_instrs_append)
  qed
qed


lemma causal_program_history_Write:
  assumes causal_Write: "causal_program_history (Write volatile a sop A L R W# issb) sb"
  shows  "causal_program_history issb (sb @ [Writesb volatile a sop v A L R W])"
proof
  fix sb1 sb2
  assume sb: "sb @ [Writesb volatile a sop v A L R W] = sb1 @ sb2"
  from causal_Write
  interpret causal_program_history "Write volatile a sop A L R W# issb" "sb" .
  show "is. instrs sb2 @ issb = is @ prog_instrs sb2"
  proof (cases sb2)
    case Nil
    thus ?thesis
      by simp
  next
    case (Cons r sb')
    from Cons_to_snoc [of r sb'] Cons obtain ys y where sb2_snoc: "sb2=ys@[y]"
      by auto
    with sb obtain y: "y = Writesb volatile a sop v A L R W" and sb: "sb = sb1@ys"
      by simp

    from causal_program_history [OF sb] obtain "is" where
      "instrs ys @ Write volatile a sop A L R W# issb = is @ prog_instrs ys"
      by auto
    then show ?thesis
      by (simp add: sb2_snoc y instrs_append prog_instrs_append)
  qed
qed

lemma causal_program_history_Progsb:
  assumes causal_Write: "causal_program_history issb sb"
  shows  "causal_program_history (issb@mis) (sb @ [Progsb p1 p2 mis])"
proof
  fix sb1 sb2
  assume sb: "sb @ [Progsb p1 p2 mis] = sb1 @ sb2"
  from causal_Write
  interpret causal_program_history "issb" "sb" .
  show "is. instrs sb2 @ (issb@mis) = is @ prog_instrs sb2"
  proof (cases sb2)
    case Nil
    thus ?thesis
      by simp
  next
    case (Cons r sb')
    from Cons_to_snoc [of r sb'] Cons obtain ys y where sb2_snoc: "sb2=ys@[y]"
      by auto
    with sb obtain y: "y = Progsb p1 p2 mis" and sb: "sb = sb1@ys"
      by simp

    from causal_program_history [OF sb] obtain "is" where
      "instrs ys @ (issb @ mis) = is @ prog_instrs (ys@[Progsb p1 p2 mis])"
      by (auto simp add: prog_instrs_append)
    then show ?thesis
      by (simp add: sb2_snoc y instrs_append prog_instrs_append)
  qed
qed

lemma causal_program_history_Ghost:
  assumes causal_Ghostsb: "causal_program_history (Ghost A L R W # issb) sb"
  shows  "causal_program_history issb (sb @ [Ghostsb A L R W])"
proof
  fix sb1 sb2
  assume sb: "sb @ [Ghostsb A L R W] = sb1 @ sb2"
  from causal_Ghostsb
  interpret causal_program_history "Ghost A L R W # issb" "sb" .
  show "is. instrs sb2 @ issb = is @ prog_instrs sb2"
  proof (cases sb2)
    case Nil
    thus ?thesis
      by simp
  next
    case (Cons r sb')
    from Cons_to_snoc [of r sb'] Cons obtain ys y where sb2_snoc: "sb2=ys@[y]"
      by auto
    with sb obtain y: "y = Ghostsb A L R W" and sb: "sb = sb1@ys"
      by simp

    from causal_program_history [OF sb] obtain "is" where
      "instrs ys @ Ghost A L R W # issb = is @ prog_instrs ys"
      by auto
    then show ?thesis
      by (simp add: sb2_snoc y instrs_append prog_instrs_append)
  qed
qed

lemma hd_prog_last_prog_end: "p = hd_prog p sb ; last_prog p sb = psb  p = hd_prog psb sb"
  by (induct sb) (auto split: memref.splits)

lemma hd_prog_idem: "hd_prog (hd_prog p xs) xs = hd_prog p xs"
  by (induct xs) (auto split: memref.splits)

lemma last_prog_idem: "last_prog (last_prog p sb) sb = last_prog p sb"
  by (induct sb) (auto split: memref.splits)


lemma last_prog_hd_prog_append:
  "last_prog (hd_prog psb (sb@sb')) sb =last_prog (hd_prog psb sb') sb"
apply (induct sb)
apply (auto split: memref.splits)
done

lemma last_prog_hd_prog: "last_prog (hd_prog p xs) xs = last_prog p xs"
  by (induct xs) (auto split: memref.splits)


lemma last_prog_append_Readsb: 
  "p. last_prog p (sb @ [Readsb volatile a t v]) = last_prog p sb"
  by (induct sb) (auto split: memref.splits)


lemma last_prog_append_Writesb: 
  "p. last_prog p (sb @ [Writesb volatile a sop v A L R W]) = last_prog p sb"
  by (induct sb) (auto split: memref.splits)


lemma last_prog_append_Progsb:
  "x. last_prog x (sb@[Progsb p p' mis]) = p'"
  by (induct sb) (auto split: memref.splits)

lemma hd_prog_append_Progsb: "hd_prog x (sb @ [Progsb p p' mis]) = hd_prog p sb"
  by (induct sb) (auto split: memref.splits)


lemma hd_prog_last_prog_append_Progsb:
  "p'. hd_prog p' xs = p'  last_prog p' xs = p1  
       hd_prog p' (xs @ [Progsb p1 p2 mis]) = p'"
apply (induct xs)
apply (auto split: memref.splits)
done

lemma hd_prog_append_Ghostsb:
  "hd_prog p (sb@[Ghostsb A  R L W]) = hd_prog p sb"
  by (induct sb) (auto split: memref.splits)

lemma last_prog_append_Ghostsb: 
  "p. last_prog p (sb @ [Ghostsb A L R W]) = last_prog p sb"
  by (induct sb) (auto split: memref.splits)

lemma dropWhile_all_False_conv:  
"x  set xs. ¬ P x  dropWhile P xs = xs"
by (induct xs) auto

lemma dropWhile_append_all_False: 
"y  set ys. ¬ P y  
  dropWhile P (xs@ys) = dropWhile P xs @ ys"
apply (induct xs)
apply (auto simp add: dropWhile_all_False_conv)
done


lemma reads_consistent_append_first:
  "m ys. reads_consistent pending_write 𝒪 m (xs @ ys)  reads_consistent pending_write 𝒪 m xs"
  by (clarsimp simp add: reads_consistent_append)

lemma reads_consistent_takeWhile:
assumes consis: "reads_consistent pending_write 𝒪 m sb" 
shows "reads_consistent pending_write 𝒪 m (takeWhile P sb)"
using reads_consistent_append [where xs="(takeWhile P sb)" and ys="(dropWhile P sb)"] consis
apply (simp add: reads_consistent_append)
done

lemma flush_flush_all_until_volatile_write_Writesb_volatile_commute:
  "i m. i < length ts; ts!i=(p,is,xs,Writesb True a sop v A L R W#sb,𝒟,𝒪,);
        i < length ts. (j < length ts. i  j 
                  (let (_,_,_,sbi,_,_,_) = ts!i;
                      (_,_,_,sbj,_,_,_) = ts!j
                   in outstanding_refs is_Writesb sbi  
                      outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj) = {}));
        j < length ts. i  j 
                (let (_,_,_,sbj,_,_,_) = ts!j in a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))
        
       flush (takeWhile (Not  is_volatile_Writesb) sb)
         ((flush_all_until_volatile_write ts m)(a := v)) =
       flush_all_until_volatile_write (ts[i := (p,is,xs, sb, 𝒟', 𝒪',ℛ')])
         (m(a := v))"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons l ts)
  note i_bound =  i < length (l#ts)
  note ith = (l#ts)!i = (p,is,xs,Writesb True a sop v A L R W#sb,𝒟,𝒪,)
  note disj = i < length (l#ts). (j < length (l#ts). i  j 
                  (let (_,_,_,sbi,_,_,_) = (l#ts)!i;
                      (_,_,_,sbj,_,_,_) = (l#ts)!j
                   in outstanding_refs is_Writesb sbi  
                      outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj) = {}))
  note a_notin = j < length (l#ts). i  j 
                (let (_,_,_,sbj,_,_,_) = (l#ts)!j 
                 in a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,is,xs,Writesb True a sop v A L R W#sb,𝒟,𝒪,)"
      by simp
    have a_notin_ts:
      "a  ((λ(_,_,_,sb,_,_,_). outstanding_refs is_Writesb 
                            (takeWhile (Not  is_volatile_Writesb) sb)) ` set ts)" (is "a  ?U")
    proof 
      assume "a  ?U"
      from in_Union_image_nth_conv [OF this]
      obtain j pj "isj" "𝒪j" j 𝒟j "xsj" "sbj" where 
	j_bound: "j < length ts" and
	jth: "ts!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)" and
	a_in_j: "a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	by fastforce
      from a_notin [rule_format, of "Suc j"] j_bound 0 a_in_j
      show False
	by (auto simp add: jth)
    qed

    from a_notin_ts
    have "(flush_all_until_volatile_write ts m)(a := v) =
                flush_all_until_volatile_write ts (m(a := v))"
      apply -
      apply (rule update_commute' [where F="{a}" and G="?U" and 
	g="flush_all_until_volatile_write ts"])
      apply (auto intro: flush_all_until_volatile_wirte_mem_independent
             flush_all_until_volatile_write_unchanged_addresses)
      done
    
    moreover

    let ?SB = "outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sb)"

    have U_SB_disj: "?U  ?SB = {}"
    proof -
      {
	fix a'
	assume a'_in_U: "a'  ?U"
	have "a'  ?SB"
	proof 
	  assume a'_in_SB: "a'  ?SB"
	  hence a'_in_SB': "a'  outstanding_refs is_Writesb sb"
	    apply (clarsimp simp add: outstanding_refs_conv)
	    apply (drule set_takeWhileD)
      subgoal for x
	    apply (rule_tac x=x in exI)
	    apply (auto simp add: is_Writesb_def split:memref.splits)
	    done
	    done
	  from in_Union_image_nth_conv [OF a'_in_U]
	  obtain j pj "isj" "𝒪j" j 𝒟j "xsj" "sbj" where 
	    j_bound: "j < length ts" and
	    jth: "ts!j = (pj,isj,xsj,sbj,𝒟j,𝒪j,j)" and
	    a'_in_j: "a'  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	    by fastforce

	  from disj [rule_format, of 0 "Suc j"] 0 j_bound a'_in_SB' a'_in_j jth l
	  show False
	    by auto
	qed
      }
      moreover
      {
	fix a'
	assume a'_in_SB: "a'  ?SB"
	hence a'_in_SB': "a'  outstanding_refs is_Writesb sb"
	    apply (clarsimp simp add: outstanding_refs_conv)
	    apply (drule set_takeWhileD)
	    subgoal for x
	    apply (rule_tac x=x in exI)
	    apply (auto simp add: is_Writesb_def split:memref.splits)
	    done
	    done
	have "a'  ?U"
	proof 
	  assume "a'  ?U"
	  from in_Union_image_nth_conv [OF this]
	  obtain j pj "isj" "𝒪j" j 𝒟j "xsj" "sbj" where 
	    j_bound: "j < length ts" and
	    jth: "ts!j = (pj,isj,xsj,sbj,𝒟j,j,𝒪j)" and
	    a'_in_j: "a'  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	    by fastforce

	  from disj [rule_format, of 0 "Suc j"] j_bound a'_in_SB' a'_in_j jth  l
	  show False
	    by auto
	qed
      }
      ultimately
      show ?thesis by blast
    qed

    have "flush (takeWhile (Not  is_volatile_Writesb) sb)
           (flush_all_until_volatile_write ts (m(a := v))) = 
          flush_all_until_volatile_write ts 
           (flush (takeWhile (Not  is_volatile_Writesb) sb) (m(a := v)))"
      apply (rule update_commute' [where g = "flush_all_until_volatile_write ts ",
             OF _ _ _ _ U_SB_disj])
      apply (auto intro: flush_all_until_volatile_wirte_mem_independent
             flush_all_until_volatile_write_unchanged_addresses
             flush_unchanged_addresses
             flushed_values_mem_independent simp del: o_apply)
      done
      
    ultimately
    have "flush (takeWhile (Not  is_volatile_Writesb) sb)
           ((flush_all_until_volatile_write ts m)(a := v)) =
          flush_all_until_volatile_write ts
           (flush (takeWhile (Not  is_volatile_Writesb) sb) (m(a := v)))"
      by simp

    then show ?thesis 
      by (auto simp add: l 0 o_def simp del: fun_upd_apply)
  next
    case (Suc n)

    obtain pl "isl" 𝒪l l 𝒟j xsl sbl where l: "l = (pl,isl,xsl,sbl,𝒟j,𝒪l,l)"
      by (cases l)

    from i_bound ith disj a_notin
    have  
      "flush (takeWhile (Not  is_volatile_Writesb) sb)
        ((flush_all_until_volatile_write ts
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m))
          (a := v)) =
       flush_all_until_volatile_write (ts[n := (p,is, xs, sb,𝒟', 𝒪',ℛ')])
        ((flush (takeWhile (Not  is_volatile_Writesb) sbl) m)(a := v))"
      apply -
      apply (rule Cons.hyps)
      apply (force simp add: Suc Let_def simp del: o_apply)+
      done

    moreover

    let ?SB = "outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbl)"
    have "a  ?SB"
    proof 
      assume "a  ?SB"
      with a_notin [rule_format, of 0]
      show False
	by (auto simp add: l Suc)
    qed
    then
    have "((flush (takeWhile (Not  is_volatile_Writesb) sbl) m)(a := v)) =
          (flush (takeWhile (Not  is_volatile_Writesb) sbl) (m(a := v)))"
      apply -
      apply (rule update_commute' [where m=m and F="{a}" and G="?SB"])
      apply (auto intro: 
             flush_unchanged_addresses
             flushed_values_mem_independent simp del: o_apply)
      done

    ultimately
    show ?thesis
      by (simp add: l Suc del: fun_upd_apply o_apply)
  qed
qed



   




lemma (in program)
"sb' p. history_consistent θ (hd_prog p (sb@sb')) (sb@sb') 
          last_prog p (sb@sb') = p  
 last_prog (hd_prog p (sb@sb')) sb = hd_prog p sb'"
proof (induct sb)
  case Nil thus ?case by simp
next
  case (Cons r sb1)
  have consis: "history_consistent θ (hd_prog p ((r # sb1) @ sb')) ((r # sb1) @ sb')" 
    by fact
  have last_prog: "last_prog p ((r # sb1) @ sb') = p" by fact
  show ?case
  proof (cases r)
    case Writesb with Cons show ?thesis by auto
  next
    case Readsb with Cons show ?thesis by (auto split: option.splits)
  next
    case (Progsb p1 p2 "is")
    from last_prog have last_prog_p2: "last_prog p2 (sb1 @ sb') = p"
      by (simp add: Progsb)
    from consis obtain consis': "history_consistent θ p2 (sb1 @ sb')"
      by (simp add: Progsb)

    hence "history_consistent θ (hd_prog p2 (sb1 @ sb')) (sb1 @ sb')"
      by (rule history_consistent_hd_prog)
    from Cons.hyps [OF this ]
    have "last_prog p2 sb1 = hd_prog p sb'"
      oops

lemma last_prog_to_last_prog_same: "p'. last_prog p' sb = p  last_prog p sb = p"
  by (induct sb) (auto split: memref.splits)

lemma last_prog_hd_prog_same: "last_prog p' sb = p; hd_prog p' sb = p'  hd_prog p sb = p'"
  by (induct sb) (auto split : memref.splits)

lemma last_prog_hd_prog_last_prog:   
  "last_prog p' (sb@sb') = p  hd_prog p' (sb@sb') = p' 
   last_prog (hd_prog p sb') sb = last_prog p' sb"
apply (induct sb)
apply (simp add: last_prog_hd_prog_same)
apply (auto split : memref.splits)
done

lemma (in program) last_prog_hd_prog_append':
"sb' p. history_consistent θ (hd_prog p (sb@sb')) (sb@sb') 
          last_prog p (sb@sb') = p  
 last_prog (hd_prog p sb') sb = hd_prog p sb'"
proof (induct sb)
  case Nil thus ?case by simp
next
  case (Cons r sb1)
  have consis: "history_consistent θ (hd_prog p ((r # sb1) @ sb')) ((r # sb1) @ sb')" 
    by fact
  have last_prog: "last_prog p ((r # sb1) @ sb') = p" by fact
  show ?case
  proof (cases r)
    case Writesb with Cons show ?thesis by auto
  next
    case Readsb with Cons show ?thesis by (auto split: option.splits)
  next
    case (Progsb p1 p2 "is")
    from last_prog have last_prog_p2: "last_prog p2 (sb1 @ sb') = p"
      by (simp add: Progsb)
    from last_prog_to_last_prog_same [OF this]
    have last_prog_p: "last_prog p (sb1 @ sb') = p".
    from consis obtain consis': "history_consistent θ p2 (sb1 @ sb')"
      by (simp add: Progsb)
    from history_consistent_hd_prog_p [OF consis']
    have hd_prog_p2: "hd_prog p2 (sb1 @ sb') = p2" by simp
    from consis' have "history_consistent θ (hd_prog p (sb1 @ sb')) (sb1 @ sb')"
      by (rule history_consistent_hd_prog)
    from Cons.hyps [OF this last_prog_p]
    have "last_prog (hd_prog p sb') sb1 = hd_prog p sb'".
    moreover
    from last_prog_hd_prog_last_prog [OF last_prog_p2 hd_prog_p2]
    have "last_prog (hd_prog p sb') sb1 = last_prog p2 sb1".
    ultimately
    have "last_prog p2 sb1 = hd_prog p sb'"
      by simp
    thus ?thesis
      by (simp add: Progsb)
  next
    case Ghostsb with Cons show ?thesis by (auto split: option.splits)
  qed
qed

lemma flush_all_until_volatile_write_Writesb_non_volatile_commute:
  "i m. i < length ts; ts!i=(p,is,xs,Writesb False a sop v A L R W#sb,𝒟,𝒪,);
        i < length ts. (j < length ts. i  j 
                  (let (_,_,_,sbi,_,_,_) = ts!i;
                      (_,_,_,sbj,_,_,_) = ts!j
                   in outstanding_refs is_Writesb sbi  
                      outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj) = {}));
        j < length ts. i  j 
                (let (_,_,_,sbj,_,_,_) = ts!j in a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))
        flush_all_until_volatile_write (ts[i := (p,is, xs, sb,𝒟', 𝒪,ℛ')])(m(a := v))  =
               flush_all_until_volatile_write ts m"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons l ts)
  note i_bound =  i < length (l#ts)
  note ith = (l#ts)!i = (p,is,xs,Writesb False a sop v A L R W#sb,𝒟,𝒪,)
  note disj = i < length (l#ts). (j < length (l#ts). i  j 
                  (let (_,_,_,sbi,_,_,_) = (l#ts)!i;
                      (_,_,_,sbj,_,_,_) = (l#ts)!j
                   in outstanding_refs is_Writesb sbi  
                      outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj) = {}))
  note a_notin = j < length (l#ts). i  j 
                (let (_,_,_,sbj,_,_,_) = (l#ts)!j 
    in a  outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbj))
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,is,xs,Writesb False a sop v A L R W#sb,𝒟,𝒪,)"
      by simp
    thus ?thesis 
      by (simp add: 0 del: fun_upd_apply)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l xsl sbl where l: "l = (pl,isl,xsl,sbl,𝒟l,𝒪l,l)"
      by (cases l)

    from i_bound ith disj a_notin 
    have
      "flush_all_until_volatile_write (ts[n := (p,is,xs, sb, 𝒟', 𝒪,ℛ')])
          ((flush (takeWhile (Not  is_volatile_Writesb) sbl) m)(a := v)) =
       flush_all_until_volatile_write ts
          (flush (takeWhile (Not  is_volatile_Writesb) sbl) m)"
      apply -
      apply (rule Cons.hyps)
      apply (force simp add: Suc Let_def simp del: o_apply)+
      done

    moreover

    let ?SB = "outstanding_refs is_Writesb (takeWhile (Not  is_volatile_Writesb) sbl)"
    have "a  ?SB"
    proof 
      assume "a  ?SB"
      with a_notin [rule_format, of 0]
      show False
	by (auto simp add: l Suc)
    qed
    then
    have "((flush (takeWhile (Not  is_volatile_Writesb) sbl) m)(a := v)) =
          (flush (takeWhile (Not  is_volatile_Writesb) sbl) (m(a := v)))"
      apply -
      apply (rule update_commute' [where m=m and F="{a}" and G="?SB"])
      apply (auto intro: 
             flush_unchanged_addresses
             flushed_values_mem_independent simp del: o_apply)
      done

    ultimately
    show ?thesis
      by (simp add: l Suc del: fun_upd_apply o_apply)
  qed
qed   

lemma (in program) history_consistent_access_last_read': 
  "p. history_consistent θ p (sb @ [Readsb volatile a t v]) 
        θ t = Some v"
apply (induct sb)
apply (auto  split: memref.splits option.splits)
done

lemma (in program) history_consistent_access_last_read:
  "history_consistent θ p (rev (Readsb volatile a t v # sb))  θ t = Some v"
  by (simp add: history_consistent_access_last_read')

lemma flush_all_until_volatile_write_Readsb_commute:
  "i m. i < length ts; ts!i=(p,is,θ,Readsb volatile a t v#sb,𝒟,𝒪,)
        flush_all_until_volatile_write (ts[i := (p,is,θ, sb, 𝒟', 𝒪,ℛ')]) m
       = flush_all_until_volatile_write ts m"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons l ts)
  note i_bound =  i < length (l#ts)
  note ith = (l#ts)!i = (p,is,θ,Readsb volatile a t v#sb,𝒟,𝒪,)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,is,θ,Readsb volatile a t v#sb,𝒟,𝒪,)"
      by simp
    thus ?thesis 
      by (simp add: 0 del: fun_upd_apply)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l,l)"
      by (cases l)

    from i_bound ith
    have "flush_all_until_volatile_write (ts[n := (p,is,θ, sb, 𝒟', 𝒪,ℛ')])
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m) =
         flush_all_until_volatile_write ts
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m)"
      apply -
      apply (rule Cons.hyps)
      apply (auto simp add: Suc l)
      done

    then show ?thesis
      by (simp add: Suc l)
  qed
qed

lemma flush_all_until_volatile_write_Ghostsb_commute:
  "i m. i < length ts; ts!i=(p,is,θ,Ghostsb A L R W#sb,𝒟,𝒪,)
        flush_all_until_volatile_write (ts[i := (p',is',θ', sb, 𝒟', 𝒪',ℛ')]) m
       = flush_all_until_volatile_write ts m"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons l ts)
  note i_bound =  i < length (l#ts)
  note ith = (l#ts)!i = (p,is,θ,Ghostsb A L R W#sb,𝒟,𝒪,)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,is,θ,Ghostsb A L R W#sb,𝒟,𝒪,)"
      by simp
    thus ?thesis 
      by (simp add: 0 del: fun_upd_apply)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l,l)"
      by (cases l)

    from i_bound ith
    have "flush_all_until_volatile_write (ts[n := (p',is',θ', sb, 𝒟', 𝒪',ℛ')])
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m) =
         flush_all_until_volatile_write ts
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m)"
      apply -
      apply (rule Cons.hyps)
      apply (auto simp add: Suc l)
      done

    then show ?thesis
      by (simp add: Suc l)
  qed
qed

lemma flush_all_until_volatile_write_Progsb_commute:
  "i m. i < length ts; ts!i=(p,is,θ,Progsb p1 p2 mis#sb,𝒟,𝒪,)
        flush_all_until_volatile_write (ts[i := (p,is, θ, sb,𝒟', 𝒪,ℛ')]) m
       = flush_all_until_volatile_write ts m"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons l ts)
  note i_bound =  i < length (l#ts)
  note ith = (l#ts)!i = (p,is,θ,Progsb p1 p2 mis#sb,𝒟,𝒪,)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,is,θ,Progsb p1 p2 mis#sb,𝒟,𝒪,)"
      by simp
    thus ?thesis 
      by (simp add: 0 del: fun_upd_apply)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l,l)"
      by (cases l)

    from i_bound ith
    have "flush_all_until_volatile_write (ts[n := (p,is, θ, sb,𝒟', 𝒪,ℛ')])
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m) =
         flush_all_until_volatile_write ts
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m)"
      apply -
      apply (rule Cons.hyps)
      apply (auto simp add: Suc l)
      done

    then show ?thesis
      by (simp add: Suc l)
  qed
qed


lemma flush_all_until_volatile_write_append_Progsb_commute:
  "i m. i < length ts; ts!i=(p,is,θ,sb,𝒟,𝒪,)
        flush_all_until_volatile_write (ts[i := (p2,is@mis, θ, sb@[Progsb p1 p2 mis],𝒟', 𝒪,ℛ')]) m
       = flush_all_until_volatile_write ts m"
proof (induct ts)
  case Nil thus ?case
    by simp
next
  case (Cons l ts)
  note i_bound =  i < length (l#ts)
  note ith = (l#ts)!i = (p,is,θ,sb,𝒟,𝒪,)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,is,θ,sb,𝒟,𝒪,)"
      by simp
    thus ?thesis 
      by (simp add: 0 flush_append_Progsb del: fun_upd_apply)
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l,l)"
      by (cases l)

    from i_bound ith
    have "flush_all_until_volatile_write 
              (ts[n := (p2,is@mis,θ, sb@[Progsb p1 p2 mis], 𝒟', 𝒪,ℛ')])
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m) =
         flush_all_until_volatile_write ts
           (flush (takeWhile (Not  is_volatile_Writesb) sbl) m)"
      apply -
      apply (rule Cons.hyps)
      apply (auto simp add: Suc l)
      done

    then show ?thesis
      by (simp add: Suc l)
  qed
qed




lemma (in program) history_consistent_append_Progsb:
  assumes step: "θ p p (p', mis)"
  shows "history_consistent θ (hd_prog p xs) xs  last_prog p xs = p 
       history_consistent θ (hd_prog p' (xs@[Progsb p p' mis])) (xs@[Progsb p p' mis])"
proof (induct xs)
  case Nil with step show ?case by simp
next
  case (Cons x xs)
  note consis = history_consistent θ (hd_prog p (x # xs)) (x # xs)
  note last = last_prog p (x#xs) = p
  show ?case
  proof (cases x)
    case Writesb with Cons show ?thesis by (auto simp add: read_tmps_append)
  next
    case Readsb with Cons show ?thesis by (auto split: option.splits)
  next
    case (Progsb p1 p2 mis')
    from consis obtain
      step: "θ |`(dom θ - read_tmps (xs @ [Progsb p p' mis])) p1 p (p2, mis')" and
      consis': "history_consistent θ p2 xs"
      by (auto simp add: Progsb read_tmps_append)
    from last have last_p2: "last_prog p2 xs = p"
      by (simp add: Progsb)
    from last_prog_to_last_prog_same [OF this]
    have last_prog': "last_prog p xs = p".
    from history_consistent_hd_prog [OF consis']
    have consis'': "history_consistent θ (hd_prog p xs) xs".
    from Cons.hyps [OF this last_prog']
    have "history_consistent θ (hd_prog p' (xs @ [Progsb p p' mis]))
            (xs @ [Progsb p p' mis])".
    from history_consistent_hd_prog [OF this]
    have "history_consistent θ (hd_prog p2 (xs @ [Progsb p p' mis])) 
           (xs @ [Progsb p p' mis])".
    moreover
    from history_consistent_hd_prog_p [OF consis'] 
    have "p2 = hd_prog p2 xs".
    from hd_prog_last_prog_append_Progsb [OF this [symmetric] last_p2]
    have "hd_prog p2 (xs @ [Progsb p p' mis]) = p2"
      by simp
    ultimately
    have "history_consistent θ p2 (xs @ [Progsb p p' mis])"
      by simp
    thus ?thesis
      by (simp add: Progsb step)
  next
    case Ghostsb with Cons show ?thesis by (auto)
  qed
qed


(* FIXME: consistent naming: acquired vs. acquire; released vs. release *)
(* augment_rels, really only depends on the owned part of dom 𝒮. *)
primrec release :: "'a memref list  addr set  rels  rels"
where
"release [] S  = "
| "release (x#xs) S  =
  (case x of
     Writesb volatile _ _ _ A L R W  
        (if volatile then release xs (S  R - L) Map.empty 
         else release xs S )
   | Ghostsb A L R W  release xs (S  R - L) (augment_rels S R ) 
   | _  release xs S )"


lemma augment_rels_shared_exchange: "a  R. (a  S') = (a  S)  augment_rels S R  = augment_rels S' R "
apply (rule ext)
apply (auto simp add: augment_rels_def split: option.splits)
done


lemma sharing_consistent_shared_exchange: 
assumes shared_eq: "a  all_acquired sb. 𝒮' a = 𝒮 a"
assumes consis: "sharing_consistent 𝒮 𝒪 sb" 
shows "sharing_consistent 𝒮' 𝒪 sb"
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 
	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
        shared_eq: "a  A  all_acquired sb. 𝒮' a = 𝒮 a"
	by (clarsimp simp add: Writesb True )
      from shared_eq
      have shared_eq': "a 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 "sharing_consistent (𝒮' ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb".
      thus ?thesis
      using A_shared_owns L_A A_R R_owns shared_eq
        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 
      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
      shared_eq: "a  A  all_acquired sb. 𝒮' a = 𝒮 a"
      by (clarsimp simp add: Ghostsb )
    from shared_eq
    have shared_eq': "aall_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 "sharing_consistent (𝒮' ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb".
    thus ?thesis
    using A_shared_owns L_A A_R R_owns shared_eq
      by (auto  simp add: Ghostsb domIff)
  qed
qed



lemma release_shared_exchange: 
assumes shared_eq: "a  𝒪  all_acquired sb. 𝒮' a = 𝒮 a"
assumes consis: "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 
	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
        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 
      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
      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 A_shared_owns 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 release_append: 
"𝒮 . release (sb@xs) (dom 𝒮)  = release xs (dom (share sb 𝒮)) (release sb (dom (𝒮)) )"
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
      from Cons.hyps [of "(𝒮 ⊕⇘WR ⊖⇘AL)" "Map.empty"]
      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) 
    with Cons.hyps [of "(𝒮 ⊕⇘WR ⊖⇘AL)" "augment_rels (dom 𝒮) R "]
    show ?thesis
      by (clarsimp simp add: Ghostsb)
  qed
qed

locale xvalid_program = valid_program +
  fixes valid
  assumes valid_implies_valid_prog:     
     "i < length ts; 
      ts!i = (p,is,θ,sb,𝒟,𝒪,); valid ts  valid_prog p" 

  assumes valid_implies_valid_prog_hd:     
     "i < length ts; 
      ts!i = (p,is,θ,sb,𝒟,𝒪,); valid ts  valid_prog (hd_prog p sb)" 
  assumes distinct_load_tmps_prog_step: 
    "i < length ts; 
      ts!i = (p,is,θ,sb,𝒟,𝒪,); θp p (p',is'); valid ts 
     
      distinct_load_tmps is'  
      (load_tmps is'  load_tmps is = {}) 
      (load_tmps is'  read_tmps sb) = {}"

  assumes valid_data_dependency_prog_step: 
    "i < length ts;
      ts!i = (p,is,θ,sb,𝒟,𝒪,); θp p (p',is'); valid ts 
      
     data_dependency_consistent_instrs (dom θ  load_tmps is) is'  
     load_tmps is'  (fst ` store_sops is)  = {} 
     load_tmps is'  (fst ` write_sops sb)  = {}"

  assumes load_tmps_fresh_prog_step:
  "i < length ts;
      ts!i = (p,is,θ,sb,𝒟,𝒪,); θp p (p',is'); valid ts 
    
   load_tmps is'  dom θ = {}"

  assumes valid_sops_prog_step:
      "θp p (p',is'); valid_prog p sopstore_sops is'. valid_sop sop"

  assumes prog_step_preserves_valid:
      "i < length ts;
        ts!i = (p,is,θ,sb,𝒟,𝒪,); θp p (p',is'); valid ts 
        valid (ts[i:=(p',is@is',θ,sb@[Progsb p p' is'],𝒟,𝒪,)])"

  assumes flush_step_preserves_valid:
      "i < length ts;
        ts!i = (p,is,θ,sb,𝒟,𝒪,); (m,sb,𝒪,,𝒮) f (m',sb',𝒪',ℛ',𝒮'); valid ts 
        valid (ts[i:=(p,is,θ,sb',𝒟,𝒪',ℛ')])"

  assumes sbh_step_preserves_valid:
      "i < length ts;
        ts!i = (p,is,θ,sb,𝒟,𝒪,); 
        (is,θ,sb,m,𝒟,𝒪,,𝒮) sbh (is',θ',sb',m',𝒟',𝒪',ℛ',𝒮'); 
       valid ts 
       
       valid (ts[i:=(p,is',θ',sb',𝒟',𝒪',ℛ')])"



lemma refl': "x = y  r^** x y"
  by auto



lemma no_volatile_Readsb_volatile_reads_consistent:
  "m. outstanding_refs is_volatile_Readsb sb = {}  volatile_reads_consistent m sb"
  apply (induct sb)
  apply  simp
  subgoal for a sb m
  apply (case_tac a)
  apply (auto split: if_split_asm)
  done
  done


theorem (in program) flush_store_buffer_append:
shows "ts p m θ 𝒪  𝒟 𝒮 is  𝒪'. 
 i < length ts;
  instrs (sb@sb') @ issb = is @ prog_instrs (sb@sb');
  causal_program_history issb (sb@sb');
  ts!i = (p,is,θ |` (dom θ - read_tmps (sb@sb')),x,𝒟,𝒪,);
  p=hd_prog psb (sb@sb');
  (last_prog psb (sb@sb')) = psb;
  reads_consistent True 𝒪' m sb; 
  history_consistent θ p (sb@sb');
  sop  write_sops sb. valid_sop sop;
  distinct_read_tmps (sb@sb');
  volatile_reads_consistent m sb

  
  is'. instrs sb' @ issb = is' @ prog_instrs sb' 
     (ts,m,𝒮) d* 
     (ts[i:=(last_prog (hd_prog psb sb') sb,is',θ|` (dom θ - read_tmps sb'),x,
       (𝒟  outstanding_refs is_volatile_Writesb sb  {}),
       acquired True sb 𝒪, release sb (dom 𝒮) )], flush sb m,share sb 𝒮)"
proof (induct sb)
  case Nil
  thus ?case by (auto simp add: list_update_id' split: if_split_asm)
next
case (Cons r sb)
  interpret direct_computation:
    computation direct_memop_step empty_storebuffer_step program_step "λp p' is sb. sb".
  have ts_i: 
    "ts!i = (p,is,θ |` (dom θ - read_tmps ((r#sb)@sb')),x,𝒟,𝒪,)"
    by fact
  have "is": "instrs ((r # sb) @ sb') @ issb = is @ prog_instrs ((r # sb) @ sb')" by fact

 
  have i_bound: "i < length ts" by fact
  have causal: "causal_program_history issb ((r # sb) @ sb')" by fact
  hence causal': "causal_program_history issb (sb @ sb')" 
    by (auto simp add: causal_program_history_def)

  note reads_consis = reads_consistent True 𝒪' m (r#sb)
  note p = p = hd_prog psb ((r#sb)@sb')
  note psb = last_prog psb ((r # sb) @ sb') = psb
  note hist_consis = history_consistent θ p ((r#sb)@sb')
  note valid_sops = sop  write_sops (r#sb). valid_sop sop
  note dist = distinct_read_tmps ((r#sb)@sb')
  note vol_read_consis = volatile_reads_consistent m (r#sb)

  show ?case
  proof (cases r)
    case (Progsb p1 p2 pis)

    from vol_read_consis
    have vol_read_consis': "volatile_reads_consistent m sb"
      by (auto simp add: Progsb)

    from hist_consis  obtain
      prog_step: "θ|` (dom θ - read_tmps (sb @ sb')) p1 p (p2, pis)" and
      hist_consis': "history_consistent θ p2 (sb @ sb')" 
      by (auto simp add: Progsb)
    from p obtain  p: "p = p1" 
      by (simp add: Progsb)

    from history_consistent_hd_prog [OF hist_consis']
    have hist_consis'': "history_consistent θ (hd_prog p2 (sb @ sb')) (sb @ sb')" .
    from "is"
    have "is": "instrs (sb @ sb') @ issb = (is @ pis) @ prog_instrs (sb @ sb')"
      by (simp add: Progsb)
    
    
    from ts_i "is" have 
      ts_i: "ts!i = (p, is,θ |` (dom θ - read_tmps (sb @ sb')), x, 𝒟, 𝒪,)"
      by (simp add: Progsb)
   
    let ?ts'= "ts[i:=(p2,is@pis,θ |` (dom θ - read_tmps (sb @ sb')), x,𝒟,𝒪,)]"
    from direct_computation.Program [OF i_bound ts_i prog_step [simplified p[symmetric]]]
    have "(ts,m,𝒮) d (?ts',m,𝒮)" by simp    

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

    from i_bound
    have ts'_i: "?ts'!i = (p2,is@pis,(θ |` (dom θ - read_tmps (sb @ sb'))),x, 𝒟, 𝒪,)"
      by auto

    from history_consistent_hd_prog_p [OF hist_consis'] 
    have p2_hd_prog: " p2 = hd_prog p2 (sb @ sb')".

    from reads_consis have reads_consis': "reads_consistent True 𝒪' m sb"
      by (simp add: Progsb)

    from valid_sops have valid_sops': "sop  write_sops sb. valid_sop sop"
      by (simp add: Progsb)

    from dist have dist': "distinct_read_tmps (sb@sb')"
      by (simp add: Progsb)
    

    from psb have last_prog_p2: "last_prog p2 (sb @ sb') = psb"
      by (simp add: Progsb)
    from hd_prog_last_prog_end [OF p2_hd_prog this]
    have p2_hd_prog': "p2 = hd_prog psb (sb @ sb')".
    from last_prog_p2 [symmetric] have last_prog': "last_prog psb (sb @ sb') = psb"
      by (simp add: last_prog_idem)
    

    from Cons.hyps [OF i_bound'  "is" causal' ts'_i p2_hd_prog' last_prog' reads_consis' 
      hist_consis' valid_sops' dist' vol_read_consis'] i_bound 
    obtain is' where
      is': "instrs sb' @ issb = is' @ prog_instrs sb'" and
      step: "(?ts', m,𝒮) d*
         (ts[i := (last_prog (hd_prog psb sb') sb, is',
          θ |` (dom θ - read_tmps sb'), x, 𝒟  outstanding_refs is_volatile_Writesb sb  {},
          acquired True sb 𝒪,release sb (dom 𝒮) )],
          flush sb m,share sb 𝒮 )"
      by (auto)
    from p2_hd_prog' 
    have last_prog_eq: "last_prog (hd_prog psb sb') sb = last_prog p2 sb"
      by (simp add: last_prog_hd_prog_append)
    note step
    finally show ?thesis
      using is' 
      by (simp add: Progsb last_prog_eq)
  next
    case (Writesb volatile a sop v A L R W)
    obtain D f where sop: "sop=(D,f)"
      by (cases sop)


    from vol_read_consis
    have vol_read_consis': "volatile_reads_consistent (m(a:=v)) sb"
      by (auto simp add: Writesb)

    from hist_consis obtain 
      D_tmps: "D  dom θ" and
      f_v: "f θ = v" and
      dep: "D  read_tmps (sb@sb') = {}" and
      hist_consis': "history_consistent θ p (sb@sb')"
      by (simp add: Writesb sop split: option.splits)

    from dist have dist': "distinct_read_tmps (sb@sb')" by (auto simp add: Writesb)

    from valid_sops obtain "valid_sop sop" and
      valid_sops': "sop  write_sops sb. valid_sop sop" 
      by (simp add: Writesb)
    interpret valid_sop sop by fact
    from valid_sop [OF sop D_tmps]
    have "f θ = f (θ |` D)" .
    moreover
    from dep D_tmps have D_subset: "D  (dom θ - read_tmps (sb@sb'))"
      by auto
    moreover from D_subset have "(θ|`(dom θ - read_tmps (sb@sb')) |` D) = θ |` D"
      apply -
      apply (rule ext)
      apply (auto simp add: restrict_map_def)
      done
    moreover from D_subset D_tmps have "D  dom (θ |` (dom θ - read_tmps (sb@sb')))"
      by simp
    moreover
    note valid_sop [OF sop this] 
    ultimately have f_v': "f (θ|`(dom θ - read_tmps (sb@sb'))) = v"
      by (simp add: f_v)

    interpret causal': causal_program_history "issb" "sb@sb'" by fact

    from "is"
    have "Write volatile a sop A L R W# instrs (sb @ sb') @ issb = is @ prog_instrs (sb @ sb')"
      by (simp add: Writesb)
    with causal'.causal_program_history [of "[]", simplified, OF refl]    
    obtain is' where "is": "is=Write volatile a sop A L R W#is'" and
      is': "instrs (sb @ sb') @ issb = is' @ prog_instrs (sb @ sb')"
      by auto

    from ts_i "is"
    have ts_i: "ts!i = (p,Write volatile a sop A L R W#is',
      θ |` (dom θ - read_tmps (sb@sb')),x,𝒟,𝒪,)"
      by (simp add: Writesb)
    
    from p have p': "p = hd_prog psb (sb@sb')"
      by (auto simp add: Writesb hd_prog_idem)

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

    show ?thesis
    proof (cases volatile)
      case False
      have memop_step:
	"(Write volatile a sop A L R W#is',θ|`(dom θ - read_tmps (sb@sb')),
                 x,m,𝒟,𝒪,,𝒮)  
           (is',θ|` (dom θ - read_tmps (sb@sb')),x,m(a:=v),𝒟,𝒪,,𝒮)"
	using D_subset
	apply (simp only: sop f_v' [symmetric] False)
	apply (rule direct_memop_step.WriteNonVolatile)
	done
    
      let ?ts' = "ts[i := (p, is',θ |` (dom θ - read_tmps (sb @ sb')),x, 𝒟, 𝒪,)]"
      from direct_computation.Memop [OF i_bound ts_i  memop_step]
      have "(ts, m, 𝒮) d (?ts', m(a := v), 𝒮)".

      also
      from reads_consis have reads_consis': "reads_consistent True 𝒪' (m(a:=v)) sb"
	by (auto simp add: Writesb False)
      from i_bound have i_bound': "i < length ?ts'"
	by auto
    
      from i_bound
      have ts'_i: "?ts' ! i = (p, is',θ |` (dom θ - read_tmps (sb @ sb')), x, 𝒟, 𝒪,)"
	by simp      

      from Cons.hyps [OF i_bound' is' causal' ts'_i p' psb' reads_consis' hist_consis' 
	valid_sops' dist' vol_read_consis'] i_bound
      obtain is'' where
	is'': "instrs sb' @ issb = is'' @ prog_instrs sb'" and
	steps: "(?ts',m(a:=v),𝒮) d* 
        (ts[i := (last_prog (hd_prog psb sb') sb, is'',
	    θ |` (dom θ - read_tmps sb'), x, 
            𝒟  outstanding_refs is_volatile_Writesb sb  {}, acquired True sb 𝒪, release sb (dom 𝒮) )],
         flush sb (m(a := v)),share sb 𝒮)"
	by (auto simp del: fun_upd_apply)
      note steps
      finally
      show ?thesis
	using is''
	by (simp add: Writesb False)
    next
      case True
      have memop_step:
	"(Write volatile a sop A L R W#is',θ|`(dom θ - read_tmps (sb@sb')),
                 x,m,𝒟,𝒪,,𝒮 )  
           (is',θ|` (dom θ - read_tmps (sb@sb')),x,m(a:=v),True,𝒪  A - R,Map.empty,𝒮 ⊕⇘WR ⊖⇘AL)"
	using D_subset
	apply (simp only: sop f_v' [symmetric] True)
	apply (rule direct_memop_step.WriteVolatile)
	done

      let ?ts' = "ts[i := (p, is', θ |` (dom θ - read_tmps (sb @ sb')),x, True, 𝒪  A - R,Map.empty)]"
      from direct_computation.Memop [OF i_bound ts_i  memop_step]
      have "(ts, m, 𝒮) d (?ts', m(a := v), 𝒮 ⊕⇘WR ⊖⇘AL)".

      also
      from reads_consis have reads_consis': "reads_consistent True (𝒪'  A - R)(m(a:=v)) sb"
	by (auto simp add: Writesb True)
      from i_bound have i_bound': "i < length ?ts'"
	by auto
    
      from i_bound
      have ts'_i: "?ts' ! i = (p, is',θ |` (dom θ - read_tmps (sb @ sb')), x, True, 𝒪  A - R,Map.empty)"
	by simp      

      from Cons.hyps [OF i_bound' is' causal' ts'_i p' psb' reads_consis' hist_consis' 
	valid_sops' dist' vol_read_consis', of "(𝒮 ⊕⇘WR ⊖⇘AL)"] i_bound
      obtain is'' where
	is'': "instrs sb' @ issb = is'' @ prog_instrs sb'" and
	steps: "(?ts',m(a:=v),𝒮 ⊕⇘WR ⊖⇘AL) d* 
        (ts[i := (last_prog (hd_prog psb sb') sb, is'',
            θ |` (dom θ - read_tmps sb'), x, 
	    True, acquired True sb (𝒪  A - R),release sb (dom (𝒮 ⊕⇘WR ⊖⇘AL)) Map.empty)],
         flush sb (m(a := v)), share sb (𝒮 ⊕⇘WR ⊖⇘AL))"
	by (auto simp del: fun_upd_apply)
      note steps
      finally
      show ?thesis
	using is''
	by (simp add: Writesb True)
    qed
  next
    case (Readsb volatile a t v)

    from vol_read_consis reads_consis obtain v: "v=m a" and r_consis: "reads_consistent True 𝒪' m sb" and
      vol_read_consis': "volatile_reads_consistent m sb"
      by (cases volatile) (auto simp add: Readsb)

    from valid_sops have valid_sops': "sop  write_sops sb. valid_sop sop"  
      by (simp add: Readsb)

    from hist_consis obtain θ: "θ t = Some v" and 
      hist_consis': "history_consistent θ p (sb@sb')" 
      by (simp add: Readsb split: option.splits)
    from dist obtain t_notin: "t  read_tmps (sb@sb')" and
      dist': "distinct_read_tmps (sb@sb')" by (simp add: Readsb)
    from θ t_notin have restrict_commute:
      "(θ|` (dom θ - read_tmps (sb@sb')))(tv) =
        θ|` (dom θ - read_tmps (sb@sb'))"
      apply -
      apply (rule ext)
      apply (auto simp add: restrict_map_def domIff)
      done
    from θ t_notin 
    have restrict_commute': 
      "((θ |` (dom θ - insert t (read_tmps (sb@sb'))))(t  v)) =
          θ|` (dom θ - read_tmps (sb@sb'))"
      apply -
      apply (rule ext)
      apply (auto simp add: restrict_map_def domIff)
      done

    interpret causal': causal_program_history "issb" "sb@sb'" by fact

    from "is"
    have "Read volatile a t # instrs (sb @ sb') @ issb = is @ prog_instrs (sb @ sb')"
      by (simp add: Readsb)

    with causal'.causal_program_history [of "[]", simplified, OF refl]    
    obtain is' where "is": "is=Read volatile a t#is'" and
      is': "instrs (sb @ sb') @ issb = is' @ prog_instrs (sb @ sb')"
      by auto

    from ts_i "is"
    have ts_i: "ts!i = (p,Read volatile a t#is',
                 θ |` (dom θ - insert t (read_tmps (sb@sb'))),x,𝒟,𝒪,)"
      by (simp add: Readsb)

    from direct_memop_step.Read [of volatile a t "is'" "θ|` (dom θ - insert t (read_tmps (sb@sb')))" x m 𝒟 𝒪  𝒮]
    have memop_step: " 
          (Read volatile a t # is',
            θ |` (dom θ - insert t (read_tmps (sb @ sb'))), x, m, 𝒟, 𝒪,,𝒮)  
          (is',
             θ |` (dom θ - (read_tmps (sb @ sb'))), x, m, 𝒟, 𝒪, ,𝒮)"
      by (simp add: v [symmetric] restrict_commute restrict_commute')

    let ?ts' = "ts[i := (p, is',
                 θ |` (dom θ - read_tmps (sb @ sb')),x, 𝒟, 𝒪,)]"

    from direct_computation.Memop [OF i_bound ts_i memop_step]
    have "(ts, m, 𝒮) d (?ts', m, 𝒮)".

    also

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

    from i_bound
    have ts'_i: "?ts'!i = (p,is', (θ |` (dom θ - read_tmps (sb @ sb'))),x,𝒟, 𝒪, )"
      by auto

    from p have p': "p = hd_prog psb (sb@sb')"
      by (auto simp add: Readsb hd_prog_idem)

    from psb have psb': "last_prog psb (sb @ sb') = psb"
      by (simp add: Readsb)

    
    from Cons.hyps [OF i_bound' is' causal' ts'_i p' psb' r_consis  hist_consis'
    valid_sops' dist' vol_read_consis']
    
    obtain is'' where
      is'': "instrs sb' @ issb = is'' @ prog_instrs sb'" and
      steps: "(?ts',m,𝒮) d* 
          (ts[i := (last_prog (hd_prog psb sb') sb, is'',
             θ |` (dom θ - read_tmps sb'),x, 𝒟  outstanding_refs is_volatile_Writesb sb  {}, 
             acquired True sb 𝒪, release sb (dom 𝒮) )],
             flush sb m,share sb 𝒮)"
      by (auto simp del: fun_upd_apply)

    note steps
    finally
    show ?thesis
      using is''
      by (simp add: Readsb)
  next
    case (Ghostsb A L R W)

    from vol_read_consis
    have vol_read_consis': "volatile_reads_consistent m sb"
      by (auto simp add: Ghostsb)

    from reads_consis have  r_consis: "reads_consistent True (𝒪'  A - R) m sb"
      by (auto simp add: Ghostsb)

    from valid_sops have valid_sops': "sop  write_sops sb. valid_sop sop"  
      by (simp add: Ghostsb)

    from hist_consis obtain 
      hist_consis': "history_consistent θ p (sb@sb')" 
      by (simp add: Ghostsb)

    from dist obtain 
      dist': "distinct_read_tmps (sb@sb')" by (simp add: Ghostsb)

    interpret causal': causal_program_history "issb" "sb@sb'" by fact

    from "is"
    have "Ghost A L R W# instrs (sb @ sb') @ issb = is @ prog_instrs (sb @ sb')"
      by (simp add: Ghostsb)

    with causal'.causal_program_history [of "[]", simplified, OF refl]    
    obtain is' where "is": "is=Ghost A L R W#is'" and
      is': "instrs (sb @ sb') @ issb = is' @ prog_instrs (sb @ sb')"
      by auto

    from ts_i "is"
    have ts_i: "ts!i = (p,Ghost A L R W#is',
                 θ |` (dom θ - (read_tmps (sb@sb'))),x,𝒟,𝒪,)"
      by (simp add: Ghostsb)

    from direct_memop_step.Ghost [of A L R W "is'" 
      "θ|` (dom θ - (read_tmps (sb@sb')))" x  m 𝒟 "𝒪"  𝒮]
    have memop_step:"
      (Ghost A L R W# is',θ |` (dom θ - read_tmps (sb @ sb')), x, m, 𝒟, 𝒪, , 𝒮) 
       (is',θ |` (dom θ - read_tmps (sb @ sb')), x, m, 𝒟, 𝒪  A - R , augment_rels (dom 𝒮) R , 
      𝒮 ⊕⇘WR ⊖⇘AL)".

    let ?ts' = "ts[i := (p, is',
                 θ |` (dom θ - read_tmps (sb @ sb')),x, 𝒟, 𝒪  A - R, augment_rels (dom 𝒮) R )]"
    from direct_computation.Memop [OF i_bound ts_i memop_step]
    have "(ts, m, 𝒮) d (?ts', m, 𝒮 ⊕⇘WR ⊖⇘AL)".

    also

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

    from i_bound
    have ts'_i: "?ts'!i = (p,is',(θ |` (dom θ - read_tmps (sb @ sb'))),x, 𝒟, 𝒪  A - R,augment_rels (dom 𝒮) R  )"
      by auto

    from p have p': "p = hd_prog psb (sb@sb')"
      by (auto simp add: Ghostsb hd_prog_idem)

    from psb have psb': "last_prog psb (sb @ sb') = psb"
      by (simp add: Ghostsb)

    from Cons.hyps [OF   i_bound' is' causal' ts'_i p' psb' r_consis hist_consis' 
      valid_sops' dist' vol_read_consis', of "𝒮 ⊕⇘WR ⊖⇘AL"] 
    obtain is'' where
      is'': "instrs sb' @ issb = is'' @ prog_instrs sb'" and
      steps: "(?ts',m,𝒮 ⊕⇘WR ⊖⇘AL) d* 
          (ts[i := (last_prog (hd_prog psb sb') sb, is'',
             θ |` (dom θ - read_tmps sb'),x, 
             𝒟  outstanding_refs is_volatile_Writesb sb  {}, acquired True sb (𝒪  A - R), 
              release sb (dom (𝒮 ⊕⇘WR ⊖⇘AL)) (augment_rels (dom 𝒮) R ))],
           flush sb m,share sb (𝒮 ⊕⇘WR ⊖⇘AL))"
      by (auto simp add: list_update_overwrite simp del: fun_upd_apply)

    note steps
    finally
    show ?thesis
      using is''
      by (simp add: Ghostsb)
  qed
qed

corollary  (in program) flush_store_buffer:
  assumes i_bound: "i < length ts"
  assumes instrs: "instrs sb @ issb = is @ prog_instrs sb"
  assumes cph: "causal_program_history issb sb"
  assumes ts_i: "ts!i = (p,is,θ |` (dom θ - read_tmps sb),x,𝒟,𝒪,)"
  assumes p: "p=hd_prog psb sb"
  assumes last_prog: "(last_prog psb sb) = psb"
  assumes reads_consis: "reads_consistent True 𝒪' m sb"
  assumes hist_consis: "history_consistent θ p sb"
  assumes valid_sops: "sop  write_sops sb. valid_sop sop"
  assumes dist: "distinct_read_tmps sb"
  assumes vol_read_consis: "volatile_reads_consistent m sb"
  shows "(ts,m,𝒮) d* 
         (ts[i:=(psb,issb, θ,x,
            𝒟  outstanding_refs is_volatile_Writesb sb  {},acquired True sb 𝒪, release sb (dom 𝒮) )],
             flush sb m,share sb 𝒮)"
using flush_store_buffer_append [where sb'="[]", simplified, OF i_bound instrs cph ts_i [simplified] p last_prog reads_consis hist_consis valid_sops dist vol_read_consis] last_prog
by simp


lemma last_prog_same_append: "xs psb. last_prog psb (sb@xs) = psb  last_prog psb xs = psb"
  apply (induct sb)
  apply  simp
  subgoal for a sb xs psb
  apply (case_tac a)
  apply     simp
  apply    simp
  apply   simp
  apply  (drule last_prog_to_last_prog_same)
  apply  simp
  apply simp
  done
  done


lemma reads_consistent_drop_volatile_writes_no_volatile_reads: 
  "pending_write 𝒪 m. reads_consistent pending_write 𝒪 m sb  
  outstanding_refs is_volatile_Readsb ((dropWhile (Not  is_volatile_Writesb)) sb) = {}"
  apply (induct sb)
  apply (auto split: memref.splits)
  done

(* cf reads_consistent_append
lemma reads_consistent_flush: 
"⋀m. reads_consistent m sb ⟹ 
  reads_consistent (flush (takeWhile (Not ∘ is_volatile_Writesb) sb) m) 
   (dropWhile (Not ∘ is_volatile_Writesb) sb)"
  apply (induct sb)
  apply  simp
  apply (case_tac a)
  apply auto
  done
*)


lemma reads_consistent_flush_other: 
  assumes no_volatile_Writesb_sb: "outstanding_refs is_volatile_Writesb sb = {}"
  shows "m pending_write 𝒪. 
  outstanding_refs (Not  is_volatile_Readsb) xs  outstanding_refs is_non_volatile_Writesb sb = {};
       reads_consistent pending_write 𝒪 m xs  reads_consistent pending_write 𝒪 (flush sb m) xs"
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x xs)
  note no_inter = outstanding_refs (Not  is_volatile_Readsb) (x # xs)  
    outstanding_refs is_non_volatile_Writesb sb = {}
  hence no_inter': "outstanding_refs (Not  is_volatile_Readsb) xs  outstanding_refs is_non_volatile_Writesb sb = {}"
    by (auto)
  note consis = reads_consistent pending_write 𝒪 m (x # xs)
  show ?case
  proof (cases x)
    case (Writesb volatile a sop v A L R)
    show ?thesis
    proof (cases volatile)
      case False
      from consis obtain consis': "reads_consistent pending_write 𝒪 (m(a := v)) xs" 
	by (simp add: Writesb False)
      from Cons.hyps [OF no_inter' consis']
      have "reads_consistent pending_write 𝒪 (flush sb (m(a := v))) xs".
      moreover
      from no_inter have "a  outstanding_refs is_non_volatile_Writesb sb"
	by (auto simp add: Writesb split: if_split_asm)
    
      from flush_update_other' [OF this no_volatile_Writesb_sb]
      have "(flush sb (m(a := v))) = (flush sb m)(a := v)".
      ultimately
      show ?thesis
	by (simp add: Writesb False)
    next
      case True
      from consis obtain consis': "reads_consistent True (𝒪  A - R) (m(a := v)) xs" and
	no_read: "(outstanding_refs is_volatile_Readsb xs = {} )"
	by (simp add: Writesb True)
      from Cons.hyps [OF no_inter' consis']
      have "reads_consistent True (𝒪  A - R) (flush sb (m(a := v))) xs".
      moreover
      from no_inter have "a  outstanding_refs is_non_volatile_Writesb sb"
	by (auto simp add: Writesb split: if_split_asm)
    
      from flush_update_other' [OF this no_volatile_Writesb_sb]
      have "(flush sb (m(a := v))) = (flush sb m)(a := v)".
      ultimately
      show ?thesis
	using no_read
	by (simp add: Writesb True)
    qed
  next
    case (Readsb volatile a t v)
    from consis obtain val: "(¬ volatile  (pending_write  a  𝒪)  v = m a)" and
      consis': "reads_consistent pending_write 𝒪 m xs"
      by (simp add: Readsb)
    from Cons.hyps [OF no_inter' consis']
    have hyp: "reads_consistent pending_write 𝒪 (flush sb m) xs"
      by simp
    show ?thesis
    proof (cases volatile)
      case False
      from no_inter False have "a  outstanding_refs is_non_volatile_Writesb sb"
	by (auto simp add: Readsb split: if_split_asm)
      with no_volatile_Writesb_sb 
      have "a  outstanding_refs is_Writesb sb"
	apply (clarsimp simp add: outstanding_refs_conv is_Writesb_def split: memref.splits)
	apply force
	done
      with hyp val flush_unchanged_addresses  [OF this]
      show ?thesis
	by (simp add: Readsb)
    next
      case True
      with hyp val show ?thesis
	by (simp add: Readsb)
    qed
  next
    case Progsb with Cons show ?thesis by auto
  next
    case Ghostsb with Cons show ?thesis by auto
  qed
qed

lemma reads_consistent_flush_independent: 
  assumes no_volatile_Writesb_sb: "outstanding_refs is_Writesb sb  outstanding_refs is_non_volatile_Readsb xs = {}"
  assumes consis: "reads_consistent pending_write 𝒪 m xs" 
  shows "reads_consistent pending_write 𝒪 (flush sb m) xs"
proof -
  from flush_unchanged_addresses [where sb=sb and m=m] no_volatile_Writesb_sb
  have "a  outstanding_refs is_non_volatile_Readsb xs. flush sb m a = m a"
    by auto
  from reads_consistent_mem_eq_on_non_volatile_reads [OF this subset_refl consis]
  show ?thesis .
qed


lemma reads_consistent_flush_all_until_volatile_write_aux:
  assumes no_reads: "outstanding_refs is_volatile_Readsb xs = {}" 
  shows "m pending_write 𝒪'.  reads_consistent pending_write 𝒪' m xs; i < length ts. 
    let (p,is,θ,sb,𝒟,𝒪,) = ts!i in
      outstanding_refs (Not  is_volatile_Readsb) xs  
      outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sb) = {} 
  reads_consistent pending_write 𝒪' (flush_all_until_volatile_write ts m) xs"
proof (induct ts)
  case Nil thus ?case by simp
next
  case (Cons t ts)
  have consis: "reads_consistent pending_write 𝒪' m xs" by fact


  obtain pt "ist" 𝒪t t 𝒟t θt sbt 
    where t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
    by (cases t)

  from Cons.prems t obtain
    no_inter: "outstanding_refs (Not  is_volatile_Readsb) xs  
      outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbt) = {}" and
    no_inter': "i < length ts. 
    let (p,is,θ,sb,𝒟,𝒪,) = ts!i in
      outstanding_refs (Not  is_volatile_Readsb) xs  
      outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sb) = {}"
    by (force simp add: Let_def simp del: o_apply)


  have out1: "outstanding_refs is_volatile_Writesb 
    (takeWhile (Not  is_volatile_Writesb) sbt) = {}"
    by (auto simp add: outstanding_refs_conv dest: set_takeWhileD)

  from no_inter have "outstanding_refs (Not  is_volatile_Readsb) xs  
    outstanding_refs is_non_volatile_Writesb  (takeWhile (Not  is_volatile_Writesb) sbt) = {}" 
    by auto

  from reads_consistent_flush_other [OF out1 this consis]
 
  have "reads_consistent pending_write 𝒪' (flush (takeWhile (Not  is_volatile_Writesb) sbt) m) xs".
  from Cons.hyps [OF this no_inter']
  show ?case
    by (simp add: t)
qed

(* FIXME: delete
lemma read_only_reads_takeWhile_owns:
  "⋀𝒪. read_only_reads 𝒪 (takeWhile (Not ∘ is_volatile_Writesb) sb) ∩ 𝒪 = {}"
apply (induct sb)
apply clarsimp
apply (case_tac a)
apply auto
done
*)

lemma reads_consistent_flush_other': 
  assumes no_volatile_Writesb_sb: "outstanding_refs is_volatile_Writesb sb = {}"
  shows "m  𝒪. 
  outstanding_refs is_non_volatile_Writesb sb  
     (outstanding_refs is_volatile_Writesb xs   
         outstanding_refs is_non_volatile_Writesb xs  
         outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) xs)  
         (outstanding_refs is_non_volatile_Readsb (takeWhile (Not  is_volatile_Writesb) xs) - RO)  
         (𝒪  all_acquired (takeWhile (Not  is_volatile_Writesb) xs))  
   )  = {};
  reads_consistent False 𝒪 m xs;
   read_only_reads 𝒪 (takeWhile (Not  is_volatile_Writesb) xs)  RO 
   reads_consistent False 𝒪 (flush sb m) xs"
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x xs)

  note no_inter = Cons.prems (1)

  note consis = reads_consistent False 𝒪 m (x # xs)
  have aargh: "(Not  is_volatile_Writesb) = (λa. ¬ is_volatile_Writesb a)"
    by (rule ext) auto


  note RO = read_only_reads 𝒪 (takeWhile (Not  is_volatile_Writesb) (x#xs))  RO


  show ?case
  proof (cases x)
    case (Writesb volatile a sop v A L R)
    show ?thesis
    proof (cases volatile)
      case False
      from consis obtain consis': "reads_consistent False 𝒪 (m(a := v)) xs" 
	by (simp add: Writesb False)

      from no_inter
      have no_inter': "outstanding_refs is_non_volatile_Writesb sb  
       (outstanding_refs is_volatile_Writesb xs   
         outstanding_refs is_non_volatile_Writesb xs  
         outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) xs)  
         (outstanding_refs is_non_volatile_Readsb (takeWhile (Not  is_volatile_Writesb) xs) - RO)  
         (𝒪  all_acquired (takeWhile (Not  is_volatile_Writesb) xs))
        ) = {}"
	by (clarsimp simp add: Writesb False aargh)


      from RO
      have RO': "read_only_reads 𝒪 (takeWhile (Not  is_volatile_Writesb) xs)  RO" 
	by (auto simp add: Writesb False)

      from Cons.hyps [OF no_inter' consis' RO']
      have "reads_consistent False 𝒪 (flush sb (m(a := v))) xs".
      moreover
      from no_inter have "a  outstanding_refs is_non_volatile_Writesb sb"
	by (auto simp add: Writesb split: if_split_asm)
    
      from flush_update_other' [OF this no_volatile_Writesb_sb]
      have "(flush sb (m(a := v))) = (flush sb m)(a := v)".
      ultimately
      show ?thesis
	by (simp add: Writesb False)
    next
      case True
      from consis obtain consis': "reads_consistent True (𝒪  A - R) (m(a := v)) xs" and
	no_read: "(outstanding_refs is_volatile_Readsb xs = {})"
	by (simp add: Writesb True)

      from no_inter obtain
	a_notin: "a  outstanding_refs is_non_volatile_Writesb sb" and
	disj: "(outstanding_refs (Not  is_volatile_Readsb) xs) 
	        outstanding_refs is_non_volatile_Writesb sb = {}"
	by (auto simp add: Writesb True aargh misc_outstanding_refs_convs)

      from reads_consistent_flush_other [OF no_volatile_Writesb_sb disj consis']

      have "reads_consistent True (𝒪  A - R) (flush sb (m(a := v))) xs".
      moreover
      note a_notin
    
      from flush_update_other' [OF this no_volatile_Writesb_sb]
      have "(flush sb (m(a := v))) = (flush sb m)(a := v)".
      ultimately
      show ?thesis
	using no_read
	by (simp add: Writesb True)
    qed
  next
    case (Readsb volatile a t v)
    from consis obtain val: "(¬ volatile  a  𝒪  v = m a)" and
      consis': "reads_consistent False 𝒪 m xs"
      by (simp add: Readsb)


    from RO
    have RO': "read_only_reads 𝒪 (takeWhile (Not  is_volatile_Writesb) xs)  RO"
      by (auto simp add: Readsb )

    from no_inter
    have no_inter': "outstanding_refs is_non_volatile_Writesb sb  
       (outstanding_refs is_volatile_Writesb xs   
        outstanding_refs is_non_volatile_Writesb xs  
        outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) xs)  
        (outstanding_refs is_non_volatile_Readsb (takeWhile (Not  is_volatile_Writesb) xs) - RO)  
        (𝒪  all_acquired (takeWhile (Not  is_volatile_Writesb) xs))
        ) = {}"
      by (fastforce simp add: Readsb aargh)


    show ?thesis
    proof (cases volatile)
      case True

      from Cons.hyps [OF no_inter' consis' RO'] 
      show ?thesis
	by (simp add: Readsb True)
    next
      case False
      note non_volatile=this

      from Cons.hyps [OF no_inter' consis' RO'] 
      have hyp: "reads_consistent False 𝒪 (flush sb m) xs".

      show ?thesis
      proof (cases "a  𝒪")
	case False
	with hyp show ?thesis
	  by (simp add: Readsb non_volatile False)
      next
	case True
	from no_inter True have a_notin: "a  outstanding_refs is_non_volatile_Writesb sb"
	  by blast

	with no_volatile_Writesb_sb 
	have "a  outstanding_refs is_Writesb sb"
	  apply (clarsimp simp add: outstanding_refs_conv is_Writesb_def split: memref.splits)
	  apply force
	  done

	from flush_unchanged_addresses  [OF this] hyp val 
	
	show ?thesis
	  by (simp add: Readsb non_volatile True)
      qed
    qed
  next
    case Progsb with Cons show ?thesis
      by auto
  next
    case (Ghostsb A L R W)
    from consis obtain consis': "reads_consistent False (𝒪  A - R) m xs" 
      by (simp add: Ghostsb)

    from RO
    have RO': "read_only_reads (𝒪  A - R) (takeWhile (Not  is_volatile_Writesb) xs)  RO"
      by (auto simp add: Ghostsb)


    from no_inter
    have no_inter': "outstanding_refs is_non_volatile_Writesb sb   
      (outstanding_refs is_volatile_Writesb xs   
        outstanding_refs is_non_volatile_Writesb xs  
         outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) xs)  
         (outstanding_refs is_non_volatile_Readsb (takeWhile (Not  is_volatile_Writesb) xs) - RO)  
         (𝒪  A - R  all_acquired (takeWhile (Not  is_volatile_Writesb) xs))
        ) = {}"
      by (fastforce simp add: Ghostsb aargh)

    from Cons.hyps [OF no_inter' consis' RO' ] 
    show ?thesis
      by (clarsimp simp add: Ghostsb)
  qed
qed

lemma reads_consistent_flush_all_until_volatile_write_aux':
  assumes no_reads: "outstanding_refs is_volatile_Readsb xs = {}" 
  assumes read_only_reads_RO: "read_only_reads 𝒪' (takeWhile (Not  is_volatile_Writesb) xs)  RO"
  shows "m.  reads_consistent False 𝒪' m xs; i < length ts. 
    let (p,is,θ,sb,𝒟,𝒪) = ts!i in
      outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sb)  
       (outstanding_refs is_volatile_Writesb xs   
         outstanding_refs is_non_volatile_Writesb xs  
         outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) xs)  
         (outstanding_refs is_non_volatile_Readsb (takeWhile (Not  is_volatile_Writesb) xs) - RO)  
         (𝒪'  all_acquired (takeWhile (Not  is_volatile_Writesb) xs))
        )  
       = {}
 
  reads_consistent False 𝒪' (flush_all_until_volatile_write ts m) xs"
proof (induct ts)
  case Nil thus ?case by simp
next
  case (Cons t ts)
  have consis: "reads_consistent False 𝒪' m xs" by fact


  obtain pt "ist" 𝒪t t 𝒟t θt sbt 
    where t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
    by (cases t)

  obtain
    no_inter: "outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbt)  
    (outstanding_refs is_volatile_Writesb xs   
         outstanding_refs is_non_volatile_Writesb xs  
         outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) xs)  
         (outstanding_refs is_non_volatile_Readsb (takeWhile (Not  is_volatile_Writesb) xs) - RO)  
         (𝒪'  all_acquired (takeWhile (Not  is_volatile_Writesb) xs))

        )  
       = {}" and
    no_inter': "i < length ts. 
    let (p,is,θ,sb,𝒟,𝒪) = ts!i in
      outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sb)  
       (outstanding_refs is_volatile_Writesb xs   
         outstanding_refs is_non_volatile_Writesb xs  
         outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) xs)  
         (outstanding_refs is_non_volatile_Readsb (takeWhile (Not  is_volatile_Writesb) xs) - RO)  
         (𝒪'  all_acquired (takeWhile (Not  is_volatile_Writesb) xs))
        )  
    = {}"
  proof -
    show ?thesis
      apply (rule that)
      using  Cons.prems (2) [rule_format, of 0]
      apply  (clarsimp simp add: t)
      apply clarsimp
      using Cons.prems (2)
      apply -
      subgoal for i
      apply (drule_tac x="Suc i" in spec)
      apply (clarsimp simp add: Let_def simp del: o_apply)
      done
      done
  qed


  have out1: "outstanding_refs is_volatile_Writesb 
    (takeWhile (Not  is_volatile_Writesb) sbt) = {}"
    by (auto simp add: outstanding_refs_conv dest: set_takeWhileD)

  from reads_consistent_flush_other' [OF out1 no_inter consis read_only_reads_RO ] 
  have "reads_consistent False 𝒪' (flush (takeWhile (Not  is_volatile_Writesb) sbt) m) xs".
  from Cons.hyps [OF this no_inter']
  show ?case
    by (simp add: t)
qed





(* cf. reads_consistent_drop_volatile_writes_no_volatile_reads
lemma reads_consistent_no_volatile_Readsb_drop: "⋀m. reads_consistent m sb ⟹ 
   outstanding_refs is_volatile_Readsb (dropWhile (Not ∘ is_volatile_Writesb) sb) = {}"
apply (induct sb)
apply  simp
apply (case_tac a)
apply auto
done
*)

lemma in_outstanding_refs_cases [consumes 1, case_names Writesb Readsb]:
  "a  outstanding_refs P xs 
       (volatile sop v A L R W.  (Writesb volatile a sop v A L R W)  set xs  P (Writesb volatile a sop v A L R W)  C) 
       (volatile t v.  (Readsb volatile a t v)  set xs  P (Readsb volatile a t v)  C)
        C"
  apply (clarsimp simp add: outstanding_refs_conv)
  subgoal for x
  apply (case_tac x)
  apply fastforce+
  done
  done

lemma dropWhile_Cons: "(dropWhile P xs) = x#ys  ¬ P x"
apply (induct xs)
apply (auto split: if_split_asm)
done

lemma reads_consistent_dropWhile: 
  "reads_consistent pending_write 𝒪 m (dropWhile (Not  is_volatile_Writesb) sb) = 
       reads_consistent True 𝒪 m  (dropWhile (Not  is_volatile_Writesb) sb)"
apply (case_tac "(dropWhile (Not  is_volatile_Writesb) sb)")
apply (simp only:)
apply  simp
apply (frule dropWhile_Cons)
apply (auto split: memref.splits)
done


theorem 
  reads_consistent_flush_all_until_volatile_write: 
  "i m pending_write. valid_ownership_and_sharing 𝒮 ts; 
  i < length ts; ts!i = (p, is,θ, sb, 𝒟, 𝒪,); 
  reads_consistent pending_write 𝒪 m sb  
   reads_consistent True (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪)
  (flush_all_until_volatile_write ts m) (dropWhile (Not  is_volatile_Writesb) sb)"
proof (induct ts)
  case Nil thus ?case by simp
next
  case (Cons t ts)
  note i_bound = i < length (t # ts)
  note ts_i = (t # ts) ! i = (p, is,θ, sb, 𝒟, 𝒪,)
  note consis = reads_consistent pending_write 𝒪 m sb
  note valid = valid_ownership_and_sharing 𝒮 (t#ts)
  then interpret valid_ownership_and_sharing 𝒮 "t#ts".
  from valid_ownership_and_sharing_tl [OF valid] have valid': "valid_ownership_and_sharing 𝒮 ts".
    
  obtain pt "ist" 𝒪t t 𝒟t θt sbt 
    where t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
    by (cases t)
  show ?case
  proof (cases i)
    case 0
    with ts_i t have sb_eq: "sb=sbt"
      by simp

    let ?take_sb = "(takeWhile (Not  is_volatile_Writesb) sb)"
    let ?drop_sb = "(dropWhile (Not  is_volatile_Writesb) sb)"

    from reads_consistent_append [of pending_write 𝒪 m ?take_sb ?drop_sb] consis
    have consis': "reads_consistent True (acquired True ?take_sb 𝒪) (flush ?take_sb m) ?drop_sb"
      apply (cases "outstanding_refs is_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sb)  {}")
      apply  clarsimp
      apply clarsimp
      apply (simp add: reads_consistent_dropWhile [of pending_write])
      done


    from reads_consistent_drop_volatile_writes_no_volatile_reads [OF consis]
    have no_vol_Readsb: "outstanding_refs is_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) sb) = {}".
    hence "outstanding_refs (Not  is_volatile_Readsb) (dropWhile (Not  is_volatile_Writesb) sb)
           =
           outstanding_refs (λs. True) (dropWhile (Not  is_volatile_Writesb) sb)"
      by (auto simp add: outstanding_refs_conv)

    have "i<length ts.
     let (p, is,θ, sb', 𝒟, 𝒪,) = ts ! i
     in outstanding_refs (Not  is_volatile_Readsb) (dropWhile (Not  is_volatile_Writesb) sb) 
        outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sb') = {}"
    proof - 
      {
	fix j pj "isj" 𝒪j j 𝒟j θj sbj x
	assume j_bound: "j < length ts"
	assume ts_j: "ts!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
	assume x_in_sb: "x  outstanding_refs (Not  is_volatile_Readsb) (dropWhile (Not  is_volatile_Writesb) sb)"
	assume x_in_j: "x  outstanding_refs is_non_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbj)"
	have False
	proof -
	  from outstanding_non_volatile_write_not_volatile_read_disj [rule_format, of "Suc j" 0, simplified, OF j_bound ts_j t]
	  sb_eq x_in_sb x_in_j 
	  show ?thesis
	    by auto
	qed
      }
      thus ?thesis
	by (auto simp add: Let_def)
    qed
    from reads_consistent_flush_all_until_volatile_write_aux [OF no_vol_Readsb consis' this]
    show ?thesis
      by (simp add: t sb_eq del: o_apply)
  next
    case (Suc k)
    with i_bound have k_bound: "k < length ts"
      by auto
      
    from ts_i Suc have ts_k: "ts ! k = (p, is,θ, sb, 𝒟, 𝒪,)"
      by simp



    have "reads_consistent False 𝒪 (flush (takeWhile (Not  is_volatile_Writesb) sbt) m) sb"
    proof -
      have no_vW: 
	"outstanding_refs is_volatile_Writesb (takeWhile (Not  is_volatile_Writesb) sbt) = {}"
	apply (clarsimp simp add: outstanding_refs_conv )
	apply (drule set_takeWhileD)
	apply simp
	done
      
      
      from consis have consis': "reads_consistent False 𝒪 m sb"
	by (cases pending_write) (auto intro: reads_consistent_pending_write_antimono)
      note disj = outstanding_non_volatile_write_disj [where i=0, OF _  i_bound [simplified Suc], simplified, OF t ts_k ]
      
      
      from reads_consistent_flush_other' [OF no_vW disj consis' subset_refl] 
      show ?thesis .
    qed
    from Cons.hyps [OF valid' k_bound ts_k this]
    show ?thesis
      by (simp add: t)
  qed
qed


lemma split_volatile_Writesb_in_outstanding_refs:
  "a  outstanding_refs is_volatile_Writesb xs  (sop v ys zs A L R W. xs = ys@(Writesb True a sop v A L R W#zs))"
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x xs)
  have a_in: "a  outstanding_refs is_volatile_Writesb (x # xs)" by fact
  show ?case
  proof (cases x)
    case (Writesb volatile a' sop v A L R W)
    show ?thesis
    proof (cases volatile)
      case False
      from a_in have "a  outstanding_refs is_volatile_Writesb xs"
	by (auto simp add: False Writesb)
      from Cons.hyps [OF this] obtain sop'' v'' A'' L'' R'' W'' ys zs
	where "xs=ys@Writesb True a sop'' v'' A'' L'' R'' W''#zs"
	by auto
      hence "x#xs = (x#ys)@Writesb True a sop'' v'' A'' L'' R'' W''#zs"
	by auto
      thus ?thesis
	by blast
    next
      case True
      note volatile = this
      show ?thesis
      proof (cases "a'=a")
	case False
	with a_in have "a  outstanding_refs is_volatile_Writesb xs"
	  by (auto simp add: volatile Writesb)
	from Cons.hyps [OF this] obtain sop'' v'' A'' L'' R'' W'' ys zs
	  where "xs=ys@Writesb True a sop'' v'' A'' L'' R'' W''#zs"
	  by auto
	hence "x#xs = (x#ys)@Writesb True a sop'' v'' A'' L'' R'' W''#zs"
	  by auto
	thus ?thesis
	  by blast
      next
	case True
	then have "x#xs=[]@(Writesb True a sop v A L R W#xs)"
	  by (simp add: Writesb volatile True)
	thus ?thesis
	  by blast
      qed
    qed
  next
    case Readsb 
    from a_in have "a  outstanding_refs is_volatile_Writesb xs"
      by (auto simp add: Readsb)
    from Cons.hyps [OF this] obtain sop'' v'' A'' L'' R'' W'' ys zs
      where "xs=ys@Writesb True a sop'' v'' A'' L'' R'' W''#zs"
      by auto
    hence "x#xs = (x#ys)@Writesb True a sop'' v'' A'' L'' R'' W''#zs"
      by auto
    thus ?thesis
      by blast
  next
    case Progsb
    from a_in have "a  outstanding_refs is_volatile_Writesb xs"
      by (auto simp add: Progsb)
    from Cons.hyps [OF this] obtain sop'' v'' A'' L'' R'' W'' ys zs
      where "xs=ys@Writesb True a sop'' v'' A'' L'' R'' W''#zs"
      by auto
    hence "x#xs = (x#ys)@Writesb True a sop'' v'' A'' L'' R'' W''#zs"
      by auto
    thus ?thesis
      by blast
  next
    case Ghostsb
    from a_in have "a  outstanding_refs is_volatile_Writesb xs"
      by (auto simp add: Ghostsb)
    from Cons.hyps [OF this] obtain sop'' v'' A'' L'' R'' W'' ys zs
      where "xs=ys@Writesb True a sop'' v'' A'' L'' R'' W''#zs"
      by auto
    hence "x#xs = (x#ys)@Writesb True a sop'' v'' A'' L'' R'' W''#zs"
      by auto
    thus ?thesis
      by blast
  qed
qed

lemma sharing_consistent_mono_shared:
"𝒮 𝒮' 𝒪.
  dom 𝒮  dom 𝒮'  sharing_consistent 𝒮 𝒪 sb  sharing_consistent 𝒮' 𝒪 sb"
apply (induct sb)
apply simp
subgoal for a sb 𝒮 𝒮' 𝒪
apply (case_tac a)
apply    clarsimp
         subgoal for volatile a D f v A L R W  
         apply (frule_tac A="𝒮" and B="𝒮'" and C="R" and x="W" in augment_mono_aux)
         apply (frule_tac A="𝒮 ⊕⇘WR" and B="𝒮' ⊕⇘WR" and C="L" in restrict_mono_aux)
         apply blast
         done
apply   clarsimp
apply  clarsimp
apply clarsimp
subgoal for A L R W
apply (frule_tac A="𝒮" and B="𝒮'" and C="R" and x="W" in augment_mono_aux)
apply (frule_tac A="𝒮 ⊕⇘WR" and B="𝒮' ⊕⇘WR" and C="L" in restrict_mono_aux)
apply blast
done
done
done

lemma sharing_consistent_mono_owns:
"𝒪 𝒪' 𝒮.
  𝒪  𝒪'  sharing_consistent 𝒮 𝒪 sb  sharing_consistent 𝒮 𝒪' sb"
apply (induct sb)
apply simp
subgoal for a sb 𝒪 𝒪' 𝒮
apply (case_tac a)
apply    clarsimp
         subgoal for volatile a D f v A L R W
         apply (frule_tac A="𝒪" and B="𝒪'" and C="A" in union_mono_aux)
         apply (frule_tac A="𝒪  A" and B="𝒪'  A" and C="R" in set_minus_mono_aux)
         apply fastforce
         done
apply   clarsimp
apply  clarsimp
apply clarsimp
subgoal for A L R W
apply (frule_tac A="𝒪" and B="𝒪'" and C="A" in union_mono_aux)
apply (frule_tac A="𝒪  A" and B="𝒪'  A" and C="R" in set_minus_mono_aux)
apply fastforce
done
done
done


(* FIXME: move up *)
primrec all_shared :: "'a memref list  addr set"
where 
  "all_shared [] = {}"
| "all_shared (i#is) =
    (case i of
      Writesb volatile _ _ _ A L R W  (if volatile then R  all_shared is else all_shared is)
     | Ghostsb A L R W  R  all_shared is
     | _  all_shared is)"

lemma sharing_consistent_all_shared:
  "𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb  all_shared sb  dom 𝒮  𝒪"
  apply (induct sb)
  apply  clarsimp
  subgoal for a
  apply (case_tac a) 
  apply    (fastforce split: memref.splits if_split_asm)
  apply   clarsimp
  apply  clarsimp
  apply fastforce
  done
  done

lemma sharing_consistent_share_all_shared:
  "𝒮. dom (share sb 𝒮)  dom 𝒮  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 t A L R W)
    show ?thesis
    proof (cases volatile)
      case True
      from Cons.hyps [of "(𝒮 ⊕⇘WR ⊖⇘AL)"]
      show ?thesis
        by (auto simp add: Writesb True)
    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.hyps [of "(𝒮 ⊕⇘WR ⊖⇘AL)"]
    show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed




primrec all_unshared :: "'a memref list  addr set"
where 
  "all_unshared [] = {}"
| "all_unshared (i#is) =
    (case i of
      Writesb volatile _ _ _  A L R W  (if volatile then L  all_unshared is else all_unshared is)
     | Ghostsb A L R W  L  all_unshared is
     | _  all_unshared is)"

lemma all_unshared_append: "all_unshared (xs @ ys) = all_unshared xs  all_unshared ys"
  apply (induct xs)
  apply  simp
  subgoal for a
  apply (case_tac a)
  apply auto
  done
  done


lemma freshly_shared_owned:
  "𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb  dom (share sb 𝒮) - dom 𝒮  𝒪"
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.hyps [where 𝒮="(𝒮 ⊕⇘WR ⊖⇘AL)" and 𝒪="(𝒪  A - R)"] Cons.prems
      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)       
    with Cons.hyps [where 𝒮="(𝒮 ⊕⇘WR ⊖⇘AL)" and 𝒪="(𝒪  A - R)"] Cons.prems show ?thesis by auto
  qed
qed

lemma unshared_all_unshared:
  "𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb  dom 𝒮 - dom (share sb 𝒮)  all_unshared 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.hyps [where 𝒮="(𝒮 ⊕⇘WR ⊖⇘AL)" and 𝒪="(𝒪  A - R)"] Cons.prems
      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)       
    with Cons.hyps [where 𝒮="(𝒮 ⊕⇘WR ⊖⇘AL)" and 𝒪="(𝒪  A - R)"] Cons.prems show ?thesis by auto
  qed
qed

lemma unshared_acquired_or_owned: 
  "𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb  all_unshared sb  all_acquired sb  𝒪"
  apply (induct sb)
  apply  simp
  subgoal for a
  apply (case_tac a)
  apply auto+
  done
  done

lemma all_shared_acquired_or_owned: 
  "𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb  all_shared sb  all_acquired sb  𝒪"
  apply (induct sb)
  apply  simp
  subgoal for a
  apply (case_tac a)
  apply auto+
  done
  done

(* FIXME: replace by thm sharing_consistent_shared_exchange? *)
lemma sharing_consistent_preservation:
"𝒮 𝒮' 𝒪. 
sharing_consistent 𝒮 𝒪 sb;
 all_acquired sb  dom 𝒮 - dom 𝒮' = {};
 all_unshared sb  dom 𝒮' - dom 𝒮 = {}
  sharing_consistent 𝒮' 𝒪 sb"
proof (induct sb)
  case Nil thus ?case by simp
next
  case (Cons x sb)
  have consis: "sharing_consistent 𝒮 𝒪 (x # sb)" by fact
  have removed_cond: "all_acquired (x # sb)  dom 𝒮 - dom 𝒮' = {}" by fact
  have new_cond: "all_unshared (x # sb)  dom 𝒮' - dom 𝒮 = {}" by fact
  show ?case
  proof (cases x)
    case (Writesb volatile a sop v A L R W)
    show ?thesis
    proof (cases volatile)
      case False with Writesb Cons show ?thesis
	by auto
    next
      case True
      from consis obtain 
	A: "A  dom 𝒮  𝒪" and
	L: "L  A" and
        A_R: "A  R = {}" and
	R: "R  𝒪" and
	consis': "sharing_consistent (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb"
	by (clarsimp simp add: Writesb True)
      

      from removed_cond obtain rem_cond: "(A  all_acquired sb)  dom 𝒮  dom 𝒮'" by (clarsimp simp add: Writesb True)
      hence rem_cond': "all_acquired sb  dom (𝒮 ⊕⇘WR ⊖⇘AL) - dom (𝒮' ⊕⇘WR ⊖⇘AL) = {}"
	by auto

      from new_cond obtain "(L  all_unshared sb)  dom 𝒮'  dom 𝒮" by (clarsimp simp add: Writesb True)
      hence new_cond': "all_unshared sb  dom (𝒮' ⊕⇘WR ⊖⇘AL) - dom (𝒮 ⊕⇘WR ⊖⇘AL) = {}"     
	by auto
      
      from Cons.hyps [OF consis' rem_cond' new_cond']
      have "sharing_consistent (𝒮' ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb".
      moreover
      from A rem_cond have "A  dom 𝒮'  𝒪"
	by auto
      moreover note L A_R R
      ultimately show ?thesis
	by (auto simp add: Writesb True)
    qed
  next
    case (Ghostsb A L R W)
    from consis obtain 
      A: "A  dom 𝒮  𝒪" and
      L: "L  A" and
      A_R: "A  R = {}" and
      R: "R  𝒪" and
      consis': "sharing_consistent (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb"
      by (clarsimp simp add: Ghostsb)
      

    from removed_cond obtain rem_cond: "(A  all_acquired sb)  dom 𝒮  dom 𝒮'" by (clarsimp simp add: Ghostsb)
    hence rem_cond': "all_acquired sb  dom (𝒮 ⊕⇘WR ⊖⇘AL) - dom (𝒮' ⊕⇘WR ⊖⇘AL) = {}"
      by auto

    from new_cond obtain "(L  all_unshared sb)  dom 𝒮'  dom 𝒮" by (clarsimp simp add: Ghostsb)
    hence new_cond': "all_unshared sb  dom (𝒮' ⊕⇘WR ⊖⇘AL) - dom (𝒮 ⊕⇘WR ⊖⇘AL) = {}"     
      by auto
      
    from Cons.hyps [OF consis' rem_cond' new_cond']
    have "sharing_consistent (𝒮' ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb".
    moreover
    from A rem_cond have "A  dom 𝒮'  𝒪"
      by auto
    moreover note L A_R R
    ultimately show ?thesis
      by (auto simp add: Ghostsb)
  qed (insert Cons, auto)
qed 

lemma (in sharing_consis) sharing_consis_preservation:
assumes dist: 
        "i < length ts. let (_,_,_,sb,_,_,_) = ts!i in 
          all_acquired sb  dom 𝒮 - dom 𝒮' = {}  all_unshared sb  dom 𝒮' - dom 𝒮 = {}"   
shows "sharing_consis 𝒮' ts"   
proof 
  fix i p "is" 𝒪  𝒟 θ sb
  assume i_bound: "i < length ts"
  assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)"
  show "sharing_consistent 𝒮' 𝒪 sb"
  proof -
    from sharing_consis [OF i_bound ts_i]
    have consis: "sharing_consistent 𝒮 𝒪 sb".
    from dist [rule_format, OF i_bound] ts_i
    obtain 
      acq: "all_acquired sb  dom 𝒮 - dom 𝒮' = {}" and
      uns: "all_unshared sb  dom 𝒮' - dom 𝒮 = {}"
      by auto
    from sharing_consistent_preservation [OF consis acq uns]
    show ?thesis .
  qed
qed

lemma (in sharing_consis) sharing_consis_shared_exchange:
assumes dist: 
        "i < length ts. let (_,_,_,sb,_,_,_) = ts!i in 
          a  all_acquired sb. 𝒮' a = 𝒮 a"   
shows "sharing_consis 𝒮' ts"   
proof 
  fix i p "is" 𝒪  𝒟 θ sb
  assume i_bound: "i < length ts"
  assume ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)"
  show "sharing_consistent 𝒮' 𝒪 sb"
  proof -
    from sharing_consis [OF i_bound ts_i]
    have consis: "sharing_consistent 𝒮 𝒪 sb".
    from dist [rule_format, OF i_bound] ts_i
    obtain 
      dist_sb: "a  all_acquired sb. 𝒮' a = 𝒮 a"
      by auto
    from sharing_consistent_shared_exchange [OF dist_sb consis]
    show ?thesis .
  qed
qed



lemma all_acquired_takeWhile: "all_acquired (takeWhile P sb)  all_acquired sb"
proof -
  from all_acquired_append [of "takeWhile P sb" "dropWhile P sb"] 
  show ?thesis
    by auto
qed

lemma all_acquired_dropWhile: "all_acquired (dropWhile P sb)  all_acquired sb"
proof -
  from all_acquired_append [of "takeWhile P sb" "dropWhile P sb"] 
  show ?thesis
    by auto
qed

lemma acquired_share_owns_shared:
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  shows "acquired pending_write sb 𝒪  dom (share sb 𝒮)  𝒪  dom 𝒮"
proof -
  from acquired_all_acquired have "acquired pending_write sb 𝒪  𝒪  all_acquired sb".
  moreover
  from sharing_consistent_all_acquired  [OF consis] have "all_acquired sb  dom 𝒮  𝒪".
  moreover
  from sharing_consistent_share_all_shared have "dom (share sb 𝒮)  dom 𝒮  all_shared sb".
  moreover
  from sharing_consistent_all_shared [OF consis] have "all_shared sb  dom 𝒮  𝒪".
  ultimately
  show ?thesis
    by blast
qed
   
lemma acquired_owns_shared:
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  shows "acquired True sb 𝒪  𝒪  dom 𝒮"
using acquired_share_owns_shared [OF consis]
by blast

lemma share_owns_shared:
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  shows "dom (share sb 𝒮)  𝒪  dom 𝒮"
using acquired_share_owns_shared [OF consis]
by blast

lemma all_shared_append: "all_shared (xs@ys) = all_shared xs  all_shared ys"
  by (induct xs) (auto split: memref.splits)

lemma acquired_union_notin_first:
  " pending_write A B. a  acquired pending_write sb (A  B)  a  A  a  acquired pending_write sb B"
proof (induct sb)
  case Nil thus ?case by (auto split: if_split_asm)
next
  case (Cons x sb)
  then obtain a_notin_A: "a  A" and
    a_acq: "a  acquired pending_write (x # sb) (A  B)" 
    by blast
  show ?case
  proof (cases x)
    case (Writesb volatile a' sop v A' L R W)
    show ?thesis
    proof (cases volatile)
      case False
      with Writesb Cons show ?thesis by simp
    next
      case True
      note volatile = this
      show ?thesis
      proof (cases pending_write)
	case True
	from a_acq have a_acq': "a  acquired True sb (A  B  A' - R)"
	  by (simp add: Writesb volatile True)
	have "(A  B  A' - R)   (A  (B  A' - R))"
	  by auto
	from acquired_mono_in [OF a_acq' this]
	have "a  acquired True sb (A  (B  A' - R))".
	from Cons.hyps [OF this a_notin_A]

	have "a  acquired True sb (B  A' - R)".
	then
	show ?thesis by (simp add: Writesb volatile True)
      next
	case False
	from a_acq have a_acq': "a  acquired True sb (A' - R)"
	  by (simp add: Writesb volatile False)
	then
	show ?thesis
	  by (simp add: Writesb volatile False)
      qed
    qed
  next
    case (Ghostsb A' L R W)
    show ?thesis
    proof (cases pending_write)
      case True
      from a_acq have a_acq': "a  acquired True sb (A  B  A' - R)"
        by (simp add: Ghostsb True)
      have "(A  B  A' - R)   (A  (B  A' - R))"
        by auto
      from acquired_mono_in [OF a_acq' this]
      have "a  acquired True sb (A  (B  A' - R))".
      from Cons.hyps [OF this a_notin_A]

      have "a  acquired True sb (B  A' - R)".
      then
      show ?thesis by (simp add: Ghostsb True)
    next
      case False
      from a_acq have a_acq': "a  acquired False sb (A  B)"
	by (simp add: Ghostsb False)
      from Cons.hyps [OF this a_notin_A]
      show ?thesis
	by (simp add: Ghostsb False)
    qed
  qed (insert Cons, auto)
qed







(* FIXME: move up *)
lemma split_all_acquired_in:
"a  all_acquired xs 
(sop a' v ys zs A L R W. xs = ys @ Writesb True a' sop v A L R W# zs  a  A) 
 (A L R W ys zs. xs = ys @ Ghostsb A L R W# zs  a  A)"
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x xs)
  have a_in: "a  all_acquired (x # xs)" by fact
  show ?case
  proof (cases x)
    case (Writesb volatile a' sop v A L R W)
    show ?thesis
    proof (cases volatile)
      case False
      from a_in have "a  all_acquired xs"
	by (auto simp add: False Writesb)
      from Cons.hyps [OF this] 
      have "(sop a' v ys zs A L R W. xs = ys @ Writesb True a' sop v A L R W# zs  a  A) 
            (A L R W ys zs. xs = ys @ Ghostsb A L R W # zs  a  A)" (is "?write  ?ghst").
      then 
      show ?thesis
      proof 
	assume ?write
	then
	obtain sop'' a'' v'' A'' L'' R'' W'' ys zs
	  where "xs=ys@Writesb True a'' sop'' v'' A'' L'' R'' W''#zs" and a_in: "a  A''"
	  by auto
	hence "x#xs = (x#ys)@Writesb True a'' sop'' v'' A'' L'' R'' W''#zs"
	  by auto
	thus ?thesis
	  using a_in
	  by blast
      next
	assume ?ghst
	then obtain A'' L'' R'' W'' ys zs where
	  "xs=ys@Ghostsb A'' L'' R'' W''#zs" and a_in: "a  A''"
	  by auto
	hence "x#xs = (x#ys)@Ghostsb A'' L'' R'' W''#zs"
	  by auto
	thus ?thesis
	  using a_in
	  by blast
      qed
    next
      case True
      note volatile = this
      show ?thesis
      proof (cases "a  A")
	case False
	with a_in have "a  all_acquired xs"
	  by (auto simp add: volatile Writesb)
	from Cons.hyps [OF this] 
	have "(sop a' v ys zs A L R W. xs = ys @ Writesb True a' sop v A L R W # zs  a  A) 
              (A L R W ys zs. xs = ys @ Ghostsb A L R W# zs  a  A)" (is "?write  ?ghst").
	then 
	show ?thesis
	proof 
	  assume ?write
	  then
	  obtain sop'' a'' v'' A'' L'' R'' W'' ys zs
	    where "xs=ys@Writesb True a'' sop'' v'' A'' L'' R'' W'' #zs" and a_in: "a  A''"
	    by auto
	  hence "x#xs = (x#ys)@Writesb True a'' sop'' v'' A'' L'' R'' W''#zs"
	    by auto
	  thus ?thesis
	    using a_in
	    by blast
	next
	  assume ?ghst
	  then obtain A'' L'' R'' W'' ys zs where
	    "xs=ys @Ghostsb A'' L'' R'' W''#zs" and a_in: "a  A''"
	    by auto 
	  hence "x#xs  = (x#ys)@Ghostsb A'' L'' R'' W''#zs"
	    by auto
	  thus ?thesis
	    using a_in
	    by blast
	qed
      next
	case True
	then have "x#xs=[]@(Writesb True a' sop v A L R W#xs)"
	  by (simp add: Writesb volatile True)
	thus ?thesis
	  using True
	  by blast
      qed
    qed
  next
    case Readsb 
    from a_in have "a  all_acquired xs"
      by (auto simp add: Readsb)
    from Cons.hyps [OF this] 
    have "(sop a' v ys zs A L R W. xs = ys @ Writesb True a' sop v A L R W# zs  a  A) 
            (A L R W ys zs. xs = ys @ Ghostsb A L R W# zs  a  A)" (is "?write  ?ghst").
    then 
    show ?thesis
    proof 
      assume ?write
      then
      obtain sop'' a'' v'' A'' L'' R'' W'' ys zs
	where "xs=ys@Writesb True a'' sop'' v'' A'' L'' R'' W''#zs" and a_in: "a  A''"
	by auto
      hence "x#xs = (x#ys)@Writesb True a'' sop'' v'' A'' L'' R'' W''#zs"
	by auto
      thus ?thesis
	using a_in
	by blast
    next
      assume ?ghst
      then obtain A'' L'' R'' W'' ys zs where
	"xs=ys@Ghostsb A'' L'' R'' W''#zs" and a_in: "a  A''"
	by auto
      hence "x#xs = (x#ys)@Ghostsb A'' L'' R'' W''#zs"
	by auto
      thus ?thesis
	using a_in
	by blast
    qed
  next
    case Progsb
    from a_in have "a  all_acquired xs"
      by (auto simp add: Progsb)
    from Cons.hyps [OF this] 
    have "(sop a' v ys zs A L R W. xs = ys @ Writesb True a' sop v A L R W# zs  a  A) 
            (A L R W ys zs. xs = ys @ Ghostsb A L R W# zs  a  A)" (is "?write  ?ghst").
    then 
    show ?thesis
    proof 
      assume ?write
      then
      obtain sop'' a'' v'' A'' L'' R'' W'' ys zs
	where "xs=ys@Writesb True a'' sop'' v'' A'' L'' R'' W''#zs" and a_in: "a  A''"
	by auto
      hence "x#xs = (x#ys)@Writesb True a'' sop'' v'' A'' L'' R'' W''#zs"
	by auto
      thus ?thesis
	using a_in
	by blast
    next
      assume ?ghst
      then obtain A'' L'' R'' W'' ys zs where
	"xs=ys@Ghostsb A'' L'' R'' W''#zs" and a_in: "a  A''"
	by auto
      hence "x#xs = (x#ys)@Ghostsb A'' L'' R'' W''#zs"
	by auto
      thus ?thesis
	using a_in
	by blast
    qed
  next
    case (Ghostsb A L R W)
    show ?thesis
    proof (cases "a  A")
      case False
      with a_in have "a  all_acquired xs"
	by (auto simp add: Ghostsb)
      from Cons.hyps [OF this] 
      have "(sop a' v ys zs A L R W. xs = ys @ Writesb True a' sop v A L R W # zs  a  A) 
            (A L R W ys zs. xs = ys @ Ghostsb A L R W# zs  a  A)" (is "?write  ?ghst").
      then 
      show ?thesis
      proof 
	assume ?write
	then
	obtain sop'' a'' v'' A'' L'' R'' W'' ys zs
	  where "xs=ys@Writesb True a'' sop'' v'' A'' L'' R'' W''#zs" and a_in: "a  A''"
	  by auto
	hence "x#xs = (x#ys)@Writesb True a'' sop'' v'' A'' L'' R'' W''#zs"
	  by auto
	thus ?thesis
	  using a_in
	  by blast
      next
	assume ?ghst
	then obtain A'' L'' R'' W'' ys zs where
	  "xs=ys@Ghostsb A'' L'' R'' W''#zs" and a_in: "a  A''"
	  by auto
	hence "x#xs = (x#ys)@Ghostsb A'' L'' R'' W''#zs"
	  by auto
	thus ?thesis
	  using a_in
	  by blast
      qed
    next
      case True
      
      then have "x#xs=[]@(Ghostsb A L R W#xs)"
	by (simp add: Ghostsb True)
      thus ?thesis
	using True
	by blast
    qed
  qed
qed


lemma split_Writesb_in_outstanding_refs:
  "a  outstanding_refs is_Writesb xs  (sop volatile v ys zs A L R W. xs = ys@(Writesb volatile a sop v A L R W#zs))"
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x xs)
  have a_in: "a  outstanding_refs is_Writesb (x # xs)" by fact
  show ?case
  proof (cases x)
    case (Writesb volatile a' sop v A L R W)
    show ?thesis
    proof (cases "a'=a")
      case False
      with a_in have "a  outstanding_refs is_Writesb xs"
	by (auto simp add: Writesb)
      from Cons.hyps [OF this] obtain sop'' volatile'' v'' A'' L'' R'' W'' ys zs
	where "xs=ys@Writesb volatile'' a sop'' v'' A'' L'' R'' W''#zs"
	by auto
      hence "x#xs = (x#ys)@Writesb volatile'' a sop'' v'' A'' L'' R'' W''#zs"
	by auto
      thus ?thesis
	by blast
    next
      case True
      then have "x#xs=[]@(Writesb volatile a sop v A L R W#xs)"
	by (simp add: Writesb True)
      thus ?thesis
	by blast
    qed
  next
    case Readsb 
    from a_in have "a  outstanding_refs is_Writesb xs"
      by (auto simp add: Readsb)
    from Cons.hyps [OF this] obtain sop'' volatile'' v'' A'' L'' R'' W'' ys zs
      where "xs=ys@Writesb volatile'' a sop'' v'' A'' L'' R'' W'' #zs"
      by auto
    hence "x#xs = (x#ys)@Writesb volatile'' a sop'' v'' A'' L'' R'' W''#zs"
      by auto
    thus ?thesis
      by blast
  next
    case Progsb
    from a_in have "a  outstanding_refs is_Writesb xs"
      by (auto simp add: Progsb)
    from Cons.hyps [OF this] obtain sop'' volatile'' v'' A'' L'' R'' W'' ys zs
      where "xs=ys@Writesb volatile'' a sop'' v'' A'' L'' R'' W''#zs"
      by auto
    hence "x#xs = (x#ys)@Writesb volatile'' a sop'' v'' A'' L'' R'' W''#zs"
      by auto
    thus ?thesis
      by blast
  next
    case Ghostsb
    from a_in have "a  outstanding_refs is_Writesb xs"
      by (auto simp add: Ghostsb)
    from Cons.hyps [OF this] obtain sop'' volatile'' v'' A'' L'' R'' W'' ys zs
      where "xs=ys@Writesb volatile'' a sop'' v'' A'' L'' R'' W''#zs"
      by auto
    hence "x#xs = (x#ys)@Writesb volatile'' a sop'' v'' A'' L'' R'' W''#zs"
      by auto
    thus ?thesis
      by blast
  qed
qed

lemma outstanding_refs_is_Writesb_union:
  "outstanding_refs is_Writesb xs = 
    (outstanding_refs is_volatile_Writesb xs  outstanding_refs is_non_volatile_Writesb xs)"
apply (induct xs)
apply  simp
subgoal for a
apply (case_tac a)
apply auto
done
done

  
lemma rtranclp_r_rtranclp: "r** x y; r y z  r** x z"
  by auto

lemma r_rtranclp_rtranclp: "r x y; r** y z  r** x z"
  by auto

lemma unshared_is_non_volatile_Writesb: "𝒮.
  non_volatile_writes_unshared 𝒮 sb; a  dom 𝒮; a  all_unshared sb  
  a  outstanding_refs is_non_volatile_Writesb 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.hyps [where 𝒮="(𝒮 ⊕⇘WR ⊖⇘AL)" ] Cons.prems
      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)       
    with Cons.hyps [where 𝒮="(𝒮 ⊕⇘WR ⊖⇘AL)"] Cons.prems show ?thesis by auto
  qed
qed

lemma outstanding_non_volatile_Readsb_acquired_or_read_only_reads:
  "𝒪 𝒮 pending_write. 
  non_volatile_owned_or_read_only pending_write 𝒮 𝒪 sb;
 a  outstanding_refs is_non_volatile_Readsb sb
 a  acquired_reads True sb 𝒪  a  read_only_reads 𝒪 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 True "(𝒮 ⊕⇘WR ⊖⇘AL)" "(𝒪  A - R)"] Cons.prems
      show ?thesis by auto
    next
      case False
      with Cons show ?thesis
	by (auto simp add: Writesb)
    qed
  next
    case (Readsb volatile a' t v) 
    show ?thesis
    proof (cases volatile)
      case False with Readsb Cons show ?thesis by auto
    next
      case True
      with Readsb Cons show ?thesis by auto
    qed
  next
    case Progsb with Cons show ?thesis by auto
  next
    case (Ghostsb A L R W) with Cons.hyps [of pending_write "(𝒮 ⊕⇘WR ⊖⇘AL)" "𝒪  A - R"] Cons.prems
    show ?thesis
      by auto
  qed
qed

lemma acquired_reads_union: "pending_writes A B. 
  a  acquired_reads pending_writes sb (A  B); a  A  a  acquired_reads pending_writes sb B"
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
      show ?thesis
      proof (cases pending_writes)
	case True
	from Cons.prems obtain 
	  a_in: "a  acquired_reads True sb (A  B  A' - R')" and
	  a_notin: "a  A"
	  by (simp add: Writesb volatile True)
	have "(A  B  A' - R')   (A  (B  A' - R'))"
	  by auto
	from acquired_reads_mono [OF this ] a_in
	have "a  acquired_reads True sb (A  (B  A' - R'))"
	  by auto

	from Cons.hyps [OF this a_notin]
	have "a  acquired_reads True sb (B  A' - R')".
	then show ?thesis
	  by (simp add: Writesb volatile True)
      next
	case False
	with Cons show ?thesis
	  by (auto simp add: Writesb volatile False)
      qed
    next
      case False
      with Cons show ?thesis
	by (auto simp add: Writesb False)
    qed
  next
    case Readsb with Cons show ?thesis
      by (auto split: if_split_asm)
  next
    case Progsb with Cons show ?thesis
      by (auto)
  next
    case (Ghostsb A' L' R' W')
    show ?thesis
    proof -
      from Cons.prems obtain 
	a_in: "a  acquired_reads pending_writes sb (A  B  A' - R')" and
	a_notin: "a  A"
        by (simp add: Ghostsb )
      have "(A  B  A' - R')   (A  (B  A' - R'))"
        by auto
      from acquired_reads_mono [OF this ] a_in
      have "a  acquired_reads pending_writes sb (A  (B  A' - R'))"
        by auto

      from Cons.hyps [OF this a_notin]
      have "a  acquired_reads pending_writes sb (B  A' - R')".
      then show ?thesis
        by (simp add: Ghostsb)
    qed
  qed
qed

   
lemma non_volatile_writes_unshared_no_outstanding_non_volatile_Writesb: "𝒮 𝒮'. 
  non_volatile_writes_unshared 𝒮 sb; 
  a  dom 𝒮' - dom 𝒮. a  outstanding_refs is_non_volatile_Writesb sb 
  non_volatile_writes_unshared 𝒮' 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
      from Cons.prems obtain 
	unshared_sb: "non_volatile_writes_unshared (𝒮 ⊕⇘WR ⊖⇘AL) sb" and
        no_refs_sb: "adom 𝒮' - dom 𝒮. a  outstanding_refs is_non_volatile_Writesb sb"
	by (simp add: Writesb True)
      from no_refs_sb have "adom (𝒮' ⊕⇘WR ⊖⇘AL) - dom (𝒮 ⊕⇘WR ⊖⇘AL). 
	a  outstanding_refs is_non_volatile_Writesb sb"
	by auto
      from Cons.hyps [OF unshared_sb this]
      show ?thesis
	by (simp add: Writesb True)
    next
      case False
      with Cons show ?thesis
	by (auto simp add: Writesb False)
    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 
      unshared_sb: "non_volatile_writes_unshared (𝒮 ⊕⇘WR ⊖⇘AL) sb" and
      no_refs_sb: "adom 𝒮' - dom 𝒮. a  outstanding_refs is_non_volatile_Writesb sb"
      by (simp add: Ghostsb)
    from no_refs_sb have "adom (𝒮' ⊕⇘WR ⊖⇘AL) - dom (𝒮 ⊕⇘WR ⊖⇘AL). 
      a  outstanding_refs is_non_volatile_Writesb sb"
      by auto
    from Cons.hyps [OF unshared_sb this]
    show ?thesis
      by (simp add: Ghostsb)
  qed
qed



theorem sharing_consis_share_all_until_volatile_write:
  "𝒮 ts'. ownership_distinct ts; sharing_consis 𝒮 ts; length ts' = length ts;
        i < length ts. 
              (let (_,_,_,sb,_,𝒪,_) = ts!i;
                   (_,_,_,sb',_,𝒪',_) = ts'!i 
               in 𝒪' = acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪  
                  sb' = dropWhile (Not  is_volatile_Writesb) sb) 
       sharing_consis (share_all_until_volatile_write ts 𝒮) ts' 
       dom (share_all_until_volatile_write ts 𝒮) - dom 𝒮  
           ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts) 
       dom 𝒮 - dom (share_all_until_volatile_write ts 𝒮)  
           ((λ(_,_,_,sb,_,𝒪,_). all_acquired sb  𝒪) ` set ts)"
proof (induct ts)
  case Nil thus ?case by auto
next
  case (Cons t ts)
  have leq: "length ts' = length (t#ts)" by fact
  have sim: "i < length (t#ts). 
              (let (_,_,_,sb,_,𝒪,_) = (t#ts)!i;
                   (_,_,_,sb',_,𝒪',_) = ts'!i 
               in 𝒪' = acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪  
                  sb' = dropWhile (Not  is_volatile_Writesb) sb)" 
    by fact
  obtain p "is" 𝒪  𝒟 θ sb 
    where t: "t = (p,is,θ,sb,𝒟,𝒪,)"
    by (cases t)

  from leq obtain t' ts'' where ts': "ts'=t'#ts''" and leq': "length ts'' = length ts"
    by (cases ts') force+
  
  obtain p' "is'" 𝒪' ℛ' 𝒟' θ' sb' 
    where t': "t' = (p',is',θ',sb',𝒟',𝒪',ℛ')"
    by (cases t')

  from sim [rule_format, of 0] t t' ts'
  obtain 𝒪': "𝒪' = acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪" and
         sb': "sb' = dropWhile (Not  is_volatile_Writesb) sb"
    by auto

  from sim ts'
  have sim': "i < length ts. 
              (let (_,_,_,sb,_,𝒪,) = ts!i;
                   (_,_,_,sb',_,𝒪',) = ts''!i 
               in 𝒪' = acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪  
                  sb' = dropWhile (Not  is_volatile_Writesb) sb)"
    by auto
    

  have consis: "sharing_consis 𝒮 (t#ts)" by fact
  then interpret sharing_consis 𝒮 "(t#ts)".
  from sharing_consis [of 0] t
  have consis_sb: "sharing_consistent 𝒮 𝒪 sb"
    by fastforce
  from sharing_consistent_takeWhile [OF this]
  have consis': "sharing_consistent 𝒮 𝒪 (takeWhile (Not  is_volatile_Writesb) sb)"
    by simp
  
  let ?𝒮' = "(share (takeWhile (Not  is_volatile_Writesb) sb) 𝒮)"
  from freshly_shared_owned [OF consis']
  have fresh_owned: "dom ?𝒮' - dom 𝒮  𝒪".
  from unshared_all_unshared [OF consis'] unshared_acquired_or_owned [OF consis']
  have unshared_acq_owned: "dom 𝒮 - dom ?𝒮'
                  all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  𝒪"
    by simp

    
  have dist: "ownership_distinct (t#ts)" by fact
  from ownership_distinct_tl [OF this]
  have dist': "ownership_distinct ts" .

  
  from sharing_consis_tl [OF consis]
  interpret consis': sharing_consis 𝒮 "ts".
    

  from dist interpret ownership_distinct "(t#ts)".

  (* FIXME: this proof appears in several times. Maybe make a lemma *)
  have sep: 
    "i < length ts. let (_,_,_,sb',_,_,_) = ts!i in 
          all_acquired sb'  dom 𝒮 - dom ?𝒮' = {}  
          all_unshared sb'  dom ?𝒮' - dom 𝒮 = {}"
  proof -
    {
      fix i pi "isi" 𝒪i i 𝒟i θi sbi
      assume i_bound: "i < length ts"
      assume ts_i: "ts ! i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
      have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}  
            all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
      proof -
	from ownership_distinct [of "0" "Suc i"] ts_i t i_bound
	have dist: "(𝒪  all_acquired sb)  (𝒪i  all_acquired sbi) = {}"
	  by force


	from dist unshared_acq_owned all_acquired_takeWhile [of "(Not  is_volatile_Writesb)" sb]
	have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}"
	  by blast

	moreover
	
	from sharing_consis [of "Suc i"] ts_i i_bound
	have "sharing_consistent 𝒮 𝒪i sbi"
	  by force
	from unshared_acquired_or_owned [OF this]
	have "all_unshared sbi  all_acquired sbi  𝒪i".      
	with dist fresh_owned
	have "all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
	  by blast
      
	ultimately show ?thesis by simp
      qed
    }
    thus ?thesis
      by (fastforce simp add: Let_def)
  qed

      
      
  from consis'.sharing_consis_preservation [OF sep]
  have consis_ts: "sharing_consis ?𝒮' ts".


  from Cons.hyps [OF dist' this leq' sim']
  obtain consis_ts'':
    "sharing_consis (share_all_until_volatile_write ts ?𝒮') ts''" and

    fresh: "dom (share_all_until_volatile_write ts ?𝒮') - dom ?𝒮'  
            ((λ(_,_,_,_,_,𝒪,).  𝒪) ` set ts)" and
    unshared: "dom ?𝒮' - dom (share_all_until_volatile_write ts ?𝒮')  
           ((λ(_,_,_,sb,_,𝒪,). all_acquired sb  𝒪)` set ts)"
    by auto


  from sharing_consistent_append [of _ _ "(takeWhile (Not  is_volatile_Writesb) sb)" 
  "(dropWhile (Not  is_volatile_Writesb) sb)"] consis_sb
  have consis_t': "sharing_consistent ?𝒮' 𝒪' sb'"
    by (simp add: 𝒪' sb')


  have fresh_dist: "all_acquired sb'  dom ?𝒮' - dom (share_all_until_volatile_write ts ?𝒮') = {}"
  proof -
    have "all_acquired sb'   ((λ(_,_,_,sb,_,𝒪,_). all_acquired sb  𝒪)` set ts) = {}"
    proof -
      {
	fix x
	assume x_sb': "x  all_acquired sb'"
	assume x_ts: "x   ((λ(_,_,_,sb,_,𝒪,_). all_acquired sb  𝒪)` set ts)"
	have False
	proof -
	  from x_ts
	  obtain i pi isi 𝒪i i 𝒟i θi sbi where 
	    i_bound: "i < length ts" and
            ts_i: "ts!i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)" and
	    x_in: "x  all_acquired sbi  𝒪i"
	    by (force simp add:  in_set_conv_nth)
	  from ownership_distinct [of "0" "Suc i"] ts_i t i_bound
	  have dist: "(𝒪  all_acquired sb)  (𝒪i  all_acquired sbi) = {}"
	    by force
	  with x_sb' x_in all_acquired_dropWhile [of "(Not  is_volatile_Writesb)" "sb"] show False
	    by (auto simp add: sb')
	qed
      }
      thus ?thesis by blast
    qed
    with unshared show ?thesis
      by blast
  qed

  have unshared_dist: "all_unshared sb'  dom (share_all_until_volatile_write ts ?𝒮') - dom ?𝒮' = {}"
  proof -
    from unshared_acquired_or_owned [OF consis_t']
    have "all_unshared sb'  all_acquired sb'  𝒪'".
    also
    from all_acquired_dropWhile [of "(Not  is_volatile_Writesb)" "sb"]
    acquired_all_acquired [of True "takeWhile (Not  is_volatile_Writesb) sb" 𝒪]
    all_acquired_takeWhile [of "(Not  is_volatile_Writesb)" "sb"]
    have "all_acquired sb'  𝒪'  all_acquired sb  𝒪"
      by (auto simp add: sb' 𝒪')
    finally
    have "all_unshared sb'  (all_acquired sb  𝒪)".

    moreover

    have "(all_acquired sb  𝒪)   ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts) = {}"
    proof -
      {
	fix x
	assume x_sb': "x  all_acquired sb  𝒪"
	assume x_ts: "x   ((λ(_,_,_,_,_,𝒪,_). 𝒪)` set ts)"
	have False
	proof -
	  from x_ts
	  obtain i pi isi 𝒪i i 𝒟i θi sbi where 
	    i_bound: "i < length ts" and
            ts_i: "ts!i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)" and
	    x_in: "x  𝒪i"
	    by (force simp add:  in_set_conv_nth)
	  from ownership_distinct [of "0" "Suc i"] ts_i t i_bound
	  have dist: "(𝒪  all_acquired sb)  (𝒪i  all_acquired sbi) = {}"
	    by force
	  with x_sb' x_in show False
	    by (auto simp add: sb')
	qed
      }
      thus ?thesis by blast
    qed
    ultimately show ?thesis
      using fresh by fastforce
  qed

  from sharing_consistent_preservation [OF consis_t' fresh_dist unshared_dist]
  have consis_ts: "sharing_consistent (share_all_until_volatile_write ts ?𝒮') 𝒪' sb'".
  note sharing_consis_Cons [OF consis_ts'' consis_ts, of p' is' θ' 𝒟' ]
  moreover
  from fresh fresh_owned
  have "dom (share_all_until_volatile_write ts ?𝒮') - dom 𝒮  
          𝒪   ((λ(_,_,_,_,_,𝒪,_).  𝒪) ` set ts)"  
    by auto
  moreover
  from unshared unshared_acq_owned all_acquired_takeWhile [of "(Not  is_volatile_Writesb)" sb]
  have "dom 𝒮 - dom (share_all_until_volatile_write ts ?𝒮')  
          all_acquired sb  𝒪   ((λ(_,_,_,sb,_,𝒪,_). all_acquired sb  𝒪) ` set ts)"
    by auto
  ultimately

  show ?case
    by (auto simp add: t ts' t')
qed


corollary sharing_consistent_share_all_until_volatile_write: 
assumes dist: "ownership_distinct ts" 
assumes consis: "sharing_consis 𝒮 ts" 
assumes i_bound: "i < length ts" 
assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)" 
shows "sharing_consistent (share_all_until_volatile_write ts 𝒮) 
                          (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪) 
                          (dropWhile (Not  is_volatile_Writesb) sb)"
proof -
  define ts' where "ts' == map (λ(p,is,θ,sb,𝒟,𝒪,).  
                    (p,is,θ,
                          dropWhile (Not  is_volatile_Writesb) sb,𝒟,acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪,)) ts"
  have leq: "length ts' = length ts"
    by (simp add: ts'_def)

  have flush: "i < length ts. 
              (let (_,_,_,sb,_,𝒪,_) = ts!i;
                   (_,_,_,sb',_,𝒪',_) = ts'!i 
               in 𝒪' = acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪  
                  sb' = dropWhile (Not  is_volatile_Writesb) sb)"
    by (auto simp add: ts'_def Let_def)

  from sharing_consis_share_all_until_volatile_write [OF dist consis leq flush]
  interpret sharing_consis "(share_all_until_volatile_write ts 𝒮)" "ts'" by simp
  from i_bound leq ts_i sharing_consis [of i]
  show ?thesis
    by (force simp add: ts'_def)
qed



lemma restrict_map_UNIV [simp]: "S |` UNIV = S"
  by (auto simp add: restrict_map_def)

(*
lemma share_takeWhile_non_volatile_Writesb: 
  "⋀S.  (share (takeWhile (Not ∘ is_volatile_Writesb) sb) S) =
  S ⊖(all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sb)) (all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb))"
apply (induct sb)
apply  simp
apply (case_tac a)
apply (auto intro: share_mono_in simp add: restrict_shared_fuse)
done
*)

(*
lemma dom_share_takeWhile_non_volatile_Writesb: 
  "dom (share (takeWhile (Not ∘ is_volatile_Writesb) sb) S) = dom S - all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb)"
  apply (auto simp add: share_takeWhile_non_volatile_Writesb)
  done
*)
(*
lemma dom_share_takeWhile_non_volatile_Writesb: 
  "⋀S.  dom (share (takeWhile (Not ∘ is_volatile_Writesb) sb) S) ⊆ dom S - all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb)"
apply (induct sb)
apply  simp
apply (case_tac a)
apply (auto intro: share_mono_in simp add: restrict_shared_fuse)
  apply (auto simp add: )
  done
*)
lemma share_all_until_volatile_write_Read_commute:
  shows "S i.  i < length ls; ls!i=(p,Read volatile a t#is,θ,sb,𝒟,𝒪)
      
    
    share_all_until_volatile_write 
       (ls[i := (p,is, θ(tv), sb @ [Readsb volatile a t v],𝒟', 𝒪)]) S =
    share_all_until_volatile_write ls S"
proof (induct ls)
  case Nil thus ?case
    by simp
next
  case (Cons l ls)
  note i_bound =  i < length (l#ls)
  note ith = (l#ls)!i = (p,Read volatile a t#is,θ,sb,𝒟,𝒪)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,Read volatile a t#is,θ,sb,𝒟,𝒪)"
      by simp
    thus ?thesis 
      by (simp add: 0 share_append_Readsb del: fun_upd_apply )
  next
    case (Suc n)
    obtain pl "isl" 𝒪l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l)"
      by (cases l)
    from i_bound ith
    have "share_all_until_volatile_write
      (ls[n := (p,is, θ(tv), sb @ [Readsb volatile a t v],𝒟', 𝒪)]) 
      (share (takeWhile (Not  is_volatile_Writesb) sbl) S) =
      share_all_until_volatile_write ls (share (takeWhile (Not  is_volatile_Writesb) sbl) S)"
      apply -
      apply (rule Cons.hyps)
      apply (auto simp add: Suc l)
      done
    
    then
    show ?thesis
      by (simp add: Suc l del: fun_upd_apply)
  qed
qed

lemma share_all_until_volatile_write_Write_commute:
  shows "S i.  i < length ls; ls!i=(p,Write volatile a (D,f) A L R W#is,θ,sb,𝒟,𝒪)
      
    
    share_all_until_volatile_write 
       (ls[i := (p,is,θ, sb @ [Writesb volatile a t (f θ) A L R W], 𝒟', 𝒪)]) S =
    share_all_until_volatile_write ls S"
proof (induct ls)
  case Nil thus ?case
    by simp
next
  case (Cons l ls)
  note i_bound =  i < length (l#ls)
  note ith = (l#ls)!i = (p,Write volatile a (D,f) A L R W#is,θ,sb,𝒟,𝒪)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,Write volatile a (D,f) A L R W#is,θ,sb,𝒟,𝒪)"
      by simp
    thus ?thesis 
      by (simp add: 0 share_append_Writesb del: fun_upd_apply )
  next
    case (Suc n)
    obtain pl "isl" 𝒪l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l)"
      by (cases l)
    from i_bound ith
    have "share_all_until_volatile_write
      (ls[n := (p,is, θ, sb @ [Writesb volatile a t (f θ) A L R W],𝒟', 𝒪)]) 
      (share (takeWhile (Not  is_volatile_Writesb) sbl) S) =
      share_all_until_volatile_write ls (share (takeWhile (Not  is_volatile_Writesb) sbl) S)"
      apply -
      apply (rule Cons.hyps)
      apply (auto simp add: Suc l)
      done
    
    then
    show ?thesis
      by (simp add: Suc l del: fun_upd_apply)
  qed
qed

lemma share_all_until_volatile_write_RMW_commute:
  shows "S i.  i < length ls; ls!i=(p,RMW a t (D,f) cond ret A L R W#is,θ,[],𝒟,𝒪)
      
    
    share_all_until_volatile_write (ls[i := (p',is, θ', [],𝒟', 𝒪')]) S =
    share_all_until_volatile_write ls S"
proof (induct ls)
  case Nil thus ?case
    by simp
next
  case (Cons l ls)
  note i_bound =  i < length (l#ls)
  note ith = (l#ls)!i = (p,RMW a t (D,f) cond ret A L R W#is,θ,[],𝒟,𝒪)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,RMW a t (D,f) cond ret A L R W#is,θ,[],𝒟,𝒪)"
      by simp
    thus ?thesis 
      by (simp add: 0 share_append_Writesb del: fun_upd_apply )
  next
    case (Suc n)
    obtain pl "isl" 𝒪l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l)"
      by (cases l)
    from i_bound ith
    have "share_all_until_volatile_write
      (ls[n := (p',is,θ', [], 𝒟', 𝒪')]) 
      (share (takeWhile (Not  is_volatile_Writesb) sbl) S) =
      share_all_until_volatile_write ls (share (takeWhile (Not  is_volatile_Writesb) sbl) S)"
      apply -
      apply (rule Cons.hyps)
      apply (auto simp add: Suc l)
      done
    
    then
    show ?thesis
      by (simp add: Suc l del: fun_upd_apply)
  qed
qed

lemma share_all_until_volatile_write_Fence_commute:
  shows "S i.  i < length ls; ls!i=(p,Fence#is,θ,[],𝒟,𝒪,)
      
    
    share_all_until_volatile_write (ls[i := (p,is,θ, [], 𝒟', 𝒪,ℛ')]) S =
    share_all_until_volatile_write ls S"
proof (induct ls)
  case Nil thus ?case
    by simp
next
  case (Cons l ls)
  note i_bound =  i < length (l#ls)
  note ith = (l#ls)!i = (p,Fence#is,θ,[],𝒟,𝒪,)
  show ?case
  proof (cases i)
    case 0
    from ith 0 have l: "l = (p,Fence#is,θ,[],𝒟,𝒪,)"
      by simp
    thus ?thesis 
      by (simp add: 0 share_append_Writesb del: fun_upd_apply )
  next
    case (Suc n)
    obtain pl "isl" 𝒪l l 𝒟l θl sbl where l: "l = (pl,isl,θl,sbl,𝒟l,𝒪l,l)"
      by (cases l)
    from i_bound ith
    have "share_all_until_volatile_write
      (ls[n := (p,is, θ, [],𝒟', 𝒪,ℛ')]) 
      (share (takeWhile (Not  is_volatile_Writesb) sbl) S) =
      share_all_until_volatile_write ls (share (takeWhile (Not  is_volatile_Writesb) sbl) S)"
      apply -
      apply (rule Cons.hyps)
      apply (auto simp add: Suc l)
      done
    
    then
    show ?thesis
      by (simp add: Suc l del: fun_upd_apply)
  qed
qed



(*
lemma unshared_share_takeWhile:  
  assumes unshared: "non_volatile_writes_unshared 𝒮 xs"
  shows "non_volatile_writes_unshared (share (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒮) xs"
proof -
  from share_takeWhile_non_volatile_Writesb [of sb 𝒮]
  have "dom (share (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒮) ⊆ dom 𝒮"
    by auto
  from non_volatile_writes_unshared_antimono [OF this unshared]
  show ?thesis
    by simp
qed
*)
(*
lemma non_volatile_writes_unshared_share_all_until_volatile_write: "⋀𝒮. non_volatile_writes_unshared 𝒮 xs ⟹ 
       non_volatile_writes_unshared (share_all_until_volatile_write ts 𝒮) xs"
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) 

  from unshared_share_takeWhile [OF Cons.prems]
  have "non_volatile_writes_unshared (share (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒮) xs".
  from Cons.hyps [OF this]
  show ?case
    by (simp add: t)
qed
*)
(*
lemma share_takeWhile_non_volatile_Writesb_restrict:
  "dom (share (takeWhile (Not ∘ is_volatile_Writesb) sb) S) ⊆ dom S"
  apply (auto simp add: share_takeWhile_non_volatile_Writesb)
  done
*)

(*
lemma share_all_until_volatile_write_restrict:
  "⋀S. dom (share_all_until_volatile_write ts S) ⊆ dom S"
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) 

  from share_takeWhile_non_volatile_Writesb_restrict
  have "dom (share (takeWhile (Not ∘ is_volatile_Writesb) sb) S) ⊆ dom S".
  with Cons.hyps
  show ?case
    by (auto simp add: t)
qed
*)
(*
lemma share_all_until_volatile_all_unshared:
  "⋀S. dom (share_all_until_volatile_write ts S) = 
        dom S - ⋃(λ(_,_,_,sb,_,_,_). all_unshared (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) 
  from Cons.hyps 
  have "dom (share_all_until_volatile_write ts (share (takeWhile (Not ∘ is_volatile_Writesb) sb) S)) = 
          dom (share (takeWhile (Not ∘ is_volatile_Writesb) sb) S) 
           - ⋃(λ(_,_,_,sb,_,_,_). all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb)) ` set ts".
  moreover
  from dom_share_takeWhile_non_volatile_Writesb
  have "dom (share (takeWhile (Not ∘ is_volatile_Writesb) sb) S) = dom S - all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb)".
  ultimately
  have "dom (share_all_until_volatile_write (t#ts) S) = 
       dom S - all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb) - 
       ⋃(λ(_,_,_,sb,_,_). all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb)) ` set ts"
    by (simp add: t)
  also
  have "… = dom S - ⋃(λ(_,_,_,sb,_,_). all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb)) ` set (t#ts)"
    by (auto simp add: t)
  finally
  show ?case
    by (auto simp del: o_apply)
qed
    
lemma share_all_until_volatile_all_unshared_i:
  assumes i_bound: "i < length ts" 
  assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪)"
  assumes a_in: "a ∈ dom (share_all_until_volatile_write ts S)" 
  shows "a ∉ all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb)"
proof -
  from nth_mem [OF i_bound]  ts_i have "(p,is,θ,sb,𝒟,𝒪) ∈ set ts"
    by auto
  with a_in
  show ?thesis
    by (auto simp add: share_all_until_volatile_all_unshared)
qed
*)
    
lemma unshared_share_in: "S. a  dom S  a   all_unshared sb  a  dom (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
      show ?thesis
      proof -
	from Cons.prems obtain a_S: "a  dom S" and a_L: "a  L" and  a_sb: "a  all_unshared sb" 
	  by (clarsimp simp add: Writesb True)
	from a_S a_L have "a  dom (S ⊕⇘WR ⊖⇘AL)"
	  by auto
	from Cons.hyps [OF this a_sb]
	show ?thesis  
	  by (clarsimp simp add: Writesb True)
      qed
    next
      case False
      with Cons show ?thesis
	by (auto simp add: Writesb False)
    qed
  next
    case Readsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case Progsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case Ghostsb
    with Cons show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed


lemma dom_eq_dom_share_eq: "S S'. dom S = dom S'  dom (share sb S) = dom (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
      from Cons.prems
      have "dom (S ⊕⇘WR ⊖⇘A'L) = dom (S' ⊕⇘WR ⊖⇘A'L)"
	by auto
      from Cons.hyps [OF this]
      show ?thesis
	by (clarsimp simp add: Writesb True)
    next
      case False with Cons.hyps [of S S'] Cons.prems Writesb show ?thesis by auto
    qed
  next
    case Readsb with Cons.hyps [of S S'] Cons.prems show ?thesis by auto
  next
    case Progsb with Cons.hyps [of S S'] Cons.prems show ?thesis by auto
  next
    case (Ghostsb A' L R W)
    from Cons.prems
    have "dom (S ⊕⇘WR  ⊖⇘A'L) = dom (S' ⊕⇘WR  ⊖⇘A'L)"
      by auto
    from Cons.hyps [OF this]
    show ?thesis
      by (clarsimp simp add: Ghostsb)
  qed
qed
   
lemma share_union:
  "A B. a  dom (share sb (A ⊕⇘zB)); a  dom A  a  dom (share sb (Map.empty ⊕⇘zB))"
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_in: "a  dom (share sb ((A ⊕⇘zB) ⊕⇘WR ⊖⇘A'L))" and a_A: "a  dom A"
	by (clarsimp simp add: Writesb True)
      have "dom ((A ⊕⇘zB) ⊕⇘WR ⊖⇘A'L)  dom (A ⊕⇘z(B  R - L))"
	by auto
      from share_mono [OF this] a_in
      have "a  dom (share sb (A ⊕⇘z(B  R - L)))"
	by blast
      from Cons.hyps [OF this] a_A
      have "a  dom (share sb (Map.empty ⊕⇘z(B  R - L)))"
	by blast
      moreover
      have "dom (Map.empty ⊕⇘zB  R - L) = dom ((Map.empty ⊕⇘zB) ⊕⇘WR ⊖⇘A'L)"
	by auto
      note dom_eq_dom_share_eq [OF this, of sb]
      ultimately
      show ?thesis
	by (clarsimp simp add: Writesb True)
    next
      case False
      with Cons show ?thesis
	by (auto simp add: Writesb False)
    qed
  next
    case Readsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case Progsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case (Ghostsb A' L R W)
    from Cons.prems 
    obtain a_in: "a  dom (share sb ((A ⊕⇘zB) ⊕⇘WR ⊖⇘A'L))" and a_A: "a  dom A"
      by (clarsimp simp add: Ghostsb)
    have "dom ((A ⊕⇘zB) ⊕⇘WR ⊖⇘A'L)  dom (A ⊕⇘z(B  R - L))"
      by auto
    from share_mono [OF this] a_in
    have "a  dom (share sb (A ⊕⇘z(B  R - L)))"
      by blast
    from Cons.hyps [OF this] a_A
    have "a  dom (share sb (Map.empty ⊕⇘z(B  R - L)))"
      by blast
    moreover
    have "dom (Map.empty ⊕⇘zB  R - L) = dom ((Map.empty ⊕⇘zB) ⊕⇘WR ⊖⇘A'L)"
      by auto
    note dom_eq_dom_share_eq [OF this, of sb]
    ultimately
    show ?thesis
      by (clarsimp simp add: Ghostsb)
  qed
qed

   
lemma share_unshared_in: 
  "S. a  dom (share sb S)  a  dom (share sb Map.empty)  (a  dom S  a  all_unshared 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
      note volatile=this
      from Cons.prems
      have a_in: "a  dom (share sb (S ⊕⇘WR ⊖⇘AL))"
	by (clarsimp simp add: Writesb True)
      show ?thesis
      proof (cases "a  dom S")
	case True
	from Cons.hyps [OF a_in]
	have "a  dom (share sb Map.empty)  a  dom (S ⊕⇘WR ⊖⇘AL)  a  all_unshared sb".
	then show ?thesis
	proof 
	  assume "a  dom (share sb Map.empty)"
	  from share_mono_in [OF this]
	  have "a  dom (share sb (Map.empty ⊕⇘WR ⊖⇘AL))" by auto
	  then show ?thesis
	    by (clarsimp simp add: Writesb volatile True)
	next
	  assume "a  dom (S ⊕⇘WR ⊖⇘AL)  a  all_unshared sb"
	  then obtain "a  L" "a  all_unshared sb"
	    by auto
	  then show ?thesis by (clarsimp simp add: Writesb volatile True)
	qed
      next
	case False
	have "dom (S ⊕⇘WR ⊖⇘AL)  dom (S ⊕⇘W(R - L))"
	  by auto
	from share_mono [OF this] a_in
	have "a  dom (share sb (S ⊕⇘W(R - L)))" by blast
	from share_union [OF this False]
	have "a  dom (share sb (Map.empty ⊕⇘W(R - L)))".
	moreover
	have "dom (Map.empty ⊕⇘W(R - L)) = dom (Map.empty ⊕⇘WR ⊖⇘AL)"
	  by auto
	note dom_eq_dom_share_eq [OF this, of sb]
	ultimately
	show ?thesis
	  by (clarsimp simp add: Writesb True)
      qed
    next
      case False
      with Cons show ?thesis
	by (auto simp add: Writesb False)
    qed
  next
    case Readsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case Progsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case (Ghostsb A L R W)
    from Cons.prems
    have a_in: "a  dom (share sb (S ⊕⇘WR ⊖⇘AL))"
      by (clarsimp simp add: Ghostsb)
    show ?thesis
    proof (cases "a  dom S")
      case True
      from Cons.hyps [OF a_in]
      have "a  dom (share sb Map.empty)  a  dom (S ⊕⇘WR ⊖⇘AL)  a  all_unshared sb".
      then show ?thesis
      proof 
	assume "a  dom (share sb Map.empty)"
	from share_mono_in [OF this]
	have "a  dom (share sb (Map.empty ⊕⇘WR ⊖⇘AL))" by auto
        then show ?thesis
	  by (clarsimp simp add: Ghostsb True)
      next
	assume "a  dom (S ⊕⇘WR ⊖⇘AL)  a  all_unshared sb"
	then obtain "a  L" "a  all_unshared sb"
	  by auto
	then show ?thesis by (clarsimp simp add: Ghostsb True)
      qed
    next
      case False
      have "dom (S ⊕⇘WR ⊖⇘AL)  dom (S ⊕⇘W(R - L))"
        by auto
      from share_mono [OF this] a_in
      have "a  dom (share sb (S ⊕⇘W(R - L)))" by blast
      from share_union [OF this False]
      have "a  dom (share sb (Map.empty ⊕⇘W(R - L)))".
      moreover
      have "dom (Map.empty ⊕⇘W(R - L)) = dom (Map.empty ⊕⇘WR ⊖⇘AL)"
        by auto
      note dom_eq_dom_share_eq [OF this, of sb]
      ultimately
      show ?thesis
        by (clarsimp simp add: Ghostsb False)
    qed
  qed
qed

(* FIXME: move up *)
lemma dom_augment_rels_shared_eq: "dom (augment_rels S R ) = dom (augment_rels S' R )"
  by (auto simp add: augment_rels_def domIff split: option.splits if_split_asm)

lemma dom_eq_SomeD1: "dom m = dom n  m x = Some y  n x  None"
  by (auto simp add: dom_def)

lemma dom_eq_SomeD2: "dom m = dom n  n x = Some y  m x  None"
  by (auto simp add: dom_def)

lemma dom_augment_rels_rels_eq: "dom ℛ'  = dom   dom (augment_rels S R ℛ') = dom (augment_rels S R )"
  by (auto simp add: augment_rels_def domIff split: option.splits if_split_asm dest: dom_eq_SomeD1 dom_eq_SomeD2)


lemma dom_release_rels_eq: "𝒮  ℛ'. dom ℛ' = dom   
  dom (release sb 𝒮 ℛ') = dom (release sb 𝒮 )"
proof (induct sb)
  case Nil thus ?case by simp
next 
  case (Cons x sb)
  hence dr: "dom ℛ' = dom "
    by simp
  show ?case
  proof (cases x)
    case Writesb with Cons.hyps [OF dr] show ?thesis by (clarsimp)
  next 
    case Readsb with Cons.hyps [OF dr] show ?thesis by (clarsimp)
  next 
    case Progsb with Cons.hyps [OF dr] show ?thesis by (clarsimp)
  next
    case (Ghostsb A L R W)
    from Cons.hyps [OF dom_augment_rels_rels_eq [OF dr]]
    show ?thesis
     by (simp add: Ghostsb)
 qed
qed



lemma dom_release_shared_eq: "𝒮 𝒮' . dom (release sb 𝒮' ) = dom (release sb 𝒮 )"
proof (induct sb)
  case Nil thus ?case by simp
next 
  case (Cons x sb)
  show ?case
  proof (cases x)
    case Writesb with Cons.hyps show ?thesis by (clarsimp)
  next 
    case Readsb with Cons.hyps show ?thesis by (clarsimp)
  next 
    case Progsb with Cons.hyps show ?thesis by (clarsimp)
  next
    case (Ghostsb A L R W)
    have dr: "dom (augment_rels 𝒮' R ) = dom (augment_rels 𝒮 R )"
      by(rule dom_augment_rels_shared_eq)
    have "dom (release sb (𝒮'  R - L) (augment_rels 𝒮' R )) =
          dom (release sb (𝒮  R - L) (augment_rels 𝒮' R ))"
      by (rule Cons.hyps)
    also have "... = dom (release sb (𝒮  R - L) (augment_rels 𝒮 R ))"
      by (rule dom_release_rels_eq [OF dr])
    finally show ?thesis
      by (clarsimp simp add: Ghostsb)
  qed
qed


lemma share_other_untouched:
  "𝒪 𝒮. sharing_consistent 𝒮 𝒪 sb  a  𝒪  all_acquired sb share sb 𝒮 a = 𝒮 a"
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_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_owns: "a  𝒪" and a_A: "a  A" and a_sb: "a  all_acquired sb"
	by ( simp add: Writesb True )

      from a_owns a_A a_sb 
      have "a  𝒪  A - R  all_acquired sb"
        by auto
      from Cons.hyps [OF consis' this]
      have "share sb (𝒮 ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a".
      moreover have "(𝒮 ⊕⇘WR ⊖⇘AL) a = 𝒮 a"
      using L_A A_R R_owns a_owns a_A
        by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
      ultimately show ?thesis
        by (simp add: Writesb True)
    next
      case False with Cons show ?thesis
        by (auto simp add: Writesb False)
    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_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_owns: "a  𝒪" and a_A: "a  A" and a_sb: "a  all_acquired sb"
      by ( simp add: Ghostsb )

    from a_owns a_A a_sb 
    have "a  𝒪  A - R  all_acquired sb"
      by auto
    from Cons.hyps [OF consis' this]
    have "share sb (𝒮 ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a".
    moreover have "(𝒮 ⊕⇘WR ⊖⇘AL) a = 𝒮 a"
    using L_A A_R R_owns a_owns a_A
      by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
    ultimately show ?thesis
      by (simp add: Ghostsb)
  qed
qed

lemma shared_owned: "𝒪 𝒮. sharing_consistent 𝒮 𝒪 sb  a  dom 𝒮  a  dom (share sb 𝒮)  
    a  𝒪  all_acquired 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

      from Cons.prems obtain 
	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_notin: "a  dom 𝒮" and a_in: "a  dom (share sb (𝒮 ⊕⇘WR ⊖⇘AL))"
	by ( simp add: Writesb True )
      
      show ?thesis
      proof (cases "a  𝒪")
        case True thus ?thesis by auto
      next
        case False
        with a_notin R_owns A_shared_owns L_A A_R have "a  dom (𝒮 ⊕⇘WR ⊖⇘AL)"
          by (auto)
        from Cons.hyps [OF consis' this a_in]
        show ?thesis
          by (auto simp add: Writesb True)
      qed
    next
      case False with Cons show ?thesis
        by (auto simp add: Writesb False)
    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_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_notin: "a  dom 𝒮" and a_in: "a  dom (share sb (𝒮 ⊕⇘WR ⊖⇘AL))"
      by (simp add: Ghostsb)
      
    show ?thesis
    proof (cases "a  𝒪")
      case True thus ?thesis by auto
    next
      case False
      with a_notin R_owns A_shared_owns L_A A_R have "a  dom (𝒮 ⊕⇘WR ⊖⇘AL)"
        by (auto)
      from Cons.hyps [OF consis' this a_in]
      show ?thesis
        by (auto simp add: Ghostsb)
    qed
  qed
qed

(*
      sharing consistent:
      a ∈ dom (share (takeWhile (Not ∘ is_volatile_Writesb) (sb!i)) empty)
      ⟹ a ∈ dom (share_all_until_volatile_write ts 𝒮)
(it should even be the same boolean value )
generalization

      a ∉ dom S ⟹
      a ∈ dom (share (takeWhile (Not ∘ is_volatile_Writesb) (sb!i)) S)
      ⟹ a ∈ dom (share_all_until_volatile_write ts S')

also generalize freshly_shared_owned; cf. share_owns_shared

if a is owned by a thread, no other thread can mess around with sharing!


a ∉ 𝒪; sharing_consistent 𝒮 𝒪 sb;
share sb 𝒮 a = share a
*)

(* FIXME: move up *)
lemma share_all_shared_in: "a  dom (share sb 𝒮)  a  dom 𝒮  a  all_shared sb"
using sharing_consistent_share_all_shared [of sb 𝒮]
  by auto

lemma share_all_until_volatile_write_unowned:
  assumes dist: "ownership_distinct ts"
  assumes consis: "sharing_consis 𝒮 ts"
  assumes other: "i p is θ sb 𝒟 𝒪 . i < length ts  ts!i = (p,is,θ,sb,𝒟,𝒪,) 
              a  𝒪  all_acquired sb"
  shows "share_all_until_volatile_write ts 𝒮 a = 𝒮 a"
using dist consis other 
proof (induct ts arbitrary: 𝒮)
  case Nil thus ?case by simp
next
  case (Cons t ts)
  obtain pt "ist" 𝒪t t 𝒟t θt sbt where
    t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
    by (cases t)

  from Cons.prems t obtain
    other': "i p is θ sb 𝒟 𝒪 . i < length ts  ts!i = (p,is,θ,sb,𝒟,𝒪,) 
              a  𝒪  all_acquired sb" and
    a_notin: "a  𝒪t  all_acquired sbt"
  apply -
  apply (rule that)
  apply  clarsimp
         subgoal for i p "is" θ sb 𝒟 𝒪 ℛ
         apply (drule_tac x="Suc i" in spec)
         apply clarsimp
         done
  apply (drule_tac x="0" in spec)
  apply clarsimp
  done 

  have dist: "ownership_distinct (t#ts)" by fact
  then interpret ownership_distinct "t#ts".
  have consis: "sharing_consis 𝒮 (t#ts)" by fact
  then interpret sharing_consis 𝒮 "t#ts".

  from ownership_distinct_tl [OF dist]
  have dist': "ownership_distinct ts".
  
  from sharing_consis_tl [OF consis]
  have consis': "sharing_consis 𝒮 ts".
  then
  interpret consis': sharing_consis 𝒮 "ts".
  
  let ?𝒮' = "(share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮)"

  from sharing_consis [of 0, simplified, OF t]
  have "sharing_consistent 𝒮 𝒪t sbt".
  from sharing_consistent_takeWhile [OF this]
  have consis_sb: "sharing_consistent 𝒮 𝒪t (takeWhile (Not  is_volatile_Writesb) sbt)".
  from freshly_shared_owned [OF consis_sb]
  have fresh_owned: "dom ?𝒮' - dom 𝒮  𝒪t".
  from unshared_all_unshared [OF consis_sb] unshared_acquired_or_owned [OF consis_sb]
  have unshared_acq_owned: "dom 𝒮 - dom ?𝒮'
                  all_acquired (takeWhile (Not  is_volatile_Writesb) sbt)  𝒪t"
    by simp

  have sep: 
    "i < length ts. let (_,_,_,sb',_,_,_) = ts!i in 
          all_acquired sb'  dom 𝒮 - dom ?𝒮' = {}  
          all_unshared sb'  dom ?𝒮' - dom 𝒮 = {}"
  proof -
    {
      fix i pi "isi" 𝒪i i 𝒟i θi sbi
      assume i_bound: "i < length ts"
      assume ts_i: "ts ! i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
      have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}  
            all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
      proof -
	from ownership_distinct [of "0" "Suc i"] ts_i t i_bound
	have dist: "(𝒪t  all_acquired sbt)  (𝒪i  all_acquired sbi) = {}"
	  by force


	from dist unshared_acq_owned all_acquired_takeWhile [of "(Not  is_volatile_Writesb)" sbt]
	have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}"
	  by blast

	moreover
	
	from sharing_consis [of "Suc i"] ts_i i_bound
	have "sharing_consistent 𝒮 𝒪i sbi"
	  by force
	from unshared_acquired_or_owned [OF this]
	have "all_unshared sbi  all_acquired sbi  𝒪i".      
	with dist fresh_owned
	have "all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
	  by blast
      
	ultimately show ?thesis by simp
      qed
    }
    thus ?thesis
      by (fastforce simp add: Let_def)
  qed

  from consis'.sharing_consis_preservation [OF this]
  have "sharing_consis ?𝒮' ts".

  from Cons.hyps [OF dist' this other']
  have "share_all_until_volatile_write ts ?𝒮' a =
    share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮 a" .
  moreover
  from share_other_untouched [OF consis_sb] a_notin 
    all_acquired_append [of "(takeWhile (Not  is_volatile_Writesb) sbt)" "(dropWhile (Not  is_volatile_Writesb) sbt)"]
  have "share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮 a = 𝒮 a"
    by auto
  ultimately
  show ?case
    by (simp add: t)
qed

lemma share_shared_eq: "𝒮' 𝒮. 𝒮' a = 𝒮 a  share sb 𝒮' a = share sb 𝒮 a"
proof (induct sb)
  case Nil thus ?case by simp
next
  case (Cons x sb)
  have eq: "𝒮' a = 𝒮 a" by fact
  show ?case
  proof (cases x)
    case (Writesb volatile a' sop v A L R W)
    show ?thesis
    proof (cases volatile)
      case True
      
      have "(𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
      using eq by (auto simp add: augment_shared_def restrict_shared_def)
      from Cons.hyps [of "(𝒮' ⊕⇘WR ⊖⇘AL)" "(𝒮 ⊕⇘WR ⊖⇘AL)", OF this]
      show ?thesis
        by (clarsimp simp add: Writesb True)
    next
      case False
      with Cons.hyps [of 𝒮' 𝒮] Cons.prems show ?thesis
	by (auto simp add: Writesb False)
    qed
  next
    case Readsb
    with Cons.hyps [of 𝒮' 𝒮] Cons.prems show ?thesis
      by (auto simp add: Readsb)
  next
    case Progsb
    with Cons.hyps [of 𝒮' 𝒮] Cons.prems show ?thesis
      by (auto simp add: Readsb)
  next
    case (Ghostsb A L R W)
    have "(𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
    using eq by (auto simp add: augment_shared_def restrict_shared_def)
    from Cons.hyps [of "(𝒮' ⊕⇘WR ⊖⇘AL)" "(𝒮 ⊕⇘WR ⊖⇘AL)", OF this]
    show ?thesis
      by (clarsimp simp add: Ghostsb)
  qed
qed

lemma  share_all_until_volatile_write_thread_local:
  assumes dist: "ownership_distinct ts"
  assumes consis: "sharing_consis 𝒮 ts"
  assumes i_bound: "i < length ts"
  assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)"
  assumes a_owned: "a  𝒪  all_acquired sb"
  shows "share_all_until_volatile_write ts 𝒮 a = share (takeWhile (Not  is_volatile_Writesb) sb) 𝒮 a"
using dist consis i_bound ts_i
proof (induct ts arbitrary: 𝒮 i)
  case Nil thus ?case by simp
next
  case (Cons t ts)


  obtain pt "ist" 𝒪t t 𝒟t θt sbt where
    t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
    by (cases t)

  have dist: "ownership_distinct (t#ts)" by fact
  then interpret ownership_distinct "t#ts".
  have consis: "sharing_consis 𝒮 (t#ts)" by fact
  then interpret sharing_consis 𝒮 "t#ts".
  
  from ownership_distinct_tl [OF dist]
  have dist': "ownership_distinct ts".

  from sharing_consis_tl [OF consis]
  have consis': "sharing_consis 𝒮 ts".
  then
  interpret consis': sharing_consis 𝒮 "ts".
  let ?𝒮' = "(share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮)"

  from sharing_consis [of 0, simplified, OF t]
  have "sharing_consistent 𝒮 𝒪t sbt".
  from sharing_consistent_takeWhile [OF this]
  have consis_sb: "sharing_consistent 𝒮 𝒪t (takeWhile (Not  is_volatile_Writesb) sbt)".
  from freshly_shared_owned [OF consis_sb]
  have fresh_owned: "dom ?𝒮' - dom 𝒮  𝒪t".
  from unshared_all_unshared [OF consis_sb] unshared_acquired_or_owned [OF consis_sb]
  have unshared_acq_owned: "dom 𝒮 - dom ?𝒮'
                  all_acquired (takeWhile (Not  is_volatile_Writesb) sbt)  𝒪t"
    by simp

  have sep: 
    "i < length ts. let (_,_,_,sb',_,_,_) = ts!i in 
          all_acquired sb'  dom 𝒮 - dom ?𝒮' = {}  
          all_unshared sb'  dom ?𝒮' - dom 𝒮 = {}"
  proof -
    {
      fix i pi "isi" 𝒪i i 𝒟i θi sbi
      assume i_bound: "i < length ts"
      assume ts_i: "ts ! i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
      have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}  
            all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
      proof -
	from ownership_distinct [of "0" "Suc i"] ts_i t i_bound
	have dist: "(𝒪t  all_acquired sbt)  (𝒪i  all_acquired sbi) = {}"
	  by force


	from dist unshared_acq_owned all_acquired_takeWhile [of "(Not  is_volatile_Writesb)" sbt]
	have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}"
	  by blast

	moreover
	
	from sharing_consis [of "Suc i"] ts_i i_bound
	have "sharing_consistent 𝒮 𝒪i sbi"
	  by force
	from unshared_acquired_or_owned [OF this]
	have "all_unshared sbi  all_acquired sbi  𝒪i".      
	with dist fresh_owned
	have "all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
	  by blast
      
	ultimately show ?thesis by simp
      qed
    }
    thus ?thesis
      by (fastforce simp add: Let_def)
  qed

  from consis'.sharing_consis_preservation [OF this]
  have consis_shared': "sharing_consis ?𝒮' ts".


  have aargh: "(Not  is_volatile_Writesb) = (λa. ¬ is_volatile_Writesb a)"
    by (rule ext) auto

  show ?case
  proof (cases "i")
    case 0
    with Cons.prems
    have t': "t = (p, is, θ, sb, 𝒟, 𝒪, )" 
      by simp
    
    {
      fix j pj "isj" θj sbj 𝒟j 𝒪j j
      assume j_bound: "j < length ts"
      assume ts_j: "ts ! j = (pj, isj, θj, sbj, 𝒟j, 𝒪j, j)"
      have "a  𝒪j  all_acquired sbj"
      proof -
        from ownership_distinct [of "0" "Suc j", simplified, OF j_bound t ts_j] t a_owned t' 0
        show ?thesis
          by auto
      qed
    }
    
    with share_all_until_volatile_write_unowned [OF dist' consis_shared', of a] 
    have "share_all_until_volatile_write ts ?𝒮' a = ?𝒮' a"
      by fastforce
    then show ?thesis
    using t t' 0
      by (auto simp add: Cons t aargh)
  next
    case (Suc n)
    with Cons.prems obtain n_bound: "n < length ts" and ts_n: "ts!n = (p,is,θ,sb,𝒟,𝒪,)"
      by auto
    from Cons.hyps [OF dist' consis_shared' n_bound ts_n]
    have "share_all_until_volatile_write ts ?𝒮' a =
            share (takeWhile (Not  is_volatile_Writesb) sb) ?𝒮' a" .
    moreover 
    from ownership_distinct [of "0" "Suc n"] t a_owned ts_n n_bound
    have "a  𝒪t  all_acquired sbt"
      by fastforce
    with share_other_untouched [OF consis_sb, of a]  
      all_acquired_append [of "(takeWhile (Not  is_volatile_Writesb) sbt)" "(dropWhile (Not  is_volatile_Writesb) sbt)"]
    have "?𝒮' a = 𝒮 a"
      by auto
    from share_shared_eq [of ?𝒮' a 𝒮,OF this ]
    have "share (takeWhile (Not  is_volatile_Writesb) sb) ?𝒮' a = 
           share (takeWhile (Not  is_volatile_Writesb) sb) 𝒮 a" .
    ultimately show ?thesis
    using t Suc 
      by (auto simp add: aargh)
  qed
qed

lemma share_all_until_volatile_write_thread_local':
  assumes dist: "ownership_distinct ts"
  assumes consis: "sharing_consis 𝒮 ts"
  assumes i_bound: "i < length ts"
  assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)"
  assumes a_owned: "a  𝒪  all_acquired sb"
  shows "share (dropWhile (Not  is_volatile_Writesb) sb) (share_all_until_volatile_write ts 𝒮) a = 
          share sb 𝒮 a"
proof -
  let ?take = "takeWhile (Not  is_volatile_Writesb) sb"
  let ?drop = "dropWhile (Not  is_volatile_Writesb) sb"
  from share_all_until_volatile_write_thread_local [OF dist consis i_bound ts_i a_owned]
  have "share_all_until_volatile_write ts 𝒮 a = share ?take 𝒮 a" .
  moreover
  from share_shared_eq [of "share_all_until_volatile_write ts 𝒮" a "share ?take 𝒮", OF this]
  have "share ?drop (share_all_until_volatile_write ts 𝒮) a = share ?drop (share ?take 𝒮) a" .
  thus ?thesis
  using share_append [of ?take ?drop 𝒮]
    by simp
qed

lemma (in ownership_distinct) in_shared_sb_share_all_until_volatile_write:
  assumes consis: "sharing_consis 𝒮 ts"
  assumes i_bound: "i < length ts"
  assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)"
  assumes a_owned: "a  𝒪  all_acquired sb"
  assumes a_share: "a  dom (share sb 𝒮)"
  shows "a  dom (share (dropWhile (Not  is_volatile_Writesb) sb) 
                    (share_all_until_volatile_write ts 𝒮))"
proof -
  have dist: "ownership_distinct ts" 
  using assms ownership_distinct
    apply -
    apply (rule ownership_distinct.intro)
    apply auto
    done
  from share_all_until_volatile_write_thread_local' [OF dist consis i_bound ts_i a_owned]
    a_share
  show ?thesis
    by (auto simp add: domIff)
qed

lemma owns_unshared_share_acquired:
  "𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb; a  𝒪; a  all_unshared sb
   a  dom (share sb 𝒮)  acquired True 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_owns: "a  𝒪" and A_shared_onws: "A  dom 𝒮  𝒪" and
	a_L: "a  L" and a_unsh: " a  all_unshared sb" and L_A: "L  A" and
        A_R: "A  R = {}" and R_owns: "R  𝒪" and
	consis': "sharing_consistent (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb"
	by (clarsimp simp add: Writesb volatile)
      have "a  dom (share sb (𝒮 ⊕⇘WR ⊖⇘AL))  acquired True sb (𝒪  A - R)"
      proof (cases "a  R")
	case True
	with a_L have "a  dom (𝒮 ⊕⇘WR ⊖⇘AL)"
	  by auto
	from unshared_share_in [OF this a_unsh]
	show ?thesis by blast
      next
	case False
	hence "a  𝒪  A - R"
	  using a_owns
	  by auto
	from Cons.hyps [OF consis' this a_unsh]
	show ?thesis .
      qed
      then
      show ?thesis
	by (clarsimp simp add: Writesb volatile) 
    next
      case False
      with Cons
      show ?thesis
	by (auto simp add: Writesb)
    qed
  next
    case Readsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case Progsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case (Ghostsb A L R W)
    from Cons.prems obtain 
      a_owns: "a  𝒪" and A_shared_onws: "A  dom 𝒮  𝒪" and
      a_L: "a  L" and a_unsh: " a  all_unshared sb" and L_A: "L  A" and
      A_R: "A  R = {}" and R_owns: "R  𝒪" and
      consis': "sharing_consistent (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb"
      by (clarsimp simp add: Ghostsb)
    have "a  dom (share sb (𝒮 ⊕⇘WR ⊖⇘AL))  acquired True sb (𝒪  A - R)"
    proof (cases "a  R")
      case True
      with a_L have "a  dom (𝒮 ⊕⇘WR ⊖⇘AL)"
        by auto
      from unshared_share_in [OF this a_unsh]
      show ?thesis by blast
    next
      case False
      hence "a  𝒪  A - R"
        using a_owns
	by auto
      from Cons.hyps [OF consis' this a_unsh]
      show ?thesis .
    qed
    then show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed

lemma shared_share_acquired: "𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb  
  a  dom 𝒮  a  dom (share sb 𝒮)  acquired True 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_shared: "a  dom 𝒮" 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"
	by (clarsimp simp add: Writesb True) 
      show ?thesis
      proof (cases "a  L")
	case False with a_shared
	have "a  dom (𝒮 ⊕⇘WR ⊖⇘AL)"
	  by auto
	from Cons.hyps [OF consis' this]
	show ?thesis
	  by (clarsimp simp add: Writesb volatile)
      next
	case True
	with L_A have a_A: "a  A"
	  by blast
	from sharing_consistent_mono_shared [OF _ consis', where 𝒮'="(𝒮 ⊕⇘WR)"]
	have "sharing_consistent (𝒮 ⊕⇘WR) (𝒪  A - R) sb"
	  by auto
	from Cons.hyps [OF this] a_shared
	have hyp: "a  dom (share sb (𝒮 ⊕⇘WR))  acquired True sb (𝒪  A - R)"
	  by auto
	{
	  assume "a  dom (share sb (𝒮 ⊕⇘WR))"
	  from share_unshared_in [OF this]
	  have "a  dom (share sb (𝒮 ⊕⇘WR ⊖⇘AL))  acquired True sb (𝒪  A - R)"
	  proof 
	    assume "a  dom (share sb Map.empty)"
	    from share_mono_in [OF this]
	    have "a  dom (share sb (𝒮 ⊕⇘WR ⊖⇘AL))"
	      by auto
	    thus ?thesis by blast
	  next
	    assume "a  dom (𝒮 ⊕⇘WR)  a  all_unshared sb"
	    hence a_unsh: "a  all_unshared sb" by blast
	    from a_A A_R have "a  𝒪  A - R"
	      by auto
	    from owns_unshared_share_acquired [OF consis' this a_unsh]
	    show ?thesis .
	  qed
	}
	with hyp show ?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 simp add: Readsb)
  next
    case Progsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case (Ghostsb A L R W)
    from Cons.prems obtain 
      a_shared: "a  dom 𝒮" 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"
      by (clarsimp simp add: Ghostsb) 
    show ?thesis
    proof (cases "a  L")
      case False with a_shared
      have "a  dom (𝒮 ⊕⇘WR ⊖⇘AL)"
        by auto
      from Cons.hyps [OF consis' this]
      show ?thesis
        by (clarsimp simp add: Ghostsb)
    next
      case True
      with L_A have a_A: "a  A"
        by blast
      from sharing_consistent_mono_shared [OF _ consis', where 𝒮'="(𝒮 ⊕⇘WR)"]
      have "sharing_consistent (𝒮 ⊕⇘WR) (𝒪  A - R) sb"
        by auto
      from Cons.hyps [OF this] a_shared
      have hyp: "a  dom (share sb (𝒮 ⊕⇘WR))  acquired True sb (𝒪  A - R)"
        by auto
      {
	assume "a  dom (share sb (𝒮 ⊕⇘WR))"
	from share_unshared_in [OF this]
	have "a  dom (share sb (𝒮 ⊕⇘WR ⊖⇘AL))  acquired True sb (𝒪  A - R)"
        proof 
	  assume "a  dom (share sb Map.empty)"
	  from share_mono_in [OF this]
	  have "a  dom (share sb (𝒮 ⊕⇘WR ⊖⇘AL))"
	    by auto
	  thus ?thesis by blast
        next
	  assume "a  dom (𝒮 ⊕⇘WR)  a  all_unshared sb"
	  hence a_unsh: "a  all_unshared sb" by blast
	  from a_A A_R have "a  𝒪  A - R"
	    by auto
	  from owns_unshared_share_acquired [OF consis' this a_unsh]
	  show ?thesis .
        qed
      }
      with hyp show ?thesis
        by (auto simp add: Ghostsb)
    qed
  qed
qed

lemma dom_release_takeWhile:
  "S .
  dom (release (takeWhile (Not  is_volatile_Writesb) sb) S ) = 
  dom   all_shared (takeWhile (Not  is_volatile_Writesb) sb)"
apply (induct sb)
apply  (clarsimp)
subgoal for a sb S ℛ
apply (case_tac a)
apply (auto simp add: augment_rels_def domIff split: if_split_asm option.splits)
done
done

lemma  share_all_until_volatile_write_share_acquired:
  assumes dist: "ownership_distinct ts"
  assumes consis: "sharing_consis 𝒮 ts"
  assumes a_notin: "a  dom 𝒮"
  assumes a_in: "a  dom (share_all_until_volatile_write ts 𝒮)" 
  shows "i < length ts. 
           let (_,_,_,sb,_,_,_) = ts!i 
           in a  all_shared (takeWhile (Not  is_volatile_Writesb) sb)"
using dist consis a_notin a_in 
proof (induct ts arbitrary: 𝒮 i)
  case Nil thus ?case by simp
next
  case (Cons t ts)

  have a_notin: "a  dom 𝒮" by fact
  obtain pt "ist" 𝒪t t 𝒟t θt sbt where
    t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
    by (cases t)    

  let ?take = "(takeWhile (Not  is_volatile_Writesb) sbt)"
  from t Cons.prems 
  have a_in: "a  dom (share_all_until_volatile_write ts (share ?take 𝒮))"
    by auto

  have dist: "ownership_distinct (t#ts)" by fact
  then interpret ownership_distinct "t#ts".
  have consis: "sharing_consis 𝒮 (t#ts)" by fact
  then interpret sharing_consis 𝒮 "t#ts".
  
  from ownership_distinct_tl [OF dist]
  have dist': "ownership_distinct ts".

  from sharing_consis_tl [OF consis]
  have consis': "sharing_consis 𝒮 ts".
  then
  interpret consis': sharing_consis 𝒮 "ts".
  let ?𝒮' = "(share ?take 𝒮)"

  from sharing_consis [of 0, simplified, OF t]
  have "sharing_consistent 𝒮 𝒪t sbt".
  from sharing_consistent_takeWhile [OF this]
  have consis_sb: "sharing_consistent 𝒮 𝒪t ?take".
  from freshly_shared_owned [OF consis_sb]
  have fresh_owned: "dom ?𝒮' - dom 𝒮  𝒪t".
  from unshared_all_unshared [OF consis_sb] unshared_acquired_or_owned [OF consis_sb]
  have unshared_acq_owned: "dom 𝒮 - dom ?𝒮'
                  all_acquired ?take  𝒪t"
    by simp

  have sep: 
    "i < length ts. let (_,_,_,sb',_,_,_) = ts!i in 
          all_acquired sb'  dom 𝒮 - dom ?𝒮' = {}  
          all_unshared sb'  dom ?𝒮' - dom 𝒮 = {}"
  proof -
    {
      fix i pi "isi" 𝒪i i 𝒟i θi sbi
      assume i_bound: "i < length ts"
      assume ts_i: "ts ! i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
      have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}  
            all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
      proof -
	from ownership_distinct [of "0" "Suc i"] ts_i t i_bound
	have dist: "(𝒪t  all_acquired sbt)  (𝒪i  all_acquired sbi) = {}"
	  by force


	from dist unshared_acq_owned all_acquired_takeWhile [of "(Not  is_volatile_Writesb)" sbt]
	have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}"
	  by blast

	moreover
	
	from sharing_consis [of "Suc i"] ts_i i_bound
	have "sharing_consistent 𝒮 𝒪i sbi"
	  by force
	from unshared_acquired_or_owned [OF this]
	have "all_unshared sbi  all_acquired sbi  𝒪i".      
	with dist fresh_owned
	have "all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
	  by blast
      
	ultimately show ?thesis by simp
      qed
    }
    thus ?thesis
      by (fastforce simp add: Let_def)
  qed

  from consis'.sharing_consis_preservation [OF this]
  have consis_shared': "sharing_consis ?𝒮' ts".


  have aargh: "(Not  is_volatile_Writesb) = (λa. ¬ is_volatile_Writesb a)"
    by (rule ext) auto

  show ?case
  proof (cases "a  all_shared ?take")
    case True
    thus ?thesis
    apply -
    apply (rule_tac x=0 in exI)
    apply (auto simp add: t aargh)
    done
  next
    case False

    have a_notin': "a  dom ?𝒮'"
    proof 
      assume "a  dom ?𝒮'"
      from share_all_shared_in [OF this] False a_notin
      show False
        by auto
    qed
    from Cons.hyps [OF dist' consis_shared' a_notin' a_in]
    obtain i where "i < length ts" and 
      rel: "let (p,is,θ,sb,𝒟,𝒪,) = ts!i 
            in a  all_shared (takeWhile (Not  is_volatile_Writesb) sb)"
      by (auto simp add: Let_def aargh)
    then show ?thesis
      apply -
      apply (rule_tac x = "Suc i" in exI)
      apply (auto simp add: Let_def aargh)
      done
  qed
qed

lemma all_shared_share_acquired: "𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb  
  a  all_shared sb  a  dom (share sb 𝒮)  acquired True 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_shared: "a  R  all_shared sb" 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"
	by (clarsimp simp add: Writesb True) 
      show ?thesis
      proof (cases "a  all_shared sb")
        case True
        from Cons.hyps [OF consis' True]
        show ?thesis
          by (clarsimp simp add: Writesb volatile)
      next
        case False
        with a_shared have "a  R"
          by auto
        with L_A A_R R_owns have "a  dom (𝒮 ⊕⇘WR ⊖⇘AL)"
          by auto
        from shared_share_acquired [OF consis' this]
        show ?thesis
          by (clarsimp 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 simp add: Readsb)
  next
    case Progsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case (Ghostsb A L R W)
    from Cons.prems obtain 
      a_shared: "a  R  all_shared sb" 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"
      by (clarsimp simp add: Ghostsb) 
    show ?thesis
    proof (cases "a  all_shared sb")
      case True
      from Cons.hyps [OF consis' True]
      show ?thesis
        by (clarsimp simp add: Ghostsb)
    next
      case False
      with a_shared have "a  R"
        by auto
      with L_A A_R R_owns have "a  dom (𝒮 ⊕⇘WR ⊖⇘AL)"
        by auto
      from shared_share_acquired [OF consis' this]
      show ?thesis
        by (clarsimp simp add: Ghostsb)
    qed
  qed
qed

lemma (in ownership_distinct) share_all_until_volatile_write_share_acquired:
  assumes consis: "sharing_consis 𝒮 ts"
  assumes i_bound: "i < length ts"
  assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)"
  assumes a_in: "a  dom (share_all_until_volatile_write ts 𝒮)" 
  shows "a  dom (share sb 𝒮)  a  acquired True sb 𝒪     
          (j < length ts.  j  i 
           (let (_,_,_,sbj,_,_,_) = ts!j 
            in a  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)))"
proof -
  from assms ownership_distinct have dist: "ownership_distinct ts" 
    apply - 
    apply (rule ownership_distinct.intro)
    apply simp
    done
  from consis
  interpret sharing_consis 𝒮 ts .
  from sharing_consis [OF i_bound ts_i]
  have consis_sb: "sharing_consistent 𝒮 𝒪 sb".

  let ?take_sb = "takeWhile (Not  is_volatile_Writesb) sb"
  let ?drop_sb = "dropWhile (Not  is_volatile_Writesb) sb"

  show ?thesis
  proof (cases "a  dom 𝒮")
    case True
    from shared_share_acquired [OF consis_sb True]
    have "a  dom (share sb 𝒮)  acquired True sb 𝒪".
    thus ?thesis by auto
  next
    case False
    from share_all_until_volatile_write_share_acquired [OF dist consis False a_in]
    obtain j where j_bound: "j < length ts" and 
      rel: "let (_,_,_,sbj,_,_,_) = ts!j 
            in a  all_shared (takeWhile (Not  is_volatile_Writesb) sbj)" 
      by auto
    show ?thesis
    proof (cases "j=i")
      case False
      with j_bound rel 
      show ?thesis
        by blast
    next
      case True
      with rel ts_i have "a  all_shared ?take_sb"
        by (auto simp add: Let_def)
      hence "a  all_shared sb" 
      using all_shared_append [of ?take_sb ?drop_sb]
        by auto
      from all_shared_share_acquired [OF consis_sb this]
      have "a  dom (share sb 𝒮)  acquired True sb 𝒪".
      thus ?thesis
        by auto
    qed
  qed
qed
    
(*
lemma all_unshared_acquired: 
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  assumes a_in: "a ∈ all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb)" 
  shows "a ∈ acquired True (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒪" 
proof -
  from unshared_acquired_or_owned [OF sharing_consistent_takeWhile [OF consis]]
  have "all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb) ⊆ 
    𝒪 ∪ all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sb)"
    by auto
  with a_in acquired_takeWhile_non_volatile_Writesb [of sb 𝒪]
  show ?thesis
    apply (auto )
qed
*)

(*
lemma not_acquired_all_unshared: 
  assumes consis: "sharing_consistent 𝒮 𝒪 sb"
  assumes a_notin: "a ∉ acquired True (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒪" 
  shows "a ∉ all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb)"
proof -
  from a_notin
  have "a ∉ 𝒪 ∪ all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sb)"
    by (simp add: acquired_takeWhile_non_volatile_Writesb) 
  moreover
  
  from unshared_acquired_or_owned [OF sharing_consistent_takeWhile [OF consis]]
  have "all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb) ⊆ 
    𝒪 ∪ all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sb)"
    by auto
  
  ultimately show ?thesis
    by auto
qed

*)

(*
lemma (in valid_sharing) unacquired_share_all_until_volatile_write:
shows
  "- ⋃ (λ(_,_,_,sb,_,𝒪,_).  acquired True (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒪) ` set ts 
    ⊆ dom (share_all_until_volatile_write ts 𝒮)" (is "- ?U ⊆ ?S")
proof
  fix a
  assume a_U: "a ∈ - ?U"
  show "a ∈ ?S"
  proof -
    from a_U have a_not_U: "a ∉ ?U"
      by auto
    with acquired_takeWhile_non_volatile_Writesb
    have a_notin: "a ∉ ⋃ (λ(_,_,_,sb,_,𝒪).  𝒪 ∪ all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sb)) ` set ts "
      by force
    from a_notin have a_unowned: "a ∉ ⋃ (λ(_,_,_,_,_,𝒪).  𝒪) ` set ts"
      by blast
    from a_unowned unowned_shared
    have a_shared: "a ∈ dom 𝒮"
      by auto

    moreover

    have "a ∉ ⋃ (λ(_,_,_,sb,_,𝒪).  all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb)) ` set ts "
          (is "a ∉ ?UNSH")
    proof
      assume "a ∈ ?UNSH"
      from in_Union_image_nth_conv [OF this]
      obtain i p "is" 𝒪 𝒟 θ sb where
	i_bound: "i < length ts" and
	ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪)" and
	a_in: "a ∈ all_unshared (takeWhile (Not ∘ is_volatile_Writesb) sb)"
	by fastforce
      from all_unshared_acquired [OF sharing_consis [OF i_bound ts_i] a_in]
      have "a ∈ acquired True (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒪".
      with a_not_U nth_mem [OF i_bound] ts_i
      show False

	by (auto)
    qed

    ultimately
    show ?thesis
      apply (simp only: share_all_until_volatile_all_unshared)
      apply blast
      done
  qed
qed
*)
lemma acquired_all_shared_in: 
  "A. a  acquired True sb A  a  acquired True sb {}  (a  A  a  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)
    show ?thesis
    proof (cases volatile)
      case True
      note volatile=this
      from Cons.prems
      have a_in: "a  acquired True sb (A  A' - R)"
	by (clarsimp simp add: Writesb True)
      show ?thesis
      proof (cases "a  A")
	case True
	from Cons.hyps [OF a_in]
	have "a  acquired True sb {}  a  A  A' - R  a  all_shared sb".
	then show ?thesis
	proof 
	  assume "a  acquired True sb {}"
	  from acquired_mono_in [OF this]
	  have "a  acquired True sb (A' - R)" by auto
	  then show ?thesis
	    by (clarsimp simp add: Writesb volatile True)
	next
	  assume "a  A  A' - R  a  all_shared sb"
	  then obtain "a  R" "a  all_shared sb"
	    by blast
	  then show ?thesis by (clarsimp simp add: Writesb volatile True)
	qed
      next
	case False
	have "(A  A' - R)  A  (A' - R)"
	  by blast
	from acquired_mono [OF this] a_in
	have "a  acquired True sb (A  (A' - R))" by blast
	from acquired_union_notin_first [OF this False]
	have "a  acquired True sb (A' - R)".
	then show ?thesis
	  by (clarsimp simp add: Writesb True)
      qed
    next
      case False
      with Cons show ?thesis
	by (auto simp add: Writesb False)
    qed
  next
    case Readsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case Progsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case (Ghostsb A' L R W)
    from Cons.prems
    have a_in: "a  acquired True sb (A  A' - R)"
      by (clarsimp simp add: Ghostsb)
    show ?thesis
    proof (cases "a  A")
      case True
      from Cons.hyps [OF a_in]
      have "a  acquired True sb {}  a  A  A' - R  a  all_shared sb".
      then show ?thesis
      proof 
	assume "a  acquired True sb {}"
	from acquired_mono_in [OF this]
	have "a  acquired True sb (A' - R)" by auto
        then show ?thesis
	  by (clarsimp simp add: Ghostsb True)
      next
	assume "a  A  A' - R  a  all_shared sb"
	then obtain "a  R" "a  all_shared sb"
	  by blast
	then show ?thesis by (clarsimp simp add: Ghostsb True)
      qed
    next
      case False
      have "(A  A' - R)  A  (A' - R)"
        by blast
      from acquired_mono [OF this] a_in
      have "a  acquired True sb (A  (A' - R))" by blast
      from acquired_union_notin_first [OF this False]
      have "a  acquired True sb (A' - R)".
      then show ?thesis
        by (clarsimp simp add: Ghostsb)
    qed
  qed
qed
   

lemma all_shared_acquired_in: "A. a  A  a  all_shared sb  a  acquired True sb A"
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
      show ?thesis
      proof -
	from Cons.prems obtain a_A: "a  A" and a_R: "a  R" and  a_sb: "a  all_shared sb" 
	  by (clarsimp simp add: Writesb True)
	from a_A a_R have "a  A  A' - R"
	  by blast
	from Cons.hyps [OF this a_sb]
	show ?thesis  
	  by (clarsimp simp add: Writesb True)
      qed
    next
      case False
      with Cons show ?thesis
	by (auto simp add: Writesb False)
    qed
  next
    case Readsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case Progsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case Ghostsb
    with Cons show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed

lemma owned_share_acquired: "𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb  
  a  𝒪  a  dom (share sb 𝒮)  acquired True 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_owned: "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"
	by (clarsimp simp add: Writesb True) 
      show ?thesis
      proof (cases "a  R")
	case False with a_owned
	have "a  𝒪  A - R"
	  by auto
	from Cons.hyps [OF consis' this]
	show ?thesis
	  by (clarsimp simp add: Writesb volatile)
      next
	case True
	from True L_A A_R have "a  dom (𝒮 ⊕⇘WR ⊖⇘AL)"
	  by auto
	from shared_share_acquired [OF consis' this]
	show ?thesis 
	  by (clarsimp simp add: Writesb volatile True)
      qed
    next
      case False
      with Cons
      show ?thesis
	by (auto simp add: Writesb)
    qed
  next
    case Readsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case Progsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case (Ghostsb A L R W)
    from Cons.prems obtain 
      a_owned: "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"
      by (clarsimp simp add: Ghostsb) 
    show ?thesis
    proof (cases "a  R")
      case False with a_owned
      have "a  𝒪  A - R"
        by auto
      from Cons.hyps [OF consis' this]
      show ?thesis
        by (clarsimp simp add: Ghostsb)
    next
      case True
      from True L_A A_R have "a  dom (𝒮 ⊕⇘WR ⊖⇘AL)"
        by auto
      from shared_share_acquired [OF consis' this]
      show ?thesis 
        by (clarsimp simp add: Ghostsb True)
    qed
  qed
qed

(*
lemma (in valid_sharing) a_unowned_by_others_owned_or_shared:
  assumes dist: "ownership_distinct ts"
  assumes i_bound: "i < length ts"
  assumes ts_i: "ts!i = (p,is,tsmp,sb,𝒟,𝒪)"
  assumes a_unowned_others:
        "∀j<length (map owns_sb ts). i ≠ j ⟶ 
          (let (𝒪j,sbj) = (map owns_sb ts)!j 
           in a ∉ acquired True (takeWhile (Not ∘ is_volatile_Writesb) sbj) 𝒪j)" 

  shows "a ∈ acquired True sb 𝒪 ∨ 
         a ∈ dom (share (dropWhile (Not ∘ is_volatile_Writesb) sb) (share_all_until_volatile_write ts 𝒮))"
proof -
  
  let ?take_sb = "(takeWhile (Not ∘ is_volatile_Writesb) sb)"
  let ?drop_sb = "(dropWhile (Not ∘ is_volatile_Writesb) sb)"

{
    fix j pj isj 𝒪j 𝒟j xsj sbj
    assume a_unowned: "a ∉ acquired True ?take_sb 𝒪"
    assume j_bound: "j < length ts"
    assume jth: "ts!j = (pj,isj,xsj, sbj, 𝒟j, 𝒪j)"
    have "a ∉ acquired True (takeWhile (Not ∘ is_volatile_Writesb) sbj) 𝒪j "
    proof (cases "i=j")
      case True
      from a_unowned ts_i jth True
      show ?thesis
	by auto
    next
      case False
      from a_unowned_others [rule_format, of j] j_bound jth False
      have "a ∉ acquired True (takeWhile (Not ∘ is_volatile_Writesb) sbj) 𝒪j"
	by auto
      then
      show ?thesis
	by auto
    qed
  } note lem = this

  from prems have consis: "sharing_consis 𝒮 ts"
    by (simp add: valid_sharing_def)
  
  from sharing_consistent_share_all_until_volatile_write [OF dist consis i_bound ts_i]
  have consis':
    "sharing_consistent (share_all_until_volatile_write ts 𝒮) (acquired True ?take_sb 𝒪) ?drop_sb"
    by simp

  {
    assume a_notin: "a ∉ acquired True sb 𝒪"
    have ?thesis
    proof (cases "a ∈ acquired True ?take_sb 𝒪")
      case True

      from owned_share_acquired [OF consis' True]
      have "a ∈ dom (share ?drop_sb (share_all_until_volatile_write ts 𝒮)) ∪ 
	        acquired True ?drop_sb (acquired True ?take_sb 𝒪)".
      with acquired_append [of True "?take_sb" "?drop_sb" 𝒪]
      show ?thesis
	by auto
    next
      case False
      assume "a ∉ acquired True (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒪"
      from lem [OF this]
      have "a ∈  - ⋃ (λ(_,_,_,sb,_,𝒪). acquired True (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒪) ` set ts"
	by (fastforce simp add: in_set_conv_nth)

      with unacquired_share_all_until_volatile_write 
      have a_in: "a ∈ dom (share_all_until_volatile_write ts 𝒮)"
	by auto

      from shared_share_acquired [OF consis' a_in] acquired_append [of True "?take_sb" "?drop_sb" 𝒪]
      show ?thesis
	by auto
    qed
  } 
  then
  show ?thesis
    by auto
qed

*)
lemma outstanding_refs_non_volatile_Readsb_all_acquired: 
"m 𝒮 𝒪 pending_write. 
  reads_consistent pending_write 𝒪 m sb;non_volatile_owned_or_read_only pending_write 𝒮 𝒪 sb;  
a  outstanding_refs is_non_volatile_Readsb sb
 a   𝒪  a  all_acquired sb 
    a  read_only_reads 𝒪 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
      note volatile=this
      from Cons.prems obtain 
	non_vo: "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) 
	            (𝒪  A - R) sb" and
        out_vol: "outstanding_refs is_volatile_Readsb sb = {}" and
	out: "a  outstanding_refs is_non_volatile_Readsb sb"
	by (clarsimp simp add: Writesb True) 
      show ?thesis
      proof (cases "a  𝒪")
	case True
	show ?thesis
	  by (clarsimp simp add: Writesb True volatile)
      next
	case False
	from outstanding_non_volatile_Readsb_acquired_or_read_only_reads [OF non_vo out]
	have a_in: "a  acquired_reads True sb (𝒪  A - R) 
                    a  read_only_reads (𝒪  A - R) sb"
	  by auto
	with acquired_reads_all_acquired [of True sb "(𝒪  A - R)"]
	show ?thesis 
	  by (auto simp add: Writesb volatile)
        qed
     next
      case False
      with Cons show ?thesis
	by (auto simp add: Writesb False)
    qed
  next
    case Readsb
    with Cons show ?thesis
      apply (clarsimp simp del: o_apply simp add: Readsb 
	acquired_takeWhile_non_volatile_Writesb split: if_split_asm)
      apply auto
      done
  next
    case Progsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case (Ghostsb A L)
    with Cons show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed



lemma outstanding_refs_non_volatile_Readsb_all_acquired_dropWhile: 
assumes consis: "reads_consistent pending_write 𝒪 m sb" 
assumes nvo: "non_volatile_owned_or_read_only pending_write 𝒮 𝒪 sb"
assumes out: "a  outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) sb)"
shows "a  𝒪  a  all_acquired sb 
       a  read_only_reads 𝒪 sb"
using outstanding_refs_append [of is_non_volatile_Readsb "takeWhile (Not  is_volatile_Writesb) sb" 
  "dropWhile (Not  is_volatile_Writesb) sb"] 
  outstanding_refs_non_volatile_Readsb_all_acquired [OF consis nvo, of a] out
by (auto)


lemma share_commute: 
  "L R 𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb; 
all_shared sb  L = {}; all_shared sb  A = {}; all_acquired sb  R = {};
all_unshared sb  R = {}; all_shared sb  R = {}  
  (share sb (𝒮 ⊕⇘WR ⊖⇘AL)) =
  (share sb 𝒮) ⊕⇘WR ⊖⇘AL"
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 
	L_prop: "(R'  all_shared sb)  L = {}" and 
	A_prop: "(R'  all_shared sb)  A = {}" and 
	R_acq_prop: "(A'  all_acquired sb)  R = {}" and
	R_prop:"(L'  all_unshared sb)  R = {}" and  
	R_prop_sh: "(R'  all_shared sb)  R = {}" and
	A'_shared_owns: "A'  dom 𝒮  𝒪" and L'_A': " L'  A'" and A'_R': "A'  R' = {}" and 
	R'_owns: "R'  𝒪" and
        consis': "sharing_consistent (𝒮 ⊕⇘W'R' ⊖⇘A'L') (𝒪  A' - R') sb" 
	by (clarsimp simp add: Writesb volatile)


      from L_prop obtain R'_L: "R'  L = {}" and acq_L: "all_shared sb  L = {}"
	by blast
      from A_prop obtain R'_A: "R'  A = {}" and acq_A: "all_shared sb  A = {}"
	by blast
      from R_acq_prop obtain A'_R: "A'  R = {}" and acq_R:"all_acquired sb  R = {}" 
	by blast
      from R_prop obtain L'_R: "L'  R = {}" and unsh_R: "all_unshared sb  R = {}"
	by blast
      from R_prop_sh obtain R'_R: "R'  R = {}" and sh_R: "all_shared sb  R = {}"
	by blast

      from Cons.hyps [OF consis' acq_L acq_A acq_R unsh_R sh_R ]
      have "share sb ((𝒮 ⊕⇘W'R' ⊖⇘A'L') ⊕⇘WR ⊖⇘AL) = share sb (𝒮 ⊕⇘W'R' ⊖⇘A'L') ⊕⇘WR ⊖⇘AL".

      moreover

      from R'_L L'_R R'_R R'_A A'_R (*A'_R' L'_A'*)
      have "((𝒮 ⊕⇘WR ⊖⇘AL) ⊕⇘W'R' ⊖⇘A'L') = ((𝒮 ⊕⇘W'R' ⊖⇘A'L') ⊕⇘WR ⊖⇘AL)"
	apply -
	apply (rule ext)
	apply (clarsimp simp add: augment_shared_def restrict_shared_def)
	apply (auto split: if_split_asm option.splits)
	done
      
      ultimately
      have "share sb ((𝒮 ⊕⇘WR ⊖⇘AL) ⊕⇘W'R' ⊖⇘A'L') = share sb (𝒮 ⊕⇘W'R' ⊖⇘A'L') ⊕⇘WR ⊖⇘AL"
	by simp
      then
      show ?thesis
	by (clarsimp simp add: Writesb volatile)
    next
      case False with Cons show ?thesis
	by (clarsimp simp add: Writesb False)
    qed
  next
    case Readsb with Cons show ?thesis
      by (clarsimp simp add: Readsb)
  next
    case Progsb with Cons show ?thesis
      by (clarsimp simp add: Progsb)
  next
    case (Ghostsb A' L' R' W') 
    from Cons.prems obtain 
      L_prop: "(R'  all_shared sb)  L = {}" and 
      A_prop: "(R'  all_shared sb)  A = {}" and 
      R_acq_prop: "(A'  all_acquired sb)  R = {}" and
      R_prop:"(L'  all_unshared sb)  R = {}" and  
      R_prop_sh: "(R'  all_shared sb)  R = {}" and
      A'_shared_owns: "A'  dom 𝒮  𝒪" and L'_A': " L'  A'" and A'_R': "A'  R' = {}" and 
      R'_owns: "R'  𝒪" and
      consis': "sharing_consistent (𝒮 ⊕⇘W'R' ⊖⇘A'L') (𝒪  A' - R') sb" 
      by (clarsimp simp add: Ghostsb)


    from L_prop obtain R'_L: "R'  L = {}" and acq_L: "all_shared sb  L = {}"
      by blast
    from A_prop obtain R'_A: "R'  A = {}" and acq_A: "all_shared sb  A = {}"
      by blast
    from R_acq_prop obtain A'_R: "A'  R = {}" and acq_R:"all_acquired sb  R = {}" 
      by blast
    from R_prop obtain L'_R: "L'  R = {}" and unsh_R: "all_unshared sb  R = {}"
      by blast
    from R_prop_sh obtain R'_R: "R'  R = {}" and sh_R: "all_shared sb  R = {}"
      by blast

    from Cons.hyps [OF consis' acq_L acq_A acq_R unsh_R sh_R ]
    have "share sb ((𝒮 ⊕⇘W'R' ⊖⇘A'L') ⊕⇘WR ⊖⇘AL) = share sb (𝒮 ⊕⇘W'R' ⊖⇘A'L') ⊕⇘WR ⊖⇘AL".

    moreover

    from R'_L L'_R R'_R R'_A A'_R (*A'_R' L'_A'*)
    have "((𝒮 ⊕⇘WR ⊖⇘AL) ⊕⇘W'R' ⊖⇘A'L') = ((𝒮 ⊕⇘W'R' ⊖⇘A'L') ⊕⇘WR ⊖⇘AL)"
      apply -
      apply (rule ext)
      apply (clarsimp simp add: augment_shared_def restrict_shared_def)
      apply (auto split: if_split_asm option.splits)
      done
      
    ultimately
    have "share sb ((𝒮 ⊕⇘WR ⊖⇘AL) ⊕⇘W'R' ⊖⇘A'L') = share sb (𝒮 ⊕⇘W'R' ⊖⇘A'L') ⊕⇘WR ⊖⇘AL"
      by simp
    then
    show ?thesis
      by (clarsimp simp add: Ghostsb)
  qed
qed


lemma share_all_until_volatile_write_commute:
" 𝒮 R L. ownership_distinct ts; sharing_consis 𝒮 ts;
        i p is 𝒪  𝒟 θ sb. i < length ts  ts!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {};
        i p is 𝒪  𝒟 θ sb. i < length ts  ts!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  A = {};
       i p is 𝒪  𝒟 θ sb.  i < length ts  ts!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  R = {};
       i p is 𝒪  𝒟 θ sb.  i < length ts  ts!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_unshared (takeWhile (Not  is_volatile_Writesb) sb)  R = {};
       i p is 𝒪  𝒟 θ sb.  i < length ts  ts!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  R = {} 
 
share_all_until_volatile_write ts 𝒮 ⊕⇘WR ⊖⇘AL = share_all_until_volatile_write ts (𝒮 ⊕⇘WR ⊖⇘AL)"
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".
  have consis: "sharing_consis 𝒮 (t#ts)" by fact
  then interpret sharing_consis 𝒮 "t#ts".

  have L_prop: "i p is 𝒪  𝒟 θ sb. i < length (t#ts)  (t#ts)!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {}" by fact
  hence L_prop': "i p is 𝒪  𝒟 θ sb. i < length (ts)  (ts)!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {}"
    by force
  have A_prop: "i p is 𝒪  𝒟 θ sb. i < length (t#ts)  (t#ts)!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  A = {}" by fact
  hence A_prop': "i p is 𝒪  𝒟 θ sb. i < length (ts)  (ts)!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  A = {}"
    by force
  have  R_prop_acq: "i p is 𝒪  𝒟 θ sb.  i < length (t#ts)  (t#ts)!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" by fact
  hence R_prop_acq': "i p is 𝒪  𝒟 θ sb.  i < length (ts)  (ts)!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  R = {}"
    by force

  have  R_prop: "i p is 𝒪  𝒟 θ sb.  i < length (t#ts)  (t#ts)!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_unshared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" by fact
  hence R_prop': "i p is 𝒪  𝒟 θ sb.  i < length (ts)  (ts)!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_unshared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}"
    by force

  have  R_prop_sh: "i p is 𝒪  𝒟 θ sb.  i < length (t#ts)  (t#ts)!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" by fact
  hence R_prop_sh': "i p is 𝒪  𝒟 θ sb.  i < length (ts)  (ts)!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}"
    by force

  from ownership_distinct_tl [OF dist]
  have dist': "ownership_distinct ts".

  from sharing_consis_tl [OF consis]
  have consis': "sharing_consis 𝒮 ts".
  then
  interpret consis': sharing_consis 𝒮 "ts".

  from L_prop [rule_format, of 0 p "is" θ sb 𝒟 𝒪 ] t 
  have sh_L: "all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {}" 
    by simp

  from A_prop [rule_format, of 0 p "is" θ sb 𝒟 𝒪 ] t 
  have sh_A: "all_shared (takeWhile (Not  is_volatile_Writesb) sb)  A = {}" 
    by simp

  from R_prop_acq [rule_format, of 0 p "is" θ sb 𝒟 𝒪 ] t 
  have acq_R: "all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" 
    by simp

  from R_prop [rule_format, of 0 p "is" θ sb 𝒟 𝒪  ] t 
  have unsh_R: "all_unshared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" 
    by simp


  from R_prop_sh [rule_format, of 0 p "is"  θ sb 𝒟 𝒪] t 
  have sh_R: "all_shared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" 
    by simp

  from sharing_consis [of 0, simplified, OF t]
  have "sharing_consistent 𝒮 𝒪 sb".
  from sharing_consistent_takeWhile [OF this]
  have consis_sb: "sharing_consistent 𝒮 𝒪 (takeWhile (Not  is_volatile_Writesb) sb)".

  from share_commute [OF consis_sb sh_L sh_A acq_R unsh_R sh_R]
  have share_eq: 
    "(share (takeWhile (Not  is_volatile_Writesb) sb) (𝒮 ⊕⇘WR ⊖⇘AL)) =
        (share (takeWhile (Not  is_volatile_Writesb) sb) 𝒮) ⊕⇘WR ⊖⇘AL".
    
  let ?𝒮' = "(share (takeWhile (Not  is_volatile_Writesb) sb) 𝒮)"

  from freshly_shared_owned [OF consis_sb]
  have fresh_owned: "dom ?𝒮' - dom 𝒮  𝒪".
  from unshared_all_unshared [OF consis_sb] unshared_acquired_or_owned [OF consis_sb]
  have unshared_acq_owned: "dom 𝒮 - dom ?𝒮'
                  all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  𝒪"
    by simp


  have sep: 
    "i < length ts. let (_,_,_,sb',_,_,_) = ts!i in 
          all_acquired sb'  dom 𝒮 - dom ?𝒮' = {}  
          all_unshared sb'  dom ?𝒮' - dom 𝒮 = {}"
  proof -
    {
      fix i pi "isi" 𝒪i i 𝒟i θi sbi
      assume i_bound: "i < length ts"
      assume ts_i: "ts ! i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
      have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}  
            all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
      proof -
	from ownership_distinct [of "0" "Suc i"] ts_i t i_bound
	have dist: "(𝒪  all_acquired sb)  (𝒪i  all_acquired sbi) = {}"
	  by force


	from dist unshared_acq_owned all_acquired_takeWhile [of "(Not  is_volatile_Writesb)" sb]
	have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}"
	  by blast

	moreover
	
	from sharing_consis [of "Suc i"] ts_i i_bound
	have "sharing_consistent 𝒮 𝒪i sbi"
	  by force
	from unshared_acquired_or_owned [OF this]
	have "all_unshared sbi  all_acquired sbi  𝒪i".      
	with dist fresh_owned
	have "all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
	  by blast
      
	ultimately show ?thesis by simp
      qed
    }
    thus ?thesis
      by (fastforce simp add: Let_def)
  qed

  from consis'.sharing_consis_preservation [OF sep]
  have sharing_consis': "sharing_consis (share (takeWhile (Not  is_volatile_Writesb) sb) 𝒮) ts".

  from Cons.hyps [OF dist' sharing_consis' L_prop' A_prop' R_prop_acq' R_prop' R_prop_sh']
  have "share_all_until_volatile_write ts ?𝒮' ⊕⇘WR ⊖⇘AL =
        share_all_until_volatile_write ts (?𝒮' ⊕⇘WR ⊖⇘AL)".

  then
  have "share_all_until_volatile_write ts
         ?𝒮' ⊕⇘WR ⊖⇘AL =
        share_all_until_volatile_write ts
          (share (takeWhile (Not  is_volatile_Writesb) sb) (𝒮 ⊕⇘WR ⊖⇘AL))"
    by (simp add: share_eq)
  then
  show ?case
    by (simp add: t)
qed

lemma share_append_Ghostsb: 
  "𝒮. outstanding_refs is_volatile_Writesb sb = {}  (share (sb @ [Ghostsb A L R W]) 𝒮) = (share sb 𝒮) ⊕⇘WR ⊖⇘AL"
apply (induct sb)
apply simp
subgoal for a sb 𝒮
apply (case_tac a)
apply auto
done
done

lemma share_append_Ghostsb':
  "𝒮. outstanding_refs is_volatile_Writesb sb  {}  
     (share (takeWhile (Not  is_volatile_Writesb) (sb @ [Ghostsb A L R W])) 𝒮) = 
      (share (takeWhile (Not  is_volatile_Writesb) sb) 𝒮)"
apply (induct sb)
apply  simp
subgoal for a sb 𝒮
apply (case_tac a)
apply force+
done
done

lemma share_all_until_volatile_write_append_Ghostsb: 
assumes no_out_VWritesb: "outstanding_refs is_volatile_Writesb sb = {}"
shows "𝒮 i. ownership_distinct ts; sharing_consis 𝒮 ts;
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,);
   j p is 𝒪  𝒟 θ sb. j < length ts  ij  ts!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {};
   j p is 𝒪  𝒟 θ sb. j < length ts  ij  ts!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  A = {};
       j p is 𝒪  𝒟 θ sb.  j < length ts  ij  ts!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  R = {};
       j p is 𝒪  𝒟 θ sb.  j < length ts  ij  ts!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_unshared (takeWhile (Not  is_volatile_Writesb) sb)  R = {};
       j p is 𝒪  𝒟 θ sb.  j < length ts  ij  ts!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}
  
  share_all_until_volatile_write (ts[i := (p', is',θ', sb @ [Ghostsb A L R W], 𝒟', 𝒪')]) 𝒮
                = share_all_until_volatile_write ts 𝒮  ⊕⇘WR ⊖⇘AL"
proof (induct ts)
  case Nil
  thus ?case by simp
next
  case (Cons t ts)
  obtain pt "ist" 𝒪t t 𝒟t acqt θt sbt where
    t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
    by (cases t) 
  have dist: "ownership_distinct (t#ts)" by fact
  then interpret ownership_distinct "t#ts".
  have consis: "sharing_consis 𝒮 (t#ts)" by fact
  then interpret sharing_consis 𝒮 "t#ts".

  have L_prop: "j p is 𝒪  𝒟 θ sb. j < length (t#ts)  ij  (t#ts)!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {}" by fact

  have A_prop: "j p is 𝒪  𝒟 θ sb. j < length (t#ts)  ij  (t#ts)!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  A = {}" by fact

  have  R_prop_acq: "j p is 𝒪  𝒟 θ sb.  j < length (t#ts)  ij (t#ts)!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" by fact
  have  R_prop: "j p is 𝒪  𝒟 θ sb.  j < length (t#ts)  ij (t#ts)!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_unshared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" by fact

  have  R_prop_sh: "j p is 𝒪  𝒟 θ sb.  j < length (t#ts)  ij  (t#ts)!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" by fact

  from ownership_distinct_tl [OF dist]
  have dist': "ownership_distinct ts".

  from sharing_consis_tl [OF consis]
  have consis': "sharing_consis 𝒮 ts".
  then
  interpret consis': sharing_consis 𝒮 "ts".


  from sharing_consis [of 0, simplified, OF t]
  have "sharing_consistent 𝒮 𝒪t sbt" .

  from sharing_consistent_takeWhile [OF this]
  have consis_sb: "sharing_consistent 𝒮 𝒪t (takeWhile (Not  is_volatile_Writesb) sbt)".

  let ?𝒮' = "(share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮)"

  from freshly_shared_owned [OF consis_sb]
  have fresh_owned: "dom ?𝒮' - dom 𝒮  𝒪t".
  from unshared_all_unshared [OF consis_sb] unshared_acquired_or_owned [OF consis_sb]
  have unshared_acq_owned: "dom 𝒮 - dom ?𝒮'
                  all_acquired (takeWhile (Not  is_volatile_Writesb) sbt)  𝒪t"
    by simp


  have sep: 
    "i < length ts. let (_,_,_,sb',_,_,_) = ts!i in 
          all_acquired sb'  dom 𝒮 - dom ?𝒮' = {}  
          all_unshared sb'  dom ?𝒮' - dom 𝒮 = {}"
  proof -
    {
      fix i pi "isi" 𝒪i i 𝒟i acqi θi sbi
      assume i_bound: "i < length ts"
      assume ts_i: "ts ! i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
      have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}  
            all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
      proof -
	from ownership_distinct [of "0" "Suc i"] ts_i t i_bound
	have dist: "(𝒪t  all_acquired sbt)  (𝒪i  all_acquired sbi) = {}"
	  by force


	from dist unshared_acq_owned all_acquired_takeWhile [of "(Not  is_volatile_Writesb)" sbt]
	have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}"
	  by blast

	moreover
	
	from sharing_consis [of "Suc i"] ts_i i_bound
	have "sharing_consistent 𝒮 𝒪i sbi"
	  by force
	from unshared_acquired_or_owned [OF this]
	have "all_unshared sbi  all_acquired sbi  𝒪i".      
	with dist fresh_owned
	have "all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
	  by blast
      
	ultimately show ?thesis by simp
      qed
    }
    thus ?thesis
      by (fastforce simp add: Let_def)
  qed

  from consis'.sharing_consis_preservation [OF sep]
  have sharing_consis': "sharing_consis (share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮) ts".

  show ?case
  proof (cases i)
    case 0
    with t Cons.prems have eqs: "pt=p" "ist=is" "𝒪t=𝒪" "t=" "θt=θ" "sbt=sb" "𝒟t=𝒟" 
      by auto

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


    from no_out_VWritesb
    have flush_all': "takeWhile (Not  is_volatile_Writesb) (sb@[Ghostsb A L R W]) = sb@[Ghostsb A L R W]"
      by (auto simp add: outstanding_refs_conv)

    have share_eq:
      "(share (takeWhile (Not  is_volatile_Writesb) (sb @ [Ghostsb A L R W])) 𝒮) = 
            (share (takeWhile (Not  is_volatile_Writesb) sb) 𝒮) ⊕⇘WR ⊖⇘AL"
      apply (simp only: flush_all flush_all')
      apply (rule share_append_Ghostsb [OF no_out_VWritesb])
      done

    from L_prop 0 have L_prop': 
    "i p is 𝒪  𝒟 θ sb.
      i < length ts 
      ts ! i = (p, is,θ, sb, 𝒟, 𝒪,) 
      all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {}"
      apply clarsimp
      subgoal for i1 p "is" 𝒪 ℛ 𝒟 θ sb
      apply (drule_tac x="Suc i1" in spec)
      apply auto
      done
      done
    from A_prop 0 have A_prop': 
    "i p is 𝒪  𝒟 θ sb.
      i < length ts 
      ts ! i = (p, is,θ, sb, 𝒟, 𝒪,) 
      all_shared (takeWhile (Not  is_volatile_Writesb) sb)  A = {}"
      apply clarsimp
      subgoal for i1 p "is" 𝒪 ℛ 𝒟 θ sb
      apply (drule_tac x="Suc i1" in spec)
      apply auto
      done
      done
    from R_prop_acq 0 have R_prop_acq':
        "i p is 𝒪  𝒟 θ sb.  i < length ts  ts!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" 
      apply clarsimp
      subgoal for i1 p "is" 𝒪 ℛ 𝒟 θ sb
      apply (drule_tac x="Suc i1" in spec)
      apply auto
      done
      done
    from R_prop 0 
    have  R_prop': 
      "i p is 𝒪  𝒟 θ sb.  i < length ts  ts!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_unshared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" 
      apply clarsimp
      subgoal for i1 p "is" 𝒪 ℛ 𝒟 θ sb
      apply (drule_tac x="Suc i1" in spec)
      apply auto
      done
      done
    from R_prop_sh 0 have R_prop_sh': 
      "i p is 𝒪  𝒟 θ sb.  i < length ts  ts!i=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" 
      apply clarsimp
      subgoal for i1 p "is" 𝒪 ℛ 𝒟 θ sb
      apply (drule_tac x="Suc i1" in spec)
      apply auto
      done
      done


    from share_all_until_volatile_write_commute [OF dist' sharing_consis' L_prop' A_prop' R_prop_acq' R_prop' 
      R_prop_sh']

    have "share_all_until_volatile_write ts (share (takeWhile (Not  is_volatile_Writesb) sb) 𝒮  ⊕⇘WR ⊖⇘AL) =
          share_all_until_volatile_write ts (share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮)  ⊕⇘WR ⊖⇘AL"
      by (simp add: eqs)
    with share_eq
    show ?thesis
      by (clarsimp simp add: 0 t)
  next
    case (Suc k)
    from L_prop Suc
    have L_prop': "j p is 𝒪  𝒟 θ sb. j < length (ts)  kj  (ts)!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  L = {}" by force

    from A_prop Suc
    have A_prop': "j p is 𝒪  𝒟 θ sb. j < length (ts)  kj  (ts)!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  A = {}" by force
    from R_prop_acq Suc have R_prop_acq':
        "j p is 𝒪  𝒟 θ sb.  j < length ts  kj  ts!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_acquired (takeWhile (Not  is_volatile_Writesb) sb)  R = {}"  by force

    from R_prop Suc
    have  R_prop': 
      "j p is 𝒪  𝒟 θ sb.  j < length ts  kj  ts!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_unshared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" by force

    from R_prop_sh Suc have R_prop_sh': 
      "j p is 𝒪  𝒟 θ sb.  j < length ts  kj  ts!j=(p,is,θ,sb,𝒟,𝒪,)  
                    all_shared (takeWhile (Not  is_volatile_Writesb) sb)  R = {}" by force

    from Cons.prems Suc obtain k_bound: "k < length ts" and ts_k: "ts!k =  (p, is,θ, sb, 𝒟, 𝒪,)"
      by auto

    from Cons.hyps [OF dist' sharing_consis' k_bound ts_k L_prop' A_prop' R_prop_acq' R_prop' R_prop_sh']
    show ?thesis
      by (clarsimp simp add: t Suc)
  qed
qed




(*
I think this is what should work:
share_all_until_volatile_write (ts[i :=(p',is',θ',sb',𝒟',𝒪',ℛ')]) (𝒮 ⊕W R ⊖A L) =
share (takeWhile (Not ∘ is_volatile_Writesb) sb') (share_all_until_volatile_write ts 𝒮 ⊕W R ⊖A L)
*)

lemma share_domain_changes:
  "𝒮 𝒮'. a  all_shared sb  all_unshared sb  share sb 𝒮' a = share sb 𝒮 a "
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_in: "a  R  all_shared sb  L  all_unshared sb"
        by (clarsimp simp add: Writesb True)
      show ?thesis
      proof (cases "a  R")
        case True
        from True have "(𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
          by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
        from  share_shared_eq [where 𝒮'="𝒮' ⊕⇘WR ⊖⇘AL" and 𝒮="𝒮 ⊕⇘WR ⊖⇘AL",  OF this]
        have "share sb (𝒮' ⊕⇘WR ⊖⇘AL) a = share sb (𝒮 ⊕⇘WR ⊖⇘AL) a"
          by auto
        then show ?thesis
          by (clarsimp simp add: Writesb volatile)
      next
        case False
        note not_R = this
        show ?thesis
        proof (cases "a  L")
          case True        
          from not_R True have "(𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
            by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
          from  share_shared_eq [where 𝒮'="𝒮' ⊕⇘WR ⊖⇘AL" and 𝒮="𝒮 ⊕⇘WR ⊖⇘AL",  OF this]
          have "share sb (𝒮' ⊕⇘WR ⊖⇘AL) a = share sb (𝒮 ⊕⇘WR ⊖⇘AL) a"
            by auto
          then show ?thesis
            by (clarsimp simp add: Writesb volatile)
        next
          case False
          with not_R a_in have "a  all_shared sb  all_unshared sb"
            by auto
          from Cons.hyps [OF this]
          show ?thesis by (clarsimp simp add: Writesb volatile)
        qed
      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)
  next
    case (Ghostsb A L R W)
    from Cons.prems obtain a_in: "a  R  all_shared sb  L  all_unshared sb"
      by (clarsimp simp add: Ghostsb)
    show ?thesis
    proof (cases "a  R")
      case True
      from True have "(𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
        by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
      from  share_shared_eq [where 𝒮'="𝒮' ⊕⇘WR ⊖⇘AL" and 𝒮="𝒮 ⊕⇘WR ⊖⇘AL",  OF this]
      have "share sb (𝒮' ⊕⇘WR ⊖⇘AL) a = share sb (𝒮 ⊕⇘WR ⊖⇘AL) a"
        by auto
      then show ?thesis
       by (clarsimp simp add: Ghostsb)
    next
      case False
      note not_R = this
      show ?thesis
      proof (cases "a  L")
        case True        
        from not_R True have "(𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
          by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
        from  share_shared_eq [where 𝒮'="𝒮' ⊕⇘WR ⊖⇘AL" and 𝒮="𝒮 ⊕⇘WR ⊖⇘AL",  OF this]
        have "share sb (𝒮' ⊕⇘WR ⊖⇘AL) a = share sb (𝒮 ⊕⇘WR ⊖⇘AL) a"
          by auto
        then show ?thesis
          by (clarsimp simp add: Ghostsb)
      next
        case False
        with not_R a_in have "a  all_shared sb  all_unshared sb"
          by auto
        from Cons.hyps [OF this]
        show ?thesis by (clarsimp simp add: Ghostsb)
      qed
    qed
  qed
qed

lemma share_domain_changesX:
  "𝒮 𝒮' X. a  X. 𝒮' a = 𝒮 a 
   a  all_shared sb  all_unshared sb  X  share sb 𝒮' a = share sb 𝒮 a "
proof (induct sb)
  case Nil thus ?case by simp
next
  case (Cons x sb)
  then have shared_eq: "a  X. 𝒮' a = 𝒮 a"
    by auto
  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  R  all_shared sb  L  all_unshared sb  X"
        by (clarsimp simp add: Writesb True)
      show ?thesis
      proof (cases "a  R")
        case True
        from True have "(𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
          by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
        from  share_shared_eq [where 𝒮'="𝒮' ⊕⇘WR ⊖⇘AL" and 𝒮="𝒮 ⊕⇘WR ⊖⇘AL",  OF this]
        have "share sb (𝒮' ⊕⇘WR ⊖⇘AL) a = share sb (𝒮 ⊕⇘WR ⊖⇘AL) a"
          by auto
        then show ?thesis
          by (clarsimp simp add: Writesb volatile)
      next
        case False
        note not_R = this
        show ?thesis
        proof (cases "a  L")
          case True        
          from not_R True have "(𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
            by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
          from  share_shared_eq [where 𝒮'="𝒮' ⊕⇘WR ⊖⇘AL" and 𝒮="𝒮 ⊕⇘WR ⊖⇘AL",  OF this]
          have "share sb (𝒮' ⊕⇘WR ⊖⇘AL) a = share sb (𝒮 ⊕⇘WR ⊖⇘AL) a"
            by auto
          then show ?thesis
            by (clarsimp simp add: Writesb volatile)
        next
          case False
          from shared_eq have shared_eq': "a  X. (𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
            by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
          from False not_R a_in have "a  all_shared sb  all_unshared sb  X"
            by auto
          from Cons.hyps [OF shared_eq' this]
          show ?thesis by (clarsimp simp add: Writesb volatile)
        qed
      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)
  next
    case (Ghostsb A L R W)
    from Cons.prems obtain a_in: "a  R  all_shared sb  L  all_unshared sb  X"
      by (clarsimp simp add: Ghostsb)
    show ?thesis
    proof (cases "a  R")
      case True
      from True have "(𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
        by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
      from  share_shared_eq [where 𝒮'="𝒮' ⊕⇘WR ⊖⇘AL" and 𝒮="𝒮 ⊕⇘WR ⊖⇘AL",  OF this]
      have "share sb (𝒮' ⊕⇘WR ⊖⇘AL) a = share sb (𝒮 ⊕⇘WR ⊖⇘AL) a"
        by auto
      then show ?thesis
       by (clarsimp simp add: Ghostsb)
    next
      case False
      note not_R = this
      show ?thesis
      proof (cases "a  L")
        case True        
        from not_R True have "(𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
          by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
        from  share_shared_eq [where 𝒮'="𝒮' ⊕⇘WR ⊖⇘AL" and 𝒮="𝒮 ⊕⇘WR ⊖⇘AL",  OF this]
        have "share sb (𝒮' ⊕⇘WR ⊖⇘AL) a = share sb (𝒮 ⊕⇘WR ⊖⇘AL) a"
          by auto
        then show ?thesis
          by (clarsimp simp add: Ghostsb)
      next
        case False
        from shared_eq have shared_eq': "a  X. (𝒮' ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
          by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
        from False not_R a_in have "a  all_shared sb  all_unshared sb  X"
          by auto
        from Cons.hyps [OF shared_eq' this]
        show ?thesis by (clarsimp simp add: Ghostsb)
      qed
    qed
  qed
qed

lemma share_unchanged: 
  "𝒮. a  all_shared sb  all_unshared sb  all_acquired sb  share sb 𝒮 a = 𝒮 a "
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_R: "a  R" and a_L: "a  L"  and a_A: "a  A" 
        and a': "a  all_shared sb  all_unshared sb  all_acquired sb"
        by (clarsimp simp add: Writesb True)
      from Cons.hyps [OF a']
      have "share sb (𝒮 ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a" .
      moreover
      from a_R a_L a_A have "(𝒮 ⊕⇘WR ⊖⇘AL) a = 𝒮 a"
        by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
      ultimately
      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_R: "a  R" and a_L: "a  L"  and a_A: "a  A" 
      and a': "a  all_shared sb  all_unshared sb  all_acquired sb"
      by (clarsimp simp add: Ghostsb)
    from Cons.hyps [OF a']
    have "share sb (𝒮 ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a" .
    moreover
    from a_R a_L a_A have "(𝒮 ⊕⇘WR ⊖⇘AL) a = 𝒮 a"
      by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
    ultimately
    show ?thesis
      by (clarsimp simp add: Ghostsb)
  qed
qed

(* FIXME: duplication with share_commute *)
lemma share_augment_release_commute: 
assumes dist: "(R  L  A)  (all_shared sb  all_unshared sb  all_acquired sb) = {}"
shows "(share sb 𝒮 ⊕⇘WR ⊖⇘AL) = share sb (𝒮 ⊕⇘WR ⊖⇘AL)"
proof -
  from dist have shared_eq: "a  all_acquired sb. (𝒮 ⊕⇘WR ⊖⇘AL) a = 𝒮 a"
    by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)  
  {
    fix a
    assume a_in: "a  all_shared sb  all_unshared sb  all_acquired sb"
    from share_domain_changesX [OF shared_eq this]
    have "share sb (𝒮 ⊕⇘WR ⊖⇘AL) a = share sb 𝒮 a".
    also
    from dist a_in have "... = (share sb 𝒮 ⊕⇘WR ⊖⇘AL) a"
      by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)  
    finally have  "share sb (𝒮 ⊕⇘WR ⊖⇘AL) a = (share sb 𝒮 ⊕⇘WR ⊖⇘AL) a".
  }
  moreover
  {
    fix a
    assume a_notin: "a  all_shared sb  all_unshared sb  all_acquired sb"
    from share_unchanged [OF a_notin]
    have "share sb (𝒮 ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a".
    moreover
    from share_unchanged [OF a_notin]
    have "share sb 𝒮 a = 𝒮 a".
    hence "(share sb 𝒮 ⊕⇘WR ⊖⇘AL) a = (𝒮 ⊕⇘WR ⊖⇘AL) a"
      by (auto simp add: augment_shared_def restrict_shared_def split: option.splits)
    ultimately have "share sb (𝒮 ⊕⇘WR ⊖⇘AL) a = (share sb 𝒮 ⊕⇘WR ⊖⇘AL) a"
      by simp
  }
  ultimately show ?thesis
    apply -
    apply (rule ext)
    subgoal for x
    apply (case_tac "x  all_shared sb  all_unshared sb  all_acquired sb")
    apply auto
    done
    done
qed

lemma share_append_commute: 
  "ys 𝒮. (all_shared xs  all_unshared xs  all_acquired xs)  
             (all_shared ys  all_unshared ys  all_acquired ys) = {} 
 share xs (share ys 𝒮) = share ys (share xs 𝒮)"
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x xs)
  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 have 
        dist': "(all_shared xs  all_unshared xs  all_acquired xs)  
                (all_shared ys   all_unshared ys  all_acquired ys) = {} "
        apply (clarsimp  simp add: Writesb True)
        apply blast
        done
      from Cons.prems have
        dist: "(R  L  A)  (all_shared ys  all_unshared ys  all_acquired ys) = {}"
        apply (clarsimp  simp add: Writesb True)
        apply blast
        done
      from share_augment_release_commute [OF dist]
      have "(share ys 𝒮 ⊕⇘WR ⊖⇘AL) = share ys (𝒮 ⊕⇘WR ⊖⇘AL)".
      
      with Cons.hyps [OF dist']
      show ?thesis
        by (clarsimp simp add: Writesb True)
    next
      case False with Cons show ?thesis
        by (clarsimp simp add: Writesb False)
    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 have 
      dist': "(all_shared xs  all_unshared xs  all_acquired xs)  
                (all_shared ys   all_unshared ys  all_acquired ys) = {} "
      apply (clarsimp  simp add: Ghostsb)
      apply blast
      done
    from Cons.prems have
      dist: "(R  L  A)  (all_shared ys  all_unshared ys  all_acquired ys) = {}"
      apply (clarsimp  simp add: Ghostsb)
      apply blast
      done
    from share_augment_release_commute [OF dist]
    have "(share ys 𝒮 ⊕⇘WR ⊖⇘AL) = share ys (𝒮 ⊕⇘WR ⊖⇘AL)".
      
    with Cons.hyps [OF dist']
    show ?thesis
      by (clarsimp simp add: Ghostsb)
  qed
qed

lemma share_append_commute': 
  assumes dist: "(all_shared xs  all_unshared xs  all_acquired xs)  
             (all_shared ys  all_unshared ys  all_acquired ys) = {} "
  shows "share (ys@xs) 𝒮 = share (xs@ys) 𝒮"
proof -
  from share_append_commute [OF dist] share_append [of xs ys] share_append [of ys xs]
  show ?thesis
    by simp
qed

lemma share_all_until_volatile_write_share_commute:
shows " 𝒮 (sb'::'a memref list). ownership_distinct ts; sharing_consis 𝒮 ts; 
        i p is 𝒪  𝒟 θ (sb::'a memref list). i < length ts 
             ts!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 sb'  all_unshared sb'  all_acquired sb') = {} 
 
share_all_until_volatile_write ts (share sb' 𝒮) =
share sb' (share_all_until_volatile_write ts 𝒮)"
proof (induct ts)
  case Nil
  thus ?case by simp
next
  case (Cons t ts)
  obtain pt "ist" 𝒪t t 𝒟t θt sbt where
    t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
    by (cases t)

  let ?take = "(takeWhile (Not  is_volatile_Writesb) sbt)"
  have dist: "ownership_distinct (t#ts)" by fact
  then interpret ownership_distinct "t#ts" .
  have consis: "sharing_consis 𝒮 (t#ts)" by fact
  then interpret sharing_consis 𝒮 "t#ts" .

  have dist_prop: "i p is 𝒪  𝒟 θ sb. i < length (t#ts) 
             (t#ts)!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 sb'  all_unshared sb'  all_acquired sb') = {}" by fact
  from dist_prop [rule_format, of 0] t
  have dist_t: "(all_shared ?take  all_unshared ?take  all_acquired ?take)  
         (all_shared sb'  all_unshared sb'  all_acquired sb') = {}"
    apply clarsimp
    done
  from dist_prop have 
  dist_prop':"i p is 𝒪  𝒟 θ sb. i < length ts 
             ts!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 sb'  all_unshared sb'  all_acquired sb') = {}"
    apply (clarsimp)
    subgoal for i p "is" 𝒪 ℛ 𝒟 θ sb
    apply (drule_tac x="Suc i" in spec)
    apply clarsimp
    done
    done

  from ownership_distinct_tl [OF dist]
  have dist': "ownership_distinct ts".

  from sharing_consis_tl [OF consis]
  have consis': "sharing_consis 𝒮 ts".
  then
  interpret consis': sharing_consis 𝒮 "ts" .

  from sharing_consis [of 0, simplified, OF t]
  have "sharing_consistent 𝒮 𝒪t sbt" .

  from sharing_consistent_takeWhile [OF this]
  have consis_sb: "sharing_consistent 𝒮 𝒪t ?take".

  let ?𝒮' = "(share ?take 𝒮)"

  from freshly_shared_owned [OF consis_sb]
  have fresh_owned: "dom ?𝒮' - dom 𝒮  𝒪t".
  from unshared_all_unshared [OF consis_sb] unshared_acquired_or_owned [OF consis_sb]
  have unshared_acq_owned: "dom 𝒮 - dom ?𝒮'
                  all_acquired (takeWhile (Not  is_volatile_Writesb) sbt)  𝒪t"
    by simp


  have sep: 
    "i < length ts. let (_,_,_,sb',_,_,_) = ts!i in 
          all_acquired sb'  dom 𝒮 - dom ?𝒮' = {}  
          all_unshared sb'  dom ?𝒮' - dom 𝒮 = {}"
  proof -
    {
      fix i pi "isi" 𝒪i i 𝒟i acqi θi sbi
      assume i_bound: "i < length ts"
      assume ts_i: "ts ! i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
      have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}  
            all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
      proof -
	from ownership_distinct [of "0" "Suc i"] ts_i t i_bound
	have dist: "(𝒪t  all_acquired sbt)  (𝒪i  all_acquired sbi) = {}"
	  by force


	from dist unshared_acq_owned all_acquired_takeWhile [of "(Not  is_volatile_Writesb)" sbt]
	have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}"
	  by blast

	moreover
	
	from sharing_consis [of "Suc i"] ts_i i_bound
	have "sharing_consistent 𝒮 𝒪i sbi"
	  by force
	from unshared_acquired_or_owned [OF this]
	have "all_unshared sbi  all_acquired sbi  𝒪i".      
	with dist fresh_owned
	have "all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
	  by blast
      
	ultimately show ?thesis by simp
      qed
    }
    thus ?thesis
      by (fastforce simp add: Let_def)
  qed

  from consis'.sharing_consis_preservation [OF sep]
  have sharing_consis': "sharing_consis ?𝒮' ts".

  have "share_all_until_volatile_write ts (share ?take (share sb' 𝒮)) =
        share sb' (share_all_until_volatile_write ts (share ?take 𝒮))"
  proof -
    from share_append_commute [OF dist_t]
    have "(share ?take (share sb' 𝒮)) = (share sb' (share ?take 𝒮))" .
    then
    have "share_all_until_volatile_write ts (share ?take (share sb' 𝒮)) =
          share_all_until_volatile_write ts (share sb' (share ?take 𝒮))"
      by (simp)
    also
    from Cons.hyps [OF dist' sharing_consis' dist_prop']
    have "... = share sb' (share_all_until_volatile_write ts (share ?take 𝒮))".
    finally show ?thesis .
  qed
  then show ?case
    by (clarsimp simp add: t)
qed
  
  
(* FIXME: move up*)
lemma all_shared_takeWhile_subset: "all_shared (takeWhile P sb)  all_shared sb"
using all_shared_append [of "(takeWhile P sb)" "(dropWhile P sb)"]
  by auto
lemma all_shared_dropWhile_subset: "all_shared (dropWhile P sb)  all_shared sb"
using all_shared_append [of "(takeWhile P sb)" "(dropWhile P sb)"]
  by auto

lemma all_unshared_takeWhile_subset: "all_unshared (takeWhile P sb)  all_unshared sb"
using all_unshared_append [of "(takeWhile P sb)" "(dropWhile P sb)"]
  by auto
lemma all_unshared_dropWhile_subset: "all_unshared (dropWhile P sb)  all_unshared sb"
using all_unshared_append [of "(takeWhile P sb)" "(dropWhile P sb)"]
  by auto

lemma all_acquired_takeWhile_subset: "all_acquired (takeWhile P sb)  all_acquired sb"
using all_acquired_append [of "(takeWhile P sb)" "(dropWhile P sb)"]
  by auto
lemma all_acquired_dropWhile_subset: "all_acquired (dropWhile P sb)  all_acquired sb"
using all_acquired_append [of "(takeWhile P sb)" "(dropWhile P sb)"]
  by auto

lemma share_all_until_volatile_write_flush_commute:
assumes takeWhile_empty: "(takeWhile (Not  is_volatile_Writesb) sb) = []"
shows " 𝒮 R L W A i. ownership_distinct ts; sharing_consis 𝒮 ts; i < length ts;
        ts!i = (p,is,θ,sb,𝒟,𝒪,); 
        i p is 𝒪  𝒟 θ (sb::'a memref list). i < length ts 
             ts!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')) = {};
        j p is 𝒪  𝒟 θ (sb::'a memref list). j < length ts  ij 
             ts!j=(p,is,θ,sb,𝒟,𝒪,)  
                    (all_shared sb  all_unshared sb  all_acquired sb)  
                    (R  L  A) = {} 
 
share_all_until_volatile_write (ts[i :=(p',is',θ',sb',𝒟',𝒪',ℛ')]) (𝒮 ⊕⇘WR ⊖⇘AL) =
share (takeWhile (Not  is_volatile_Writesb) sb') (share_all_until_volatile_write ts 𝒮 ⊕⇘WR ⊖⇘AL)"
proof (induct ts)
  case Nil
  thus ?case by simp
next
  case (Cons t ts)
  obtain pt "ist" 𝒪t t 𝒟t θt sbt where
    t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
    by (cases t)

  let ?take = "(takeWhile (Not  is_volatile_Writesb) sbt)"
  let ?take_sb' = "(takeWhile (Not  is_volatile_Writesb) sb')"
  let ?drop = "(dropWhile (Not  is_volatile_Writesb) sbt)"
  have dist: "ownership_distinct (t#ts)" by fact
  then interpret ownership_distinct "t#ts" .
  have consis: "sharing_consis 𝒮 (t#ts)" by fact
  then interpret sharing_consis 𝒮 "t#ts" .
  have dist_prop: "i p is 𝒪  𝒟 θ sb. i < length (t#ts) 
             (t#ts)!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 ?take_sb'  all_unshared ?take_sb'  all_acquired ?take_sb') = {}" by fact
  from dist_prop [rule_format, of 0] t
  have dist_t: "(all_shared ?take  all_unshared ?take  all_acquired ?take)  
         (all_shared ?take_sb'  all_unshared ?take_sb'  all_acquired ?take_sb') = {}"
    apply clarsimp
    done
  from dist_prop have 
  dist_prop':"i p is 𝒪  𝒟 θ sb. i < length ts 
             ts!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 ?take_sb'  all_unshared ?take_sb'  all_acquired ?take_sb') = {}"
    apply (clarsimp)
    subgoal for i p "is" 𝒪 ℛ 𝒟 θ sb
    apply (drule_tac x="Suc i" in spec)
    apply clarsimp
    done
    done
  have dist_prop_R_L_A: "j p is 𝒪  𝒟 θ sb. j < length (t#ts)  i  j
             (t#ts)!j=(p,is,θ,sb,𝒟,𝒪,)  
                    (all_shared sb  all_unshared sb  all_acquired sb)  
                    (R  L  A) = {}" by fact

    
  from ownership_distinct_tl [OF dist]
  have dist': "ownership_distinct ts".

  from sharing_consis_tl [OF consis]
  have consis': "sharing_consis 𝒮 ts".
  then
  interpret consis': sharing_consis 𝒮 "ts" .

  from sharing_consis [of 0, simplified, OF t]
  have "sharing_consistent 𝒮 𝒪t sbt" .

  from sharing_consistent_takeWhile [OF this]
  have consis_sb: "sharing_consistent 𝒮 𝒪t (takeWhile (Not  is_volatile_Writesb) sbt)".

  have aargh: "(Not  is_volatile_Writesb) = (λa. ¬ is_volatile_Writesb a)"
    by (rule ext) auto


  show ?case
  proof (cases i)
    case 0

    with t Cons.prems have eqs: "pt=p" "ist=is" "𝒪t=𝒪" "t=" "θt=θ" "sbt=sb" "𝒟t=𝒟" 
      by auto

    let ?𝒮' = "𝒮 ⊕⇘WR ⊖⇘AL"

    from dist_prop_R_L_A 0 have 
      dist_prop_R_L_A':"i p is 𝒪  𝒟 θ sb. i < length ts 
             ts!i=(p,is,θ,sb,𝒟,𝒪,)  
                    (all_shared sb  all_unshared sb  all_acquired sb)  
                    (R  L  A) = {}"
      apply (clarsimp)
      subgoal for i1 p "is" 𝒪 ℛ 𝒟 θ sb
      apply (drule_tac x="Suc i1" in spec)
      apply clarsimp
      done 
      done
    then 
    have dist_prop_R_L_A'':"i p is 𝒪  𝒟 θ sb. i < length ts 
         ts!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))  
      (R  L  A) = {}"
      apply (clarsimp)
      subgoal for i p "is" 𝒪 ℛ 𝒟 θ sb
      apply (cut_tac sb=sb in all_shared_takeWhile_subset [where P="Not  is_volatile_Writesb"])
      apply (cut_tac sb=sb in all_unshared_takeWhile_subset [where P="Not  is_volatile_Writesb"])
      apply (cut_tac sb=sb in all_acquired_takeWhile_subset [where P="Not  is_volatile_Writesb" ])
      apply fastforce
      done
      done

    have  sep: "i<length ts.
      let (_, _, _, sb, _, _, _) = ts ! i
      in aall_acquired sb. ?𝒮' a = 𝒮 a"
    proof -
      {
        fix i pi "isi" 𝒪i i 𝒟i acqi θi sbi a
        assume i_bound: "i < length ts"
        assume ts_i: "ts ! i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
        assume a_in: "a  all_acquired sbi"
        have "?𝒮' a = 𝒮 a"
        proof -
          from dist_prop_R_L_A' [rule_format, OF i_bound ts_i] a_in
          show ?thesis
            by (auto simp add: augment_shared_def restrict_shared_def split: option.splits) 
        qed
      }
      thus ?thesis by auto
    qed
    from consis'.sharing_consis_shared_exchange [OF sep]
    have sharing_consis': "sharing_consis ?𝒮' ts".
    
    from share_all_until_volatile_write_share_commute [of ts "(𝒮 ⊕⇘WR ⊖⇘AL)" "(takeWhile (Not  is_volatile_Writesb) sb')", OF dist' sharing_consis' dist_prop']

    have "share_all_until_volatile_write ts (share ?take_sb' ?𝒮') =
          share ?take_sb' (share_all_until_volatile_write ts ?𝒮')" .

    moreover 

    from dist_prop_R_L_A''
    have "(share_all_until_volatile_write ts (𝒮 ⊕⇘WR ⊖⇘AL)) =
          (share_all_until_volatile_write ts 𝒮 ⊕⇘WR ⊖⇘AL)"

      apply -
      apply (rule  share_all_until_volatile_write_commute [OF dist' consis', of L A R W,symmetric])
      apply (clarsimp,blast)+
      done
    ultimately
    show ?thesis
      using takeWhile_empty
      by (clarsimp simp add: t 0  aargh eqs)
  next
    case (Suc k)
    from Cons.prems Suc obtain k_bound: "k < length ts" and ts_k: "ts!k =  (p, is,θ, sb, 𝒟, 𝒪,)"
      by auto

    let ?𝒮' = "(share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮)"
    from freshly_shared_owned [OF consis_sb]
    have fresh_owned: "dom ?𝒮' - dom 𝒮  𝒪t".
    from unshared_all_unshared [OF consis_sb] unshared_acquired_or_owned [OF consis_sb]
    have unshared_acq_owned: "dom 𝒮 - dom ?𝒮'
                  all_acquired (takeWhile (Not  is_volatile_Writesb) sbt)  𝒪t"
      by simp



    from freshly_shared_owned [OF consis_sb]
    have fresh_owned: "dom ?𝒮' - dom 𝒮  𝒪t".
    from unshared_all_unshared [OF consis_sb] unshared_acquired_or_owned [OF consis_sb]
    have unshared_acq_owned: "dom 𝒮 - dom ?𝒮'
                  all_acquired (takeWhile (Not  is_volatile_Writesb) sbt)  𝒪t"
      by simp


    have sep: 
      "i < length ts. let (_,_,_,sb',_,_,_) = ts!i in 
          all_acquired sb'  dom 𝒮 - dom ?𝒮' = {}  
          all_unshared sb'  dom ?𝒮' - dom 𝒮 = {}"
    proof -
      {
        fix i pi "isi" 𝒪i i 𝒟i acqi θi sbi
        assume i_bound: "i < length ts"
        assume ts_i: "ts ! i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
        have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}  
              all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
        proof -
	  from ownership_distinct [of "0" "Suc i"] ts_i t i_bound
	  have dist: "(𝒪t  all_acquired sbt)  (𝒪i  all_acquired sbi) = {}"
	    by force


	  from dist unshared_acq_owned all_acquired_takeWhile [of "(Not  is_volatile_Writesb)" sbt]
	  have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}"
	    by blast

	  moreover
	
	  from sharing_consis [of "Suc i"] ts_i i_bound
	  have "sharing_consistent 𝒮 𝒪i sbi"
	    by force
	  from unshared_acquired_or_owned [OF this]
	  have "all_unshared sbi  all_acquired sbi  𝒪i".      
	  with dist fresh_owned
	  have "all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
	    by blast
      
	  ultimately show ?thesis by simp
        qed
      }
      thus ?thesis
        by (fastforce simp add: Let_def)
    qed
    from consis'.sharing_consis_preservation [OF sep]
    have sharing_consis': "sharing_consis ?𝒮' ts".


    from dist_prop_R_L_A [rule_format, of 0] Suc t
    have dist_t_R_L_A: "(all_shared sbt  all_unshared sbt  all_acquired sbt)  
         (R  L  A) = {}"
      apply clarsimp
      done
    from dist_t_R_L_A 
    have "(R  L  A)  (all_shared ?take  all_unshared ?take  all_acquired ?take) = {}"
    using all_shared_append [of ?take ?drop] all_unshared_append [of ?take ?drop] all_acquired_append [of ?take ?drop]
      by auto

    from share_augment_release_commute [OF this]
    have "share ?take 𝒮 ⊕⇘WR ⊖⇘AL = share ?take (𝒮 ⊕⇘WR ⊖⇘AL)" .
    moreover
    
    from dist_prop_R_L_A Suc
    have "j p is 𝒪  𝒟 θ sb. j < length (ts)  k  j
       (ts)!j=(p,is,θ,sb,𝒟,𝒪,)  
                    (all_shared sb  all_unshared sb  all_acquired sb)  
                    (R  L  A) = {}" 
      apply (clarsimp)
      subgoal for j p "is" 𝒪 ℛ 𝒟 θ sb
      apply (drule_tac x="Suc j" in spec)
      apply clarsimp
      done
      done
    note Cons.hyps [OF dist' sharing_consis' k_bound ts_k dist_prop' this, of W]
    ultimately
    show ?thesis
      by (clarsimp simp add: t Suc )
  qed
qed


lemma share_all_until_volatile_write_Ghostsb_commute:
shows " 𝒮 i. ownership_distinct ts; sharing_consis 𝒮 ts; i < length ts;
        ts!i = (p,is,θ,Ghostsb A L R W#sb,𝒟,𝒪,); 
        j p is 𝒪  𝒟 θ sb. j < length ts  ij  ts!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) = {} 
 
share_all_until_volatile_write (ts[i :=(p',is',θ',sb,𝒟',𝒪',ℛ')]) (𝒮 ⊕⇘WR ⊖⇘AL) =
share_all_until_volatile_write ts 𝒮"
proof (induct ts)
  case Nil
  thus ?case by simp
next
  case (Cons t ts)
  obtain pt "ist" 𝒪t t 𝒟t θt sbt where
    t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
    by (cases t)
  have dist: "ownership_distinct (t#ts)" by fact
  then interpret ownership_distinct "t#ts" .
  have consis: "sharing_consis 𝒮 (t#ts)" by fact
  then interpret sharing_consis 𝒮 "t#ts" .
  have dist_prop:  "j p is 𝒪  𝒟 θ sb. j < length (t#ts)  ij  (t#ts)!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) = {}" by fact

  from ownership_distinct_tl [OF dist]
  have dist': "ownership_distinct ts".

  from sharing_consis_tl [OF consis]
  have consis': "sharing_consis 𝒮 ts".
  then
  interpret consis': sharing_consis 𝒮 "ts" .

  from sharing_consis [of 0, simplified, OF t]
  have "sharing_consistent 𝒮 𝒪t sbt" .

  from sharing_consistent_takeWhile [OF this]
  have consis_sb: "sharing_consistent 𝒮 𝒪t (takeWhile (Not  is_volatile_Writesb) sbt)".

  let ?𝒮' = "(share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮)"

  from freshly_shared_owned [OF consis_sb]
  have fresh_owned: "dom ?𝒮' - dom 𝒮  𝒪t".
  from unshared_all_unshared [OF consis_sb] unshared_acquired_or_owned [OF consis_sb]
  have unshared_acq_owned: "dom 𝒮 - dom ?𝒮'
                  all_acquired (takeWhile (Not  is_volatile_Writesb) sbt)  𝒪t"
    by simp


  have sep: 
    "i < length ts. let (_,_,_,sb',_,_,_) = ts!i in 
          all_acquired sb'  dom 𝒮 - dom ?𝒮' = {}  
          all_unshared sb'  dom ?𝒮' - dom 𝒮 = {}"
  proof -
    {
      fix i pi "isi" 𝒪i i 𝒟i θi sbi
      assume i_bound: "i < length ts"
      assume ts_i: "ts ! i = (pi,isi,θi,sbi,𝒟i,𝒪i,i)"
      have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}  
            all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
      proof -
	from ownership_distinct [of "0" "Suc i"] ts_i t i_bound
	have dist: "(𝒪t  all_acquired sbt)  (𝒪i  all_acquired sbi) = {}"
	  by force


	from dist unshared_acq_owned all_acquired_takeWhile [of "(Not  is_volatile_Writesb)" sbt]
	have "all_acquired sbi  dom 𝒮 - dom ?𝒮' = {}"
	  by blast

	moreover
	
	from sharing_consis [of "Suc i"] ts_i i_bound
	have "sharing_consistent 𝒮 𝒪i sbi"
	  by force
	from unshared_acquired_or_owned [OF this]
	have "all_unshared sbi  all_acquired sbi  𝒪i".      
	with dist fresh_owned
	have "all_unshared sbi  dom ?𝒮' - dom 𝒮 = {}"
	  by blast
      
	ultimately show ?thesis by simp
      qed
    }
    thus ?thesis
      by (fastforce simp add: Let_def)
  qed

  from consis'.sharing_consis_preservation [OF sep]
  have sharing_consis': "sharing_consis (share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮) ts".

  show ?case
  proof (cases i)
    case 0

    with t Cons.prems have eqs: "pt=p" "ist=is" "𝒪t=𝒪" "t=" "θt=θ" "sbt=Ghostsb A L R W#sb" "𝒟t=𝒟" 
      by auto

    show ?thesis
      by (clarsimp simp add: 0 t eqs)
  next
    case (Suc k)
    from Cons.prems Suc obtain k_bound: "k < length ts" and ts_k: "ts!k =  (p, is,θ, Ghostsb A L R W#sb, 𝒟, 𝒪,)"
      by auto

    from dist_prop Suc 
    have dist_prop':  "j p is 𝒪  𝒟 θ sb. j < length ts  kj  ts!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) = {}" 
      apply (clarsimp)
      subgoal for j p "is" 𝒪 ℛ 𝒟 θ sb
      apply (drule_tac x="Suc j" in spec)
      apply auto
      done
      done

    from Cons.hyps [OF dist' sharing_consis' k_bound ts_k dist_prop']
    have "share_all_until_volatile_write (ts[k := (p', is', θ', sb, 𝒟', 𝒪', ℛ')])
        (share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮 ⊕⇘WR ⊖⇘AL) =
     share_all_until_volatile_write ts
       (share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮)" .
    
    moreover
    from dist_prop [rule_format, of 0 pt "ist" θt sbt 𝒟t  𝒪t t ] t Suc
    have "(R  L  A)  (all_shared (takeWhile (Not  is_volatile_Writesb) sbt)  all_unshared (takeWhile (Not  is_volatile_Writesb) sbt)  all_acquired (takeWhile (Not  is_volatile_Writesb) sbt)) = {}"
      apply clarsimp
      apply blast
      done
    from share_augment_release_commute [OF this]
    have "share (takeWhile (Not  is_volatile_Writesb) sbt) 𝒮 ⊕⇘WR ⊖⇘AL =
      share (takeWhile (Not  is_volatile_Writesb) sbt) (𝒮 ⊕⇘WR ⊖⇘AL)".
    ultimately
      show ?thesis
      by (clarsimp simp add: Suc t)
  qed
qed

lemma share_all_until_volatile_write_update_sb:
assumes congr: "S. share (takeWhile (Not  is_volatile_Writesb) sb') S = share (takeWhile (Not  is_volatile_Writesb) sb) S"
shows  "𝒮 i. i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,) 
  
  share_all_until_volatile_write ts 𝒮 =
    share_all_until_volatile_write (ts[i := (p', is',θ', sb', 𝒟', 𝒪',ℛ')]) 𝒮"
proof (induct ts)
  case Nil
  thus ?case by simp
next
  case (Cons t ts)
  obtain pt "ist" 𝒪t t 𝒟t θt sbt where
    t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
    by (cases t)

  show ?case
  proof (cases i)
    case 0
    with t Cons.prems have eqs: "pt=p" "ist=is" "𝒪t=𝒪" "t=" "θt=θ" "sbt=sb" "𝒟t=𝒟" 
      by auto

    show ?thesis
      by (clarsimp simp add: 0 t eqs congr)
  next
    case (Suc k)
    from Cons.prems Suc obtain k_bound: "k < length ts" and ts_k: "ts!k =  (p, is,θ, sb, 𝒟, 𝒪, )"
      by auto
    from Cons.hyps [OF k_bound ts_k ]
    show ?thesis
      by (clarsimp simp add: t Suc)
  qed
qed

lemma share_all_until_volatile_write_append_Ghostsb':
assumes out_VWritesb: "outstanding_refs is_volatile_Writesb sb  {}"
assumes i_bound: "i < length ts"
assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)"
shows "share_all_until_volatile_write ts 𝒮 =
    share_all_until_volatile_write
     (ts[i := (p', is',θ', sb @ [Ghostsb A L R W], 𝒟', 𝒪',ℛ')]) 𝒮"
proof -
  from out_VWritesb 
  have "S. share (takeWhile (Not  is_volatile_Writesb) (sb @ [Ghostsb A L R W])) S = 
             share (takeWhile (Not  is_volatile_Writesb) sb) S"
    by (simp add: outstanding_vol_write_takeWhile_append)
  from share_all_until_volatile_write_update_sb [OF this i_bound ts_i]
  show ?thesis
    by simp
qed

lemma acquired_append_Progsb:
"S. (acquired pending_write (takeWhile (Not  is_volatile_Writesb) (sb @ [Progsb p1 p2 mis])) S) = 
       (acquired pending_write (takeWhile (Not  is_volatile_Writesb) sb) S) "
  by (induct sb) (auto split: memref.splits)

(* FIXME: move up *)
lemma outstanding_refs_non_empty_dropWhile: 
  "outstanding_refs P xs  {}  outstanding_refs P (dropWhile (Not  P) xs)  {}"
apply (induct xs)
apply simp
apply (simp split: if_split_asm)
done

lemma ex_not: "Ex Not"
  by blast



(*
lemma read_only_share_all_until_volatile_write:
  "⋀𝒮. read_only (share_all_until_volatile_write ts 𝒮) ⊆ read_only 𝒮"
proof (induct ts)
  case Nil thus ?case by simp
next
  case (Cons t ts)
  obtain pt "ist" 𝒪tt 𝒟t θt sbt where
    t: "t=(pt,istt,sbt,𝒟t,𝒪t,ℛt)"
    by (cases t)
  let ?take_sbt ="(takeWhile (Not ∘ is_volatile_Writesb) sbt)"
 
  from read_only_share_takeWhile 
  have "read_only (share ?take_sbt 𝒮) ⊆ read_only 𝒮".
  moreover
  from Cons.hyps 
  have "read_only (share_all_until_volatile_write ts (share ?take_sbt 𝒮)) ⊆ 
    read_only (share ?take_sbt 𝒮)".
  ultimately
  show ?case
    by (simp add: t)
qed
*)

(*       
lemma read_only_takeWhile_share_all_until_volatile_write:
  "⋀i 𝒮. ⟦i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪)⟧
   ⟹ read_only (share_all_until_volatile_write ts 𝒮)
       ⊆ read_only (share (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒮)"
proof (induct ts)
  case Nil thus ?case by simp
next
  case (Cons t ts)
  obtain pt "ist" 𝒪t 𝒟t θt sbt where
    t: "t=(pt,istt,sbt,𝒟t,𝒪t)"
    by (cases t)

  let ?take_sb = "(takeWhile (Not ∘ is_volatile_Writesb) sb)"
  let ?take_sbt ="(takeWhile (Not ∘ is_volatile_Writesb) sbt)"

  note i_bound = `i < length (t#ts)`
  note ts_i = `(t#ts)!i = (p,is,θ,sb,𝒟,𝒪)`
  show ?case
  proof (cases i)
    case 0
    from read_only_share_all_until_volatile_write
    have "read_only (share_all_until_volatile_write ts (share ?take_sb 𝒮))
           ⊆ read_only (share ?take_sb 𝒮)".
    with ts_i
    show ?thesis
      by (simp add: t 0 del: o_apply)
  next
    case (Suc k)
    from i_bound Suc have k_bound: "k < length ts"
      by auto
    from ts_i Suc have ts_k: "ts!k = (p,is,θ,sb,𝒟,𝒪)"
      by auto

    from Cons.hyps [OF k_bound ts_k]
    have "read_only (share_all_until_volatile_write ts (share ?take_sbt 𝒮)) ⊆ 
      read_only (share (takeWhile (Not ∘ is_volatile_Writesb) sb) 
           (share ?take_sbt 𝒮))".
    moreover
    from read_only_share_takeWhile
    have "read_only (share ?take_sbt 𝒮) ⊆ read_only 𝒮".
    from share_read_only_mono [OF this, of ?take_sb]
    have "read_only (share ?take_sb (share ?take_sbt 𝒮)) ⊆ 
      read_only (share ?take_sb 𝒮)".
    ultimately
    show ?thesis
      by (simp add: t del: o_apply)
  qed
qed
*)

(*
lemma read_only_takeWhile_dropWhile_share_all_until_volatile_write:
  assumes i_bound: "i < length ts" 
  assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪)"
  shows "read_only (share (dropWhile (Not ∘ is_volatile_Writesb) sb) 
           (share_all_until_volatile_write ts 𝒮))
          ⊆ read_only (share sb 𝒮)"
proof -
  let ?take_sb = "(takeWhile (Not ∘ is_volatile_Writesb) sb)"
  let ?drop_sb = "(dropWhile (Not ∘ is_volatile_Writesb) sb)"

  from read_only_takeWhile_share_all_until_volatile_write [OF i_bound ts_i]
  have "read_only (share_all_until_volatile_write ts 𝒮)
        ⊆ read_only (share ?take_sb 𝒮)".
  from share_read_only_mono [OF this]
  have "read_only (share ?drop_sb (share_all_until_volatile_write ts 𝒮)) ⊆ 
        read_only (share ?drop_sb (share ?take_sb 𝒮))".
  with share_append [of ?take_sb ?drop_sb 𝒮]
  show ?thesis
    by simp
qed
*)


(* FIXME: unused ? *)
lemma (in computation) concurrent_step_append:
  assumes step: "(ts,m,𝒮)  (ts',m',𝒮')"
  shows "(xs@ts,m,𝒮)  (xs@ts',m',𝒮')"
using step
proof (cases)
  case (Program i p "is" θ sb 𝒟 𝒪  p' is'   )
  then obtain
    i_bound: "i < length ts" and
    ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)" and
    prog_step:   "θp p (p',is')" and
    ts': "ts'=ts[i:=(p',is@is',θ,record p p' is' sb,𝒟,𝒪,)]" and
    𝒮': "𝒮'=𝒮" and
    m': "m'=m"
    by auto
  
  from i_bound have i_bound': "i + length xs < length (xs@ts)"
    by auto

  from ts_i i_bound have ts_i': "(xs@ts)!(i + length xs) = (p,is,θ,sb,𝒟,𝒪,)"
    by (auto simp add: nth_append)
	
  from concurrent_step.Program [OF i_bound' ts_i' prog_step, of m 𝒮 ] ts' i_bound
  show ?thesis    
    by (auto simp add: ts' list_update_append 𝒮' m')
next
  case (Memop i p "is" θ sb  𝒟 𝒪  is' θ' sb' 𝒟' 𝒪' ℛ' )
  then obtain
    i_bound: "i < length ts" and
    ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)" and
    memop_step: "(is,θ,sb,m,𝒟,𝒪,,𝒮) m (is',θ',sb',m',𝒟',𝒪',ℛ',𝒮')" and
    ts': "ts'=ts[i:=(p,is',θ',sb',𝒟',𝒪',ℛ')]"
    by auto

  from i_bound have i_bound': "i + length xs < length (xs@ts)"
    by auto

  from ts_i i_bound have ts_i': "(xs@ts)!(i + length xs) = (p,is,θ,sb,𝒟,𝒪,)"
    by (auto simp add: nth_append)
  
  from concurrent_step.Memop [OF i_bound' ts_i' memop_step] ts' i_bound
  show ?thesis
    by (auto simp add: ts' list_update_append)
next
  case (StoreBuffer i p "is" θ sb 𝒟 𝒪  sb' 𝒪' ℛ')
  then obtain
    i_bound: "i < length ts" and
    ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,)" and
    sb_step: "(m,sb,𝒪,,𝒮) sb (m',sb',𝒪',ℛ',𝒮')" and
    ts': "ts'=ts[i:=(p,is,θ,sb',𝒟,𝒪',ℛ')]"
    by auto
  from i_bound have i_bound': "i + length xs < length (xs@ts)"
    by auto

  from ts_i i_bound have ts_i': "(xs@ts)!(i + length xs) = (p,is,θ,sb,𝒟,𝒪,)"
    by (auto simp add: nth_append)
  
  from concurrent_step.StoreBuffer [OF i_bound' ts_i' sb_step] ts' i_bound
  show ?thesis
    by (auto simp add: ts' list_update_append)
qed

primrec weak_sharing_consistent:: "owns  'a memref list   bool"
where
"weak_sharing_consistent 𝒪 [] = True"
| "weak_sharing_consistent 𝒪 (r#rs) =
   (case r of
     Writesb volatile _ _ _ A L R W  
      (if volatile then L  A  A  R = {}  R  𝒪  
                       weak_sharing_consistent (𝒪  A - R) rs
      else weak_sharing_consistent 𝒪 rs)  
   | Ghostsb A L R W  L  A  A  R = {}  R  𝒪  weak_sharing_consistent (𝒪  A - R) rs
   | _  weak_sharing_consistent 𝒪 rs)"

lemma sharing_consistent_weak_sharing_consistent:
  "𝒮 𝒪. sharing_consistent 𝒮 𝒪 sb  weak_sharing_consistent 𝒪 sb"
apply (induct sb)
apply (auto split: memref.splits)
done

lemma weak_sharing_consistent_append: 
"𝒪. weak_sharing_consistent 𝒪 (xs @ ys) =
  (weak_sharing_consistent 𝒪 xs  weak_sharing_consistent (acquired True xs 𝒪) ys)"
apply (induct xs)
apply (auto split: memref.splits)
done

lemma read_only_share_unowned: "𝒪 𝒮.
  weak_sharing_consistent 𝒪 sb; a  𝒪  all_acquired sb; a  read_only (share sb 𝒮) 
   a  read_only 𝒮"
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.hyps [where 𝒮="(𝒮 ⊕⇘WR ⊖⇘AL)" and 𝒪="(𝒪  A - R)"] Cons.prems
      show ?thesis
	by (auto simp add: Writesb True in_read_only_restrict_conv in_read_only_augment_conv
	split: if_split_asm)
    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 [where 𝒮="(𝒮 ⊕⇘WR  ⊖⇘AL)" and 𝒪="(𝒪  A - R)"] Cons.prems show ?thesis 
      by (auto simp add: in_read_only_restrict_conv in_read_only_augment_conv split: if_split_asm)
  qed
qed


(*
lemma read_only_share_not_acquired_takeWhile: "⋀𝒮. 
  read_only (share (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒮) ⊆ 
  - all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sb)"
apply (induct sb)
apply simp
apply (case_tac a)
apply (auto dest: read_only_share_takeWhile_in simp add: in_read_only_restrict_conv)
done

*)
(*
lemma read_only_share_all_until_volatile_Writesb_not_acquired_takeWhile: 
"⋀𝒮. read_only (share_all_until_volatile_write ts 𝒮 ) ⊆ 
               (- ⋃(λ(_, _, _, sb, _, _).
                   all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sb)) ` set ts)"
proof (induct ts)
  case Nil thus ?case by simp
next
  case (Cons t ts)
  
  obtain pt "ist" 𝒪t 𝒟t θt sbt 
    where t: "t=(pt,istt,sbt,𝒟t,𝒪t)"
    by (cases t)

  let "?take" = "(takeWhile (Not ∘ is_volatile_Writesb) sbt)"

  have aargh: "(Not ∘ is_volatile_Writesb) = (λa. ¬ is_volatile_Writesb a)"
    by (rule ext) auto


  from read_only_share_all_until_volatile_write
  have "read_only (share_all_until_volatile_write ts (share ?take 𝒮)) ⊆
        read_only (share ?take 𝒮)".
  also
  from read_only_share_not_acquired_takeWhile 
  have "read_only (share ?take 𝒮) ⊆ - all_acquired ?take"
    by blast
  finally have "read_only (share_all_until_volatile_write ts (share ?take 𝒮)) ⊆ - all_acquired ?take".

  moreover

  from Cons.hyps [of "share ?take 𝒮"]
  have hyp: "read_only (share_all_until_volatile_write ts (share ?take 𝒮)) ⊆ 
              (- ⋃(λ(_, _, _, sb, _, _).
                   all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sb)) ` set ts)".

  ultimately
  show ?case
    by (force simp add: t aargh)
qed
*)

(*
lemma read_only_share_all_until_volatile_Writesb_not_acquired_takeWhile_in:
assumes a_in: "a ∈ read_only (share_all_until_volatile_write ts 𝒮 )"
assumes i_bound: "i < length ts"
assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪)"
shows "a ∉ all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sb)"
proof -
  from nth_mem [OF i_bound] ts_i read_only_share_all_until_volatile_Writesb_not_acquired_takeWhile [of ts 𝒮] a_in
  show ?thesis
    by fastforce
qed
*)

lemma share_read_only_mono_in: 
  assumes a_in: "a  read_only (share sb 𝒮)"
  assumes ss: "read_only 𝒮  read_only 𝒮'"
  shows "a  read_only (share sb 𝒮')"
using share_read_only_mono [OF ss] a_in
by auto


lemma read_only_unacquired_share:
  "S 𝒪. 𝒪  read_only S = {}; weak_sharing_consistent 𝒪 sb; a  read_only S; 
  a  all_acquired 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_A: "a  A" and a_unacq: "a  all_acquired 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"
	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 a_ro a_A 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_unacq]
      show ?thesis
	by (clarsimp simp add: Writesb True)
    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_A: "a  A" and a_unacq: "a  all_acquired 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"
      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 a_ro a_A 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_unacq]
    show ?thesis
      by (clarsimp simp add: Ghostsb)
  qed
qed



lemma read_only_share_unacquired: " 𝒪 S. 𝒪  read_only S = {}  weak_sharing_consistent 𝒪 sb  
 a  read_only (share sb S)  a  acquired True 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 False
      with Cons Writesb show ?thesis by auto
    next
      case True 
      note volatile=this
      from Cons.prems
      obtain a_ro: "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 volatile)

      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_ro]
      show ?thesis
        by (auto simp add: Writesb volatile) 
    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_ro: "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_ro]
    show ?thesis
      by (auto simp add: Ghostsb) 
  qed
qed


lemma read_only_share_all_acquired_in: 
  "S 𝒪. 𝒪  read_only S = {}; weak_sharing_consistent 𝒪 sb; a  read_only (share sb S) 
   a  read_only (share sb Map.empty)  (a  read_only S  a  all_acquired 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
      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  all_acquired sb".

      have "a  read_only (share sb (Map.empty ⊕⇘WR ⊖⇘AL))  (a  read_only S  a  A  a  all_acquired 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_unacq: "a  all_acquired sb"
	  have ?thesis
	  proof (cases "a  read_only S")
	    case True
	    with a_ro obtain "a  A"
	      by (auto simp add: in_read_only_convs)
	    with True a_unacq 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_unacq]
	    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  all_acquired sb".

    have "a  read_only (share sb (Map.empty ⊕⇘WR ⊖⇘AL))  (a  read_only S  a  A  a  all_acquired 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_unacq: "a  all_acquired sb"
	have ?thesis
        proof (cases "a  read_only S")
	  case True
	  with a_ro obtain "a  A"
	    by (auto simp add: in_read_only_convs)
	  with True a_unacq 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_unacq]
	  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 weak_sharing_consistent_preserves_distinct:
  "𝒪 𝒮. weak_sharing_consistent 𝒪 sb  𝒪  read_only 𝒮 = {} 
           acquired True sb 𝒪  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 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"
	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']
      show ?thesis
	by (auto simp add: Writesb True)
    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"
      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']
    show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed


locale weak_sharing_consis =
fixes ts::"('p,'p store_buffer,bool,owns,rels) thread_config list"
assumes weak_sharing_consis:
  "i p is 𝒪  𝒟 θ sb. 
   i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)  
   
   weak_sharing_consistent 𝒪 sb"

sublocale sharing_consis  weak_sharing_consis
proof
  fix i p "is" 𝒪  𝒟 θ sb
  assume i_bound: "i < length ts" 
  assume ts_i: "ts ! i = (p, is, θ, sb, 𝒟, 𝒪,)"
  from sharing_consistent_weak_sharing_consistent [OF sharing_consis [OF i_bound ts_i]]
  show "weak_sharing_consistent 𝒪 sb".
qed

    
lemma weak_sharing_consis_tl: "weak_sharing_consis (t#ts)  weak_sharing_consis ts"
apply (unfold weak_sharing_consis_def)
apply force
done


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  all_acquired (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  all_acquired (takeWhile (Not  is_volatile_Writesb) sb))" and
    unacq_sb: "a  all_acquired (takeWhile (Not  is_volatile_Writesb) sb)" 
    by (force simp add: t aargh)


  
  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 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 read_only_share_unowned_in: 
"weak_sharing_consistent 𝒪 sb; a  read_only (share sb 𝒮)
 a  read_only 𝒮  𝒪  all_acquired sb" 
using read_only_share_unowned [of 𝒪 sb]
by auto


lemma read_only_shared_all_until_volatile_write_subset:
"𝒮. ownership_distinct ts;  
           weak_sharing_consis ts 
read_only (share_all_until_volatile_write ts 𝒮)  
  read_only 𝒮  (((λ(_, _, _, sb, _, 𝒪,_). 𝒪  all_acquired (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 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 consis: "weak_sharing_consis (t#ts)" by fact
  then interpret weak_sharing_consis "t#ts" .

  have aargh: "(Not  is_volatile_Writesb) = (λa. ¬ is_volatile_Writesb a)"
    by (rule ext) auto

  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

  {
    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_acquired (takeWhile (Not  is_volatile_Writesb) sb)) ` set ts))"
    have "a  𝒪  all_acquired (takeWhile (Not  is_volatile_Writesb) sb)"
    proof -
      from Cons.hyps [OF dist' consis', 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_unowned_in [OF consis_take this] a_notin_shared
      show ?thesis by auto
    qed
  }
      
  then show ?case
    by (auto simp add: t aargh)
qed
(*
lemma weak_sharing_consistent_preserves_distinct_share_all_until_volatile_write:
  assumes dist: "ownership_distinct ts" 
  assumes ro: "read_only_unowned 𝒮 ts"  
  assumes consis: "weak_sharing_consis ts" 
  assumes i_bound: "i < length ts" 
  assumes ts_i: "ts!i = (p,is,θ,sb,𝒟,𝒪,ℛ)"
  shows "acquired True (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒪 ∩ 
     read_only (share_all_until_volatile_write ts 𝒮) = {}"
proof -
  {
    fix a
    assume a_in: "a ∈ acquired True (takeWhile (Not ∘ is_volatile_Writesb) sb) 𝒪"
    assume a_in_share: "a ∈  read_only (share_all_until_volatile_write ts 𝒮)"
    have False
    proof -
      from read_only_shared_all_until_volatile_write_subset [OF dist consis, of 𝒮] a_in_share
      have "a ∈ read_only 𝒮 ∪ 
                  (⋃(λ(_, _, _, sb, _, 𝒪,_). 𝒪 ∪ all_acquired (takeWhile (Not ∘ is_volatile_Writesb) sb)) ` set ts)"
        by fastforce
      moreover
      from acquired_all_acquired [of True "(takeWhile (Not ∘ is_volatile_Writesb) sb)" 𝒪] a_in
      have "a ∈ 𝒪 ∪ 
      find_theorems acquired all_acquired
      from a_in 
*)
lemma weak_sharing_consistent_preserves_distinct_share_all_until_volatile_write:
  "𝒮 i. ownership_distinct ts; read_only_unowned 𝒮 ts;weak_sharing_consis ts; 
 i < length ts; ts!i = (p,is,θ,sb,𝒟,𝒪,)
  acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪  
  read_only (share_all_until_volatile_write ts 𝒮) = {}"
proof (induct ts)
  case Nil thus ?case by simp
next
  case (Cons t ts)
  note read_only_unowned 𝒮 (t#ts)
  then interpret read_only_unowned 𝒮 "t#ts" .
  note i_bound = i < length (t # ts)
  note ith = (t # ts) ! i = (p,is,θ, sb, 𝒟, 𝒪,)

  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 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)"


  have aargh: "(Not  is_volatile_Writesb) = (λa. ¬ is_volatile_Writesb a)"
    by (rule ext) auto
  show ?case
  proof (cases i)
    case 0
    from read_only_unowned [of 0] ith 0
    have owns_ro: "𝒪  read_only 𝒮 = {}"
      by force
    from weak_sharing_consis [of 0] ith 0
    have "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
    from weak_sharing_consistent_preserves_distinct [OF this owns_ro]
    have dist_t: "acquired True ?take_sb 𝒪  read_only (share ?take_sb 𝒮) = {}".
    from read_only_shared_all_until_volatile_write_subset [OF dist' consis', of "(share ?take_sb 𝒮)"]
    have ro_rest: "read_only (share_all_until_volatile_write ts (share ?take_sb 𝒮))  
            read_only (share ?take_sb 𝒮)  
            (((λ(_, _, _, sb, _, 𝒪,_). 𝒪  all_acquired (takeWhile (Not  is_volatile_Writesb) sb)) ` set ts))"
      by auto
    {
      fix a
      assume a_in_sb: "a  acquired True ?take_sb 𝒪"
      assume a_in_ro: "a  read_only (share_all_until_volatile_write ts (share ?take_sb 𝒮))"
      have False
      proof -
        
        from Set.in_mono [rule_format, OF ro_rest a_in_ro] dist_t a_in_sb
        
        have "a  (((λ(_, _, _, sb, _, 𝒪,_). 𝒪  all_acquired (takeWhile (Not  is_volatile_Writesb) sb)) ` set ts))"
          by auto
        then obtain j pj "isj" θj sbj 𝒟j 𝒪j j
            where j_bound: "j < length ts" and ts_j: "ts!j = (pj,isj,θj,sbj,𝒟j,𝒪j,j)"
            and a_in_j: "a  𝒪j  all_acquired (takeWhile (Not  is_volatile_Writesb) sbj)"
          by (fastforce simp add: in_set_conv_nth)
        from ownership_distinct [of 0 "Suc j"] ith ts_j j_bound 0
        have dist: "(𝒪  all_acquired sb)  (𝒪j  all_acquired sbj) = {}"
          by fastforce
        moreover
        from acquired_all_acquired [of True ?take_sb 𝒪] a_in_sb all_acquired_append [of ?take_sb ?drop_sb]
        have "a  𝒪  all_acquired sb"
          by auto
        with a_in_j all_acquired_append [of "(takeWhile (Not  is_volatile_Writesb) sbj)" "(dropWhile (Not  is_volatile_Writesb) sbj)"] 
          dist
        have False by fastforce
        thus ?thesis ..
     qed
   }
   then show ?thesis
   using 0 ith
     by (auto simp add: aargh)
  next
    case (Suc k)
    from i_bound Suc have k_bound: "k < length ts"
      by auto
    from ith Suc have kth: "ts!k = (p, is, θ, sb, 𝒟, 𝒪, )"
      by auto

    obtain pt "ist" 𝒪t t 𝒟t θt sbt 
      where t: "t=(pt,ist,θt,sbt,𝒟t,𝒪t,t)"
      by (cases t)

    let ?take_sbt = "(takeWhile (Not  is_volatile_Writesb) sbt)"
    let ?drop_sbt = "(dropWhile (Not  is_volatile_Writesb) sbt)"

    have ro_unowned': "read_only_unowned (share ?take_sbt 𝒮) 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_sbt 𝒮) = {}"
      proof -
	from read_only_unowned [of "Suc j"] j_bound jth
	have dist: "𝒪j  read_only 𝒮 = {}" by force
        
        from weak_sharing_consis [of 0] t
        have "weak_sharing_consistent 𝒪t sbt" 
          by fastforce
        with weak_sharing_consistent_append [of 𝒪t ?take_sbt ?drop_sbt]
        have consis_t: "weak_sharing_consistent 𝒪t ?take_sbt" 
          by auto
        {
          fix a
          assume a_in_j: "a  𝒪j"
          assume a_ro: "a  read_only (share ?take_sbt 𝒮)"
          have False
          proof -
            from a_in_j ownership_distinct [of 0 "Suc j"] j_bound t jth
            have "(𝒪t  all_acquired sbt)  (𝒪j  all_acquired sbj) = {}"
              by fastforce
            with a_in_j all_acquired_append [of ?take_sbt ?drop_sbt]
            have "a  (𝒪t  all_acquired ?take_sbt)"
              by fastforce
            from  read_only_share_unowned [OF consis_t this a_ro]
            have "a  read_only 𝒮" .
            with a_in_j dist 
            show False by auto
          qed
        }
        then
	show ?thesis
	  by auto
      qed
    qed

    from Cons.hyps [OF dist' ro_unowned' consis' k_bound kth]
    show ?thesis
      by (simp add: t)
  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  all_acquired (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]
  have "a  read_only (share sb Map.empty)  a  read_only 𝒮  a  all_acquired 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  all_acquired sb"
    with all_acquired_append [of ?take_sb ?drop_sb]
    obtain "a  all_acquired ?take_sb" and a_notin_drop: "a  all_acquired ?drop_sb"
      by auto
    with a_unacquired_others i_bound ts_i
    have a_unacq: "j < length ts. 
            (let (_,_,_,sbj,_,_,_) = ts!j in
            a  all_acquired (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]
    have ?thesis .
  }
  ultimately show ?thesis by blast
qed

lemma all_acquired_dropWhile_in: "x  all_acquired (dropWhile P sb)  x  all_acquired sb"	
  using all_acquired_append [of "takeWhile P sb" "dropWhile P sb"]
  by auto


lemma all_acquired_takeWhile_in: "x  all_acquired (takeWhile P sb)  x  all_acquired sb"	
  using all_acquired_append [of "takeWhile P sb" "dropWhile P sb"]
  by auto

lemmas all_acquired_takeWhile_dropWhile_in = all_acquired_takeWhile_in all_acquired_dropWhile_in



lemma split_in_read_only_reads: 
  "𝒪. a  read_only_reads 𝒪 xs  
  (t v ys zs. xs=ys @ Readsb False a t v # zs  a  acquired True ys 𝒪)"
proof (induct xs)
  case Nil thus ?case by simp
next
  case (Cons x xs)
  have a_in: "a  read_only_reads 𝒪 (x#xs)" by fact
  show ?case
  proof (cases x)
    case (Writesb volatile a' sop v A L R W)
    show ?thesis
    proof (cases volatile)
      case False
      from a_in have "a  read_only_reads 𝒪 xs"
	by (auto simp add: Writesb False)
      from Cons.hyps [OF this] obtain t v ys zs where
	xs: "xs=ys@Readsb False a t v # zs" and a_notin: "a  acquired True ys 𝒪"
	by auto
      with xs a_notin obtain "x#xs=(x#ys)@Readsb False a t v # zs" "a  acquired True (x#ys) 𝒪"
	by (simp add: Writesb False)
      then show ?thesis
	by blast
    next
      case True
      from a_in have "a  read_only_reads (𝒪  A - R) xs"
	by (auto simp add: Writesb True)
      from Cons.hyps [OF this] obtain t v ys zs where
	xs: "xs=ys@Readsb False a t v # zs" and a_notin: "a  acquired True ys (𝒪  A - R)"
	by auto
      with xs a_notin obtain "x#xs=(x#ys)@Readsb False a t v # zs" "a  acquired True (x#ys) 𝒪"
	by (simp add: Writesb True)
      then show ?thesis
	by blast
    qed
  next
    case (Readsb volatile a' t' v')
    show ?thesis
    proof (cases "¬ volatile  a  𝒪  a'=a")
      case True
      with Readsb show ?thesis
	by fastforce
    next
      case False
      with a_in have "a  read_only_reads 𝒪 xs"	
	by (auto simp add: Readsb  split: if_split_asm)
      from Cons.hyps [OF this] obtain t v ys zs where
	xs: "xs=ys@Readsb False a t v # zs" and a_notin: "a  acquired True ys 𝒪"
	by auto
      with xs a_notin obtain "x#xs=(x#ys)@Readsb False a t v # zs" "a  acquired True (x#ys) 𝒪"
	by (simp add: Readsb)
      then show ?thesis
	by blast
    qed
  next
    case Progsb
    with a_in have "a  read_only_reads 𝒪 xs"	
      by (auto)
    from Cons.hyps [OF this] obtain t v ys zs where
      xs: "xs=ys@Readsb False a t v # zs" and a_notin: "a  acquired True ys 𝒪"
      by auto
    with xs a_notin obtain "x#xs=(x#ys)@Readsb False a t v # zs" "a  acquired True (x#ys) 𝒪"
      by (simp add: Progsb)
    then show ?thesis
      by blast
  next
    case (Ghostsb A L R W)
    with a_in have "a  read_only_reads (𝒪  A - R) xs"	
      by (auto)
    from Cons.hyps [OF this] obtain t v ys zs where
      xs: "xs=ys@Readsb False a t v # zs" and a_notin: "a  acquired True ys (𝒪  A -R)"
      by auto
    with xs a_notin obtain "x#xs=(x#ys)@Readsb False a t v # zs" "a  acquired True (x#ys) 𝒪"
      by (simp add: Ghostsb)
    then show ?thesis
      by blast
  qed
qed


lemma insert_monoD: "W  W'  insert a W  insert a W'"
  by blast


primrec unforwarded_non_volatile_reads:: "'a memref list  addr set  addr set"
where
"unforwarded_non_volatile_reads [] W = {}"
| "unforwarded_non_volatile_reads (x#xs) W =
  (case x of
     Readsb volatile a _ _  (if a  W  ¬ volatile 
                             then insert a (unforwarded_non_volatile_reads xs W) 
                             else  (unforwarded_non_volatile_reads xs W))
  | Writesb _ a _ _ _ _ _ _  unforwarded_non_volatile_reads xs (insert a W)
  | _  unforwarded_non_volatile_reads xs W)"


lemma unforwarded_non_volatile_reads_non_volatile_Readsb:
  "W. unforwarded_non_volatile_reads sb W  outstanding_refs is_non_volatile_Readsb sb"
apply (induct sb)
apply (auto split: memref.splits if_split_asm)
done

lemma in_unforwarded_non_volatile_reads_non_volatile_Readsb:
  "a  unforwarded_non_volatile_reads sb W  a  outstanding_refs is_non_volatile_Readsb sb"
using unforwarded_non_volatile_reads_non_volatile_Readsb
by blast


lemma unforwarded_non_volatile_reads_antimono:
 "W W'. W  W'  unforwarded_non_volatile_reads xs W'  unforwarded_non_volatile_reads xs W"
apply (induct xs)
apply (auto split: memref.splits dest: insert_monoD)
done

lemma unforwarded_non_volatile_reads_antimono_in:
 "x  unforwarded_non_volatile_reads xs W'  W  W'
   x  unforwarded_non_volatile_reads xs W"
  using unforwarded_non_volatile_reads_antimono
  by blast

lemma unforwarded_non_volatile_reads_append: "W. unforwarded_non_volatile_reads (xs@ys) W =
 (unforwarded_non_volatile_reads xs W  
  unforwarded_non_volatile_reads ys (W  outstanding_refs is_Writesb xs))"
apply (induct xs)
apply  clarsimp
apply (auto split: memref.splits)
done

lemma reads_consistent_mem_eq_on_unforwarded_non_volatile_reads:
  assumes mem_eq: "a  A  W. m' a = m a"
  assumes subset: "unforwarded_non_volatile_reads sb W  A"
  assumes consis_m: "reads_consistent pending_write 𝒪 m sb"
  shows "reads_consistent pending_write 𝒪 m' sb"
using mem_eq subset consis_m 
proof (induct sb arbitrary: A W m' m pending_write 𝒪)
  case Nil thus ?case by simp
next
  case (Cons r sb)
  note mem_eq = a  A  W. m' a = m a
  note subset = unforwarded_non_volatile_reads (r#sb) W  A
  note consis_m = reads_consistent pending_write 𝒪 m (r#sb)

  show ?case
  proof (cases r)
    case (Writesb volatile a sop v A' L R W')
    from subset obtain
      subset': "unforwarded_non_volatile_reads sb (insert a W)  A"
      by (auto simp add: Writesb)
    from mem_eq
    have mem_eq': 
      "a'  (A  (insert a W)). (m'(a:=v)) a' = (m(a:=v)) a'"
      by (auto)
    show ?thesis
    proof (cases volatile)
      case True
      from consis_m obtain
	consis': "reads_consistent True (𝒪  A' - R) (m(a := v)) sb" and
        no_volatile_Readsb: "outstanding_refs is_volatile_Readsb sb = {}" 
	by (simp add: Writesb True)

      from Cons.hyps [OF mem_eq' subset' consis']
      have "reads_consistent True (𝒪  A' - R) (m'(a := v)) sb".
      with no_volatile_Readsb 
      show ?thesis
	by (simp add: Writesb True)
    next
      case False
      from consis_m obtain consis': "reads_consistent pending_write 𝒪 (m(a := v)) sb" 
	by (simp add: Writesb False)
      from Cons.hyps [OF mem_eq' subset' consis']
      have "reads_consistent pending_write 𝒪 (m'(a := v)) sb".
      then
      show ?thesis
	by (simp add: Writesb False)
    qed
  next
    case (Readsb volatile a t v)
    from mem_eq
    have mem_eq': 
      "a'  A  W. m' a' = m a'"
      by (auto)
    show ?thesis
    proof (cases volatile)
      case True
      note volatile=this
      from consis_m obtain	
	consis': "reads_consistent pending_write 𝒪 m sb"  
	by (simp add: Readsb True)

      show ?thesis
      proof (cases "a  W")
	case False
	from subset obtain
	  subset': "unforwarded_non_volatile_reads sb W  A"
	  using False
	  by (auto simp add: Readsb True split: if_split_asm)
	from Cons.hyps [OF mem_eq' subset' consis']
	show ?thesis
	  by (simp add: Readsb True)
      next
	case True
	from subset have
	  subset': "unforwarded_non_volatile_reads sb W  
	             insert a A "
	  using True
	  apply (auto simp add: Readsb volatile split: if_split_asm)
	  done
	from mem_eq True have mem_eq': "a'  (insert a A)  W. m' a' = m a'"
	  by auto
	from Cons.hyps [OF mem_eq' subset' consis']
	show ?thesis
	  by (simp add: Readsb volatile)
      qed
    next
      case False
      note non_vol = this
      from consis_m obtain	
	consis': "reads_consistent pending_write 𝒪 m sb"  and 
	v: "(pending_write  a  𝒪)  v=m a" 
	by (simp add: Readsb False)
      show ?thesis
      proof (cases "a  W")
	case True
	from mem_eq subset Readsb True non_vol have "m' a = m a"
	  by (auto simp add: False)
	from subset obtain
	  subset': "unforwarded_non_volatile_reads sb W  insert a A"
	  using False
	  by (auto simp add: Readsb non_vol split: if_split_asm)
	from mem_eq True have mem_eq': "a'  (insert a A)  W. m' a' = m a'"
	  by auto

	with Cons.hyps [OF mem_eq' subset' consis'] v
	show ?thesis
	  by (simp add: Readsb non_vol)
      next
	case False
	from mem_eq subset Readsb False non_vol have meq: "m' a = m a"
	  by (clarsimp )
	from subset obtain
	  subset': "unforwarded_non_volatile_reads sb W  A"
	  using non_vol False
	  by (auto simp add: Readsb non_vol split: if_split_asm)
	from mem_eq non_vol have mem_eq': "a'  A  W. m' a' = m a'"
	  by auto
	with Cons.hyps [OF mem_eq' subset' consis'] v meq
	show ?thesis
	  by (simp add: Readsb non_vol False)
      qed
    qed
  next
    case Progsb with Cons show ?thesis by auto
  next
    case Ghostsb with Cons show ?thesis by auto
  qed
qed


lemma reads_consistent_mem_eq_on_unforwarded_non_volatile_reads_drop:
  assumes mem_eq: "a  A  W. m' a = m a"
  assumes subset: "unforwarded_non_volatile_reads (dropWhile (Not  is_volatile_Writesb) sb) W  A"
  assumes subset_acq: "acquired_reads True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪  A"
  assumes consis_m: "reads_consistent False 𝒪 m sb"
  shows "reads_consistent False 𝒪 m' sb"
using mem_eq subset subset_acq consis_m 
proof (induct sb arbitrary: A W m' m 𝒪)
  case Nil thus ?case by simp
next
  case (Cons r sb)
  note mem_eq = a  A  W. m' a = m a
  note subset = unforwarded_non_volatile_reads 
    (dropWhile (Not  is_volatile_Writesb) (r#sb)) W  A
  note subset_acq = acquired_reads True (takeWhile (Not  is_volatile_Writesb)(r#sb)) 𝒪  A
  note consis_m = reads_consistent False 𝒪 m (r#sb)

  show ?case
  proof (cases r)
    case (Writesb volatile a sop v A' L R W')
    show ?thesis
    proof (cases volatile)
      case True
      from mem_eq
      have mem_eq': 
	"a'  (A  (insert a W)). (m'(a:=v)) a' = (m(a:=v)) a'"
	by (auto)

      from consis_m obtain
	consis': "reads_consistent True (𝒪  A' - R) (m(a := v)) sb" and
        no_volatile_Readsb: "outstanding_refs is_volatile_Readsb sb = {}" 
	by (simp add: Writesb True)

      from subset obtain "unforwarded_non_volatile_reads sb (insert a W)  A" 
	by (clarsimp simp add: Writesb True)

      from reads_consistent_mem_eq_on_unforwarded_non_volatile_reads [OF mem_eq' this consis']
      have "reads_consistent True (𝒪  A' - R) (m'(a := v)) sb".
      with no_volatile_Readsb 
      show ?thesis
	by (simp add: Writesb True)
    next
      case False
      from mem_eq
      have mem_eq': 
	"a'  (A  W). (m'(a:=v)) a' = (m(a:=v)) a'"
	by (auto)
      from subset obtain
	subset': "unforwarded_non_volatile_reads (dropWhile (Not  is_volatile_Writesb) sb) W  A"
	by (auto simp add: Writesb False)
      from subset_acq have 
	subset_acq': "acquired_reads True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪  A"
	by (auto simp add: Writesb False)
	
      from consis_m obtain consis': "reads_consistent False 𝒪 (m(a := v)) sb" 
	by (simp add: Writesb False)
      from Cons.hyps [OF mem_eq' subset' subset_acq' consis']
      have "reads_consistent False 𝒪 (m'(a := v)) sb".
      then
      show ?thesis
	by (simp add: Writesb False)
    qed
  next
    case (Readsb volatile a t v)
    from mem_eq
    have mem_eq': 
      "a'  A  W. m' a' = m a'"
      by (auto)
    from subset obtain
      subset': "unforwarded_non_volatile_reads (dropWhile (Not  is_volatile_Writesb) sb) W  A"
      by (clarsimp simp add: Readsb)
    from subset_acq obtain
      a_A: "¬ volatile  a  𝒪  a  A" and
      subset_acq': "acquired_reads True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪  A"
      by (auto simp add: Readsb split: if_split_asm)
    show ?thesis
    proof (cases volatile)
      case True
      note volatile=this
      from consis_m obtain	
	consis': "reads_consistent False 𝒪 m sb"  
	by (simp add: Readsb True)
      
      from Cons.hyps [OF mem_eq' subset' subset_acq'  consis']
      show ?thesis
	by (simp add: Readsb True)
    next
      case False
      note non_vol = this
      from consis_m obtain	
	consis': "reads_consistent False 𝒪 m sb"  and 
	v: "a  𝒪  v=m a" 
	by (simp add: Readsb False)

      from mem_eq a_A v have v': "a  𝒪  v=m' a"
	by (auto simp add: non_vol)
      from Cons.hyps [OF mem_eq' subset' subset_acq'  consis'] v'
      show ?thesis
	by (simp add: Readsb False)
    qed
  next
    case Progsb with Cons show ?thesis by auto
  next
    case Ghostsb with Cons show ?thesis by auto
  qed
qed





(* FIXME: unused?*)
lemma read_only_read_witness:"𝒮 𝒪.
  non_volatile_owned_or_read_only True 𝒮 𝒪 sb;
   a  read_only_reads 𝒪 sb
   
  xs ys t v. sb=xs@ Readsb False a t v # ys  a  read_only (share xs 𝒮)  a  read_only_reads 𝒪 xs"
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_ro: "a  read_only_reads (𝒪  A - R) sb" and
	nvo': "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb" 
	by (clarsimp simp add: Writesb True)

      from Cons.hyps [OF nvo' a_ro]
      obtain xs ys t v where
	"sb = xs @ Readsb False a t v # ys  a  read_only (share xs (𝒮 ⊕⇘WR ⊖⇘AL))  
	a  read_only_reads (𝒪  A - R) xs"
	by blast
      
      thus ?thesis
	apply -
	apply (rule_tac x="(x#xs)" in exI)
	apply (rule_tac x=ys in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v in exI)
	apply (clarsimp simp add: Writesb True)
	done
    next
      case False
      from Cons.prems obtain 
	a_ro: "a  read_only_reads 𝒪 sb" and
	nvo': "non_volatile_owned_or_read_only True 𝒮 𝒪 sb" 
	by (clarsimp simp add: Writesb False)

      from Cons.hyps [OF nvo' a_ro]
      obtain xs ys t v where
	"sb = xs @ Readsb False a t v # ys  a  read_only (share xs 𝒮)  a  read_only_reads 𝒪 xs"
	by blast
      
      thus ?thesis
	apply -
	apply (rule_tac x="(x#xs)" in exI)
	apply (rule_tac x=ys in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v in exI)
	apply (clarsimp simp add: Writesb False)
	done
    qed
  next
    case (Readsb volatile a' t v)
    show ?thesis
    proof (cases "a'=a  a  𝒪  ¬ volatile")
      case True
      with Cons.prems have "a  read_only 𝒮"
	by (simp add: Readsb)
      
      with True show ?thesis
	apply -
	apply (rule_tac x="[]" in exI)
	apply (rule_tac x="sb" in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v in exI)
	apply (clarsimp simp add: Readsb)
	done
    next
      case False
      with Cons.prems obtain 	
	a_ro: "a  read_only_reads 𝒪 sb" and
	nvo': "non_volatile_owned_or_read_only True 𝒮 𝒪 sb" 
	by (auto simp add: Readsb split: if_split_asm)
      from Cons.hyps [OF nvo' a_ro]
      obtain xs ys t' v' where
	"sb = xs @ Readsb False a t' v' # ys  a  read_only (share xs 𝒮)  a  read_only_reads 𝒪 xs"
	by blast
      
      with False show ?thesis
	apply -
	apply (rule_tac x="(x#xs)" in exI)
	apply (rule_tac x=ys in exI)
	apply (rule_tac x=t' in exI)
	apply (rule_tac x=v' in exI)
	apply (clarsimp simp add: Readsb )
	done
    qed
  next
    case Progsb
    from Cons.prems obtain 
      a_ro: "a  read_only_reads 𝒪 sb" and
      nvo': "non_volatile_owned_or_read_only True 𝒮 𝒪 sb" 
      by (clarsimp simp add: Progsb)
    
    from Cons.hyps [OF nvo' a_ro]
    obtain xs ys t v where
      "sb = xs @ Readsb False a t v # ys  a  read_only (share xs 𝒮)  a  read_only_reads 𝒪 xs"
      by blast
    
    thus ?thesis
      apply -
      apply (rule_tac x="(x#xs)" in exI)
      apply (rule_tac x=ys in exI)
      apply (rule_tac x=t in exI)
      apply (rule_tac x=v in exI)
      apply (clarsimp simp add: Progsb)
      done
  next
    case (Ghostsb A L R W)
    from Cons.prems obtain 
      a_ro: "a  read_only_reads (𝒪  A - R) sb" and
      nvo': "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb" 
      by (clarsimp simp add: Ghostsb)
    
    from Cons.hyps [OF nvo' a_ro]
    obtain xs ys t v where
      "sb = xs @ Readsb False a t v # ys  a  read_only (share xs (𝒮 ⊕⇘WR ⊖⇘AL))  a  read_only_reads (𝒪  A - R) xs"
      by blast
    
    thus ?thesis
      apply -
      apply (rule_tac x="(x#xs)" in exI)
      apply (rule_tac x=ys in exI)
      apply (rule_tac x=t in exI)
      apply (rule_tac x=v in exI)
      apply (clarsimp simp add: Ghostsb)
      done
  qed
qed

(* FIXME: unused?*)
lemma read_only_read_acquired_witness: "𝒮 𝒪.
  non_volatile_owned_or_read_only True 𝒮 𝒪 sb; sharing_consistent 𝒮 𝒪 sb;
  a  read_only 𝒮; a  𝒪; a  read_only_reads 𝒪 sb
   
  xs ys t v. sb=xs@ Readsb False a t v # ys  a  all_acquired xs  a  read_only (share xs 𝒮) 
              a  read_only_reads 𝒪 xs"
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_ro': "a  read_only_reads (𝒪  A - R) sb" 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"
	by (clarsimp simp add: Writesb True)

      from R_owns a_unowned
      have a_R: "a  R"
	by auto
      show ?thesis
      proof (cases "a  A")
	case True
	from read_only_read_witness [OF nvo' a_ro']
	obtain xs ys t v' where
	  sb: "sb = xs @ Readsb False a t v' # ys" and
	  ro: "a  read_only (share xs (𝒮 ⊕⇘WR ⊖⇘AL))" and
	  a_ro_xs: "a  read_only_reads (𝒪  A - R) xs"
	  by blast

	with True show ?thesis
	  apply -
	  apply (rule_tac x="x#xs" in exI)
	  apply (rule_tac x=ys in exI)
	  apply (rule_tac x=t in exI)
	  apply (rule_tac x=v' in exI)
	  apply (clarsimp simp add: Writesb volatile)
	  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_ro']
	obtain xs ys t v' where "sb = xs @ Readsb False a t v' # ys 
	  a  all_acquired xs  a  read_only (share xs (𝒮 ⊕⇘WR ⊖⇘AL))  
	  a  read_only_reads (𝒪  A - R) xs"
	  by blast

	then show ?thesis
	  apply -
	  apply (rule_tac x="x#xs" in exI)
	  apply (rule_tac x=ys in exI)
	  apply (rule_tac x=t in exI)
	  apply (rule_tac x=v' in exI)
	  apply (clarsimp simp add: Writesb volatile)
	  done
      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  read_only_reads 𝒪 sb" and
	"a'  𝒪" and
	nvo': "non_volatile_owned_or_read_only True 𝒮 𝒪 sb"
	by (clarsimp simp add: Writesb False)
      
      from Cons.hyps [OF nvo' consis' a_nro' a_unowned a_ro']
      obtain xs ys t v' where
	"sb = xs @ Readsb False a t v' # ys 
         a  all_acquired xs  a  read_only (share xs 𝒮)  a  read_only_reads 𝒪 xs"
	by blast

      then show ?thesis
	apply -
	apply (rule_tac x="x#xs" in exI)
	apply (rule_tac x=ys in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v' in exI)
	apply (clarsimp simp add: Writesb False)
	done
    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
      a_ro': "a  read_only_reads 𝒪 sb" and
      nvo': "non_volatile_owned_or_read_only True 𝒮 𝒪 sb"
      by (auto simp add: Readsb split: if_split_asm)

    from Cons.hyps [OF nvo' consis' a_nro' a_unowned a_ro']
    obtain xs ys t v' where
      "sb = xs @ Readsb False a t v' # ys 
      a  all_acquired xs  a  read_only (share xs 𝒮)  a  read_only_reads 𝒪 xs"
      by blast

    with Cons.prems show ?thesis
      apply -
      apply (rule_tac x="x#xs" in exI)
      apply (rule_tac x=ys in exI)
      apply (rule_tac x=t in exI)
      apply (rule_tac x=v' in exI)
      apply (clarsimp simp add: Readsb)
      done
  next
    case Progsb
    from Cons.prems
    obtain 	
      consis': "sharing_consistent 𝒮 𝒪 sb" and
      a_nro': "a  read_only 𝒮" and
      a_unowned: "a  𝒪" and
      a_ro': "a  read_only_reads 𝒪 sb" and
      nvo': "non_volatile_owned_or_read_only True 𝒮 𝒪 sb"
      by (auto simp add: Progsb)

    from Cons.hyps [OF nvo' consis' a_nro' a_unowned a_ro']
    obtain xs ys t v where
      "sb = xs @ Readsb False a t v # ys 
      a  all_acquired xs  a  read_only (share xs 𝒮)  a  read_only_reads 𝒪 xs"
      by blast

    then show ?thesis
      apply -
      apply (rule_tac x="x#xs" in exI)
      apply (rule_tac x=ys in exI)
      apply (rule_tac x=t in exI)
      apply (rule_tac x=v in exI)
      apply (clarsimp simp add: Progsb)
      done
  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_ro': "a  read_only_reads (𝒪  A - R) sb" 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"
      by (clarsimp simp add: Ghostsb)

    from R_owns a_unowned
    have a_R: "a  R"
      by auto
    show ?thesis
    proof (cases "a  A")
      case True
      from read_only_read_witness [OF nvo' a_ro']
      obtain xs ys t v' where
	sb: "sb = xs @ Readsb False a t v' # ys" and
	ro: "a  read_only (share xs (𝒮 ⊕⇘WR ⊖⇘AL))" and
	a_ro_xs: "a  read_only_reads (𝒪  A - R) xs"
        by blast

      with True show ?thesis
        apply -
        apply (rule_tac x="x#xs" in exI)
        apply (rule_tac x=ys in exI)
        apply (rule_tac x=t in exI)
        apply (rule_tac x=v' in exI)
        apply (clarsimp simp add: Ghostsb)
        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_ro']
      obtain xs ys t v' where "sb = xs @ Readsb False a t v' # ys 
	a  all_acquired xs  a  read_only (share xs (𝒮 ⊕⇘WR ⊖⇘AL))  
	a  read_only_reads (𝒪  A - R) xs"
        by blast

      then show ?thesis
        apply -
	apply (rule_tac x="x#xs" in exI)
	apply (rule_tac x=ys in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v' in exI)
	apply (clarsimp simp add: Ghostsb)
	done
    qed
  qed
qed
    


lemma unforwarded_not_written: "W. a  unforwarded_non_volatile_reads sb W  a  W"
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')
    from Cons.prems
    have "a  unforwarded_non_volatile_reads sb (insert a' W)"
      by (clarsimp simp add: Writesb )
    from Cons.hyps [OF this]
    have "a  insert a' W".
    then show ?thesis
      by blast
  next
    case (Readsb volatile a' t v)
    with Cons.hyps [of W] Cons.prems show ?thesis
      by (auto split: if_split_asm)
  next
    case Progsb
    with Cons.hyps [of W] Cons.prems show ?thesis
      by (auto split: if_split_asm)
  next
    case Ghostsb
    with Cons.hyps [of W] Cons.prems show ?thesis
      by (auto split: if_split_asm)
  qed
qed

   
lemma unforwarded_witness:"X.
  a  unforwarded_non_volatile_reads sb X
   
  xs ys t v. sb=xs@ Readsb False a t v # ys  a  outstanding_refs is_Writesb xs"
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_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'_a: "a'a"
	by auto

      from Cons.hyps [OF a_unforw]
      obtain xs ys t v where
	"sb = xs @ Readsb False a t v # ys 
	a  outstanding_refs is_Writesb xs"
	by blast
      
      thus ?thesis
	using a'_a
	apply -
	apply (rule_tac x="(x#xs)" in exI)
	apply (rule_tac x=ys in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v in exI)
	apply (clarsimp simp add: Writesb True)
	done
    next
      case False
      from Cons.prems obtain 
	a_unforw: "a  unforwarded_non_volatile_reads sb (insert a' X)" 
	by (clarsimp simp add: Writesb False)

      from unforwarded_not_written  [OF a_unforw]
      have a'_a: "a'a"
	by auto

      from Cons.hyps [OF a_unforw] 
      obtain xs ys t v where
	"sb = xs @ Readsb False a t v # ys 
	a  outstanding_refs is_Writesb xs"
	by blast
      
      thus ?thesis
	using a'_a
	apply -
	apply (rule_tac x="(x#xs)" in exI)
	apply (rule_tac x=ys in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v in exI)
	apply (clarsimp simp add: Writesb False)
	done
    qed
  next
    case (Readsb volatile a' t v)
    show ?thesis
    proof (cases "a'=a  a  X  ¬ volatile")
      case True
      
      with True show ?thesis
	apply -
	apply (rule_tac x="[]" in exI)
	apply (rule_tac x="sb" in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v in exI)
	apply (clarsimp simp add: Readsb)
	done
    next
      case False
      note not_ror = this
      with Cons.prems obtain a_unforw: "a  unforwarded_non_volatile_reads sb X" 
	by (auto simp add: Readsb split: if_split_asm)

      from Cons.hyps [OF a_unforw]
      obtain xs ys t v where
	"sb = xs @ Readsb False a t v # ys 
	a  outstanding_refs is_Writesb xs"
	by blast
      
      thus ?thesis
	apply -
	apply (rule_tac x="(x#xs)" in exI)
	apply (rule_tac x=ys in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v in exI)
	apply (clarsimp simp add: Readsb)
	done
    qed
  next
    case Progsb
    from Cons.prems obtain a_unforw: "a  unforwarded_non_volatile_reads sb X" 
      by (auto simp add: Progsb)

    from Cons.hyps [OF a_unforw]
    obtain xs ys t v where
      "sb = xs @ Readsb False a t v # ys 
      a  outstanding_refs is_Writesb xs"
      by blast
    
    thus ?thesis
      apply -
      apply (rule_tac x="(x#xs)" in exI)
      apply (rule_tac x=ys in exI)
      apply (rule_tac x=t in exI)
      apply (rule_tac x=v in exI)
      apply (clarsimp simp add: Progsb)
      done
  next
    case (Ghostsb A L R W)
    from Cons.prems obtain a_unforw: "a  unforwarded_non_volatile_reads sb X" 
      by (auto simp add: Ghostsb)

    from Cons.hyps [OF a_unforw]
    obtain xs ys t v where
      "sb = xs @ Readsb False a t v # ys 
      a  outstanding_refs is_Writesb xs"
      by blast
    
    thus ?thesis
      apply -
      apply (rule_tac x="(x#xs)" in exI)
      apply (rule_tac x=ys in exI)
      apply (rule_tac x=t in exI)
      apply (rule_tac x=v in exI)
      apply (clarsimp simp add: Ghostsb)
      done
  qed
qed


lemma read_only_read_acquired_unforwarded_witness: "𝒮 𝒪 X.
  non_volatile_owned_or_read_only True 𝒮 𝒪 sb; sharing_consistent 𝒮 𝒪 sb;
  a  read_only 𝒮; a  𝒪; a  read_only_reads 𝒪 sb;
  a  unforwarded_non_volatile_reads sb X 
   
  xs ys t v. sb=xs@ Readsb False a t v # ys  a  all_acquired xs 
              a  outstanding_refs is_Writesb xs"
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_ro': "a  read_only_reads (𝒪  A - R) sb" 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".
      from R_owns a_unowned
      have a_R: "a  R"
	by auto
      show ?thesis
      proof (cases "a  A")
	case True

	from unforwarded_witness [OF a_unforw]
	obtain xs ys t v' where
	  sb: "sb = xs @ Readsb False a t v' # ys" and
	  a_xs: "a  outstanding_refs is_Writesb xs"
	  by blast

	with True a_notin show ?thesis
	  apply -
	  apply (rule_tac x="x#xs" in exI)
	  apply (rule_tac x=ys in exI)
	  apply (rule_tac x=t in exI)
	  apply (rule_tac x=v' in exI)
	  apply (clarsimp simp add: Writesb volatile)
	  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_ro' a_unforw]
	obtain xs ys t v' where "sb = xs @ Readsb False a t v' # ys 
	  a  all_acquired xs  
	  a  outstanding_refs is_Writesb xs"
	  by blast

	with a_notin show ?thesis
	  apply -
	  apply (rule_tac x="x#xs" in exI)
	  apply (rule_tac x=ys in exI)
	  apply (rule_tac x=t in exI)
	  apply (rule_tac x=v' in exI)
	  apply (clarsimp simp add: Writesb volatile)
	  done
      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  read_only_reads 𝒪 sb" and
	"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_ro' a_unforw']
      obtain xs ys t v' where
	"sb = xs @ Readsb False a t v' # ys 
         a  all_acquired xs  a  outstanding_refs is_Writesb xs"
	by blast

      with a_notin show ?thesis
	apply -
	apply (rule_tac x="x#xs" in exI)
	apply (rule_tac x=ys in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v' in exI)
	apply (clarsimp simp add: Writesb False)
	done
    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
      a_ro': "a  read_only_reads 𝒪 sb" 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_ro' a_unforw]
    obtain xs ys t v' where
      "sb = xs @ Readsb False a t v' # ys 
      a  all_acquired xs  a  outstanding_refs is_Writesb xs"
      by blast

    with Cons.prems show ?thesis
      apply -
      apply (rule_tac x="x#xs" in exI)
      apply (rule_tac x=ys in exI)
      apply (rule_tac x=t in exI)
      apply (rule_tac x=v' in exI)
      apply (clarsimp simp add: Readsb)
      done
  next
    case Progsb
    from Cons.prems
    obtain 	
      consis': "sharing_consistent 𝒮 𝒪 sb" and
      a_nro': "a  read_only 𝒮" and
      a_unowned: "a  𝒪" and
      a_ro': "a  read_only_reads 𝒪 sb" 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_ro' a_unforw]
    obtain xs ys t v where
      "sb = xs @ Readsb False a t v # ys 
      a  all_acquired xs  a  outstanding_refs is_Writesb xs"
      by blast

    then show ?thesis
      apply -
      apply (rule_tac x="x#xs" in exI)
      apply (rule_tac x=ys in exI)
      apply (rule_tac x=t in exI)
      apply (rule_tac x=v in exI)
      apply (clarsimp simp add: Progsb)
      done
  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_ro': "a  read_only_reads (𝒪  A - R) sb" 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)

    from unforwarded_not_written [OF a_unforw]
    have a_notin: "a  X".
    from R_owns a_unowned
    have a_R: "a  R"
      by auto
    show ?thesis
    proof (cases "a  A")
      case True

      from unforwarded_witness [OF a_unforw]
      obtain xs ys t v' where
	sb: "sb = xs @ Readsb False a t v' # ys" and
	a_xs: "a  outstanding_refs is_Writesb xs"
        by blast

      with True a_notin show ?thesis
        apply -
        apply (rule_tac x="x#xs" in exI)
        apply (rule_tac x=ys in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v' in exI)
	apply (clarsimp simp add: Ghostsb)
	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_ro' a_unforw]
      obtain xs ys t v' where "sb = xs @ Readsb False a t v' # ys 
	a  all_acquired xs  
	a  outstanding_refs is_Writesb xs"
        by blast

      with a_notin show ?thesis
        apply -
	apply (rule_tac x="x#xs" in exI)
	apply (rule_tac x=ys in exI)
	apply (rule_tac x=t in exI)
	apply (rule_tac x=v' in exI)
	apply (clarsimp simp add: Ghostsb)
	done
    qed
  qed
qed


lemma takeWhile_prefix: "ys. takeWhile P xs @ ys = xs"
apply (induct xs)
apply auto
done

lemma unforwarded_empty_extend: 
  "W. x  unforwarded_non_volatile_reads sb {}   x  W  x  unforwarded_non_volatile_reads sb W"
apply (induct sb)
apply  clarsimp
subgoal for a sb W
apply (case_tac a) 
apply    clarsimp
apply    (frule unforwarded_not_written)
apply    (drule_tac W="{}" in unforwarded_non_volatile_reads_antimono_in)
apply    blast
apply   (auto split: if_split_asm)
done
done

lemma notin_unforwarded_empty: 
  "W. a  unforwarded_non_volatile_reads sb W  a  W  a  unforwarded_non_volatile_reads sb {}"
using unforwarded_empty_extend
by blast

lemma 
  assumes ro: "a  read_only 𝒮  a  read_only 𝒮'"
  assumes a_in: "a  read_only (𝒮 ⊕⇘WR) "
  shows "a  read_only (𝒮' ⊕⇘WR) "
  using ro a_in
  by (auto simp add: in_read_only_convs)

lemma 
  assumes ro: "a  read_only 𝒮  a  read_only 𝒮'"
  assumes a_in: "a  read_only (𝒮 ⊖⇘AL) "
  shows "a  read_only (𝒮' ⊖⇘AL) "
  using ro a_in
  by (auto simp add: in_read_only_convs)

lemma non_volatile_owned_or_read_only_read_only_reads_eq:
  "𝒮 𝒮' 𝒪 pending_write.
  non_volatile_owned_or_read_only pending_write 𝒮 𝒪 sb;
   a  read_only_reads 𝒪 sb. a  read_only 𝒮  a  read_only 𝒮' 
  
   non_volatile_owned_or_read_only pending_write 𝒮' 𝒪 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
      note volatile=this
      from Cons.prems obtain 
	nvo': "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb" and
	ro': "aread_only_reads (𝒪  A - R) sb. a  read_only 𝒮  a  read_only 𝒮'"
	by (clarsimp simp add: Writesb volatile)

      from ro'
      have ro'':"aread_only_reads (𝒪  A - R) sb.
        a  read_only (𝒮 ⊕⇘WR ⊖⇘AL)  a  read_only (𝒮' ⊕⇘WR ⊖⇘AL)"
	by (auto simp add: in_read_only_convs)
      from Cons.hyps [OF nvo' ro'']
      show ?thesis
	by (clarsimp simp add: Writesb volatile)
    next
      case False
      with Cons.hyps [of pending_write 𝒮 𝒪 𝒮'] Cons.prems show ?thesis
	by (auto simp add: Writesb)
    qed
  next
    case (Readsb volatile a t v)
    show ?thesis
    proof (cases volatile)
      case True
      with Cons.hyps [of pending_write 𝒮 𝒪 𝒮'] Cons.prems show ?thesis
	by (auto simp add: Readsb)
    next
      case False
      note non_vol = this
      show ?thesis
      proof (cases "a  𝒪")
	case True
	with Cons.hyps [of pending_write 𝒮 𝒪 𝒮'] Cons.prems show ?thesis
	  by (auto simp add: Readsb non_vol)
      next
	case False
	from Cons.prems Cons.hyps [of pending_write 𝒮 𝒪 𝒮'] show ?thesis 
	  by (clarsimp simp add: Readsb non_vol False)
      qed
    qed
  next
    case Progsb
    with Cons.hyps [of pending_write 𝒮 𝒪 𝒮'] Cons.prems show ?thesis
      by (auto)
  next
    case (Ghostsb A L R W)
    from Cons.hyps [of pending_write "(𝒮 ⊕⇘WR ⊖⇘AL)" "𝒪  A - R" "𝒮' ⊕⇘WR ⊖⇘AL"] Cons.prems
    show ?thesis
      by (auto simp add: Ghostsb in_read_only_convs)
  qed
qed


lemma non_volatile_owned_or_read_only_read_only_reads_eq':
  "𝒮 𝒮' 𝒪.
  non_volatile_owned_or_read_only False 𝒮 𝒪 sb;
   a  read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪) 
         (dropWhile (Not  is_volatile_Writesb) sb). a  read_only 𝒮  a  read_only 𝒮' 
  
   non_volatile_owned_or_read_only False 𝒮' 𝒪 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
      note volatile=this
      from Cons.prems obtain 
	nvo': "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) (𝒪  A - R) sb" and
	ro': "aread_only_reads (𝒪  A - R) sb. a  read_only 𝒮  a  read_only 𝒮'"
	by (clarsimp simp add: Writesb volatile)

      from ro'
      have ro'':"aread_only_reads (𝒪  A - R) sb.
        a  read_only (𝒮 ⊕⇘WR ⊖⇘AL)  a  read_only (𝒮' ⊕⇘WR ⊖⇘AL)"
	by (auto simp add: in_read_only_convs)
      from non_volatile_owned_or_read_only_read_only_reads_eq [OF nvo' ro'']
      show ?thesis
	by (clarsimp simp add: Writesb volatile)
    next
      case False
      with Cons.hyps [of 𝒮 𝒪 𝒮'] Cons.prems show ?thesis
	by (auto simp add: Writesb)
    qed
  next
    case (Readsb volatile a t v)
    show ?thesis
    proof (cases volatile)
      case True
      with Cons.hyps [of 𝒮 𝒪 𝒮'] Cons.prems show ?thesis
	by (auto simp add: Readsb)
    next
      case False
      note non_vol = this
      show ?thesis
      proof (cases "a  𝒪")
	case True
	with Cons.hyps [of 𝒮 𝒪 𝒮'] Cons.prems show ?thesis
	  by (auto simp add: Readsb non_vol)
      next
	case False
	from Cons.prems Cons.hyps [of 𝒮 𝒪 𝒮'] show ?thesis 
	  by (clarsimp simp add: Readsb non_vol False)
      qed
    qed
  next
    case Progsb
    with Cons.hyps [of 𝒮 𝒪 𝒮'] Cons.prems show ?thesis
      by (auto)
  next
    case (Ghostsb A L R W)
    from Cons.hyps [of "(𝒮 ⊕⇘WR ⊖⇘AL)" "𝒪  A - R" "𝒮' ⊕⇘WR ⊖⇘AL"] Cons.prems
    show ?thesis
      by (auto simp add: Ghostsb in_read_only_convs)
  qed
qed


lemma no_write_to_read_only_memory_read_only_reads_eq:
  "𝒮 𝒮'.
  no_write_to_read_only_memory 𝒮 sb;
   a  outstanding_refs is_Writesb sb. a  read_only 𝒮'  a  read_only 𝒮 
  
   no_write_to_read_only_memory 𝒮' 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
      note volatile=this
      from Cons.prems obtain 
	nvo': "no_write_to_read_only_memory (𝒮 ⊕⇘WR ⊖⇘AL) sb" and
	ro': "aoutstanding_refs is_Writesb sb. a  read_only 𝒮'  a  read_only 𝒮" and
	not_ro: "a  read_only 𝒮'"
	by (auto simp add: Writesb volatile)

      from ro'
      have ro'':"aoutstanding_refs is_Writesb sb.
        a  read_only (𝒮' ⊕⇘WR ⊖⇘AL)  a  read_only (𝒮 ⊕⇘WR ⊖⇘AL)"
	by (auto simp add: in_read_only_convs)
      from Cons.hyps [OF nvo' ro''] not_ro
      show ?thesis
	by (clarsimp simp add: Writesb volatile)
    next
      case False
      with Cons.hyps [of 𝒮 𝒮'] Cons.prems show ?thesis
	by (auto simp add: Writesb)
    qed
  next
    case (Readsb volatile a t v)
    with Cons.hyps [of 𝒮 𝒮'] Cons.prems show ?thesis
      by (auto simp add: Readsb)  
  next
    case Progsb
    with Cons.hyps [of 𝒮 𝒮'] Cons.prems show ?thesis
      by (auto)
  next
    case (Ghostsb A L R W)
    from Cons.hyps [of "(𝒮 ⊕⇘WR ⊖⇘AL)" "𝒮' ⊕⇘WR ⊖⇘AL"] Cons.prems
    show ?thesis
      by (auto simp add: Ghostsb in_read_only_convs)
  qed
qed


lemma reads_consistent_drop:
  "reads_consistent False 𝒪 m sb
   reads_consistent True  
      (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪) 
      (flush (takeWhile (Not  is_volatile_Writesb) sb) m)
      (dropWhile (Not  is_volatile_Writesb) sb)"
using reads_consistent_append [of False 𝒪 m "(takeWhile (Not  is_volatile_Writesb) sb)" 
  "(dropWhile (Not  is_volatile_Writesb) sb)"]
apply (cases "outstanding_refs is_volatile_Writesb sb = {}")
apply  (clarsimp simp add: outstanding_vol_write_take_drop_appends 
  takeWhile_not_vol_write_outstanding_refs dropWhile_not_vol_write_empty)
apply(clarsimp simp add: outstanding_vol_write_take_drop_appends 
  takeWhile_not_vol_write_outstanding_refs dropWhile_not_vol_write_empty )
apply (case_tac "(dropWhile (Not  is_volatile_Writesb) sb)")
apply  (fastforce simp add: outstanding_refs_conv)
apply (frule dropWhile_ConsD)
apply (clarsimp split: memref.splits)
done

(* FIXME: subsumes outstanding_refs_non_volatile_Readsb_all_acquired_drop *)
lemma outstanding_refs_non_volatile_Readsb_all_acquired_dropWhile': 
"m 𝒮 𝒪 pending_write. 
  reads_consistent pending_write 𝒪 m sb;non_volatile_owned_or_read_only pending_write 𝒮 𝒪 sb;  
a  outstanding_refs is_non_volatile_Readsb (dropWhile (Not  is_volatile_Writesb) sb)
 a  𝒪  a  all_acquired sb 
    a  read_only_reads (acquired True (takeWhile (Not  is_volatile_Writesb) sb) 𝒪) 
      (dropWhile (Not  is_volatile_Writesb) 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
      note volatile=this
      from Cons.prems obtain 
	non_vo: "non_volatile_owned_or_read_only True (𝒮 ⊕⇘WR ⊖⇘AL) 
	            (𝒪  A - R) sb" and
        out_vol: "outstanding_refs is_volatile_Readsb sb = {}" and
	out: "a  outstanding_refs is_non_volatile_Readsb sb"
	by (clarsimp simp add: Writesb True) 
      show ?thesis
      proof (cases "a  𝒪")
	case True
	show ?thesis
	by (clarsimp simp add: Writesb True volatile)
      next
	case False
	from outstanding_non_volatile_Readsb_acquired_or_read_only_reads [OF non_vo out]
	have a_in: "a  acquired_reads True sb (𝒪  A - R) 
                    a  read_only_reads (𝒪  A - R) sb"
	  by auto
	with acquired_reads_all_acquired [of True sb "(𝒪  A - R)"]
	show ?thesis 
	  by (auto simp add: Writesb volatile)
      qed

    next
      case False
      with Cons show ?thesis
	by (auto simp add: Writesb False)
    qed
  next
    case Readsb
    with Cons show ?thesis
      apply (clarsimp simp del: o_apply simp add: Readsb 
	acquired_takeWhile_non_volatile_Writesb split: if_split_asm)
      apply auto
      done
  next
    case Progsb
    with Cons show ?thesis
      by (auto simp add: Readsb)
  next
    case (Ghostsb A L R W)
    with Cons.hyps [of pending_write "𝒪  A - R" "m" "𝒮 ⊕⇘WR  ⊖⇘AL"]  read_only_reads_antimono [of 𝒪 "𝒪  A - R"]
      Cons.prems show ?thesis
      by (auto simp add: Ghostsb)
  qed
qed



end