clarified directory structure;
authorwenzelm
Thu, 31 Dec 2015 12:43:09 +0100
changeset 62008 cbedaddc9351
parent 62007 3f8b97ceedb2
child 62009 ecb5212d5885
clarified directory structure;
src/HOL/HOLCF/IOA/ABP/Abschannel.thy
src/HOL/HOLCF/IOA/ABP/Abschannel_finite.thy
src/HOL/HOLCF/IOA/ABP/Correctness.thy
src/HOL/HOLCF/IOA/ABP/Env.thy
src/HOL/HOLCF/IOA/ABP/Receiver.thy
src/HOL/HOLCF/IOA/ABP/Sender.thy
src/HOL/HOLCF/IOA/Abstraction.thy
src/HOL/HOLCF/IOA/Asig.thy
src/HOL/HOLCF/IOA/Automata.thy
src/HOL/HOLCF/IOA/CompoExecs.thy
src/HOL/HOLCF/IOA/CompoScheds.thy
src/HOL/HOLCF/IOA/CompoTraces.thy
src/HOL/HOLCF/IOA/Compositionality.thy
src/HOL/HOLCF/IOA/Deadlock.thy
src/HOL/HOLCF/IOA/IOA.thy
src/HOL/HOLCF/IOA/LiveIOA.thy
src/HOL/HOLCF/IOA/NTP/Abschannel.thy
src/HOL/HOLCF/IOA/NTP/Receiver.thy
src/HOL/HOLCF/IOA/NTP/Sender.thy
src/HOL/HOLCF/IOA/NTP/Spec.thy
src/HOL/HOLCF/IOA/Pred.thy
src/HOL/HOLCF/IOA/RefCorrectness.thy
src/HOL/HOLCF/IOA/RefMappings.thy
src/HOL/HOLCF/IOA/Seq.thy
src/HOL/HOLCF/IOA/Sequence.thy
src/HOL/HOLCF/IOA/ShortExecutions.thy
src/HOL/HOLCF/IOA/SimCorrectness.thy
src/HOL/HOLCF/IOA/Simulations.thy
src/HOL/HOLCF/IOA/Storage/Spec.thy
src/HOL/HOLCF/IOA/TL.thy
src/HOL/HOLCF/IOA/TLS.thy
src/HOL/HOLCF/IOA/Traces.thy
src/HOL/HOLCF/IOA/meta_theory/Abstraction.thy
src/HOL/HOLCF/IOA/meta_theory/Asig.thy
src/HOL/HOLCF/IOA/meta_theory/Automata.thy
src/HOL/HOLCF/IOA/meta_theory/CompoExecs.thy
src/HOL/HOLCF/IOA/meta_theory/CompoScheds.thy
src/HOL/HOLCF/IOA/meta_theory/CompoTraces.thy
src/HOL/HOLCF/IOA/meta_theory/Compositionality.thy
src/HOL/HOLCF/IOA/meta_theory/Deadlock.thy
src/HOL/HOLCF/IOA/meta_theory/IOA.thy
src/HOL/HOLCF/IOA/meta_theory/LiveIOA.thy
src/HOL/HOLCF/IOA/meta_theory/Pred.thy
src/HOL/HOLCF/IOA/meta_theory/RefCorrectness.thy
src/HOL/HOLCF/IOA/meta_theory/RefMappings.thy
src/HOL/HOLCF/IOA/meta_theory/Seq.thy
src/HOL/HOLCF/IOA/meta_theory/Sequence.thy
src/HOL/HOLCF/IOA/meta_theory/ShortExecutions.thy
src/HOL/HOLCF/IOA/meta_theory/SimCorrectness.thy
src/HOL/HOLCF/IOA/meta_theory/Simulations.thy
src/HOL/HOLCF/IOA/meta_theory/TL.thy
src/HOL/HOLCF/IOA/meta_theory/TLS.thy
src/HOL/HOLCF/IOA/meta_theory/Traces.thy
src/HOL/ROOT
--- a/src/HOL/HOLCF/IOA/ABP/Abschannel.thy	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Abschannel.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -5,7 +5,7 @@
 section \<open>The transmission channel\<close>
 
 theory Abschannel
-imports "~~/src/HOL/HOLCF/IOA/meta_theory/IOA" Action Lemmas
+imports "~~/src/HOL/HOLCF/IOA/IOA" Action Lemmas
 begin
 
 datatype 'a abs_action = S 'a | R 'a
--- a/src/HOL/HOLCF/IOA/ABP/Abschannel_finite.thy	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Abschannel_finite.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -5,7 +5,7 @@
 section \<open>The transmission channel -- finite version\<close>
 
 theory Abschannel_finite
-imports Abschannel "~~/src/HOL/HOLCF/IOA/meta_theory/IOA" Action Lemmas
+imports Abschannel "~~/src/HOL/HOLCF/IOA/IOA" Action Lemmas
 begin
 
 primrec reverse :: "'a list => 'a list"
--- a/src/HOL/HOLCF/IOA/ABP/Correctness.thy	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Correctness.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -5,7 +5,7 @@
 section \<open>The main correctness proof: System_fin implements System\<close>
 
 theory Correctness
-imports "~~/src/HOL/HOLCF/IOA/meta_theory/IOA" Env Impl Impl_finite
+imports "~~/src/HOL/HOLCF/IOA/IOA" Env Impl Impl_finite
 begin
 
 ML_file "Check.ML"
--- a/src/HOL/HOLCF/IOA/ABP/Env.thy	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Env.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -5,7 +5,7 @@
 section \<open>The environment\<close>
 
 theory Env
-imports "~~/src/HOL/HOLCF/IOA/meta_theory/IOA" Action
+imports "~~/src/HOL/HOLCF/IOA/IOA" Action
 begin
 
 type_synonym
--- a/src/HOL/HOLCF/IOA/ABP/Receiver.thy	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Receiver.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -5,7 +5,7 @@
 section \<open>The implementation: receiver\<close>
 
 theory Receiver
-imports "~~/src/HOL/HOLCF/IOA/meta_theory/IOA" Action Lemmas
+imports "~~/src/HOL/HOLCF/IOA/IOA" Action Lemmas
 begin
 
 type_synonym
--- a/src/HOL/HOLCF/IOA/ABP/Sender.thy	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/HOLCF/IOA/ABP/Sender.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -5,7 +5,7 @@
 section \<open>The implementation: sender\<close>
 
 theory Sender
-imports "~~/src/HOL/HOLCF/IOA/meta_theory/IOA" Action Lemmas
+imports "~~/src/HOL/HOLCF/IOA/IOA" Action Lemmas
 begin
 
 type_synonym
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/Abstraction.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,614 @@
+(*  Title:      HOL/HOLCF/IOA/Abstraction.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Abstraction Theory -- tailored for I/O automata\<close>
+
+theory Abstraction
+imports LiveIOA
+begin
+
+default_sort type
+
+definition
+  cex_abs :: "('s1 => 's2) => ('a,'s1)execution => ('a,'s2)execution" where
+  "cex_abs f ex = (f (fst ex), Map (%(a,t). (a,f t))$(snd ex))"
+definition
+  \<comment> \<open>equals cex_abs on Sequences -- after ex2seq application\<close>
+  cex_absSeq :: "('s1 => 's2) => ('a option,'s1)transition Seq => ('a option,'s2)transition Seq" where
+  "cex_absSeq f = (%s. Map (%(s,a,t). (f s,a,f t))$s)"
+
+definition
+  is_abstraction ::"[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool" where
+  "is_abstraction f C A =
+   ((!s:starts_of(C). f(s):starts_of(A)) &
+   (!s t a. reachable C s & s \<midarrow>a\<midarrow>C\<rightarrow> t
+            --> (f s) \<midarrow>a\<midarrow>A\<rightarrow> (f t)))"
+
+definition
+  weakeningIOA :: "('a,'s2)ioa => ('a,'s1)ioa => ('s1 => 's2) => bool" where
+  "weakeningIOA A C h = (!ex. ex : executions C --> cex_abs h ex : executions A)"
+definition
+  temp_strengthening :: "('a,'s2)ioa_temp => ('a,'s1)ioa_temp => ('s1 => 's2) => bool" where
+  "temp_strengthening Q P h = (!ex. (cex_abs h ex \<TTurnstile> Q) --> (ex \<TTurnstile> P))"
+definition
+  temp_weakening :: "('a,'s2)ioa_temp => ('a,'s1)ioa_temp => ('s1 => 's2) => bool" where
+  "temp_weakening Q P h = temp_strengthening (\<^bold>\<not> Q) (\<^bold>\<not> P) h"
+
+definition
+  state_strengthening :: "('a,'s2)step_pred => ('a,'s1)step_pred => ('s1 => 's2) => bool" where
+  "state_strengthening Q P h = (!s t a.  Q (h(s),a,h(t)) --> P (s,a,t))"
+definition
+  state_weakening :: "('a,'s2)step_pred => ('a,'s1)step_pred => ('s1 => 's2) => bool" where
+  "state_weakening Q P h = state_strengthening (\<^bold>\<not>Q) (\<^bold>\<not>P) h"
+
+definition
+  is_live_abstraction :: "('s1 => 's2) => ('a,'s1)live_ioa => ('a,'s2)live_ioa => bool" where
+  "is_live_abstraction h CL AM =
+     (is_abstraction h (fst CL) (fst AM) &
+      temp_weakening (snd AM) (snd CL) h)"
+
+
+axiomatization where
+(* thm about ex2seq which is not provable by induction as ex2seq is not continous *)
+ex2seq_abs_cex:
+  "ex2seq (cex_abs h ex) = cex_absSeq h (ex2seq ex)"
+
+axiomatization where
+(* analog to the proved thm strength_Box - proof skipped as trivial *)
+weak_Box:
+"temp_weakening P Q h
+ ==> temp_weakening (\<box>P) (\<box>Q) h"
+
+axiomatization where
+(* analog to the proved thm strength_Next - proof skipped as trivial *)
+weak_Next:
+"temp_weakening P Q h
+ ==> temp_weakening (Next P) (Next Q) h"
+
+
+subsection "cex_abs"
+
+lemma cex_abs_UU: "cex_abs f (s,UU) = (f s, UU)"
+  by (simp add: cex_abs_def)
+
+lemma cex_abs_nil: "cex_abs f (s,nil) = (f s, nil)"
+  by (simp add: cex_abs_def)
+
+lemma cex_abs_cons: "cex_abs f (s,(a,t)\<leadsto>ex) = (f s, (a,f t) \<leadsto> (snd (cex_abs f (t,ex))))"
+  by (simp add: cex_abs_def)
+
+declare cex_abs_UU [simp] cex_abs_nil [simp] cex_abs_cons [simp]
+
+
+subsection "lemmas"
+
+lemma temp_weakening_def2: "temp_weakening Q P h = (! ex. (ex \<TTurnstile> P) --> (cex_abs h ex \<TTurnstile> Q))"
+  apply (simp add: temp_weakening_def temp_strengthening_def NOT_def temp_sat_def satisfies_def)
+  apply auto
+  done
+
+lemma state_weakening_def2: "state_weakening Q P h = (! s t a. P (s,a,t) --> Q (h(s),a,h(t)))"
+  apply (simp add: state_weakening_def state_strengthening_def NOT_def)
+  apply auto
+  done
+
+
+subsection "Abstraction Rules for Properties"
+
+lemma exec_frag_abstraction [rule_format]:
+ "[| is_abstraction h C A |] ==>
+  !s. reachable C s & is_exec_frag C (s,xs)
+  --> is_exec_frag A (cex_abs h (s,xs))"
+apply (unfold cex_abs_def)
+apply simp
+apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1\<close>)
+txt \<open>main case\<close>
+apply (auto dest: reachable.reachable_n simp add: is_abstraction_def)
+done
+
+
+lemma abs_is_weakening: "is_abstraction h C A ==> weakeningIOA A C h"
+apply (simp add: weakeningIOA_def)
+apply auto
+apply (simp add: executions_def)
+txt \<open>start state\<close>
+apply (rule conjI)
+apply (simp add: is_abstraction_def cex_abs_def)
+txt \<open>is-execution-fragment\<close>
+apply (erule exec_frag_abstraction)
+apply (simp add: reachable.reachable_0)
+done
+
+
+lemma AbsRuleT1: "[|is_abstraction h C A; validIOA A Q; temp_strengthening Q P h |]
+          ==> validIOA C P"
+apply (drule abs_is_weakening)
+apply (simp add: weakeningIOA_def validIOA_def temp_strengthening_def)
+apply (auto simp add: split_paired_all)
+done
+
+
+(* FIX: Nach TLS.ML *)
+
+lemma IMPLIES_temp_sat: "(ex \<TTurnstile> P \<^bold>\<longrightarrow> Q) = ((ex \<TTurnstile> P) --> (ex \<TTurnstile> Q))"
+  by (simp add: IMPLIES_def temp_sat_def satisfies_def)
+
+lemma AND_temp_sat: "(ex \<TTurnstile> P \<^bold>\<and> Q) = ((ex \<TTurnstile> P) & (ex \<TTurnstile> Q))"
+  by (simp add: AND_def temp_sat_def satisfies_def)
+
+lemma OR_temp_sat: "(ex \<TTurnstile> P \<^bold>\<or> Q) = ((ex \<TTurnstile> P) | (ex \<TTurnstile> Q))"
+  by (simp add: OR_def temp_sat_def satisfies_def)
+
+lemma NOT_temp_sat: "(ex \<TTurnstile> \<^bold>\<not> P) = (~ (ex \<TTurnstile> P))"
+  by (simp add: NOT_def temp_sat_def satisfies_def)
+
+declare IMPLIES_temp_sat [simp] AND_temp_sat [simp] OR_temp_sat [simp] NOT_temp_sat [simp]
+
+
+lemma AbsRuleT2:
+   "[|is_live_abstraction h (C,L) (A,M);
+          validLIOA (A,M) Q;  temp_strengthening Q P h |]
+          ==> validLIOA (C,L) P"
+apply (unfold is_live_abstraction_def)
+apply auto
+apply (drule abs_is_weakening)
+apply (simp add: weakeningIOA_def temp_weakening_def2 validLIOA_def validIOA_def temp_strengthening_def)
+apply (auto simp add: split_paired_all)
+done
+
+
+lemma AbsRuleTImprove:
+   "[|is_live_abstraction h (C,L) (A,M);
+          validLIOA (A,M) (H1 \<^bold>\<longrightarrow> Q);  temp_strengthening Q P h;
+          temp_weakening H1 H2 h; validLIOA (C,L) H2 |]
+          ==> validLIOA (C,L) P"
+apply (unfold is_live_abstraction_def)
+apply auto
+apply (drule abs_is_weakening)
+apply (simp add: weakeningIOA_def temp_weakening_def2 validLIOA_def validIOA_def temp_strengthening_def)
+apply (auto simp add: split_paired_all)
+done
+
+
+subsection "Correctness of safe abstraction"
+
+lemma abstraction_is_ref_map:
+"is_abstraction h C A ==> is_ref_map h C A"
+apply (unfold is_abstraction_def is_ref_map_def)
+apply auto
+apply (rule_tac x = "(a,h t) \<leadsto>nil" in exI)
+apply (simp add: move_def)
+done
+
+
+lemma abs_safety: "[| inp(C)=inp(A); out(C)=out(A);
+                   is_abstraction h C A |]
+                ==> C =<| A"
+apply (simp add: ioa_implements_def)
+apply (rule trace_inclusion)
+apply (simp (no_asm) add: externals_def)
+apply (auto)[1]
+apply (erule abstraction_is_ref_map)
+done
+
+
+subsection "Correctness of life abstraction"
+
+(* Reduces to Filter (Map fst x) = Filter (Map fst (Map (%(a,t). (a,x)) x),
+   that is to special Map Lemma *)
+lemma traces_coincide_abs:
+  "ext C = ext A
+         ==> mk_trace C$xs = mk_trace A$(snd (cex_abs f (s,xs)))"
+apply (unfold cex_abs_def mk_trace_def filter_act_def)
+apply simp
+apply (tactic \<open>pair_induct_tac @{context} "xs" [] 1\<close>)
+done
+
+
+(* Does not work with abstraction_is_ref_map as proof of abs_safety, because
+   is_live_abstraction includes temp_strengthening which is necessarily based
+   on cex_abs and not on corresp_ex. Thus, the proof is redoone in a more specific
+   way for cex_abs *)
+lemma abs_liveness: "[| inp(C)=inp(A); out(C)=out(A);
+                   is_live_abstraction h (C,M) (A,L) |]
+                ==> live_implements (C,M) (A,L)"
+apply (simp add: is_live_abstraction_def live_implements_def livetraces_def liveexecutions_def)
+apply auto
+apply (rule_tac x = "cex_abs h ex" in exI)
+apply auto
+  (* Traces coincide *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (rule traces_coincide_abs)
+  apply (simp (no_asm) add: externals_def)
+  apply (auto)[1]
+
+  (* cex_abs is execution *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (simp add: executions_def)
+  (* start state *)
+  apply (rule conjI)
+  apply (simp add: is_abstraction_def cex_abs_def)
+  (* is-execution-fragment *)
+  apply (erule exec_frag_abstraction)
+  apply (simp add: reachable.reachable_0)
+
+ (* Liveness *)
+apply (simp add: temp_weakening_def2)
+ apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+done
+
+(* FIX: NAch Traces.ML bringen *)
+
+lemma implements_trans:
+"[| A =<| B; B =<| C|] ==> A =<| C"
+by (auto simp add: ioa_implements_def)
+
+
+subsection "Abstraction Rules for Automata"
+
+lemma AbsRuleA1: "[| inp(C)=inp(A); out(C)=out(A);
+                   inp(Q)=inp(P); out(Q)=out(P);
+                   is_abstraction h1 C A;
+                   A =<| Q ;
+                   is_abstraction h2 Q P |]
+                ==> C =<| P"
+apply (drule abs_safety)
+apply assumption+
+apply (drule abs_safety)
+apply assumption+
+apply (erule implements_trans)
+apply (erule implements_trans)
+apply assumption
+done
+
+
+lemma AbsRuleA2: "!!LC. [| inp(C)=inp(A); out(C)=out(A);
+                   inp(Q)=inp(P); out(Q)=out(P);
+                   is_live_abstraction h1 (C,LC) (A,LA);
+                   live_implements (A,LA) (Q,LQ) ;
+                   is_live_abstraction h2 (Q,LQ) (P,LP) |]
+                ==> live_implements (C,LC) (P,LP)"
+apply (drule abs_liveness)
+apply assumption+
+apply (drule abs_liveness)
+apply assumption+
+apply (erule live_implements_trans)
+apply (erule live_implements_trans)
+apply assumption
+done
+
+
+declare split_paired_All [simp del]
+
+
+subsection "Localizing Temporal Strengthenings and Weakenings"
+
+lemma strength_AND:
+"[| temp_strengthening P1 Q1 h;
+          temp_strengthening P2 Q2 h |]
+       ==> temp_strengthening (P1 \<^bold>\<and> P2) (Q1 \<^bold>\<and> Q2) h"
+apply (unfold temp_strengthening_def)
+apply auto
+done
+
+lemma strength_OR:
+"[| temp_strengthening P1 Q1 h;
+          temp_strengthening P2 Q2 h |]
+       ==> temp_strengthening (P1 \<^bold>\<or> P2) (Q1 \<^bold>\<or> Q2) h"
+apply (unfold temp_strengthening_def)
+apply auto
+done
+
+lemma strength_NOT:
+"[| temp_weakening P Q h |]
+       ==> temp_strengthening (\<^bold>\<not> P) (\<^bold>\<not> Q) h"
+apply (unfold temp_strengthening_def)
+apply (simp add: temp_weakening_def2)
+apply auto
+done
+
+lemma strength_IMPLIES:
+"[| temp_weakening P1 Q1 h;
+          temp_strengthening P2 Q2 h |]
+       ==> temp_strengthening (P1 \<^bold>\<longrightarrow> P2) (Q1 \<^bold>\<longrightarrow> Q2) h"
+apply (unfold temp_strengthening_def)
+apply (simp add: temp_weakening_def2)
+done
+
+
+lemma weak_AND:
+"[| temp_weakening P1 Q1 h;
+          temp_weakening P2 Q2 h |]
+       ==> temp_weakening (P1 \<^bold>\<and> P2) (Q1 \<^bold>\<and> Q2) h"
+apply (simp add: temp_weakening_def2)
+done
+
+lemma weak_OR:
+"[| temp_weakening P1 Q1 h;
+          temp_weakening P2 Q2 h |]
+       ==> temp_weakening (P1 \<^bold>\<or> P2) (Q1 \<^bold>\<or> Q2) h"
+apply (simp add: temp_weakening_def2)
+done
+
+lemma weak_NOT:
+"[| temp_strengthening P Q h |]
+       ==> temp_weakening (\<^bold>\<not> P) (\<^bold>\<not> Q) h"
+apply (unfold temp_strengthening_def)
+apply (simp add: temp_weakening_def2)
+apply auto
+done
+
+lemma weak_IMPLIES:
+"[| temp_strengthening P1 Q1 h;
+          temp_weakening P2 Q2 h |]
+       ==> temp_weakening (P1 \<^bold>\<longrightarrow> P2) (Q1 \<^bold>\<longrightarrow> Q2) h"
+apply (unfold temp_strengthening_def)
+apply (simp add: temp_weakening_def2)
+done
+
+
+subsubsection \<open>Box\<close>
+
+(* FIX: should be same as nil_is_Conc2 when all nils are turned to right side !! *)
+lemma UU_is_Conc: "(UU = x @@ y) = (((x::'a Seq)= UU) | (x=nil & y=UU))"
+apply (tactic \<open>Seq_case_simp_tac @{context} "x" 1\<close>)
+done
+
+lemma ex2seqConc [rule_format]:
+"Finite s1 -->
+  (! ex. (s~=nil & s~=UU & ex2seq ex = s1 @@ s) --> (? ex'. s = ex2seq ex'))"
+apply (rule impI)
+apply (tactic \<open>Seq_Finite_induct_tac @{context} 1\<close>)
+apply blast
+(* main case *)
+apply clarify
+apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
+(* UU case *)
+apply (simp add: nil_is_Conc)
+(* nil case *)
+apply (simp add: nil_is_Conc)
+(* cons case *)
+apply (tactic \<open>pair_tac @{context} "aa" 1\<close>)
+apply auto
+done
+
+(* important property of ex2seq: can be shiftet, as defined "pointwise" *)
+
+lemma ex2seq_tsuffix:
+"tsuffix s (ex2seq ex) ==> ? ex'. s = (ex2seq ex')"
+apply (unfold tsuffix_def suffix_def)
+apply auto
+apply (drule ex2seqConc)
+apply auto
+done
+
+
+(* FIX: NAch Sequence.ML bringen *)
+
+lemma Mapnil: "(Map f$s = nil) = (s=nil)"
+apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
+done
+
+lemma MapUU: "(Map f$s = UU) = (s=UU)"
+apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
+done
+
+
+(* important property of cex_absSeq: As it is a 1to1 correspondence,
+  properties carry over *)
+
+lemma cex_absSeq_tsuffix:
+"tsuffix s t ==> tsuffix (cex_absSeq h s) (cex_absSeq h t)"
+apply (unfold tsuffix_def suffix_def cex_absSeq_def)
+apply auto
+apply (simp add: Mapnil)
+apply (simp add: MapUU)
+apply (rule_tac x = "Map (% (s,a,t) . (h s,a, h t))$s1" in exI)
+apply (simp add: Map2Finite MapConc)
+done
+
+
+lemma strength_Box:
+"[| temp_strengthening P Q h |]
+       ==> temp_strengthening (\<box>P) (\<box>Q) h"
+apply (unfold temp_strengthening_def state_strengthening_def temp_sat_def satisfies_def Box_def)
+apply clarify
+apply (frule ex2seq_tsuffix)
+apply clarify
+apply (drule_tac h = "h" in cex_absSeq_tsuffix)
+apply (simp add: ex2seq_abs_cex)
+done
+
+
+subsubsection \<open>Init\<close>
+
+lemma strength_Init:
+"[| state_strengthening P Q h |]
+       ==> temp_strengthening (Init P) (Init Q) h"
+apply (unfold temp_strengthening_def state_strengthening_def
+  temp_sat_def satisfies_def Init_def unlift_def)
+apply auto
+apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+done
+
+
+subsubsection \<open>Next\<close>
+
+lemma TL_ex2seq_UU:
+"(TL$(ex2seq (cex_abs h ex))=UU) = (TL$(ex2seq ex)=UU)"
+apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+done
+
+lemma TL_ex2seq_nil:
+"(TL$(ex2seq (cex_abs h ex))=nil) = (TL$(ex2seq ex)=nil)"
+apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+done
+
+(* FIX: put to Sequence Lemmas *)
+lemma MapTL: "Map f$(TL$s) = TL$(Map f$s)"
+apply (tactic \<open>Seq_induct_tac @{context} "s" [] 1\<close>)
+done
+
+(* important property of cex_absSeq: As it is a 1to1 correspondence,
+  properties carry over *)
+
+lemma cex_absSeq_TL:
+"cex_absSeq h (TL$s) = (TL$(cex_absSeq h s))"
+apply (unfold cex_absSeq_def)
+apply (simp add: MapTL)
+done
+
+(* important property of ex2seq: can be shiftet, as defined "pointwise" *)
+
+lemma TLex2seq: "[| (snd ex)~=UU ; (snd ex)~=nil |] ==> (? ex'. TL$(ex2seq ex) = ex2seq ex')"
+apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+apply auto
+done
+
+
+lemma ex2seqnilTL: "(TL$(ex2seq ex)~=nil) = ((snd ex)~=nil & (snd ex)~=UU)"
+apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+done
+
+
+lemma strength_Next:
+"[| temp_strengthening P Q h |]
+       ==> temp_strengthening (Next P) (Next Q) h"
+apply (unfold temp_strengthening_def state_strengthening_def temp_sat_def satisfies_def Next_def)
+apply simp
+apply auto
+apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
+apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
+apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
+apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
+(* cons case *)
+apply (simp add: TL_ex2seq_nil TL_ex2seq_UU ex2seq_abs_cex cex_absSeq_TL [symmetric] ex2seqnilTL)
+apply (erule conjE)
+apply (drule TLex2seq)
+apply assumption
+apply auto
+done
+
+
+text \<open>Localizing Temporal Weakenings     - 2\<close>
+
+lemma weak_Init:
+"[| state_weakening P Q h |]
+       ==> temp_weakening (Init P) (Init Q) h"
+apply (simp add: temp_weakening_def2 state_weakening_def2
+  temp_sat_def satisfies_def Init_def unlift_def)
+apply auto
+apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+done
+
+
+text \<open>Localizing Temproal Strengthenings - 3\<close>
+
+lemma strength_Diamond:
+"[| temp_strengthening P Q h |]
+       ==> temp_strengthening (\<diamond>P) (\<diamond>Q) h"
+apply (unfold Diamond_def)
+apply (rule strength_NOT)
+apply (rule weak_Box)
+apply (erule weak_NOT)
+done
+
+lemma strength_Leadsto:
+"[| temp_weakening P1 P2 h;
+          temp_strengthening Q1 Q2 h |]
+       ==> temp_strengthening (P1 \<leadsto> Q1) (P2 \<leadsto> Q2) h"
+apply (unfold Leadsto_def)
+apply (rule strength_Box)
+apply (erule strength_IMPLIES)
+apply (erule strength_Diamond)
+done
+
+
+text \<open>Localizing Temporal Weakenings - 3\<close>
+
+lemma weak_Diamond:
+"[| temp_weakening P Q h |]
+       ==> temp_weakening (\<diamond>P) (\<diamond>Q) h"
+apply (unfold Diamond_def)
+apply (rule weak_NOT)
+apply (rule strength_Box)
+apply (erule strength_NOT)
+done
+
+lemma weak_Leadsto:
+"[| temp_strengthening P1 P2 h;
+          temp_weakening Q1 Q2 h |]
+       ==> temp_weakening (P1 \<leadsto> Q1) (P2 \<leadsto> Q2) h"
+apply (unfold Leadsto_def)
+apply (rule weak_Box)
+apply (erule weak_IMPLIES)
+apply (erule weak_Diamond)
+done
+
+lemma weak_WF:
+  " !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|]
+    ==> temp_weakening (WF A acts) (WF C acts) h"
+apply (unfold WF_def)
+apply (rule weak_IMPLIES)
+apply (rule strength_Diamond)
+apply (rule strength_Box)
+apply (rule strength_Init)
+apply (rule_tac [2] weak_Box)
+apply (rule_tac [2] weak_Diamond)
+apply (rule_tac [2] weak_Init)
+apply (auto simp add: state_weakening_def state_strengthening_def
+  xt2_def plift_def option_lift_def NOT_def)
+done
+
+lemma weak_SF:
+  " !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|]
+    ==> temp_weakening (SF A acts) (SF C acts) h"
+apply (unfold SF_def)
+apply (rule weak_IMPLIES)
+apply (rule strength_Box)
+apply (rule strength_Diamond)
+apply (rule strength_Init)
+apply (rule_tac [2] weak_Box)
+apply (rule_tac [2] weak_Diamond)
+apply (rule_tac [2] weak_Init)
+apply (auto simp add: state_weakening_def state_strengthening_def
+  xt2_def plift_def option_lift_def NOT_def)
+done
+
+
+lemmas weak_strength_lemmas =
+  weak_OR weak_AND weak_NOT weak_IMPLIES weak_Box weak_Next weak_Init
+  weak_Diamond weak_Leadsto strength_OR strength_AND strength_NOT
+  strength_IMPLIES strength_Box strength_Next strength_Init
+  strength_Diamond strength_Leadsto weak_WF weak_SF
+
+ML \<open>
+fun abstraction_tac ctxt =
+  SELECT_GOAL (auto_tac
+    (ctxt addSIs @{thms weak_strength_lemmas}
+      addsimps [@{thm state_strengthening_def}, @{thm state_weakening_def}]))
+\<close>
+
+method_setup abstraction = \<open>Scan.succeed (SIMPLE_METHOD' o abstraction_tac)\<close>
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/Asig.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,101 @@
+(*  Title:      HOL/HOLCF/IOA/Asig.thy
+    Author:     Olaf Müller, Tobias Nipkow & Konrad Slind
+*)
+
+section \<open>Action signatures\<close>
+
+theory Asig
+imports Main
+begin
+
+type_synonym
+  'a signature = "('a set * 'a set * 'a set)"
+
+definition
+  inputs :: "'action signature => 'action set" where
+  asig_inputs_def: "inputs = fst"
+
+definition
+  outputs :: "'action signature => 'action set" where
+  asig_outputs_def: "outputs = (fst o snd)"
+
+definition
+  internals :: "'action signature => 'action set" where
+  asig_internals_def: "internals = (snd o snd)"
+
+definition
+  actions :: "'action signature => 'action set" where
+  "actions(asig) = (inputs(asig) Un outputs(asig) Un internals(asig))"
+
+definition
+  externals :: "'action signature => 'action set" where
+  "externals(asig) = (inputs(asig) Un outputs(asig))"
+
+definition
+  locals :: "'action signature => 'action set" where
+  "locals asig = ((internals asig) Un (outputs asig))"
+
+definition
+  is_asig :: "'action signature => bool" where
+  "is_asig(triple) =
+     ((inputs(triple) Int outputs(triple) = {}) &
+      (outputs(triple) Int internals(triple) = {}) &
+      (inputs(triple) Int internals(triple) = {}))"
+
+definition
+  mk_ext_asig :: "'action signature => 'action signature" where
+  "mk_ext_asig(triple) = (inputs(triple), outputs(triple), {})"
+
+
+lemmas asig_projections = asig_inputs_def asig_outputs_def asig_internals_def
+
+lemma asig_triple_proj:
+ "(outputs    (a,b,c) = b)   &
+  (inputs     (a,b,c) = a) &
+  (internals  (a,b,c) = c)"
+  apply (simp add: asig_projections)
+  done
+
+lemma int_and_ext_is_act: "[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)"
+apply (simp add: externals_def actions_def)
+done
+
+lemma ext_is_act: "[|a:externals(S)|] ==> a:actions(S)"
+apply (simp add: externals_def actions_def)
+done
+
+lemma int_is_act: "[|a:internals S|] ==> a:actions S"
+apply (simp add: asig_internals_def actions_def)
+done
+
+lemma inp_is_act: "[|a:inputs S|] ==> a:actions S"
+apply (simp add: asig_inputs_def actions_def)
+done
+
+lemma out_is_act: "[|a:outputs S|] ==> a:actions S"
+apply (simp add: asig_outputs_def actions_def)
+done
+
+lemma ext_and_act: "(x: actions S & x : externals S) = (x: externals S)"
+apply (fast intro!: ext_is_act)
+done
+
+lemma not_ext_is_int: "[|is_asig S;x: actions S|] ==> (x~:externals S) = (x: internals S)"
+apply (simp add: actions_def is_asig_def externals_def)
+apply blast
+done
+
+lemma not_ext_is_int_or_not_act: "is_asig S ==> (x~:externals S) = (x: internals S | x~:actions S)"
+apply (simp add: actions_def is_asig_def externals_def)
+apply blast
+done
+
+lemma int_is_not_ext:
+ "[| is_asig (S); x:internals S |] ==> x~:externals S"
+apply (unfold externals_def actions_def is_asig_def)
+apply simp
+apply blast
+done
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/Automata.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,615 @@
+(*  Title:      HOL/HOLCF/IOA/Automata.thy
+    Author:     Olaf Müller, Konrad Slind, Tobias Nipkow
+*)
+
+section \<open>The I/O automata of Lynch and Tuttle in HOLCF\<close>
+
+theory Automata
+imports Asig
+begin
+
+default_sort type
+
+type_synonym ('a, 's) transition = "'s * 'a * 's"
+type_synonym ('a, 's) ioa =
+  "'a signature * 's set * ('a,'s)transition set * ('a set set) * ('a set set)"
+
+
+(* --------------------------------- IOA ---------------------------------*)
+
+(* IO automata *)
+
+definition asig_of :: "('a, 's)ioa \<Rightarrow> 'a signature"
+  where "asig_of = fst"
+
+definition starts_of :: "('a, 's) ioa \<Rightarrow> 's set"
+  where "starts_of = (fst \<circ> snd)"
+
+definition trans_of :: "('a, 's) ioa \<Rightarrow> ('a, 's) transition set"
+  where "trans_of = (fst \<circ> snd \<circ> snd)"
+
+abbreviation trans_of_syn  ("_ \<midarrow>_\<midarrow>_\<rightarrow> _" [81, 81, 81, 81] 100)
+  where "s \<midarrow>a\<midarrow>A\<rightarrow> t \<equiv> (s, a, t) \<in> trans_of A"
+
+definition wfair_of :: "('a, 's) ioa \<Rightarrow> 'a set set"
+  where "wfair_of = (fst \<circ> snd \<circ> snd \<circ> snd)"
+
+definition sfair_of :: "('a, 's) ioa \<Rightarrow> 'a set set"
+  where "sfair_of = (snd \<circ> snd \<circ> snd \<circ> snd)"
+
+definition is_asig_of :: "('a, 's) ioa \<Rightarrow> bool"
+  where "is_asig_of A = is_asig (asig_of A)"
+
+definition is_starts_of :: "('a, 's) ioa \<Rightarrow> bool"
+  where "is_starts_of A \<longleftrightarrow> starts_of A \<noteq> {}"
+
+definition is_trans_of :: "('a, 's) ioa \<Rightarrow> bool"
+  where "is_trans_of A \<longleftrightarrow>
+    (\<forall>triple. triple \<in> trans_of A \<longrightarrow> fst (snd triple) \<in> actions (asig_of A))"
+
+definition input_enabled :: "('a, 's) ioa \<Rightarrow> bool"
+  where "input_enabled A \<longleftrightarrow>
+    (\<forall>a. a \<in> inputs (asig_of A) \<longrightarrow> (\<forall>s1. \<exists>s2. (s1, a, s2) \<in> trans_of A))"
+
+definition IOA :: "('a, 's) ioa \<Rightarrow> bool"
+  where "IOA A \<longleftrightarrow>
+    is_asig_of A \<and>
+    is_starts_of A \<and>
+    is_trans_of A \<and>
+    input_enabled A"
+
+abbreviation "act A == actions (asig_of A)"
+abbreviation "ext A == externals (asig_of A)"
+abbreviation int where "int A == internals (asig_of A)"
+abbreviation "inp A == inputs (asig_of A)"
+abbreviation "out A == outputs (asig_of A)"
+abbreviation "local A == locals (asig_of A)"
+
+(* invariants *)
+inductive reachable :: "('a, 's) ioa \<Rightarrow> 's \<Rightarrow> bool"
+  for C :: "('a, 's) ioa"
+where
+  reachable_0:  "s \<in> starts_of C \<Longrightarrow> reachable C s"
+| reachable_n:  "\<lbrakk>reachable C s; (s, a, t) \<in> trans_of C\<rbrakk> \<Longrightarrow> reachable C t"
+
+definition invariant :: "[('a, 's) ioa, 's \<Rightarrow> bool] \<Rightarrow> bool"
+  where "invariant A P \<longleftrightarrow> (\<forall>s. reachable A s \<longrightarrow> P s)"
+
+
+(* ------------------------- parallel composition --------------------------*)
+
+(* binary composition of action signatures and automata *)
+
+definition compatible :: "[('a, 's) ioa, ('a, 't) ioa] \<Rightarrow> bool"
+where
+  "compatible A B \<longleftrightarrow>
+  (((out A \<inter> out B) = {}) \<and>
+   ((int A \<inter> act B) = {}) \<and>
+   ((int B \<inter> act A) = {}))"
+
+definition asig_comp :: "['a signature, 'a signature] \<Rightarrow> 'a signature"
+where
+  "asig_comp a1 a2 =
+     (((inputs(a1) \<union> inputs(a2)) - (outputs(a1) \<union> outputs(a2)),
+       (outputs(a1) \<union> outputs(a2)),
+       (internals(a1) \<union> internals(a2))))"
+
+definition par :: "[('a, 's) ioa, ('a, 't) ioa] \<Rightarrow> ('a, 's * 't) ioa"  (infixr "\<parallel>" 10)
+where
+  "(A \<parallel> B) =
+      (asig_comp (asig_of A) (asig_of B),
+       {pr. fst(pr) \<in> starts_of(A) \<and> snd(pr) \<in> starts_of(B)},
+       {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))
+            in (a \<in> act A | a:act B) \<and>
+               (if a \<in> act A then
+                  (fst(s), a, fst(t)) \<in> trans_of(A)
+                else fst(t) = fst(s))
+               &
+               (if a \<in> act B then
+                  (snd(s), a, snd(t)) \<in> trans_of(B)
+                else snd(t) = snd(s))},
+        wfair_of A \<union> wfair_of B,
+        sfair_of A \<union> sfair_of B)"
+
+
+(* ------------------------ hiding -------------------------------------------- *)
+
+(* hiding and restricting *)
+
+definition restrict_asig :: "['a signature, 'a set] \<Rightarrow> 'a signature"
+where
+  "restrict_asig asig actns =
+    (inputs(asig) Int actns,
+     outputs(asig) Int actns,
+     internals(asig) Un (externals(asig) - actns))"
+
+(* Notice that for wfair_of and sfair_of nothing has to be changed, as
+   changes from the outputs to the internals does not touch the locals as
+   a whole, which is of importance for fairness only *)
+definition restrict :: "[('a, 's) ioa, 'a set] \<Rightarrow> ('a, 's) ioa"
+where
+  "restrict A actns =
+    (restrict_asig (asig_of A) actns,
+     starts_of A,
+     trans_of A,
+     wfair_of A,
+     sfair_of A)"
+
+definition hide_asig :: "['a signature, 'a set] \<Rightarrow> 'a signature"
+where
+  "hide_asig asig actns =
+    (inputs(asig) - actns,
+     outputs(asig) - actns,
+     internals(asig) \<union> actns)"
+
+definition hide :: "[('a, 's) ioa, 'a set] \<Rightarrow> ('a, 's) ioa"
+where
+  "hide A actns =
+    (hide_asig (asig_of A) actns,
+     starts_of A,
+     trans_of A,
+     wfair_of A,
+     sfair_of A)"
+
+(* ------------------------- renaming ------------------------------------------- *)
+
+definition rename_set :: "'a set \<Rightarrow> ('c \<Rightarrow> 'a option) \<Rightarrow> 'c set"
+  where "rename_set A ren = {b. \<exists>x. Some x = ren b \<and> x \<in> A}"
+
+definition rename :: "('a, 'b) ioa \<Rightarrow> ('c \<Rightarrow> 'a option) \<Rightarrow> ('c, 'b) ioa"
+where
+  "rename ioa ren =
+    ((rename_set (inp ioa) ren,
+      rename_set (out ioa) ren,
+      rename_set (int ioa) ren),
+     starts_of ioa,
+     {tr. let s = fst(tr); a = fst(snd(tr));  t = snd(snd(tr))
+          in
+          \<exists>x. Some(x) = ren(a) \<and> (s,x,t):trans_of ioa},
+     {rename_set s ren | s. s \<in> wfair_of ioa},
+     {rename_set s ren | s. s \<in> sfair_of ioa})"
+
+
+(* ------------------------- fairness ----------------------------- *)
+
+(* enabledness of actions and action sets *)
+
+definition enabled :: "('a, 's) ioa \<Rightarrow> 'a \<Rightarrow> 's \<Rightarrow> bool"
+  where "enabled A a s \<longleftrightarrow> (\<exists>t. s \<midarrow>a\<midarrow>A\<rightarrow> t)"
+
+definition Enabled :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> 's \<Rightarrow> bool"
+  where "Enabled A W s \<longleftrightarrow> (\<exists>w \<in> W. enabled A w s)"
+
+
+(* action set keeps enabled until probably disabled by itself *)
+
+definition en_persistent :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> bool"
+where
+  "en_persistent A W \<longleftrightarrow>
+    (\<forall>s a t. Enabled A W s \<and> a \<notin> W \<and> s \<midarrow>a\<midarrow>A\<rightarrow> t \<longrightarrow> Enabled A W t)"
+
+
+(* post_conditions for actions and action sets *)
+
+definition was_enabled :: "('a, 's) ioa \<Rightarrow> 'a \<Rightarrow> 's \<Rightarrow> bool"
+  where "was_enabled A a t \<longleftrightarrow> (\<exists>s. s \<midarrow>a\<midarrow>A\<rightarrow> t)"
+
+definition set_was_enabled :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> 's \<Rightarrow> bool"
+  where "set_was_enabled A W t \<longleftrightarrow> (\<exists>w \<in> W. was_enabled A w t)"
+
+
+(* constraints for fair IOA *)
+
+definition fairIOA :: "('a, 's) ioa \<Rightarrow> bool"
+  where "fairIOA A \<longleftrightarrow> (\<forall>S \<in> wfair_of A. S \<subseteq> local A) \<and> (\<forall>S \<in> sfair_of A. S \<subseteq> local A)"
+
+definition input_resistant :: "('a, 's) ioa \<Rightarrow> bool"
+where
+  "input_resistant A \<longleftrightarrow>
+    (\<forall>W \<in> sfair_of A. \<forall>s a t.
+      reachable A s \<and> reachable A t \<and> a \<in> inp A \<and>
+      Enabled A W s \<and> s \<midarrow>a\<midarrow>A\<rightarrow> t \<longrightarrow> Enabled A W t)"
+
+
+declare split_paired_Ex [simp del]
+
+lemmas ioa_projections = asig_of_def starts_of_def trans_of_def wfair_of_def sfair_of_def
+
+
+subsection "asig_of, starts_of, trans_of"
+
+lemma ioa_triple_proj:
+ "((asig_of (x,y,z,w,s)) = x)   &
+  ((starts_of (x,y,z,w,s)) = y) &
+  ((trans_of (x,y,z,w,s)) = z)  &
+  ((wfair_of (x,y,z,w,s)) = w) &
+  ((sfair_of (x,y,z,w,s)) = s)"
+  apply (simp add: ioa_projections)
+  done
+
+lemma trans_in_actions:
+  "[| is_trans_of A; (s1,a,s2):trans_of(A) |] ==> a:act A"
+  apply (unfold is_trans_of_def actions_def is_asig_def)
+    apply (erule allE, erule impE, assumption)
+    apply simp
+  done
+
+lemma starts_of_par: "starts_of(A \<parallel> B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}"
+  by (simp add: par_def ioa_projections)
+
+lemma trans_of_par:
+"trans_of(A \<parallel> B) = {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))
+             in (a:act A | a:act B) &
+                (if a:act A then
+                   (fst(s),a,fst(t)):trans_of(A)
+                 else fst(t) = fst(s))
+                &
+                (if a:act B then
+                   (snd(s),a,snd(t)):trans_of(B)
+                 else snd(t) = snd(s))}"
+  by (simp add: par_def ioa_projections)
+
+
+subsection "actions and par"
+
+lemma actions_asig_comp: "actions(asig_comp a b) = actions(a) Un actions(b)"
+  by (auto simp add: actions_def asig_comp_def asig_projections)
+
+lemma asig_of_par: "asig_of(A \<parallel> B) = asig_comp (asig_of A) (asig_of B)"
+  by (simp add: par_def ioa_projections)
+
+
+lemma externals_of_par: "ext (A1\<parallel>A2) = (ext A1) Un (ext A2)"
+  apply (simp add: externals_def asig_of_par asig_comp_def
+    asig_inputs_def asig_outputs_def Un_def set_diff_eq)
+  apply blast
+  done
+
+lemma actions_of_par: "act (A1\<parallel>A2) = (act A1) Un (act A2)"
+  apply (simp add: actions_def asig_of_par asig_comp_def
+    asig_inputs_def asig_outputs_def asig_internals_def Un_def set_diff_eq)
+  apply blast
+  done
+
+lemma inputs_of_par: "inp (A1\<parallel>A2) = ((inp A1) Un (inp A2)) - ((out A1) Un (out A2))"
+  by (simp add: actions_def asig_of_par asig_comp_def
+    asig_inputs_def asig_outputs_def Un_def set_diff_eq)
+
+lemma outputs_of_par: "out (A1\<parallel>A2) = (out A1) Un (out A2)"
+  by (simp add: actions_def asig_of_par asig_comp_def
+    asig_outputs_def Un_def set_diff_eq)
+
+lemma internals_of_par: "int (A1\<parallel>A2) = (int A1) Un (int A2)"
+  by (simp add: actions_def asig_of_par asig_comp_def
+    asig_inputs_def asig_outputs_def asig_internals_def Un_def set_diff_eq)
+
+
+subsection "actions and compatibility"
+
+lemma compat_commute: "compatible A B = compatible B A"
+  by (auto simp add: compatible_def Int_commute)
+
+lemma ext1_is_not_int2: "[| compatible A1 A2; a:ext A1|] ==> a~:int A2"
+  apply (unfold externals_def actions_def compatible_def)
+  apply simp
+  apply blast
+  done
+
+(* just commuting the previous one: better commute compatible *)
+lemma ext2_is_not_int1: "[| compatible A2 A1 ; a:ext A1|] ==> a~:int A2"
+  apply (unfold externals_def actions_def compatible_def)
+  apply simp
+  apply blast
+  done
+
+lemmas ext1_ext2_is_not_act2 = ext1_is_not_int2 [THEN int_and_ext_is_act]
+lemmas ext1_ext2_is_not_act1 = ext2_is_not_int1 [THEN int_and_ext_is_act]
+
+lemma intA_is_not_extB: "[| compatible A B; x:int A |] ==> x~:ext B"
+  apply (unfold externals_def actions_def compatible_def)
+  apply simp
+  apply blast
+  done
+
+lemma intA_is_not_actB: "[| compatible A B; a:int A |] ==> a ~: act B"
+  apply (unfold externals_def actions_def compatible_def is_asig_def asig_of_def)
+  apply simp
+  apply blast
+  done
+
+(* the only one that needs disjointness of outputs and of internals and _all_ acts *)
+lemma outAactB_is_inpB: "[| compatible A B; a:out A ;a:act B|] ==> a : inp B"
+  apply (unfold asig_outputs_def asig_internals_def actions_def asig_inputs_def
+      compatible_def is_asig_def asig_of_def)
+  apply simp
+  apply blast
+  done
+
+(* needed for propagation of input_enabledness from A,B to A\<parallel>B *)
+lemma inpAAactB_is_inpBoroutB:
+  "[| compatible A B; a:inp A ;a:act B|] ==> a : inp B | a: out B"
+  apply (unfold asig_outputs_def asig_internals_def actions_def asig_inputs_def
+      compatible_def is_asig_def asig_of_def)
+  apply simp
+  apply blast
+  done
+
+
+subsection "input_enabledness and par"
+
+(* ugly case distinctions. Heart of proof:
+     1. inpAAactB_is_inpBoroutB ie. internals are really hidden.
+     2. inputs_of_par: outputs are no longer inputs of par. This is important here *)
+lemma input_enabled_par:
+  "[| compatible A B; input_enabled A; input_enabled B|]
+        ==> input_enabled (A\<parallel>B)"
+  apply (unfold input_enabled_def)
+  apply (simp add: Let_def inputs_of_par trans_of_par)
+  apply (tactic "safe_tac (Context.raw_transfer @{theory} @{theory_context Fun})")
+  apply (simp add: inp_is_act)
+  prefer 2
+  apply (simp add: inp_is_act)
+  (* a: inp A *)
+  apply (case_tac "a:act B")
+  (* a:act B *)
+  apply (erule_tac x = "a" in allE)
+  apply simp
+  apply (drule inpAAactB_is_inpBoroutB)
+  apply assumption
+  apply assumption
+  apply (erule_tac x = "a" in allE)
+  apply simp
+  apply (erule_tac x = "aa" in allE)
+  apply (erule_tac x = "b" in allE)
+  apply (erule exE)
+  apply (erule exE)
+  apply (rule_tac x = " (s2,s2a) " in exI)
+  apply (simp add: inp_is_act)
+  (* a~: act B*)
+  apply (simp add: inp_is_act)
+  apply (erule_tac x = "a" in allE)
+  apply simp
+  apply (erule_tac x = "aa" in allE)
+  apply (erule exE)
+  apply (rule_tac x = " (s2,b) " in exI)
+  apply simp
+  
+  (* a:inp B *)
+  apply (case_tac "a:act A")
+  (* a:act A *)
+  apply (erule_tac x = "a" in allE)
+  apply (erule_tac x = "a" in allE)
+  apply (simp add: inp_is_act)
+  apply (frule_tac A1 = "A" in compat_commute [THEN iffD1])
+  apply (drule inpAAactB_is_inpBoroutB)
+  back
+  apply assumption
+  apply assumption
+  apply simp
+  apply (erule_tac x = "aa" in allE)
+  apply (erule_tac x = "b" in allE)
+  apply (erule exE)
+  apply (erule exE)
+  apply (rule_tac x = " (s2,s2a) " in exI)
+  apply (simp add: inp_is_act)
+  (* a~: act B*)
+  apply (simp add: inp_is_act)
+  apply (erule_tac x = "a" in allE)
+  apply (erule_tac x = "a" in allE)
+  apply simp
+  apply (erule_tac x = "b" in allE)
+  apply (erule exE)
+  apply (rule_tac x = " (aa,s2) " in exI)
+  apply simp
+  done
+
+
+subsection "invariants"
+
+lemma invariantI:
+  "[| !!s. s:starts_of(A) ==> P(s);
+      !!s t a. [|reachable A s; P(s)|] ==> (s,a,t): trans_of(A) --> P(t) |]
+   ==> invariant A P"
+  apply (unfold invariant_def)
+  apply (rule allI)
+  apply (rule impI)
+  apply (rule_tac x = "s" in reachable.induct)
+  apply assumption
+  apply blast
+  apply blast
+  done
+
+lemma invariantI1:
+ "[| !!s. s : starts_of(A) ==> P(s);
+     !!s t a. reachable A s ==> P(s) --> (s,a,t):trans_of(A) --> P(t)
+  |] ==> invariant A P"
+  apply (blast intro: invariantI)
+  done
+
+lemma invariantE: "[| invariant A P; reachable A s |] ==> P(s)"
+  apply (unfold invariant_def)
+  apply blast
+  done
+
+
+subsection "restrict"
+
+
+lemmas reachable_0 = reachable.reachable_0
+  and reachable_n = reachable.reachable_n
+
+lemma cancel_restrict_a: "starts_of(restrict ioa acts) = starts_of(ioa) &
+          trans_of(restrict ioa acts) = trans_of(ioa)"
+  by (simp add: restrict_def ioa_projections)
+
+lemma cancel_restrict_b: "reachable (restrict ioa acts) s = reachable ioa s"
+  apply (rule iffI)
+  apply (erule reachable.induct)
+  apply (simp add: cancel_restrict_a reachable_0)
+  apply (erule reachable_n)
+  apply (simp add: cancel_restrict_a)
+  (* <--  *)
+  apply (erule reachable.induct)
+  apply (rule reachable_0)
+  apply (simp add: cancel_restrict_a)
+  apply (erule reachable_n)
+  apply (simp add: cancel_restrict_a)
+  done
+
+lemma acts_restrict: "act (restrict A acts) = act A"
+  apply (simp (no_asm) add: actions_def asig_internals_def
+    asig_outputs_def asig_inputs_def externals_def asig_of_def restrict_def restrict_asig_def)
+  apply auto
+  done
+
+lemma cancel_restrict: "starts_of(restrict ioa acts) = starts_of(ioa) &
+          trans_of(restrict ioa acts) = trans_of(ioa) &
+          reachable (restrict ioa acts) s = reachable ioa s &
+          act (restrict A acts) = act A"
+  by (simp add: cancel_restrict_a cancel_restrict_b acts_restrict)
+
+
+subsection "rename"
+
+lemma trans_rename: "s \<midarrow>a\<midarrow>(rename C f)\<rightarrow> t ==> (? x. Some(x) = f(a) & s \<midarrow>x\<midarrow>C\<rightarrow> t)"
+  by (simp add: Let_def rename_def trans_of_def)
+
+
+lemma reachable_rename: "[| reachable (rename C g) s |] ==> reachable C s"
+  apply (erule reachable.induct)
+  apply (rule reachable_0)
+  apply (simp add: rename_def ioa_projections)
+  apply (drule trans_rename)
+  apply (erule exE)
+  apply (erule conjE)
+  apply (erule reachable_n)
+  apply assumption
+  done
+
+
+subsection "trans_of(A\<parallel>B)"
+
+lemma trans_A_proj: "[|(s,a,t):trans_of (A\<parallel>B); a:act A|]
+              ==> (fst s,a,fst t):trans_of A"
+  by (simp add: Let_def par_def trans_of_def)
+
+lemma trans_B_proj: "[|(s,a,t):trans_of (A\<parallel>B); a:act B|]
+              ==> (snd s,a,snd t):trans_of B"
+  by (simp add: Let_def par_def trans_of_def)
+
+lemma trans_A_proj2: "[|(s,a,t):trans_of (A\<parallel>B); a~:act A|]
+              ==> fst s = fst t"
+  by (simp add: Let_def par_def trans_of_def)
+
+lemma trans_B_proj2: "[|(s,a,t):trans_of (A\<parallel>B); a~:act B|]
+              ==> snd s = snd t"
+  by (simp add: Let_def par_def trans_of_def)
+
+lemma trans_AB_proj: "(s,a,t):trans_of (A\<parallel>B)
+               ==> a :act A | a :act B"
+  by (simp add: Let_def par_def trans_of_def)
+
+lemma trans_AB: "[|a:act A;a:act B;
+       (fst s,a,fst t):trans_of A;(snd s,a,snd t):trans_of B|]
+   ==> (s,a,t):trans_of (A\<parallel>B)"
+  by (simp add: Let_def par_def trans_of_def)
+
+lemma trans_A_notB: "[|a:act A;a~:act B;
+       (fst s,a,fst t):trans_of A;snd s=snd t|]
+   ==> (s,a,t):trans_of (A\<parallel>B)"
+  by (simp add: Let_def par_def trans_of_def)
+
+lemma trans_notA_B: "[|a~:act A;a:act B;
+       (snd s,a,snd t):trans_of B;fst s=fst t|]
+   ==> (s,a,t):trans_of (A\<parallel>B)"
+  by (simp add: Let_def par_def trans_of_def)
+
+lemmas trans_of_defs1 = trans_AB trans_A_notB trans_notA_B
+  and trans_of_defs2 = trans_A_proj trans_B_proj trans_A_proj2 trans_B_proj2 trans_AB_proj
+
+
+lemma trans_of_par4:
+"((s,a,t) : trans_of(A \<parallel> B \<parallel> C \<parallel> D)) =
+  ((a:actions(asig_of(A)) | a:actions(asig_of(B)) | a:actions(asig_of(C)) |
+    a:actions(asig_of(D))) &
+   (if a:actions(asig_of(A)) then (fst(s),a,fst(t)):trans_of(A)
+    else fst t=fst s) &
+   (if a:actions(asig_of(B)) then (fst(snd(s)),a,fst(snd(t))):trans_of(B)
+    else fst(snd(t))=fst(snd(s))) &
+   (if a:actions(asig_of(C)) then
+      (fst(snd(snd(s))),a,fst(snd(snd(t)))):trans_of(C)
+    else fst(snd(snd(t)))=fst(snd(snd(s)))) &
+   (if a:actions(asig_of(D)) then
+      (snd(snd(snd(s))),a,snd(snd(snd(t)))):trans_of(D)
+    else snd(snd(snd(t)))=snd(snd(snd(s)))))"
+  by (simp add: par_def actions_asig_comp prod_eq_iff Let_def ioa_projections)
+
+
+subsection "proof obligation generator for IOA requirements"
+
+(* without assumptions on A and B because is_trans_of is also incorporated in \<parallel>def *)
+lemma is_trans_of_par: "is_trans_of (A\<parallel>B)"
+  by (simp add: is_trans_of_def Let_def actions_of_par trans_of_par)
+
+lemma is_trans_of_restrict: "is_trans_of A ==> is_trans_of (restrict A acts)"
+  by (simp add: is_trans_of_def cancel_restrict acts_restrict)
+
+lemma is_trans_of_rename: "is_trans_of A ==> is_trans_of (rename A f)"
+  apply (unfold is_trans_of_def restrict_def restrict_asig_def)
+  apply (simp add: Let_def actions_def trans_of_def asig_internals_def
+    asig_outputs_def asig_inputs_def externals_def asig_of_def rename_def rename_set_def)
+  apply blast
+  done
+
+lemma is_asig_of_par: "[| is_asig_of A; is_asig_of B; compatible A B|]
+          ==> is_asig_of (A\<parallel>B)"
+  apply (simp add: is_asig_of_def asig_of_par asig_comp_def compatible_def
+    asig_internals_def asig_outputs_def asig_inputs_def actions_def is_asig_def)
+  apply (simp add: asig_of_def)
+  apply auto
+  done
+
+lemma is_asig_of_restrict: "is_asig_of A ==> is_asig_of (restrict A f)"
+  apply (unfold is_asig_of_def is_asig_def asig_of_def restrict_def restrict_asig_def
+             asig_internals_def asig_outputs_def asig_inputs_def externals_def o_def)
+  apply simp
+  apply auto
+  done
+
+lemma is_asig_of_rename: "is_asig_of A ==> is_asig_of (rename A f)"
+  apply (simp add: is_asig_of_def rename_def rename_set_def asig_internals_def
+    asig_outputs_def asig_inputs_def actions_def is_asig_def asig_of_def)
+  apply auto
+  apply (drule_tac [!] s = "Some _" in sym)
+  apply auto
+  done
+
+lemmas [simp] = is_asig_of_par is_asig_of_restrict
+  is_asig_of_rename is_trans_of_par is_trans_of_restrict is_trans_of_rename
+
+
+lemma compatible_par: "[|compatible A B; compatible A C |]==> compatible A (B\<parallel>C)"
+  apply (unfold compatible_def)
+  apply (simp add: internals_of_par outputs_of_par actions_of_par)
+  apply auto
+  done
+
+(*  better derive by previous one and compat_commute *)
+lemma compatible_par2: "[|compatible A C; compatible B C |]==> compatible (A\<parallel>B) C"
+  apply (unfold compatible_def)
+  apply (simp add: internals_of_par outputs_of_par actions_of_par)
+  apply auto
+  done
+
+lemma compatible_restrict:
+  "[| compatible A B; (ext B - S) Int ext A = {}|]
+        ==> compatible A (restrict B S)"
+  apply (unfold compatible_def)
+  apply (simp add: ioa_triple_proj asig_triple_proj externals_def
+    restrict_def restrict_asig_def actions_def)
+  apply auto
+  done
+
+declare split_paired_Ex [simp]
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/CompoExecs.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,303 @@
+(*  Title:      HOL/HOLCF/IOA/CompoExecs.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Compositionality on Execution level\<close>
+
+theory CompoExecs
+imports Traces
+begin
+
+definition
+  ProjA2 :: "('a,'s * 't)pairs -> ('a,'s)pairs" where
+  "ProjA2 = Map (%x.(fst x,fst(snd x)))"
+
+definition
+  ProjA :: "('a,'s * 't)execution => ('a,'s)execution" where
+  "ProjA ex = (fst (fst ex), ProjA2$(snd ex))"
+
+definition
+  ProjB2 :: "('a,'s * 't)pairs -> ('a,'t)pairs" where
+  "ProjB2 = Map (%x.(fst x,snd(snd x)))"
+
+definition
+  ProjB :: "('a,'s * 't)execution => ('a,'t)execution" where
+  "ProjB ex = (snd (fst ex), ProjB2$(snd ex))"
+
+definition
+  Filter_ex2 :: "'a signature => ('a,'s)pairs -> ('a,'s)pairs" where
+  "Filter_ex2 sig = Filter (%x. fst x:actions sig)"
+
+definition
+  Filter_ex :: "'a signature => ('a,'s)execution => ('a,'s)execution" where
+  "Filter_ex sig ex = (fst ex,Filter_ex2 sig$(snd ex))"
+
+definition
+  stutter2 :: "'a signature => ('a,'s)pairs -> ('s => tr)" where
+  "stutter2 sig = (fix$(LAM h ex. (%s. case ex of
+      nil => TT
+    | x##xs => (flift1
+            (%p.(If Def ((fst p)~:actions sig)
+                 then Def (s=(snd p))
+                 else TT)
+                andalso (h$xs) (snd p))
+             $x)
+   )))"
+
+definition
+  stutter :: "'a signature => ('a,'s)execution => bool" where
+  "stutter sig ex = ((stutter2 sig$(snd ex)) (fst ex) ~= FF)"
+
+definition
+  par_execs :: "[('a,'s)execution_module,('a,'t)execution_module] => ('a,'s*'t)execution_module" where
+  "par_execs ExecsA ExecsB =
+      (let exA = fst ExecsA; sigA = snd ExecsA;
+           exB = fst ExecsB; sigB = snd ExecsB
+       in
+       (    {ex. Filter_ex sigA (ProjA ex) : exA}
+        Int {ex. Filter_ex sigB (ProjB ex) : exB}
+        Int {ex. stutter sigA (ProjA ex)}
+        Int {ex. stutter sigB (ProjB ex)}
+        Int {ex. Forall (%x. fst x:(actions sigA Un actions sigB)) (snd ex)},
+        asig_comp sigA sigB))"
+
+
+lemmas [simp del] = split_paired_All
+
+
+section "recursive equations of operators"
+
+
+(* ---------------------------------------------------------------- *)
+(*                               ProjA2                             *)
+(* ---------------------------------------------------------------- *)
+
+
+lemma ProjA2_UU: "ProjA2$UU = UU"
+apply (simp add: ProjA2_def)
+done
+
+lemma ProjA2_nil: "ProjA2$nil = nil"
+apply (simp add: ProjA2_def)
+done
+
+lemma ProjA2_cons: "ProjA2$((a,t)\<leadsto>xs) = (a,fst t) \<leadsto> ProjA2$xs"
+apply (simp add: ProjA2_def)
+done
+
+
+(* ---------------------------------------------------------------- *)
+(*                               ProjB2                             *)
+(* ---------------------------------------------------------------- *)
+
+
+lemma ProjB2_UU: "ProjB2$UU = UU"
+apply (simp add: ProjB2_def)
+done
+
+lemma ProjB2_nil: "ProjB2$nil = nil"
+apply (simp add: ProjB2_def)
+done
+
+lemma ProjB2_cons: "ProjB2$((a,t)\<leadsto>xs) = (a,snd t) \<leadsto> ProjB2$xs"
+apply (simp add: ProjB2_def)
+done
+
+
+
+(* ---------------------------------------------------------------- *)
+(*                             Filter_ex2                           *)
+(* ---------------------------------------------------------------- *)
+
+
+lemma Filter_ex2_UU: "Filter_ex2 sig$UU=UU"
+apply (simp add: Filter_ex2_def)
+done
+
+lemma Filter_ex2_nil: "Filter_ex2 sig$nil=nil"
+apply (simp add: Filter_ex2_def)
+done
+
+lemma Filter_ex2_cons: "Filter_ex2 sig$(at \<leadsto> xs) =
+             (if (fst at:actions sig)
+                  then at \<leadsto> (Filter_ex2 sig$xs)
+                  else        Filter_ex2 sig$xs)"
+
+apply (simp add: Filter_ex2_def)
+done
+
+
+(* ---------------------------------------------------------------- *)
+(*                             stutter2                             *)
+(* ---------------------------------------------------------------- *)
+
+
+lemma stutter2_unfold: "stutter2 sig = (LAM ex. (%s. case ex of
+       nil => TT
+     | x##xs => (flift1
+             (%p.(If Def ((fst p)~:actions sig)
+                  then Def (s=(snd p))
+                  else TT)
+                 andalso (stutter2 sig$xs) (snd p))
+              $x)
+            ))"
+apply (rule trans)
+apply (rule fix_eq2)
+apply (simp only: stutter2_def)
+apply (rule beta_cfun)
+apply (simp add: flift1_def)
+done
+
+lemma stutter2_UU: "(stutter2 sig$UU) s=UU"
+apply (subst stutter2_unfold)
+apply simp
+done
+
+lemma stutter2_nil: "(stutter2 sig$nil) s = TT"
+apply (subst stutter2_unfold)
+apply simp
+done
+
+lemma stutter2_cons: "(stutter2 sig$(at\<leadsto>xs)) s =
+               ((if (fst at)~:actions sig then Def (s=snd at) else TT)
+                 andalso (stutter2 sig$xs) (snd at))"
+apply (rule trans)
+apply (subst stutter2_unfold)
+apply (simp add: Consq_def flift1_def If_and_if)
+apply simp
+done
+
+
+declare stutter2_UU [simp] stutter2_nil [simp] stutter2_cons [simp]
+
+
+(* ---------------------------------------------------------------- *)
+(*                             stutter                              *)
+(* ---------------------------------------------------------------- *)
+
+lemma stutter_UU: "stutter sig (s, UU)"
+apply (simp add: stutter_def)
+done
+
+lemma stutter_nil: "stutter sig (s, nil)"
+apply (simp add: stutter_def)
+done
+
+lemma stutter_cons: "stutter sig (s, (a,t)\<leadsto>ex) =
+      ((a~:actions sig --> (s=t)) & stutter sig (t,ex))"
+apply (simp add: stutter_def)
+done
+
+(* ----------------------------------------------------------------------------------- *)
+
+declare stutter2_UU [simp del] stutter2_nil [simp del] stutter2_cons [simp del]
+
+lemmas compoex_simps = ProjA2_UU ProjA2_nil ProjA2_cons
+  ProjB2_UU ProjB2_nil ProjB2_cons
+  Filter_ex2_UU Filter_ex2_nil Filter_ex2_cons
+  stutter_UU stutter_nil stutter_cons
+
+declare compoex_simps [simp]
+
+
+
+(* ------------------------------------------------------------------ *)
+(*                      The following lemmata aim for                 *)
+(*             COMPOSITIONALITY   on    EXECUTION     Level           *)
+(* ------------------------------------------------------------------ *)
+
+
+(* --------------------------------------------------------------------- *)
+(*  Lemma_1_1a : is_ex_fr propagates from A\<parallel>B to Projections A and B    *)
+(* --------------------------------------------------------------------- *)
+
+lemma lemma_1_1a: "!s. is_exec_frag (A\<parallel>B) (s,xs)
+       -->  is_exec_frag A (fst s, Filter_ex2 (asig_of A)$(ProjA2$xs)) &
+            is_exec_frag B (snd s, Filter_ex2 (asig_of B)$(ProjB2$xs))"
+
+apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1\<close>)
+(* main case *)
+apply (auto simp add: trans_of_defs2)
+done
+
+
+(* --------------------------------------------------------------------- *)
+(*  Lemma_1_1b : is_ex_fr (A\<parallel>B) implies stuttering on Projections       *)
+(* --------------------------------------------------------------------- *)
+
+lemma lemma_1_1b: "!s. is_exec_frag (A\<parallel>B) (s,xs)
+       --> stutter (asig_of A) (fst s,ProjA2$xs)  &
+           stutter (asig_of B) (snd s,ProjB2$xs)"
+
+apply (tactic \<open>pair_induct_tac @{context} "xs"
+  [@{thm stutter_def}, @{thm is_exec_frag_def}] 1\<close>)
+(* main case *)
+apply (auto simp add: trans_of_defs2)
+done
+
+
+(* --------------------------------------------------------------------- *)
+(*  Lemma_1_1c : Executions of A\<parallel>B have only  A- or B-actions           *)
+(* --------------------------------------------------------------------- *)
+
+lemma lemma_1_1c: "!s. (is_exec_frag (A\<parallel>B) (s,xs)
+   --> Forall (%x. fst x:act (A\<parallel>B)) xs)"
+
+apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm Forall_def}, @{thm sforall_def},
+  @{thm is_exec_frag_def}] 1\<close>)
+(* main case *)
+apply auto
+apply (simp add: trans_of_defs2 actions_asig_comp asig_of_par)
+done
+
+
+(* ----------------------------------------------------------------------- *)
+(*  Lemma_1_2 : ex A, exB, stuttering and forall a:A|B implies ex (A\<parallel>B)   *)
+(* ----------------------------------------------------------------------- *)
+
+lemma lemma_1_2:
+"!s. is_exec_frag A (fst s,Filter_ex2 (asig_of A)$(ProjA2$xs)) &
+     is_exec_frag B (snd s,Filter_ex2 (asig_of B)$(ProjB2$xs)) &
+     stutter (asig_of A) (fst s,(ProjA2$xs)) &
+     stutter (asig_of B) (snd s,(ProjB2$xs)) &
+     Forall (%x. fst x:act (A\<parallel>B)) xs
+     --> is_exec_frag (A\<parallel>B) (s,xs)"
+
+apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm Forall_def}, @{thm sforall_def},
+  @{thm is_exec_frag_def}, @{thm stutter_def}] 1\<close>)
+apply (auto simp add: trans_of_defs1 actions_asig_comp asig_of_par)
+done
+
+
+subsection \<open>COMPOSITIONALITY on EXECUTION Level -- Main Theorem\<close>
+
+lemma compositionality_ex:
+"(ex:executions(A\<parallel>B)) =
+ (Filter_ex (asig_of A) (ProjA ex) : executions A &
+  Filter_ex (asig_of B) (ProjB ex) : executions B &
+  stutter (asig_of A) (ProjA ex) & stutter (asig_of B) (ProjB ex) &
+  Forall (%x. fst x:act (A\<parallel>B)) (snd ex))"
+
+apply (simp (no_asm) add: executions_def ProjB_def Filter_ex_def ProjA_def starts_of_par)
+apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+apply (rule iffI)
+(* ==>  *)
+apply (erule conjE)+
+apply (simp add: lemma_1_1a lemma_1_1b lemma_1_1c)
+(* <==  *)
+apply (erule conjE)+
+apply (simp add: lemma_1_2)
+done
+
+
+subsection \<open>COMPOSITIONALITY on EXECUTION Level -- for Modules\<close>
+
+lemma compositionality_ex_modules:
+  "Execs (A\<parallel>B) = par_execs (Execs A) (Execs B)"
+apply (unfold Execs_def par_execs_def)
+apply (simp add: asig_of_par)
+apply (rule set_eqI)
+apply (simp add: compositionality_ex actions_of_par)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/CompoScheds.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,541 @@
+(*  Title:      HOL/HOLCF/IOA/CompoScheds.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Compositionality on Schedule level\<close>
+
+theory CompoScheds
+imports CompoExecs
+begin
+
+definition
+  mkex2 :: "('a,'s)ioa => ('a,'t)ioa => 'a Seq ->
+              ('a,'s)pairs -> ('a,'t)pairs ->
+              ('s => 't => ('a,'s*'t)pairs)" where
+  "mkex2 A B = (fix$(LAM h sch exA exB. (%s t. case sch of
+       nil => nil
+    | x##xs =>
+      (case x of
+        UU => UU
+      | Def y =>
+         (if y:act A then
+             (if y:act B then
+                (case HD$exA of
+                   UU => UU
+                 | Def a => (case HD$exB of
+                              UU => UU
+                            | Def b =>
+                   (y,(snd a,snd b))\<leadsto>
+                     (h$xs$(TL$exA)$(TL$exB)) (snd a) (snd b)))
+              else
+                (case HD$exA of
+                   UU => UU
+                 | Def a =>
+                   (y,(snd a,t))\<leadsto>(h$xs$(TL$exA)$exB) (snd a) t)
+              )
+          else
+             (if y:act B then
+                (case HD$exB of
+                   UU => UU
+                 | Def b =>
+                   (y,(s,snd b))\<leadsto>(h$xs$exA$(TL$exB)) s (snd b))
+             else
+               UU
+             )
+         )
+       ))))"
+
+definition
+  mkex :: "('a,'s)ioa => ('a,'t)ioa => 'a Seq =>
+              ('a,'s)execution => ('a,'t)execution =>('a,'s*'t)execution" where
+  "mkex A B sch exA exB =
+       ((fst exA,fst exB),
+        (mkex2 A B$sch$(snd exA)$(snd exB)) (fst exA) (fst exB))"
+
+definition
+  par_scheds ::"['a schedule_module,'a schedule_module] => 'a schedule_module" where
+  "par_scheds SchedsA SchedsB =
+      (let schA = fst SchedsA; sigA = snd SchedsA;
+           schB = fst SchedsB; sigB = snd SchedsB
+       in
+       (    {sch. Filter (%a. a:actions sigA)$sch : schA}
+        Int {sch. Filter (%a. a:actions sigB)$sch : schB}
+        Int {sch. Forall (%x. x:(actions sigA Un actions sigB)) sch},
+        asig_comp sigA sigB))"
+
+
+subsection "mkex rewrite rules"
+
+
+lemma mkex2_unfold:
+"mkex2 A B = (LAM sch exA exB. (%s t. case sch of
+      nil => nil
+   | x##xs =>
+     (case x of
+       UU => UU
+     | Def y =>
+        (if y:act A then
+            (if y:act B then
+               (case HD$exA of
+                  UU => UU
+                | Def a => (case HD$exB of
+                             UU => UU
+                           | Def b =>
+                  (y,(snd a,snd b))\<leadsto>
+                    (mkex2 A B$xs$(TL$exA)$(TL$exB)) (snd a) (snd b)))
+             else
+               (case HD$exA of
+                  UU => UU
+                | Def a =>
+                  (y,(snd a,t))\<leadsto>(mkex2 A B$xs$(TL$exA)$exB) (snd a) t)
+             )
+         else
+            (if y:act B then
+               (case HD$exB of
+                  UU => UU
+                | Def b =>
+                  (y,(s,snd b))\<leadsto>(mkex2 A B$xs$exA$(TL$exB)) s (snd b))
+            else
+              UU
+            )
+        )
+      )))"
+apply (rule trans)
+apply (rule fix_eq2)
+apply (simp only: mkex2_def)
+apply (rule beta_cfun)
+apply simp
+done
+
+lemma mkex2_UU: "(mkex2 A B$UU$exA$exB) s t = UU"
+apply (subst mkex2_unfold)
+apply simp
+done
+
+lemma mkex2_nil: "(mkex2 A B$nil$exA$exB) s t= nil"
+apply (subst mkex2_unfold)
+apply simp
+done
+
+lemma mkex2_cons_1: "[| x:act A; x~:act B; HD$exA=Def a|]
+    ==> (mkex2 A B$(x\<leadsto>sch)$exA$exB) s t =
+        (x,snd a,t) \<leadsto> (mkex2 A B$sch$(TL$exA)$exB) (snd a) t"
+apply (rule trans)
+apply (subst mkex2_unfold)
+apply (simp add: Consq_def If_and_if)
+apply (simp add: Consq_def)
+done
+
+lemma mkex2_cons_2: "[| x~:act A; x:act B; HD$exB=Def b|]
+    ==> (mkex2 A B$(x\<leadsto>sch)$exA$exB) s t =
+        (x,s,snd b) \<leadsto> (mkex2 A B$sch$exA$(TL$exB)) s (snd b)"
+apply (rule trans)
+apply (subst mkex2_unfold)
+apply (simp add: Consq_def If_and_if)
+apply (simp add: Consq_def)
+done
+
+lemma mkex2_cons_3: "[| x:act A; x:act B; HD$exA=Def a;HD$exB=Def b|]
+    ==> (mkex2 A B$(x\<leadsto>sch)$exA$exB) s t =
+         (x,snd a,snd b) \<leadsto>
+            (mkex2 A B$sch$(TL$exA)$(TL$exB)) (snd a) (snd b)"
+apply (rule trans)
+apply (subst mkex2_unfold)
+apply (simp add: Consq_def If_and_if)
+apply (simp add: Consq_def)
+done
+
+declare mkex2_UU [simp] mkex2_nil [simp] mkex2_cons_1 [simp]
+  mkex2_cons_2 [simp] mkex2_cons_3 [simp]
+
+
+subsection \<open>mkex\<close>
+
+lemma mkex_UU: "mkex A B UU  (s,exA) (t,exB) = ((s,t),UU)"
+apply (simp add: mkex_def)
+done
+
+lemma mkex_nil: "mkex A B nil (s,exA) (t,exB) = ((s,t),nil)"
+apply (simp add: mkex_def)
+done
+
+lemma mkex_cons_1: "[| x:act A; x~:act B |]
+    ==> mkex A B (x\<leadsto>sch) (s,a\<leadsto>exA) (t,exB)  =
+        ((s,t), (x,snd a,t) \<leadsto> snd (mkex A B sch (snd a,exA) (t,exB)))"
+apply (simp (no_asm) add: mkex_def)
+apply (cut_tac exA = "a\<leadsto>exA" in mkex2_cons_1)
+apply auto
+done
+
+lemma mkex_cons_2: "[| x~:act A; x:act B |]
+    ==> mkex A B (x\<leadsto>sch) (s,exA) (t,b\<leadsto>exB) =
+        ((s,t), (x,s,snd b) \<leadsto> snd (mkex A B sch (s,exA) (snd b,exB)))"
+apply (simp (no_asm) add: mkex_def)
+apply (cut_tac exB = "b\<leadsto>exB" in mkex2_cons_2)
+apply auto
+done
+
+lemma mkex_cons_3: "[| x:act A; x:act B |]
+    ==>  mkex A B (x\<leadsto>sch) (s,a\<leadsto>exA) (t,b\<leadsto>exB) =
+         ((s,t), (x,snd a,snd b) \<leadsto> snd (mkex A B sch (snd a,exA) (snd b,exB)))"
+apply (simp (no_asm) add: mkex_def)
+apply (cut_tac exB = "b\<leadsto>exB" and exA = "a\<leadsto>exA" in mkex2_cons_3)
+apply auto
+done
+
+declare mkex2_UU [simp del] mkex2_nil [simp del]
+  mkex2_cons_1 [simp del] mkex2_cons_2 [simp del] mkex2_cons_3 [simp del]
+
+lemmas composch_simps = mkex_UU mkex_nil mkex_cons_1 mkex_cons_2 mkex_cons_3
+
+declare composch_simps [simp]
+
+
+subsection \<open>COMPOSITIONALITY on SCHEDULE Level\<close>
+
+subsubsection "Lemmas for ==>"
+
+(* --------------------------------------------------------------------- *)
+(*    Lemma_2_1 :  tfilter(ex) and filter_act are commutative            *)
+(* --------------------------------------------------------------------- *)
+
+lemma lemma_2_1a:
+   "filter_act$(Filter_ex2 (asig_of A)$xs)=
+    Filter (%a. a:act A)$(filter_act$xs)"
+
+apply (unfold filter_act_def Filter_ex2_def)
+apply (simp (no_asm) add: MapFilter o_def)
+done
+
+
+(* --------------------------------------------------------------------- *)
+(*    Lemma_2_2 : State-projections do not affect filter_act             *)
+(* --------------------------------------------------------------------- *)
+
+lemma lemma_2_1b:
+   "filter_act$(ProjA2$xs) =filter_act$xs &
+    filter_act$(ProjB2$xs) =filter_act$xs"
+apply (tactic \<open>pair_induct_tac @{context} "xs" [] 1\<close>)
+done
+
+
+(* --------------------------------------------------------------------- *)
+(*             Schedules of A\<parallel>B have only  A- or B-actions              *)
+(* --------------------------------------------------------------------- *)
+
+(* very similar to lemma_1_1c, but it is not checking if every action element of
+   an ex is in A or B, but after projecting it onto the action schedule. Of course, this
+   is the same proposition, but we cannot change this one, when then rather lemma_1_1c  *)
+
+lemma sch_actions_in_AorB: "!s. is_exec_frag (A\<parallel>B) (s,xs)
+   --> Forall (%x. x:act (A\<parallel>B)) (filter_act$xs)"
+
+apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}, @{thm Forall_def},
+  @{thm sforall_def}] 1\<close>)
+(* main case *)
+apply auto
+apply (simp add: trans_of_defs2 actions_asig_comp asig_of_par)
+done
+
+
+subsubsection "Lemmas for <=="
+
+(*---------------------------------------------------------------------------
+    Filtering actions out of mkex(sch,exA,exB) yields the oracle sch
+                             structural induction
+  --------------------------------------------------------------------------- *)
+
+lemma Mapfst_mkex_is_sch: "! exA exB s t.
+  Forall (%x. x:act (A\<parallel>B)) sch  &
+  Filter (%a. a:act A)$sch << filter_act$exA &
+  Filter (%a. a:act B)$sch << filter_act$exB
+  --> filter_act$(snd (mkex A B sch (s,exA) (t,exB))) = sch"
+
+apply (tactic \<open>Seq_induct_tac @{context} "sch" [@{thm Filter_def}, @{thm Forall_def},
+  @{thm sforall_def}, @{thm mkex_def}] 1\<close>)
+
+(* main case *)
+(* splitting into 4 cases according to a:A, a:B *)
+apply auto
+
+(* Case y:A, y:B *)
+apply (tactic \<open>Seq_case_simp_tac @{context} "exA" 1\<close>)
+(* Case exA=UU, Case exA=nil*)
+(* These UU and nil cases are the only places where the assumption filter A sch<<f_act exA
+   is used! --> to generate a contradiction using  ~a\<leadsto>ss<< UU(nil), using theorems
+   Cons_not_less_UU and Cons_not_less_nil  *)
+apply (tactic \<open>Seq_case_simp_tac @{context} "exB" 1\<close>)
+(* Case exA=a\<leadsto>x, exB=b\<leadsto>y *)
+(* here it is important that Seq_case_simp_tac uses no !full!_simp_tac for the cons case,
+   as otherwise mkex_cons_3 would  not be rewritten without use of rotate_tac: then tactic
+   would not be generally applicable *)
+apply simp
+
+(* Case y:A, y~:B *)
+apply (tactic \<open>Seq_case_simp_tac @{context} "exA" 1\<close>)
+apply simp
+
+(* Case y~:A, y:B *)
+apply (tactic \<open>Seq_case_simp_tac @{context} "exB" 1\<close>)
+apply simp
+
+(* Case y~:A, y~:B *)
+apply (simp add: asig_of_par actions_asig_comp)
+done
+
+
+(* generalizing the proof above to a proof method *)
+
+ML \<open>
+fun mkex_induct_tac ctxt sch exA exB =
+  EVERY'[Seq_induct_tac ctxt sch @{thms Filter_def Forall_def sforall_def mkex_def stutter_def},
+         asm_full_simp_tac ctxt,
+         SELECT_GOAL
+          (safe_tac (Context.raw_transfer (Proof_Context.theory_of ctxt) @{theory_context Fun})),
+         Seq_case_simp_tac ctxt exA,
+         Seq_case_simp_tac ctxt exB,
+         asm_full_simp_tac ctxt,
+         Seq_case_simp_tac ctxt exA,
+         asm_full_simp_tac ctxt,
+         Seq_case_simp_tac ctxt exB,
+         asm_full_simp_tac ctxt,
+         asm_full_simp_tac (ctxt addsimps @{thms asig_of_par actions_asig_comp})
+        ]
+\<close>
+
+method_setup mkex_induct = \<open>
+  Scan.lift (Args.name -- Args.name -- Args.name)
+    >> (fn ((sch, exA), exB) => fn ctxt => SIMPLE_METHOD' (mkex_induct_tac ctxt sch exA exB))
+\<close>
+
+
+(*---------------------------------------------------------------------------
+               Projection of mkex(sch,exA,exB) onto A stutters on A
+                             structural induction
+  --------------------------------------------------------------------------- *)
+
+lemma stutterA_mkex: "! exA exB s t.
+  Forall (%x. x:act (A\<parallel>B)) sch &
+  Filter (%a. a:act A)$sch << filter_act$exA &
+  Filter (%a. a:act B)$sch << filter_act$exB
+  --> stutter (asig_of A) (s,ProjA2$(snd (mkex A B sch (s,exA) (t,exB))))"
+  by (mkex_induct sch exA exB)
+
+lemma stutter_mkex_on_A: "[|
+  Forall (%x. x:act (A\<parallel>B)) sch ;
+  Filter (%a. a:act A)$sch << filter_act$(snd exA) ;
+  Filter (%a. a:act B)$sch << filter_act$(snd exB) |]
+  ==> stutter (asig_of A) (ProjA (mkex A B sch exA exB))"
+
+apply (cut_tac stutterA_mkex)
+apply (simp add: stutter_def ProjA_def mkex_def)
+apply (erule allE)+
+apply (drule mp)
+prefer 2 apply (assumption)
+apply simp
+done
+
+
+(*---------------------------------------------------------------------------
+               Projection of mkex(sch,exA,exB) onto B stutters on B
+                             structural induction
+  --------------------------------------------------------------------------- *)
+
+lemma stutterB_mkex: "! exA exB s t.
+  Forall (%x. x:act (A\<parallel>B)) sch &
+  Filter (%a. a:act A)$sch << filter_act$exA &
+  Filter (%a. a:act B)$sch << filter_act$exB
+  --> stutter (asig_of B) (t,ProjB2$(snd (mkex A B sch (s,exA) (t,exB))))"
+  by (mkex_induct sch exA exB)
+
+
+lemma stutter_mkex_on_B: "[|
+  Forall (%x. x:act (A\<parallel>B)) sch ;
+  Filter (%a. a:act A)$sch << filter_act$(snd exA) ;
+  Filter (%a. a:act B)$sch << filter_act$(snd exB) |]
+  ==> stutter (asig_of B) (ProjB (mkex A B sch exA exB))"
+apply (cut_tac stutterB_mkex)
+apply (simp add: stutter_def ProjB_def mkex_def)
+apply (erule allE)+
+apply (drule mp)
+prefer 2 apply (assumption)
+apply simp
+done
+
+
+(*---------------------------------------------------------------------------
+     Filter of mkex(sch,exA,exB) to A after projection onto A is exA
+        --  using zip$(proj1$exA)$(proj2$exA) instead of exA    --
+        --           because of admissibility problems          --
+                             structural induction
+  --------------------------------------------------------------------------- *)
+
+lemma filter_mkex_is_exA_tmp: "! exA exB s t.
+  Forall (%x. x:act (A\<parallel>B)) sch &
+  Filter (%a. a:act A)$sch << filter_act$exA  &
+  Filter (%a. a:act B)$sch << filter_act$exB
+  --> Filter_ex2 (asig_of A)$(ProjA2$(snd (mkex A B sch (s,exA) (t,exB)))) =
+      Zip$(Filter (%a. a:act A)$sch)$(Map snd$exA)"
+  by (mkex_induct sch exB exA)
+
+(*---------------------------------------------------------------------------
+                      zip$(proj1$y)$(proj2$y) = y   (using the lift operations)
+                    lemma for admissibility problems
+  --------------------------------------------------------------------------- *)
+
+lemma Zip_Map_fst_snd: "Zip$(Map fst$y)$(Map snd$y) = y"
+apply (tactic \<open>Seq_induct_tac @{context} "y" [] 1\<close>)
+done
+
+
+(*---------------------------------------------------------------------------
+      filter A$sch = proj1$ex   -->  zip$(filter A$sch)$(proj2$ex) = ex
+         lemma for eliminating non admissible equations in assumptions
+  --------------------------------------------------------------------------- *)
+
+lemma trick_against_eq_in_ass: "!! sch ex.
+  Filter (%a. a:act AB)$sch = filter_act$ex
+  ==> ex = Zip$(Filter (%a. a:act AB)$sch)$(Map snd$ex)"
+apply (simp add: filter_act_def)
+apply (rule Zip_Map_fst_snd [symmetric])
+done
+
+(*---------------------------------------------------------------------------
+     Filter of mkex(sch,exA,exB) to A after projection onto A is exA
+                       using the above trick
+  --------------------------------------------------------------------------- *)
+
+
+lemma filter_mkex_is_exA: "!!sch exA exB.
+  [| Forall (%a. a:act (A\<parallel>B)) sch ;
+  Filter (%a. a:act A)$sch = filter_act$(snd exA)  ;
+  Filter (%a. a:act B)$sch = filter_act$(snd exB) |]
+  ==> Filter_ex (asig_of A) (ProjA (mkex A B sch exA exB)) = exA"
+apply (simp add: ProjA_def Filter_ex_def)
+apply (tactic \<open>pair_tac @{context} "exA" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "exB" 1\<close>)
+apply (rule conjI)
+apply (simp (no_asm) add: mkex_def)
+apply (simplesubst trick_against_eq_in_ass)
+back
+apply assumption
+apply (simp add: filter_mkex_is_exA_tmp)
+done
+
+
+(*---------------------------------------------------------------------------
+     Filter of mkex(sch,exA,exB) to B after projection onto B is exB
+        --  using zip$(proj1$exB)$(proj2$exB) instead of exB    --
+        --           because of admissibility problems          --
+                             structural induction
+  --------------------------------------------------------------------------- *)
+
+lemma filter_mkex_is_exB_tmp: "! exA exB s t.
+  Forall (%x. x:act (A\<parallel>B)) sch &
+  Filter (%a. a:act A)$sch << filter_act$exA  &
+  Filter (%a. a:act B)$sch << filter_act$exB
+  --> Filter_ex2 (asig_of B)$(ProjB2$(snd (mkex A B sch (s,exA) (t,exB)))) =
+      Zip$(Filter (%a. a:act B)$sch)$(Map snd$exB)"
+
+(* notice necessary change of arguments exA and exB *)
+  by (mkex_induct sch exA exB)
+
+
+(*---------------------------------------------------------------------------
+     Filter of mkex(sch,exA,exB) to A after projection onto B is exB
+                       using the above trick
+  --------------------------------------------------------------------------- *)
+
+
+lemma filter_mkex_is_exB: "!!sch exA exB.
+  [| Forall (%a. a:act (A\<parallel>B)) sch ;
+  Filter (%a. a:act A)$sch = filter_act$(snd exA)  ;
+  Filter (%a. a:act B)$sch = filter_act$(snd exB) |]
+  ==> Filter_ex (asig_of B) (ProjB (mkex A B sch exA exB)) = exB"
+apply (simp add: ProjB_def Filter_ex_def)
+apply (tactic \<open>pair_tac @{context} "exA" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "exB" 1\<close>)
+apply (rule conjI)
+apply (simp (no_asm) add: mkex_def)
+apply (simplesubst trick_against_eq_in_ass)
+back
+apply assumption
+apply (simp add: filter_mkex_is_exB_tmp)
+done
+
+(* --------------------------------------------------------------------- *)
+(*                    mkex has only  A- or B-actions                    *)
+(* --------------------------------------------------------------------- *)
+
+
+lemma mkex_actions_in_AorB: "!s t exA exB.
+  Forall (%x. x : act (A \<parallel> B)) sch &
+  Filter (%a. a:act A)$sch << filter_act$exA  &
+  Filter (%a. a:act B)$sch << filter_act$exB
+   --> Forall (%x. fst x : act (A \<parallel>B))
+         (snd (mkex A B sch (s,exA) (t,exB)))"
+  by (mkex_induct sch exA exB)
+
+
+(* ------------------------------------------------------------------ *)
+(*           COMPOSITIONALITY   on    SCHEDULE      Level             *)
+(*                          Main Theorem                              *)
+(* ------------------------------------------------------------------ *)
+
+lemma compositionality_sch:
+"(sch : schedules (A\<parallel>B)) =
+  (Filter (%a. a:act A)$sch : schedules A &
+   Filter (%a. a:act B)$sch : schedules B &
+   Forall (%x. x:act (A\<parallel>B)) sch)"
+apply (simp add: schedules_def has_schedule_def)
+apply auto
+(* ==> *)
+apply (rule_tac x = "Filter_ex (asig_of A) (ProjA ex) " in bexI)
+prefer 2
+apply (simp add: compositionality_ex)
+apply (simp (no_asm) add: Filter_ex_def ProjA_def lemma_2_1a lemma_2_1b)
+apply (rule_tac x = "Filter_ex (asig_of B) (ProjB ex) " in bexI)
+prefer 2
+apply (simp add: compositionality_ex)
+apply (simp (no_asm) add: Filter_ex_def ProjB_def lemma_2_1a lemma_2_1b)
+apply (simp add: executions_def)
+apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+apply (erule conjE)
+apply (simp add: sch_actions_in_AorB)
+
+(* <== *)
+
+(* mkex is exactly the construction of exA\<parallel>B out of exA, exB, and the oracle sch,
+   we need here *)
+apply (rename_tac exA exB)
+apply (rule_tac x = "mkex A B sch exA exB" in bexI)
+(* mkex actions are just the oracle *)
+apply (tactic \<open>pair_tac @{context} "exA" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "exB" 1\<close>)
+apply (simp add: Mapfst_mkex_is_sch)
+
+(* mkex is an execution -- use compositionality on ex-level *)
+apply (simp add: compositionality_ex)
+apply (simp add: stutter_mkex_on_A stutter_mkex_on_B filter_mkex_is_exB filter_mkex_is_exA)
+apply (tactic \<open>pair_tac @{context} "exA" 1\<close>)
+apply (tactic \<open>pair_tac @{context} "exB" 1\<close>)
+apply (simp add: mkex_actions_in_AorB)
+done
+
+
+subsection \<open>COMPOSITIONALITY on SCHEDULE Level -- for Modules\<close>
+
+lemma compositionality_sch_modules:
+  "Scheds (A\<parallel>B) = par_scheds (Scheds A) (Scheds B)"
+
+apply (unfold Scheds_def par_scheds_def)
+apply (simp add: asig_of_par)
+apply (rule set_eqI)
+apply (simp add: compositionality_sch actions_of_par)
+done
+
+
+declare compoex_simps [simp del]
+declare composch_simps [simp del]
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/CompoTraces.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,969 @@
+(*  Title:      HOL/HOLCF/IOA/CompoTraces.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Compositionality on Trace level\<close>
+
+theory CompoTraces
+imports CompoScheds ShortExecutions
+begin
+
+definition mksch :: "('a,'s)ioa => ('a,'t)ioa => 'a Seq -> 'a Seq -> 'a Seq -> 'a Seq"
+where
+  "mksch A B = (fix$(LAM h tr schA schB. case tr of
+       nil => nil
+    | x##xs =>
+      (case x of
+        UU => UU
+      | Def y =>
+         (if y:act A then
+             (if y:act B then
+                   ((Takewhile (%a. a:int A)$schA)
+                      @@ (Takewhile (%a. a:int B)$schB)
+                           @@ (y\<leadsto>(h$xs
+                                    $(TL$(Dropwhile (%a. a:int A)$schA))
+                                    $(TL$(Dropwhile (%a. a:int B)$schB))
+                    )))
+              else
+                 ((Takewhile (%a. a:int A)$schA)
+                  @@ (y\<leadsto>(h$xs
+                           $(TL$(Dropwhile (%a. a:int A)$schA))
+                           $schB)))
+              )
+          else
+             (if y:act B then
+                 ((Takewhile (%a. a:int B)$schB)
+                     @@ (y\<leadsto>(h$xs
+                              $schA
+                              $(TL$(Dropwhile (%a. a:int B)$schB))
+                              )))
+             else
+               UU
+             )
+         )
+       )))"
+
+definition par_traces ::"['a trace_module,'a trace_module] => 'a trace_module"
+where
+  "par_traces TracesA TracesB =
+      (let trA = fst TracesA; sigA = snd TracesA;
+           trB = fst TracesB; sigB = snd TracesB
+       in
+       (    {tr. Filter (%a. a:actions sigA)$tr : trA}
+        Int {tr. Filter (%a. a:actions sigB)$tr : trB}
+        Int {tr. Forall (%x. x:(externals sigA Un externals sigB)) tr},
+        asig_comp sigA sigB))"
+
+axiomatization
+where
+  finiteR_mksch:
+    "Finite (mksch A B$tr$x$y) \<Longrightarrow> Finite tr"
+
+lemma finiteR_mksch': "\<not> Finite tr \<Longrightarrow> \<not> Finite (mksch A B$tr$x$y)"
+  by (blast intro: finiteR_mksch)
+
+
+declaration \<open>fn _ => Simplifier.map_ss (Simplifier.set_mksym (K (K NONE)))\<close>
+
+
+subsection "mksch rewrite rules"
+
+lemma mksch_unfold:
+"mksch A B = (LAM tr schA schB. case tr of
+       nil => nil
+    | x##xs =>
+      (case x of
+        UU => UU
+      | Def y =>
+         (if y:act A then
+             (if y:act B then
+                   ((Takewhile (%a. a:int A)$schA)
+                         @@(Takewhile (%a. a:int B)$schB)
+                              @@(y\<leadsto>(mksch A B$xs
+                                       $(TL$(Dropwhile (%a. a:int A)$schA))
+                                       $(TL$(Dropwhile (%a. a:int B)$schB))
+                    )))
+              else
+                 ((Takewhile (%a. a:int A)$schA)
+                      @@ (y\<leadsto>(mksch A B$xs
+                              $(TL$(Dropwhile (%a. a:int A)$schA))
+                              $schB)))
+              )
+          else
+             (if y:act B then
+                 ((Takewhile (%a. a:int B)$schB)
+                       @@ (y\<leadsto>(mksch A B$xs
+                              $schA
+                              $(TL$(Dropwhile (%a. a:int B)$schB))
+                              )))
+             else
+               UU
+             )
+         )
+       ))"
+apply (rule trans)
+apply (rule fix_eq4)
+apply (rule mksch_def)
+apply (rule beta_cfun)
+apply simp
+done
+
+lemma mksch_UU: "mksch A B$UU$schA$schB = UU"
+apply (subst mksch_unfold)
+apply simp
+done
+
+lemma mksch_nil: "mksch A B$nil$schA$schB = nil"
+apply (subst mksch_unfold)
+apply simp
+done
+
+lemma mksch_cons1: "[|x:act A;x~:act B|]
+    ==> mksch A B$(x\<leadsto>tr)$schA$schB =
+          (Takewhile (%a. a:int A)$schA)
+          @@ (x\<leadsto>(mksch A B$tr$(TL$(Dropwhile (%a. a:int A)$schA))
+                              $schB))"
+apply (rule trans)
+apply (subst mksch_unfold)
+apply (simp add: Consq_def If_and_if)
+apply (simp add: Consq_def)
+done
+
+lemma mksch_cons2: "[|x~:act A;x:act B|]
+    ==> mksch A B$(x\<leadsto>tr)$schA$schB =
+         (Takewhile (%a. a:int B)$schB)
+          @@ (x\<leadsto>(mksch A B$tr$schA$(TL$(Dropwhile (%a. a:int B)$schB))
+                             ))"
+apply (rule trans)
+apply (subst mksch_unfold)
+apply (simp add: Consq_def If_and_if)
+apply (simp add: Consq_def)
+done
+
+lemma mksch_cons3: "[|x:act A;x:act B|]
+    ==> mksch A B$(x\<leadsto>tr)$schA$schB =
+             (Takewhile (%a. a:int A)$schA)
+          @@ ((Takewhile (%a. a:int B)$schB)
+          @@ (x\<leadsto>(mksch A B$tr$(TL$(Dropwhile (%a. a:int A)$schA))
+                             $(TL$(Dropwhile (%a. a:int B)$schB))))
+              )"
+apply (rule trans)
+apply (subst mksch_unfold)
+apply (simp add: Consq_def If_and_if)
+apply (simp add: Consq_def)
+done
+
+lemmas compotr_simps = mksch_UU mksch_nil mksch_cons1 mksch_cons2 mksch_cons3
+
+declare compotr_simps [simp]
+
+
+subsection \<open>COMPOSITIONALITY on TRACE Level\<close>
+
+subsubsection "Lemmata for ==>"
+
+(* Consequence out of ext1_ext2_is_not_act1(2), which in turn are consequences out of
+   the compatibility of IOA, in particular out of the condition that internals are
+   really hidden. *)
+
+lemma compatibility_consequence1: "(eB & ~eA --> ~A) -->
+          (A & (eA | eB)) = (eA & A)"
+apply fast
+done
+
+
+(* very similar to above, only the commutativity of | is used to make a slight change *)
+
+lemma compatibility_consequence2: "(eB & ~eA --> ~A) -->
+          (A & (eB | eA)) = (eA & A)"
+apply fast
+done
+
+
+subsubsection "Lemmata for <=="
+
+(* Lemma for substitution of looping assumption in another specific assumption *)
+lemma subst_lemma1: "[| f << (g x) ; x=(h x) |] ==> f << g (h x)"
+by (erule subst)
+
+(* Lemma for substitution of looping assumption in another specific assumption *)
+lemma subst_lemma2: "[| (f x) = y \<leadsto> g; x=(h x) |] ==> (f (h x)) = y \<leadsto> g"
+by (erule subst)
+
+lemma ForallAorB_mksch [rule_format]:
+  "!!A B. compatible A B ==>
+    ! schA schB. Forall (%x. x:act (A\<parallel>B)) tr
+    --> Forall (%x. x:act (A\<parallel>B)) (mksch A B$tr$schA$schB)"
+apply (tactic \<open>Seq_induct_tac @{context} "tr"
+  [@{thm Forall_def}, @{thm sforall_def}, @{thm mksch_def}] 1\<close>)
+apply auto
+apply (simp add: actions_of_par)
+apply (case_tac "a:act A")
+apply (case_tac "a:act B")
+(* a:A, a:B *)
+apply simp
+apply (rule Forall_Conc_impl [THEN mp])
+apply (simp add: intA_is_not_actB int_is_act)
+apply (rule Forall_Conc_impl [THEN mp])
+apply (simp add: intA_is_not_actB int_is_act)
+(* a:A,a~:B *)
+apply simp
+apply (rule Forall_Conc_impl [THEN mp])
+apply (simp add: intA_is_not_actB int_is_act)
+apply (case_tac "a:act B")
+(* a~:A, a:B *)
+apply simp
+apply (rule Forall_Conc_impl [THEN mp])
+apply (simp add: intA_is_not_actB int_is_act)
+(* a~:A,a~:B *)
+apply auto
+done
+
+lemma ForallBnAmksch [rule_format (no_asm)]: "!!A B. compatible B A  ==>
+    ! schA schB.  (Forall (%x. x:act B & x~:act A) tr
+    --> Forall (%x. x:act B & x~:act A) (mksch A B$tr$schA$schB))"
+apply (tactic \<open>Seq_induct_tac @{context} "tr"
+  [@{thm Forall_def}, @{thm sforall_def}, @{thm mksch_def}] 1\<close>)
+apply auto
+apply (rule Forall_Conc_impl [THEN mp])
+apply (simp add: intA_is_not_actB int_is_act)
+done
+
+lemma ForallAnBmksch [rule_format (no_asm)]: "!!A B. compatible A B ==>
+    ! schA schB.  (Forall (%x. x:act A & x~:act B) tr
+    --> Forall (%x. x:act A & x~:act B) (mksch A B$tr$schA$schB))"
+apply (tactic \<open>Seq_induct_tac @{context} "tr"
+  [@{thm Forall_def}, @{thm sforall_def}, @{thm mksch_def}] 1\<close>)
+apply auto
+apply (rule Forall_Conc_impl [THEN mp])
+apply (simp add: intA_is_not_actB int_is_act)
+done
+
+(* safe-tac makes too many case distinctions with this lemma in the next proof *)
+declare FiniteConc [simp del]
+
+lemma FiniteL_mksch [rule_format (no_asm)]: "[| Finite tr; is_asig(asig_of A); is_asig(asig_of B) |] ==>
+    ! x y. Forall (%x. x:act A) x & Forall (%x. x:act B) y &
+           Filter (%a. a:ext A)$x = Filter (%a. a:act A)$tr &
+           Filter (%a. a:ext B)$y = Filter (%a. a:act B)$tr &
+           Forall (%x. x:ext (A\<parallel>B)) tr
+           --> Finite (mksch A B$tr$x$y)"
+
+apply (erule Seq_Finite_ind)
+apply simp
+(* main case *)
+apply simp
+apply auto
+
+(* a: act A; a: act B *)
+apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
+apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
+back
+apply (erule conjE)+
+(* Finite (tw iA x) and Finite (tw iB y) *)
+apply (simp add: not_ext_is_int_or_not_act FiniteConc)
+(* now for conclusion IH applicable, but assumptions have to be transformed *)
+apply (drule_tac x = "x" and g = "Filter (%a. a:act A) $s" in subst_lemma2)
+apply assumption
+apply (drule_tac x = "y" and g = "Filter (%a. a:act B) $s" in subst_lemma2)
+apply assumption
+(* IH *)
+apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
+
+(* a: act B; a~: act A *)
+apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
+
+apply (erule conjE)+
+(* Finite (tw iB y) *)
+apply (simp add: not_ext_is_int_or_not_act FiniteConc)
+(* now for conclusion IH applicable, but assumptions have to be transformed *)
+apply (drule_tac x = "y" and g = "Filter (%a. a:act B) $s" in subst_lemma2)
+apply assumption
+(* IH *)
+apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
+
+(* a~: act B; a: act A *)
+apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
+
+apply (erule conjE)+
+(* Finite (tw iA x) *)
+apply (simp add: not_ext_is_int_or_not_act FiniteConc)
+(* now for conclusion IH applicable, but assumptions have to be transformed *)
+apply (drule_tac x = "x" and g = "Filter (%a. a:act A) $s" in subst_lemma2)
+apply assumption
+(* IH *)
+apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
+
+(* a~: act B; a~: act A *)
+apply (fastforce intro!: ext_is_act simp: externals_of_par)
+done
+
+declare FiniteConc [simp]
+
+declare FilterConc [simp del]
+
+lemma reduceA_mksch1 [rule_format (no_asm)]: " [| Finite bs; is_asig(asig_of A); is_asig(asig_of B);compatible A B|] ==>
+ ! y. Forall (%x. x:act B) y & Forall (%x. x:act B & x~:act A) bs &
+     Filter (%a. a:ext B)$y = Filter (%a. a:act B)$(bs @@ z)
+     --> (? y1 y2.  (mksch A B$(bs @@ z)$x$y) = (y1 @@ (mksch A B$z$x$y2)) &
+                    Forall (%x. x:act B & x~:act A) y1 &
+                    Finite y1 & y = (y1 @@ y2) &
+                    Filter (%a. a:ext B)$y1 = bs)"
+apply (frule_tac A1 = "A" in compat_commute [THEN iffD1])
+apply (erule Seq_Finite_ind)
+apply (rule allI)+
+apply (rule impI)
+apply (rule_tac x = "nil" in exI)
+apply (rule_tac x = "y" in exI)
+apply simp
+(* main case *)
+apply (rule allI)+
+apply (rule impI)
+apply simp
+apply (erule conjE)+
+apply simp
+(* divide_Seq on s *)
+apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
+apply (erule conjE)+
+(* transform assumption f eB y = f B (s@z) *)
+apply (drule_tac x = "y" and g = "Filter (%a. a:act B) $ (s@@z) " in subst_lemma2)
+apply assumption
+apply (simp add: not_ext_is_int_or_not_act FilterConc)
+(* apply IH *)
+apply (erule_tac x = "TL$ (Dropwhile (%a. a:int B) $y) " in allE)
+apply (simp add: ForallTL ForallDropwhile FilterConc)
+apply (erule exE)+
+apply (erule conjE)+
+apply (simp add: FilterConc)
+(* for replacing IH in conclusion *)
+apply (rotate_tac -2)
+(* instantiate y1a and y2a *)
+apply (rule_tac x = "Takewhile (%a. a:int B) $y @@ a\<leadsto>y1" in exI)
+apply (rule_tac x = "y2" in exI)
+(* elminate all obligations up to two depending on Conc_assoc *)
+apply (simp add: intA_is_not_actB int_is_act int_is_not_ext FilterConc)
+apply (simp (no_asm) add: Conc_assoc FilterConc)
+done
+
+lemmas reduceA_mksch = conjI [THEN [6] conjI [THEN [5] reduceA_mksch1]]
+
+lemma reduceB_mksch1 [rule_format]:
+" [| Finite a_s; is_asig(asig_of A); is_asig(asig_of B);compatible A B|] ==>
+ ! x. Forall (%x. x:act A) x & Forall (%x. x:act A & x~:act B) a_s &
+     Filter (%a. a:ext A)$x = Filter (%a. a:act A)$(a_s @@ z)
+     --> (? x1 x2.  (mksch A B$(a_s @@ z)$x$y) = (x1 @@ (mksch A B$z$x2$y)) &
+                    Forall (%x. x:act A & x~:act B) x1 &
+                    Finite x1 & x = (x1 @@ x2) &
+                    Filter (%a. a:ext A)$x1 = a_s)"
+apply (frule_tac A1 = "A" in compat_commute [THEN iffD1])
+apply (erule Seq_Finite_ind)
+apply (rule allI)+
+apply (rule impI)
+apply (rule_tac x = "nil" in exI)
+apply (rule_tac x = "x" in exI)
+apply simp
+(* main case *)
+apply (rule allI)+
+apply (rule impI)
+apply simp
+apply (erule conjE)+
+apply simp
+(* divide_Seq on s *)
+apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
+apply (erule conjE)+
+(* transform assumption f eA x = f A (s@z) *)
+apply (drule_tac x = "x" and g = "Filter (%a. a:act A) $ (s@@z) " in subst_lemma2)
+apply assumption
+apply (simp add: not_ext_is_int_or_not_act FilterConc)
+(* apply IH *)
+apply (erule_tac x = "TL$ (Dropwhile (%a. a:int A) $x) " in allE)
+apply (simp add: ForallTL ForallDropwhile FilterConc)
+apply (erule exE)+
+apply (erule conjE)+
+apply (simp add: FilterConc)
+(* for replacing IH in conclusion *)
+apply (rotate_tac -2)
+(* instantiate y1a and y2a *)
+apply (rule_tac x = "Takewhile (%a. a:int A) $x @@ a\<leadsto>x1" in exI)
+apply (rule_tac x = "x2" in exI)
+(* elminate all obligations up to two depending on Conc_assoc *)
+apply (simp add: intA_is_not_actB int_is_act int_is_not_ext FilterConc)
+apply (simp (no_asm) add: Conc_assoc FilterConc)
+done
+
+lemmas reduceB_mksch = conjI [THEN [6] conjI [THEN [5] reduceB_mksch1]]
+
+declare FilterConc [simp]
+
+
+subsection "Filtering external actions out of mksch(tr,schA,schB) yields the oracle tr"
+
+lemma FilterA_mksch_is_tr:
+"!! A B. [| compatible A B; compatible B A;
+            is_asig(asig_of A); is_asig(asig_of B) |] ==>
+  ! schA schB. Forall (%x. x:act A) schA & Forall (%x. x:act B) schB &
+  Forall (%x. x:ext (A\<parallel>B)) tr &
+  Filter (%a. a:act A)$tr << Filter (%a. a:ext A)$schA &
+  Filter (%a. a:act B)$tr << Filter (%a. a:ext B)$schB
+  --> Filter (%a. a:ext (A\<parallel>B))$(mksch A B$tr$schA$schB) = tr"
+
+apply (tactic \<open>Seq_induct_tac @{context} "tr"
+  [@{thm Forall_def}, @{thm sforall_def}, @{thm mksch_def}] 1\<close>)
+(* main case *)
+(* splitting into 4 cases according to a:A, a:B *)
+apply auto
+
+(* Case a:A, a:B *)
+apply (frule divide_Seq)
+apply (frule divide_Seq)
+back
+apply (erule conjE)+
+(* filtering internals of A in schA and of B in schB is nil *)
+apply (simp add: not_ext_is_int_or_not_act externals_of_par intA_is_not_extB int_is_not_ext)
+(* conclusion of IH ok, but assumptions of IH have to be transformed *)
+apply (drule_tac x = "schA" in subst_lemma1)
+apply assumption
+apply (drule_tac x = "schB" in subst_lemma1)
+apply assumption
+(* IH *)
+apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
+
+(* Case a:A, a~:B *)
+apply (frule divide_Seq)
+apply (erule conjE)+
+(* filtering internals of A is nil *)
+apply (simp add: not_ext_is_int_or_not_act externals_of_par intA_is_not_extB int_is_not_ext)
+apply (drule_tac x = "schA" in subst_lemma1)
+apply assumption
+apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
+
+(* Case a:B, a~:A *)
+apply (frule divide_Seq)
+apply (erule conjE)+
+(* filtering internals of A is nil *)
+apply (simp add: not_ext_is_int_or_not_act externals_of_par intA_is_not_extB int_is_not_ext)
+apply (drule_tac x = "schB" in subst_lemma1)
+back
+apply assumption
+apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
+
+(* Case a~:A, a~:B *)
+apply (fastforce intro!: ext_is_act simp: externals_of_par)
+done
+
+
+subsection" Filter of mksch(tr,schA,schB) to A is schA -- take lemma proof"
+
+lemma FilterAmksch_is_schA: "!! A B. [| compatible A B; compatible B A;
+  is_asig(asig_of A); is_asig(asig_of B) |] ==>
+  Forall (%x. x:ext (A\<parallel>B)) tr &
+  Forall (%x. x:act A) schA & Forall (%x. x:act B) schB &
+  Filter (%a. a:ext A)$schA = Filter (%a. a:act A)$tr &
+  Filter (%a. a:ext B)$schB = Filter (%a. a:act B)$tr &
+  LastActExtsch A schA & LastActExtsch B schB
+  --> Filter (%a. a:act A)$(mksch A B$tr$schA$schB) = schA"
+apply (intro strip)
+apply (rule seq.take_lemma)
+apply (rule mp)
+prefer 2 apply assumption
+back back back back
+apply (rule_tac x = "schA" in spec)
+apply (rule_tac x = "schB" in spec)
+apply (rule_tac x = "tr" in spec)
+apply (tactic "thin_tac' @{context} 5 1")
+apply (rule nat_less_induct)
+apply (rule allI)+
+apply (rename_tac tr schB schA)
+apply (intro strip)
+apply (erule conjE)+
+
+apply (case_tac "Forall (%x. x:act B & x~:act A) tr")
+
+apply (rule seq_take_lemma [THEN iffD2, THEN spec])
+apply (tactic "thin_tac' @{context} 5 1")
+
+
+apply (case_tac "Finite tr")
+
+(* both sides of this equation are nil *)
+apply (subgoal_tac "schA=nil")
+apply (simp (no_asm_simp))
+(* first side: mksch = nil *)
+apply (auto intro!: ForallQFilterPnil ForallBnAmksch FiniteL_mksch)[1]
+(* second side: schA = nil *)
+apply (erule_tac A = "A" in LastActExtimplnil)
+apply (simp (no_asm_simp))
+apply (erule_tac Q = "%x. x:act B & x~:act A" in ForallQFilterPnil)
+apply assumption
+apply fast
+
+(* case ~ Finite s *)
+
+(* both sides of this equation are UU *)
+apply (subgoal_tac "schA=UU")
+apply (simp (no_asm_simp))
+(* first side: mksch = UU *)
+apply (auto intro!: ForallQFilterPUU finiteR_mksch' ForallBnAmksch)[1]
+(* schA = UU *)
+apply (erule_tac A = "A" in LastActExtimplUU)
+apply (simp (no_asm_simp))
+apply (erule_tac Q = "%x. x:act B & x~:act A" in ForallQFilterPUU)
+apply assumption
+apply fast
+
+(* case" ~ Forall (%x.x:act B & x~:act A) s" *)
+
+apply (drule divide_Seq3)
+
+apply (erule exE)+
+apply (erule conjE)+
+apply hypsubst
+
+(* bring in lemma reduceA_mksch *)
+apply (frule_tac x = "schA" and y = "schB" and A = "A" and B = "B" in reduceA_mksch)
+apply assumption+
+apply (erule exE)+
+apply (erule conjE)+
+
+(* use reduceA_mksch to rewrite conclusion *)
+apply hypsubst
+apply simp
+
+(* eliminate the B-only prefix *)
+
+apply (subgoal_tac " (Filter (%a. a :act A) $y1) = nil")
+apply (erule_tac [2] ForallQFilterPnil)
+prefer 2 apply assumption
+prefer 2 apply fast
+
+(* Now real recursive step follows (in y) *)
+
+apply simp
+apply (case_tac "x:act A")
+apply (case_tac "x~:act B")
+apply (rotate_tac -2)
+apply simp
+
+apply (subgoal_tac "Filter (%a. a:act A & a:ext B) $y1=nil")
+apply (rotate_tac -1)
+apply simp
+(* eliminate introduced subgoal 2 *)
+apply (erule_tac [2] ForallQFilterPnil)
+prefer 2 apply assumption
+prefer 2 apply fast
+
+(* bring in divide Seq for s *)
+apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
+apply (erule conjE)+
+
+(* subst divide_Seq in conclusion, but only at the righest occurrence *)
+apply (rule_tac t = "schA" in ssubst)
+back
+back
+back
+apply assumption
+
+(* reduce trace_takes from n to strictly smaller k *)
+apply (rule take_reduction)
+
+(* f A (tw iA) = tw ~eA *)
+apply (simp add: int_is_act not_ext_is_int_or_not_act)
+apply (rule refl)
+apply (simp add: int_is_act not_ext_is_int_or_not_act)
+apply (rotate_tac -11)
+
+(* now conclusion fulfills induction hypothesis, but assumptions are not ready *)
+
+(* assumption Forall tr *)
+(* assumption schB *)
+apply (simp add: ext_and_act)
+(* assumption schA *)
+apply (drule_tac x = "schA" and g = "Filter (%a. a:act A) $rs" in subst_lemma2)
+apply assumption
+apply (simp add: int_is_not_ext)
+(* assumptions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping  *)
+apply (drule_tac sch = "schA" and P = "%a. a:int A" in LastActExtsmall1)
+apply (frule_tac ?sch1.0 = "y1" in LastActExtsmall2)
+apply assumption
+
+(* assumption Forall schA *)
+apply (drule_tac s = "schA" and P = "Forall (%x. x:act A) " in subst)
+apply assumption
+apply (simp add: int_is_act)
+
+(* case x:actions(asig_of A) & x: actions(asig_of B) *)
+
+
+apply (rotate_tac -2)
+apply simp
+
+apply (subgoal_tac "Filter (%a. a:act A & a:ext B) $y1=nil")
+apply (rotate_tac -1)
+apply simp
+(* eliminate introduced subgoal 2 *)
+apply (erule_tac [2] ForallQFilterPnil)
+prefer 2 apply (assumption)
+prefer 2 apply (fast)
+
+(* bring in divide Seq for s *)
+apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
+apply (erule conjE)+
+
+(* subst divide_Seq in conclusion, but only at the rightmost occurrence *)
+apply (rule_tac t = "schA" in ssubst)
+back
+back
+back
+apply assumption
+
+(* f A (tw iA) = tw ~eA *)
+apply (simp add: int_is_act not_ext_is_int_or_not_act)
+
+(* rewrite assumption forall and schB *)
+apply (rotate_tac 13)
+apply (simp add: ext_and_act)
+
+(* divide_Seq for schB2 *)
+apply (frule_tac y = "y2" in sym [THEN eq_imp_below, THEN divide_Seq])
+apply (erule conjE)+
+(* assumption schA *)
+apply (drule_tac x = "schA" and g = "Filter (%a. a:act A) $rs" in subst_lemma2)
+apply assumption
+apply (simp add: int_is_not_ext)
+
+(* f A (tw iB schB2) = nil *)
+apply (simp add: int_is_not_ext not_ext_is_int_or_not_act intA_is_not_actB)
+
+
+(* reduce trace_takes from n to strictly smaller k *)
+apply (rule take_reduction)
+apply (rule refl)
+apply (rule refl)
+
+(* now conclusion fulfills induction hypothesis, but assumptions are not all ready *)
+
+(* assumption schB *)
+apply (drule_tac x = "y2" and g = "Filter (%a. a:act B) $rs" in subst_lemma2)
+apply assumption
+apply (simp add: intA_is_not_actB int_is_not_ext)
+
+(* conclusions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping  *)
+apply (drule_tac sch = "schA" and P = "%a. a:int A" in LastActExtsmall1)
+apply (frule_tac ?sch1.0 = "y1" in LastActExtsmall2)
+apply assumption
+apply (drule_tac sch = "y2" and P = "%a. a:int B" in LastActExtsmall1)
+
+(* assumption Forall schA, and Forall schA are performed by ForallTL,ForallDropwhile *)
+apply (simp add: ForallTL ForallDropwhile)
+
+(* case x~:A & x:B  *)
+(* cannot occur, as just this case has been scheduled out before as the B-only prefix *)
+apply (case_tac "x:act B")
+apply fast
+
+(* case x~:A & x~:B  *)
+(* cannot occur because of assumption: Forall (a:ext A | a:ext B) *)
+apply (rotate_tac -9)
+(* reduce forall assumption from tr to (x\<leadsto>rs) *)
+apply (simp add: externals_of_par)
+apply (fast intro!: ext_is_act)
+
+done
+
+
+
+subsection" Filter of mksch(tr,schA,schB) to B is schB -- take lemma proof"
+
+lemma FilterBmksch_is_schB: "!! A B. [| compatible A B; compatible B A;
+  is_asig(asig_of A); is_asig(asig_of B) |] ==>
+  Forall (%x. x:ext (A\<parallel>B)) tr &
+  Forall (%x. x:act A) schA & Forall (%x. x:act B) schB &
+  Filter (%a. a:ext A)$schA = Filter (%a. a:act A)$tr &
+  Filter (%a. a:ext B)$schB = Filter (%a. a:act B)$tr &
+  LastActExtsch A schA & LastActExtsch B schB
+  --> Filter (%a. a:act B)$(mksch A B$tr$schA$schB) = schB"
+apply (intro strip)
+apply (rule seq.take_lemma)
+apply (rule mp)
+prefer 2 apply assumption
+back back back back
+apply (rule_tac x = "schA" in spec)
+apply (rule_tac x = "schB" in spec)
+apply (rule_tac x = "tr" in spec)
+apply (tactic "thin_tac' @{context} 5 1")
+apply (rule nat_less_induct)
+apply (rule allI)+
+apply (rename_tac tr schB schA)
+apply (intro strip)
+apply (erule conjE)+
+
+apply (case_tac "Forall (%x. x:act A & x~:act B) tr")
+
+apply (rule seq_take_lemma [THEN iffD2, THEN spec])
+apply (tactic "thin_tac' @{context} 5 1")
+
+apply (case_tac "Finite tr")
+
+(* both sides of this equation are nil *)
+apply (subgoal_tac "schB=nil")
+apply (simp (no_asm_simp))
+(* first side: mksch = nil *)
+apply (auto intro!: ForallQFilterPnil ForallAnBmksch FiniteL_mksch)[1]
+(* second side: schA = nil *)
+apply (erule_tac A = "B" in LastActExtimplnil)
+apply (simp (no_asm_simp))
+apply (erule_tac Q = "%x. x:act A & x~:act B" in ForallQFilterPnil)
+apply assumption
+apply fast
+
+(* case ~ Finite tr *)
+
+(* both sides of this equation are UU *)
+apply (subgoal_tac "schB=UU")
+apply (simp (no_asm_simp))
+(* first side: mksch = UU *)
+apply (force intro!: ForallQFilterPUU finiteR_mksch' ForallAnBmksch)
+(* schA = UU *)
+apply (erule_tac A = "B" in LastActExtimplUU)
+apply (simp (no_asm_simp))
+apply (erule_tac Q = "%x. x:act A & x~:act B" in ForallQFilterPUU)
+apply assumption
+apply fast
+
+(* case" ~ Forall (%x.x:act B & x~:act A) s" *)
+
+apply (drule divide_Seq3)
+
+apply (erule exE)+
+apply (erule conjE)+
+apply hypsubst
+
+(* bring in lemma reduceB_mksch *)
+apply (frule_tac y = "schB" and x = "schA" and A = "A" and B = "B" in reduceB_mksch)
+apply assumption+
+apply (erule exE)+
+apply (erule conjE)+
+
+(* use reduceB_mksch to rewrite conclusion *)
+apply hypsubst
+apply simp
+
+(* eliminate the A-only prefix *)
+
+apply (subgoal_tac "(Filter (%a. a :act B) $x1) = nil")
+apply (erule_tac [2] ForallQFilterPnil)
+prefer 2 apply (assumption)
+prefer 2 apply (fast)
+
+(* Now real recursive step follows (in x) *)
+
+apply simp
+apply (case_tac "x:act B")
+apply (case_tac "x~:act A")
+apply (rotate_tac -2)
+apply simp
+
+apply (subgoal_tac "Filter (%a. a:act B & a:ext A) $x1=nil")
+apply (rotate_tac -1)
+apply simp
+(* eliminate introduced subgoal 2 *)
+apply (erule_tac [2] ForallQFilterPnil)
+prefer 2 apply (assumption)
+prefer 2 apply (fast)
+
+(* bring in divide Seq for s *)
+apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
+apply (erule conjE)+
+
+(* subst divide_Seq in conclusion, but only at the rightmost occurrence *)
+apply (rule_tac t = "schB" in ssubst)
+back
+back
+back
+apply assumption
+
+(* reduce trace_takes from n to strictly smaller k *)
+apply (rule take_reduction)
+
+(* f B (tw iB) = tw ~eB *)
+apply (simp add: int_is_act not_ext_is_int_or_not_act)
+apply (rule refl)
+apply (simp add: int_is_act not_ext_is_int_or_not_act)
+apply (rotate_tac -11)
+
+(* now conclusion fulfills induction hypothesis, but assumptions are not ready *)
+
+(* assumption schA *)
+apply (simp add: ext_and_act)
+(* assumption schB *)
+apply (drule_tac x = "schB" and g = "Filter (%a. a:act B) $rs" in subst_lemma2)
+apply assumption
+apply (simp add: int_is_not_ext)
+(* assumptions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping  *)
+apply (drule_tac sch = "schB" and P = "%a. a:int B" in LastActExtsmall1)
+apply (frule_tac ?sch1.0 = "x1" in LastActExtsmall2)
+apply assumption
+
+(* assumption Forall schB *)
+apply (drule_tac s = "schB" and P = "Forall (%x. x:act B) " in subst)
+apply assumption
+apply (simp add: int_is_act)
+
+(* case x:actions(asig_of A) & x: actions(asig_of B) *)
+
+apply (rotate_tac -2)
+apply simp
+
+apply (subgoal_tac "Filter (%a. a:act B & a:ext A) $x1=nil")
+apply (rotate_tac -1)
+apply simp
+(* eliminate introduced subgoal 2 *)
+apply (erule_tac [2] ForallQFilterPnil)
+prefer 2 apply (assumption)
+prefer 2 apply (fast)
+
+(* bring in divide Seq for s *)
+apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
+apply (erule conjE)+
+
+(* subst divide_Seq in conclusion, but only at the rightmost occurrence *)
+apply (rule_tac t = "schB" in ssubst)
+back
+back
+back
+apply assumption
+
+(* f B (tw iB) = tw ~eB *)
+apply (simp add: int_is_act not_ext_is_int_or_not_act)
+
+(* rewrite assumption forall and schB *)
+apply (rotate_tac 13)
+apply (simp add: ext_and_act)
+
+(* divide_Seq for schB2 *)
+apply (frule_tac y = "x2" in sym [THEN eq_imp_below, THEN divide_Seq])
+apply (erule conjE)+
+(* assumption schA *)
+apply (drule_tac x = "schB" and g = "Filter (%a. a:act B) $rs" in subst_lemma2)
+apply assumption
+apply (simp add: int_is_not_ext)
+
+(* f B (tw iA schA2) = nil *)
+apply (simp add: int_is_not_ext not_ext_is_int_or_not_act intA_is_not_actB)
+
+
+(* reduce trace_takes from n to strictly smaller k *)
+apply (rule take_reduction)
+apply (rule refl)
+apply (rule refl)
+
+(* now conclusion fulfills induction hypothesis, but assumptions are not all ready *)
+
+(* assumption schA *)
+apply (drule_tac x = "x2" and g = "Filter (%a. a:act A) $rs" in subst_lemma2)
+apply assumption
+apply (simp add: intA_is_not_actB int_is_not_ext)
+
+(* conclusions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping  *)
+apply (drule_tac sch = "schB" and P = "%a. a:int B" in LastActExtsmall1)
+apply (frule_tac ?sch1.0 = "x1" in LastActExtsmall2)
+apply assumption
+apply (drule_tac sch = "x2" and P = "%a. a:int A" in LastActExtsmall1)
+
+(* assumption Forall schA, and Forall schA are performed by ForallTL,ForallDropwhile *)
+apply (simp add: ForallTL ForallDropwhile)
+
+(* case x~:B & x:A  *)
+(* cannot occur, as just this case has been scheduled out before as the B-only prefix *)
+apply (case_tac "x:act A")
+apply fast
+
+(* case x~:B & x~:A  *)
+(* cannot occur because of assumption: Forall (a:ext A | a:ext B) *)
+apply (rotate_tac -9)
+(* reduce forall assumption from tr to (x\<leadsto>rs) *)
+apply (simp add: externals_of_par)
+apply (fast intro!: ext_is_act)
+
+done
+
+
+subsection "COMPOSITIONALITY on TRACE Level -- Main Theorem"
+
+lemma compositionality_tr:
+"!! A B. [| is_trans_of A; is_trans_of B; compatible A B; compatible B A;
+            is_asig(asig_of A); is_asig(asig_of B)|]
+        ==>  (tr: traces(A\<parallel>B)) =
+             (Filter (%a. a:act A)$tr : traces A &
+              Filter (%a. a:act B)$tr : traces B &
+              Forall (%x. x:ext(A\<parallel>B)) tr)"
+
+apply (simp (no_asm) add: traces_def has_trace_def)
+apply auto
+
+(* ==> *)
+(* There is a schedule of A *)
+apply (rule_tac x = "Filter (%a. a:act A) $sch" in bexI)
+prefer 2
+apply (simp add: compositionality_sch)
+apply (simp add: compatibility_consequence1 externals_of_par ext1_ext2_is_not_act1)
+(* There is a schedule of B *)
+apply (rule_tac x = "Filter (%a. a:act B) $sch" in bexI)
+prefer 2
+apply (simp add: compositionality_sch)
+apply (simp add: compatibility_consequence2 externals_of_par ext1_ext2_is_not_act2)
+(* Traces of A\<parallel>B have only external actions from A or B *)
+apply (rule ForallPFilterP)
+
+(* <== *)
+
+(* replace schA and schB by Cut(schA) and Cut(schB) *)
+apply (drule exists_LastActExtsch)
+apply assumption
+apply (drule exists_LastActExtsch)
+apply assumption
+apply (erule exE)+
+apply (erule conjE)+
+(* Schedules of A(B) have only actions of A(B) *)
+apply (drule scheds_in_sig)
+apply assumption
+apply (drule scheds_in_sig)
+apply assumption
+
+apply (rename_tac h1 h2 schA schB)
+(* mksch is exactly the construction of trA\<parallel>B out of schA, schB, and the oracle tr,
+   we need here *)
+apply (rule_tac x = "mksch A B$tr$schA$schB" in bexI)
+
+(* External actions of mksch are just the oracle *)
+apply (simp add: FilterA_mksch_is_tr)
+
+(* mksch is a schedule -- use compositionality on sch-level *)
+apply (simp add: compositionality_sch)
+apply (simp add: FilterAmksch_is_schA FilterBmksch_is_schB)
+apply (erule ForallAorB_mksch)
+apply (erule ForallPForallQ)
+apply (erule ext_is_act)
+done
+
+
+
+subsection \<open>COMPOSITIONALITY on TRACE Level -- for Modules\<close>
+
+lemma compositionality_tr_modules:
+
+"!! A B. [| is_trans_of A; is_trans_of B; compatible A B; compatible B A;
+            is_asig(asig_of A); is_asig(asig_of B)|]
+ ==> Traces (A\<parallel>B) = par_traces (Traces A) (Traces B)"
+
+apply (unfold Traces_def par_traces_def)
+apply (simp add: asig_of_par)
+apply (rule set_eqI)
+apply (simp add: compositionality_tr externals_of_par)
+done
+
+
+declaration \<open>fn _ => Simplifier.map_ss (Simplifier.set_mksym Simplifier.default_mk_sym)\<close>
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/Compositionality.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,75 @@
+(*  Title:      HOL/HOLCF/IOA/Compositionality.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Compositionality of I/O automata\<close>
+theory Compositionality
+imports CompoTraces
+begin
+
+lemma compatibility_consequence3: "[|eA --> A ; eB & ~eA --> ~A|] ==> (eA | eB) --> A=eA"
+apply auto
+done
+
+
+lemma Filter_actAisFilter_extA: 
+"!! A B. [| compatible A B; Forall (%a. a: ext A | a: ext B) tr |] ==>  
+            Filter (%a. a: act A)$tr= Filter (%a. a: ext A)$tr"
+apply (rule ForallPFilterQR)
+(* i.e.: [| (! x. P x --> (Q x = R x)) ; Forall P tr |] ==> Filter Q$tr = Filter R$tr *)
+prefer 2 apply (assumption)
+apply (rule compatibility_consequence3)
+apply (simp_all add: ext_is_act ext1_ext2_is_not_act1)
+done
+
+
+(* the next two theorems are only necessary, as there is no theorem ext (A\<parallel>B) = ext (B\<parallel>A) *)
+
+lemma compatibility_consequence4: "[|eA --> A ; eB & ~eA --> ~A|] ==> (eB | eA) --> A=eA"
+apply auto
+done
+
+lemma Filter_actAisFilter_extA2: "[| compatible A B; Forall (%a. a: ext B | a: ext A) tr |] ==>  
+            Filter (%a. a: act A)$tr= Filter (%a. a: ext A)$tr"
+apply (rule ForallPFilterQR)
+prefer 2 apply (assumption)
+apply (rule compatibility_consequence4)
+apply (simp_all add: ext_is_act ext1_ext2_is_not_act1)
+done
+
+
+subsection " Main Compositionality Theorem "
+
+lemma compositionality: "[| is_trans_of A1; is_trans_of A2; is_trans_of B1; is_trans_of B2; 
+             is_asig_of A1; is_asig_of A2;  
+             is_asig_of B1; is_asig_of B2;  
+             compatible A1 B1; compatible A2 B2;  
+             A1 =<| A2;  
+             B1 =<| B2 |]  
+         ==> (A1 \<parallel> B1) =<| (A2 \<parallel> B2)"
+apply (simp add: is_asig_of_def)
+apply (frule_tac A1 = "A1" in compat_commute [THEN iffD1])
+apply (frule_tac A1 = "A2" in compat_commute [THEN iffD1])
+apply (simp add: ioa_implements_def inputs_of_par outputs_of_par externals_of_par)
+apply auto
+apply (simp add: compositionality_tr)
+apply (subgoal_tac "ext A1 = ext A2 & ext B1 = ext B2")
+prefer 2
+apply (simp add: externals_def)
+apply (erule conjE)+
+(* rewrite with proven subgoal *)
+apply (simp add: externals_of_par)
+apply auto
+
+(* 2 goals, the 3rd has been solved automatically *)
+(* 1: Filter A2 x : traces A2 *)
+apply (drule_tac A = "traces A1" in subsetD)
+apply assumption
+apply (simp add: Filter_actAisFilter_extA)
+(* 2: Filter B2 x : traces B2 *)
+apply (drule_tac A = "traces B1" in subsetD)
+apply assumption
+apply (simp add: Filter_actAisFilter_extA2)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/Deadlock.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,92 @@
+(*  Title:      HOL/HOLCF/IOA/Deadlock.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Deadlock freedom of I/O Automata\<close>
+
+theory Deadlock
+imports RefCorrectness CompoScheds
+begin
+
+text \<open>input actions may always be added to a schedule\<close>
+
+lemma scheds_input_enabled:
+  "[| Filter (%x. x:act A)$sch : schedules A; a:inp A; input_enabled A; Finite sch|]  
+          ==> Filter (%x. x:act A)$sch @@ a\<leadsto>nil : schedules A"
+apply (simp add: schedules_def has_schedule_def)
+apply auto
+apply (frule inp_is_act)
+apply (simp add: executions_def)
+apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+apply (rename_tac s ex)
+apply (subgoal_tac "Finite ex")
+prefer 2
+apply (simp add: filter_act_def)
+defer
+apply (rule_tac [2] Map2Finite [THEN iffD1])
+apply (rule_tac [2] t = "Map fst$ex" in subst)
+prefer 2 apply (assumption)
+apply (erule_tac [2] FiniteFilter)
+(* subgoal 1 *)
+apply (frule exists_laststate)
+apply (erule allE)
+apply (erule exE)
+(* using input-enabledness *)
+apply (simp add: input_enabled_def)
+apply (erule conjE)+
+apply (erule_tac x = "a" in allE)
+apply simp
+apply (erule_tac x = "u" in allE)
+apply (erule exE)
+(* instantiate execution *)
+apply (rule_tac x = " (s,ex @@ (a,s2) \<leadsto>nil) " in exI)
+apply (simp add: filter_act_def MapConc)
+apply (erule_tac t = "u" in lemma_2_1)
+apply simp
+apply (rule sym)
+apply assumption
+done
+
+text \<open>
+               Deadlock freedom: component B cannot block an out or int action
+                                 of component A in every schedule.
+    Needs compositionality on schedule level, input-enabledness, compatibility
+                    and distributivity of is_exec_frag over @@
+\<close>
+
+declare split_if [split del]
+lemma IOA_deadlock_free: "[| a : local A; Finite sch; sch : schedules (A\<parallel>B);  
+             Filter (%x. x:act A)$(sch @@ a\<leadsto>nil) : schedules A; compatible A B; input_enabled B |]  
+           ==> (sch @@ a\<leadsto>nil) : schedules (A\<parallel>B)"
+apply (simp add: compositionality_sch locals_def)
+apply (rule conjI)
+(* a : act (A\<parallel>B) *)
+prefer 2
+apply (simp add: actions_of_par)
+apply (blast dest: int_is_act out_is_act)
+
+(* Filter B (sch@@[a]) : schedules B *)
+
+apply (case_tac "a:int A")
+apply (drule intA_is_not_actB)
+apply (assumption) (* --> a~:act B *)
+apply simp
+
+(* case a~:int A , i.e. a:out A *)
+apply (case_tac "a~:act B")
+apply simp
+(* case a:act B *)
+apply simp
+apply (subgoal_tac "a:out A")
+prefer 2 apply (blast)
+apply (drule outAactB_is_inpB)
+apply assumption
+apply assumption
+apply (rule scheds_input_enabled)
+apply simp
+apply assumption+
+done
+
+declare split_if [split]
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/IOA.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,11 @@
+(*  Title:      HOL/HOLCF/IOA/IOA.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>The theory of I/O automata in HOLCF\<close>
+
+theory IOA
+imports SimCorrectness Compositionality Deadlock
+begin
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/LiveIOA.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,82 @@
+(*  Title:      HOL/HOLCF/IOA/LiveIOA.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Live I/O automata -- specified by temproal formulas\<close>
+
+theory LiveIOA
+imports TLS
+begin
+
+default_sort type
+
+type_synonym
+  ('a, 's) live_ioa = "('a,'s)ioa * ('a,'s)ioa_temp"
+
+definition
+  validLIOA :: "('a,'s)live_ioa => ('a,'s)ioa_temp  => bool" where
+  "validLIOA AL P = validIOA (fst AL) ((snd AL) \<^bold>\<longrightarrow> P)"
+
+definition
+  WF :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp" where
+  "WF A acts = (\<diamond>\<box>\<langle>%(s,a,t). Enabled A acts s\<rangle> \<^bold>\<longrightarrow> \<box>\<diamond>\<langle>xt2 (plift (%a. a : acts))\<rangle>)"
+definition
+  SF :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp" where
+  "SF A acts = (\<box>\<diamond>\<langle>%(s,a,t). Enabled A acts s\<rangle> \<^bold>\<longrightarrow> \<box>\<diamond>\<langle>xt2 (plift (%a. a : acts))\<rangle>)"
+
+definition
+  liveexecutions :: "('a,'s)live_ioa => ('a,'s)execution set" where
+  "liveexecutions AP = {exec. exec : executions (fst AP) & (exec \<TTurnstile> (snd AP))}"
+definition
+  livetraces :: "('a,'s)live_ioa => 'a trace set" where
+  "livetraces AP = {mk_trace (fst AP)$(snd ex) | ex. ex:liveexecutions AP}"
+definition
+  live_implements :: "('a,'s1)live_ioa => ('a,'s2)live_ioa => bool" where
+  "live_implements CL AM = ((inp (fst CL) = inp (fst AM)) &
+                            (out (fst CL) = out (fst AM)) &
+                            livetraces CL <= livetraces AM)"
+definition
+  is_live_ref_map :: "('s1 => 's2) => ('a,'s1)live_ioa => ('a,'s2)live_ioa => bool" where
+  "is_live_ref_map f CL AM =
+           (is_ref_map f (fst CL ) (fst AM) &
+            (! exec : executions (fst CL). (exec \<TTurnstile> (snd CL)) -->
+                                           ((corresp_ex (fst AM) f exec) \<TTurnstile> (snd AM))))"
+
+
+lemma live_implements_trans:
+"!!LC. [| live_implements (A,LA) (B,LB); live_implements (B,LB) (C,LC) |]
+      ==> live_implements (A,LA) (C,LC)"
+apply (unfold live_implements_def)
+apply auto
+done
+
+
+subsection "Correctness of live refmap"
+
+lemma live_implements: "[| inp(C)=inp(A); out(C)=out(A);
+                   is_live_ref_map f (C,M) (A,L) |]
+                ==> live_implements (C,M) (A,L)"
+apply (simp add: is_live_ref_map_def live_implements_def livetraces_def liveexecutions_def)
+apply auto
+apply (rule_tac x = "corresp_ex A f ex" in exI)
+apply auto
+  (* Traces coincide, Lemma 1 *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (erule lemma_1 [THEN spec, THEN mp])
+  apply (simp (no_asm) add: externals_def)
+  apply (auto)[1]
+  apply (simp add: executions_def reachable.reachable_0)
+
+  (* corresp_ex is execution, Lemma 2 *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (simp add: executions_def)
+  (* start state *)
+  apply (rule conjI)
+  apply (simp add: is_ref_map_def corresp_ex_def)
+  (* is-execution-fragment *)
+  apply (erule lemma_2 [THEN spec, THEN mp])
+  apply (simp add: reachable.reachable_0)
+
+done
+
+end
--- a/src/HOL/HOLCF/IOA/NTP/Abschannel.thy	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Abschannel.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -5,7 +5,7 @@
 section \<open>The (faulty) transmission channel (both directions)\<close>
 
 theory Abschannel
-imports "~~/src/HOL/HOLCF/IOA/meta_theory/IOA" Action
+imports "~~/src/HOL/HOLCF/IOA/IOA" Action
 begin
 
 datatype 'a abs_action = S 'a | R 'a
--- a/src/HOL/HOLCF/IOA/NTP/Receiver.thy	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Receiver.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -5,7 +5,7 @@
 section \<open>The implementation: receiver\<close>
 
 theory Receiver
-imports "~~/src/HOL/HOLCF/IOA/meta_theory/IOA" Action
+imports "~~/src/HOL/HOLCF/IOA/IOA" Action
 begin
 
 type_synonym
--- a/src/HOL/HOLCF/IOA/NTP/Sender.thy	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Sender.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -5,7 +5,7 @@
 section \<open>The implementation: sender\<close>
 
 theory Sender
-imports "~~/src/HOL/HOLCF/IOA/meta_theory/IOA" Action
+imports "~~/src/HOL/HOLCF/IOA/IOA" Action
 begin
 
 type_synonym
--- a/src/HOL/HOLCF/IOA/NTP/Spec.thy	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/HOLCF/IOA/NTP/Spec.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -5,7 +5,7 @@
 section \<open>The specification of reliable transmission\<close>
 
 theory Spec
-imports "~~/src/HOL/HOLCF/IOA/meta_theory/IOA" Action
+imports "~~/src/HOL/HOLCF/IOA/IOA" Action
 begin
 
 definition
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/Pred.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,33 @@
+(*  Title:      HOL/HOLCF/IOA/Pred.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Logical Connectives lifted to predicates\<close>
+
+theory Pred
+imports Main
+begin
+
+default_sort type
+
+type_synonym 'a predicate = "'a \<Rightarrow> bool"
+
+definition satisfies :: "'a \<Rightarrow> 'a predicate \<Rightarrow> bool"  ("_ \<Turnstile> _" [100,9] 8)
+  where "(s \<Turnstile> P) \<longleftrightarrow> P s"
+
+definition valid :: "'a predicate \<Rightarrow> bool"  (*  ("|-") *)
+  where "valid P \<longleftrightarrow> (\<forall>s. (s \<Turnstile> P))"
+
+definition NOT :: "'a predicate \<Rightarrow> 'a predicate"  ("\<^bold>\<not> _" [40] 40)
+  where "NOT P s \<longleftrightarrow> ~ (P s)"
+
+definition AND :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<and>" 35)
+  where "(P \<^bold>\<and> Q) s \<longleftrightarrow> P s \<and> Q s"
+
+definition OR :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<or>" 30)
+  where "(P \<^bold>\<or> Q) s \<longleftrightarrow> P s \<or> Q s"
+
+definition IMPLIES :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<longrightarrow>" 25)
+  where "(P \<^bold>\<longrightarrow> Q) s \<longleftrightarrow> P s \<longrightarrow> Q s"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/RefCorrectness.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,370 @@
+(*  Title:      HOL/HOLCF/IOA/RefCorrectness.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Correctness of Refinement Mappings in HOLCF/IOA\<close>
+
+theory RefCorrectness
+imports RefMappings
+begin
+
+definition
+  corresp_exC :: "('a,'s2)ioa => ('s1 => 's2) => ('a,'s1)pairs
+                   -> ('s1 => ('a,'s2)pairs)" where
+  "corresp_exC A f = (fix$(LAM h ex. (%s. case ex of
+      nil =>  nil
+    | x##xs => (flift1 (%pr. (@cex. move A cex (f s) (fst pr) (f (snd pr)))
+                              @@ ((h$xs) (snd pr)))
+                        $x) )))"
+definition
+  corresp_ex :: "('a,'s2)ioa => ('s1 => 's2) =>
+                  ('a,'s1)execution => ('a,'s2)execution" where
+  "corresp_ex A f ex = (f (fst ex),(corresp_exC A f$(snd ex)) (fst ex))"
+
+definition
+  is_fair_ref_map :: "('s1 => 's2) => ('a,'s1)ioa => ('a,'s2)ioa => bool" where
+  "is_fair_ref_map f C A =
+      (is_ref_map f C A &
+       (! ex : executions(C). fair_ex C ex --> fair_ex A (corresp_ex A f ex)))"
+
+(* Axioms for fair trace inclusion proof support, not for the correctness proof
+   of refinement mappings!
+   Note: Everything is superseded by LiveIOA.thy! *)
+
+axiomatization where
+corresp_laststate:
+  "Finite ex ==> laststate (corresp_ex A f (s,ex)) = f (laststate (s,ex))"
+
+axiomatization where
+corresp_Finite:
+  "Finite (snd (corresp_ex A f (s,ex))) = Finite ex"
+
+axiomatization where
+FromAtoC:
+  "fin_often (%x. P (snd x)) (snd (corresp_ex A f (s,ex))) ==> fin_often (%y. P (f (snd y))) ex"
+
+axiomatization where
+FromCtoA:
+  "inf_often (%y. P (fst y)) ex ==> inf_often (%x. P (fst x)) (snd (corresp_ex A f (s,ex)))"
+
+
+(* Proof by case on inf W in ex: If so, ok. If not, only fin W in ex, ie there is
+   an index i from which on no W in ex. But W inf enabled, ie at least once after i
+   W is enabled. As W does not occur after i and W is enabling_persistent, W keeps
+   enabled until infinity, ie. indefinitely *)
+axiomatization where
+persistent:
+  "[|inf_often (%x. Enabled A W (snd x)) ex; en_persistent A W|]
+   ==> inf_often (%x. fst x :W) ex | fin_often (%x. ~Enabled A W (snd x)) ex"
+
+axiomatization where
+infpostcond:
+  "[| is_exec_frag A (s,ex); inf_often (%x. fst x:W) ex|]
+    ==> inf_often (% x. set_was_enabled A W (snd x)) ex"
+
+
+subsection "corresp_ex"
+
+lemma corresp_exC_unfold: "corresp_exC A f  = (LAM ex. (%s. case ex of
+       nil =>  nil
+     | x##xs => (flift1 (%pr. (@cex. move A cex (f s) (fst pr) (f (snd pr)))
+                               @@ ((corresp_exC A f $xs) (snd pr)))
+                         $x) ))"
+apply (rule trans)
+apply (rule fix_eq2)
+apply (simp only: corresp_exC_def)
+apply (rule beta_cfun)
+apply (simp add: flift1_def)
+done
+
+lemma corresp_exC_UU: "(corresp_exC A f$UU) s=UU"
+apply (subst corresp_exC_unfold)
+apply simp
+done
+
+lemma corresp_exC_nil: "(corresp_exC A f$nil) s = nil"
+apply (subst corresp_exC_unfold)
+apply simp
+done
+
+lemma corresp_exC_cons: "(corresp_exC A f$(at\<leadsto>xs)) s =
+           (@cex. move A cex (f s) (fst at) (f (snd at)))
+           @@ ((corresp_exC A f$xs) (snd at))"
+apply (rule trans)
+apply (subst corresp_exC_unfold)
+apply (simp add: Consq_def flift1_def)
+apply simp
+done
+
+
+declare corresp_exC_UU [simp] corresp_exC_nil [simp] corresp_exC_cons [simp]
+
+
+
+subsection "properties of move"
+
+lemma move_is_move:
+   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
+      move A (@x. move A x (f s) a (f t)) (f s) a (f t)"
+apply (unfold is_ref_map_def)
+apply (subgoal_tac "? ex. move A ex (f s) a (f t) ")
+prefer 2
+apply simp
+apply (erule exE)
+apply (rule someI)
+apply assumption
+done
+
+lemma move_subprop1:
+   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
+     is_exec_frag A (f s,@x. move A x (f s) a (f t))"
+apply (cut_tac move_is_move)
+defer
+apply assumption+
+apply (simp add: move_def)
+done
+
+lemma move_subprop2:
+   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
+     Finite ((@x. move A x (f s) a (f t)))"
+apply (cut_tac move_is_move)
+defer
+apply assumption+
+apply (simp add: move_def)
+done
+
+lemma move_subprop3:
+   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
+     laststate (f s,@x. move A x (f s) a (f t)) = (f t)"
+apply (cut_tac move_is_move)
+defer
+apply assumption+
+apply (simp add: move_def)
+done
+
+lemma move_subprop4:
+   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
+      mk_trace A$((@x. move A x (f s) a (f t))) =
+        (if a:ext A then a\<leadsto>nil else nil)"
+apply (cut_tac move_is_move)
+defer
+apply assumption+
+apply (simp add: move_def)
+done
+
+
+(* ------------------------------------------------------------------ *)
+(*                   The following lemmata contribute to              *)
+(*                 TRACE INCLUSION Part 1: Traces coincide            *)
+(* ------------------------------------------------------------------ *)
+
+section "Lemmata for <=="
+
+(* --------------------------------------------------- *)
+(*   Lemma 1.1: Distribution of mk_trace and @@        *)
+(* --------------------------------------------------- *)
+
+lemma mk_traceConc: "mk_trace C$(ex1 @@ ex2)= (mk_trace C$ex1) @@ (mk_trace C$ex2)"
+apply (simp add: mk_trace_def filter_act_def MapConc)
+done
+
+
+
+(* ------------------------------------------------------
+                 Lemma 1 :Traces coincide
+   ------------------------------------------------------- *)
+declare split_if [split del]
+
+lemma lemma_1:
+  "[|is_ref_map f C A; ext C = ext A|] ==>
+         !s. reachable C s & is_exec_frag C (s,xs) -->
+             mk_trace C$xs = mk_trace A$(snd (corresp_ex A f (s,xs)))"
+apply (unfold corresp_ex_def)
+apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1\<close>)
+(* cons case *)
+apply (auto simp add: mk_traceConc)
+apply (frule reachable.reachable_n)
+apply assumption
+apply (auto simp add: move_subprop4 split add: split_if) 
+done
+
+declare split_if [split]
+
+(* ------------------------------------------------------------------ *)
+(*                   The following lemmata contribute to              *)
+(*              TRACE INCLUSION Part 2: corresp_ex is execution       *)
+(* ------------------------------------------------------------------ *)
+
+section "Lemmata for ==>"
+
+(* -------------------------------------------------- *)
+(*                   Lemma 2.1                        *)
+(* -------------------------------------------------- *)
+
+lemma lemma_2_1 [rule_format (no_asm)]:
+"Finite xs -->
+ (!s .is_exec_frag A (s,xs) & is_exec_frag A (t,ys) &
+      t = laststate (s,xs)
+  --> is_exec_frag A (s,xs @@ ys))"
+
+apply (rule impI)
+apply (tactic \<open>Seq_Finite_induct_tac @{context} 1\<close>)
+(* main case *)
+apply (auto simp add: split_paired_all)
+done
+
+
+(* ----------------------------------------------------------- *)
+(*               Lemma 2 : corresp_ex is execution             *)
+(* ----------------------------------------------------------- *)
+
+
+
+lemma lemma_2:
+ "[| is_ref_map f C A |] ==>
+  !s. reachable C s & is_exec_frag C (s,xs)
+  --> is_exec_frag A (corresp_ex A f (s,xs))"
+
+apply (unfold corresp_ex_def)
+
+apply simp
+apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1\<close>)
+(* main case *)
+apply auto
+apply (rule_tac t = "f x2" in lemma_2_1)
+
+(* Finite *)
+apply (erule move_subprop2)
+apply assumption+
+apply (rule conjI)
+
+(* is_exec_frag *)
+apply (erule move_subprop1)
+apply assumption+
+apply (rule conjI)
+
+(* Induction hypothesis  *)
+(* reachable_n looping, therefore apply it manually *)
+apply (erule_tac x = "x2" in allE)
+apply simp
+apply (frule reachable.reachable_n)
+apply assumption
+apply simp
+(* laststate *)
+apply (erule move_subprop3 [symmetric])
+apply assumption+
+done
+
+
+subsection "Main Theorem: TRACE - INCLUSION"
+
+lemma trace_inclusion:
+  "[| ext C = ext A; is_ref_map f C A |]
+           ==> traces C <= traces A"
+
+  apply (unfold traces_def)
+
+  apply (simp (no_asm) add: has_trace_def2)
+  apply auto
+
+  (* give execution of abstract automata *)
+  apply (rule_tac x = "corresp_ex A f ex" in bexI)
+
+  (* Traces coincide, Lemma 1 *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (erule lemma_1 [THEN spec, THEN mp])
+  apply assumption+
+  apply (simp add: executions_def reachable.reachable_0)
+
+  (* corresp_ex is execution, Lemma 2 *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (simp add: executions_def)
+  (* start state *)
+  apply (rule conjI)
+  apply (simp add: is_ref_map_def corresp_ex_def)
+  (* is-execution-fragment *)
+  apply (erule lemma_2 [THEN spec, THEN mp])
+  apply (simp add: reachable.reachable_0)
+  done
+
+
+subsection "Corollary:  FAIR TRACE - INCLUSION"
+
+lemma fininf: "(~inf_often P s) = fin_often P s"
+apply (unfold fin_often_def)
+apply auto
+done
+
+
+lemma WF_alt: "is_wfair A W (s,ex) =
+  (fin_often (%x. ~Enabled A W (snd x)) ex --> inf_often (%x. fst x :W) ex)"
+apply (simp add: is_wfair_def fin_often_def)
+apply auto
+done
+
+lemma WF_persistent: "[|is_wfair A W (s,ex); inf_often (%x. Enabled A W (snd x)) ex;
+          en_persistent A W|]
+    ==> inf_often (%x. fst x :W) ex"
+apply (drule persistent)
+apply assumption
+apply (simp add: WF_alt)
+apply auto
+done
+
+
+lemma fair_trace_inclusion: "!! C A.
+          [| is_ref_map f C A; ext C = ext A;
+          !! ex. [| ex:executions C; fair_ex C ex|] ==> fair_ex A (corresp_ex A f ex) |]
+          ==> fairtraces C <= fairtraces A"
+apply (simp (no_asm) add: fairtraces_def fairexecutions_def)
+apply auto
+apply (rule_tac x = "corresp_ex A f ex" in exI)
+apply auto
+
+  (* Traces coincide, Lemma 1 *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (erule lemma_1 [THEN spec, THEN mp])
+  apply assumption+
+  apply (simp add: executions_def reachable.reachable_0)
+
+  (* corresp_ex is execution, Lemma 2 *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (simp add: executions_def)
+  (* start state *)
+  apply (rule conjI)
+  apply (simp add: is_ref_map_def corresp_ex_def)
+  (* is-execution-fragment *)
+  apply (erule lemma_2 [THEN spec, THEN mp])
+  apply (simp add: reachable.reachable_0)
+
+done
+
+lemma fair_trace_inclusion2: "!! C A.
+          [| inp(C) = inp(A); out(C)=out(A);
+             is_fair_ref_map f C A |]
+          ==> fair_implements C A"
+apply (simp add: is_fair_ref_map_def fair_implements_def fairtraces_def fairexecutions_def)
+apply auto
+apply (rule_tac x = "corresp_ex A f ex" in exI)
+apply auto
+
+  (* Traces coincide, Lemma 1 *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (erule lemma_1 [THEN spec, THEN mp])
+  apply (simp (no_asm) add: externals_def)
+  apply (auto)[1]
+  apply (simp add: executions_def reachable.reachable_0)
+
+  (* corresp_ex is execution, Lemma 2 *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (simp add: executions_def)
+  (* start state *)
+  apply (rule conjI)
+  apply (simp add: is_ref_map_def corresp_ex_def)
+  (* is-execution-fragment *)
+  apply (erule lemma_2 [THEN spec, THEN mp])
+  apply (simp add: reachable.reachable_0)
+
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/RefMappings.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,129 @@
+(*  Title:      HOL/HOLCF/IOA/RefMappings.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Refinement Mappings in HOLCF/IOA\<close>
+
+theory RefMappings
+imports Traces
+begin
+
+default_sort type
+
+definition
+  move :: "[('a,'s)ioa,('a,'s)pairs,'s,'a,'s] => bool" where
+  "move ioa ex s a t =
+    (is_exec_frag ioa (s,ex) &  Finite ex &
+     laststate (s,ex)=t  &
+     mk_trace ioa$ex = (if a:ext(ioa) then a\<leadsto>nil else nil))"
+
+definition
+  is_ref_map :: "[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool" where
+  "is_ref_map f C A =
+   ((!s:starts_of(C). f(s):starts_of(A)) &
+   (!s t a. reachable C s &
+            s \<midarrow>a\<midarrow>C\<rightarrow> t
+            --> (? ex. move A ex (f s) a (f t))))"
+
+definition
+  is_weak_ref_map :: "[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool" where
+  "is_weak_ref_map f C A =
+   ((!s:starts_of(C). f(s):starts_of(A)) &
+   (!s t a. reachable C s &
+            s \<midarrow>a\<midarrow>C\<rightarrow> t
+            --> (if a:ext(C)
+                 then (f s) \<midarrow>a\<midarrow>A\<rightarrow> (f t)
+                 else (f s)=(f t))))"
+
+
+subsection "transitions and moves"
+
+
+lemma transition_is_ex: "s \<midarrow>a\<midarrow>A\<rightarrow> t ==> ? ex. move A ex s a t"
+apply (rule_tac x = " (a,t) \<leadsto>nil" in exI)
+apply (simp add: move_def)
+done
+
+
+lemma nothing_is_ex: "(~a:ext A) & s=t ==> ? ex. move A ex s a t"
+apply (rule_tac x = "nil" in exI)
+apply (simp add: move_def)
+done
+
+
+lemma ei_transitions_are_ex: "(s \<midarrow>a\<midarrow>A\<rightarrow> s') & (s' \<midarrow>a'\<midarrow>A\<rightarrow> s'') & (~a':ext A)  
+         ==> ? ex. move A ex s a s''"
+apply (rule_tac x = " (a,s') \<leadsto> (a',s'') \<leadsto>nil" in exI)
+apply (simp add: move_def)
+done
+
+
+lemma eii_transitions_are_ex: "(s1 \<midarrow>a1\<midarrow>A\<rightarrow> s2) & (s2 \<midarrow>a2\<midarrow>A\<rightarrow> s3) & (s3 \<midarrow>a3\<midarrow>A\<rightarrow> s4) & 
+      (~a2:ext A) & (~a3:ext A) ==>  
+      ? ex. move A ex s1 a1 s4"
+apply (rule_tac x = " (a1,s2) \<leadsto> (a2,s3) \<leadsto> (a3,s4) \<leadsto>nil" in exI)
+apply (simp add: move_def)
+done
+
+
+subsection "weak_ref_map and ref_map"
+
+lemma weak_ref_map2ref_map:
+  "[| ext C = ext A;  
+     is_weak_ref_map f C A |] ==> is_ref_map f C A"
+apply (unfold is_weak_ref_map_def is_ref_map_def)
+apply auto
+apply (case_tac "a:ext A")
+apply (auto intro: transition_is_ex nothing_is_ex)
+done
+
+
+lemma imp_conj_lemma: "(P ==> Q-->R) ==> P&Q --> R"
+  by blast
+
+declare split_if [split del]
+declare if_weak_cong [cong del]
+
+lemma rename_through_pmap: "[| is_weak_ref_map f C A |]  
+      ==> (is_weak_ref_map f (rename C g) (rename A g))"
+apply (simp add: is_weak_ref_map_def)
+apply (rule conjI)
+(* 1: start states *)
+apply (simp add: rename_def rename_set_def starts_of_def)
+(* 2: reachable transitions *)
+apply (rule allI)+
+apply (rule imp_conj_lemma)
+apply (simp (no_asm) add: rename_def rename_set_def)
+apply (simp add: externals_def asig_inputs_def asig_outputs_def asig_of_def trans_of_def)
+apply safe
+apply (simplesubst split_if)
+ apply (rule conjI)
+ apply (rule impI)
+ apply (erule disjE)
+ apply (erule exE)
+apply (erule conjE)
+(* x is input *)
+ apply (drule sym)
+ apply (drule sym)
+apply simp
+apply hypsubst+
+apply (frule reachable_rename)
+apply simp
+(* x is output *)
+ apply (erule exE)
+apply (erule conjE)
+ apply (drule sym)
+ apply (drule sym)
+apply simp
+apply hypsubst+
+apply (frule reachable_rename)
+apply simp
+(* x is internal *)
+apply (frule reachable_rename)
+apply auto
+done
+
+declare split_if [split]
+declare if_weak_cong [cong]
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/Seq.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,328 @@
+(*  Title:      HOL/HOLCF/IOA/Seq.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Partial, Finite and Infinite Sequences (lazy lists), modeled as domain\<close>
+
+theory Seq
+imports "../../HOLCF"
+begin
+
+default_sort pcpo
+
+domain (unsafe) 'a seq = nil  ("nil") | cons (HD :: 'a) (lazy TL :: "'a seq")  (infixr "##" 65)
+
+(*
+   sfilter       :: "('a -> tr) -> 'a seq -> 'a seq"
+   smap          :: "('a -> 'b) -> 'a seq -> 'b seq"
+   sforall       :: "('a -> tr) => 'a seq => bool"
+   sforall2      :: "('a -> tr) -> 'a seq -> tr"
+   slast         :: "'a seq     -> 'a"
+   sconc         :: "'a seq     -> 'a seq -> 'a seq"
+   sdropwhile    :: "('a -> tr)  -> 'a seq -> 'a seq"
+   stakewhile    :: "('a -> tr)  -> 'a seq -> 'a seq"
+   szip          :: "'a seq      -> 'b seq -> ('a*'b) seq"
+   sflat        :: "('a seq) seq  -> 'a seq"
+
+   sfinite       :: "'a seq set"
+   Partial       :: "'a seq => bool"
+   Infinite      :: "'a seq => bool"
+
+   nproj        :: "nat => 'a seq => 'a"
+   sproj        :: "nat => 'a seq => 'a seq"
+*)
+
+inductive
+  Finite :: "'a seq => bool"
+  where
+    sfinite_0:  "Finite nil"
+  | sfinite_n:  "[| Finite tr; a~=UU |] ==> Finite (a##tr)"
+
+declare Finite.intros [simp]
+
+definition
+  Partial :: "'a seq => bool"
+where
+  "Partial x  == (seq_finite x) & ~(Finite x)"
+
+definition
+  Infinite :: "'a seq => bool"
+where
+  "Infinite x == ~(seq_finite x)"
+
+
+subsection \<open>recursive equations of operators\<close>
+
+subsubsection \<open>smap\<close>
+
+fixrec
+  smap :: "('a -> 'b) -> 'a seq -> 'b seq"
+where
+  smap_nil: "smap$f$nil=nil"
+| smap_cons: "[|x~=UU|] ==> smap$f$(x##xs)= (f$x)##smap$f$xs"
+
+lemma smap_UU [simp]: "smap$f$UU=UU"
+by fixrec_simp
+
+subsubsection \<open>sfilter\<close>
+
+fixrec
+   sfilter :: "('a -> tr) -> 'a seq -> 'a seq"
+where
+  sfilter_nil: "sfilter$P$nil=nil"
+| sfilter_cons:
+    "x~=UU ==> sfilter$P$(x##xs)=
+              (If P$x then x##(sfilter$P$xs) else sfilter$P$xs)"
+
+lemma sfilter_UU [simp]: "sfilter$P$UU=UU"
+by fixrec_simp
+
+subsubsection \<open>sforall2\<close>
+
+fixrec
+  sforall2 :: "('a -> tr) -> 'a seq -> tr"
+where
+  sforall2_nil: "sforall2$P$nil=TT"
+| sforall2_cons:
+    "x~=UU ==> sforall2$P$(x##xs)= ((P$x) andalso sforall2$P$xs)"
+
+lemma sforall2_UU [simp]: "sforall2$P$UU=UU"
+by fixrec_simp
+
+definition
+  sforall_def: "sforall P t == (sforall2$P$t ~=FF)"
+
+subsubsection \<open>stakewhile\<close>
+
+fixrec
+  stakewhile :: "('a -> tr)  -> 'a seq -> 'a seq"
+where
+  stakewhile_nil: "stakewhile$P$nil=nil"
+| stakewhile_cons:
+    "x~=UU ==> stakewhile$P$(x##xs) =
+              (If P$x then x##(stakewhile$P$xs) else nil)"
+
+lemma stakewhile_UU [simp]: "stakewhile$P$UU=UU"
+by fixrec_simp
+
+subsubsection \<open>sdropwhile\<close>
+
+fixrec
+  sdropwhile :: "('a -> tr) -> 'a seq -> 'a seq"
+where
+  sdropwhile_nil: "sdropwhile$P$nil=nil"
+| sdropwhile_cons:
+    "x~=UU ==> sdropwhile$P$(x##xs) =
+              (If P$x then sdropwhile$P$xs else x##xs)"
+
+lemma sdropwhile_UU [simp]: "sdropwhile$P$UU=UU"
+by fixrec_simp
+
+subsubsection \<open>slast\<close>
+
+fixrec
+  slast :: "'a seq -> 'a"
+where
+  slast_nil: "slast$nil=UU"
+| slast_cons:
+    "x~=UU ==> slast$(x##xs)= (If is_nil$xs then x else slast$xs)"
+
+lemma slast_UU [simp]: "slast$UU=UU"
+by fixrec_simp
+
+subsubsection \<open>sconc\<close>
+
+fixrec
+  sconc :: "'a seq -> 'a seq -> 'a seq"
+where
+  sconc_nil: "sconc$nil$y = y"
+| sconc_cons':
+    "x~=UU ==> sconc$(x##xs)$y = x##(sconc$xs$y)"
+
+abbreviation
+  sconc_syn :: "'a seq => 'a seq => 'a seq"  (infixr "@@" 65) where
+  "xs @@ ys == sconc $ xs $ ys"
+
+lemma sconc_UU [simp]: "UU @@ y=UU"
+by fixrec_simp
+
+lemma sconc_cons [simp]: "(x##xs) @@ y=x##(xs @@ y)"
+apply (cases "x=UU")
+apply simp_all
+done
+
+declare sconc_cons' [simp del]
+
+subsubsection \<open>sflat\<close>
+
+fixrec
+  sflat :: "('a seq) seq -> 'a seq"
+where
+  sflat_nil: "sflat$nil=nil"
+| sflat_cons': "x~=UU ==> sflat$(x##xs)= x@@(sflat$xs)"
+
+lemma sflat_UU [simp]: "sflat$UU=UU"
+by fixrec_simp
+
+lemma sflat_cons [simp]: "sflat$(x##xs)= x@@(sflat$xs)"
+by (cases "x=UU", simp_all)
+
+declare sflat_cons' [simp del]
+
+subsubsection \<open>szip\<close>
+
+fixrec
+  szip :: "'a seq -> 'b seq -> ('a*'b) seq"
+where
+  szip_nil: "szip$nil$y=nil"
+| szip_cons_nil: "x~=UU ==> szip$(x##xs)$nil=UU"
+| szip_cons:
+    "[| x~=UU; y~=UU|] ==> szip$(x##xs)$(y##ys) = (x,y)##szip$xs$ys"
+
+lemma szip_UU1 [simp]: "szip$UU$y=UU"
+by fixrec_simp
+
+lemma szip_UU2 [simp]: "x~=nil ==> szip$x$UU=UU"
+by (cases x, simp_all, fixrec_simp)
+
+
+subsection "scons, nil"
+
+lemma scons_inject_eq:
+ "[|x~=UU;y~=UU|]==> (x##xs=y##ys) = (x=y & xs=ys)"
+by simp
+
+lemma nil_less_is_nil: "nil<<x ==> nil=x"
+apply (cases x)
+apply simp
+apply simp
+apply simp
+done
+
+subsection "sfilter, sforall, sconc"
+
+lemma if_and_sconc [simp]: "(if b then tr1 else tr2) @@ tr
+        = (if b then tr1 @@ tr else tr2 @@ tr)"
+by simp
+
+
+lemma sfiltersconc: "sfilter$P$(x @@ y) = (sfilter$P$x @@ sfilter$P$y)"
+apply (induct x)
+(* adm *)
+apply simp
+(* base cases *)
+apply simp
+apply simp
+(* main case *)
+apply (rule_tac p="P$a" in trE)
+apply simp
+apply simp
+apply simp
+done
+
+lemma sforallPstakewhileP: "sforall P (stakewhile$P$x)"
+apply (simp add: sforall_def)
+apply (induct x)
+(* adm *)
+apply simp
+(* base cases *)
+apply simp
+apply simp
+(* main case *)
+apply (rule_tac p="P$a" in trE)
+apply simp
+apply simp
+apply simp
+done
+
+lemma forallPsfilterP: "sforall P (sfilter$P$x)"
+apply (simp add: sforall_def)
+apply (induct x)
+(* adm *)
+apply simp
+(* base cases *)
+apply simp
+apply simp
+(* main case *)
+apply (rule_tac p="P$a" in trE)
+apply simp
+apply simp
+apply simp
+done
+
+
+subsection "Finite"
+
+(* ----------------------------------------------------  *)
+(* Proofs of rewrite rules for Finite:                  *)
+(* 1. Finite(nil),   (by definition)                    *)
+(* 2. ~Finite(UU),                                      *)
+(* 3. a~=UU==> Finite(a##x)=Finite(x)                  *)
+(* ----------------------------------------------------  *)
+
+lemma Finite_UU_a: "Finite x --> x~=UU"
+apply (rule impI)
+apply (erule Finite.induct)
+ apply simp
+apply simp
+done
+
+lemma Finite_UU [simp]: "~(Finite UU)"
+apply (cut_tac x="UU" in Finite_UU_a)
+apply fast
+done
+
+lemma Finite_cons_a: "Finite x --> a~=UU --> x=a##xs --> Finite xs"
+apply (intro strip)
+apply (erule Finite.cases)
+apply fastforce
+apply simp
+done
+
+lemma Finite_cons: "a~=UU ==>(Finite (a##x)) = (Finite x)"
+apply (rule iffI)
+apply (erule (1) Finite_cons_a [rule_format])
+apply fast
+apply simp
+done
+
+lemma Finite_upward: "\<lbrakk>Finite x; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> Finite y"
+apply (induct arbitrary: y set: Finite)
+apply (case_tac y, simp, simp, simp)
+apply (case_tac y, simp, simp)
+apply simp
+done
+
+lemma adm_Finite [simp]: "adm Finite"
+by (rule adm_upward, rule Finite_upward)
+
+
+subsection "induction"
+
+
+(*--------------------------------   *)
+(* Extensions to Induction Theorems  *)
+(*--------------------------------   *)
+
+
+lemma seq_finite_ind_lemma:
+  assumes "(!!n. P(seq_take n$s))"
+  shows "seq_finite(s) -->P(s)"
+apply (unfold seq.finite_def)
+apply (intro strip)
+apply (erule exE)
+apply (erule subst)
+apply (rule assms)
+done
+
+
+lemma seq_finite_ind: "!!P.[|P(UU);P(nil);
+   !! x s1.[|x~=UU;P(s1)|] ==> P(x##s1)
+   |] ==> seq_finite(s) --> P(s)"
+apply (rule seq_finite_ind_lemma)
+apply (erule seq.finite_induct)
+ apply assumption
+apply simp
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/Sequence.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,1084 @@
+(*  Title:      HOL/HOLCF/IOA/Sequence.thy
+    Author:     Olaf Müller
+
+Sequences over flat domains with lifted elements.
+*)
+
+theory Sequence
+imports Seq
+begin
+
+default_sort type
+
+type_synonym 'a Seq = "'a lift seq"
+
+definition Consq :: "'a \<Rightarrow> 'a Seq \<rightarrow> 'a Seq"
+  where "Consq a = (LAM s. Def a ## s)"
+
+definition Filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq"
+  where "Filter P = sfilter $ (flift2 P)"
+
+definition Map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a Seq \<rightarrow> 'b Seq"
+  where "Map f = smap $ (flift2 f)"
+
+definition Forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<Rightarrow> bool"
+  where "Forall P = sforall (flift2 P)"
+
+definition Last :: "'a Seq \<rightarrow> 'a lift"
+  where "Last = slast"
+
+definition Dropwhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq"
+  where "Dropwhile P = sdropwhile $ (flift2 P)"
+
+definition Takewhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq"
+  where "Takewhile P = stakewhile $ (flift2 P)"
+
+definition Zip :: "'a Seq \<rightarrow> 'b Seq \<rightarrow> ('a * 'b) Seq"
+where
+  "Zip = (fix$(LAM h t1 t2. case t1 of
+               nil   => nil
+             | x##xs => (case t2 of
+                          nil => UU
+                        | y##ys => (case x of
+                                      UU  => UU
+                                    | Def a => (case y of
+                                                  UU => UU
+                                                | Def b => Def (a,b)##(h$xs$ys))))))"
+
+definition Flat :: "'a Seq seq \<rightarrow> 'a Seq"
+  where "Flat = sflat"
+
+definition Filter2 :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq"
+where
+  "Filter2 P = (fix $ (LAM h t. case t of
+            nil \<Rightarrow> nil
+          | x##xs \<Rightarrow> (case x of UU \<Rightarrow> UU | Def y \<Rightarrow> (if P y
+                     then x##(h$xs)
+                     else     h$xs))))"
+
+abbreviation Consq_syn  ("(_/\<leadsto>_)" [66,65] 65)
+  where "a\<leadsto>s \<equiv> Consq a $ s"
+
+
+text \<open>List enumeration\<close>
+syntax
+  "_totlist"      :: "args => 'a Seq"              ("[(_)!]")
+  "_partlist"     :: "args => 'a Seq"              ("[(_)?]")
+translations
+  "[x, xs!]"     == "x\<leadsto>[xs!]"
+  "[x!]"         == "x\<leadsto>nil"
+  "[x, xs?]"     == "x\<leadsto>[xs?]"
+  "[x?]"         == "x\<leadsto>CONST bottom"
+
+
+declare andalso_and [simp]
+declare andalso_or [simp]
+
+
+subsection "recursive equations of operators"
+
+subsubsection "Map"
+
+lemma Map_UU: "Map f$UU =UU"
+  by (simp add: Map_def)
+
+lemma Map_nil: "Map f$nil =nil"
+  by (simp add: Map_def)
+
+lemma Map_cons: "Map f$(x\<leadsto>xs)=(f x) \<leadsto> Map f$xs"
+  by (simp add: Map_def Consq_def flift2_def)
+
+
+subsubsection \<open>Filter\<close>
+
+lemma Filter_UU: "Filter P$UU =UU"
+  by (simp add: Filter_def)
+
+lemma Filter_nil: "Filter P$nil =nil"
+  by (simp add: Filter_def)
+
+lemma Filter_cons:
+  "Filter P$(x\<leadsto>xs)= (if P x then x\<leadsto>(Filter P$xs) else Filter P$xs)"
+  by (simp add: Filter_def Consq_def flift2_def If_and_if)
+
+
+subsubsection \<open>Forall\<close>
+
+lemma Forall_UU: "Forall P UU"
+  by (simp add: Forall_def sforall_def)
+
+lemma Forall_nil: "Forall P nil"
+  by (simp add: Forall_def sforall_def)
+
+lemma Forall_cons: "Forall P (x\<leadsto>xs)= (P x & Forall P xs)"
+  by (simp add: Forall_def sforall_def Consq_def flift2_def)
+
+
+subsubsection \<open>Conc\<close>
+
+lemma Conc_cons: "(x\<leadsto>xs) @@ y = x\<leadsto>(xs @@y)"
+  by (simp add: Consq_def)
+
+
+subsubsection \<open>Takewhile\<close>
+
+lemma Takewhile_UU: "Takewhile P$UU =UU"
+  by (simp add: Takewhile_def)
+
+lemma Takewhile_nil: "Takewhile P$nil =nil"
+  by (simp add: Takewhile_def)
+
+lemma Takewhile_cons:
+  "Takewhile P$(x\<leadsto>xs)= (if P x then x\<leadsto>(Takewhile P$xs) else nil)"
+  by (simp add: Takewhile_def Consq_def flift2_def If_and_if)
+
+
+subsubsection \<open>DropWhile\<close>
+
+lemma Dropwhile_UU: "Dropwhile P$UU =UU"
+  by (simp add: Dropwhile_def)
+
+lemma Dropwhile_nil: "Dropwhile P$nil =nil"
+  by (simp add: Dropwhile_def)
+
+lemma Dropwhile_cons:
+  "Dropwhile P$(x\<leadsto>xs)= (if P x then Dropwhile P$xs else x\<leadsto>xs)"
+  by (simp add: Dropwhile_def Consq_def flift2_def If_and_if)
+
+
+subsubsection \<open>Last\<close>
+
+lemma Last_UU: "Last$UU =UU"
+  by (simp add: Last_def)
+
+lemma Last_nil: "Last$nil =UU"
+  by (simp add: Last_def)
+
+lemma Last_cons: "Last$(x\<leadsto>xs)= (if xs=nil then Def x else Last$xs)"
+  apply (simp add: Last_def Consq_def)
+  apply (cases xs)
+  apply simp_all
+  done
+
+
+subsubsection \<open>Flat\<close>
+
+lemma Flat_UU: "Flat$UU =UU"
+  by (simp add: Flat_def)
+
+lemma Flat_nil: "Flat$nil =nil"
+  by (simp add: Flat_def)
+
+lemma Flat_cons: "Flat$(x##xs)= x @@ (Flat$xs)"
+  by (simp add: Flat_def Consq_def)
+
+
+subsubsection \<open>Zip\<close>
+
+lemma Zip_unfold:
+  "Zip = (LAM t1 t2. case t1 of
+                  nil   => nil
+                | x##xs => (case t2 of
+                             nil => UU
+                           | y##ys => (case x of
+                                         UU  => UU
+                                       | Def a => (case y of
+                                                     UU => UU
+                                                   | Def b => Def (a,b)##(Zip$xs$ys)))))"
+  apply (rule trans)
+  apply (rule fix_eq4)
+  apply (rule Zip_def)
+  apply (rule beta_cfun)
+  apply simp
+  done
+
+lemma Zip_UU1: "Zip$UU$y =UU"
+  apply (subst Zip_unfold)
+  apply simp
+  done
+
+lemma Zip_UU2: "x~=nil ==> Zip$x$UU =UU"
+  apply (subst Zip_unfold)
+  apply simp
+  apply (cases x)
+  apply simp_all
+  done
+
+lemma Zip_nil: "Zip$nil$y =nil"
+  apply (subst Zip_unfold)
+  apply simp
+  done
+
+lemma Zip_cons_nil: "Zip$(x\<leadsto>xs)$nil= UU"
+  apply (subst Zip_unfold)
+  apply (simp add: Consq_def)
+  done
+
+lemma Zip_cons: "Zip$(x\<leadsto>xs)$(y\<leadsto>ys)= (x,y) \<leadsto> Zip$xs$ys"
+  apply (rule trans)
+  apply (subst Zip_unfold)
+  apply simp
+  apply (simp add: Consq_def)
+  done
+
+lemmas [simp del] =
+  sfilter_UU sfilter_nil sfilter_cons
+  smap_UU smap_nil smap_cons
+  sforall2_UU sforall2_nil sforall2_cons
+  slast_UU slast_nil slast_cons
+  stakewhile_UU  stakewhile_nil  stakewhile_cons
+  sdropwhile_UU  sdropwhile_nil  sdropwhile_cons
+  sflat_UU sflat_nil sflat_cons
+  szip_UU1 szip_UU2 szip_nil szip_cons_nil szip_cons
+
+lemmas [simp] =
+  Filter_UU Filter_nil Filter_cons
+  Map_UU Map_nil Map_cons
+  Forall_UU Forall_nil Forall_cons
+  Last_UU Last_nil Last_cons
+  Conc_cons
+  Takewhile_UU Takewhile_nil Takewhile_cons
+  Dropwhile_UU Dropwhile_nil Dropwhile_cons
+  Zip_UU1 Zip_UU2 Zip_nil Zip_cons_nil Zip_cons
+
+
+
+section "Cons"
+
+lemma Consq_def2: "a\<leadsto>s = (Def a)##s"
+  by (simp add: Consq_def)
+
+lemma Seq_exhaust: "x = UU | x = nil | (? a s. x = a \<leadsto> s)"
+  apply (simp add: Consq_def2)
+  apply (cut_tac seq.nchotomy)
+  apply (fast dest: not_Undef_is_Def [THEN iffD1])
+  done
+
+
+lemma Seq_cases: "!!P. [| x = UU ==> P; x = nil ==> P; !!a s. x = a \<leadsto> s  ==> P |] ==> P"
+  apply (cut_tac x="x" in Seq_exhaust)
+  apply (erule disjE)
+  apply simp
+  apply (erule disjE)
+  apply simp
+  apply (erule exE)+
+  apply simp
+  done
+
+(*
+fun Seq_case_tac s i = rule_tac x",s)] Seq_cases i
+          THEN hyp_subst_tac i THEN hyp_subst_tac (i+1) THEN hyp_subst_tac (i+2);
+*)
+(* on a\<leadsto>s only simp_tac, as full_simp_tac is uncomplete and often causes errors *)
+(*
+fun Seq_case_simp_tac s i = Seq_case_tac s i THEN Asm_simp_tac (i+2)
+                                             THEN Asm_full_simp_tac (i+1)
+                                             THEN Asm_full_simp_tac i;
+*)
+
+lemma Cons_not_UU: "a\<leadsto>s ~= UU"
+  apply (subst Consq_def2)
+  apply simp
+  done
+
+
+lemma Cons_not_less_UU: "~(a\<leadsto>x) << UU"
+  apply (rule notI)
+  apply (drule below_antisym)
+  apply simp
+  apply (simp add: Cons_not_UU)
+  done
+
+lemma Cons_not_less_nil: "~a\<leadsto>s << nil"
+  by (simp add: Consq_def2)
+
+lemma Cons_not_nil: "a\<leadsto>s ~= nil"
+  by (simp add: Consq_def2)
+
+lemma Cons_not_nil2: "nil ~= a\<leadsto>s"
+  by (simp add: Consq_def2)
+
+lemma Cons_inject_eq: "(a\<leadsto>s = b\<leadsto>t) = (a = b & s = t)"
+  apply (simp only: Consq_def2)
+  apply (simp add: scons_inject_eq)
+  done
+
+lemma Cons_inject_less_eq: "(a\<leadsto>s<<b\<leadsto>t) = (a = b & s<<t)"
+  by (simp add: Consq_def2)
+
+lemma seq_take_Cons: "seq_take (Suc n)$(a\<leadsto>x) = a\<leadsto> (seq_take n$x)"
+  by (simp add: Consq_def)
+
+lemmas [simp] =
+  Cons_not_nil2 Cons_inject_eq Cons_inject_less_eq seq_take_Cons
+  Cons_not_UU Cons_not_less_UU Cons_not_less_nil Cons_not_nil
+
+
+subsection "induction"
+
+lemma Seq_induct: "!! P. [| adm P; P UU; P nil; !! a s. P s ==> P (a\<leadsto>s)|] ==> P x"
+  apply (erule (2) seq.induct)
+  apply defined
+  apply (simp add: Consq_def)
+  done
+
+lemma Seq_FinitePartial_ind:
+  "!! P.[|P UU;P nil; !! a s. P s ==> P(a\<leadsto>s) |]
+                  ==> seq_finite x --> P x"
+  apply (erule (1) seq_finite_ind)
+  apply defined
+  apply (simp add: Consq_def)
+  done
+
+lemma Seq_Finite_ind:
+  "!! P.[| Finite x; P nil; !! a s. [| Finite s; P s|] ==> P (a\<leadsto>s) |] ==> P x"
+  apply (erule (1) Finite.induct)
+  apply defined
+  apply (simp add: Consq_def)
+  done
+
+
+(* rws are definitions to be unfolded for admissibility check *)
+(*
+fun Seq_induct_tac s rws i = rule_tac x",s)] Seq_induct i
+                         THEN (REPEAT_DETERM (CHANGED (Asm_simp_tac (i+1))))
+                         THEN simp add: rws) i;
+
+fun Seq_Finite_induct_tac i = erule Seq_Finite_ind i
+                              THEN (REPEAT_DETERM (CHANGED (Asm_simp_tac i)));
+
+fun pair_tac s = rule_tac p",s)] prod.exhaust
+                          THEN' hyp_subst_tac THEN' Simp_tac;
+*)
+(* induction on a sequence of pairs with pairsplitting and simplification *)
+(*
+fun pair_induct_tac s rws i =
+           rule_tac x",s)] Seq_induct i
+           THEN pair_tac "a" (i+3)
+           THEN (REPEAT_DETERM (CHANGED (Simp_tac (i+1))))
+           THEN simp add: rws) i;
+*)
+
+
+(* ------------------------------------------------------------------------------------ *)
+
+subsection "HD,TL"
+
+lemma HD_Cons [simp]: "HD$(x\<leadsto>y) = Def x"
+  by (simp add: Consq_def)
+
+lemma TL_Cons [simp]: "TL$(x\<leadsto>y) = y"
+  by (simp add: Consq_def)
+
+(* ------------------------------------------------------------------------------------ *)
+
+subsection "Finite, Partial, Infinite"
+
+lemma Finite_Cons [simp]: "Finite (a\<leadsto>xs) = Finite xs"
+  by (simp add: Consq_def2 Finite_cons)
+
+lemma FiniteConc_1: "Finite (x::'a Seq) ==> Finite y --> Finite (x@@y)"
+  apply (erule Seq_Finite_ind)
+  apply simp_all
+  done
+
+lemma FiniteConc_2: "Finite (z::'a Seq) ==> !x y. z= x@@y --> (Finite x & Finite y)"
+  apply (erule Seq_Finite_ind)
+  (* nil*)
+  apply (intro strip)
+  apply (rule_tac x="x" in Seq_cases, simp_all)
+  (* cons *)
+  apply (intro strip)
+  apply (rule_tac x="x" in Seq_cases, simp_all)
+  apply (rule_tac x="y" in Seq_cases, simp_all)
+  done
+
+lemma FiniteConc [simp]: "Finite(x@@y) = (Finite (x::'a Seq) & Finite y)"
+  apply (rule iffI)
+  apply (erule FiniteConc_2 [rule_format])
+  apply (rule refl)
+  apply (rule FiniteConc_1 [rule_format])
+  apply auto
+  done
+
+
+lemma FiniteMap1: "Finite s ==> Finite (Map f$s)"
+  apply (erule Seq_Finite_ind)
+  apply simp_all
+  done
+
+lemma FiniteMap2: "Finite s ==> ! t. (s = Map f$t) --> Finite t"
+  apply (erule Seq_Finite_ind)
+  apply (intro strip)
+  apply (rule_tac x="t" in Seq_cases, simp_all)
+  (* main case *)
+  apply auto
+  apply (rule_tac x="t" in Seq_cases, simp_all)
+  done
+
+lemma Map2Finite: "Finite (Map f$s) = Finite s"
+  apply auto
+  apply (erule FiniteMap2 [rule_format])
+  apply (rule refl)
+  apply (erule FiniteMap1)
+  done
+
+
+lemma FiniteFilter: "Finite s ==> Finite (Filter P$s)"
+  apply (erule Seq_Finite_ind)
+  apply simp_all
+  done
+
+
+(* ----------------------------------------------------------------------------------- *)
+
+subsection "Conc"
+
+lemma Conc_cong: "!! x::'a Seq. Finite x ==> ((x @@ y) = (x @@ z)) = (y = z)"
+  apply (erule Seq_Finite_ind)
+  apply simp_all
+  done
+
+lemma Conc_assoc: "(x @@ y) @@ z = (x::'a Seq) @@ y @@ z"
+  apply (rule_tac x="x" in Seq_induct)
+  apply simp_all
+  done
+
+lemma nilConc [simp]: "s@@ nil = s"
+  apply (induct s)
+  apply simp
+  apply simp
+  apply simp
+  apply simp
+  done
+
+
+(* should be same as nil_is_Conc2 when all nils are turned to right side !! *)
+lemma nil_is_Conc: "(nil = x @@ y) = ((x::'a Seq)= nil & y = nil)"
+  apply (rule_tac x="x" in Seq_cases)
+  apply auto
+  done
+
+lemma nil_is_Conc2: "(x @@ y = nil) = ((x::'a Seq)= nil & y = nil)"
+  apply (rule_tac x="x" in Seq_cases)
+  apply auto
+  done
+
+
+(* ------------------------------------------------------------------------------------ *)
+
+subsection "Last"
+
+lemma Finite_Last1: "Finite s ==> s~=nil --> Last$s~=UU"
+  apply (erule Seq_Finite_ind, simp_all)
+  done
+
+lemma Finite_Last2: "Finite s ==> Last$s=UU --> s=nil"
+  apply (erule Seq_Finite_ind, simp_all)
+  apply fast
+  done
+
+
+(* ------------------------------------------------------------------------------------ *)
+
+
+subsection "Filter, Conc"
+
+
+lemma FilterPQ: "Filter P$(Filter Q$s) = Filter (%x. P x & Q x)$s"
+  apply (rule_tac x="s" in Seq_induct, simp_all)
+  done
+
+lemma FilterConc: "Filter P$(x @@ y) = (Filter P$x @@ Filter P$y)"
+  apply (simp add: Filter_def sfiltersconc)
+  done
+
+(* ------------------------------------------------------------------------------------ *)
+
+subsection "Map"
+
+lemma MapMap: "Map f$(Map g$s) = Map (f o g)$s"
+  apply (rule_tac x="s" in Seq_induct, simp_all)
+  done
+
+lemma MapConc: "Map f$(x@@y) = (Map f$x) @@ (Map f$y)"
+  apply (rule_tac x="x" in Seq_induct, simp_all)
+  done
+
+lemma MapFilter: "Filter P$(Map f$x) = Map f$(Filter (P o f)$x)"
+  apply (rule_tac x="x" in Seq_induct, simp_all)
+  done
+
+lemma nilMap: "nil = (Map f$s) --> s= nil"
+  apply (rule_tac x="s" in Seq_cases, simp_all)
+  done
+
+
+lemma ForallMap: "Forall P (Map f$s) = Forall (P o f) s"
+  apply (rule_tac x="s" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+
+
+
+(* ------------------------------------------------------------------------------------ *)
+
+subsection "Forall"
+
+lemma ForallPForallQ1: "Forall P ys & (! x. P x --> Q x) --> Forall Q ys"
+  apply (rule_tac x="ys" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+lemmas ForallPForallQ =
+  ForallPForallQ1 [THEN mp, OF conjI, OF _ allI, OF _ impI]
+
+lemma Forall_Conc_impl: "(Forall P x & Forall P y) --> Forall P (x @@ y)"
+  apply (rule_tac x="x" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+lemma Forall_Conc [simp]:
+  "Finite x ==> Forall P (x @@ y) = (Forall P x & Forall P y)"
+  apply (erule Seq_Finite_ind, simp_all)
+  done
+
+lemma ForallTL1: "Forall P s  --> Forall P (TL$s)"
+  apply (rule_tac x="s" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+lemmas ForallTL = ForallTL1 [THEN mp]
+
+lemma ForallDropwhile1: "Forall P s  --> Forall P (Dropwhile Q$s)"
+  apply (rule_tac x="s" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+lemmas ForallDropwhile = ForallDropwhile1 [THEN mp]
+
+
+(* only admissible in t, not if done in s *)
+
+lemma Forall_prefix: "! s. Forall P s --> t<<s --> Forall P t"
+  apply (rule_tac x="t" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  apply (intro strip)
+  apply (rule_tac x="sa" in Seq_cases)
+  apply simp
+  apply auto
+  done
+
+lemmas Forall_prefixclosed = Forall_prefix [rule_format]
+
+lemma Forall_postfixclosed: "[| Finite h; Forall P s; s= h @@ t |] ==> Forall P t"
+  by auto
+
+
+lemma ForallPFilterQR1:
+  "((! x. P x --> (Q x = R x)) & Forall P tr) --> Filter Q$tr = Filter R$tr"
+  apply (rule_tac x="tr" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+lemmas ForallPFilterQR = ForallPFilterQR1 [THEN mp, OF conjI, OF allI]
+
+
+(* ------------------------------------------------------------------------------------- *)
+
+subsection "Forall, Filter"
+
+
+lemma ForallPFilterP: "Forall P (Filter P$x)"
+  by (simp add: Filter_def Forall_def forallPsfilterP)
+
+(* holds also in other direction, then equal to forallPfilterP *)
+lemma ForallPFilterPid1: "Forall P x --> Filter P$x = x"
+  apply (rule_tac x="x" in Seq_induct)
+  apply (simp add: Forall_def sforall_def Filter_def)
+  apply simp_all
+  done
+
+lemmas ForallPFilterPid = ForallPFilterPid1 [THEN mp]
+
+
+(* holds also in other direction *)
+lemma ForallnPFilterPnil1: "!! ys . Finite ys ==>
+   Forall (%x. ~P x) ys --> Filter P$ys = nil "
+  apply (erule Seq_Finite_ind, simp_all)
+  done
+
+lemmas ForallnPFilterPnil = ForallnPFilterPnil1 [THEN mp]
+
+
+(* holds also in other direction *)
+lemma ForallnPFilterPUU1: "~Finite ys & Forall (%x. ~P x) ys --> Filter P$ys = UU"
+  apply (rule_tac x="ys" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+lemmas ForallnPFilterPUU = ForallnPFilterPUU1 [THEN mp, OF conjI]
+
+
+(* inverse of ForallnPFilterPnil *)
+
+lemma FilternPnilForallP [rule_format]: "Filter P$ys = nil -->
+   (Forall (%x. ~P x) ys & Finite ys)"
+  apply (rule_tac x="ys" in Seq_induct)
+  (* adm *)
+  apply (simp add: Forall_def sforall_def)
+  (* base cases *)
+  apply simp
+  apply simp
+  (* main case *)
+  apply simp
+  done
+
+
+(* inverse of ForallnPFilterPUU *)
+
+lemma FilternPUUForallP:
+  assumes "Filter P$ys = UU"
+  shows "Forall (%x. ~P x) ys  & ~Finite ys"
+proof
+  show "Forall (%x. ~P x) ys"
+  proof (rule classical)
+    assume "\<not> ?thesis"
+    then have "Filter P$ys ~= UU"
+      apply (rule rev_mp)
+      apply (induct ys rule: Seq_induct)
+      apply (simp add: Forall_def sforall_def)
+      apply simp_all
+      done
+    with assms show ?thesis by contradiction
+  qed
+  show "~ Finite ys"
+  proof
+    assume "Finite ys"
+    then have "Filter P$ys ~= UU"
+      by (rule Seq_Finite_ind) simp_all
+    with assms show False by contradiction
+  qed
+qed
+
+
+lemma ForallQFilterPnil:
+  "!! Q P.[| Forall Q ys; Finite ys; !!x. Q x ==> ~P x|]
+    ==> Filter P$ys = nil"
+  apply (erule ForallnPFilterPnil)
+  apply (erule ForallPForallQ)
+  apply auto
+  done
+
+lemma ForallQFilterPUU:
+ "!! Q P. [| ~Finite ys; Forall Q ys;  !!x. Q x ==> ~P x|]
+    ==> Filter P$ys = UU "
+  apply (erule ForallnPFilterPUU)
+  apply (erule ForallPForallQ)
+  apply auto
+  done
+
+
+
+(* ------------------------------------------------------------------------------------- *)
+
+subsection "Takewhile, Forall, Filter"
+
+
+lemma ForallPTakewhileP [simp]: "Forall P (Takewhile P$x)"
+  by (simp add: Forall_def Takewhile_def sforallPstakewhileP)
+
+
+lemma ForallPTakewhileQ [simp]: "!! P. [| !!x. Q x==> P x |] ==> Forall P (Takewhile Q$x)"
+  apply (rule ForallPForallQ)
+  apply (rule ForallPTakewhileP)
+  apply auto
+  done
+
+
+lemma FilterPTakewhileQnil [simp]:
+  "!! Q P.[| Finite (Takewhile Q$ys); !!x. Q x ==> ~P x |]
+   ==> Filter P$(Takewhile Q$ys) = nil"
+  apply (erule ForallnPFilterPnil)
+  apply (rule ForallPForallQ)
+  apply (rule ForallPTakewhileP)
+  apply auto
+  done
+
+lemma FilterPTakewhileQid [simp]:
+ "!! Q P. [| !!x. Q x ==> P x |] ==>
+            Filter P$(Takewhile Q$ys) = (Takewhile Q$ys)"
+  apply (rule ForallPFilterPid)
+  apply (rule ForallPForallQ)
+  apply (rule ForallPTakewhileP)
+  apply auto
+  done
+
+
+lemma Takewhile_idempotent: "Takewhile P$(Takewhile P$s) = Takewhile P$s"
+  apply (rule_tac x="s" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+lemma ForallPTakewhileQnP [simp]: "Forall P s --> Takewhile (%x. Q x | (~P x))$s = Takewhile Q$s"
+  apply (rule_tac x="s" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+lemma ForallPDropwhileQnP [simp]: "Forall P s --> Dropwhile (%x. Q x | (~P x))$s = Dropwhile Q$s"
+  apply (rule_tac x="s" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+
+lemma TakewhileConc1: "Forall P s --> Takewhile P$(s @@ t) = s @@ (Takewhile P$t)"
+  apply (rule_tac x="s" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+lemmas TakewhileConc = TakewhileConc1 [THEN mp]
+
+lemma DropwhileConc1: "Finite s ==> Forall P s --> Dropwhile P$(s @@ t) = Dropwhile P$t"
+  apply (erule Seq_Finite_ind, simp_all)
+  done
+
+lemmas DropwhileConc = DropwhileConc1 [THEN mp]
+
+
+
+(* ----------------------------------------------------------------------------------- *)
+
+subsection "coinductive characterizations of Filter"
+
+lemma divide_Seq_lemma:
+ "HD$(Filter P$y) = Def x
+    --> y = ((Takewhile (%x. ~P x)$y) @@ (x \<leadsto> TL$(Dropwhile (%a. ~P a)$y)))
+             & Finite (Takewhile (%x. ~ P x)$y)  & P x"
+  (* FIX: pay attention: is only admissible with chain-finite package to be added to
+          adm test and Finite f x admissibility *)
+  apply (rule_tac x="y" in Seq_induct)
+  apply (simp add: adm_subst [OF _ adm_Finite])
+  apply simp
+  apply simp
+  apply (case_tac "P a")
+   apply simp
+   apply blast
+  (* ~ P a *)
+  apply simp
+  done
+
+lemma divide_Seq: "(x\<leadsto>xs) << Filter P$y
+   ==> y = ((Takewhile (%a. ~ P a)$y) @@ (x \<leadsto> TL$(Dropwhile (%a. ~ P a)$y)))
+      & Finite (Takewhile (%a. ~ P a)$y)  & P x"
+  apply (rule divide_Seq_lemma [THEN mp])
+  apply (drule_tac f="HD" and x="x\<leadsto>xs" in  monofun_cfun_arg)
+  apply simp
+  done
+
+
+lemma nForall_HDFilter: "~Forall P y --> (? x. HD$(Filter (%a. ~P a)$y) = Def x)"
+  unfolding not_Undef_is_Def [symmetric]
+  apply (induct y rule: Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+
+lemma divide_Seq2: "~Forall P y
+  ==> ? x. y= (Takewhile P$y @@ (x \<leadsto> TL$(Dropwhile P$y))) &
+      Finite (Takewhile P$y) & (~ P x)"
+  apply (drule nForall_HDFilter [THEN mp])
+  apply safe
+  apply (rule_tac x="x" in exI)
+  apply (cut_tac P1="%x. ~ P x" in divide_Seq_lemma [THEN mp])
+  apply auto
+  done
+
+
+lemma divide_Seq3: "~Forall P y
+  ==> ? x bs rs. y= (bs @@ (x\<leadsto>rs)) & Finite bs & Forall P bs & (~ P x)"
+  apply (drule divide_Seq2)
+  apply fastforce
+  done
+
+lemmas [simp] = FilterPQ FilterConc Conc_cong
+
+
+(* ------------------------------------------------------------------------------------- *)
+
+
+subsection "take_lemma"
+
+lemma seq_take_lemma: "(!n. seq_take n$x = seq_take n$x') = (x = x')"
+  apply (rule iffI)
+  apply (rule seq.take_lemma)
+  apply auto
+  done
+
+lemma take_reduction1:
+  "\<forall>n. ((! k. k < n --> seq_take k$y1 = seq_take k$y2)
+    --> seq_take n$(x @@ (t\<leadsto>y1)) =  seq_take n$(x @@ (t\<leadsto>y2)))"
+  apply (rule_tac x="x" in Seq_induct)
+  apply simp_all
+  apply (intro strip)
+  apply (case_tac "n")
+  apply auto
+  apply (case_tac "n")
+  apply auto
+  done
+
+
+lemma take_reduction:
+  "!! n.[| x=y; s=t; !! k. k<n ==> seq_take k$y1 = seq_take k$y2|]
+    ==> seq_take n$(x @@ (s\<leadsto>y1)) =  seq_take n$(y @@ (t\<leadsto>y2))"
+  by (auto intro!: take_reduction1 [rule_format])
+
+(* ------------------------------------------------------------------
+          take-lemma and take_reduction for << instead of =
+   ------------------------------------------------------------------ *)
+
+lemma take_reduction_less1:
+  "\<forall>n. ((! k. k < n --> seq_take k$y1 << seq_take k$y2)
+    --> seq_take n$(x @@ (t\<leadsto>y1)) <<  seq_take n$(x @@ (t\<leadsto>y2)))"
+  apply (rule_tac x="x" in Seq_induct)
+  apply simp_all
+  apply (intro strip)
+  apply (case_tac "n")
+  apply auto
+  apply (case_tac "n")
+  apply auto
+  done
+
+
+lemma take_reduction_less:
+  "\<And>n.[| x=y; s=t;!! k. k<n ==> seq_take k$y1 << seq_take k$y2|]
+    ==> seq_take n$(x @@ (s\<leadsto>y1)) <<  seq_take n$(y @@ (t\<leadsto>y2))"
+  by (auto intro!: take_reduction_less1 [rule_format])
+
+lemma take_lemma_less1:
+  assumes "!! n. seq_take n$s1 << seq_take n$s2"
+  shows "s1<<s2"
+  apply (rule_tac t="s1" in seq.reach [THEN subst])
+  apply (rule_tac t="s2" in seq.reach [THEN subst])
+  apply (rule lub_mono)
+  apply (rule seq.chain_take [THEN ch2ch_Rep_cfunL])
+  apply (rule seq.chain_take [THEN ch2ch_Rep_cfunL])
+  apply (rule assms)
+  done
+
+
+lemma take_lemma_less: "(!n. seq_take n$x << seq_take n$x') = (x << x')"
+  apply (rule iffI)
+  apply (rule take_lemma_less1)
+  apply auto
+  apply (erule monofun_cfun_arg)
+  done
+
+(* ------------------------------------------------------------------
+          take-lemma proof principles
+   ------------------------------------------------------------------ *)
+
+lemma take_lemma_principle1:
+  "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
+            !! s1 s2 y. [| Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2)|]
+                          ==> (f (s1 @@ y\<leadsto>s2)) = (g (s1 @@ y\<leadsto>s2)) |]
+               ==> A x --> (f x)=(g x)"
+  apply (case_tac "Forall Q x")
+  apply (auto dest!: divide_Seq3)
+  done
+
+lemma take_lemma_principle2:
+  "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
+           !! s1 s2 y. [| Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2)|]
+                          ==> ! n. seq_take n$(f (s1 @@ y\<leadsto>s2))
+                                 = seq_take n$(g (s1 @@ y\<leadsto>s2)) |]
+               ==> A x --> (f x)=(g x)"
+  apply (case_tac "Forall Q x")
+  apply (auto dest!: divide_Seq3)
+  apply (rule seq.take_lemma)
+  apply auto
+  done
+
+
+(* Note: in the following proofs the ordering of proof steps is very
+         important, as otherwise either (Forall Q s1) would be in the IH as
+         assumption (then rule useless) or it is not possible to strengthen
+         the IH apply doing a forall closure of the sequence t (then rule also useless).
+         This is also the reason why the induction rule (nat_less_induct or nat_induct) has to
+         to be imbuilt into the rule, as induction has to be done early and the take lemma
+         has to be used in the trivial direction afterwards for the (Forall Q x) case.  *)
+
+lemma take_lemma_induct:
+"!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
+         !! s1 s2 y n. [| ! t. A t --> seq_take n$(f t) = seq_take n$(g t);
+                          Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2) |]
+                          ==>   seq_take (Suc n)$(f (s1 @@ y\<leadsto>s2))
+                              = seq_take (Suc n)$(g (s1 @@ y\<leadsto>s2)) |]
+               ==> A x --> (f x)=(g x)"
+  apply (rule impI)
+  apply (rule seq.take_lemma)
+  apply (rule mp)
+  prefer 2 apply assumption
+  apply (rule_tac x="x" in spec)
+  apply (rule nat.induct)
+  apply simp
+  apply (rule allI)
+  apply (case_tac "Forall Q xa")
+  apply (force intro!: seq_take_lemma [THEN iffD2, THEN spec])
+  apply (auto dest!: divide_Seq3)
+  done
+
+
+lemma take_lemma_less_induct:
+"!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
+        !! s1 s2 y n. [| ! t m. m < n --> A t --> seq_take m$(f t) = seq_take m$(g t);
+                          Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2) |]
+                          ==>   seq_take n$(f (s1 @@ y\<leadsto>s2))
+                              = seq_take n$(g (s1 @@ y\<leadsto>s2)) |]
+               ==> A x --> (f x)=(g x)"
+  apply (rule impI)
+  apply (rule seq.take_lemma)
+  apply (rule mp)
+  prefer 2 apply assumption
+  apply (rule_tac x="x" in spec)
+  apply (rule nat_less_induct)
+  apply (rule allI)
+  apply (case_tac "Forall Q xa")
+  apply (force intro!: seq_take_lemma [THEN iffD2, THEN spec])
+  apply (auto dest!: divide_Seq3)
+  done
+
+
+
+lemma take_lemma_in_eq_out:
+"!! Q. [| A UU  ==> (f UU) = (g UU) ;
+          A nil ==> (f nil) = (g nil) ;
+          !! s y n. [| ! t. A t --> seq_take n$(f t) = seq_take n$(g t);
+                     A (y\<leadsto>s) |]
+                     ==>   seq_take (Suc n)$(f (y\<leadsto>s))
+                         = seq_take (Suc n)$(g (y\<leadsto>s)) |]
+               ==> A x --> (f x)=(g x)"
+  apply (rule impI)
+  apply (rule seq.take_lemma)
+  apply (rule mp)
+  prefer 2 apply assumption
+  apply (rule_tac x="x" in spec)
+  apply (rule nat.induct)
+  apply simp
+  apply (rule allI)
+  apply (rule_tac x="xa" in Seq_cases)
+  apply simp_all
+  done
+
+
+(* ------------------------------------------------------------------------------------ *)
+
+subsection "alternative take_lemma proofs"
+
+
+(* --------------------------------------------------------------- *)
+(*              Alternative Proof of FilterPQ                      *)
+(* --------------------------------------------------------------- *)
+
+declare FilterPQ [simp del]
+
+
+(* In general: How to do this case without the same adm problems
+   as for the entire proof ? *)
+lemma Filter_lemma1: "Forall (%x. ~(P x & Q x)) s
+          --> Filter P$(Filter Q$s) =
+              Filter (%x. P x & Q x)$s"
+
+  apply (rule_tac x="s" in Seq_induct)
+  apply (simp add: Forall_def sforall_def)
+  apply simp_all
+  done
+
+lemma Filter_lemma2: "Finite s ==>
+          (Forall (%x. (~P x) | (~ Q x)) s
+          --> Filter P$(Filter Q$s) = nil)"
+  apply (erule Seq_Finite_ind, simp_all)
+  done
+
+lemma Filter_lemma3: "Finite s ==>
+          Forall (%x. (~P x) | (~ Q x)) s
+          --> Filter (%x. P x & Q x)$s = nil"
+  apply (erule Seq_Finite_ind, simp_all)
+  done
+
+
+lemma FilterPQ_takelemma: "Filter P$(Filter Q$s) = Filter (%x. P x & Q x)$s"
+  apply (rule_tac A1="%x. True" and
+                   Q1="%x. ~(P x & Q x)" and x1="s" in
+                   take_lemma_induct [THEN mp])
+  (* better support for A = %x. True *)
+  apply (simp add: Filter_lemma1)
+  apply (simp add: Filter_lemma2 Filter_lemma3)
+  apply simp
+  done
+
+declare FilterPQ [simp]
+
+
+(* --------------------------------------------------------------- *)
+(*              Alternative Proof of MapConc                       *)
+(* --------------------------------------------------------------- *)
+
+
+
+lemma MapConc_takelemma: "Map f$(x@@y) = (Map f$x) @@ (Map f$y)"
+  apply (rule_tac A1="%x. True" and x1="x" in
+      take_lemma_in_eq_out [THEN mp])
+  apply auto
+  done
+
+
+ML \<open>
+
+fun Seq_case_tac ctxt s i =
+  Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_cases} i
+  THEN hyp_subst_tac ctxt i THEN hyp_subst_tac ctxt (i+1) THEN hyp_subst_tac ctxt (i+2);
+
+(* on a\<leadsto>s only simp_tac, as full_simp_tac is uncomplete and often causes errors *)
+fun Seq_case_simp_tac ctxt s i =
+  Seq_case_tac ctxt s i
+  THEN asm_simp_tac ctxt (i+2)
+  THEN asm_full_simp_tac ctxt (i+1)
+  THEN asm_full_simp_tac ctxt i;
+
+(* rws are definitions to be unfolded for admissibility check *)
+fun Seq_induct_tac ctxt s rws i =
+  Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_induct} i
+  THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ctxt (i+1))))
+  THEN simp_tac (ctxt addsimps rws) i;
+
+fun Seq_Finite_induct_tac ctxt i =
+  eresolve_tac ctxt @{thms Seq_Finite_ind} i
+  THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ctxt i)));
+
+fun pair_tac ctxt s =
+  Rule_Insts.res_inst_tac ctxt [((("y", 0), Position.none), s)] [] @{thm prod.exhaust}
+  THEN' hyp_subst_tac ctxt THEN' asm_full_simp_tac ctxt;
+
+(* induction on a sequence of pairs with pairsplitting and simplification *)
+fun pair_induct_tac ctxt s rws i =
+  Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_induct} i
+  THEN pair_tac ctxt "a" (i+3)
+  THEN (REPEAT_DETERM (CHANGED (simp_tac ctxt (i+1))))
+  THEN simp_tac (ctxt addsimps rws) i;
+\<close>
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/ShortExecutions.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,278 @@
+(*  Title:      HOL/HOLCF/IOA/ShortExecutions.thy
+    Author:     Olaf Müller
+*)
+
+theory ShortExecutions
+imports Traces
+begin
+
+text \<open>
+  Some properties about \<open>Cut ex\<close>, defined as follows:
+
+  For every execution ex there is another shorter execution \<open>Cut ex\<close>
+  that has the same trace as ex, but its schedule ends with an external action.
+\<close>
+
+definition
+  oraclebuild :: "('a => bool) => 'a Seq -> 'a Seq -> 'a Seq" where
+  "oraclebuild P = (fix$(LAM h s t. case t of
+       nil => nil
+    | x##xs =>
+      (case x of
+        UU => UU
+      | Def y => (Takewhile (%x. \<not>P x)$s)
+                  @@ (y\<leadsto>(h$(TL$(Dropwhile (%x. \<not> P x)$s))$xs))
+      )
+    ))"
+
+definition
+  Cut :: "('a => bool) => 'a Seq => 'a Seq" where
+  "Cut P s = oraclebuild P$s$(Filter P$s)"
+
+definition
+  LastActExtsch :: "('a,'s)ioa => 'a Seq => bool" where
+  "LastActExtsch A sch = (Cut (%x. x: ext A) sch = sch)"
+
+(* LastActExtex      ::"('a,'s)ioa => ('a,'s) pairs  => bool"*)
+(* LastActExtex_def:
+  "LastActExtex A ex == LastActExtsch A (filter_act$ex)" *)
+
+axiomatization where
+  Cut_prefixcl_Finite: "Finite s ==> (? y. s = Cut P s @@ y)"
+
+axiomatization where
+  LastActExtsmall1: "LastActExtsch A sch ==> LastActExtsch A (TL$(Dropwhile P$sch))"
+
+axiomatization where
+  LastActExtsmall2: "[| Finite sch1; LastActExtsch A (sch1 @@ sch2) |] ==> LastActExtsch A sch2"
+
+
+ML \<open>
+fun thin_tac' ctxt j =
+  rotate_tac (j - 1) THEN'
+  eresolve_tac ctxt [thin_rl] THEN'
+  rotate_tac (~ (j - 1))
+\<close>
+
+
+subsection "oraclebuild rewrite rules"
+
+
+lemma oraclebuild_unfold:
+"oraclebuild P = (LAM s t. case t of
+       nil => nil
+    | x##xs =>
+      (case x of
+        UU => UU
+      | Def y => (Takewhile (%a. \<not> P a)$s)
+                  @@ (y\<leadsto>(oraclebuild P$(TL$(Dropwhile (%a. \<not> P a)$s))$xs))
+      )
+    )"
+apply (rule trans)
+apply (rule fix_eq2)
+apply (simp only: oraclebuild_def)
+apply (rule beta_cfun)
+apply simp
+done
+
+lemma oraclebuild_UU: "oraclebuild P$sch$UU = UU"
+apply (subst oraclebuild_unfold)
+apply simp
+done
+
+lemma oraclebuild_nil: "oraclebuild P$sch$nil = nil"
+apply (subst oraclebuild_unfold)
+apply simp
+done
+
+lemma oraclebuild_cons: "oraclebuild P$s$(x\<leadsto>t) =
+          (Takewhile (%a. \<not> P a)$s)
+           @@ (x\<leadsto>(oraclebuild P$(TL$(Dropwhile (%a. \<not> P a)$s))$t))"
+apply (rule trans)
+apply (subst oraclebuild_unfold)
+apply (simp add: Consq_def)
+apply (simp add: Consq_def)
+done
+
+
+subsection "Cut rewrite rules"
+
+lemma Cut_nil:
+"[| Forall (%a. \<not> P a) s; Finite s|]
+            ==> Cut P s =nil"
+apply (unfold Cut_def)
+apply (subgoal_tac "Filter P$s = nil")
+apply (simp (no_asm_simp) add: oraclebuild_nil)
+apply (rule ForallQFilterPnil)
+apply assumption+
+done
+
+lemma Cut_UU:
+"[| Forall (%a. \<not> P a) s; ~Finite s|]
+            ==> Cut P s =UU"
+apply (unfold Cut_def)
+apply (subgoal_tac "Filter P$s= UU")
+apply (simp (no_asm_simp) add: oraclebuild_UU)
+apply (rule ForallQFilterPUU)
+apply assumption+
+done
+
+lemma Cut_Cons:
+"[| P t;  Forall (%x. \<not> P x) ss; Finite ss|]
+            ==> Cut P (ss @@ (t\<leadsto> rs))
+                 = ss @@ (t \<leadsto> Cut P rs)"
+apply (unfold Cut_def)
+apply (simp add: ForallQFilterPnil oraclebuild_cons TakewhileConc DropwhileConc)
+done
+
+
+subsection "Cut lemmas for main theorem"
+
+lemma FilterCut: "Filter P$s = Filter P$(Cut P s)"
+apply (rule_tac A1 = "%x. True" and Q1 = "%x. \<not> P x" and x1 = "s" in take_lemma_induct [THEN mp])
+prefer 3 apply (fast)
+apply (case_tac "Finite s")
+apply (simp add: Cut_nil ForallQFilterPnil)
+apply (simp add: Cut_UU ForallQFilterPUU)
+(* main case *)
+apply (simp add: Cut_Cons ForallQFilterPnil)
+done
+
+
+lemma Cut_idemp: "Cut P (Cut P s) = (Cut P s)"
+apply (rule_tac A1 = "%x. True" and Q1 = "%x. \<not> P x" and x1 = "s" in
+  take_lemma_less_induct [THEN mp])
+prefer 3 apply (fast)
+apply (case_tac "Finite s")
+apply (simp add: Cut_nil ForallQFilterPnil)
+apply (simp add: Cut_UU ForallQFilterPUU)
+(* main case *)
+apply (simp add: Cut_Cons ForallQFilterPnil)
+apply (rule take_reduction)
+apply auto
+done
+
+
+lemma MapCut: "Map f$(Cut (P o f) s) = Cut P (Map f$s)"
+apply (rule_tac A1 = "%x. True" and Q1 = "%x. \<not> P (f x) " and x1 = "s" in
+  take_lemma_less_induct [THEN mp])
+prefer 3 apply (fast)
+apply (case_tac "Finite s")
+apply (simp add: Cut_nil)
+apply (rule Cut_nil [symmetric])
+apply (simp add: ForallMap o_def)
+apply (simp add: Map2Finite)
+(* csae ~ Finite s *)
+apply (simp add: Cut_UU)
+apply (rule Cut_UU)
+apply (simp add: ForallMap o_def)
+apply (simp add: Map2Finite)
+(* main case *)
+apply (simp add: Cut_Cons MapConc ForallMap FiniteMap1 o_def)
+apply (rule take_reduction)
+apply auto
+done
+
+
+lemma Cut_prefixcl_nFinite [rule_format (no_asm)]: "~Finite s --> Cut P s << s"
+apply (intro strip)
+apply (rule take_lemma_less [THEN iffD1])
+apply (intro strip)
+apply (rule mp)
+prefer 2 apply (assumption)
+apply (tactic "thin_tac' @{context} 1 1")
+apply (rule_tac x = "s" in spec)
+apply (rule nat_less_induct)
+apply (intro strip)
+apply (rename_tac na n s)
+apply (case_tac "Forall (%x. ~ P x) s")
+apply (rule take_lemma_less [THEN iffD2, THEN spec])
+apply (simp add: Cut_UU)
+(* main case *)
+apply (drule divide_Seq3)
+apply (erule exE)+
+apply (erule conjE)+
+apply hypsubst
+apply (simp add: Cut_Cons)
+apply (rule take_reduction_less)
+(* auto makes also reasoning about Finiteness of parts of s ! *)
+apply auto
+done
+
+
+lemma execThruCut: "!!ex .is_exec_frag A (s,ex) ==> is_exec_frag A (s,Cut P ex)"
+apply (case_tac "Finite ex")
+apply (cut_tac s = "ex" and P = "P" in Cut_prefixcl_Finite)
+apply assumption
+apply (erule exE)
+apply (rule exec_prefix2closed)
+apply (erule_tac s = "ex" and t = "Cut P ex @@ y" in subst)
+apply assumption
+apply (erule exec_prefixclosed)
+apply (erule Cut_prefixcl_nFinite)
+done
+
+
+subsection "Main Cut Theorem"
+
+lemma exists_LastActExtsch:
+ "[|sch : schedules A ; tr = Filter (%a. a:ext A)$sch|]
+    ==> ? sch. sch : schedules A &
+               tr = Filter (%a. a:ext A)$sch &
+               LastActExtsch A sch"
+
+apply (unfold schedules_def has_schedule_def [abs_def])
+apply auto
+apply (rule_tac x = "filter_act$ (Cut (%a. fst a:ext A) (snd ex))" in exI)
+apply (simp add: executions_def)
+apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+apply auto
+apply (rule_tac x = " (x1,Cut (%a. fst a:ext A) x2) " in exI)
+apply (simp (no_asm_simp))
+
+(* Subgoal 1: Lemma:  propagation of execution through Cut *)
+
+apply (simp add: execThruCut)
+
+(* Subgoal 2:  Lemma:  Filter P s = Filter P (Cut P s) *)
+
+apply (simp (no_asm) add: filter_act_def)
+apply (subgoal_tac "Map fst$ (Cut (%a. fst a: ext A) x2) = Cut (%a. a:ext A) (Map fst$x2) ")
+
+apply (rule_tac [2] MapCut [unfolded o_def])
+apply (simp add: FilterCut [symmetric])
+
+(* Subgoal 3: Lemma: Cut idempotent  *)
+
+apply (simp (no_asm) add: LastActExtsch_def filter_act_def)
+apply (subgoal_tac "Map fst$ (Cut (%a. fst a: ext A) x2) = Cut (%a. a:ext A) (Map fst$x2) ")
+apply (rule_tac [2] MapCut [unfolded o_def])
+apply (simp add: Cut_idemp)
+done
+
+
+subsection "Further Cut lemmas"
+
+lemma LastActExtimplnil:
+  "[| LastActExtsch A sch; Filter (%x. x:ext A)$sch = nil |]
+    ==> sch=nil"
+apply (unfold LastActExtsch_def)
+apply (drule FilternPnilForallP)
+apply (erule conjE)
+apply (drule Cut_nil)
+apply assumption
+apply simp
+done
+
+lemma LastActExtimplUU:
+  "[| LastActExtsch A sch; Filter (%x. x:ext A)$sch = UU |]
+    ==> sch=UU"
+apply (unfold LastActExtsch_def)
+apply (drule FilternPUUForallP)
+apply (erule conjE)
+apply (drule Cut_UU)
+apply assumption
+apply simp
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/SimCorrectness.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,291 @@
+(*  Title:      HOL/HOLCF/IOA/SimCorrectness.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Correctness of Simulations in HOLCF/IOA\<close>
+
+theory SimCorrectness
+imports Simulations
+begin
+
+definition
+  (* Note: s2 instead of s1 in last argument type !! *)
+  corresp_ex_simC :: "('a,'s2)ioa => (('s1 * 's2)set) => ('a,'s1)pairs
+                   -> ('s2 => ('a,'s2)pairs)" where
+  "corresp_ex_simC A R = (fix$(LAM h ex. (%s. case ex of
+      nil =>  nil
+    | x##xs => (flift1 (%pr. let a = (fst pr); t = (snd pr);
+                                 T' = @t'. ? ex1. (t,t'):R & move A ex1 s a t'
+                             in
+                                (@cex. move A cex s a T')
+                                 @@ ((h$xs) T'))
+                        $x) )))"
+
+definition
+  corresp_ex_sim :: "('a,'s2)ioa => (('s1 *'s2)set) =>
+                      ('a,'s1)execution => ('a,'s2)execution" where
+  "corresp_ex_sim A R ex == let S'= (@s'.(fst ex,s'):R & s': starts_of A)
+                            in
+                               (S',(corresp_ex_simC A R$(snd ex)) S')"
+
+
+subsection "corresp_ex_sim"
+
+lemma corresp_ex_simC_unfold: "corresp_ex_simC A R  = (LAM ex. (%s. case ex of
+       nil =>  nil
+     | x##xs => (flift1 (%pr. let a = (fst pr); t = (snd pr);
+                                  T' = @t'. ? ex1. (t,t'):R & move A ex1 s a t'
+                              in
+                                 (@cex. move A cex s a T')
+                               @@ ((corresp_ex_simC A R $xs) T'))
+                         $x) ))"
+apply (rule trans)
+apply (rule fix_eq2)
+apply (simp only: corresp_ex_simC_def)
+apply (rule beta_cfun)
+apply (simp add: flift1_def)
+done
+
+lemma corresp_ex_simC_UU: "(corresp_ex_simC A R$UU) s=UU"
+apply (subst corresp_ex_simC_unfold)
+apply simp
+done
+
+lemma corresp_ex_simC_nil: "(corresp_ex_simC A R$nil) s = nil"
+apply (subst corresp_ex_simC_unfold)
+apply simp
+done
+
+lemma corresp_ex_simC_cons: "(corresp_ex_simC A R$((a,t)\<leadsto>xs)) s =
+           (let T' = @t'. ? ex1. (t,t'):R & move A ex1 s a t'
+            in
+             (@cex. move A cex s a T')
+              @@ ((corresp_ex_simC A R$xs) T'))"
+apply (rule trans)
+apply (subst corresp_ex_simC_unfold)
+apply (simp add: Consq_def flift1_def)
+apply simp
+done
+
+
+declare corresp_ex_simC_UU [simp] corresp_ex_simC_nil [simp] corresp_ex_simC_cons [simp]
+
+
+subsection "properties of move"
+
+declare Let_def [simp del]
+
+lemma move_is_move_sim:
+   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
+      let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
+      (t,T'): R & move A (@ex2. move A ex2 s' a T') s' a T'"
+apply (unfold is_simulation_def)
+
+(* Does not perform conditional rewriting on assumptions automatically as
+   usual. Instantiate all variables per hand. Ask Tobias?? *)
+apply (subgoal_tac "? t' ex. (t,t') :R & move A ex s' a t'")
+prefer 2
+apply simp
+apply (erule conjE)
+apply (erule_tac x = "s" in allE)
+apply (erule_tac x = "s'" in allE)
+apply (erule_tac x = "t" in allE)
+apply (erule_tac x = "a" in allE)
+apply simp
+(* Go on as usual *)
+apply (erule exE)
+apply (drule_tac x = "t'" and P = "%t'. ? ex. (t,t') :R & move A ex s' a t'" in someI)
+apply (erule exE)
+apply (erule conjE)
+apply (simp add: Let_def)
+apply (rule_tac x = "ex" in someI)
+apply assumption
+done
+
+declare Let_def [simp]
+
+lemma move_subprop1_sim:
+   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
+    let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
+     is_exec_frag A (s',@x. move A x s' a T')"
+apply (cut_tac move_is_move_sim)
+defer
+apply assumption+
+apply (simp add: move_def)
+done
+
+lemma move_subprop2_sim:
+   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
+    let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
+    Finite (@x. move A x s' a T')"
+apply (cut_tac move_is_move_sim)
+defer
+apply assumption+
+apply (simp add: move_def)
+done
+
+lemma move_subprop3_sim:
+   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
+    let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
+     laststate (s',@x. move A x s' a T') = T'"
+apply (cut_tac move_is_move_sim)
+defer
+apply assumption+
+apply (simp add: move_def)
+done
+
+lemma move_subprop4_sim:
+   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
+    let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
+      mk_trace A$((@x. move A x s' a T')) =
+        (if a:ext A then a\<leadsto>nil else nil)"
+apply (cut_tac move_is_move_sim)
+defer
+apply assumption+
+apply (simp add: move_def)
+done
+
+lemma move_subprop5_sim:
+   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
+    let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
+      (t,T'):R"
+apply (cut_tac move_is_move_sim)
+defer
+apply assumption+
+apply (simp add: move_def)
+done
+
+
+subsection \<open>TRACE INCLUSION Part 1: Traces coincide\<close>
+
+subsubsection "Lemmata for <=="
+
+(* ------------------------------------------------------
+                 Lemma 1 :Traces coincide
+   ------------------------------------------------------- *)
+
+declare split_if [split del]
+lemma traces_coincide_sim [rule_format (no_asm)]:
+  "[|is_simulation R C A; ext C = ext A|] ==>
+         !s s'. reachable C s & is_exec_frag C (s,ex) & (s,s'): R -->
+             mk_trace C$ex = mk_trace A$((corresp_ex_simC A R$ex) s')"
+
+apply (tactic \<open>pair_induct_tac @{context} "ex" [@{thm is_exec_frag_def}] 1\<close>)
+(* cons case *)
+apply auto
+apply (rename_tac ex a t s s')
+apply (simp add: mk_traceConc)
+apply (frule reachable.reachable_n)
+apply assumption
+apply (erule_tac x = "t" in allE)
+apply (erule_tac x = "@t'. ? ex1. (t,t') :R & move A ex1 s' a t'" in allE)
+apply (simp add: move_subprop5_sim [unfolded Let_def]
+  move_subprop4_sim [unfolded Let_def] split add: split_if)
+done
+declare split_if [split]
+
+
+(* ----------------------------------------------------------- *)
+(*               Lemma 2 : corresp_ex_sim is execution         *)
+(* ----------------------------------------------------------- *)
+
+
+lemma correspsim_is_execution [rule_format (no_asm)]:
+ "[| is_simulation R C A |] ==>
+  !s s'. reachable C s & is_exec_frag C (s,ex) & (s,s'):R
+  --> is_exec_frag A (s',(corresp_ex_simC A R$ex) s')"
+
+apply (tactic \<open>pair_induct_tac @{context} "ex" [@{thm is_exec_frag_def}] 1\<close>)
+(* main case *)
+apply auto
+apply (rename_tac ex a t s s')
+apply (rule_tac t = "@t'. ? ex1. (t,t') :R & move A ex1 s' a t'" in lemma_2_1)
+
+(* Finite *)
+apply (erule move_subprop2_sim [unfolded Let_def])
+apply assumption+
+apply (rule conjI)
+
+(* is_exec_frag *)
+apply (erule move_subprop1_sim [unfolded Let_def])
+apply assumption+
+apply (rule conjI)
+
+(* Induction hypothesis  *)
+(* reachable_n looping, therefore apply it manually *)
+apply (erule_tac x = "t" in allE)
+apply (erule_tac x = "@t'. ? ex1. (t,t') :R & move A ex1 s' a t'" in allE)
+apply simp
+apply (frule reachable.reachable_n)
+apply assumption
+apply (simp add: move_subprop5_sim [unfolded Let_def])
+(* laststate *)
+apply (erule move_subprop3_sim [unfolded Let_def, symmetric])
+apply assumption+
+done
+
+
+subsection "Main Theorem: TRACE - INCLUSION"
+
+(* -------------------------------------------------------------------------------- *)
+
+  (* generate condition (s,S'):R & S':starts_of A, the first being intereting
+     for the induction cases concerning the two lemmas correpsim_is_execution and
+     traces_coincide_sim, the second for the start state case.
+     S':= @s'. (s,s'):R & s':starts_of A, where s:starts_of C  *)
+
+lemma simulation_starts:
+"[| is_simulation R C A; s:starts_of C |]
+  ==> let S' = @s'. (s,s'):R & s':starts_of A in
+      (s,S'):R & S':starts_of A"
+  apply (simp add: is_simulation_def corresp_ex_sim_def Int_non_empty Image_def)
+  apply (erule conjE)+
+  apply (erule ballE)
+  prefer 2 apply (blast)
+  apply (erule exE)
+  apply (rule someI2)
+  apply assumption
+  apply blast
+  done
+
+lemmas sim_starts1 = simulation_starts [unfolded Let_def, THEN conjunct1]
+lemmas sim_starts2 = simulation_starts [unfolded Let_def, THEN conjunct2]
+
+
+lemma trace_inclusion_for_simulations:
+  "[| ext C = ext A; is_simulation R C A |]
+           ==> traces C <= traces A"
+
+  apply (unfold traces_def)
+
+  apply (simp (no_asm) add: has_trace_def2)
+  apply auto
+
+  (* give execution of abstract automata *)
+  apply (rule_tac x = "corresp_ex_sim A R ex" in bexI)
+
+  (* Traces coincide, Lemma 1 *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (rename_tac s ex)
+  apply (simp (no_asm) add: corresp_ex_sim_def)
+  apply (rule_tac s = "s" in traces_coincide_sim)
+  apply assumption+
+  apply (simp add: executions_def reachable.reachable_0 sim_starts1)
+
+  (* corresp_ex_sim is execution, Lemma 2 *)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (simp add: executions_def)
+  apply (rename_tac s ex)
+
+  (* start state *)
+  apply (rule conjI)
+  apply (simp add: sim_starts2 corresp_ex_sim_def)
+
+  (* is-execution-fragment *)
+  apply (simp add: corresp_ex_sim_def)
+  apply (rule_tac s = s in correspsim_is_execution)
+  apply assumption
+  apply (simp add: reachable.reachable_0 sim_starts1)
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/Simulations.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,85 @@
+(*  Title:      HOL/HOLCF/IOA/Simulations.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Simulations in HOLCF/IOA\<close>
+
+theory Simulations
+imports RefCorrectness
+begin
+
+default_sort type
+
+definition
+  is_simulation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
+  "is_simulation R C A =
+   ((!s:starts_of C. R``{s} Int starts_of A ~= {}) &
+   (!s s' t a. reachable C s &
+               s \<midarrow>a\<midarrow>C\<rightarrow> t   &
+               (s,s') : R
+               --> (? t' ex. (t,t'):R & move A ex s' a t')))"
+
+definition
+  is_backward_simulation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
+  "is_backward_simulation R C A =
+   ((!s:starts_of C. R``{s} <= starts_of A) &
+   (!s t t' a. reachable C s &
+               s \<midarrow>a\<midarrow>C\<rightarrow> t   &
+               (t,t') : R
+               --> (? ex s'. (s,s'):R & move A ex s' a t')))"
+
+definition
+  is_forw_back_simulation :: "[('s1 * 's2 set)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
+  "is_forw_back_simulation R C A =
+   ((!s:starts_of C. ? S'. (s,S'):R & S'<= starts_of A) &
+   (!s S' t a. reachable C s &
+               s \<midarrow>a\<midarrow>C\<rightarrow> t   &
+               (s,S') : R
+               --> (? T'. (t,T'):R & (! t':T'. ? s':S'. ? ex. move A ex s' a t'))))"
+
+definition
+  is_back_forw_simulation :: "[('s1 * 's2 set)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
+  "is_back_forw_simulation R C A =
+   ((!s:starts_of C. ! S'. (s,S'):R --> S' Int starts_of A ~={}) &
+   (!s t T' a. reachable C s &
+               s \<midarrow>a\<midarrow>C\<rightarrow> t   &
+               (t,T') : R
+               --> (? S'. (s,S'):R & (! s':S'. ? t':T'. ? ex. move A ex s' a t'))))"
+
+definition
+  is_history_relation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
+  "is_history_relation R C A = (is_simulation R C A &
+                                is_ref_map (%x.(@y. (x,y):(R^-1))) A C)"
+
+definition
+  is_prophecy_relation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
+  "is_prophecy_relation R C A = (is_backward_simulation R C A &
+                                 is_ref_map (%x.(@y. (x,y):(R^-1))) A C)"
+
+
+lemma set_non_empty: "(A~={}) = (? x. x:A)"
+apply auto
+done
+
+lemma Int_non_empty: "(A Int B ~= {}) = (? x. x: A & x:B)"
+apply (simp add: set_non_empty)
+done
+
+
+lemma Sim_start_convert:
+"(R``{x} Int S ~= {}) = (? y. (x,y):R & y:S)"
+apply (unfold Image_def)
+apply (simp add: Int_non_empty)
+done
+
+declare Sim_start_convert [simp]
+
+
+lemma ref_map_is_simulation:
+"!! f. is_ref_map f C A ==> is_simulation {p. (snd p) = f (fst p)} C A"
+
+apply (unfold is_ref_map_def is_simulation_def)
+apply simp
+done
+
+end
--- a/src/HOL/HOLCF/IOA/Storage/Spec.thy	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/HOLCF/IOA/Storage/Spec.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -5,7 +5,7 @@
 section \<open>The specification of a memory\<close>
 
 theory Spec
-imports "~~/src/HOL/HOLCF/IOA/meta_theory/IOA" Action
+imports "~~/src/HOL/HOLCF/IOA/IOA" Action
 begin
 
 definition
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/TL.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,179 @@
+(*  Title:      HOL/HOLCF/IOA/TL.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>A General Temporal Logic\<close>
+
+theory TL
+imports Pred Sequence
+begin
+
+default_sort type
+
+type_synonym 'a temporal = "'a Seq predicate"
+
+definition validT :: "'a Seq predicate \<Rightarrow> bool"
+  where "validT P \<longleftrightarrow> (\<forall>s. s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> (s \<Turnstile> P))"
+
+definition unlift :: "'a lift \<Rightarrow> 'a"
+  where "unlift x = (case x of Def y \<Rightarrow> y)"
+
+definition Init :: "'a predicate \<Rightarrow> 'a temporal"  ("\<langle>_\<rangle>" [0] 1000)
+  where "Init P s = P (unlift (HD $ s))"
+    \<comment> \<open>this means that for \<open>nil\<close> and \<open>UU\<close> the effect is unpredictable\<close>
+
+definition Next :: "'a temporal \<Rightarrow> 'a temporal"
+  where "(Next P) s \<longleftrightarrow> (if TL $ s = UU \<or> TL $ s = nil then P s else P (TL $ s))"
+
+definition suffix :: "'a Seq \<Rightarrow> 'a Seq \<Rightarrow> bool"
+  where "suffix s2 s \<longleftrightarrow> (\<exists>s1. Finite s1 \<and> s = s1 @@ s2)"
+
+definition tsuffix :: "'a Seq \<Rightarrow> 'a Seq \<Rightarrow> bool"
+  where "tsuffix s2 s \<longleftrightarrow> s2 \<noteq> nil \<and> s2 \<noteq> UU \<and> suffix s2 s"
+
+definition Box :: "'a temporal \<Rightarrow> 'a temporal"  ("\<box>(_)" [80] 80)
+  where "(\<box>P) s \<longleftrightarrow> (\<forall>s2. tsuffix s2 s \<longrightarrow> P s2)"
+
+definition Diamond :: "'a temporal \<Rightarrow> 'a temporal"  ("\<diamond>(_)" [80] 80)
+  where "\<diamond>P = (\<^bold>\<not> (\<box>(\<^bold>\<not> P)))"
+
+definition Leadsto :: "'a temporal \<Rightarrow> 'a temporal \<Rightarrow> 'a temporal"  (infixr "\<leadsto>" 22)
+  where "(P \<leadsto> Q) = (\<box>(P \<^bold>\<longrightarrow> (\<diamond>Q)))"
+
+
+lemma simple: "\<box>\<diamond>(\<^bold>\<not> P) = (\<^bold>\<not> \<diamond>\<box>P)"
+apply (rule ext)
+apply (simp add: Diamond_def NOT_def Box_def)
+done
+
+lemma Boxnil: "nil \<Turnstile> \<box>P"
+apply (simp add: satisfies_def Box_def tsuffix_def suffix_def nil_is_Conc)
+done
+
+lemma Diamondnil: "~(nil \<Turnstile> \<diamond>P)"
+apply (simp add: Diamond_def satisfies_def NOT_def)
+apply (cut_tac Boxnil)
+apply (simp add: satisfies_def)
+done
+
+lemma Diamond_def2: "(\<diamond>F) s = (? s2. tsuffix s2 s & F s2)"
+apply (simp add: Diamond_def NOT_def Box_def)
+done
+
+
+
+subsection "TLA Axiomatization by Merz"
+
+lemma suffix_refl: "suffix s s"
+apply (simp add: suffix_def)
+apply (rule_tac x = "nil" in exI)
+apply auto
+done
+
+lemma reflT: "s~=UU & s~=nil --> (s \<Turnstile> \<box>F \<^bold>\<longrightarrow> F)"
+apply (simp add: satisfies_def IMPLIES_def Box_def)
+apply (rule impI)+
+apply (erule_tac x = "s" in allE)
+apply (simp add: tsuffix_def suffix_refl)
+done
+
+
+lemma suffix_trans: "[| suffix y x ; suffix z y |]  ==> suffix z x"
+apply (simp add: suffix_def)
+apply auto
+apply (rule_tac x = "s1 @@ s1a" in exI)
+apply auto
+apply (simp (no_asm) add: Conc_assoc)
+done
+
+lemma transT: "s \<Turnstile> \<box>F \<^bold>\<longrightarrow> \<box>\<box>F"
+apply (simp (no_asm) add: satisfies_def IMPLIES_def Box_def tsuffix_def)
+apply auto
+apply (drule suffix_trans)
+apply assumption
+apply (erule_tac x = "s2a" in allE)
+apply auto
+done
+
+
+lemma normalT: "s \<Turnstile> \<box>(F \<^bold>\<longrightarrow> G) \<^bold>\<longrightarrow> \<box>F \<^bold>\<longrightarrow> \<box>G"
+apply (simp (no_asm) add: satisfies_def IMPLIES_def Box_def)
+done
+
+
+subsection "TLA Rules by Lamport"
+
+lemma STL1a: "validT P ==> validT (\<box>P)"
+apply (simp add: validT_def satisfies_def Box_def tsuffix_def)
+done
+
+lemma STL1b: "valid P ==> validT (Init P)"
+apply (simp add: valid_def validT_def satisfies_def Init_def)
+done
+
+lemma STL1: "valid P ==> validT (\<box>(Init P))"
+apply (rule STL1a)
+apply (erule STL1b)
+done
+
+(* Note that unlift and HD is not at all used !!! *)
+lemma STL4: "valid (P \<^bold>\<longrightarrow> Q)  ==> validT (\<box>(Init P) \<^bold>\<longrightarrow> \<box>(Init Q))"
+apply (simp add: valid_def validT_def satisfies_def IMPLIES_def Box_def Init_def)
+done
+
+
+subsection "LTL Axioms by Manna/Pnueli"
+
+lemma tsuffix_TL [rule_format (no_asm)]:
+"s~=UU & s~=nil --> tsuffix s2 (TL$s) --> tsuffix s2 s"
+apply (unfold tsuffix_def suffix_def)
+apply auto
+apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
+apply (rule_tac x = "a\<leadsto>s1" in exI)
+apply auto
+done
+
+lemmas tsuffix_TL2 = conjI [THEN tsuffix_TL]
+
+declare split_if [split del]
+lemma LTL1:
+   "s~=UU & s~=nil --> (s \<Turnstile> \<box>F \<^bold>\<longrightarrow> (F \<^bold>\<and> (Next (\<box>F))))"
+apply (unfold Next_def satisfies_def NOT_def IMPLIES_def AND_def Box_def)
+apply auto
+(* \<box>F \<^bold>\<longrightarrow> F *)
+apply (erule_tac x = "s" in allE)
+apply (simp add: tsuffix_def suffix_refl)
+(* \<box>F \<^bold>\<longrightarrow> Next \<box>F *)
+apply (simp split add: split_if)
+apply auto
+apply (drule tsuffix_TL2)
+apply assumption+
+apply auto
+done
+declare split_if [split]
+
+
+lemma LTL2a:
+    "s \<Turnstile> \<^bold>\<not> (Next F) \<^bold>\<longrightarrow> (Next (\<^bold>\<not> F))"
+apply (unfold Next_def satisfies_def NOT_def IMPLIES_def)
+apply simp
+done
+
+lemma LTL2b:
+    "s \<Turnstile> (Next (\<^bold>\<not> F)) \<^bold>\<longrightarrow> (\<^bold>\<not> (Next F))"
+apply (unfold Next_def satisfies_def NOT_def IMPLIES_def)
+apply simp
+done
+
+lemma LTL3:
+"ex \<Turnstile> (Next (F \<^bold>\<longrightarrow> G)) \<^bold>\<longrightarrow> (Next F) \<^bold>\<longrightarrow> (Next G)"
+apply (unfold Next_def satisfies_def NOT_def IMPLIES_def)
+apply simp
+done
+
+
+lemma ModusPonens: "[| validT (P \<^bold>\<longrightarrow> Q); validT P |] ==> validT Q"
+apply (simp add: validT_def satisfies_def IMPLIES_def)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/TLS.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,176 @@
+(*  Title:      HOL/HOLCF/IOA/TLS.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Temporal Logic of Steps -- tailored for I/O automata\<close>
+
+theory TLS
+imports IOA TL
+begin
+
+default_sort type
+
+type_synonym ('a, 's) ioa_temp  = "('a option, 's) transition temporal"
+
+type_synonym ('a, 's) step_pred = "('a option, 's) transition predicate"
+
+type_synonym 's state_pred = "'s predicate"
+
+definition mkfin :: "'a Seq \<Rightarrow> 'a Seq"
+  where "mkfin s = (if Partial s then SOME t. Finite t \<and> s = t @@ UU else s)"
+
+definition option_lift :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a option \<Rightarrow> 'b"
+  where "option_lift f s y = (case y of None \<Rightarrow> s | Some x \<Rightarrow> f x)"
+
+definition plift :: "('a \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> bool"
+(* plift is used to determine that None action is always false in
+   transition predicates *)
+  where "plift P = option_lift P False"
+
+definition xt1 :: "'s predicate \<Rightarrow> ('a, 's) step_pred"
+  where "xt1 P tr = P (fst tr)"
+
+definition xt2 :: "'a option predicate \<Rightarrow> ('a, 's) step_pred"
+  where "xt2 P tr = P (fst (snd tr))"
+
+definition ex2seqC :: "('a, 's) pairs \<rightarrow> ('s \<Rightarrow> ('a option, 's) transition Seq)"
+where
+  "ex2seqC = (fix$(LAM h ex. (%s. case ex of
+      nil =>  (s,None,s)\<leadsto>nil
+    | x##xs => (flift1 (%pr.
+                (s,Some (fst pr), snd pr)\<leadsto> (h$xs) (snd pr))
+                $x)
+      )))"
+
+definition ex2seq :: "('a, 's) execution \<Rightarrow> ('a option, 's) transition Seq"
+  where "ex2seq ex = (ex2seqC $ (mkfin (snd ex))) (fst ex)"
+
+definition temp_sat :: "('a, 's) execution \<Rightarrow> ('a, 's) ioa_temp \<Rightarrow> bool"  (infixr "\<TTurnstile>" 22)
+  where "(ex \<TTurnstile> P) \<longleftrightarrow> ((ex2seq ex) \<Turnstile> P)"
+
+definition validTE :: "('a, 's) ioa_temp \<Rightarrow> bool"
+  where "validTE P \<longleftrightarrow> (\<forall>ex. (ex \<TTurnstile> P))"
+
+definition validIOA :: "('a, 's) ioa \<Rightarrow> ('a, 's) ioa_temp \<Rightarrow> bool"
+  where "validIOA A P \<longleftrightarrow> (\<forall>ex \<in> executions A. (ex \<TTurnstile> P))"
+
+
+axiomatization
+where
+
+mkfin_UU:
+  "mkfin UU = nil" and
+
+mkfin_nil:
+  "mkfin nil =nil" and
+
+mkfin_cons:
+  "(mkfin (a\<leadsto>s)) = (a\<leadsto>(mkfin s))"
+
+
+lemmas [simp del] = HOL.ex_simps HOL.all_simps split_paired_Ex
+
+setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close>
+
+
+subsection \<open>ex2seqC\<close>
+
+lemma ex2seqC_unfold: "ex2seqC  = (LAM ex. (%s. case ex of
+       nil =>  (s,None,s)\<leadsto>nil
+     | x##xs => (flift1 (%pr.
+                 (s,Some (fst pr), snd pr)\<leadsto> (ex2seqC$xs) (snd pr))
+                 $x)
+       ))"
+  apply (rule trans)
+  apply (rule fix_eq4)
+  apply (rule ex2seqC_def)
+  apply (rule beta_cfun)
+  apply (simp add: flift1_def)
+  done
+
+lemma ex2seqC_UU: "(ex2seqC $UU) s=UU"
+  apply (subst ex2seqC_unfold)
+  apply simp
+  done
+
+lemma ex2seqC_nil: "(ex2seqC $nil) s = (s,None,s)\<leadsto>nil"
+  apply (subst ex2seqC_unfold)
+  apply simp
+  done
+
+lemma ex2seqC_cons: "(ex2seqC $((a,t)\<leadsto>xs)) s = (s,Some a,t)\<leadsto> ((ex2seqC$xs) t)"
+  apply (rule trans)
+  apply (subst ex2seqC_unfold)
+  apply (simp add: Consq_def flift1_def)
+  apply (simp add: Consq_def flift1_def)
+  done
+
+declare ex2seqC_UU [simp] ex2seqC_nil [simp] ex2seqC_cons [simp]
+
+
+
+declare mkfin_UU [simp] mkfin_nil [simp] mkfin_cons [simp]
+
+lemma ex2seq_UU: "ex2seq (s, UU) = (s,None,s)\<leadsto>nil"
+  by (simp add: ex2seq_def)
+
+lemma ex2seq_nil: "ex2seq (s, nil) = (s,None,s)\<leadsto>nil"
+  by (simp add: ex2seq_def)
+
+lemma ex2seq_cons: "ex2seq (s, (a,t)\<leadsto>ex) = (s,Some a,t) \<leadsto> ex2seq (t, ex)"
+  by (simp add: ex2seq_def)
+
+declare ex2seqC_UU [simp del] ex2seqC_nil [simp del] ex2seqC_cons [simp del]
+declare ex2seq_UU [simp] ex2seq_nil [simp] ex2seq_cons [simp]
+
+
+lemma ex2seq_nUUnnil: "ex2seq exec ~= UU & ex2seq exec ~= nil"
+  apply (tactic \<open>pair_tac @{context} "exec" 1\<close>)
+  apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
+  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+  done
+
+
+subsection \<open>Interface TL -- TLS\<close>
+
+(* uses the fact that in executions states overlap, which is lost in
+   after the translation via ex2seq !! *)
+
+lemma TL_TLS:
+  "[| ! s a t. (P s) & s \<midarrow>a\<midarrow>A\<rightarrow> t --> (Q t) |]
+    ==> ex \<TTurnstile> (Init (%(s,a,t). P s) \<^bold>\<and> Init (%(s,a,t). s \<midarrow>a\<midarrow>A\<rightarrow> t)
+              \<^bold>\<longrightarrow> (Next (Init (%(s,a,t).Q s))))"
+  apply (unfold Init_def Next_def temp_sat_def satisfies_def IMPLIES_def AND_def)
+
+  apply clarify
+  apply (simp split add: split_if)
+  (* TL = UU *)
+  apply (rule conjI)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
+  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+  apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
+  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+  (* TL = nil *)
+  apply (rule conjI)
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (tactic \<open>Seq_case_tac @{context} "x2" 1\<close>)
+  apply (simp add: unlift_def)
+  apply fast
+  apply (simp add: unlift_def)
+  apply fast
+  apply (simp add: unlift_def)
+  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+  apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
+  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+  (* TL =cons *)
+  apply (simp add: unlift_def)
+
+  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
+  apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
+  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+  apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
+  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/Traces.thy	Thu Dec 31 12:43:09 2015 +0100
@@ -0,0 +1,356 @@
+(*  Title:      HOL/HOLCF/IOA/Traces.thy
+    Author:     Olaf Müller
+*)
+
+section \<open>Executions and Traces of I/O automata in HOLCF\<close>
+
+theory Traces
+imports Sequence Automata
+begin
+
+default_sort type
+
+type_synonym ('a, 's) pairs = "('a * 's) Seq"
+type_synonym ('a, 's) execution = "'s * ('a, 's) pairs"
+type_synonym 'a trace = "'a Seq"
+type_synonym ('a, 's) execution_module = "('a, 's) execution set * 'a signature"
+type_synonym 'a schedule_module = "'a trace set * 'a signature"
+type_synonym 'a trace_module = "'a trace set * 'a signature"
+
+
+(*  ------------------- Executions ------------------------------ *)
+
+definition is_exec_fragC :: "('a, 's) ioa \<Rightarrow> ('a, 's) pairs \<rightarrow> 's \<Rightarrow> tr"
+where
+  "is_exec_fragC A = (fix $ (LAM h ex. (%s. case ex of
+      nil => TT
+    | x##xs => (flift1
+            (%p. Def ((s,p):trans_of A) andalso (h$xs) (snd p))
+             $x)
+   )))"
+
+definition is_exec_frag :: "[('a,'s)ioa, ('a,'s)execution] \<Rightarrow> bool"
+  where "is_exec_frag A ex = ((is_exec_fragC A $ (snd ex)) (fst ex) ~= FF)"
+
+definition executions :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution set"
+  where "executions ioa = {e. ((fst e) \<in> starts_of(ioa)) \<and> is_exec_frag ioa e}"
+
+
+(*  ------------------- Schedules ------------------------------ *)
+
+definition filter_act :: "('a, 's) pairs \<rightarrow> 'a trace"
+  where "filter_act = Map fst"
+
+definition has_schedule :: "[('a, 's) ioa, 'a trace] \<Rightarrow> bool"
+  where "has_schedule ioa sch \<longleftrightarrow> (\<exists>ex \<in> executions ioa. sch = filter_act $ (snd ex))"
+
+definition schedules :: "('a, 's) ioa \<Rightarrow> 'a trace set"
+  where "schedules ioa = {sch. has_schedule ioa sch}"
+
+
+(*  ------------------- Traces ------------------------------ *)
+
+definition has_trace :: "[('a, 's) ioa, 'a trace] \<Rightarrow> bool"
+  where "has_trace ioa tr = (\<exists>sch \<in> schedules ioa. tr = Filter (\<lambda>a. a \<in> ext ioa) $ sch)"
+
+definition traces :: "('a, 's) ioa \<Rightarrow> 'a trace set"
+  where "traces ioa \<equiv> {tr. has_trace ioa tr}"
+
+definition mk_trace :: "('a, 's) ioa \<Rightarrow> ('a, 's) pairs \<rightarrow> 'a trace"
+  where "mk_trace ioa = (LAM tr. Filter (\<lambda>a. a \<in> ext ioa) $ (filter_act $ tr))"
+
+
+(*  ------------------- Fair Traces ------------------------------ *)
+
+definition laststate :: "('a, 's) execution \<Rightarrow> 's"
+where
+  "laststate ex = (case Last $ (snd ex) of
+                    UU  => fst ex
+                  | Def at => snd at)"
+
+(* A predicate holds infinitely (finitely) often in a sequence *)
+
+definition inf_often :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<Rightarrow> bool"
+  where "inf_often P s \<longleftrightarrow> Infinite (Filter P $ s)"
+
+(*  filtering P yields a finite or partial sequence *)
+definition fin_often :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<Rightarrow> bool"
+  where "fin_often P s \<longleftrightarrow> \<not> inf_often P s"
+
+
+(* fairness of executions *)
+
+(* Note that partial execs cannot be wfair as the inf_often predicate in the
+   else branch prohibits it. However they can be sfair in the case when all W
+   are only finitely often enabled: Is this the right model?
+   See LiveIOA for solution conforming with the literature and superseding this one *)
+definition is_wfair :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> ('a, 's) execution \<Rightarrow> bool"
+where
+  "is_wfair A W ex \<longleftrightarrow>
+    (inf_often (\<lambda>x. fst x \<in> W) (snd ex) \<or> inf_often (\<lambda>x. \<not> Enabled A W (snd x)) (snd ex))"
+definition wfair_ex :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution \<Rightarrow> bool"
+where
+  "wfair_ex A ex \<longleftrightarrow> (\<forall>W \<in> wfair_of A.
+                      if Finite (snd ex)
+                      then \<not> Enabled A W (laststate ex)
+                      else is_wfair A W ex)"
+
+definition is_sfair :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> ('a, 's) execution \<Rightarrow> bool"
+where
+  "is_sfair A W ex \<longleftrightarrow>
+    (inf_often (\<lambda>x. fst x:W) (snd ex) \<or> fin_often (\<lambda>x. Enabled A W (snd x)) (snd ex))"
+definition sfair_ex :: "('a, 's)ioa \<Rightarrow> ('a, 's) execution \<Rightarrow> bool"
+where
+  "sfair_ex A ex \<longleftrightarrow> (\<forall>W \<in> sfair_of A.
+                      if   Finite (snd ex)
+                      then ~Enabled A W (laststate ex)
+                      else is_sfair A W ex)"
+
+definition fair_ex :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution \<Rightarrow> bool"
+  where "fair_ex A ex \<longleftrightarrow> wfair_ex A ex \<and> sfair_ex A ex"
+
+
+(* fair behavior sets *)
+
+definition fairexecutions :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution set"
+  where "fairexecutions A = {ex. ex \<in> executions A \<and> fair_ex A ex}"
+
+definition fairtraces :: "('a, 's) ioa \<Rightarrow> 'a trace set"
+  where "fairtraces A = {mk_trace A $ (snd ex) | ex. ex \<in> fairexecutions A}"
+
+
+(*  ------------------- Implementation ------------------------------ *)
+
+(* Notions of implementation *)
+
+definition ioa_implements :: "[('a, 's1) ioa, ('a, 's2) ioa] \<Rightarrow> bool"  (infixr "=<|" 12)
+where
+  "(ioa1 =<| ioa2) \<longleftrightarrow>
+    (((inputs(asig_of(ioa1)) = inputs(asig_of(ioa2))) \<and>
+     (outputs(asig_of(ioa1)) = outputs(asig_of(ioa2)))) \<and>
+      traces(ioa1) \<subseteq> traces(ioa2))"
+
+definition fair_implements :: "('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool"
+where
+  "fair_implements C A \<longleftrightarrow> inp C = inp A \<and> out C = out A \<and> fairtraces C \<subseteq> fairtraces A"
+
+
+(*  ------------------- Modules ------------------------------ *)
+
+(* Execution, schedule and trace modules *)
+
+definition Execs :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution_module"
+  where "Execs A = (executions A, asig_of A)"
+
+definition Scheds :: "('a, 's) ioa \<Rightarrow> 'a schedule_module"
+  where "Scheds A = (schedules A, asig_of A)"
+
+definition Traces :: "('a, 's) ioa \<Rightarrow> 'a trace_module"
+  where "Traces A = (traces A, asig_of A)"
+
+
+lemmas [simp del] = HOL.ex_simps HOL.all_simps split_paired_Ex
+declare Let_def [simp]
+setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close>
+
+lemmas exec_rws = executions_def is_exec_frag_def
+
+
+
+subsection "recursive equations of operators"
+
+(* ---------------------------------------------------------------- *)
+(*                               filter_act                         *)
+(* ---------------------------------------------------------------- *)
+
+
+lemma filter_act_UU: "filter_act$UU = UU"
+  by (simp add: filter_act_def)
+
+lemma filter_act_nil: "filter_act$nil = nil"
+  by (simp add: filter_act_def)
+
+lemma filter_act_cons: "filter_act$(x\<leadsto>xs) = (fst x) \<leadsto> filter_act$xs"
+  by (simp add: filter_act_def)
+
+declare filter_act_UU [simp] filter_act_nil [simp] filter_act_cons [simp]
+
+
+(* ---------------------------------------------------------------- *)
+(*                             mk_trace                             *)
+(* ---------------------------------------------------------------- *)
+
+lemma mk_trace_UU: "mk_trace A$UU=UU"
+  by (simp add: mk_trace_def)
+
+lemma mk_trace_nil: "mk_trace A$nil=nil"
+  by (simp add: mk_trace_def)
+
+lemma mk_trace_cons: "mk_trace A$(at \<leadsto> xs) =
+             (if ((fst at):ext A)
+                  then (fst at) \<leadsto> (mk_trace A$xs)
+                  else mk_trace A$xs)"
+
+  by (simp add: mk_trace_def)
+
+declare mk_trace_UU [simp] mk_trace_nil [simp] mk_trace_cons [simp]
+
+(* ---------------------------------------------------------------- *)
+(*                             is_exec_fragC                             *)
+(* ---------------------------------------------------------------- *)
+
+
+lemma is_exec_fragC_unfold: "is_exec_fragC A = (LAM ex. (%s. case ex of
+       nil => TT
+     | x##xs => (flift1
+             (%p. Def ((s,p):trans_of A) andalso (is_exec_fragC A$xs) (snd p))
+              $x)
+    ))"
+  apply (rule trans)
+  apply (rule fix_eq4)
+  apply (rule is_exec_fragC_def)
+  apply (rule beta_cfun)
+  apply (simp add: flift1_def)
+  done
+
+lemma is_exec_fragC_UU: "(is_exec_fragC A$UU) s=UU"
+  apply (subst is_exec_fragC_unfold)
+  apply simp
+  done
+
+lemma is_exec_fragC_nil: "(is_exec_fragC A$nil) s = TT"
+  apply (subst is_exec_fragC_unfold)
+  apply simp
+  done
+
+lemma is_exec_fragC_cons: "(is_exec_fragC A$(pr\<leadsto>xs)) s =
+                         (Def ((s,pr):trans_of A)
+                 andalso (is_exec_fragC A$xs)(snd pr))"
+  apply (rule trans)
+  apply (subst is_exec_fragC_unfold)
+  apply (simp add: Consq_def flift1_def)
+  apply simp
+  done
+
+
+declare is_exec_fragC_UU [simp] is_exec_fragC_nil [simp] is_exec_fragC_cons [simp]
+
+
+(* ---------------------------------------------------------------- *)
+(*                        is_exec_frag                              *)
+(* ---------------------------------------------------------------- *)
+
+lemma is_exec_frag_UU: "is_exec_frag A (s, UU)"
+  by (simp add: is_exec_frag_def)
+
+lemma is_exec_frag_nil: "is_exec_frag A (s, nil)"
+  by (simp add: is_exec_frag_def)
+
+lemma is_exec_frag_cons: "is_exec_frag A (s, (a,t)\<leadsto>ex) =
+                                (((s,a,t):trans_of A) &
+                                is_exec_frag A (t, ex))"
+  by (simp add: is_exec_frag_def)
+
+
+(* Delsimps [is_exec_fragC_UU,is_exec_fragC_nil,is_exec_fragC_cons]; *)
+declare is_exec_frag_UU [simp] is_exec_frag_nil [simp] is_exec_frag_cons [simp]
+
+(* ---------------------------------------------------------------------------- *)
+                           section "laststate"
+(* ---------------------------------------------------------------------------- *)
+
+lemma laststate_UU: "laststate (s,UU) = s"
+  by (simp add: laststate_def)
+
+lemma laststate_nil: "laststate (s,nil) = s"
+  by (simp add: laststate_def)
+
+lemma laststate_cons: "!! ex. Finite ex ==> laststate (s,at\<leadsto>ex) = laststate (snd at,ex)"
+  apply (simp (no_asm) add: laststate_def)
+  apply (case_tac "ex=nil")
+  apply (simp (no_asm_simp))
+  apply (simp (no_asm_simp))
+  apply (drule Finite_Last1 [THEN mp])
+  apply assumption
+  apply defined
+  done
+
+declare laststate_UU [simp] laststate_nil [simp] laststate_cons [simp]
+
+lemma exists_laststate: "!!ex. Finite ex ==> (! s. ? u. laststate (s,ex)=u)"
+  apply (tactic "Seq_Finite_induct_tac @{context} 1")
+  done
+
+
+subsection "has_trace, mk_trace"
+
+(* alternative definition of has_trace tailored for the refinement proof, as it does not
+   take the detour of schedules *)
+
+lemma has_trace_def2:
+  "has_trace A b = (? ex:executions A. b = mk_trace A$(snd ex))"
+  apply (unfold executions_def mk_trace_def has_trace_def schedules_def has_schedule_def [abs_def])
+  apply auto
+  done
+
+
+subsection "signatures and executions, schedules"
+
+(* All executions of A have only actions of A. This is only true because of the
+   predicate state_trans (part of the predicate IOA): We have no dependent types.
+   For executions of parallel automata this assumption is not needed, as in par_def
+   this condition is included once more. (see Lemmas 1.1.1c in CompoExecs for example) *)
+
+lemma execfrag_in_sig:
+  "!! A. is_trans_of A ==>
+  ! s. is_exec_frag A (s,xs) --> Forall (%a. a:act A) (filter_act$xs)"
+  apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def},
+    @{thm Forall_def}, @{thm sforall_def}] 1\<close>)
+  (* main case *)
+  apply (auto simp add: is_trans_of_def)
+  done
+
+lemma exec_in_sig:
+  "!! A.[|  is_trans_of A; x:executions A |] ==>
+  Forall (%a. a:act A) (filter_act$(snd x))"
+  apply (simp add: executions_def)
+  apply (tactic \<open>pair_tac @{context} "x" 1\<close>)
+  apply (rule execfrag_in_sig [THEN spec, THEN mp])
+  apply auto
+  done
+
+lemma scheds_in_sig:
+  "!! A.[|  is_trans_of A; x:schedules A |] ==>
+    Forall (%a. a:act A) x"
+  apply (unfold schedules_def has_schedule_def [abs_def])
+  apply (fast intro!: exec_in_sig)
+  done
+
+
+subsection "executions are prefix closed"
+
+(* only admissible in y, not if done in x !! *)
+lemma execfrag_prefixclosed: "!x s. is_exec_frag A (s,x) & y<<x  --> is_exec_frag A (s,y)"
+  apply (tactic \<open>pair_induct_tac @{context} "y" [@{thm is_exec_frag_def}] 1\<close>)
+  apply (intro strip)
+  apply (tactic \<open>Seq_case_simp_tac @{context} "x" 1\<close>)
+  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+  apply auto
+  done
+
+lemmas exec_prefixclosed =
+  conjI [THEN execfrag_prefixclosed [THEN spec, THEN spec, THEN mp]]
+
+
+(* second prefix notion for Finite x *)
+
+lemma exec_prefix2closed [rule_format]:
+  "! y s. is_exec_frag A (s,x@@y) --> is_exec_frag A (s,x)"
+  apply (tactic \<open>pair_induct_tac @{context} "x" [@{thm is_exec_frag_def}] 1\<close>)
+  apply (intro strip)
+  apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
+  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
+  apply auto
+  done
+
+end
--- a/src/HOL/HOLCF/IOA/meta_theory/Abstraction.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,614 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/Abstraction.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Abstraction Theory -- tailored for I/O automata\<close>
-
-theory Abstraction
-imports LiveIOA
-begin
-
-default_sort type
-
-definition
-  cex_abs :: "('s1 => 's2) => ('a,'s1)execution => ('a,'s2)execution" where
-  "cex_abs f ex = (f (fst ex), Map (%(a,t). (a,f t))$(snd ex))"
-definition
-  \<comment> \<open>equals cex_abs on Sequences -- after ex2seq application\<close>
-  cex_absSeq :: "('s1 => 's2) => ('a option,'s1)transition Seq => ('a option,'s2)transition Seq" where
-  "cex_absSeq f = (%s. Map (%(s,a,t). (f s,a,f t))$s)"
-
-definition
-  is_abstraction ::"[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool" where
-  "is_abstraction f C A =
-   ((!s:starts_of(C). f(s):starts_of(A)) &
-   (!s t a. reachable C s & s \<midarrow>a\<midarrow>C\<rightarrow> t
-            --> (f s) \<midarrow>a\<midarrow>A\<rightarrow> (f t)))"
-
-definition
-  weakeningIOA :: "('a,'s2)ioa => ('a,'s1)ioa => ('s1 => 's2) => bool" where
-  "weakeningIOA A C h = (!ex. ex : executions C --> cex_abs h ex : executions A)"
-definition
-  temp_strengthening :: "('a,'s2)ioa_temp => ('a,'s1)ioa_temp => ('s1 => 's2) => bool" where
-  "temp_strengthening Q P h = (!ex. (cex_abs h ex \<TTurnstile> Q) --> (ex \<TTurnstile> P))"
-definition
-  temp_weakening :: "('a,'s2)ioa_temp => ('a,'s1)ioa_temp => ('s1 => 's2) => bool" where
-  "temp_weakening Q P h = temp_strengthening (\<^bold>\<not> Q) (\<^bold>\<not> P) h"
-
-definition
-  state_strengthening :: "('a,'s2)step_pred => ('a,'s1)step_pred => ('s1 => 's2) => bool" where
-  "state_strengthening Q P h = (!s t a.  Q (h(s),a,h(t)) --> P (s,a,t))"
-definition
-  state_weakening :: "('a,'s2)step_pred => ('a,'s1)step_pred => ('s1 => 's2) => bool" where
-  "state_weakening Q P h = state_strengthening (\<^bold>\<not>Q) (\<^bold>\<not>P) h"
-
-definition
-  is_live_abstraction :: "('s1 => 's2) => ('a,'s1)live_ioa => ('a,'s2)live_ioa => bool" where
-  "is_live_abstraction h CL AM =
-     (is_abstraction h (fst CL) (fst AM) &
-      temp_weakening (snd AM) (snd CL) h)"
-
-
-axiomatization where
-(* thm about ex2seq which is not provable by induction as ex2seq is not continous *)
-ex2seq_abs_cex:
-  "ex2seq (cex_abs h ex) = cex_absSeq h (ex2seq ex)"
-
-axiomatization where
-(* analog to the proved thm strength_Box - proof skipped as trivial *)
-weak_Box:
-"temp_weakening P Q h
- ==> temp_weakening (\<box>P) (\<box>Q) h"
-
-axiomatization where
-(* analog to the proved thm strength_Next - proof skipped as trivial *)
-weak_Next:
-"temp_weakening P Q h
- ==> temp_weakening (Next P) (Next Q) h"
-
-
-subsection "cex_abs"
-
-lemma cex_abs_UU: "cex_abs f (s,UU) = (f s, UU)"
-  by (simp add: cex_abs_def)
-
-lemma cex_abs_nil: "cex_abs f (s,nil) = (f s, nil)"
-  by (simp add: cex_abs_def)
-
-lemma cex_abs_cons: "cex_abs f (s,(a,t)\<leadsto>ex) = (f s, (a,f t) \<leadsto> (snd (cex_abs f (t,ex))))"
-  by (simp add: cex_abs_def)
-
-declare cex_abs_UU [simp] cex_abs_nil [simp] cex_abs_cons [simp]
-
-
-subsection "lemmas"
-
-lemma temp_weakening_def2: "temp_weakening Q P h = (! ex. (ex \<TTurnstile> P) --> (cex_abs h ex \<TTurnstile> Q))"
-  apply (simp add: temp_weakening_def temp_strengthening_def NOT_def temp_sat_def satisfies_def)
-  apply auto
-  done
-
-lemma state_weakening_def2: "state_weakening Q P h = (! s t a. P (s,a,t) --> Q (h(s),a,h(t)))"
-  apply (simp add: state_weakening_def state_strengthening_def NOT_def)
-  apply auto
-  done
-
-
-subsection "Abstraction Rules for Properties"
-
-lemma exec_frag_abstraction [rule_format]:
- "[| is_abstraction h C A |] ==>
-  !s. reachable C s & is_exec_frag C (s,xs)
-  --> is_exec_frag A (cex_abs h (s,xs))"
-apply (unfold cex_abs_def)
-apply simp
-apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1\<close>)
-txt \<open>main case\<close>
-apply (auto dest: reachable.reachable_n simp add: is_abstraction_def)
-done
-
-
-lemma abs_is_weakening: "is_abstraction h C A ==> weakeningIOA A C h"
-apply (simp add: weakeningIOA_def)
-apply auto
-apply (simp add: executions_def)
-txt \<open>start state\<close>
-apply (rule conjI)
-apply (simp add: is_abstraction_def cex_abs_def)
-txt \<open>is-execution-fragment\<close>
-apply (erule exec_frag_abstraction)
-apply (simp add: reachable.reachable_0)
-done
-
-
-lemma AbsRuleT1: "[|is_abstraction h C A; validIOA A Q; temp_strengthening Q P h |]
-          ==> validIOA C P"
-apply (drule abs_is_weakening)
-apply (simp add: weakeningIOA_def validIOA_def temp_strengthening_def)
-apply (auto simp add: split_paired_all)
-done
-
-
-(* FIX: Nach TLS.ML *)
-
-lemma IMPLIES_temp_sat: "(ex \<TTurnstile> P \<^bold>\<longrightarrow> Q) = ((ex \<TTurnstile> P) --> (ex \<TTurnstile> Q))"
-  by (simp add: IMPLIES_def temp_sat_def satisfies_def)
-
-lemma AND_temp_sat: "(ex \<TTurnstile> P \<^bold>\<and> Q) = ((ex \<TTurnstile> P) & (ex \<TTurnstile> Q))"
-  by (simp add: AND_def temp_sat_def satisfies_def)
-
-lemma OR_temp_sat: "(ex \<TTurnstile> P \<^bold>\<or> Q) = ((ex \<TTurnstile> P) | (ex \<TTurnstile> Q))"
-  by (simp add: OR_def temp_sat_def satisfies_def)
-
-lemma NOT_temp_sat: "(ex \<TTurnstile> \<^bold>\<not> P) = (~ (ex \<TTurnstile> P))"
-  by (simp add: NOT_def temp_sat_def satisfies_def)
-
-declare IMPLIES_temp_sat [simp] AND_temp_sat [simp] OR_temp_sat [simp] NOT_temp_sat [simp]
-
-
-lemma AbsRuleT2:
-   "[|is_live_abstraction h (C,L) (A,M);
-          validLIOA (A,M) Q;  temp_strengthening Q P h |]
-          ==> validLIOA (C,L) P"
-apply (unfold is_live_abstraction_def)
-apply auto
-apply (drule abs_is_weakening)
-apply (simp add: weakeningIOA_def temp_weakening_def2 validLIOA_def validIOA_def temp_strengthening_def)
-apply (auto simp add: split_paired_all)
-done
-
-
-lemma AbsRuleTImprove:
-   "[|is_live_abstraction h (C,L) (A,M);
-          validLIOA (A,M) (H1 \<^bold>\<longrightarrow> Q);  temp_strengthening Q P h;
-          temp_weakening H1 H2 h; validLIOA (C,L) H2 |]
-          ==> validLIOA (C,L) P"
-apply (unfold is_live_abstraction_def)
-apply auto
-apply (drule abs_is_weakening)
-apply (simp add: weakeningIOA_def temp_weakening_def2 validLIOA_def validIOA_def temp_strengthening_def)
-apply (auto simp add: split_paired_all)
-done
-
-
-subsection "Correctness of safe abstraction"
-
-lemma abstraction_is_ref_map:
-"is_abstraction h C A ==> is_ref_map h C A"
-apply (unfold is_abstraction_def is_ref_map_def)
-apply auto
-apply (rule_tac x = "(a,h t) \<leadsto>nil" in exI)
-apply (simp add: move_def)
-done
-
-
-lemma abs_safety: "[| inp(C)=inp(A); out(C)=out(A);
-                   is_abstraction h C A |]
-                ==> C =<| A"
-apply (simp add: ioa_implements_def)
-apply (rule trace_inclusion)
-apply (simp (no_asm) add: externals_def)
-apply (auto)[1]
-apply (erule abstraction_is_ref_map)
-done
-
-
-subsection "Correctness of life abstraction"
-
-(* Reduces to Filter (Map fst x) = Filter (Map fst (Map (%(a,t). (a,x)) x),
-   that is to special Map Lemma *)
-lemma traces_coincide_abs:
-  "ext C = ext A
-         ==> mk_trace C$xs = mk_trace A$(snd (cex_abs f (s,xs)))"
-apply (unfold cex_abs_def mk_trace_def filter_act_def)
-apply simp
-apply (tactic \<open>pair_induct_tac @{context} "xs" [] 1\<close>)
-done
-
-
-(* Does not work with abstraction_is_ref_map as proof of abs_safety, because
-   is_live_abstraction includes temp_strengthening which is necessarily based
-   on cex_abs and not on corresp_ex. Thus, the proof is redoone in a more specific
-   way for cex_abs *)
-lemma abs_liveness: "[| inp(C)=inp(A); out(C)=out(A);
-                   is_live_abstraction h (C,M) (A,L) |]
-                ==> live_implements (C,M) (A,L)"
-apply (simp add: is_live_abstraction_def live_implements_def livetraces_def liveexecutions_def)
-apply auto
-apply (rule_tac x = "cex_abs h ex" in exI)
-apply auto
-  (* Traces coincide *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (rule traces_coincide_abs)
-  apply (simp (no_asm) add: externals_def)
-  apply (auto)[1]
-
-  (* cex_abs is execution *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (simp add: executions_def)
-  (* start state *)
-  apply (rule conjI)
-  apply (simp add: is_abstraction_def cex_abs_def)
-  (* is-execution-fragment *)
-  apply (erule exec_frag_abstraction)
-  apply (simp add: reachable.reachable_0)
-
- (* Liveness *)
-apply (simp add: temp_weakening_def2)
- apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-done
-
-(* FIX: NAch Traces.ML bringen *)
-
-lemma implements_trans:
-"[| A =<| B; B =<| C|] ==> A =<| C"
-by (auto simp add: ioa_implements_def)
-
-
-subsection "Abstraction Rules for Automata"
-
-lemma AbsRuleA1: "[| inp(C)=inp(A); out(C)=out(A);
-                   inp(Q)=inp(P); out(Q)=out(P);
-                   is_abstraction h1 C A;
-                   A =<| Q ;
-                   is_abstraction h2 Q P |]
-                ==> C =<| P"
-apply (drule abs_safety)
-apply assumption+
-apply (drule abs_safety)
-apply assumption+
-apply (erule implements_trans)
-apply (erule implements_trans)
-apply assumption
-done
-
-
-lemma AbsRuleA2: "!!LC. [| inp(C)=inp(A); out(C)=out(A);
-                   inp(Q)=inp(P); out(Q)=out(P);
-                   is_live_abstraction h1 (C,LC) (A,LA);
-                   live_implements (A,LA) (Q,LQ) ;
-                   is_live_abstraction h2 (Q,LQ) (P,LP) |]
-                ==> live_implements (C,LC) (P,LP)"
-apply (drule abs_liveness)
-apply assumption+
-apply (drule abs_liveness)
-apply assumption+
-apply (erule live_implements_trans)
-apply (erule live_implements_trans)
-apply assumption
-done
-
-
-declare split_paired_All [simp del]
-
-
-subsection "Localizing Temporal Strengthenings and Weakenings"
-
-lemma strength_AND:
-"[| temp_strengthening P1 Q1 h;
-          temp_strengthening P2 Q2 h |]
-       ==> temp_strengthening (P1 \<^bold>\<and> P2) (Q1 \<^bold>\<and> Q2) h"
-apply (unfold temp_strengthening_def)
-apply auto
-done
-
-lemma strength_OR:
-"[| temp_strengthening P1 Q1 h;
-          temp_strengthening P2 Q2 h |]
-       ==> temp_strengthening (P1 \<^bold>\<or> P2) (Q1 \<^bold>\<or> Q2) h"
-apply (unfold temp_strengthening_def)
-apply auto
-done
-
-lemma strength_NOT:
-"[| temp_weakening P Q h |]
-       ==> temp_strengthening (\<^bold>\<not> P) (\<^bold>\<not> Q) h"
-apply (unfold temp_strengthening_def)
-apply (simp add: temp_weakening_def2)
-apply auto
-done
-
-lemma strength_IMPLIES:
-"[| temp_weakening P1 Q1 h;
-          temp_strengthening P2 Q2 h |]
-       ==> temp_strengthening (P1 \<^bold>\<longrightarrow> P2) (Q1 \<^bold>\<longrightarrow> Q2) h"
-apply (unfold temp_strengthening_def)
-apply (simp add: temp_weakening_def2)
-done
-
-
-lemma weak_AND:
-"[| temp_weakening P1 Q1 h;
-          temp_weakening P2 Q2 h |]
-       ==> temp_weakening (P1 \<^bold>\<and> P2) (Q1 \<^bold>\<and> Q2) h"
-apply (simp add: temp_weakening_def2)
-done
-
-lemma weak_OR:
-"[| temp_weakening P1 Q1 h;
-          temp_weakening P2 Q2 h |]
-       ==> temp_weakening (P1 \<^bold>\<or> P2) (Q1 \<^bold>\<or> Q2) h"
-apply (simp add: temp_weakening_def2)
-done
-
-lemma weak_NOT:
-"[| temp_strengthening P Q h |]
-       ==> temp_weakening (\<^bold>\<not> P) (\<^bold>\<not> Q) h"
-apply (unfold temp_strengthening_def)
-apply (simp add: temp_weakening_def2)
-apply auto
-done
-
-lemma weak_IMPLIES:
-"[| temp_strengthening P1 Q1 h;
-          temp_weakening P2 Q2 h |]
-       ==> temp_weakening (P1 \<^bold>\<longrightarrow> P2) (Q1 \<^bold>\<longrightarrow> Q2) h"
-apply (unfold temp_strengthening_def)
-apply (simp add: temp_weakening_def2)
-done
-
-
-subsubsection \<open>Box\<close>
-
-(* FIX: should be same as nil_is_Conc2 when all nils are turned to right side !! *)
-lemma UU_is_Conc: "(UU = x @@ y) = (((x::'a Seq)= UU) | (x=nil & y=UU))"
-apply (tactic \<open>Seq_case_simp_tac @{context} "x" 1\<close>)
-done
-
-lemma ex2seqConc [rule_format]:
-"Finite s1 -->
-  (! ex. (s~=nil & s~=UU & ex2seq ex = s1 @@ s) --> (? ex'. s = ex2seq ex'))"
-apply (rule impI)
-apply (tactic \<open>Seq_Finite_induct_tac @{context} 1\<close>)
-apply blast
-(* main case *)
-apply clarify
-apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
-(* UU case *)
-apply (simp add: nil_is_Conc)
-(* nil case *)
-apply (simp add: nil_is_Conc)
-(* cons case *)
-apply (tactic \<open>pair_tac @{context} "aa" 1\<close>)
-apply auto
-done
-
-(* important property of ex2seq: can be shiftet, as defined "pointwise" *)
-
-lemma ex2seq_tsuffix:
-"tsuffix s (ex2seq ex) ==> ? ex'. s = (ex2seq ex')"
-apply (unfold tsuffix_def suffix_def)
-apply auto
-apply (drule ex2seqConc)
-apply auto
-done
-
-
-(* FIX: NAch Sequence.ML bringen *)
-
-lemma Mapnil: "(Map f$s = nil) = (s=nil)"
-apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
-done
-
-lemma MapUU: "(Map f$s = UU) = (s=UU)"
-apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
-done
-
-
-(* important property of cex_absSeq: As it is a 1to1 correspondence,
-  properties carry over *)
-
-lemma cex_absSeq_tsuffix:
-"tsuffix s t ==> tsuffix (cex_absSeq h s) (cex_absSeq h t)"
-apply (unfold tsuffix_def suffix_def cex_absSeq_def)
-apply auto
-apply (simp add: Mapnil)
-apply (simp add: MapUU)
-apply (rule_tac x = "Map (% (s,a,t) . (h s,a, h t))$s1" in exI)
-apply (simp add: Map2Finite MapConc)
-done
-
-
-lemma strength_Box:
-"[| temp_strengthening P Q h |]
-       ==> temp_strengthening (\<box>P) (\<box>Q) h"
-apply (unfold temp_strengthening_def state_strengthening_def temp_sat_def satisfies_def Box_def)
-apply clarify
-apply (frule ex2seq_tsuffix)
-apply clarify
-apply (drule_tac h = "h" in cex_absSeq_tsuffix)
-apply (simp add: ex2seq_abs_cex)
-done
-
-
-subsubsection \<open>Init\<close>
-
-lemma strength_Init:
-"[| state_strengthening P Q h |]
-       ==> temp_strengthening (Init P) (Init Q) h"
-apply (unfold temp_strengthening_def state_strengthening_def
-  temp_sat_def satisfies_def Init_def unlift_def)
-apply auto
-apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-done
-
-
-subsubsection \<open>Next\<close>
-
-lemma TL_ex2seq_UU:
-"(TL$(ex2seq (cex_abs h ex))=UU) = (TL$(ex2seq ex)=UU)"
-apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-done
-
-lemma TL_ex2seq_nil:
-"(TL$(ex2seq (cex_abs h ex))=nil) = (TL$(ex2seq ex)=nil)"
-apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-done
-
-(* FIX: put to Sequence Lemmas *)
-lemma MapTL: "Map f$(TL$s) = TL$(Map f$s)"
-apply (tactic \<open>Seq_induct_tac @{context} "s" [] 1\<close>)
-done
-
-(* important property of cex_absSeq: As it is a 1to1 correspondence,
-  properties carry over *)
-
-lemma cex_absSeq_TL:
-"cex_absSeq h (TL$s) = (TL$(cex_absSeq h s))"
-apply (unfold cex_absSeq_def)
-apply (simp add: MapTL)
-done
-
-(* important property of ex2seq: can be shiftet, as defined "pointwise" *)
-
-lemma TLex2seq: "[| (snd ex)~=UU ; (snd ex)~=nil |] ==> (? ex'. TL$(ex2seq ex) = ex2seq ex')"
-apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-apply auto
-done
-
-
-lemma ex2seqnilTL: "(TL$(ex2seq ex)~=nil) = ((snd ex)~=nil & (snd ex)~=UU)"
-apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-done
-
-
-lemma strength_Next:
-"[| temp_strengthening P Q h |]
-       ==> temp_strengthening (Next P) (Next Q) h"
-apply (unfold temp_strengthening_def state_strengthening_def temp_sat_def satisfies_def Next_def)
-apply simp
-apply auto
-apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
-apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
-apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
-apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
-(* cons case *)
-apply (simp add: TL_ex2seq_nil TL_ex2seq_UU ex2seq_abs_cex cex_absSeq_TL [symmetric] ex2seqnilTL)
-apply (erule conjE)
-apply (drule TLex2seq)
-apply assumption
-apply auto
-done
-
-
-text \<open>Localizing Temporal Weakenings     - 2\<close>
-
-lemma weak_Init:
-"[| state_weakening P Q h |]
-       ==> temp_weakening (Init P) (Init Q) h"
-apply (simp add: temp_weakening_def2 state_weakening_def2
-  temp_sat_def satisfies_def Init_def unlift_def)
-apply auto
-apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-done
-
-
-text \<open>Localizing Temproal Strengthenings - 3\<close>
-
-lemma strength_Diamond:
-"[| temp_strengthening P Q h |]
-       ==> temp_strengthening (\<diamond>P) (\<diamond>Q) h"
-apply (unfold Diamond_def)
-apply (rule strength_NOT)
-apply (rule weak_Box)
-apply (erule weak_NOT)
-done
-
-lemma strength_Leadsto:
-"[| temp_weakening P1 P2 h;
-          temp_strengthening Q1 Q2 h |]
-       ==> temp_strengthening (P1 \<leadsto> Q1) (P2 \<leadsto> Q2) h"
-apply (unfold Leadsto_def)
-apply (rule strength_Box)
-apply (erule strength_IMPLIES)
-apply (erule strength_Diamond)
-done
-
-
-text \<open>Localizing Temporal Weakenings - 3\<close>
-
-lemma weak_Diamond:
-"[| temp_weakening P Q h |]
-       ==> temp_weakening (\<diamond>P) (\<diamond>Q) h"
-apply (unfold Diamond_def)
-apply (rule weak_NOT)
-apply (rule strength_Box)
-apply (erule strength_NOT)
-done
-
-lemma weak_Leadsto:
-"[| temp_strengthening P1 P2 h;
-          temp_weakening Q1 Q2 h |]
-       ==> temp_weakening (P1 \<leadsto> Q1) (P2 \<leadsto> Q2) h"
-apply (unfold Leadsto_def)
-apply (rule weak_Box)
-apply (erule weak_IMPLIES)
-apply (erule weak_Diamond)
-done
-
-lemma weak_WF:
-  " !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|]
-    ==> temp_weakening (WF A acts) (WF C acts) h"
-apply (unfold WF_def)
-apply (rule weak_IMPLIES)
-apply (rule strength_Diamond)
-apply (rule strength_Box)
-apply (rule strength_Init)
-apply (rule_tac [2] weak_Box)
-apply (rule_tac [2] weak_Diamond)
-apply (rule_tac [2] weak_Init)
-apply (auto simp add: state_weakening_def state_strengthening_def
-  xt2_def plift_def option_lift_def NOT_def)
-done
-
-lemma weak_SF:
-  " !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|]
-    ==> temp_weakening (SF A acts) (SF C acts) h"
-apply (unfold SF_def)
-apply (rule weak_IMPLIES)
-apply (rule strength_Box)
-apply (rule strength_Diamond)
-apply (rule strength_Init)
-apply (rule_tac [2] weak_Box)
-apply (rule_tac [2] weak_Diamond)
-apply (rule_tac [2] weak_Init)
-apply (auto simp add: state_weakening_def state_strengthening_def
-  xt2_def plift_def option_lift_def NOT_def)
-done
-
-
-lemmas weak_strength_lemmas =
-  weak_OR weak_AND weak_NOT weak_IMPLIES weak_Box weak_Next weak_Init
-  weak_Diamond weak_Leadsto strength_OR strength_AND strength_NOT
-  strength_IMPLIES strength_Box strength_Next strength_Init
-  strength_Diamond strength_Leadsto weak_WF weak_SF
-
-ML \<open>
-fun abstraction_tac ctxt =
-  SELECT_GOAL (auto_tac
-    (ctxt addSIs @{thms weak_strength_lemmas}
-      addsimps [@{thm state_strengthening_def}, @{thm state_weakening_def}]))
-\<close>
-
-method_setup abstraction = \<open>Scan.succeed (SIMPLE_METHOD' o abstraction_tac)\<close>
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/Asig.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,101 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/Asig.thy
-    Author:     Olaf Müller, Tobias Nipkow & Konrad Slind
-*)
-
-section \<open>Action signatures\<close>
-
-theory Asig
-imports Main
-begin
-
-type_synonym
-  'a signature = "('a set * 'a set * 'a set)"
-
-definition
-  inputs :: "'action signature => 'action set" where
-  asig_inputs_def: "inputs = fst"
-
-definition
-  outputs :: "'action signature => 'action set" where
-  asig_outputs_def: "outputs = (fst o snd)"
-
-definition
-  internals :: "'action signature => 'action set" where
-  asig_internals_def: "internals = (snd o snd)"
-
-definition
-  actions :: "'action signature => 'action set" where
-  "actions(asig) = (inputs(asig) Un outputs(asig) Un internals(asig))"
-
-definition
-  externals :: "'action signature => 'action set" where
-  "externals(asig) = (inputs(asig) Un outputs(asig))"
-
-definition
-  locals :: "'action signature => 'action set" where
-  "locals asig = ((internals asig) Un (outputs asig))"
-
-definition
-  is_asig :: "'action signature => bool" where
-  "is_asig(triple) =
-     ((inputs(triple) Int outputs(triple) = {}) &
-      (outputs(triple) Int internals(triple) = {}) &
-      (inputs(triple) Int internals(triple) = {}))"
-
-definition
-  mk_ext_asig :: "'action signature => 'action signature" where
-  "mk_ext_asig(triple) = (inputs(triple), outputs(triple), {})"
-
-
-lemmas asig_projections = asig_inputs_def asig_outputs_def asig_internals_def
-
-lemma asig_triple_proj:
- "(outputs    (a,b,c) = b)   &
-  (inputs     (a,b,c) = a) &
-  (internals  (a,b,c) = c)"
-  apply (simp add: asig_projections)
-  done
-
-lemma int_and_ext_is_act: "[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)"
-apply (simp add: externals_def actions_def)
-done
-
-lemma ext_is_act: "[|a:externals(S)|] ==> a:actions(S)"
-apply (simp add: externals_def actions_def)
-done
-
-lemma int_is_act: "[|a:internals S|] ==> a:actions S"
-apply (simp add: asig_internals_def actions_def)
-done
-
-lemma inp_is_act: "[|a:inputs S|] ==> a:actions S"
-apply (simp add: asig_inputs_def actions_def)
-done
-
-lemma out_is_act: "[|a:outputs S|] ==> a:actions S"
-apply (simp add: asig_outputs_def actions_def)
-done
-
-lemma ext_and_act: "(x: actions S & x : externals S) = (x: externals S)"
-apply (fast intro!: ext_is_act)
-done
-
-lemma not_ext_is_int: "[|is_asig S;x: actions S|] ==> (x~:externals S) = (x: internals S)"
-apply (simp add: actions_def is_asig_def externals_def)
-apply blast
-done
-
-lemma not_ext_is_int_or_not_act: "is_asig S ==> (x~:externals S) = (x: internals S | x~:actions S)"
-apply (simp add: actions_def is_asig_def externals_def)
-apply blast
-done
-
-lemma int_is_not_ext:
- "[| is_asig (S); x:internals S |] ==> x~:externals S"
-apply (unfold externals_def actions_def is_asig_def)
-apply simp
-apply blast
-done
-
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/Automata.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,615 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/Automata.thy
-    Author:     Olaf Müller, Konrad Slind, Tobias Nipkow
-*)
-
-section \<open>The I/O automata of Lynch and Tuttle in HOLCF\<close>
-
-theory Automata
-imports Asig
-begin
-
-default_sort type
-
-type_synonym ('a, 's) transition = "'s * 'a * 's"
-type_synonym ('a, 's) ioa =
-  "'a signature * 's set * ('a,'s)transition set * ('a set set) * ('a set set)"
-
-
-(* --------------------------------- IOA ---------------------------------*)
-
-(* IO automata *)
-
-definition asig_of :: "('a, 's)ioa \<Rightarrow> 'a signature"
-  where "asig_of = fst"
-
-definition starts_of :: "('a, 's) ioa \<Rightarrow> 's set"
-  where "starts_of = (fst \<circ> snd)"
-
-definition trans_of :: "('a, 's) ioa \<Rightarrow> ('a, 's) transition set"
-  where "trans_of = (fst \<circ> snd \<circ> snd)"
-
-abbreviation trans_of_syn  ("_ \<midarrow>_\<midarrow>_\<rightarrow> _" [81, 81, 81, 81] 100)
-  where "s \<midarrow>a\<midarrow>A\<rightarrow> t \<equiv> (s, a, t) \<in> trans_of A"
-
-definition wfair_of :: "('a, 's) ioa \<Rightarrow> 'a set set"
-  where "wfair_of = (fst \<circ> snd \<circ> snd \<circ> snd)"
-
-definition sfair_of :: "('a, 's) ioa \<Rightarrow> 'a set set"
-  where "sfair_of = (snd \<circ> snd \<circ> snd \<circ> snd)"
-
-definition is_asig_of :: "('a, 's) ioa \<Rightarrow> bool"
-  where "is_asig_of A = is_asig (asig_of A)"
-
-definition is_starts_of :: "('a, 's) ioa \<Rightarrow> bool"
-  where "is_starts_of A \<longleftrightarrow> starts_of A \<noteq> {}"
-
-definition is_trans_of :: "('a, 's) ioa \<Rightarrow> bool"
-  where "is_trans_of A \<longleftrightarrow>
-    (\<forall>triple. triple \<in> trans_of A \<longrightarrow> fst (snd triple) \<in> actions (asig_of A))"
-
-definition input_enabled :: "('a, 's) ioa \<Rightarrow> bool"
-  where "input_enabled A \<longleftrightarrow>
-    (\<forall>a. a \<in> inputs (asig_of A) \<longrightarrow> (\<forall>s1. \<exists>s2. (s1, a, s2) \<in> trans_of A))"
-
-definition IOA :: "('a, 's) ioa \<Rightarrow> bool"
-  where "IOA A \<longleftrightarrow>
-    is_asig_of A \<and>
-    is_starts_of A \<and>
-    is_trans_of A \<and>
-    input_enabled A"
-
-abbreviation "act A == actions (asig_of A)"
-abbreviation "ext A == externals (asig_of A)"
-abbreviation int where "int A == internals (asig_of A)"
-abbreviation "inp A == inputs (asig_of A)"
-abbreviation "out A == outputs (asig_of A)"
-abbreviation "local A == locals (asig_of A)"
-
-(* invariants *)
-inductive reachable :: "('a, 's) ioa \<Rightarrow> 's \<Rightarrow> bool"
-  for C :: "('a, 's) ioa"
-where
-  reachable_0:  "s \<in> starts_of C \<Longrightarrow> reachable C s"
-| reachable_n:  "\<lbrakk>reachable C s; (s, a, t) \<in> trans_of C\<rbrakk> \<Longrightarrow> reachable C t"
-
-definition invariant :: "[('a, 's) ioa, 's \<Rightarrow> bool] \<Rightarrow> bool"
-  where "invariant A P \<longleftrightarrow> (\<forall>s. reachable A s \<longrightarrow> P s)"
-
-
-(* ------------------------- parallel composition --------------------------*)
-
-(* binary composition of action signatures and automata *)
-
-definition compatible :: "[('a, 's) ioa, ('a, 't) ioa] \<Rightarrow> bool"
-where
-  "compatible A B \<longleftrightarrow>
-  (((out A \<inter> out B) = {}) \<and>
-   ((int A \<inter> act B) = {}) \<and>
-   ((int B \<inter> act A) = {}))"
-
-definition asig_comp :: "['a signature, 'a signature] \<Rightarrow> 'a signature"
-where
-  "asig_comp a1 a2 =
-     (((inputs(a1) \<union> inputs(a2)) - (outputs(a1) \<union> outputs(a2)),
-       (outputs(a1) \<union> outputs(a2)),
-       (internals(a1) \<union> internals(a2))))"
-
-definition par :: "[('a, 's) ioa, ('a, 't) ioa] \<Rightarrow> ('a, 's * 't) ioa"  (infixr "\<parallel>" 10)
-where
-  "(A \<parallel> B) =
-      (asig_comp (asig_of A) (asig_of B),
-       {pr. fst(pr) \<in> starts_of(A) \<and> snd(pr) \<in> starts_of(B)},
-       {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))
-            in (a \<in> act A | a:act B) \<and>
-               (if a \<in> act A then
-                  (fst(s), a, fst(t)) \<in> trans_of(A)
-                else fst(t) = fst(s))
-               &
-               (if a \<in> act B then
-                  (snd(s), a, snd(t)) \<in> trans_of(B)
-                else snd(t) = snd(s))},
-        wfair_of A \<union> wfair_of B,
-        sfair_of A \<union> sfair_of B)"
-
-
-(* ------------------------ hiding -------------------------------------------- *)
-
-(* hiding and restricting *)
-
-definition restrict_asig :: "['a signature, 'a set] \<Rightarrow> 'a signature"
-where
-  "restrict_asig asig actns =
-    (inputs(asig) Int actns,
-     outputs(asig) Int actns,
-     internals(asig) Un (externals(asig) - actns))"
-
-(* Notice that for wfair_of and sfair_of nothing has to be changed, as
-   changes from the outputs to the internals does not touch the locals as
-   a whole, which is of importance for fairness only *)
-definition restrict :: "[('a, 's) ioa, 'a set] \<Rightarrow> ('a, 's) ioa"
-where
-  "restrict A actns =
-    (restrict_asig (asig_of A) actns,
-     starts_of A,
-     trans_of A,
-     wfair_of A,
-     sfair_of A)"
-
-definition hide_asig :: "['a signature, 'a set] \<Rightarrow> 'a signature"
-where
-  "hide_asig asig actns =
-    (inputs(asig) - actns,
-     outputs(asig) - actns,
-     internals(asig) \<union> actns)"
-
-definition hide :: "[('a, 's) ioa, 'a set] \<Rightarrow> ('a, 's) ioa"
-where
-  "hide A actns =
-    (hide_asig (asig_of A) actns,
-     starts_of A,
-     trans_of A,
-     wfair_of A,
-     sfair_of A)"
-
-(* ------------------------- renaming ------------------------------------------- *)
-
-definition rename_set :: "'a set \<Rightarrow> ('c \<Rightarrow> 'a option) \<Rightarrow> 'c set"
-  where "rename_set A ren = {b. \<exists>x. Some x = ren b \<and> x \<in> A}"
-
-definition rename :: "('a, 'b) ioa \<Rightarrow> ('c \<Rightarrow> 'a option) \<Rightarrow> ('c, 'b) ioa"
-where
-  "rename ioa ren =
-    ((rename_set (inp ioa) ren,
-      rename_set (out ioa) ren,
-      rename_set (int ioa) ren),
-     starts_of ioa,
-     {tr. let s = fst(tr); a = fst(snd(tr));  t = snd(snd(tr))
-          in
-          \<exists>x. Some(x) = ren(a) \<and> (s,x,t):trans_of ioa},
-     {rename_set s ren | s. s \<in> wfair_of ioa},
-     {rename_set s ren | s. s \<in> sfair_of ioa})"
-
-
-(* ------------------------- fairness ----------------------------- *)
-
-(* enabledness of actions and action sets *)
-
-definition enabled :: "('a, 's) ioa \<Rightarrow> 'a \<Rightarrow> 's \<Rightarrow> bool"
-  where "enabled A a s \<longleftrightarrow> (\<exists>t. s \<midarrow>a\<midarrow>A\<rightarrow> t)"
-
-definition Enabled :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> 's \<Rightarrow> bool"
-  where "Enabled A W s \<longleftrightarrow> (\<exists>w \<in> W. enabled A w s)"
-
-
-(* action set keeps enabled until probably disabled by itself *)
-
-definition en_persistent :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> bool"
-where
-  "en_persistent A W \<longleftrightarrow>
-    (\<forall>s a t. Enabled A W s \<and> a \<notin> W \<and> s \<midarrow>a\<midarrow>A\<rightarrow> t \<longrightarrow> Enabled A W t)"
-
-
-(* post_conditions for actions and action sets *)
-
-definition was_enabled :: "('a, 's) ioa \<Rightarrow> 'a \<Rightarrow> 's \<Rightarrow> bool"
-  where "was_enabled A a t \<longleftrightarrow> (\<exists>s. s \<midarrow>a\<midarrow>A\<rightarrow> t)"
-
-definition set_was_enabled :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> 's \<Rightarrow> bool"
-  where "set_was_enabled A W t \<longleftrightarrow> (\<exists>w \<in> W. was_enabled A w t)"
-
-
-(* constraints for fair IOA *)
-
-definition fairIOA :: "('a, 's) ioa \<Rightarrow> bool"
-  where "fairIOA A \<longleftrightarrow> (\<forall>S \<in> wfair_of A. S \<subseteq> local A) \<and> (\<forall>S \<in> sfair_of A. S \<subseteq> local A)"
-
-definition input_resistant :: "('a, 's) ioa \<Rightarrow> bool"
-where
-  "input_resistant A \<longleftrightarrow>
-    (\<forall>W \<in> sfair_of A. \<forall>s a t.
-      reachable A s \<and> reachable A t \<and> a \<in> inp A \<and>
-      Enabled A W s \<and> s \<midarrow>a\<midarrow>A\<rightarrow> t \<longrightarrow> Enabled A W t)"
-
-
-declare split_paired_Ex [simp del]
-
-lemmas ioa_projections = asig_of_def starts_of_def trans_of_def wfair_of_def sfair_of_def
-
-
-subsection "asig_of, starts_of, trans_of"
-
-lemma ioa_triple_proj:
- "((asig_of (x,y,z,w,s)) = x)   &
-  ((starts_of (x,y,z,w,s)) = y) &
-  ((trans_of (x,y,z,w,s)) = z)  &
-  ((wfair_of (x,y,z,w,s)) = w) &
-  ((sfair_of (x,y,z,w,s)) = s)"
-  apply (simp add: ioa_projections)
-  done
-
-lemma trans_in_actions:
-  "[| is_trans_of A; (s1,a,s2):trans_of(A) |] ==> a:act A"
-  apply (unfold is_trans_of_def actions_def is_asig_def)
-    apply (erule allE, erule impE, assumption)
-    apply simp
-  done
-
-lemma starts_of_par: "starts_of(A \<parallel> B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}"
-  by (simp add: par_def ioa_projections)
-
-lemma trans_of_par:
-"trans_of(A \<parallel> B) = {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))
-             in (a:act A | a:act B) &
-                (if a:act A then
-                   (fst(s),a,fst(t)):trans_of(A)
-                 else fst(t) = fst(s))
-                &
-                (if a:act B then
-                   (snd(s),a,snd(t)):trans_of(B)
-                 else snd(t) = snd(s))}"
-  by (simp add: par_def ioa_projections)
-
-
-subsection "actions and par"
-
-lemma actions_asig_comp: "actions(asig_comp a b) = actions(a) Un actions(b)"
-  by (auto simp add: actions_def asig_comp_def asig_projections)
-
-lemma asig_of_par: "asig_of(A \<parallel> B) = asig_comp (asig_of A) (asig_of B)"
-  by (simp add: par_def ioa_projections)
-
-
-lemma externals_of_par: "ext (A1\<parallel>A2) = (ext A1) Un (ext A2)"
-  apply (simp add: externals_def asig_of_par asig_comp_def
-    asig_inputs_def asig_outputs_def Un_def set_diff_eq)
-  apply blast
-  done
-
-lemma actions_of_par: "act (A1\<parallel>A2) = (act A1) Un (act A2)"
-  apply (simp add: actions_def asig_of_par asig_comp_def
-    asig_inputs_def asig_outputs_def asig_internals_def Un_def set_diff_eq)
-  apply blast
-  done
-
-lemma inputs_of_par: "inp (A1\<parallel>A2) = ((inp A1) Un (inp A2)) - ((out A1) Un (out A2))"
-  by (simp add: actions_def asig_of_par asig_comp_def
-    asig_inputs_def asig_outputs_def Un_def set_diff_eq)
-
-lemma outputs_of_par: "out (A1\<parallel>A2) = (out A1) Un (out A2)"
-  by (simp add: actions_def asig_of_par asig_comp_def
-    asig_outputs_def Un_def set_diff_eq)
-
-lemma internals_of_par: "int (A1\<parallel>A2) = (int A1) Un (int A2)"
-  by (simp add: actions_def asig_of_par asig_comp_def
-    asig_inputs_def asig_outputs_def asig_internals_def Un_def set_diff_eq)
-
-
-subsection "actions and compatibility"
-
-lemma compat_commute: "compatible A B = compatible B A"
-  by (auto simp add: compatible_def Int_commute)
-
-lemma ext1_is_not_int2: "[| compatible A1 A2; a:ext A1|] ==> a~:int A2"
-  apply (unfold externals_def actions_def compatible_def)
-  apply simp
-  apply blast
-  done
-
-(* just commuting the previous one: better commute compatible *)
-lemma ext2_is_not_int1: "[| compatible A2 A1 ; a:ext A1|] ==> a~:int A2"
-  apply (unfold externals_def actions_def compatible_def)
-  apply simp
-  apply blast
-  done
-
-lemmas ext1_ext2_is_not_act2 = ext1_is_not_int2 [THEN int_and_ext_is_act]
-lemmas ext1_ext2_is_not_act1 = ext2_is_not_int1 [THEN int_and_ext_is_act]
-
-lemma intA_is_not_extB: "[| compatible A B; x:int A |] ==> x~:ext B"
-  apply (unfold externals_def actions_def compatible_def)
-  apply simp
-  apply blast
-  done
-
-lemma intA_is_not_actB: "[| compatible A B; a:int A |] ==> a ~: act B"
-  apply (unfold externals_def actions_def compatible_def is_asig_def asig_of_def)
-  apply simp
-  apply blast
-  done
-
-(* the only one that needs disjointness of outputs and of internals and _all_ acts *)
-lemma outAactB_is_inpB: "[| compatible A B; a:out A ;a:act B|] ==> a : inp B"
-  apply (unfold asig_outputs_def asig_internals_def actions_def asig_inputs_def
-      compatible_def is_asig_def asig_of_def)
-  apply simp
-  apply blast
-  done
-
-(* needed for propagation of input_enabledness from A,B to A\<parallel>B *)
-lemma inpAAactB_is_inpBoroutB:
-  "[| compatible A B; a:inp A ;a:act B|] ==> a : inp B | a: out B"
-  apply (unfold asig_outputs_def asig_internals_def actions_def asig_inputs_def
-      compatible_def is_asig_def asig_of_def)
-  apply simp
-  apply blast
-  done
-
-
-subsection "input_enabledness and par"
-
-(* ugly case distinctions. Heart of proof:
-     1. inpAAactB_is_inpBoroutB ie. internals are really hidden.
-     2. inputs_of_par: outputs are no longer inputs of par. This is important here *)
-lemma input_enabled_par:
-  "[| compatible A B; input_enabled A; input_enabled B|]
-        ==> input_enabled (A\<parallel>B)"
-  apply (unfold input_enabled_def)
-  apply (simp add: Let_def inputs_of_par trans_of_par)
-  apply (tactic "safe_tac (Context.raw_transfer @{theory} @{theory_context Fun})")
-  apply (simp add: inp_is_act)
-  prefer 2
-  apply (simp add: inp_is_act)
-  (* a: inp A *)
-  apply (case_tac "a:act B")
-  (* a:act B *)
-  apply (erule_tac x = "a" in allE)
-  apply simp
-  apply (drule inpAAactB_is_inpBoroutB)
-  apply assumption
-  apply assumption
-  apply (erule_tac x = "a" in allE)
-  apply simp
-  apply (erule_tac x = "aa" in allE)
-  apply (erule_tac x = "b" in allE)
-  apply (erule exE)
-  apply (erule exE)
-  apply (rule_tac x = " (s2,s2a) " in exI)
-  apply (simp add: inp_is_act)
-  (* a~: act B*)
-  apply (simp add: inp_is_act)
-  apply (erule_tac x = "a" in allE)
-  apply simp
-  apply (erule_tac x = "aa" in allE)
-  apply (erule exE)
-  apply (rule_tac x = " (s2,b) " in exI)
-  apply simp
-  
-  (* a:inp B *)
-  apply (case_tac "a:act A")
-  (* a:act A *)
-  apply (erule_tac x = "a" in allE)
-  apply (erule_tac x = "a" in allE)
-  apply (simp add: inp_is_act)
-  apply (frule_tac A1 = "A" in compat_commute [THEN iffD1])
-  apply (drule inpAAactB_is_inpBoroutB)
-  back
-  apply assumption
-  apply assumption
-  apply simp
-  apply (erule_tac x = "aa" in allE)
-  apply (erule_tac x = "b" in allE)
-  apply (erule exE)
-  apply (erule exE)
-  apply (rule_tac x = " (s2,s2a) " in exI)
-  apply (simp add: inp_is_act)
-  (* a~: act B*)
-  apply (simp add: inp_is_act)
-  apply (erule_tac x = "a" in allE)
-  apply (erule_tac x = "a" in allE)
-  apply simp
-  apply (erule_tac x = "b" in allE)
-  apply (erule exE)
-  apply (rule_tac x = " (aa,s2) " in exI)
-  apply simp
-  done
-
-
-subsection "invariants"
-
-lemma invariantI:
-  "[| !!s. s:starts_of(A) ==> P(s);
-      !!s t a. [|reachable A s; P(s)|] ==> (s,a,t): trans_of(A) --> P(t) |]
-   ==> invariant A P"
-  apply (unfold invariant_def)
-  apply (rule allI)
-  apply (rule impI)
-  apply (rule_tac x = "s" in reachable.induct)
-  apply assumption
-  apply blast
-  apply blast
-  done
-
-lemma invariantI1:
- "[| !!s. s : starts_of(A) ==> P(s);
-     !!s t a. reachable A s ==> P(s) --> (s,a,t):trans_of(A) --> P(t)
-  |] ==> invariant A P"
-  apply (blast intro: invariantI)
-  done
-
-lemma invariantE: "[| invariant A P; reachable A s |] ==> P(s)"
-  apply (unfold invariant_def)
-  apply blast
-  done
-
-
-subsection "restrict"
-
-
-lemmas reachable_0 = reachable.reachable_0
-  and reachable_n = reachable.reachable_n
-
-lemma cancel_restrict_a: "starts_of(restrict ioa acts) = starts_of(ioa) &
-          trans_of(restrict ioa acts) = trans_of(ioa)"
-  by (simp add: restrict_def ioa_projections)
-
-lemma cancel_restrict_b: "reachable (restrict ioa acts) s = reachable ioa s"
-  apply (rule iffI)
-  apply (erule reachable.induct)
-  apply (simp add: cancel_restrict_a reachable_0)
-  apply (erule reachable_n)
-  apply (simp add: cancel_restrict_a)
-  (* <--  *)
-  apply (erule reachable.induct)
-  apply (rule reachable_0)
-  apply (simp add: cancel_restrict_a)
-  apply (erule reachable_n)
-  apply (simp add: cancel_restrict_a)
-  done
-
-lemma acts_restrict: "act (restrict A acts) = act A"
-  apply (simp (no_asm) add: actions_def asig_internals_def
-    asig_outputs_def asig_inputs_def externals_def asig_of_def restrict_def restrict_asig_def)
-  apply auto
-  done
-
-lemma cancel_restrict: "starts_of(restrict ioa acts) = starts_of(ioa) &
-          trans_of(restrict ioa acts) = trans_of(ioa) &
-          reachable (restrict ioa acts) s = reachable ioa s &
-          act (restrict A acts) = act A"
-  by (simp add: cancel_restrict_a cancel_restrict_b acts_restrict)
-
-
-subsection "rename"
-
-lemma trans_rename: "s \<midarrow>a\<midarrow>(rename C f)\<rightarrow> t ==> (? x. Some(x) = f(a) & s \<midarrow>x\<midarrow>C\<rightarrow> t)"
-  by (simp add: Let_def rename_def trans_of_def)
-
-
-lemma reachable_rename: "[| reachable (rename C g) s |] ==> reachable C s"
-  apply (erule reachable.induct)
-  apply (rule reachable_0)
-  apply (simp add: rename_def ioa_projections)
-  apply (drule trans_rename)
-  apply (erule exE)
-  apply (erule conjE)
-  apply (erule reachable_n)
-  apply assumption
-  done
-
-
-subsection "trans_of(A\<parallel>B)"
-
-lemma trans_A_proj: "[|(s,a,t):trans_of (A\<parallel>B); a:act A|]
-              ==> (fst s,a,fst t):trans_of A"
-  by (simp add: Let_def par_def trans_of_def)
-
-lemma trans_B_proj: "[|(s,a,t):trans_of (A\<parallel>B); a:act B|]
-              ==> (snd s,a,snd t):trans_of B"
-  by (simp add: Let_def par_def trans_of_def)
-
-lemma trans_A_proj2: "[|(s,a,t):trans_of (A\<parallel>B); a~:act A|]
-              ==> fst s = fst t"
-  by (simp add: Let_def par_def trans_of_def)
-
-lemma trans_B_proj2: "[|(s,a,t):trans_of (A\<parallel>B); a~:act B|]
-              ==> snd s = snd t"
-  by (simp add: Let_def par_def trans_of_def)
-
-lemma trans_AB_proj: "(s,a,t):trans_of (A\<parallel>B)
-               ==> a :act A | a :act B"
-  by (simp add: Let_def par_def trans_of_def)
-
-lemma trans_AB: "[|a:act A;a:act B;
-       (fst s,a,fst t):trans_of A;(snd s,a,snd t):trans_of B|]
-   ==> (s,a,t):trans_of (A\<parallel>B)"
-  by (simp add: Let_def par_def trans_of_def)
-
-lemma trans_A_notB: "[|a:act A;a~:act B;
-       (fst s,a,fst t):trans_of A;snd s=snd t|]
-   ==> (s,a,t):trans_of (A\<parallel>B)"
-  by (simp add: Let_def par_def trans_of_def)
-
-lemma trans_notA_B: "[|a~:act A;a:act B;
-       (snd s,a,snd t):trans_of B;fst s=fst t|]
-   ==> (s,a,t):trans_of (A\<parallel>B)"
-  by (simp add: Let_def par_def trans_of_def)
-
-lemmas trans_of_defs1 = trans_AB trans_A_notB trans_notA_B
-  and trans_of_defs2 = trans_A_proj trans_B_proj trans_A_proj2 trans_B_proj2 trans_AB_proj
-
-
-lemma trans_of_par4:
-"((s,a,t) : trans_of(A \<parallel> B \<parallel> C \<parallel> D)) =
-  ((a:actions(asig_of(A)) | a:actions(asig_of(B)) | a:actions(asig_of(C)) |
-    a:actions(asig_of(D))) &
-   (if a:actions(asig_of(A)) then (fst(s),a,fst(t)):trans_of(A)
-    else fst t=fst s) &
-   (if a:actions(asig_of(B)) then (fst(snd(s)),a,fst(snd(t))):trans_of(B)
-    else fst(snd(t))=fst(snd(s))) &
-   (if a:actions(asig_of(C)) then
-      (fst(snd(snd(s))),a,fst(snd(snd(t)))):trans_of(C)
-    else fst(snd(snd(t)))=fst(snd(snd(s)))) &
-   (if a:actions(asig_of(D)) then
-      (snd(snd(snd(s))),a,snd(snd(snd(t)))):trans_of(D)
-    else snd(snd(snd(t)))=snd(snd(snd(s)))))"
-  by (simp add: par_def actions_asig_comp prod_eq_iff Let_def ioa_projections)
-
-
-subsection "proof obligation generator for IOA requirements"
-
-(* without assumptions on A and B because is_trans_of is also incorporated in \<parallel>def *)
-lemma is_trans_of_par: "is_trans_of (A\<parallel>B)"
-  by (simp add: is_trans_of_def Let_def actions_of_par trans_of_par)
-
-lemma is_trans_of_restrict: "is_trans_of A ==> is_trans_of (restrict A acts)"
-  by (simp add: is_trans_of_def cancel_restrict acts_restrict)
-
-lemma is_trans_of_rename: "is_trans_of A ==> is_trans_of (rename A f)"
-  apply (unfold is_trans_of_def restrict_def restrict_asig_def)
-  apply (simp add: Let_def actions_def trans_of_def asig_internals_def
-    asig_outputs_def asig_inputs_def externals_def asig_of_def rename_def rename_set_def)
-  apply blast
-  done
-
-lemma is_asig_of_par: "[| is_asig_of A; is_asig_of B; compatible A B|]
-          ==> is_asig_of (A\<parallel>B)"
-  apply (simp add: is_asig_of_def asig_of_par asig_comp_def compatible_def
-    asig_internals_def asig_outputs_def asig_inputs_def actions_def is_asig_def)
-  apply (simp add: asig_of_def)
-  apply auto
-  done
-
-lemma is_asig_of_restrict: "is_asig_of A ==> is_asig_of (restrict A f)"
-  apply (unfold is_asig_of_def is_asig_def asig_of_def restrict_def restrict_asig_def
-             asig_internals_def asig_outputs_def asig_inputs_def externals_def o_def)
-  apply simp
-  apply auto
-  done
-
-lemma is_asig_of_rename: "is_asig_of A ==> is_asig_of (rename A f)"
-  apply (simp add: is_asig_of_def rename_def rename_set_def asig_internals_def
-    asig_outputs_def asig_inputs_def actions_def is_asig_def asig_of_def)
-  apply auto
-  apply (drule_tac [!] s = "Some _" in sym)
-  apply auto
-  done
-
-lemmas [simp] = is_asig_of_par is_asig_of_restrict
-  is_asig_of_rename is_trans_of_par is_trans_of_restrict is_trans_of_rename
-
-
-lemma compatible_par: "[|compatible A B; compatible A C |]==> compatible A (B\<parallel>C)"
-  apply (unfold compatible_def)
-  apply (simp add: internals_of_par outputs_of_par actions_of_par)
-  apply auto
-  done
-
-(*  better derive by previous one and compat_commute *)
-lemma compatible_par2: "[|compatible A C; compatible B C |]==> compatible (A\<parallel>B) C"
-  apply (unfold compatible_def)
-  apply (simp add: internals_of_par outputs_of_par actions_of_par)
-  apply auto
-  done
-
-lemma compatible_restrict:
-  "[| compatible A B; (ext B - S) Int ext A = {}|]
-        ==> compatible A (restrict B S)"
-  apply (unfold compatible_def)
-  apply (simp add: ioa_triple_proj asig_triple_proj externals_def
-    restrict_def restrict_asig_def actions_def)
-  apply auto
-  done
-
-declare split_paired_Ex [simp]
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/CompoExecs.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,303 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/CompoExecs.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Compositionality on Execution level\<close>
-
-theory CompoExecs
-imports Traces
-begin
-
-definition
-  ProjA2 :: "('a,'s * 't)pairs -> ('a,'s)pairs" where
-  "ProjA2 = Map (%x.(fst x,fst(snd x)))"
-
-definition
-  ProjA :: "('a,'s * 't)execution => ('a,'s)execution" where
-  "ProjA ex = (fst (fst ex), ProjA2$(snd ex))"
-
-definition
-  ProjB2 :: "('a,'s * 't)pairs -> ('a,'t)pairs" where
-  "ProjB2 = Map (%x.(fst x,snd(snd x)))"
-
-definition
-  ProjB :: "('a,'s * 't)execution => ('a,'t)execution" where
-  "ProjB ex = (snd (fst ex), ProjB2$(snd ex))"
-
-definition
-  Filter_ex2 :: "'a signature => ('a,'s)pairs -> ('a,'s)pairs" where
-  "Filter_ex2 sig = Filter (%x. fst x:actions sig)"
-
-definition
-  Filter_ex :: "'a signature => ('a,'s)execution => ('a,'s)execution" where
-  "Filter_ex sig ex = (fst ex,Filter_ex2 sig$(snd ex))"
-
-definition
-  stutter2 :: "'a signature => ('a,'s)pairs -> ('s => tr)" where
-  "stutter2 sig = (fix$(LAM h ex. (%s. case ex of
-      nil => TT
-    | x##xs => (flift1
-            (%p.(If Def ((fst p)~:actions sig)
-                 then Def (s=(snd p))
-                 else TT)
-                andalso (h$xs) (snd p))
-             $x)
-   )))"
-
-definition
-  stutter :: "'a signature => ('a,'s)execution => bool" where
-  "stutter sig ex = ((stutter2 sig$(snd ex)) (fst ex) ~= FF)"
-
-definition
-  par_execs :: "[('a,'s)execution_module,('a,'t)execution_module] => ('a,'s*'t)execution_module" where
-  "par_execs ExecsA ExecsB =
-      (let exA = fst ExecsA; sigA = snd ExecsA;
-           exB = fst ExecsB; sigB = snd ExecsB
-       in
-       (    {ex. Filter_ex sigA (ProjA ex) : exA}
-        Int {ex. Filter_ex sigB (ProjB ex) : exB}
-        Int {ex. stutter sigA (ProjA ex)}
-        Int {ex. stutter sigB (ProjB ex)}
-        Int {ex. Forall (%x. fst x:(actions sigA Un actions sigB)) (snd ex)},
-        asig_comp sigA sigB))"
-
-
-lemmas [simp del] = split_paired_All
-
-
-section "recursive equations of operators"
-
-
-(* ---------------------------------------------------------------- *)
-(*                               ProjA2                             *)
-(* ---------------------------------------------------------------- *)
-
-
-lemma ProjA2_UU: "ProjA2$UU = UU"
-apply (simp add: ProjA2_def)
-done
-
-lemma ProjA2_nil: "ProjA2$nil = nil"
-apply (simp add: ProjA2_def)
-done
-
-lemma ProjA2_cons: "ProjA2$((a,t)\<leadsto>xs) = (a,fst t) \<leadsto> ProjA2$xs"
-apply (simp add: ProjA2_def)
-done
-
-
-(* ---------------------------------------------------------------- *)
-(*                               ProjB2                             *)
-(* ---------------------------------------------------------------- *)
-
-
-lemma ProjB2_UU: "ProjB2$UU = UU"
-apply (simp add: ProjB2_def)
-done
-
-lemma ProjB2_nil: "ProjB2$nil = nil"
-apply (simp add: ProjB2_def)
-done
-
-lemma ProjB2_cons: "ProjB2$((a,t)\<leadsto>xs) = (a,snd t) \<leadsto> ProjB2$xs"
-apply (simp add: ProjB2_def)
-done
-
-
-
-(* ---------------------------------------------------------------- *)
-(*                             Filter_ex2                           *)
-(* ---------------------------------------------------------------- *)
-
-
-lemma Filter_ex2_UU: "Filter_ex2 sig$UU=UU"
-apply (simp add: Filter_ex2_def)
-done
-
-lemma Filter_ex2_nil: "Filter_ex2 sig$nil=nil"
-apply (simp add: Filter_ex2_def)
-done
-
-lemma Filter_ex2_cons: "Filter_ex2 sig$(at \<leadsto> xs) =
-             (if (fst at:actions sig)
-                  then at \<leadsto> (Filter_ex2 sig$xs)
-                  else        Filter_ex2 sig$xs)"
-
-apply (simp add: Filter_ex2_def)
-done
-
-
-(* ---------------------------------------------------------------- *)
-(*                             stutter2                             *)
-(* ---------------------------------------------------------------- *)
-
-
-lemma stutter2_unfold: "stutter2 sig = (LAM ex. (%s. case ex of
-       nil => TT
-     | x##xs => (flift1
-             (%p.(If Def ((fst p)~:actions sig)
-                  then Def (s=(snd p))
-                  else TT)
-                 andalso (stutter2 sig$xs) (snd p))
-              $x)
-            ))"
-apply (rule trans)
-apply (rule fix_eq2)
-apply (simp only: stutter2_def)
-apply (rule beta_cfun)
-apply (simp add: flift1_def)
-done
-
-lemma stutter2_UU: "(stutter2 sig$UU) s=UU"
-apply (subst stutter2_unfold)
-apply simp
-done
-
-lemma stutter2_nil: "(stutter2 sig$nil) s = TT"
-apply (subst stutter2_unfold)
-apply simp
-done
-
-lemma stutter2_cons: "(stutter2 sig$(at\<leadsto>xs)) s =
-               ((if (fst at)~:actions sig then Def (s=snd at) else TT)
-                 andalso (stutter2 sig$xs) (snd at))"
-apply (rule trans)
-apply (subst stutter2_unfold)
-apply (simp add: Consq_def flift1_def If_and_if)
-apply simp
-done
-
-
-declare stutter2_UU [simp] stutter2_nil [simp] stutter2_cons [simp]
-
-
-(* ---------------------------------------------------------------- *)
-(*                             stutter                              *)
-(* ---------------------------------------------------------------- *)
-
-lemma stutter_UU: "stutter sig (s, UU)"
-apply (simp add: stutter_def)
-done
-
-lemma stutter_nil: "stutter sig (s, nil)"
-apply (simp add: stutter_def)
-done
-
-lemma stutter_cons: "stutter sig (s, (a,t)\<leadsto>ex) =
-      ((a~:actions sig --> (s=t)) & stutter sig (t,ex))"
-apply (simp add: stutter_def)
-done
-
-(* ----------------------------------------------------------------------------------- *)
-
-declare stutter2_UU [simp del] stutter2_nil [simp del] stutter2_cons [simp del]
-
-lemmas compoex_simps = ProjA2_UU ProjA2_nil ProjA2_cons
-  ProjB2_UU ProjB2_nil ProjB2_cons
-  Filter_ex2_UU Filter_ex2_nil Filter_ex2_cons
-  stutter_UU stutter_nil stutter_cons
-
-declare compoex_simps [simp]
-
-
-
-(* ------------------------------------------------------------------ *)
-(*                      The following lemmata aim for                 *)
-(*             COMPOSITIONALITY   on    EXECUTION     Level           *)
-(* ------------------------------------------------------------------ *)
-
-
-(* --------------------------------------------------------------------- *)
-(*  Lemma_1_1a : is_ex_fr propagates from A\<parallel>B to Projections A and B    *)
-(* --------------------------------------------------------------------- *)
-
-lemma lemma_1_1a: "!s. is_exec_frag (A\<parallel>B) (s,xs)
-       -->  is_exec_frag A (fst s, Filter_ex2 (asig_of A)$(ProjA2$xs)) &
-            is_exec_frag B (snd s, Filter_ex2 (asig_of B)$(ProjB2$xs))"
-
-apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1\<close>)
-(* main case *)
-apply (auto simp add: trans_of_defs2)
-done
-
-
-(* --------------------------------------------------------------------- *)
-(*  Lemma_1_1b : is_ex_fr (A\<parallel>B) implies stuttering on Projections       *)
-(* --------------------------------------------------------------------- *)
-
-lemma lemma_1_1b: "!s. is_exec_frag (A\<parallel>B) (s,xs)
-       --> stutter (asig_of A) (fst s,ProjA2$xs)  &
-           stutter (asig_of B) (snd s,ProjB2$xs)"
-
-apply (tactic \<open>pair_induct_tac @{context} "xs"
-  [@{thm stutter_def}, @{thm is_exec_frag_def}] 1\<close>)
-(* main case *)
-apply (auto simp add: trans_of_defs2)
-done
-
-
-(* --------------------------------------------------------------------- *)
-(*  Lemma_1_1c : Executions of A\<parallel>B have only  A- or B-actions           *)
-(* --------------------------------------------------------------------- *)
-
-lemma lemma_1_1c: "!s. (is_exec_frag (A\<parallel>B) (s,xs)
-   --> Forall (%x. fst x:act (A\<parallel>B)) xs)"
-
-apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm Forall_def}, @{thm sforall_def},
-  @{thm is_exec_frag_def}] 1\<close>)
-(* main case *)
-apply auto
-apply (simp add: trans_of_defs2 actions_asig_comp asig_of_par)
-done
-
-
-(* ----------------------------------------------------------------------- *)
-(*  Lemma_1_2 : ex A, exB, stuttering and forall a:A|B implies ex (A\<parallel>B)   *)
-(* ----------------------------------------------------------------------- *)
-
-lemma lemma_1_2:
-"!s. is_exec_frag A (fst s,Filter_ex2 (asig_of A)$(ProjA2$xs)) &
-     is_exec_frag B (snd s,Filter_ex2 (asig_of B)$(ProjB2$xs)) &
-     stutter (asig_of A) (fst s,(ProjA2$xs)) &
-     stutter (asig_of B) (snd s,(ProjB2$xs)) &
-     Forall (%x. fst x:act (A\<parallel>B)) xs
-     --> is_exec_frag (A\<parallel>B) (s,xs)"
-
-apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm Forall_def}, @{thm sforall_def},
-  @{thm is_exec_frag_def}, @{thm stutter_def}] 1\<close>)
-apply (auto simp add: trans_of_defs1 actions_asig_comp asig_of_par)
-done
-
-
-subsection \<open>COMPOSITIONALITY on EXECUTION Level -- Main Theorem\<close>
-
-lemma compositionality_ex:
-"(ex:executions(A\<parallel>B)) =
- (Filter_ex (asig_of A) (ProjA ex) : executions A &
-  Filter_ex (asig_of B) (ProjB ex) : executions B &
-  stutter (asig_of A) (ProjA ex) & stutter (asig_of B) (ProjB ex) &
-  Forall (%x. fst x:act (A\<parallel>B)) (snd ex))"
-
-apply (simp (no_asm) add: executions_def ProjB_def Filter_ex_def ProjA_def starts_of_par)
-apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-apply (rule iffI)
-(* ==>  *)
-apply (erule conjE)+
-apply (simp add: lemma_1_1a lemma_1_1b lemma_1_1c)
-(* <==  *)
-apply (erule conjE)+
-apply (simp add: lemma_1_2)
-done
-
-
-subsection \<open>COMPOSITIONALITY on EXECUTION Level -- for Modules\<close>
-
-lemma compositionality_ex_modules:
-  "Execs (A\<parallel>B) = par_execs (Execs A) (Execs B)"
-apply (unfold Execs_def par_execs_def)
-apply (simp add: asig_of_par)
-apply (rule set_eqI)
-apply (simp add: compositionality_ex actions_of_par)
-done
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/CompoScheds.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,541 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/CompoScheds.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Compositionality on Schedule level\<close>
-
-theory CompoScheds
-imports CompoExecs
-begin
-
-definition
-  mkex2 :: "('a,'s)ioa => ('a,'t)ioa => 'a Seq ->
-              ('a,'s)pairs -> ('a,'t)pairs ->
-              ('s => 't => ('a,'s*'t)pairs)" where
-  "mkex2 A B = (fix$(LAM h sch exA exB. (%s t. case sch of
-       nil => nil
-    | x##xs =>
-      (case x of
-        UU => UU
-      | Def y =>
-         (if y:act A then
-             (if y:act B then
-                (case HD$exA of
-                   UU => UU
-                 | Def a => (case HD$exB of
-                              UU => UU
-                            | Def b =>
-                   (y,(snd a,snd b))\<leadsto>
-                     (h$xs$(TL$exA)$(TL$exB)) (snd a) (snd b)))
-              else
-                (case HD$exA of
-                   UU => UU
-                 | Def a =>
-                   (y,(snd a,t))\<leadsto>(h$xs$(TL$exA)$exB) (snd a) t)
-              )
-          else
-             (if y:act B then
-                (case HD$exB of
-                   UU => UU
-                 | Def b =>
-                   (y,(s,snd b))\<leadsto>(h$xs$exA$(TL$exB)) s (snd b))
-             else
-               UU
-             )
-         )
-       ))))"
-
-definition
-  mkex :: "('a,'s)ioa => ('a,'t)ioa => 'a Seq =>
-              ('a,'s)execution => ('a,'t)execution =>('a,'s*'t)execution" where
-  "mkex A B sch exA exB =
-       ((fst exA,fst exB),
-        (mkex2 A B$sch$(snd exA)$(snd exB)) (fst exA) (fst exB))"
-
-definition
-  par_scheds ::"['a schedule_module,'a schedule_module] => 'a schedule_module" where
-  "par_scheds SchedsA SchedsB =
-      (let schA = fst SchedsA; sigA = snd SchedsA;
-           schB = fst SchedsB; sigB = snd SchedsB
-       in
-       (    {sch. Filter (%a. a:actions sigA)$sch : schA}
-        Int {sch. Filter (%a. a:actions sigB)$sch : schB}
-        Int {sch. Forall (%x. x:(actions sigA Un actions sigB)) sch},
-        asig_comp sigA sigB))"
-
-
-subsection "mkex rewrite rules"
-
-
-lemma mkex2_unfold:
-"mkex2 A B = (LAM sch exA exB. (%s t. case sch of
-      nil => nil
-   | x##xs =>
-     (case x of
-       UU => UU
-     | Def y =>
-        (if y:act A then
-            (if y:act B then
-               (case HD$exA of
-                  UU => UU
-                | Def a => (case HD$exB of
-                             UU => UU
-                           | Def b =>
-                  (y,(snd a,snd b))\<leadsto>
-                    (mkex2 A B$xs$(TL$exA)$(TL$exB)) (snd a) (snd b)))
-             else
-               (case HD$exA of
-                  UU => UU
-                | Def a =>
-                  (y,(snd a,t))\<leadsto>(mkex2 A B$xs$(TL$exA)$exB) (snd a) t)
-             )
-         else
-            (if y:act B then
-               (case HD$exB of
-                  UU => UU
-                | Def b =>
-                  (y,(s,snd b))\<leadsto>(mkex2 A B$xs$exA$(TL$exB)) s (snd b))
-            else
-              UU
-            )
-        )
-      )))"
-apply (rule trans)
-apply (rule fix_eq2)
-apply (simp only: mkex2_def)
-apply (rule beta_cfun)
-apply simp
-done
-
-lemma mkex2_UU: "(mkex2 A B$UU$exA$exB) s t = UU"
-apply (subst mkex2_unfold)
-apply simp
-done
-
-lemma mkex2_nil: "(mkex2 A B$nil$exA$exB) s t= nil"
-apply (subst mkex2_unfold)
-apply simp
-done
-
-lemma mkex2_cons_1: "[| x:act A; x~:act B; HD$exA=Def a|]
-    ==> (mkex2 A B$(x\<leadsto>sch)$exA$exB) s t =
-        (x,snd a,t) \<leadsto> (mkex2 A B$sch$(TL$exA)$exB) (snd a) t"
-apply (rule trans)
-apply (subst mkex2_unfold)
-apply (simp add: Consq_def If_and_if)
-apply (simp add: Consq_def)
-done
-
-lemma mkex2_cons_2: "[| x~:act A; x:act B; HD$exB=Def b|]
-    ==> (mkex2 A B$(x\<leadsto>sch)$exA$exB) s t =
-        (x,s,snd b) \<leadsto> (mkex2 A B$sch$exA$(TL$exB)) s (snd b)"
-apply (rule trans)
-apply (subst mkex2_unfold)
-apply (simp add: Consq_def If_and_if)
-apply (simp add: Consq_def)
-done
-
-lemma mkex2_cons_3: "[| x:act A; x:act B; HD$exA=Def a;HD$exB=Def b|]
-    ==> (mkex2 A B$(x\<leadsto>sch)$exA$exB) s t =
-         (x,snd a,snd b) \<leadsto>
-            (mkex2 A B$sch$(TL$exA)$(TL$exB)) (snd a) (snd b)"
-apply (rule trans)
-apply (subst mkex2_unfold)
-apply (simp add: Consq_def If_and_if)
-apply (simp add: Consq_def)
-done
-
-declare mkex2_UU [simp] mkex2_nil [simp] mkex2_cons_1 [simp]
-  mkex2_cons_2 [simp] mkex2_cons_3 [simp]
-
-
-subsection \<open>mkex\<close>
-
-lemma mkex_UU: "mkex A B UU  (s,exA) (t,exB) = ((s,t),UU)"
-apply (simp add: mkex_def)
-done
-
-lemma mkex_nil: "mkex A B nil (s,exA) (t,exB) = ((s,t),nil)"
-apply (simp add: mkex_def)
-done
-
-lemma mkex_cons_1: "[| x:act A; x~:act B |]
-    ==> mkex A B (x\<leadsto>sch) (s,a\<leadsto>exA) (t,exB)  =
-        ((s,t), (x,snd a,t) \<leadsto> snd (mkex A B sch (snd a,exA) (t,exB)))"
-apply (simp (no_asm) add: mkex_def)
-apply (cut_tac exA = "a\<leadsto>exA" in mkex2_cons_1)
-apply auto
-done
-
-lemma mkex_cons_2: "[| x~:act A; x:act B |]
-    ==> mkex A B (x\<leadsto>sch) (s,exA) (t,b\<leadsto>exB) =
-        ((s,t), (x,s,snd b) \<leadsto> snd (mkex A B sch (s,exA) (snd b,exB)))"
-apply (simp (no_asm) add: mkex_def)
-apply (cut_tac exB = "b\<leadsto>exB" in mkex2_cons_2)
-apply auto
-done
-
-lemma mkex_cons_3: "[| x:act A; x:act B |]
-    ==>  mkex A B (x\<leadsto>sch) (s,a\<leadsto>exA) (t,b\<leadsto>exB) =
-         ((s,t), (x,snd a,snd b) \<leadsto> snd (mkex A B sch (snd a,exA) (snd b,exB)))"
-apply (simp (no_asm) add: mkex_def)
-apply (cut_tac exB = "b\<leadsto>exB" and exA = "a\<leadsto>exA" in mkex2_cons_3)
-apply auto
-done
-
-declare mkex2_UU [simp del] mkex2_nil [simp del]
-  mkex2_cons_1 [simp del] mkex2_cons_2 [simp del] mkex2_cons_3 [simp del]
-
-lemmas composch_simps = mkex_UU mkex_nil mkex_cons_1 mkex_cons_2 mkex_cons_3
-
-declare composch_simps [simp]
-
-
-subsection \<open>COMPOSITIONALITY on SCHEDULE Level\<close>
-
-subsubsection "Lemmas for ==>"
-
-(* --------------------------------------------------------------------- *)
-(*    Lemma_2_1 :  tfilter(ex) and filter_act are commutative            *)
-(* --------------------------------------------------------------------- *)
-
-lemma lemma_2_1a:
-   "filter_act$(Filter_ex2 (asig_of A)$xs)=
-    Filter (%a. a:act A)$(filter_act$xs)"
-
-apply (unfold filter_act_def Filter_ex2_def)
-apply (simp (no_asm) add: MapFilter o_def)
-done
-
-
-(* --------------------------------------------------------------------- *)
-(*    Lemma_2_2 : State-projections do not affect filter_act             *)
-(* --------------------------------------------------------------------- *)
-
-lemma lemma_2_1b:
-   "filter_act$(ProjA2$xs) =filter_act$xs &
-    filter_act$(ProjB2$xs) =filter_act$xs"
-apply (tactic \<open>pair_induct_tac @{context} "xs" [] 1\<close>)
-done
-
-
-(* --------------------------------------------------------------------- *)
-(*             Schedules of A\<parallel>B have only  A- or B-actions              *)
-(* --------------------------------------------------------------------- *)
-
-(* very similar to lemma_1_1c, but it is not checking if every action element of
-   an ex is in A or B, but after projecting it onto the action schedule. Of course, this
-   is the same proposition, but we cannot change this one, when then rather lemma_1_1c  *)
-
-lemma sch_actions_in_AorB: "!s. is_exec_frag (A\<parallel>B) (s,xs)
-   --> Forall (%x. x:act (A\<parallel>B)) (filter_act$xs)"
-
-apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}, @{thm Forall_def},
-  @{thm sforall_def}] 1\<close>)
-(* main case *)
-apply auto
-apply (simp add: trans_of_defs2 actions_asig_comp asig_of_par)
-done
-
-
-subsubsection "Lemmas for <=="
-
-(*---------------------------------------------------------------------------
-    Filtering actions out of mkex(sch,exA,exB) yields the oracle sch
-                             structural induction
-  --------------------------------------------------------------------------- *)
-
-lemma Mapfst_mkex_is_sch: "! exA exB s t.
-  Forall (%x. x:act (A\<parallel>B)) sch  &
-  Filter (%a. a:act A)$sch << filter_act$exA &
-  Filter (%a. a:act B)$sch << filter_act$exB
-  --> filter_act$(snd (mkex A B sch (s,exA) (t,exB))) = sch"
-
-apply (tactic \<open>Seq_induct_tac @{context} "sch" [@{thm Filter_def}, @{thm Forall_def},
-  @{thm sforall_def}, @{thm mkex_def}] 1\<close>)
-
-(* main case *)
-(* splitting into 4 cases according to a:A, a:B *)
-apply auto
-
-(* Case y:A, y:B *)
-apply (tactic \<open>Seq_case_simp_tac @{context} "exA" 1\<close>)
-(* Case exA=UU, Case exA=nil*)
-(* These UU and nil cases are the only places where the assumption filter A sch<<f_act exA
-   is used! --> to generate a contradiction using  ~a\<leadsto>ss<< UU(nil), using theorems
-   Cons_not_less_UU and Cons_not_less_nil  *)
-apply (tactic \<open>Seq_case_simp_tac @{context} "exB" 1\<close>)
-(* Case exA=a\<leadsto>x, exB=b\<leadsto>y *)
-(* here it is important that Seq_case_simp_tac uses no !full!_simp_tac for the cons case,
-   as otherwise mkex_cons_3 would  not be rewritten without use of rotate_tac: then tactic
-   would not be generally applicable *)
-apply simp
-
-(* Case y:A, y~:B *)
-apply (tactic \<open>Seq_case_simp_tac @{context} "exA" 1\<close>)
-apply simp
-
-(* Case y~:A, y:B *)
-apply (tactic \<open>Seq_case_simp_tac @{context} "exB" 1\<close>)
-apply simp
-
-(* Case y~:A, y~:B *)
-apply (simp add: asig_of_par actions_asig_comp)
-done
-
-
-(* generalizing the proof above to a proof method *)
-
-ML \<open>
-fun mkex_induct_tac ctxt sch exA exB =
-  EVERY'[Seq_induct_tac ctxt sch @{thms Filter_def Forall_def sforall_def mkex_def stutter_def},
-         asm_full_simp_tac ctxt,
-         SELECT_GOAL
-          (safe_tac (Context.raw_transfer (Proof_Context.theory_of ctxt) @{theory_context Fun})),
-         Seq_case_simp_tac ctxt exA,
-         Seq_case_simp_tac ctxt exB,
-         asm_full_simp_tac ctxt,
-         Seq_case_simp_tac ctxt exA,
-         asm_full_simp_tac ctxt,
-         Seq_case_simp_tac ctxt exB,
-         asm_full_simp_tac ctxt,
-         asm_full_simp_tac (ctxt addsimps @{thms asig_of_par actions_asig_comp})
-        ]
-\<close>
-
-method_setup mkex_induct = \<open>
-  Scan.lift (Args.name -- Args.name -- Args.name)
-    >> (fn ((sch, exA), exB) => fn ctxt => SIMPLE_METHOD' (mkex_induct_tac ctxt sch exA exB))
-\<close>
-
-
-(*---------------------------------------------------------------------------
-               Projection of mkex(sch,exA,exB) onto A stutters on A
-                             structural induction
-  --------------------------------------------------------------------------- *)
-
-lemma stutterA_mkex: "! exA exB s t.
-  Forall (%x. x:act (A\<parallel>B)) sch &
-  Filter (%a. a:act A)$sch << filter_act$exA &
-  Filter (%a. a:act B)$sch << filter_act$exB
-  --> stutter (asig_of A) (s,ProjA2$(snd (mkex A B sch (s,exA) (t,exB))))"
-  by (mkex_induct sch exA exB)
-
-lemma stutter_mkex_on_A: "[|
-  Forall (%x. x:act (A\<parallel>B)) sch ;
-  Filter (%a. a:act A)$sch << filter_act$(snd exA) ;
-  Filter (%a. a:act B)$sch << filter_act$(snd exB) |]
-  ==> stutter (asig_of A) (ProjA (mkex A B sch exA exB))"
-
-apply (cut_tac stutterA_mkex)
-apply (simp add: stutter_def ProjA_def mkex_def)
-apply (erule allE)+
-apply (drule mp)
-prefer 2 apply (assumption)
-apply simp
-done
-
-
-(*---------------------------------------------------------------------------
-               Projection of mkex(sch,exA,exB) onto B stutters on B
-                             structural induction
-  --------------------------------------------------------------------------- *)
-
-lemma stutterB_mkex: "! exA exB s t.
-  Forall (%x. x:act (A\<parallel>B)) sch &
-  Filter (%a. a:act A)$sch << filter_act$exA &
-  Filter (%a. a:act B)$sch << filter_act$exB
-  --> stutter (asig_of B) (t,ProjB2$(snd (mkex A B sch (s,exA) (t,exB))))"
-  by (mkex_induct sch exA exB)
-
-
-lemma stutter_mkex_on_B: "[|
-  Forall (%x. x:act (A\<parallel>B)) sch ;
-  Filter (%a. a:act A)$sch << filter_act$(snd exA) ;
-  Filter (%a. a:act B)$sch << filter_act$(snd exB) |]
-  ==> stutter (asig_of B) (ProjB (mkex A B sch exA exB))"
-apply (cut_tac stutterB_mkex)
-apply (simp add: stutter_def ProjB_def mkex_def)
-apply (erule allE)+
-apply (drule mp)
-prefer 2 apply (assumption)
-apply simp
-done
-
-
-(*---------------------------------------------------------------------------
-     Filter of mkex(sch,exA,exB) to A after projection onto A is exA
-        --  using zip$(proj1$exA)$(proj2$exA) instead of exA    --
-        --           because of admissibility problems          --
-                             structural induction
-  --------------------------------------------------------------------------- *)
-
-lemma filter_mkex_is_exA_tmp: "! exA exB s t.
-  Forall (%x. x:act (A\<parallel>B)) sch &
-  Filter (%a. a:act A)$sch << filter_act$exA  &
-  Filter (%a. a:act B)$sch << filter_act$exB
-  --> Filter_ex2 (asig_of A)$(ProjA2$(snd (mkex A B sch (s,exA) (t,exB)))) =
-      Zip$(Filter (%a. a:act A)$sch)$(Map snd$exA)"
-  by (mkex_induct sch exB exA)
-
-(*---------------------------------------------------------------------------
-                      zip$(proj1$y)$(proj2$y) = y   (using the lift operations)
-                    lemma for admissibility problems
-  --------------------------------------------------------------------------- *)
-
-lemma Zip_Map_fst_snd: "Zip$(Map fst$y)$(Map snd$y) = y"
-apply (tactic \<open>Seq_induct_tac @{context} "y" [] 1\<close>)
-done
-
-
-(*---------------------------------------------------------------------------
-      filter A$sch = proj1$ex   -->  zip$(filter A$sch)$(proj2$ex) = ex
-         lemma for eliminating non admissible equations in assumptions
-  --------------------------------------------------------------------------- *)
-
-lemma trick_against_eq_in_ass: "!! sch ex.
-  Filter (%a. a:act AB)$sch = filter_act$ex
-  ==> ex = Zip$(Filter (%a. a:act AB)$sch)$(Map snd$ex)"
-apply (simp add: filter_act_def)
-apply (rule Zip_Map_fst_snd [symmetric])
-done
-
-(*---------------------------------------------------------------------------
-     Filter of mkex(sch,exA,exB) to A after projection onto A is exA
-                       using the above trick
-  --------------------------------------------------------------------------- *)
-
-
-lemma filter_mkex_is_exA: "!!sch exA exB.
-  [| Forall (%a. a:act (A\<parallel>B)) sch ;
-  Filter (%a. a:act A)$sch = filter_act$(snd exA)  ;
-  Filter (%a. a:act B)$sch = filter_act$(snd exB) |]
-  ==> Filter_ex (asig_of A) (ProjA (mkex A B sch exA exB)) = exA"
-apply (simp add: ProjA_def Filter_ex_def)
-apply (tactic \<open>pair_tac @{context} "exA" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "exB" 1\<close>)
-apply (rule conjI)
-apply (simp (no_asm) add: mkex_def)
-apply (simplesubst trick_against_eq_in_ass)
-back
-apply assumption
-apply (simp add: filter_mkex_is_exA_tmp)
-done
-
-
-(*---------------------------------------------------------------------------
-     Filter of mkex(sch,exA,exB) to B after projection onto B is exB
-        --  using zip$(proj1$exB)$(proj2$exB) instead of exB    --
-        --           because of admissibility problems          --
-                             structural induction
-  --------------------------------------------------------------------------- *)
-
-lemma filter_mkex_is_exB_tmp: "! exA exB s t.
-  Forall (%x. x:act (A\<parallel>B)) sch &
-  Filter (%a. a:act A)$sch << filter_act$exA  &
-  Filter (%a. a:act B)$sch << filter_act$exB
-  --> Filter_ex2 (asig_of B)$(ProjB2$(snd (mkex A B sch (s,exA) (t,exB)))) =
-      Zip$(Filter (%a. a:act B)$sch)$(Map snd$exB)"
-
-(* notice necessary change of arguments exA and exB *)
-  by (mkex_induct sch exA exB)
-
-
-(*---------------------------------------------------------------------------
-     Filter of mkex(sch,exA,exB) to A after projection onto B is exB
-                       using the above trick
-  --------------------------------------------------------------------------- *)
-
-
-lemma filter_mkex_is_exB: "!!sch exA exB.
-  [| Forall (%a. a:act (A\<parallel>B)) sch ;
-  Filter (%a. a:act A)$sch = filter_act$(snd exA)  ;
-  Filter (%a. a:act B)$sch = filter_act$(snd exB) |]
-  ==> Filter_ex (asig_of B) (ProjB (mkex A B sch exA exB)) = exB"
-apply (simp add: ProjB_def Filter_ex_def)
-apply (tactic \<open>pair_tac @{context} "exA" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "exB" 1\<close>)
-apply (rule conjI)
-apply (simp (no_asm) add: mkex_def)
-apply (simplesubst trick_against_eq_in_ass)
-back
-apply assumption
-apply (simp add: filter_mkex_is_exB_tmp)
-done
-
-(* --------------------------------------------------------------------- *)
-(*                    mkex has only  A- or B-actions                    *)
-(* --------------------------------------------------------------------- *)
-
-
-lemma mkex_actions_in_AorB: "!s t exA exB.
-  Forall (%x. x : act (A \<parallel> B)) sch &
-  Filter (%a. a:act A)$sch << filter_act$exA  &
-  Filter (%a. a:act B)$sch << filter_act$exB
-   --> Forall (%x. fst x : act (A \<parallel>B))
-         (snd (mkex A B sch (s,exA) (t,exB)))"
-  by (mkex_induct sch exA exB)
-
-
-(* ------------------------------------------------------------------ *)
-(*           COMPOSITIONALITY   on    SCHEDULE      Level             *)
-(*                          Main Theorem                              *)
-(* ------------------------------------------------------------------ *)
-
-lemma compositionality_sch:
-"(sch : schedules (A\<parallel>B)) =
-  (Filter (%a. a:act A)$sch : schedules A &
-   Filter (%a. a:act B)$sch : schedules B &
-   Forall (%x. x:act (A\<parallel>B)) sch)"
-apply (simp add: schedules_def has_schedule_def)
-apply auto
-(* ==> *)
-apply (rule_tac x = "Filter_ex (asig_of A) (ProjA ex) " in bexI)
-prefer 2
-apply (simp add: compositionality_ex)
-apply (simp (no_asm) add: Filter_ex_def ProjA_def lemma_2_1a lemma_2_1b)
-apply (rule_tac x = "Filter_ex (asig_of B) (ProjB ex) " in bexI)
-prefer 2
-apply (simp add: compositionality_ex)
-apply (simp (no_asm) add: Filter_ex_def ProjB_def lemma_2_1a lemma_2_1b)
-apply (simp add: executions_def)
-apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-apply (erule conjE)
-apply (simp add: sch_actions_in_AorB)
-
-(* <== *)
-
-(* mkex is exactly the construction of exA\<parallel>B out of exA, exB, and the oracle sch,
-   we need here *)
-apply (rename_tac exA exB)
-apply (rule_tac x = "mkex A B sch exA exB" in bexI)
-(* mkex actions are just the oracle *)
-apply (tactic \<open>pair_tac @{context} "exA" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "exB" 1\<close>)
-apply (simp add: Mapfst_mkex_is_sch)
-
-(* mkex is an execution -- use compositionality on ex-level *)
-apply (simp add: compositionality_ex)
-apply (simp add: stutter_mkex_on_A stutter_mkex_on_B filter_mkex_is_exB filter_mkex_is_exA)
-apply (tactic \<open>pair_tac @{context} "exA" 1\<close>)
-apply (tactic \<open>pair_tac @{context} "exB" 1\<close>)
-apply (simp add: mkex_actions_in_AorB)
-done
-
-
-subsection \<open>COMPOSITIONALITY on SCHEDULE Level -- for Modules\<close>
-
-lemma compositionality_sch_modules:
-  "Scheds (A\<parallel>B) = par_scheds (Scheds A) (Scheds B)"
-
-apply (unfold Scheds_def par_scheds_def)
-apply (simp add: asig_of_par)
-apply (rule set_eqI)
-apply (simp add: compositionality_sch actions_of_par)
-done
-
-
-declare compoex_simps [simp del]
-declare composch_simps [simp del]
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/CompoTraces.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,969 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/CompoTraces.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Compositionality on Trace level\<close>
-
-theory CompoTraces
-imports CompoScheds ShortExecutions
-begin
-
-definition mksch :: "('a,'s)ioa => ('a,'t)ioa => 'a Seq -> 'a Seq -> 'a Seq -> 'a Seq"
-where
-  "mksch A B = (fix$(LAM h tr schA schB. case tr of
-       nil => nil
-    | x##xs =>
-      (case x of
-        UU => UU
-      | Def y =>
-         (if y:act A then
-             (if y:act B then
-                   ((Takewhile (%a. a:int A)$schA)
-                      @@ (Takewhile (%a. a:int B)$schB)
-                           @@ (y\<leadsto>(h$xs
-                                    $(TL$(Dropwhile (%a. a:int A)$schA))
-                                    $(TL$(Dropwhile (%a. a:int B)$schB))
-                    )))
-              else
-                 ((Takewhile (%a. a:int A)$schA)
-                  @@ (y\<leadsto>(h$xs
-                           $(TL$(Dropwhile (%a. a:int A)$schA))
-                           $schB)))
-              )
-          else
-             (if y:act B then
-                 ((Takewhile (%a. a:int B)$schB)
-                     @@ (y\<leadsto>(h$xs
-                              $schA
-                              $(TL$(Dropwhile (%a. a:int B)$schB))
-                              )))
-             else
-               UU
-             )
-         )
-       )))"
-
-definition par_traces ::"['a trace_module,'a trace_module] => 'a trace_module"
-where
-  "par_traces TracesA TracesB =
-      (let trA = fst TracesA; sigA = snd TracesA;
-           trB = fst TracesB; sigB = snd TracesB
-       in
-       (    {tr. Filter (%a. a:actions sigA)$tr : trA}
-        Int {tr. Filter (%a. a:actions sigB)$tr : trB}
-        Int {tr. Forall (%x. x:(externals sigA Un externals sigB)) tr},
-        asig_comp sigA sigB))"
-
-axiomatization
-where
-  finiteR_mksch:
-    "Finite (mksch A B$tr$x$y) \<Longrightarrow> Finite tr"
-
-lemma finiteR_mksch': "\<not> Finite tr \<Longrightarrow> \<not> Finite (mksch A B$tr$x$y)"
-  by (blast intro: finiteR_mksch)
-
-
-declaration \<open>fn _ => Simplifier.map_ss (Simplifier.set_mksym (K (K NONE)))\<close>
-
-
-subsection "mksch rewrite rules"
-
-lemma mksch_unfold:
-"mksch A B = (LAM tr schA schB. case tr of
-       nil => nil
-    | x##xs =>
-      (case x of
-        UU => UU
-      | Def y =>
-         (if y:act A then
-             (if y:act B then
-                   ((Takewhile (%a. a:int A)$schA)
-                         @@(Takewhile (%a. a:int B)$schB)
-                              @@(y\<leadsto>(mksch A B$xs
-                                       $(TL$(Dropwhile (%a. a:int A)$schA))
-                                       $(TL$(Dropwhile (%a. a:int B)$schB))
-                    )))
-              else
-                 ((Takewhile (%a. a:int A)$schA)
-                      @@ (y\<leadsto>(mksch A B$xs
-                              $(TL$(Dropwhile (%a. a:int A)$schA))
-                              $schB)))
-              )
-          else
-             (if y:act B then
-                 ((Takewhile (%a. a:int B)$schB)
-                       @@ (y\<leadsto>(mksch A B$xs
-                              $schA
-                              $(TL$(Dropwhile (%a. a:int B)$schB))
-                              )))
-             else
-               UU
-             )
-         )
-       ))"
-apply (rule trans)
-apply (rule fix_eq4)
-apply (rule mksch_def)
-apply (rule beta_cfun)
-apply simp
-done
-
-lemma mksch_UU: "mksch A B$UU$schA$schB = UU"
-apply (subst mksch_unfold)
-apply simp
-done
-
-lemma mksch_nil: "mksch A B$nil$schA$schB = nil"
-apply (subst mksch_unfold)
-apply simp
-done
-
-lemma mksch_cons1: "[|x:act A;x~:act B|]
-    ==> mksch A B$(x\<leadsto>tr)$schA$schB =
-          (Takewhile (%a. a:int A)$schA)
-          @@ (x\<leadsto>(mksch A B$tr$(TL$(Dropwhile (%a. a:int A)$schA))
-                              $schB))"
-apply (rule trans)
-apply (subst mksch_unfold)
-apply (simp add: Consq_def If_and_if)
-apply (simp add: Consq_def)
-done
-
-lemma mksch_cons2: "[|x~:act A;x:act B|]
-    ==> mksch A B$(x\<leadsto>tr)$schA$schB =
-         (Takewhile (%a. a:int B)$schB)
-          @@ (x\<leadsto>(mksch A B$tr$schA$(TL$(Dropwhile (%a. a:int B)$schB))
-                             ))"
-apply (rule trans)
-apply (subst mksch_unfold)
-apply (simp add: Consq_def If_and_if)
-apply (simp add: Consq_def)
-done
-
-lemma mksch_cons3: "[|x:act A;x:act B|]
-    ==> mksch A B$(x\<leadsto>tr)$schA$schB =
-             (Takewhile (%a. a:int A)$schA)
-          @@ ((Takewhile (%a. a:int B)$schB)
-          @@ (x\<leadsto>(mksch A B$tr$(TL$(Dropwhile (%a. a:int A)$schA))
-                             $(TL$(Dropwhile (%a. a:int B)$schB))))
-              )"
-apply (rule trans)
-apply (subst mksch_unfold)
-apply (simp add: Consq_def If_and_if)
-apply (simp add: Consq_def)
-done
-
-lemmas compotr_simps = mksch_UU mksch_nil mksch_cons1 mksch_cons2 mksch_cons3
-
-declare compotr_simps [simp]
-
-
-subsection \<open>COMPOSITIONALITY on TRACE Level\<close>
-
-subsubsection "Lemmata for ==>"
-
-(* Consequence out of ext1_ext2_is_not_act1(2), which in turn are consequences out of
-   the compatibility of IOA, in particular out of the condition that internals are
-   really hidden. *)
-
-lemma compatibility_consequence1: "(eB & ~eA --> ~A) -->
-          (A & (eA | eB)) = (eA & A)"
-apply fast
-done
-
-
-(* very similar to above, only the commutativity of | is used to make a slight change *)
-
-lemma compatibility_consequence2: "(eB & ~eA --> ~A) -->
-          (A & (eB | eA)) = (eA & A)"
-apply fast
-done
-
-
-subsubsection "Lemmata for <=="
-
-(* Lemma for substitution of looping assumption in another specific assumption *)
-lemma subst_lemma1: "[| f << (g x) ; x=(h x) |] ==> f << g (h x)"
-by (erule subst)
-
-(* Lemma for substitution of looping assumption in another specific assumption *)
-lemma subst_lemma2: "[| (f x) = y \<leadsto> g; x=(h x) |] ==> (f (h x)) = y \<leadsto> g"
-by (erule subst)
-
-lemma ForallAorB_mksch [rule_format]:
-  "!!A B. compatible A B ==>
-    ! schA schB. Forall (%x. x:act (A\<parallel>B)) tr
-    --> Forall (%x. x:act (A\<parallel>B)) (mksch A B$tr$schA$schB)"
-apply (tactic \<open>Seq_induct_tac @{context} "tr"
-  [@{thm Forall_def}, @{thm sforall_def}, @{thm mksch_def}] 1\<close>)
-apply auto
-apply (simp add: actions_of_par)
-apply (case_tac "a:act A")
-apply (case_tac "a:act B")
-(* a:A, a:B *)
-apply simp
-apply (rule Forall_Conc_impl [THEN mp])
-apply (simp add: intA_is_not_actB int_is_act)
-apply (rule Forall_Conc_impl [THEN mp])
-apply (simp add: intA_is_not_actB int_is_act)
-(* a:A,a~:B *)
-apply simp
-apply (rule Forall_Conc_impl [THEN mp])
-apply (simp add: intA_is_not_actB int_is_act)
-apply (case_tac "a:act B")
-(* a~:A, a:B *)
-apply simp
-apply (rule Forall_Conc_impl [THEN mp])
-apply (simp add: intA_is_not_actB int_is_act)
-(* a~:A,a~:B *)
-apply auto
-done
-
-lemma ForallBnAmksch [rule_format (no_asm)]: "!!A B. compatible B A  ==>
-    ! schA schB.  (Forall (%x. x:act B & x~:act A) tr
-    --> Forall (%x. x:act B & x~:act A) (mksch A B$tr$schA$schB))"
-apply (tactic \<open>Seq_induct_tac @{context} "tr"
-  [@{thm Forall_def}, @{thm sforall_def}, @{thm mksch_def}] 1\<close>)
-apply auto
-apply (rule Forall_Conc_impl [THEN mp])
-apply (simp add: intA_is_not_actB int_is_act)
-done
-
-lemma ForallAnBmksch [rule_format (no_asm)]: "!!A B. compatible A B ==>
-    ! schA schB.  (Forall (%x. x:act A & x~:act B) tr
-    --> Forall (%x. x:act A & x~:act B) (mksch A B$tr$schA$schB))"
-apply (tactic \<open>Seq_induct_tac @{context} "tr"
-  [@{thm Forall_def}, @{thm sforall_def}, @{thm mksch_def}] 1\<close>)
-apply auto
-apply (rule Forall_Conc_impl [THEN mp])
-apply (simp add: intA_is_not_actB int_is_act)
-done
-
-(* safe-tac makes too many case distinctions with this lemma in the next proof *)
-declare FiniteConc [simp del]
-
-lemma FiniteL_mksch [rule_format (no_asm)]: "[| Finite tr; is_asig(asig_of A); is_asig(asig_of B) |] ==>
-    ! x y. Forall (%x. x:act A) x & Forall (%x. x:act B) y &
-           Filter (%a. a:ext A)$x = Filter (%a. a:act A)$tr &
-           Filter (%a. a:ext B)$y = Filter (%a. a:act B)$tr &
-           Forall (%x. x:ext (A\<parallel>B)) tr
-           --> Finite (mksch A B$tr$x$y)"
-
-apply (erule Seq_Finite_ind)
-apply simp
-(* main case *)
-apply simp
-apply auto
-
-(* a: act A; a: act B *)
-apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
-apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
-back
-apply (erule conjE)+
-(* Finite (tw iA x) and Finite (tw iB y) *)
-apply (simp add: not_ext_is_int_or_not_act FiniteConc)
-(* now for conclusion IH applicable, but assumptions have to be transformed *)
-apply (drule_tac x = "x" and g = "Filter (%a. a:act A) $s" in subst_lemma2)
-apply assumption
-apply (drule_tac x = "y" and g = "Filter (%a. a:act B) $s" in subst_lemma2)
-apply assumption
-(* IH *)
-apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
-
-(* a: act B; a~: act A *)
-apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
-
-apply (erule conjE)+
-(* Finite (tw iB y) *)
-apply (simp add: not_ext_is_int_or_not_act FiniteConc)
-(* now for conclusion IH applicable, but assumptions have to be transformed *)
-apply (drule_tac x = "y" and g = "Filter (%a. a:act B) $s" in subst_lemma2)
-apply assumption
-(* IH *)
-apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
-
-(* a~: act B; a: act A *)
-apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
-
-apply (erule conjE)+
-(* Finite (tw iA x) *)
-apply (simp add: not_ext_is_int_or_not_act FiniteConc)
-(* now for conclusion IH applicable, but assumptions have to be transformed *)
-apply (drule_tac x = "x" and g = "Filter (%a. a:act A) $s" in subst_lemma2)
-apply assumption
-(* IH *)
-apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
-
-(* a~: act B; a~: act A *)
-apply (fastforce intro!: ext_is_act simp: externals_of_par)
-done
-
-declare FiniteConc [simp]
-
-declare FilterConc [simp del]
-
-lemma reduceA_mksch1 [rule_format (no_asm)]: " [| Finite bs; is_asig(asig_of A); is_asig(asig_of B);compatible A B|] ==>
- ! y. Forall (%x. x:act B) y & Forall (%x. x:act B & x~:act A) bs &
-     Filter (%a. a:ext B)$y = Filter (%a. a:act B)$(bs @@ z)
-     --> (? y1 y2.  (mksch A B$(bs @@ z)$x$y) = (y1 @@ (mksch A B$z$x$y2)) &
-                    Forall (%x. x:act B & x~:act A) y1 &
-                    Finite y1 & y = (y1 @@ y2) &
-                    Filter (%a. a:ext B)$y1 = bs)"
-apply (frule_tac A1 = "A" in compat_commute [THEN iffD1])
-apply (erule Seq_Finite_ind)
-apply (rule allI)+
-apply (rule impI)
-apply (rule_tac x = "nil" in exI)
-apply (rule_tac x = "y" in exI)
-apply simp
-(* main case *)
-apply (rule allI)+
-apply (rule impI)
-apply simp
-apply (erule conjE)+
-apply simp
-(* divide_Seq on s *)
-apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
-apply (erule conjE)+
-(* transform assumption f eB y = f B (s@z) *)
-apply (drule_tac x = "y" and g = "Filter (%a. a:act B) $ (s@@z) " in subst_lemma2)
-apply assumption
-apply (simp add: not_ext_is_int_or_not_act FilterConc)
-(* apply IH *)
-apply (erule_tac x = "TL$ (Dropwhile (%a. a:int B) $y) " in allE)
-apply (simp add: ForallTL ForallDropwhile FilterConc)
-apply (erule exE)+
-apply (erule conjE)+
-apply (simp add: FilterConc)
-(* for replacing IH in conclusion *)
-apply (rotate_tac -2)
-(* instantiate y1a and y2a *)
-apply (rule_tac x = "Takewhile (%a. a:int B) $y @@ a\<leadsto>y1" in exI)
-apply (rule_tac x = "y2" in exI)
-(* elminate all obligations up to two depending on Conc_assoc *)
-apply (simp add: intA_is_not_actB int_is_act int_is_not_ext FilterConc)
-apply (simp (no_asm) add: Conc_assoc FilterConc)
-done
-
-lemmas reduceA_mksch = conjI [THEN [6] conjI [THEN [5] reduceA_mksch1]]
-
-lemma reduceB_mksch1 [rule_format]:
-" [| Finite a_s; is_asig(asig_of A); is_asig(asig_of B);compatible A B|] ==>
- ! x. Forall (%x. x:act A) x & Forall (%x. x:act A & x~:act B) a_s &
-     Filter (%a. a:ext A)$x = Filter (%a. a:act A)$(a_s @@ z)
-     --> (? x1 x2.  (mksch A B$(a_s @@ z)$x$y) = (x1 @@ (mksch A B$z$x2$y)) &
-                    Forall (%x. x:act A & x~:act B) x1 &
-                    Finite x1 & x = (x1 @@ x2) &
-                    Filter (%a. a:ext A)$x1 = a_s)"
-apply (frule_tac A1 = "A" in compat_commute [THEN iffD1])
-apply (erule Seq_Finite_ind)
-apply (rule allI)+
-apply (rule impI)
-apply (rule_tac x = "nil" in exI)
-apply (rule_tac x = "x" in exI)
-apply simp
-(* main case *)
-apply (rule allI)+
-apply (rule impI)
-apply simp
-apply (erule conjE)+
-apply simp
-(* divide_Seq on s *)
-apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
-apply (erule conjE)+
-(* transform assumption f eA x = f A (s@z) *)
-apply (drule_tac x = "x" and g = "Filter (%a. a:act A) $ (s@@z) " in subst_lemma2)
-apply assumption
-apply (simp add: not_ext_is_int_or_not_act FilterConc)
-(* apply IH *)
-apply (erule_tac x = "TL$ (Dropwhile (%a. a:int A) $x) " in allE)
-apply (simp add: ForallTL ForallDropwhile FilterConc)
-apply (erule exE)+
-apply (erule conjE)+
-apply (simp add: FilterConc)
-(* for replacing IH in conclusion *)
-apply (rotate_tac -2)
-(* instantiate y1a and y2a *)
-apply (rule_tac x = "Takewhile (%a. a:int A) $x @@ a\<leadsto>x1" in exI)
-apply (rule_tac x = "x2" in exI)
-(* elminate all obligations up to two depending on Conc_assoc *)
-apply (simp add: intA_is_not_actB int_is_act int_is_not_ext FilterConc)
-apply (simp (no_asm) add: Conc_assoc FilterConc)
-done
-
-lemmas reduceB_mksch = conjI [THEN [6] conjI [THEN [5] reduceB_mksch1]]
-
-declare FilterConc [simp]
-
-
-subsection "Filtering external actions out of mksch(tr,schA,schB) yields the oracle tr"
-
-lemma FilterA_mksch_is_tr:
-"!! A B. [| compatible A B; compatible B A;
-            is_asig(asig_of A); is_asig(asig_of B) |] ==>
-  ! schA schB. Forall (%x. x:act A) schA & Forall (%x. x:act B) schB &
-  Forall (%x. x:ext (A\<parallel>B)) tr &
-  Filter (%a. a:act A)$tr << Filter (%a. a:ext A)$schA &
-  Filter (%a. a:act B)$tr << Filter (%a. a:ext B)$schB
-  --> Filter (%a. a:ext (A\<parallel>B))$(mksch A B$tr$schA$schB) = tr"
-
-apply (tactic \<open>Seq_induct_tac @{context} "tr"
-  [@{thm Forall_def}, @{thm sforall_def}, @{thm mksch_def}] 1\<close>)
-(* main case *)
-(* splitting into 4 cases according to a:A, a:B *)
-apply auto
-
-(* Case a:A, a:B *)
-apply (frule divide_Seq)
-apply (frule divide_Seq)
-back
-apply (erule conjE)+
-(* filtering internals of A in schA and of B in schB is nil *)
-apply (simp add: not_ext_is_int_or_not_act externals_of_par intA_is_not_extB int_is_not_ext)
-(* conclusion of IH ok, but assumptions of IH have to be transformed *)
-apply (drule_tac x = "schA" in subst_lemma1)
-apply assumption
-apply (drule_tac x = "schB" in subst_lemma1)
-apply assumption
-(* IH *)
-apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
-
-(* Case a:A, a~:B *)
-apply (frule divide_Seq)
-apply (erule conjE)+
-(* filtering internals of A is nil *)
-apply (simp add: not_ext_is_int_or_not_act externals_of_par intA_is_not_extB int_is_not_ext)
-apply (drule_tac x = "schA" in subst_lemma1)
-apply assumption
-apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
-
-(* Case a:B, a~:A *)
-apply (frule divide_Seq)
-apply (erule conjE)+
-(* filtering internals of A is nil *)
-apply (simp add: not_ext_is_int_or_not_act externals_of_par intA_is_not_extB int_is_not_ext)
-apply (drule_tac x = "schB" in subst_lemma1)
-back
-apply assumption
-apply (simp add: not_ext_is_int_or_not_act ForallTL ForallDropwhile)
-
-(* Case a~:A, a~:B *)
-apply (fastforce intro!: ext_is_act simp: externals_of_par)
-done
-
-
-subsection" Filter of mksch(tr,schA,schB) to A is schA -- take lemma proof"
-
-lemma FilterAmksch_is_schA: "!! A B. [| compatible A B; compatible B A;
-  is_asig(asig_of A); is_asig(asig_of B) |] ==>
-  Forall (%x. x:ext (A\<parallel>B)) tr &
-  Forall (%x. x:act A) schA & Forall (%x. x:act B) schB &
-  Filter (%a. a:ext A)$schA = Filter (%a. a:act A)$tr &
-  Filter (%a. a:ext B)$schB = Filter (%a. a:act B)$tr &
-  LastActExtsch A schA & LastActExtsch B schB
-  --> Filter (%a. a:act A)$(mksch A B$tr$schA$schB) = schA"
-apply (intro strip)
-apply (rule seq.take_lemma)
-apply (rule mp)
-prefer 2 apply assumption
-back back back back
-apply (rule_tac x = "schA" in spec)
-apply (rule_tac x = "schB" in spec)
-apply (rule_tac x = "tr" in spec)
-apply (tactic "thin_tac' @{context} 5 1")
-apply (rule nat_less_induct)
-apply (rule allI)+
-apply (rename_tac tr schB schA)
-apply (intro strip)
-apply (erule conjE)+
-
-apply (case_tac "Forall (%x. x:act B & x~:act A) tr")
-
-apply (rule seq_take_lemma [THEN iffD2, THEN spec])
-apply (tactic "thin_tac' @{context} 5 1")
-
-
-apply (case_tac "Finite tr")
-
-(* both sides of this equation are nil *)
-apply (subgoal_tac "schA=nil")
-apply (simp (no_asm_simp))
-(* first side: mksch = nil *)
-apply (auto intro!: ForallQFilterPnil ForallBnAmksch FiniteL_mksch)[1]
-(* second side: schA = nil *)
-apply (erule_tac A = "A" in LastActExtimplnil)
-apply (simp (no_asm_simp))
-apply (erule_tac Q = "%x. x:act B & x~:act A" in ForallQFilterPnil)
-apply assumption
-apply fast
-
-(* case ~ Finite s *)
-
-(* both sides of this equation are UU *)
-apply (subgoal_tac "schA=UU")
-apply (simp (no_asm_simp))
-(* first side: mksch = UU *)
-apply (auto intro!: ForallQFilterPUU finiteR_mksch' ForallBnAmksch)[1]
-(* schA = UU *)
-apply (erule_tac A = "A" in LastActExtimplUU)
-apply (simp (no_asm_simp))
-apply (erule_tac Q = "%x. x:act B & x~:act A" in ForallQFilterPUU)
-apply assumption
-apply fast
-
-(* case" ~ Forall (%x.x:act B & x~:act A) s" *)
-
-apply (drule divide_Seq3)
-
-apply (erule exE)+
-apply (erule conjE)+
-apply hypsubst
-
-(* bring in lemma reduceA_mksch *)
-apply (frule_tac x = "schA" and y = "schB" and A = "A" and B = "B" in reduceA_mksch)
-apply assumption+
-apply (erule exE)+
-apply (erule conjE)+
-
-(* use reduceA_mksch to rewrite conclusion *)
-apply hypsubst
-apply simp
-
-(* eliminate the B-only prefix *)
-
-apply (subgoal_tac " (Filter (%a. a :act A) $y1) = nil")
-apply (erule_tac [2] ForallQFilterPnil)
-prefer 2 apply assumption
-prefer 2 apply fast
-
-(* Now real recursive step follows (in y) *)
-
-apply simp
-apply (case_tac "x:act A")
-apply (case_tac "x~:act B")
-apply (rotate_tac -2)
-apply simp
-
-apply (subgoal_tac "Filter (%a. a:act A & a:ext B) $y1=nil")
-apply (rotate_tac -1)
-apply simp
-(* eliminate introduced subgoal 2 *)
-apply (erule_tac [2] ForallQFilterPnil)
-prefer 2 apply assumption
-prefer 2 apply fast
-
-(* bring in divide Seq for s *)
-apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
-apply (erule conjE)+
-
-(* subst divide_Seq in conclusion, but only at the righest occurrence *)
-apply (rule_tac t = "schA" in ssubst)
-back
-back
-back
-apply assumption
-
-(* reduce trace_takes from n to strictly smaller k *)
-apply (rule take_reduction)
-
-(* f A (tw iA) = tw ~eA *)
-apply (simp add: int_is_act not_ext_is_int_or_not_act)
-apply (rule refl)
-apply (simp add: int_is_act not_ext_is_int_or_not_act)
-apply (rotate_tac -11)
-
-(* now conclusion fulfills induction hypothesis, but assumptions are not ready *)
-
-(* assumption Forall tr *)
-(* assumption schB *)
-apply (simp add: ext_and_act)
-(* assumption schA *)
-apply (drule_tac x = "schA" and g = "Filter (%a. a:act A) $rs" in subst_lemma2)
-apply assumption
-apply (simp add: int_is_not_ext)
-(* assumptions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping  *)
-apply (drule_tac sch = "schA" and P = "%a. a:int A" in LastActExtsmall1)
-apply (frule_tac ?sch1.0 = "y1" in LastActExtsmall2)
-apply assumption
-
-(* assumption Forall schA *)
-apply (drule_tac s = "schA" and P = "Forall (%x. x:act A) " in subst)
-apply assumption
-apply (simp add: int_is_act)
-
-(* case x:actions(asig_of A) & x: actions(asig_of B) *)
-
-
-apply (rotate_tac -2)
-apply simp
-
-apply (subgoal_tac "Filter (%a. a:act A & a:ext B) $y1=nil")
-apply (rotate_tac -1)
-apply simp
-(* eliminate introduced subgoal 2 *)
-apply (erule_tac [2] ForallQFilterPnil)
-prefer 2 apply (assumption)
-prefer 2 apply (fast)
-
-(* bring in divide Seq for s *)
-apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
-apply (erule conjE)+
-
-(* subst divide_Seq in conclusion, but only at the rightmost occurrence *)
-apply (rule_tac t = "schA" in ssubst)
-back
-back
-back
-apply assumption
-
-(* f A (tw iA) = tw ~eA *)
-apply (simp add: int_is_act not_ext_is_int_or_not_act)
-
-(* rewrite assumption forall and schB *)
-apply (rotate_tac 13)
-apply (simp add: ext_and_act)
-
-(* divide_Seq for schB2 *)
-apply (frule_tac y = "y2" in sym [THEN eq_imp_below, THEN divide_Seq])
-apply (erule conjE)+
-(* assumption schA *)
-apply (drule_tac x = "schA" and g = "Filter (%a. a:act A) $rs" in subst_lemma2)
-apply assumption
-apply (simp add: int_is_not_ext)
-
-(* f A (tw iB schB2) = nil *)
-apply (simp add: int_is_not_ext not_ext_is_int_or_not_act intA_is_not_actB)
-
-
-(* reduce trace_takes from n to strictly smaller k *)
-apply (rule take_reduction)
-apply (rule refl)
-apply (rule refl)
-
-(* now conclusion fulfills induction hypothesis, but assumptions are not all ready *)
-
-(* assumption schB *)
-apply (drule_tac x = "y2" and g = "Filter (%a. a:act B) $rs" in subst_lemma2)
-apply assumption
-apply (simp add: intA_is_not_actB int_is_not_ext)
-
-(* conclusions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping  *)
-apply (drule_tac sch = "schA" and P = "%a. a:int A" in LastActExtsmall1)
-apply (frule_tac ?sch1.0 = "y1" in LastActExtsmall2)
-apply assumption
-apply (drule_tac sch = "y2" and P = "%a. a:int B" in LastActExtsmall1)
-
-(* assumption Forall schA, and Forall schA are performed by ForallTL,ForallDropwhile *)
-apply (simp add: ForallTL ForallDropwhile)
-
-(* case x~:A & x:B  *)
-(* cannot occur, as just this case has been scheduled out before as the B-only prefix *)
-apply (case_tac "x:act B")
-apply fast
-
-(* case x~:A & x~:B  *)
-(* cannot occur because of assumption: Forall (a:ext A | a:ext B) *)
-apply (rotate_tac -9)
-(* reduce forall assumption from tr to (x\<leadsto>rs) *)
-apply (simp add: externals_of_par)
-apply (fast intro!: ext_is_act)
-
-done
-
-
-
-subsection" Filter of mksch(tr,schA,schB) to B is schB -- take lemma proof"
-
-lemma FilterBmksch_is_schB: "!! A B. [| compatible A B; compatible B A;
-  is_asig(asig_of A); is_asig(asig_of B) |] ==>
-  Forall (%x. x:ext (A\<parallel>B)) tr &
-  Forall (%x. x:act A) schA & Forall (%x. x:act B) schB &
-  Filter (%a. a:ext A)$schA = Filter (%a. a:act A)$tr &
-  Filter (%a. a:ext B)$schB = Filter (%a. a:act B)$tr &
-  LastActExtsch A schA & LastActExtsch B schB
-  --> Filter (%a. a:act B)$(mksch A B$tr$schA$schB) = schB"
-apply (intro strip)
-apply (rule seq.take_lemma)
-apply (rule mp)
-prefer 2 apply assumption
-back back back back
-apply (rule_tac x = "schA" in spec)
-apply (rule_tac x = "schB" in spec)
-apply (rule_tac x = "tr" in spec)
-apply (tactic "thin_tac' @{context} 5 1")
-apply (rule nat_less_induct)
-apply (rule allI)+
-apply (rename_tac tr schB schA)
-apply (intro strip)
-apply (erule conjE)+
-
-apply (case_tac "Forall (%x. x:act A & x~:act B) tr")
-
-apply (rule seq_take_lemma [THEN iffD2, THEN spec])
-apply (tactic "thin_tac' @{context} 5 1")
-
-apply (case_tac "Finite tr")
-
-(* both sides of this equation are nil *)
-apply (subgoal_tac "schB=nil")
-apply (simp (no_asm_simp))
-(* first side: mksch = nil *)
-apply (auto intro!: ForallQFilterPnil ForallAnBmksch FiniteL_mksch)[1]
-(* second side: schA = nil *)
-apply (erule_tac A = "B" in LastActExtimplnil)
-apply (simp (no_asm_simp))
-apply (erule_tac Q = "%x. x:act A & x~:act B" in ForallQFilterPnil)
-apply assumption
-apply fast
-
-(* case ~ Finite tr *)
-
-(* both sides of this equation are UU *)
-apply (subgoal_tac "schB=UU")
-apply (simp (no_asm_simp))
-(* first side: mksch = UU *)
-apply (force intro!: ForallQFilterPUU finiteR_mksch' ForallAnBmksch)
-(* schA = UU *)
-apply (erule_tac A = "B" in LastActExtimplUU)
-apply (simp (no_asm_simp))
-apply (erule_tac Q = "%x. x:act A & x~:act B" in ForallQFilterPUU)
-apply assumption
-apply fast
-
-(* case" ~ Forall (%x.x:act B & x~:act A) s" *)
-
-apply (drule divide_Seq3)
-
-apply (erule exE)+
-apply (erule conjE)+
-apply hypsubst
-
-(* bring in lemma reduceB_mksch *)
-apply (frule_tac y = "schB" and x = "schA" and A = "A" and B = "B" in reduceB_mksch)
-apply assumption+
-apply (erule exE)+
-apply (erule conjE)+
-
-(* use reduceB_mksch to rewrite conclusion *)
-apply hypsubst
-apply simp
-
-(* eliminate the A-only prefix *)
-
-apply (subgoal_tac "(Filter (%a. a :act B) $x1) = nil")
-apply (erule_tac [2] ForallQFilterPnil)
-prefer 2 apply (assumption)
-prefer 2 apply (fast)
-
-(* Now real recursive step follows (in x) *)
-
-apply simp
-apply (case_tac "x:act B")
-apply (case_tac "x~:act A")
-apply (rotate_tac -2)
-apply simp
-
-apply (subgoal_tac "Filter (%a. a:act B & a:ext A) $x1=nil")
-apply (rotate_tac -1)
-apply simp
-(* eliminate introduced subgoal 2 *)
-apply (erule_tac [2] ForallQFilterPnil)
-prefer 2 apply (assumption)
-prefer 2 apply (fast)
-
-(* bring in divide Seq for s *)
-apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
-apply (erule conjE)+
-
-(* subst divide_Seq in conclusion, but only at the rightmost occurrence *)
-apply (rule_tac t = "schB" in ssubst)
-back
-back
-back
-apply assumption
-
-(* reduce trace_takes from n to strictly smaller k *)
-apply (rule take_reduction)
-
-(* f B (tw iB) = tw ~eB *)
-apply (simp add: int_is_act not_ext_is_int_or_not_act)
-apply (rule refl)
-apply (simp add: int_is_act not_ext_is_int_or_not_act)
-apply (rotate_tac -11)
-
-(* now conclusion fulfills induction hypothesis, but assumptions are not ready *)
-
-(* assumption schA *)
-apply (simp add: ext_and_act)
-(* assumption schB *)
-apply (drule_tac x = "schB" and g = "Filter (%a. a:act B) $rs" in subst_lemma2)
-apply assumption
-apply (simp add: int_is_not_ext)
-(* assumptions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping  *)
-apply (drule_tac sch = "schB" and P = "%a. a:int B" in LastActExtsmall1)
-apply (frule_tac ?sch1.0 = "x1" in LastActExtsmall2)
-apply assumption
-
-(* assumption Forall schB *)
-apply (drule_tac s = "schB" and P = "Forall (%x. x:act B) " in subst)
-apply assumption
-apply (simp add: int_is_act)
-
-(* case x:actions(asig_of A) & x: actions(asig_of B) *)
-
-apply (rotate_tac -2)
-apply simp
-
-apply (subgoal_tac "Filter (%a. a:act B & a:ext A) $x1=nil")
-apply (rotate_tac -1)
-apply simp
-(* eliminate introduced subgoal 2 *)
-apply (erule_tac [2] ForallQFilterPnil)
-prefer 2 apply (assumption)
-prefer 2 apply (fast)
-
-(* bring in divide Seq for s *)
-apply (frule sym [THEN eq_imp_below, THEN divide_Seq])
-apply (erule conjE)+
-
-(* subst divide_Seq in conclusion, but only at the rightmost occurrence *)
-apply (rule_tac t = "schB" in ssubst)
-back
-back
-back
-apply assumption
-
-(* f B (tw iB) = tw ~eB *)
-apply (simp add: int_is_act not_ext_is_int_or_not_act)
-
-(* rewrite assumption forall and schB *)
-apply (rotate_tac 13)
-apply (simp add: ext_and_act)
-
-(* divide_Seq for schB2 *)
-apply (frule_tac y = "x2" in sym [THEN eq_imp_below, THEN divide_Seq])
-apply (erule conjE)+
-(* assumption schA *)
-apply (drule_tac x = "schB" and g = "Filter (%a. a:act B) $rs" in subst_lemma2)
-apply assumption
-apply (simp add: int_is_not_ext)
-
-(* f B (tw iA schA2) = nil *)
-apply (simp add: int_is_not_ext not_ext_is_int_or_not_act intA_is_not_actB)
-
-
-(* reduce trace_takes from n to strictly smaller k *)
-apply (rule take_reduction)
-apply (rule refl)
-apply (rule refl)
-
-(* now conclusion fulfills induction hypothesis, but assumptions are not all ready *)
-
-(* assumption schA *)
-apply (drule_tac x = "x2" and g = "Filter (%a. a:act A) $rs" in subst_lemma2)
-apply assumption
-apply (simp add: intA_is_not_actB int_is_not_ext)
-
-(* conclusions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping  *)
-apply (drule_tac sch = "schB" and P = "%a. a:int B" in LastActExtsmall1)
-apply (frule_tac ?sch1.0 = "x1" in LastActExtsmall2)
-apply assumption
-apply (drule_tac sch = "x2" and P = "%a. a:int A" in LastActExtsmall1)
-
-(* assumption Forall schA, and Forall schA are performed by ForallTL,ForallDropwhile *)
-apply (simp add: ForallTL ForallDropwhile)
-
-(* case x~:B & x:A  *)
-(* cannot occur, as just this case has been scheduled out before as the B-only prefix *)
-apply (case_tac "x:act A")
-apply fast
-
-(* case x~:B & x~:A  *)
-(* cannot occur because of assumption: Forall (a:ext A | a:ext B) *)
-apply (rotate_tac -9)
-(* reduce forall assumption from tr to (x\<leadsto>rs) *)
-apply (simp add: externals_of_par)
-apply (fast intro!: ext_is_act)
-
-done
-
-
-subsection "COMPOSITIONALITY on TRACE Level -- Main Theorem"
-
-lemma compositionality_tr:
-"!! A B. [| is_trans_of A; is_trans_of B; compatible A B; compatible B A;
-            is_asig(asig_of A); is_asig(asig_of B)|]
-        ==>  (tr: traces(A\<parallel>B)) =
-             (Filter (%a. a:act A)$tr : traces A &
-              Filter (%a. a:act B)$tr : traces B &
-              Forall (%x. x:ext(A\<parallel>B)) tr)"
-
-apply (simp (no_asm) add: traces_def has_trace_def)
-apply auto
-
-(* ==> *)
-(* There is a schedule of A *)
-apply (rule_tac x = "Filter (%a. a:act A) $sch" in bexI)
-prefer 2
-apply (simp add: compositionality_sch)
-apply (simp add: compatibility_consequence1 externals_of_par ext1_ext2_is_not_act1)
-(* There is a schedule of B *)
-apply (rule_tac x = "Filter (%a. a:act B) $sch" in bexI)
-prefer 2
-apply (simp add: compositionality_sch)
-apply (simp add: compatibility_consequence2 externals_of_par ext1_ext2_is_not_act2)
-(* Traces of A\<parallel>B have only external actions from A or B *)
-apply (rule ForallPFilterP)
-
-(* <== *)
-
-(* replace schA and schB by Cut(schA) and Cut(schB) *)
-apply (drule exists_LastActExtsch)
-apply assumption
-apply (drule exists_LastActExtsch)
-apply assumption
-apply (erule exE)+
-apply (erule conjE)+
-(* Schedules of A(B) have only actions of A(B) *)
-apply (drule scheds_in_sig)
-apply assumption
-apply (drule scheds_in_sig)
-apply assumption
-
-apply (rename_tac h1 h2 schA schB)
-(* mksch is exactly the construction of trA\<parallel>B out of schA, schB, and the oracle tr,
-   we need here *)
-apply (rule_tac x = "mksch A B$tr$schA$schB" in bexI)
-
-(* External actions of mksch are just the oracle *)
-apply (simp add: FilterA_mksch_is_tr)
-
-(* mksch is a schedule -- use compositionality on sch-level *)
-apply (simp add: compositionality_sch)
-apply (simp add: FilterAmksch_is_schA FilterBmksch_is_schB)
-apply (erule ForallAorB_mksch)
-apply (erule ForallPForallQ)
-apply (erule ext_is_act)
-done
-
-
-
-subsection \<open>COMPOSITIONALITY on TRACE Level -- for Modules\<close>
-
-lemma compositionality_tr_modules:
-
-"!! A B. [| is_trans_of A; is_trans_of B; compatible A B; compatible B A;
-            is_asig(asig_of A); is_asig(asig_of B)|]
- ==> Traces (A\<parallel>B) = par_traces (Traces A) (Traces B)"
-
-apply (unfold Traces_def par_traces_def)
-apply (simp add: asig_of_par)
-apply (rule set_eqI)
-apply (simp add: compositionality_tr externals_of_par)
-done
-
-
-declaration \<open>fn _ => Simplifier.map_ss (Simplifier.set_mksym Simplifier.default_mk_sym)\<close>
-
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/Compositionality.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,75 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/Compositionality.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Compositionality of I/O automata\<close>
-theory Compositionality
-imports CompoTraces
-begin
-
-lemma compatibility_consequence3: "[|eA --> A ; eB & ~eA --> ~A|] ==> (eA | eB) --> A=eA"
-apply auto
-done
-
-
-lemma Filter_actAisFilter_extA: 
-"!! A B. [| compatible A B; Forall (%a. a: ext A | a: ext B) tr |] ==>  
-            Filter (%a. a: act A)$tr= Filter (%a. a: ext A)$tr"
-apply (rule ForallPFilterQR)
-(* i.e.: [| (! x. P x --> (Q x = R x)) ; Forall P tr |] ==> Filter Q$tr = Filter R$tr *)
-prefer 2 apply (assumption)
-apply (rule compatibility_consequence3)
-apply (simp_all add: ext_is_act ext1_ext2_is_not_act1)
-done
-
-
-(* the next two theorems are only necessary, as there is no theorem ext (A\<parallel>B) = ext (B\<parallel>A) *)
-
-lemma compatibility_consequence4: "[|eA --> A ; eB & ~eA --> ~A|] ==> (eB | eA) --> A=eA"
-apply auto
-done
-
-lemma Filter_actAisFilter_extA2: "[| compatible A B; Forall (%a. a: ext B | a: ext A) tr |] ==>  
-            Filter (%a. a: act A)$tr= Filter (%a. a: ext A)$tr"
-apply (rule ForallPFilterQR)
-prefer 2 apply (assumption)
-apply (rule compatibility_consequence4)
-apply (simp_all add: ext_is_act ext1_ext2_is_not_act1)
-done
-
-
-subsection " Main Compositionality Theorem "
-
-lemma compositionality: "[| is_trans_of A1; is_trans_of A2; is_trans_of B1; is_trans_of B2; 
-             is_asig_of A1; is_asig_of A2;  
-             is_asig_of B1; is_asig_of B2;  
-             compatible A1 B1; compatible A2 B2;  
-             A1 =<| A2;  
-             B1 =<| B2 |]  
-         ==> (A1 \<parallel> B1) =<| (A2 \<parallel> B2)"
-apply (simp add: is_asig_of_def)
-apply (frule_tac A1 = "A1" in compat_commute [THEN iffD1])
-apply (frule_tac A1 = "A2" in compat_commute [THEN iffD1])
-apply (simp add: ioa_implements_def inputs_of_par outputs_of_par externals_of_par)
-apply auto
-apply (simp add: compositionality_tr)
-apply (subgoal_tac "ext A1 = ext A2 & ext B1 = ext B2")
-prefer 2
-apply (simp add: externals_def)
-apply (erule conjE)+
-(* rewrite with proven subgoal *)
-apply (simp add: externals_of_par)
-apply auto
-
-(* 2 goals, the 3rd has been solved automatically *)
-(* 1: Filter A2 x : traces A2 *)
-apply (drule_tac A = "traces A1" in subsetD)
-apply assumption
-apply (simp add: Filter_actAisFilter_extA)
-(* 2: Filter B2 x : traces B2 *)
-apply (drule_tac A = "traces B1" in subsetD)
-apply assumption
-apply (simp add: Filter_actAisFilter_extA2)
-done
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/Deadlock.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,92 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/Deadlock.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Deadlock freedom of I/O Automata\<close>
-
-theory Deadlock
-imports RefCorrectness CompoScheds
-begin
-
-text \<open>input actions may always be added to a schedule\<close>
-
-lemma scheds_input_enabled:
-  "[| Filter (%x. x:act A)$sch : schedules A; a:inp A; input_enabled A; Finite sch|]  
-          ==> Filter (%x. x:act A)$sch @@ a\<leadsto>nil : schedules A"
-apply (simp add: schedules_def has_schedule_def)
-apply auto
-apply (frule inp_is_act)
-apply (simp add: executions_def)
-apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-apply (rename_tac s ex)
-apply (subgoal_tac "Finite ex")
-prefer 2
-apply (simp add: filter_act_def)
-defer
-apply (rule_tac [2] Map2Finite [THEN iffD1])
-apply (rule_tac [2] t = "Map fst$ex" in subst)
-prefer 2 apply (assumption)
-apply (erule_tac [2] FiniteFilter)
-(* subgoal 1 *)
-apply (frule exists_laststate)
-apply (erule allE)
-apply (erule exE)
-(* using input-enabledness *)
-apply (simp add: input_enabled_def)
-apply (erule conjE)+
-apply (erule_tac x = "a" in allE)
-apply simp
-apply (erule_tac x = "u" in allE)
-apply (erule exE)
-(* instantiate execution *)
-apply (rule_tac x = " (s,ex @@ (a,s2) \<leadsto>nil) " in exI)
-apply (simp add: filter_act_def MapConc)
-apply (erule_tac t = "u" in lemma_2_1)
-apply simp
-apply (rule sym)
-apply assumption
-done
-
-text \<open>
-               Deadlock freedom: component B cannot block an out or int action
-                                 of component A in every schedule.
-    Needs compositionality on schedule level, input-enabledness, compatibility
-                    and distributivity of is_exec_frag over @@
-\<close>
-
-declare split_if [split del]
-lemma IOA_deadlock_free: "[| a : local A; Finite sch; sch : schedules (A\<parallel>B);  
-             Filter (%x. x:act A)$(sch @@ a\<leadsto>nil) : schedules A; compatible A B; input_enabled B |]  
-           ==> (sch @@ a\<leadsto>nil) : schedules (A\<parallel>B)"
-apply (simp add: compositionality_sch locals_def)
-apply (rule conjI)
-(* a : act (A\<parallel>B) *)
-prefer 2
-apply (simp add: actions_of_par)
-apply (blast dest: int_is_act out_is_act)
-
-(* Filter B (sch@@[a]) : schedules B *)
-
-apply (case_tac "a:int A")
-apply (drule intA_is_not_actB)
-apply (assumption) (* --> a~:act B *)
-apply simp
-
-(* case a~:int A , i.e. a:out A *)
-apply (case_tac "a~:act B")
-apply simp
-(* case a:act B *)
-apply simp
-apply (subgoal_tac "a:out A")
-prefer 2 apply (blast)
-apply (drule outAactB_is_inpB)
-apply assumption
-apply assumption
-apply (rule scheds_input_enabled)
-apply simp
-apply assumption+
-done
-
-declare split_if [split]
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/IOA.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,11 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/IOA.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>The theory of I/O automata in HOLCF\<close>
-
-theory IOA
-imports SimCorrectness Compositionality Deadlock
-begin
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/LiveIOA.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,82 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/LiveIOA.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Live I/O automata -- specified by temproal formulas\<close>
-
-theory LiveIOA
-imports TLS
-begin
-
-default_sort type
-
-type_synonym
-  ('a, 's) live_ioa = "('a,'s)ioa * ('a,'s)ioa_temp"
-
-definition
-  validLIOA :: "('a,'s)live_ioa => ('a,'s)ioa_temp  => bool" where
-  "validLIOA AL P = validIOA (fst AL) ((snd AL) \<^bold>\<longrightarrow> P)"
-
-definition
-  WF :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp" where
-  "WF A acts = (\<diamond>\<box>\<langle>%(s,a,t). Enabled A acts s\<rangle> \<^bold>\<longrightarrow> \<box>\<diamond>\<langle>xt2 (plift (%a. a : acts))\<rangle>)"
-definition
-  SF :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp" where
-  "SF A acts = (\<box>\<diamond>\<langle>%(s,a,t). Enabled A acts s\<rangle> \<^bold>\<longrightarrow> \<box>\<diamond>\<langle>xt2 (plift (%a. a : acts))\<rangle>)"
-
-definition
-  liveexecutions :: "('a,'s)live_ioa => ('a,'s)execution set" where
-  "liveexecutions AP = {exec. exec : executions (fst AP) & (exec \<TTurnstile> (snd AP))}"
-definition
-  livetraces :: "('a,'s)live_ioa => 'a trace set" where
-  "livetraces AP = {mk_trace (fst AP)$(snd ex) | ex. ex:liveexecutions AP}"
-definition
-  live_implements :: "('a,'s1)live_ioa => ('a,'s2)live_ioa => bool" where
-  "live_implements CL AM = ((inp (fst CL) = inp (fst AM)) &
-                            (out (fst CL) = out (fst AM)) &
-                            livetraces CL <= livetraces AM)"
-definition
-  is_live_ref_map :: "('s1 => 's2) => ('a,'s1)live_ioa => ('a,'s2)live_ioa => bool" where
-  "is_live_ref_map f CL AM =
-           (is_ref_map f (fst CL ) (fst AM) &
-            (! exec : executions (fst CL). (exec \<TTurnstile> (snd CL)) -->
-                                           ((corresp_ex (fst AM) f exec) \<TTurnstile> (snd AM))))"
-
-
-lemma live_implements_trans:
-"!!LC. [| live_implements (A,LA) (B,LB); live_implements (B,LB) (C,LC) |]
-      ==> live_implements (A,LA) (C,LC)"
-apply (unfold live_implements_def)
-apply auto
-done
-
-
-subsection "Correctness of live refmap"
-
-lemma live_implements: "[| inp(C)=inp(A); out(C)=out(A);
-                   is_live_ref_map f (C,M) (A,L) |]
-                ==> live_implements (C,M) (A,L)"
-apply (simp add: is_live_ref_map_def live_implements_def livetraces_def liveexecutions_def)
-apply auto
-apply (rule_tac x = "corresp_ex A f ex" in exI)
-apply auto
-  (* Traces coincide, Lemma 1 *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (erule lemma_1 [THEN spec, THEN mp])
-  apply (simp (no_asm) add: externals_def)
-  apply (auto)[1]
-  apply (simp add: executions_def reachable.reachable_0)
-
-  (* corresp_ex is execution, Lemma 2 *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (simp add: executions_def)
-  (* start state *)
-  apply (rule conjI)
-  apply (simp add: is_ref_map_def corresp_ex_def)
-  (* is-execution-fragment *)
-  apply (erule lemma_2 [THEN spec, THEN mp])
-  apply (simp add: reachable.reachable_0)
-
-done
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/Pred.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,33 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/Pred.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Logical Connectives lifted to predicates\<close>
-
-theory Pred
-imports Main
-begin
-
-default_sort type
-
-type_synonym 'a predicate = "'a \<Rightarrow> bool"
-
-definition satisfies :: "'a \<Rightarrow> 'a predicate \<Rightarrow> bool"  ("_ \<Turnstile> _" [100,9] 8)
-  where "(s \<Turnstile> P) \<longleftrightarrow> P s"
-
-definition valid :: "'a predicate \<Rightarrow> bool"  (*  ("|-") *)
-  where "valid P \<longleftrightarrow> (\<forall>s. (s \<Turnstile> P))"
-
-definition NOT :: "'a predicate \<Rightarrow> 'a predicate"  ("\<^bold>\<not> _" [40] 40)
-  where "NOT P s \<longleftrightarrow> ~ (P s)"
-
-definition AND :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<and>" 35)
-  where "(P \<^bold>\<and> Q) s \<longleftrightarrow> P s \<and> Q s"
-
-definition OR :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<or>" 30)
-  where "(P \<^bold>\<or> Q) s \<longleftrightarrow> P s \<or> Q s"
-
-definition IMPLIES :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<longrightarrow>" 25)
-  where "(P \<^bold>\<longrightarrow> Q) s \<longleftrightarrow> P s \<longrightarrow> Q s"
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/RefCorrectness.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,370 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/RefCorrectness.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Correctness of Refinement Mappings in HOLCF/IOA\<close>
-
-theory RefCorrectness
-imports RefMappings
-begin
-
-definition
-  corresp_exC :: "('a,'s2)ioa => ('s1 => 's2) => ('a,'s1)pairs
-                   -> ('s1 => ('a,'s2)pairs)" where
-  "corresp_exC A f = (fix$(LAM h ex. (%s. case ex of
-      nil =>  nil
-    | x##xs => (flift1 (%pr. (@cex. move A cex (f s) (fst pr) (f (snd pr)))
-                              @@ ((h$xs) (snd pr)))
-                        $x) )))"
-definition
-  corresp_ex :: "('a,'s2)ioa => ('s1 => 's2) =>
-                  ('a,'s1)execution => ('a,'s2)execution" where
-  "corresp_ex A f ex = (f (fst ex),(corresp_exC A f$(snd ex)) (fst ex))"
-
-definition
-  is_fair_ref_map :: "('s1 => 's2) => ('a,'s1)ioa => ('a,'s2)ioa => bool" where
-  "is_fair_ref_map f C A =
-      (is_ref_map f C A &
-       (! ex : executions(C). fair_ex C ex --> fair_ex A (corresp_ex A f ex)))"
-
-(* Axioms for fair trace inclusion proof support, not for the correctness proof
-   of refinement mappings!
-   Note: Everything is superseded by LiveIOA.thy! *)
-
-axiomatization where
-corresp_laststate:
-  "Finite ex ==> laststate (corresp_ex A f (s,ex)) = f (laststate (s,ex))"
-
-axiomatization where
-corresp_Finite:
-  "Finite (snd (corresp_ex A f (s,ex))) = Finite ex"
-
-axiomatization where
-FromAtoC:
-  "fin_often (%x. P (snd x)) (snd (corresp_ex A f (s,ex))) ==> fin_often (%y. P (f (snd y))) ex"
-
-axiomatization where
-FromCtoA:
-  "inf_often (%y. P (fst y)) ex ==> inf_often (%x. P (fst x)) (snd (corresp_ex A f (s,ex)))"
-
-
-(* Proof by case on inf W in ex: If so, ok. If not, only fin W in ex, ie there is
-   an index i from which on no W in ex. But W inf enabled, ie at least once after i
-   W is enabled. As W does not occur after i and W is enabling_persistent, W keeps
-   enabled until infinity, ie. indefinitely *)
-axiomatization where
-persistent:
-  "[|inf_often (%x. Enabled A W (snd x)) ex; en_persistent A W|]
-   ==> inf_often (%x. fst x :W) ex | fin_often (%x. ~Enabled A W (snd x)) ex"
-
-axiomatization where
-infpostcond:
-  "[| is_exec_frag A (s,ex); inf_often (%x. fst x:W) ex|]
-    ==> inf_often (% x. set_was_enabled A W (snd x)) ex"
-
-
-subsection "corresp_ex"
-
-lemma corresp_exC_unfold: "corresp_exC A f  = (LAM ex. (%s. case ex of
-       nil =>  nil
-     | x##xs => (flift1 (%pr. (@cex. move A cex (f s) (fst pr) (f (snd pr)))
-                               @@ ((corresp_exC A f $xs) (snd pr)))
-                         $x) ))"
-apply (rule trans)
-apply (rule fix_eq2)
-apply (simp only: corresp_exC_def)
-apply (rule beta_cfun)
-apply (simp add: flift1_def)
-done
-
-lemma corresp_exC_UU: "(corresp_exC A f$UU) s=UU"
-apply (subst corresp_exC_unfold)
-apply simp
-done
-
-lemma corresp_exC_nil: "(corresp_exC A f$nil) s = nil"
-apply (subst corresp_exC_unfold)
-apply simp
-done
-
-lemma corresp_exC_cons: "(corresp_exC A f$(at\<leadsto>xs)) s =
-           (@cex. move A cex (f s) (fst at) (f (snd at)))
-           @@ ((corresp_exC A f$xs) (snd at))"
-apply (rule trans)
-apply (subst corresp_exC_unfold)
-apply (simp add: Consq_def flift1_def)
-apply simp
-done
-
-
-declare corresp_exC_UU [simp] corresp_exC_nil [simp] corresp_exC_cons [simp]
-
-
-
-subsection "properties of move"
-
-lemma move_is_move:
-   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
-      move A (@x. move A x (f s) a (f t)) (f s) a (f t)"
-apply (unfold is_ref_map_def)
-apply (subgoal_tac "? ex. move A ex (f s) a (f t) ")
-prefer 2
-apply simp
-apply (erule exE)
-apply (rule someI)
-apply assumption
-done
-
-lemma move_subprop1:
-   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
-     is_exec_frag A (f s,@x. move A x (f s) a (f t))"
-apply (cut_tac move_is_move)
-defer
-apply assumption+
-apply (simp add: move_def)
-done
-
-lemma move_subprop2:
-   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
-     Finite ((@x. move A x (f s) a (f t)))"
-apply (cut_tac move_is_move)
-defer
-apply assumption+
-apply (simp add: move_def)
-done
-
-lemma move_subprop3:
-   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
-     laststate (f s,@x. move A x (f s) a (f t)) = (f t)"
-apply (cut_tac move_is_move)
-defer
-apply assumption+
-apply (simp add: move_def)
-done
-
-lemma move_subprop4:
-   "[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>
-      mk_trace A$((@x. move A x (f s) a (f t))) =
-        (if a:ext A then a\<leadsto>nil else nil)"
-apply (cut_tac move_is_move)
-defer
-apply assumption+
-apply (simp add: move_def)
-done
-
-
-(* ------------------------------------------------------------------ *)
-(*                   The following lemmata contribute to              *)
-(*                 TRACE INCLUSION Part 1: Traces coincide            *)
-(* ------------------------------------------------------------------ *)
-
-section "Lemmata for <=="
-
-(* --------------------------------------------------- *)
-(*   Lemma 1.1: Distribution of mk_trace and @@        *)
-(* --------------------------------------------------- *)
-
-lemma mk_traceConc: "mk_trace C$(ex1 @@ ex2)= (mk_trace C$ex1) @@ (mk_trace C$ex2)"
-apply (simp add: mk_trace_def filter_act_def MapConc)
-done
-
-
-
-(* ------------------------------------------------------
-                 Lemma 1 :Traces coincide
-   ------------------------------------------------------- *)
-declare split_if [split del]
-
-lemma lemma_1:
-  "[|is_ref_map f C A; ext C = ext A|] ==>
-         !s. reachable C s & is_exec_frag C (s,xs) -->
-             mk_trace C$xs = mk_trace A$(snd (corresp_ex A f (s,xs)))"
-apply (unfold corresp_ex_def)
-apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1\<close>)
-(* cons case *)
-apply (auto simp add: mk_traceConc)
-apply (frule reachable.reachable_n)
-apply assumption
-apply (auto simp add: move_subprop4 split add: split_if) 
-done
-
-declare split_if [split]
-
-(* ------------------------------------------------------------------ *)
-(*                   The following lemmata contribute to              *)
-(*              TRACE INCLUSION Part 2: corresp_ex is execution       *)
-(* ------------------------------------------------------------------ *)
-
-section "Lemmata for ==>"
-
-(* -------------------------------------------------- *)
-(*                   Lemma 2.1                        *)
-(* -------------------------------------------------- *)
-
-lemma lemma_2_1 [rule_format (no_asm)]:
-"Finite xs -->
- (!s .is_exec_frag A (s,xs) & is_exec_frag A (t,ys) &
-      t = laststate (s,xs)
-  --> is_exec_frag A (s,xs @@ ys))"
-
-apply (rule impI)
-apply (tactic \<open>Seq_Finite_induct_tac @{context} 1\<close>)
-(* main case *)
-apply (auto simp add: split_paired_all)
-done
-
-
-(* ----------------------------------------------------------- *)
-(*               Lemma 2 : corresp_ex is execution             *)
-(* ----------------------------------------------------------- *)
-
-
-
-lemma lemma_2:
- "[| is_ref_map f C A |] ==>
-  !s. reachable C s & is_exec_frag C (s,xs)
-  --> is_exec_frag A (corresp_ex A f (s,xs))"
-
-apply (unfold corresp_ex_def)
-
-apply simp
-apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1\<close>)
-(* main case *)
-apply auto
-apply (rule_tac t = "f x2" in lemma_2_1)
-
-(* Finite *)
-apply (erule move_subprop2)
-apply assumption+
-apply (rule conjI)
-
-(* is_exec_frag *)
-apply (erule move_subprop1)
-apply assumption+
-apply (rule conjI)
-
-(* Induction hypothesis  *)
-(* reachable_n looping, therefore apply it manually *)
-apply (erule_tac x = "x2" in allE)
-apply simp
-apply (frule reachable.reachable_n)
-apply assumption
-apply simp
-(* laststate *)
-apply (erule move_subprop3 [symmetric])
-apply assumption+
-done
-
-
-subsection "Main Theorem: TRACE - INCLUSION"
-
-lemma trace_inclusion:
-  "[| ext C = ext A; is_ref_map f C A |]
-           ==> traces C <= traces A"
-
-  apply (unfold traces_def)
-
-  apply (simp (no_asm) add: has_trace_def2)
-  apply auto
-
-  (* give execution of abstract automata *)
-  apply (rule_tac x = "corresp_ex A f ex" in bexI)
-
-  (* Traces coincide, Lemma 1 *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (erule lemma_1 [THEN spec, THEN mp])
-  apply assumption+
-  apply (simp add: executions_def reachable.reachable_0)
-
-  (* corresp_ex is execution, Lemma 2 *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (simp add: executions_def)
-  (* start state *)
-  apply (rule conjI)
-  apply (simp add: is_ref_map_def corresp_ex_def)
-  (* is-execution-fragment *)
-  apply (erule lemma_2 [THEN spec, THEN mp])
-  apply (simp add: reachable.reachable_0)
-  done
-
-
-subsection "Corollary:  FAIR TRACE - INCLUSION"
-
-lemma fininf: "(~inf_often P s) = fin_often P s"
-apply (unfold fin_often_def)
-apply auto
-done
-
-
-lemma WF_alt: "is_wfair A W (s,ex) =
-  (fin_often (%x. ~Enabled A W (snd x)) ex --> inf_often (%x. fst x :W) ex)"
-apply (simp add: is_wfair_def fin_often_def)
-apply auto
-done
-
-lemma WF_persistent: "[|is_wfair A W (s,ex); inf_often (%x. Enabled A W (snd x)) ex;
-          en_persistent A W|]
-    ==> inf_often (%x. fst x :W) ex"
-apply (drule persistent)
-apply assumption
-apply (simp add: WF_alt)
-apply auto
-done
-
-
-lemma fair_trace_inclusion: "!! C A.
-          [| is_ref_map f C A; ext C = ext A;
-          !! ex. [| ex:executions C; fair_ex C ex|] ==> fair_ex A (corresp_ex A f ex) |]
-          ==> fairtraces C <= fairtraces A"
-apply (simp (no_asm) add: fairtraces_def fairexecutions_def)
-apply auto
-apply (rule_tac x = "corresp_ex A f ex" in exI)
-apply auto
-
-  (* Traces coincide, Lemma 1 *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (erule lemma_1 [THEN spec, THEN mp])
-  apply assumption+
-  apply (simp add: executions_def reachable.reachable_0)
-
-  (* corresp_ex is execution, Lemma 2 *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (simp add: executions_def)
-  (* start state *)
-  apply (rule conjI)
-  apply (simp add: is_ref_map_def corresp_ex_def)
-  (* is-execution-fragment *)
-  apply (erule lemma_2 [THEN spec, THEN mp])
-  apply (simp add: reachable.reachable_0)
-
-done
-
-lemma fair_trace_inclusion2: "!! C A.
-          [| inp(C) = inp(A); out(C)=out(A);
-             is_fair_ref_map f C A |]
-          ==> fair_implements C A"
-apply (simp add: is_fair_ref_map_def fair_implements_def fairtraces_def fairexecutions_def)
-apply auto
-apply (rule_tac x = "corresp_ex A f ex" in exI)
-apply auto
-
-  (* Traces coincide, Lemma 1 *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (erule lemma_1 [THEN spec, THEN mp])
-  apply (simp (no_asm) add: externals_def)
-  apply (auto)[1]
-  apply (simp add: executions_def reachable.reachable_0)
-
-  (* corresp_ex is execution, Lemma 2 *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (simp add: executions_def)
-  (* start state *)
-  apply (rule conjI)
-  apply (simp add: is_ref_map_def corresp_ex_def)
-  (* is-execution-fragment *)
-  apply (erule lemma_2 [THEN spec, THEN mp])
-  apply (simp add: reachable.reachable_0)
-
-done
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/RefMappings.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,129 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/RefMappings.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Refinement Mappings in HOLCF/IOA\<close>
-
-theory RefMappings
-imports Traces
-begin
-
-default_sort type
-
-definition
-  move :: "[('a,'s)ioa,('a,'s)pairs,'s,'a,'s] => bool" where
-  "move ioa ex s a t =
-    (is_exec_frag ioa (s,ex) &  Finite ex &
-     laststate (s,ex)=t  &
-     mk_trace ioa$ex = (if a:ext(ioa) then a\<leadsto>nil else nil))"
-
-definition
-  is_ref_map :: "[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool" where
-  "is_ref_map f C A =
-   ((!s:starts_of(C). f(s):starts_of(A)) &
-   (!s t a. reachable C s &
-            s \<midarrow>a\<midarrow>C\<rightarrow> t
-            --> (? ex. move A ex (f s) a (f t))))"
-
-definition
-  is_weak_ref_map :: "[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool" where
-  "is_weak_ref_map f C A =
-   ((!s:starts_of(C). f(s):starts_of(A)) &
-   (!s t a. reachable C s &
-            s \<midarrow>a\<midarrow>C\<rightarrow> t
-            --> (if a:ext(C)
-                 then (f s) \<midarrow>a\<midarrow>A\<rightarrow> (f t)
-                 else (f s)=(f t))))"
-
-
-subsection "transitions and moves"
-
-
-lemma transition_is_ex: "s \<midarrow>a\<midarrow>A\<rightarrow> t ==> ? ex. move A ex s a t"
-apply (rule_tac x = " (a,t) \<leadsto>nil" in exI)
-apply (simp add: move_def)
-done
-
-
-lemma nothing_is_ex: "(~a:ext A) & s=t ==> ? ex. move A ex s a t"
-apply (rule_tac x = "nil" in exI)
-apply (simp add: move_def)
-done
-
-
-lemma ei_transitions_are_ex: "(s \<midarrow>a\<midarrow>A\<rightarrow> s') & (s' \<midarrow>a'\<midarrow>A\<rightarrow> s'') & (~a':ext A)  
-         ==> ? ex. move A ex s a s''"
-apply (rule_tac x = " (a,s') \<leadsto> (a',s'') \<leadsto>nil" in exI)
-apply (simp add: move_def)
-done
-
-
-lemma eii_transitions_are_ex: "(s1 \<midarrow>a1\<midarrow>A\<rightarrow> s2) & (s2 \<midarrow>a2\<midarrow>A\<rightarrow> s3) & (s3 \<midarrow>a3\<midarrow>A\<rightarrow> s4) & 
-      (~a2:ext A) & (~a3:ext A) ==>  
-      ? ex. move A ex s1 a1 s4"
-apply (rule_tac x = " (a1,s2) \<leadsto> (a2,s3) \<leadsto> (a3,s4) \<leadsto>nil" in exI)
-apply (simp add: move_def)
-done
-
-
-subsection "weak_ref_map and ref_map"
-
-lemma weak_ref_map2ref_map:
-  "[| ext C = ext A;  
-     is_weak_ref_map f C A |] ==> is_ref_map f C A"
-apply (unfold is_weak_ref_map_def is_ref_map_def)
-apply auto
-apply (case_tac "a:ext A")
-apply (auto intro: transition_is_ex nothing_is_ex)
-done
-
-
-lemma imp_conj_lemma: "(P ==> Q-->R) ==> P&Q --> R"
-  by blast
-
-declare split_if [split del]
-declare if_weak_cong [cong del]
-
-lemma rename_through_pmap: "[| is_weak_ref_map f C A |]  
-      ==> (is_weak_ref_map f (rename C g) (rename A g))"
-apply (simp add: is_weak_ref_map_def)
-apply (rule conjI)
-(* 1: start states *)
-apply (simp add: rename_def rename_set_def starts_of_def)
-(* 2: reachable transitions *)
-apply (rule allI)+
-apply (rule imp_conj_lemma)
-apply (simp (no_asm) add: rename_def rename_set_def)
-apply (simp add: externals_def asig_inputs_def asig_outputs_def asig_of_def trans_of_def)
-apply safe
-apply (simplesubst split_if)
- apply (rule conjI)
- apply (rule impI)
- apply (erule disjE)
- apply (erule exE)
-apply (erule conjE)
-(* x is input *)
- apply (drule sym)
- apply (drule sym)
-apply simp
-apply hypsubst+
-apply (frule reachable_rename)
-apply simp
-(* x is output *)
- apply (erule exE)
-apply (erule conjE)
- apply (drule sym)
- apply (drule sym)
-apply simp
-apply hypsubst+
-apply (frule reachable_rename)
-apply simp
-(* x is internal *)
-apply (frule reachable_rename)
-apply auto
-done
-
-declare split_if [split]
-declare if_weak_cong [cong]
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/Seq.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,328 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/Seq.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Partial, Finite and Infinite Sequences (lazy lists), modeled as domain\<close>
-
-theory Seq
-imports "../../HOLCF"
-begin
-
-default_sort pcpo
-
-domain (unsafe) 'a seq = nil  ("nil") | cons (HD :: 'a) (lazy TL :: "'a seq")  (infixr "##" 65)
-
-(*
-   sfilter       :: "('a -> tr) -> 'a seq -> 'a seq"
-   smap          :: "('a -> 'b) -> 'a seq -> 'b seq"
-   sforall       :: "('a -> tr) => 'a seq => bool"
-   sforall2      :: "('a -> tr) -> 'a seq -> tr"
-   slast         :: "'a seq     -> 'a"
-   sconc         :: "'a seq     -> 'a seq -> 'a seq"
-   sdropwhile    :: "('a -> tr)  -> 'a seq -> 'a seq"
-   stakewhile    :: "('a -> tr)  -> 'a seq -> 'a seq"
-   szip          :: "'a seq      -> 'b seq -> ('a*'b) seq"
-   sflat        :: "('a seq) seq  -> 'a seq"
-
-   sfinite       :: "'a seq set"
-   Partial       :: "'a seq => bool"
-   Infinite      :: "'a seq => bool"
-
-   nproj        :: "nat => 'a seq => 'a"
-   sproj        :: "nat => 'a seq => 'a seq"
-*)
-
-inductive
-  Finite :: "'a seq => bool"
-  where
-    sfinite_0:  "Finite nil"
-  | sfinite_n:  "[| Finite tr; a~=UU |] ==> Finite (a##tr)"
-
-declare Finite.intros [simp]
-
-definition
-  Partial :: "'a seq => bool"
-where
-  "Partial x  == (seq_finite x) & ~(Finite x)"
-
-definition
-  Infinite :: "'a seq => bool"
-where
-  "Infinite x == ~(seq_finite x)"
-
-
-subsection \<open>recursive equations of operators\<close>
-
-subsubsection \<open>smap\<close>
-
-fixrec
-  smap :: "('a -> 'b) -> 'a seq -> 'b seq"
-where
-  smap_nil: "smap$f$nil=nil"
-| smap_cons: "[|x~=UU|] ==> smap$f$(x##xs)= (f$x)##smap$f$xs"
-
-lemma smap_UU [simp]: "smap$f$UU=UU"
-by fixrec_simp
-
-subsubsection \<open>sfilter\<close>
-
-fixrec
-   sfilter :: "('a -> tr) -> 'a seq -> 'a seq"
-where
-  sfilter_nil: "sfilter$P$nil=nil"
-| sfilter_cons:
-    "x~=UU ==> sfilter$P$(x##xs)=
-              (If P$x then x##(sfilter$P$xs) else sfilter$P$xs)"
-
-lemma sfilter_UU [simp]: "sfilter$P$UU=UU"
-by fixrec_simp
-
-subsubsection \<open>sforall2\<close>
-
-fixrec
-  sforall2 :: "('a -> tr) -> 'a seq -> tr"
-where
-  sforall2_nil: "sforall2$P$nil=TT"
-| sforall2_cons:
-    "x~=UU ==> sforall2$P$(x##xs)= ((P$x) andalso sforall2$P$xs)"
-
-lemma sforall2_UU [simp]: "sforall2$P$UU=UU"
-by fixrec_simp
-
-definition
-  sforall_def: "sforall P t == (sforall2$P$t ~=FF)"
-
-subsubsection \<open>stakewhile\<close>
-
-fixrec
-  stakewhile :: "('a -> tr)  -> 'a seq -> 'a seq"
-where
-  stakewhile_nil: "stakewhile$P$nil=nil"
-| stakewhile_cons:
-    "x~=UU ==> stakewhile$P$(x##xs) =
-              (If P$x then x##(stakewhile$P$xs) else nil)"
-
-lemma stakewhile_UU [simp]: "stakewhile$P$UU=UU"
-by fixrec_simp
-
-subsubsection \<open>sdropwhile\<close>
-
-fixrec
-  sdropwhile :: "('a -> tr) -> 'a seq -> 'a seq"
-where
-  sdropwhile_nil: "sdropwhile$P$nil=nil"
-| sdropwhile_cons:
-    "x~=UU ==> sdropwhile$P$(x##xs) =
-              (If P$x then sdropwhile$P$xs else x##xs)"
-
-lemma sdropwhile_UU [simp]: "sdropwhile$P$UU=UU"
-by fixrec_simp
-
-subsubsection \<open>slast\<close>
-
-fixrec
-  slast :: "'a seq -> 'a"
-where
-  slast_nil: "slast$nil=UU"
-| slast_cons:
-    "x~=UU ==> slast$(x##xs)= (If is_nil$xs then x else slast$xs)"
-
-lemma slast_UU [simp]: "slast$UU=UU"
-by fixrec_simp
-
-subsubsection \<open>sconc\<close>
-
-fixrec
-  sconc :: "'a seq -> 'a seq -> 'a seq"
-where
-  sconc_nil: "sconc$nil$y = y"
-| sconc_cons':
-    "x~=UU ==> sconc$(x##xs)$y = x##(sconc$xs$y)"
-
-abbreviation
-  sconc_syn :: "'a seq => 'a seq => 'a seq"  (infixr "@@" 65) where
-  "xs @@ ys == sconc $ xs $ ys"
-
-lemma sconc_UU [simp]: "UU @@ y=UU"
-by fixrec_simp
-
-lemma sconc_cons [simp]: "(x##xs) @@ y=x##(xs @@ y)"
-apply (cases "x=UU")
-apply simp_all
-done
-
-declare sconc_cons' [simp del]
-
-subsubsection \<open>sflat\<close>
-
-fixrec
-  sflat :: "('a seq) seq -> 'a seq"
-where
-  sflat_nil: "sflat$nil=nil"
-| sflat_cons': "x~=UU ==> sflat$(x##xs)= x@@(sflat$xs)"
-
-lemma sflat_UU [simp]: "sflat$UU=UU"
-by fixrec_simp
-
-lemma sflat_cons [simp]: "sflat$(x##xs)= x@@(sflat$xs)"
-by (cases "x=UU", simp_all)
-
-declare sflat_cons' [simp del]
-
-subsubsection \<open>szip\<close>
-
-fixrec
-  szip :: "'a seq -> 'b seq -> ('a*'b) seq"
-where
-  szip_nil: "szip$nil$y=nil"
-| szip_cons_nil: "x~=UU ==> szip$(x##xs)$nil=UU"
-| szip_cons:
-    "[| x~=UU; y~=UU|] ==> szip$(x##xs)$(y##ys) = (x,y)##szip$xs$ys"
-
-lemma szip_UU1 [simp]: "szip$UU$y=UU"
-by fixrec_simp
-
-lemma szip_UU2 [simp]: "x~=nil ==> szip$x$UU=UU"
-by (cases x, simp_all, fixrec_simp)
-
-
-subsection "scons, nil"
-
-lemma scons_inject_eq:
- "[|x~=UU;y~=UU|]==> (x##xs=y##ys) = (x=y & xs=ys)"
-by simp
-
-lemma nil_less_is_nil: "nil<<x ==> nil=x"
-apply (cases x)
-apply simp
-apply simp
-apply simp
-done
-
-subsection "sfilter, sforall, sconc"
-
-lemma if_and_sconc [simp]: "(if b then tr1 else tr2) @@ tr
-        = (if b then tr1 @@ tr else tr2 @@ tr)"
-by simp
-
-
-lemma sfiltersconc: "sfilter$P$(x @@ y) = (sfilter$P$x @@ sfilter$P$y)"
-apply (induct x)
-(* adm *)
-apply simp
-(* base cases *)
-apply simp
-apply simp
-(* main case *)
-apply (rule_tac p="P$a" in trE)
-apply simp
-apply simp
-apply simp
-done
-
-lemma sforallPstakewhileP: "sforall P (stakewhile$P$x)"
-apply (simp add: sforall_def)
-apply (induct x)
-(* adm *)
-apply simp
-(* base cases *)
-apply simp
-apply simp
-(* main case *)
-apply (rule_tac p="P$a" in trE)
-apply simp
-apply simp
-apply simp
-done
-
-lemma forallPsfilterP: "sforall P (sfilter$P$x)"
-apply (simp add: sforall_def)
-apply (induct x)
-(* adm *)
-apply simp
-(* base cases *)
-apply simp
-apply simp
-(* main case *)
-apply (rule_tac p="P$a" in trE)
-apply simp
-apply simp
-apply simp
-done
-
-
-subsection "Finite"
-
-(* ----------------------------------------------------  *)
-(* Proofs of rewrite rules for Finite:                  *)
-(* 1. Finite(nil),   (by definition)                    *)
-(* 2. ~Finite(UU),                                      *)
-(* 3. a~=UU==> Finite(a##x)=Finite(x)                  *)
-(* ----------------------------------------------------  *)
-
-lemma Finite_UU_a: "Finite x --> x~=UU"
-apply (rule impI)
-apply (erule Finite.induct)
- apply simp
-apply simp
-done
-
-lemma Finite_UU [simp]: "~(Finite UU)"
-apply (cut_tac x="UU" in Finite_UU_a)
-apply fast
-done
-
-lemma Finite_cons_a: "Finite x --> a~=UU --> x=a##xs --> Finite xs"
-apply (intro strip)
-apply (erule Finite.cases)
-apply fastforce
-apply simp
-done
-
-lemma Finite_cons: "a~=UU ==>(Finite (a##x)) = (Finite x)"
-apply (rule iffI)
-apply (erule (1) Finite_cons_a [rule_format])
-apply fast
-apply simp
-done
-
-lemma Finite_upward: "\<lbrakk>Finite x; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> Finite y"
-apply (induct arbitrary: y set: Finite)
-apply (case_tac y, simp, simp, simp)
-apply (case_tac y, simp, simp)
-apply simp
-done
-
-lemma adm_Finite [simp]: "adm Finite"
-by (rule adm_upward, rule Finite_upward)
-
-
-subsection "induction"
-
-
-(*--------------------------------   *)
-(* Extensions to Induction Theorems  *)
-(*--------------------------------   *)
-
-
-lemma seq_finite_ind_lemma:
-  assumes "(!!n. P(seq_take n$s))"
-  shows "seq_finite(s) -->P(s)"
-apply (unfold seq.finite_def)
-apply (intro strip)
-apply (erule exE)
-apply (erule subst)
-apply (rule assms)
-done
-
-
-lemma seq_finite_ind: "!!P.[|P(UU);P(nil);
-   !! x s1.[|x~=UU;P(s1)|] ==> P(x##s1)
-   |] ==> seq_finite(s) --> P(s)"
-apply (rule seq_finite_ind_lemma)
-apply (erule seq.finite_induct)
- apply assumption
-apply simp
-done
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/Sequence.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1084 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/Sequence.thy
-    Author:     Olaf Müller
-
-Sequences over flat domains with lifted elements.
-*)
-
-theory Sequence
-imports Seq
-begin
-
-default_sort type
-
-type_synonym 'a Seq = "'a lift seq"
-
-definition Consq :: "'a \<Rightarrow> 'a Seq \<rightarrow> 'a Seq"
-  where "Consq a = (LAM s. Def a ## s)"
-
-definition Filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq"
-  where "Filter P = sfilter $ (flift2 P)"
-
-definition Map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a Seq \<rightarrow> 'b Seq"
-  where "Map f = smap $ (flift2 f)"
-
-definition Forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<Rightarrow> bool"
-  where "Forall P = sforall (flift2 P)"
-
-definition Last :: "'a Seq \<rightarrow> 'a lift"
-  where "Last = slast"
-
-definition Dropwhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq"
-  where "Dropwhile P = sdropwhile $ (flift2 P)"
-
-definition Takewhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq"
-  where "Takewhile P = stakewhile $ (flift2 P)"
-
-definition Zip :: "'a Seq \<rightarrow> 'b Seq \<rightarrow> ('a * 'b) Seq"
-where
-  "Zip = (fix$(LAM h t1 t2. case t1 of
-               nil   => nil
-             | x##xs => (case t2 of
-                          nil => UU
-                        | y##ys => (case x of
-                                      UU  => UU
-                                    | Def a => (case y of
-                                                  UU => UU
-                                                | Def b => Def (a,b)##(h$xs$ys))))))"
-
-definition Flat :: "'a Seq seq \<rightarrow> 'a Seq"
-  where "Flat = sflat"
-
-definition Filter2 :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq"
-where
-  "Filter2 P = (fix $ (LAM h t. case t of
-            nil \<Rightarrow> nil
-          | x##xs \<Rightarrow> (case x of UU \<Rightarrow> UU | Def y \<Rightarrow> (if P y
-                     then x##(h$xs)
-                     else     h$xs))))"
-
-abbreviation Consq_syn  ("(_/\<leadsto>_)" [66,65] 65)
-  where "a\<leadsto>s \<equiv> Consq a $ s"
-
-
-text \<open>List enumeration\<close>
-syntax
-  "_totlist"      :: "args => 'a Seq"              ("[(_)!]")
-  "_partlist"     :: "args => 'a Seq"              ("[(_)?]")
-translations
-  "[x, xs!]"     == "x\<leadsto>[xs!]"
-  "[x!]"         == "x\<leadsto>nil"
-  "[x, xs?]"     == "x\<leadsto>[xs?]"
-  "[x?]"         == "x\<leadsto>CONST bottom"
-
-
-declare andalso_and [simp]
-declare andalso_or [simp]
-
-
-subsection "recursive equations of operators"
-
-subsubsection "Map"
-
-lemma Map_UU: "Map f$UU =UU"
-  by (simp add: Map_def)
-
-lemma Map_nil: "Map f$nil =nil"
-  by (simp add: Map_def)
-
-lemma Map_cons: "Map f$(x\<leadsto>xs)=(f x) \<leadsto> Map f$xs"
-  by (simp add: Map_def Consq_def flift2_def)
-
-
-subsubsection \<open>Filter\<close>
-
-lemma Filter_UU: "Filter P$UU =UU"
-  by (simp add: Filter_def)
-
-lemma Filter_nil: "Filter P$nil =nil"
-  by (simp add: Filter_def)
-
-lemma Filter_cons:
-  "Filter P$(x\<leadsto>xs)= (if P x then x\<leadsto>(Filter P$xs) else Filter P$xs)"
-  by (simp add: Filter_def Consq_def flift2_def If_and_if)
-
-
-subsubsection \<open>Forall\<close>
-
-lemma Forall_UU: "Forall P UU"
-  by (simp add: Forall_def sforall_def)
-
-lemma Forall_nil: "Forall P nil"
-  by (simp add: Forall_def sforall_def)
-
-lemma Forall_cons: "Forall P (x\<leadsto>xs)= (P x & Forall P xs)"
-  by (simp add: Forall_def sforall_def Consq_def flift2_def)
-
-
-subsubsection \<open>Conc\<close>
-
-lemma Conc_cons: "(x\<leadsto>xs) @@ y = x\<leadsto>(xs @@y)"
-  by (simp add: Consq_def)
-
-
-subsubsection \<open>Takewhile\<close>
-
-lemma Takewhile_UU: "Takewhile P$UU =UU"
-  by (simp add: Takewhile_def)
-
-lemma Takewhile_nil: "Takewhile P$nil =nil"
-  by (simp add: Takewhile_def)
-
-lemma Takewhile_cons:
-  "Takewhile P$(x\<leadsto>xs)= (if P x then x\<leadsto>(Takewhile P$xs) else nil)"
-  by (simp add: Takewhile_def Consq_def flift2_def If_and_if)
-
-
-subsubsection \<open>DropWhile\<close>
-
-lemma Dropwhile_UU: "Dropwhile P$UU =UU"
-  by (simp add: Dropwhile_def)
-
-lemma Dropwhile_nil: "Dropwhile P$nil =nil"
-  by (simp add: Dropwhile_def)
-
-lemma Dropwhile_cons:
-  "Dropwhile P$(x\<leadsto>xs)= (if P x then Dropwhile P$xs else x\<leadsto>xs)"
-  by (simp add: Dropwhile_def Consq_def flift2_def If_and_if)
-
-
-subsubsection \<open>Last\<close>
-
-lemma Last_UU: "Last$UU =UU"
-  by (simp add: Last_def)
-
-lemma Last_nil: "Last$nil =UU"
-  by (simp add: Last_def)
-
-lemma Last_cons: "Last$(x\<leadsto>xs)= (if xs=nil then Def x else Last$xs)"
-  apply (simp add: Last_def Consq_def)
-  apply (cases xs)
-  apply simp_all
-  done
-
-
-subsubsection \<open>Flat\<close>
-
-lemma Flat_UU: "Flat$UU =UU"
-  by (simp add: Flat_def)
-
-lemma Flat_nil: "Flat$nil =nil"
-  by (simp add: Flat_def)
-
-lemma Flat_cons: "Flat$(x##xs)= x @@ (Flat$xs)"
-  by (simp add: Flat_def Consq_def)
-
-
-subsubsection \<open>Zip\<close>
-
-lemma Zip_unfold:
-  "Zip = (LAM t1 t2. case t1 of
-                  nil   => nil
-                | x##xs => (case t2 of
-                             nil => UU
-                           | y##ys => (case x of
-                                         UU  => UU
-                                       | Def a => (case y of
-                                                     UU => UU
-                                                   | Def b => Def (a,b)##(Zip$xs$ys)))))"
-  apply (rule trans)
-  apply (rule fix_eq4)
-  apply (rule Zip_def)
-  apply (rule beta_cfun)
-  apply simp
-  done
-
-lemma Zip_UU1: "Zip$UU$y =UU"
-  apply (subst Zip_unfold)
-  apply simp
-  done
-
-lemma Zip_UU2: "x~=nil ==> Zip$x$UU =UU"
-  apply (subst Zip_unfold)
-  apply simp
-  apply (cases x)
-  apply simp_all
-  done
-
-lemma Zip_nil: "Zip$nil$y =nil"
-  apply (subst Zip_unfold)
-  apply simp
-  done
-
-lemma Zip_cons_nil: "Zip$(x\<leadsto>xs)$nil= UU"
-  apply (subst Zip_unfold)
-  apply (simp add: Consq_def)
-  done
-
-lemma Zip_cons: "Zip$(x\<leadsto>xs)$(y\<leadsto>ys)= (x,y) \<leadsto> Zip$xs$ys"
-  apply (rule trans)
-  apply (subst Zip_unfold)
-  apply simp
-  apply (simp add: Consq_def)
-  done
-
-lemmas [simp del] =
-  sfilter_UU sfilter_nil sfilter_cons
-  smap_UU smap_nil smap_cons
-  sforall2_UU sforall2_nil sforall2_cons
-  slast_UU slast_nil slast_cons
-  stakewhile_UU  stakewhile_nil  stakewhile_cons
-  sdropwhile_UU  sdropwhile_nil  sdropwhile_cons
-  sflat_UU sflat_nil sflat_cons
-  szip_UU1 szip_UU2 szip_nil szip_cons_nil szip_cons
-
-lemmas [simp] =
-  Filter_UU Filter_nil Filter_cons
-  Map_UU Map_nil Map_cons
-  Forall_UU Forall_nil Forall_cons
-  Last_UU Last_nil Last_cons
-  Conc_cons
-  Takewhile_UU Takewhile_nil Takewhile_cons
-  Dropwhile_UU Dropwhile_nil Dropwhile_cons
-  Zip_UU1 Zip_UU2 Zip_nil Zip_cons_nil Zip_cons
-
-
-
-section "Cons"
-
-lemma Consq_def2: "a\<leadsto>s = (Def a)##s"
-  by (simp add: Consq_def)
-
-lemma Seq_exhaust: "x = UU | x = nil | (? a s. x = a \<leadsto> s)"
-  apply (simp add: Consq_def2)
-  apply (cut_tac seq.nchotomy)
-  apply (fast dest: not_Undef_is_Def [THEN iffD1])
-  done
-
-
-lemma Seq_cases: "!!P. [| x = UU ==> P; x = nil ==> P; !!a s. x = a \<leadsto> s  ==> P |] ==> P"
-  apply (cut_tac x="x" in Seq_exhaust)
-  apply (erule disjE)
-  apply simp
-  apply (erule disjE)
-  apply simp
-  apply (erule exE)+
-  apply simp
-  done
-
-(*
-fun Seq_case_tac s i = rule_tac x",s)] Seq_cases i
-          THEN hyp_subst_tac i THEN hyp_subst_tac (i+1) THEN hyp_subst_tac (i+2);
-*)
-(* on a\<leadsto>s only simp_tac, as full_simp_tac is uncomplete and often causes errors *)
-(*
-fun Seq_case_simp_tac s i = Seq_case_tac s i THEN Asm_simp_tac (i+2)
-                                             THEN Asm_full_simp_tac (i+1)
-                                             THEN Asm_full_simp_tac i;
-*)
-
-lemma Cons_not_UU: "a\<leadsto>s ~= UU"
-  apply (subst Consq_def2)
-  apply simp
-  done
-
-
-lemma Cons_not_less_UU: "~(a\<leadsto>x) << UU"
-  apply (rule notI)
-  apply (drule below_antisym)
-  apply simp
-  apply (simp add: Cons_not_UU)
-  done
-
-lemma Cons_not_less_nil: "~a\<leadsto>s << nil"
-  by (simp add: Consq_def2)
-
-lemma Cons_not_nil: "a\<leadsto>s ~= nil"
-  by (simp add: Consq_def2)
-
-lemma Cons_not_nil2: "nil ~= a\<leadsto>s"
-  by (simp add: Consq_def2)
-
-lemma Cons_inject_eq: "(a\<leadsto>s = b\<leadsto>t) = (a = b & s = t)"
-  apply (simp only: Consq_def2)
-  apply (simp add: scons_inject_eq)
-  done
-
-lemma Cons_inject_less_eq: "(a\<leadsto>s<<b\<leadsto>t) = (a = b & s<<t)"
-  by (simp add: Consq_def2)
-
-lemma seq_take_Cons: "seq_take (Suc n)$(a\<leadsto>x) = a\<leadsto> (seq_take n$x)"
-  by (simp add: Consq_def)
-
-lemmas [simp] =
-  Cons_not_nil2 Cons_inject_eq Cons_inject_less_eq seq_take_Cons
-  Cons_not_UU Cons_not_less_UU Cons_not_less_nil Cons_not_nil
-
-
-subsection "induction"
-
-lemma Seq_induct: "!! P. [| adm P; P UU; P nil; !! a s. P s ==> P (a\<leadsto>s)|] ==> P x"
-  apply (erule (2) seq.induct)
-  apply defined
-  apply (simp add: Consq_def)
-  done
-
-lemma Seq_FinitePartial_ind:
-  "!! P.[|P UU;P nil; !! a s. P s ==> P(a\<leadsto>s) |]
-                  ==> seq_finite x --> P x"
-  apply (erule (1) seq_finite_ind)
-  apply defined
-  apply (simp add: Consq_def)
-  done
-
-lemma Seq_Finite_ind:
-  "!! P.[| Finite x; P nil; !! a s. [| Finite s; P s|] ==> P (a\<leadsto>s) |] ==> P x"
-  apply (erule (1) Finite.induct)
-  apply defined
-  apply (simp add: Consq_def)
-  done
-
-
-(* rws are definitions to be unfolded for admissibility check *)
-(*
-fun Seq_induct_tac s rws i = rule_tac x",s)] Seq_induct i
-                         THEN (REPEAT_DETERM (CHANGED (Asm_simp_tac (i+1))))
-                         THEN simp add: rws) i;
-
-fun Seq_Finite_induct_tac i = erule Seq_Finite_ind i
-                              THEN (REPEAT_DETERM (CHANGED (Asm_simp_tac i)));
-
-fun pair_tac s = rule_tac p",s)] prod.exhaust
-                          THEN' hyp_subst_tac THEN' Simp_tac;
-*)
-(* induction on a sequence of pairs with pairsplitting and simplification *)
-(*
-fun pair_induct_tac s rws i =
-           rule_tac x",s)] Seq_induct i
-           THEN pair_tac "a" (i+3)
-           THEN (REPEAT_DETERM (CHANGED (Simp_tac (i+1))))
-           THEN simp add: rws) i;
-*)
-
-
-(* ------------------------------------------------------------------------------------ *)
-
-subsection "HD,TL"
-
-lemma HD_Cons [simp]: "HD$(x\<leadsto>y) = Def x"
-  by (simp add: Consq_def)
-
-lemma TL_Cons [simp]: "TL$(x\<leadsto>y) = y"
-  by (simp add: Consq_def)
-
-(* ------------------------------------------------------------------------------------ *)
-
-subsection "Finite, Partial, Infinite"
-
-lemma Finite_Cons [simp]: "Finite (a\<leadsto>xs) = Finite xs"
-  by (simp add: Consq_def2 Finite_cons)
-
-lemma FiniteConc_1: "Finite (x::'a Seq) ==> Finite y --> Finite (x@@y)"
-  apply (erule Seq_Finite_ind)
-  apply simp_all
-  done
-
-lemma FiniteConc_2: "Finite (z::'a Seq) ==> !x y. z= x@@y --> (Finite x & Finite y)"
-  apply (erule Seq_Finite_ind)
-  (* nil*)
-  apply (intro strip)
-  apply (rule_tac x="x" in Seq_cases, simp_all)
-  (* cons *)
-  apply (intro strip)
-  apply (rule_tac x="x" in Seq_cases, simp_all)
-  apply (rule_tac x="y" in Seq_cases, simp_all)
-  done
-
-lemma FiniteConc [simp]: "Finite(x@@y) = (Finite (x::'a Seq) & Finite y)"
-  apply (rule iffI)
-  apply (erule FiniteConc_2 [rule_format])
-  apply (rule refl)
-  apply (rule FiniteConc_1 [rule_format])
-  apply auto
-  done
-
-
-lemma FiniteMap1: "Finite s ==> Finite (Map f$s)"
-  apply (erule Seq_Finite_ind)
-  apply simp_all
-  done
-
-lemma FiniteMap2: "Finite s ==> ! t. (s = Map f$t) --> Finite t"
-  apply (erule Seq_Finite_ind)
-  apply (intro strip)
-  apply (rule_tac x="t" in Seq_cases, simp_all)
-  (* main case *)
-  apply auto
-  apply (rule_tac x="t" in Seq_cases, simp_all)
-  done
-
-lemma Map2Finite: "Finite (Map f$s) = Finite s"
-  apply auto
-  apply (erule FiniteMap2 [rule_format])
-  apply (rule refl)
-  apply (erule FiniteMap1)
-  done
-
-
-lemma FiniteFilter: "Finite s ==> Finite (Filter P$s)"
-  apply (erule Seq_Finite_ind)
-  apply simp_all
-  done
-
-
-(* ----------------------------------------------------------------------------------- *)
-
-subsection "Conc"
-
-lemma Conc_cong: "!! x::'a Seq. Finite x ==> ((x @@ y) = (x @@ z)) = (y = z)"
-  apply (erule Seq_Finite_ind)
-  apply simp_all
-  done
-
-lemma Conc_assoc: "(x @@ y) @@ z = (x::'a Seq) @@ y @@ z"
-  apply (rule_tac x="x" in Seq_induct)
-  apply simp_all
-  done
-
-lemma nilConc [simp]: "s@@ nil = s"
-  apply (induct s)
-  apply simp
-  apply simp
-  apply simp
-  apply simp
-  done
-
-
-(* should be same as nil_is_Conc2 when all nils are turned to right side !! *)
-lemma nil_is_Conc: "(nil = x @@ y) = ((x::'a Seq)= nil & y = nil)"
-  apply (rule_tac x="x" in Seq_cases)
-  apply auto
-  done
-
-lemma nil_is_Conc2: "(x @@ y = nil) = ((x::'a Seq)= nil & y = nil)"
-  apply (rule_tac x="x" in Seq_cases)
-  apply auto
-  done
-
-
-(* ------------------------------------------------------------------------------------ *)
-
-subsection "Last"
-
-lemma Finite_Last1: "Finite s ==> s~=nil --> Last$s~=UU"
-  apply (erule Seq_Finite_ind, simp_all)
-  done
-
-lemma Finite_Last2: "Finite s ==> Last$s=UU --> s=nil"
-  apply (erule Seq_Finite_ind, simp_all)
-  apply fast
-  done
-
-
-(* ------------------------------------------------------------------------------------ *)
-
-
-subsection "Filter, Conc"
-
-
-lemma FilterPQ: "Filter P$(Filter Q$s) = Filter (%x. P x & Q x)$s"
-  apply (rule_tac x="s" in Seq_induct, simp_all)
-  done
-
-lemma FilterConc: "Filter P$(x @@ y) = (Filter P$x @@ Filter P$y)"
-  apply (simp add: Filter_def sfiltersconc)
-  done
-
-(* ------------------------------------------------------------------------------------ *)
-
-subsection "Map"
-
-lemma MapMap: "Map f$(Map g$s) = Map (f o g)$s"
-  apply (rule_tac x="s" in Seq_induct, simp_all)
-  done
-
-lemma MapConc: "Map f$(x@@y) = (Map f$x) @@ (Map f$y)"
-  apply (rule_tac x="x" in Seq_induct, simp_all)
-  done
-
-lemma MapFilter: "Filter P$(Map f$x) = Map f$(Filter (P o f)$x)"
-  apply (rule_tac x="x" in Seq_induct, simp_all)
-  done
-
-lemma nilMap: "nil = (Map f$s) --> s= nil"
-  apply (rule_tac x="s" in Seq_cases, simp_all)
-  done
-
-
-lemma ForallMap: "Forall P (Map f$s) = Forall (P o f) s"
-  apply (rule_tac x="s" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-
-
-
-(* ------------------------------------------------------------------------------------ *)
-
-subsection "Forall"
-
-lemma ForallPForallQ1: "Forall P ys & (! x. P x --> Q x) --> Forall Q ys"
-  apply (rule_tac x="ys" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-lemmas ForallPForallQ =
-  ForallPForallQ1 [THEN mp, OF conjI, OF _ allI, OF _ impI]
-
-lemma Forall_Conc_impl: "(Forall P x & Forall P y) --> Forall P (x @@ y)"
-  apply (rule_tac x="x" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-lemma Forall_Conc [simp]:
-  "Finite x ==> Forall P (x @@ y) = (Forall P x & Forall P y)"
-  apply (erule Seq_Finite_ind, simp_all)
-  done
-
-lemma ForallTL1: "Forall P s  --> Forall P (TL$s)"
-  apply (rule_tac x="s" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-lemmas ForallTL = ForallTL1 [THEN mp]
-
-lemma ForallDropwhile1: "Forall P s  --> Forall P (Dropwhile Q$s)"
-  apply (rule_tac x="s" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-lemmas ForallDropwhile = ForallDropwhile1 [THEN mp]
-
-
-(* only admissible in t, not if done in s *)
-
-lemma Forall_prefix: "! s. Forall P s --> t<<s --> Forall P t"
-  apply (rule_tac x="t" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  apply (intro strip)
-  apply (rule_tac x="sa" in Seq_cases)
-  apply simp
-  apply auto
-  done
-
-lemmas Forall_prefixclosed = Forall_prefix [rule_format]
-
-lemma Forall_postfixclosed: "[| Finite h; Forall P s; s= h @@ t |] ==> Forall P t"
-  by auto
-
-
-lemma ForallPFilterQR1:
-  "((! x. P x --> (Q x = R x)) & Forall P tr) --> Filter Q$tr = Filter R$tr"
-  apply (rule_tac x="tr" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-lemmas ForallPFilterQR = ForallPFilterQR1 [THEN mp, OF conjI, OF allI]
-
-
-(* ------------------------------------------------------------------------------------- *)
-
-subsection "Forall, Filter"
-
-
-lemma ForallPFilterP: "Forall P (Filter P$x)"
-  by (simp add: Filter_def Forall_def forallPsfilterP)
-
-(* holds also in other direction, then equal to forallPfilterP *)
-lemma ForallPFilterPid1: "Forall P x --> Filter P$x = x"
-  apply (rule_tac x="x" in Seq_induct)
-  apply (simp add: Forall_def sforall_def Filter_def)
-  apply simp_all
-  done
-
-lemmas ForallPFilterPid = ForallPFilterPid1 [THEN mp]
-
-
-(* holds also in other direction *)
-lemma ForallnPFilterPnil1: "!! ys . Finite ys ==>
-   Forall (%x. ~P x) ys --> Filter P$ys = nil "
-  apply (erule Seq_Finite_ind, simp_all)
-  done
-
-lemmas ForallnPFilterPnil = ForallnPFilterPnil1 [THEN mp]
-
-
-(* holds also in other direction *)
-lemma ForallnPFilterPUU1: "~Finite ys & Forall (%x. ~P x) ys --> Filter P$ys = UU"
-  apply (rule_tac x="ys" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-lemmas ForallnPFilterPUU = ForallnPFilterPUU1 [THEN mp, OF conjI]
-
-
-(* inverse of ForallnPFilterPnil *)
-
-lemma FilternPnilForallP [rule_format]: "Filter P$ys = nil -->
-   (Forall (%x. ~P x) ys & Finite ys)"
-  apply (rule_tac x="ys" in Seq_induct)
-  (* adm *)
-  apply (simp add: Forall_def sforall_def)
-  (* base cases *)
-  apply simp
-  apply simp
-  (* main case *)
-  apply simp
-  done
-
-
-(* inverse of ForallnPFilterPUU *)
-
-lemma FilternPUUForallP:
-  assumes "Filter P$ys = UU"
-  shows "Forall (%x. ~P x) ys  & ~Finite ys"
-proof
-  show "Forall (%x. ~P x) ys"
-  proof (rule classical)
-    assume "\<not> ?thesis"
-    then have "Filter P$ys ~= UU"
-      apply (rule rev_mp)
-      apply (induct ys rule: Seq_induct)
-      apply (simp add: Forall_def sforall_def)
-      apply simp_all
-      done
-    with assms show ?thesis by contradiction
-  qed
-  show "~ Finite ys"
-  proof
-    assume "Finite ys"
-    then have "Filter P$ys ~= UU"
-      by (rule Seq_Finite_ind) simp_all
-    with assms show False by contradiction
-  qed
-qed
-
-
-lemma ForallQFilterPnil:
-  "!! Q P.[| Forall Q ys; Finite ys; !!x. Q x ==> ~P x|]
-    ==> Filter P$ys = nil"
-  apply (erule ForallnPFilterPnil)
-  apply (erule ForallPForallQ)
-  apply auto
-  done
-
-lemma ForallQFilterPUU:
- "!! Q P. [| ~Finite ys; Forall Q ys;  !!x. Q x ==> ~P x|]
-    ==> Filter P$ys = UU "
-  apply (erule ForallnPFilterPUU)
-  apply (erule ForallPForallQ)
-  apply auto
-  done
-
-
-
-(* ------------------------------------------------------------------------------------- *)
-
-subsection "Takewhile, Forall, Filter"
-
-
-lemma ForallPTakewhileP [simp]: "Forall P (Takewhile P$x)"
-  by (simp add: Forall_def Takewhile_def sforallPstakewhileP)
-
-
-lemma ForallPTakewhileQ [simp]: "!! P. [| !!x. Q x==> P x |] ==> Forall P (Takewhile Q$x)"
-  apply (rule ForallPForallQ)
-  apply (rule ForallPTakewhileP)
-  apply auto
-  done
-
-
-lemma FilterPTakewhileQnil [simp]:
-  "!! Q P.[| Finite (Takewhile Q$ys); !!x. Q x ==> ~P x |]
-   ==> Filter P$(Takewhile Q$ys) = nil"
-  apply (erule ForallnPFilterPnil)
-  apply (rule ForallPForallQ)
-  apply (rule ForallPTakewhileP)
-  apply auto
-  done
-
-lemma FilterPTakewhileQid [simp]:
- "!! Q P. [| !!x. Q x ==> P x |] ==>
-            Filter P$(Takewhile Q$ys) = (Takewhile Q$ys)"
-  apply (rule ForallPFilterPid)
-  apply (rule ForallPForallQ)
-  apply (rule ForallPTakewhileP)
-  apply auto
-  done
-
-
-lemma Takewhile_idempotent: "Takewhile P$(Takewhile P$s) = Takewhile P$s"
-  apply (rule_tac x="s" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-lemma ForallPTakewhileQnP [simp]: "Forall P s --> Takewhile (%x. Q x | (~P x))$s = Takewhile Q$s"
-  apply (rule_tac x="s" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-lemma ForallPDropwhileQnP [simp]: "Forall P s --> Dropwhile (%x. Q x | (~P x))$s = Dropwhile Q$s"
-  apply (rule_tac x="s" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-
-lemma TakewhileConc1: "Forall P s --> Takewhile P$(s @@ t) = s @@ (Takewhile P$t)"
-  apply (rule_tac x="s" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-lemmas TakewhileConc = TakewhileConc1 [THEN mp]
-
-lemma DropwhileConc1: "Finite s ==> Forall P s --> Dropwhile P$(s @@ t) = Dropwhile P$t"
-  apply (erule Seq_Finite_ind, simp_all)
-  done
-
-lemmas DropwhileConc = DropwhileConc1 [THEN mp]
-
-
-
-(* ----------------------------------------------------------------------------------- *)
-
-subsection "coinductive characterizations of Filter"
-
-lemma divide_Seq_lemma:
- "HD$(Filter P$y) = Def x
-    --> y = ((Takewhile (%x. ~P x)$y) @@ (x \<leadsto> TL$(Dropwhile (%a. ~P a)$y)))
-             & Finite (Takewhile (%x. ~ P x)$y)  & P x"
-  (* FIX: pay attention: is only admissible with chain-finite package to be added to
-          adm test and Finite f x admissibility *)
-  apply (rule_tac x="y" in Seq_induct)
-  apply (simp add: adm_subst [OF _ adm_Finite])
-  apply simp
-  apply simp
-  apply (case_tac "P a")
-   apply simp
-   apply blast
-  (* ~ P a *)
-  apply simp
-  done
-
-lemma divide_Seq: "(x\<leadsto>xs) << Filter P$y
-   ==> y = ((Takewhile (%a. ~ P a)$y) @@ (x \<leadsto> TL$(Dropwhile (%a. ~ P a)$y)))
-      & Finite (Takewhile (%a. ~ P a)$y)  & P x"
-  apply (rule divide_Seq_lemma [THEN mp])
-  apply (drule_tac f="HD" and x="x\<leadsto>xs" in  monofun_cfun_arg)
-  apply simp
-  done
-
-
-lemma nForall_HDFilter: "~Forall P y --> (? x. HD$(Filter (%a. ~P a)$y) = Def x)"
-  unfolding not_Undef_is_Def [symmetric]
-  apply (induct y rule: Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-
-lemma divide_Seq2: "~Forall P y
-  ==> ? x. y= (Takewhile P$y @@ (x \<leadsto> TL$(Dropwhile P$y))) &
-      Finite (Takewhile P$y) & (~ P x)"
-  apply (drule nForall_HDFilter [THEN mp])
-  apply safe
-  apply (rule_tac x="x" in exI)
-  apply (cut_tac P1="%x. ~ P x" in divide_Seq_lemma [THEN mp])
-  apply auto
-  done
-
-
-lemma divide_Seq3: "~Forall P y
-  ==> ? x bs rs. y= (bs @@ (x\<leadsto>rs)) & Finite bs & Forall P bs & (~ P x)"
-  apply (drule divide_Seq2)
-  apply fastforce
-  done
-
-lemmas [simp] = FilterPQ FilterConc Conc_cong
-
-
-(* ------------------------------------------------------------------------------------- *)
-
-
-subsection "take_lemma"
-
-lemma seq_take_lemma: "(!n. seq_take n$x = seq_take n$x') = (x = x')"
-  apply (rule iffI)
-  apply (rule seq.take_lemma)
-  apply auto
-  done
-
-lemma take_reduction1:
-  "\<forall>n. ((! k. k < n --> seq_take k$y1 = seq_take k$y2)
-    --> seq_take n$(x @@ (t\<leadsto>y1)) =  seq_take n$(x @@ (t\<leadsto>y2)))"
-  apply (rule_tac x="x" in Seq_induct)
-  apply simp_all
-  apply (intro strip)
-  apply (case_tac "n")
-  apply auto
-  apply (case_tac "n")
-  apply auto
-  done
-
-
-lemma take_reduction:
-  "!! n.[| x=y; s=t; !! k. k<n ==> seq_take k$y1 = seq_take k$y2|]
-    ==> seq_take n$(x @@ (s\<leadsto>y1)) =  seq_take n$(y @@ (t\<leadsto>y2))"
-  by (auto intro!: take_reduction1 [rule_format])
-
-(* ------------------------------------------------------------------
-          take-lemma and take_reduction for << instead of =
-   ------------------------------------------------------------------ *)
-
-lemma take_reduction_less1:
-  "\<forall>n. ((! k. k < n --> seq_take k$y1 << seq_take k$y2)
-    --> seq_take n$(x @@ (t\<leadsto>y1)) <<  seq_take n$(x @@ (t\<leadsto>y2)))"
-  apply (rule_tac x="x" in Seq_induct)
-  apply simp_all
-  apply (intro strip)
-  apply (case_tac "n")
-  apply auto
-  apply (case_tac "n")
-  apply auto
-  done
-
-
-lemma take_reduction_less:
-  "\<And>n.[| x=y; s=t;!! k. k<n ==> seq_take k$y1 << seq_take k$y2|]
-    ==> seq_take n$(x @@ (s\<leadsto>y1)) <<  seq_take n$(y @@ (t\<leadsto>y2))"
-  by (auto intro!: take_reduction_less1 [rule_format])
-
-lemma take_lemma_less1:
-  assumes "!! n. seq_take n$s1 << seq_take n$s2"
-  shows "s1<<s2"
-  apply (rule_tac t="s1" in seq.reach [THEN subst])
-  apply (rule_tac t="s2" in seq.reach [THEN subst])
-  apply (rule lub_mono)
-  apply (rule seq.chain_take [THEN ch2ch_Rep_cfunL])
-  apply (rule seq.chain_take [THEN ch2ch_Rep_cfunL])
-  apply (rule assms)
-  done
-
-
-lemma take_lemma_less: "(!n. seq_take n$x << seq_take n$x') = (x << x')"
-  apply (rule iffI)
-  apply (rule take_lemma_less1)
-  apply auto
-  apply (erule monofun_cfun_arg)
-  done
-
-(* ------------------------------------------------------------------
-          take-lemma proof principles
-   ------------------------------------------------------------------ *)
-
-lemma take_lemma_principle1:
-  "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
-            !! s1 s2 y. [| Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2)|]
-                          ==> (f (s1 @@ y\<leadsto>s2)) = (g (s1 @@ y\<leadsto>s2)) |]
-               ==> A x --> (f x)=(g x)"
-  apply (case_tac "Forall Q x")
-  apply (auto dest!: divide_Seq3)
-  done
-
-lemma take_lemma_principle2:
-  "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
-           !! s1 s2 y. [| Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2)|]
-                          ==> ! n. seq_take n$(f (s1 @@ y\<leadsto>s2))
-                                 = seq_take n$(g (s1 @@ y\<leadsto>s2)) |]
-               ==> A x --> (f x)=(g x)"
-  apply (case_tac "Forall Q x")
-  apply (auto dest!: divide_Seq3)
-  apply (rule seq.take_lemma)
-  apply auto
-  done
-
-
-(* Note: in the following proofs the ordering of proof steps is very
-         important, as otherwise either (Forall Q s1) would be in the IH as
-         assumption (then rule useless) or it is not possible to strengthen
-         the IH apply doing a forall closure of the sequence t (then rule also useless).
-         This is also the reason why the induction rule (nat_less_induct or nat_induct) has to
-         to be imbuilt into the rule, as induction has to be done early and the take lemma
-         has to be used in the trivial direction afterwards for the (Forall Q x) case.  *)
-
-lemma take_lemma_induct:
-"!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
-         !! s1 s2 y n. [| ! t. A t --> seq_take n$(f t) = seq_take n$(g t);
-                          Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2) |]
-                          ==>   seq_take (Suc n)$(f (s1 @@ y\<leadsto>s2))
-                              = seq_take (Suc n)$(g (s1 @@ y\<leadsto>s2)) |]
-               ==> A x --> (f x)=(g x)"
-  apply (rule impI)
-  apply (rule seq.take_lemma)
-  apply (rule mp)
-  prefer 2 apply assumption
-  apply (rule_tac x="x" in spec)
-  apply (rule nat.induct)
-  apply simp
-  apply (rule allI)
-  apply (case_tac "Forall Q xa")
-  apply (force intro!: seq_take_lemma [THEN iffD2, THEN spec])
-  apply (auto dest!: divide_Seq3)
-  done
-
-
-lemma take_lemma_less_induct:
-"!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
-        !! s1 s2 y n. [| ! t m. m < n --> A t --> seq_take m$(f t) = seq_take m$(g t);
-                          Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2) |]
-                          ==>   seq_take n$(f (s1 @@ y\<leadsto>s2))
-                              = seq_take n$(g (s1 @@ y\<leadsto>s2)) |]
-               ==> A x --> (f x)=(g x)"
-  apply (rule impI)
-  apply (rule seq.take_lemma)
-  apply (rule mp)
-  prefer 2 apply assumption
-  apply (rule_tac x="x" in spec)
-  apply (rule nat_less_induct)
-  apply (rule allI)
-  apply (case_tac "Forall Q xa")
-  apply (force intro!: seq_take_lemma [THEN iffD2, THEN spec])
-  apply (auto dest!: divide_Seq3)
-  done
-
-
-
-lemma take_lemma_in_eq_out:
-"!! Q. [| A UU  ==> (f UU) = (g UU) ;
-          A nil ==> (f nil) = (g nil) ;
-          !! s y n. [| ! t. A t --> seq_take n$(f t) = seq_take n$(g t);
-                     A (y\<leadsto>s) |]
-                     ==>   seq_take (Suc n)$(f (y\<leadsto>s))
-                         = seq_take (Suc n)$(g (y\<leadsto>s)) |]
-               ==> A x --> (f x)=(g x)"
-  apply (rule impI)
-  apply (rule seq.take_lemma)
-  apply (rule mp)
-  prefer 2 apply assumption
-  apply (rule_tac x="x" in spec)
-  apply (rule nat.induct)
-  apply simp
-  apply (rule allI)
-  apply (rule_tac x="xa" in Seq_cases)
-  apply simp_all
-  done
-
-
-(* ------------------------------------------------------------------------------------ *)
-
-subsection "alternative take_lemma proofs"
-
-
-(* --------------------------------------------------------------- *)
-(*              Alternative Proof of FilterPQ                      *)
-(* --------------------------------------------------------------- *)
-
-declare FilterPQ [simp del]
-
-
-(* In general: How to do this case without the same adm problems
-   as for the entire proof ? *)
-lemma Filter_lemma1: "Forall (%x. ~(P x & Q x)) s
-          --> Filter P$(Filter Q$s) =
-              Filter (%x. P x & Q x)$s"
-
-  apply (rule_tac x="s" in Seq_induct)
-  apply (simp add: Forall_def sforall_def)
-  apply simp_all
-  done
-
-lemma Filter_lemma2: "Finite s ==>
-          (Forall (%x. (~P x) | (~ Q x)) s
-          --> Filter P$(Filter Q$s) = nil)"
-  apply (erule Seq_Finite_ind, simp_all)
-  done
-
-lemma Filter_lemma3: "Finite s ==>
-          Forall (%x. (~P x) | (~ Q x)) s
-          --> Filter (%x. P x & Q x)$s = nil"
-  apply (erule Seq_Finite_ind, simp_all)
-  done
-
-
-lemma FilterPQ_takelemma: "Filter P$(Filter Q$s) = Filter (%x. P x & Q x)$s"
-  apply (rule_tac A1="%x. True" and
-                   Q1="%x. ~(P x & Q x)" and x1="s" in
-                   take_lemma_induct [THEN mp])
-  (* better support for A = %x. True *)
-  apply (simp add: Filter_lemma1)
-  apply (simp add: Filter_lemma2 Filter_lemma3)
-  apply simp
-  done
-
-declare FilterPQ [simp]
-
-
-(* --------------------------------------------------------------- *)
-(*              Alternative Proof of MapConc                       *)
-(* --------------------------------------------------------------- *)
-
-
-
-lemma MapConc_takelemma: "Map f$(x@@y) = (Map f$x) @@ (Map f$y)"
-  apply (rule_tac A1="%x. True" and x1="x" in
-      take_lemma_in_eq_out [THEN mp])
-  apply auto
-  done
-
-
-ML \<open>
-
-fun Seq_case_tac ctxt s i =
-  Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_cases} i
-  THEN hyp_subst_tac ctxt i THEN hyp_subst_tac ctxt (i+1) THEN hyp_subst_tac ctxt (i+2);
-
-(* on a\<leadsto>s only simp_tac, as full_simp_tac is uncomplete and often causes errors *)
-fun Seq_case_simp_tac ctxt s i =
-  Seq_case_tac ctxt s i
-  THEN asm_simp_tac ctxt (i+2)
-  THEN asm_full_simp_tac ctxt (i+1)
-  THEN asm_full_simp_tac ctxt i;
-
-(* rws are definitions to be unfolded for admissibility check *)
-fun Seq_induct_tac ctxt s rws i =
-  Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_induct} i
-  THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ctxt (i+1))))
-  THEN simp_tac (ctxt addsimps rws) i;
-
-fun Seq_Finite_induct_tac ctxt i =
-  eresolve_tac ctxt @{thms Seq_Finite_ind} i
-  THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ctxt i)));
-
-fun pair_tac ctxt s =
-  Rule_Insts.res_inst_tac ctxt [((("y", 0), Position.none), s)] [] @{thm prod.exhaust}
-  THEN' hyp_subst_tac ctxt THEN' asm_full_simp_tac ctxt;
-
-(* induction on a sequence of pairs with pairsplitting and simplification *)
-fun pair_induct_tac ctxt s rws i =
-  Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_induct} i
-  THEN pair_tac ctxt "a" (i+3)
-  THEN (REPEAT_DETERM (CHANGED (simp_tac ctxt (i+1))))
-  THEN simp_tac (ctxt addsimps rws) i;
-\<close>
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/ShortExecutions.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,278 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/ShortExecutions.thy
-    Author:     Olaf Müller
-*)
-
-theory ShortExecutions
-imports Traces
-begin
-
-text \<open>
-  Some properties about \<open>Cut ex\<close>, defined as follows:
-
-  For every execution ex there is another shorter execution \<open>Cut ex\<close>
-  that has the same trace as ex, but its schedule ends with an external action.
-\<close>
-
-definition
-  oraclebuild :: "('a => bool) => 'a Seq -> 'a Seq -> 'a Seq" where
-  "oraclebuild P = (fix$(LAM h s t. case t of
-       nil => nil
-    | x##xs =>
-      (case x of
-        UU => UU
-      | Def y => (Takewhile (%x. \<not>P x)$s)
-                  @@ (y\<leadsto>(h$(TL$(Dropwhile (%x. \<not> P x)$s))$xs))
-      )
-    ))"
-
-definition
-  Cut :: "('a => bool) => 'a Seq => 'a Seq" where
-  "Cut P s = oraclebuild P$s$(Filter P$s)"
-
-definition
-  LastActExtsch :: "('a,'s)ioa => 'a Seq => bool" where
-  "LastActExtsch A sch = (Cut (%x. x: ext A) sch = sch)"
-
-(* LastActExtex      ::"('a,'s)ioa => ('a,'s) pairs  => bool"*)
-(* LastActExtex_def:
-  "LastActExtex A ex == LastActExtsch A (filter_act$ex)" *)
-
-axiomatization where
-  Cut_prefixcl_Finite: "Finite s ==> (? y. s = Cut P s @@ y)"
-
-axiomatization where
-  LastActExtsmall1: "LastActExtsch A sch ==> LastActExtsch A (TL$(Dropwhile P$sch))"
-
-axiomatization where
-  LastActExtsmall2: "[| Finite sch1; LastActExtsch A (sch1 @@ sch2) |] ==> LastActExtsch A sch2"
-
-
-ML \<open>
-fun thin_tac' ctxt j =
-  rotate_tac (j - 1) THEN'
-  eresolve_tac ctxt [thin_rl] THEN'
-  rotate_tac (~ (j - 1))
-\<close>
-
-
-subsection "oraclebuild rewrite rules"
-
-
-lemma oraclebuild_unfold:
-"oraclebuild P = (LAM s t. case t of
-       nil => nil
-    | x##xs =>
-      (case x of
-        UU => UU
-      | Def y => (Takewhile (%a. \<not> P a)$s)
-                  @@ (y\<leadsto>(oraclebuild P$(TL$(Dropwhile (%a. \<not> P a)$s))$xs))
-      )
-    )"
-apply (rule trans)
-apply (rule fix_eq2)
-apply (simp only: oraclebuild_def)
-apply (rule beta_cfun)
-apply simp
-done
-
-lemma oraclebuild_UU: "oraclebuild P$sch$UU = UU"
-apply (subst oraclebuild_unfold)
-apply simp
-done
-
-lemma oraclebuild_nil: "oraclebuild P$sch$nil = nil"
-apply (subst oraclebuild_unfold)
-apply simp
-done
-
-lemma oraclebuild_cons: "oraclebuild P$s$(x\<leadsto>t) =
-          (Takewhile (%a. \<not> P a)$s)
-           @@ (x\<leadsto>(oraclebuild P$(TL$(Dropwhile (%a. \<not> P a)$s))$t))"
-apply (rule trans)
-apply (subst oraclebuild_unfold)
-apply (simp add: Consq_def)
-apply (simp add: Consq_def)
-done
-
-
-subsection "Cut rewrite rules"
-
-lemma Cut_nil:
-"[| Forall (%a. \<not> P a) s; Finite s|]
-            ==> Cut P s =nil"
-apply (unfold Cut_def)
-apply (subgoal_tac "Filter P$s = nil")
-apply (simp (no_asm_simp) add: oraclebuild_nil)
-apply (rule ForallQFilterPnil)
-apply assumption+
-done
-
-lemma Cut_UU:
-"[| Forall (%a. \<not> P a) s; ~Finite s|]
-            ==> Cut P s =UU"
-apply (unfold Cut_def)
-apply (subgoal_tac "Filter P$s= UU")
-apply (simp (no_asm_simp) add: oraclebuild_UU)
-apply (rule ForallQFilterPUU)
-apply assumption+
-done
-
-lemma Cut_Cons:
-"[| P t;  Forall (%x. \<not> P x) ss; Finite ss|]
-            ==> Cut P (ss @@ (t\<leadsto> rs))
-                 = ss @@ (t \<leadsto> Cut P rs)"
-apply (unfold Cut_def)
-apply (simp add: ForallQFilterPnil oraclebuild_cons TakewhileConc DropwhileConc)
-done
-
-
-subsection "Cut lemmas for main theorem"
-
-lemma FilterCut: "Filter P$s = Filter P$(Cut P s)"
-apply (rule_tac A1 = "%x. True" and Q1 = "%x. \<not> P x" and x1 = "s" in take_lemma_induct [THEN mp])
-prefer 3 apply (fast)
-apply (case_tac "Finite s")
-apply (simp add: Cut_nil ForallQFilterPnil)
-apply (simp add: Cut_UU ForallQFilterPUU)
-(* main case *)
-apply (simp add: Cut_Cons ForallQFilterPnil)
-done
-
-
-lemma Cut_idemp: "Cut P (Cut P s) = (Cut P s)"
-apply (rule_tac A1 = "%x. True" and Q1 = "%x. \<not> P x" and x1 = "s" in
-  take_lemma_less_induct [THEN mp])
-prefer 3 apply (fast)
-apply (case_tac "Finite s")
-apply (simp add: Cut_nil ForallQFilterPnil)
-apply (simp add: Cut_UU ForallQFilterPUU)
-(* main case *)
-apply (simp add: Cut_Cons ForallQFilterPnil)
-apply (rule take_reduction)
-apply auto
-done
-
-
-lemma MapCut: "Map f$(Cut (P o f) s) = Cut P (Map f$s)"
-apply (rule_tac A1 = "%x. True" and Q1 = "%x. \<not> P (f x) " and x1 = "s" in
-  take_lemma_less_induct [THEN mp])
-prefer 3 apply (fast)
-apply (case_tac "Finite s")
-apply (simp add: Cut_nil)
-apply (rule Cut_nil [symmetric])
-apply (simp add: ForallMap o_def)
-apply (simp add: Map2Finite)
-(* csae ~ Finite s *)
-apply (simp add: Cut_UU)
-apply (rule Cut_UU)
-apply (simp add: ForallMap o_def)
-apply (simp add: Map2Finite)
-(* main case *)
-apply (simp add: Cut_Cons MapConc ForallMap FiniteMap1 o_def)
-apply (rule take_reduction)
-apply auto
-done
-
-
-lemma Cut_prefixcl_nFinite [rule_format (no_asm)]: "~Finite s --> Cut P s << s"
-apply (intro strip)
-apply (rule take_lemma_less [THEN iffD1])
-apply (intro strip)
-apply (rule mp)
-prefer 2 apply (assumption)
-apply (tactic "thin_tac' @{context} 1 1")
-apply (rule_tac x = "s" in spec)
-apply (rule nat_less_induct)
-apply (intro strip)
-apply (rename_tac na n s)
-apply (case_tac "Forall (%x. ~ P x) s")
-apply (rule take_lemma_less [THEN iffD2, THEN spec])
-apply (simp add: Cut_UU)
-(* main case *)
-apply (drule divide_Seq3)
-apply (erule exE)+
-apply (erule conjE)+
-apply hypsubst
-apply (simp add: Cut_Cons)
-apply (rule take_reduction_less)
-(* auto makes also reasoning about Finiteness of parts of s ! *)
-apply auto
-done
-
-
-lemma execThruCut: "!!ex .is_exec_frag A (s,ex) ==> is_exec_frag A (s,Cut P ex)"
-apply (case_tac "Finite ex")
-apply (cut_tac s = "ex" and P = "P" in Cut_prefixcl_Finite)
-apply assumption
-apply (erule exE)
-apply (rule exec_prefix2closed)
-apply (erule_tac s = "ex" and t = "Cut P ex @@ y" in subst)
-apply assumption
-apply (erule exec_prefixclosed)
-apply (erule Cut_prefixcl_nFinite)
-done
-
-
-subsection "Main Cut Theorem"
-
-lemma exists_LastActExtsch:
- "[|sch : schedules A ; tr = Filter (%a. a:ext A)$sch|]
-    ==> ? sch. sch : schedules A &
-               tr = Filter (%a. a:ext A)$sch &
-               LastActExtsch A sch"
-
-apply (unfold schedules_def has_schedule_def [abs_def])
-apply auto
-apply (rule_tac x = "filter_act$ (Cut (%a. fst a:ext A) (snd ex))" in exI)
-apply (simp add: executions_def)
-apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-apply auto
-apply (rule_tac x = " (x1,Cut (%a. fst a:ext A) x2) " in exI)
-apply (simp (no_asm_simp))
-
-(* Subgoal 1: Lemma:  propagation of execution through Cut *)
-
-apply (simp add: execThruCut)
-
-(* Subgoal 2:  Lemma:  Filter P s = Filter P (Cut P s) *)
-
-apply (simp (no_asm) add: filter_act_def)
-apply (subgoal_tac "Map fst$ (Cut (%a. fst a: ext A) x2) = Cut (%a. a:ext A) (Map fst$x2) ")
-
-apply (rule_tac [2] MapCut [unfolded o_def])
-apply (simp add: FilterCut [symmetric])
-
-(* Subgoal 3: Lemma: Cut idempotent  *)
-
-apply (simp (no_asm) add: LastActExtsch_def filter_act_def)
-apply (subgoal_tac "Map fst$ (Cut (%a. fst a: ext A) x2) = Cut (%a. a:ext A) (Map fst$x2) ")
-apply (rule_tac [2] MapCut [unfolded o_def])
-apply (simp add: Cut_idemp)
-done
-
-
-subsection "Further Cut lemmas"
-
-lemma LastActExtimplnil:
-  "[| LastActExtsch A sch; Filter (%x. x:ext A)$sch = nil |]
-    ==> sch=nil"
-apply (unfold LastActExtsch_def)
-apply (drule FilternPnilForallP)
-apply (erule conjE)
-apply (drule Cut_nil)
-apply assumption
-apply simp
-done
-
-lemma LastActExtimplUU:
-  "[| LastActExtsch A sch; Filter (%x. x:ext A)$sch = UU |]
-    ==> sch=UU"
-apply (unfold LastActExtsch_def)
-apply (drule FilternPUUForallP)
-apply (erule conjE)
-apply (drule Cut_UU)
-apply assumption
-apply simp
-done
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/SimCorrectness.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,291 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/SimCorrectness.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Correctness of Simulations in HOLCF/IOA\<close>
-
-theory SimCorrectness
-imports Simulations
-begin
-
-definition
-  (* Note: s2 instead of s1 in last argument type !! *)
-  corresp_ex_simC :: "('a,'s2)ioa => (('s1 * 's2)set) => ('a,'s1)pairs
-                   -> ('s2 => ('a,'s2)pairs)" where
-  "corresp_ex_simC A R = (fix$(LAM h ex. (%s. case ex of
-      nil =>  nil
-    | x##xs => (flift1 (%pr. let a = (fst pr); t = (snd pr);
-                                 T' = @t'. ? ex1. (t,t'):R & move A ex1 s a t'
-                             in
-                                (@cex. move A cex s a T')
-                                 @@ ((h$xs) T'))
-                        $x) )))"
-
-definition
-  corresp_ex_sim :: "('a,'s2)ioa => (('s1 *'s2)set) =>
-                      ('a,'s1)execution => ('a,'s2)execution" where
-  "corresp_ex_sim A R ex == let S'= (@s'.(fst ex,s'):R & s': starts_of A)
-                            in
-                               (S',(corresp_ex_simC A R$(snd ex)) S')"
-
-
-subsection "corresp_ex_sim"
-
-lemma corresp_ex_simC_unfold: "corresp_ex_simC A R  = (LAM ex. (%s. case ex of
-       nil =>  nil
-     | x##xs => (flift1 (%pr. let a = (fst pr); t = (snd pr);
-                                  T' = @t'. ? ex1. (t,t'):R & move A ex1 s a t'
-                              in
-                                 (@cex. move A cex s a T')
-                               @@ ((corresp_ex_simC A R $xs) T'))
-                         $x) ))"
-apply (rule trans)
-apply (rule fix_eq2)
-apply (simp only: corresp_ex_simC_def)
-apply (rule beta_cfun)
-apply (simp add: flift1_def)
-done
-
-lemma corresp_ex_simC_UU: "(corresp_ex_simC A R$UU) s=UU"
-apply (subst corresp_ex_simC_unfold)
-apply simp
-done
-
-lemma corresp_ex_simC_nil: "(corresp_ex_simC A R$nil) s = nil"
-apply (subst corresp_ex_simC_unfold)
-apply simp
-done
-
-lemma corresp_ex_simC_cons: "(corresp_ex_simC A R$((a,t)\<leadsto>xs)) s =
-           (let T' = @t'. ? ex1. (t,t'):R & move A ex1 s a t'
-            in
-             (@cex. move A cex s a T')
-              @@ ((corresp_ex_simC A R$xs) T'))"
-apply (rule trans)
-apply (subst corresp_ex_simC_unfold)
-apply (simp add: Consq_def flift1_def)
-apply simp
-done
-
-
-declare corresp_ex_simC_UU [simp] corresp_ex_simC_nil [simp] corresp_ex_simC_cons [simp]
-
-
-subsection "properties of move"
-
-declare Let_def [simp del]
-
-lemma move_is_move_sim:
-   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
-      let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
-      (t,T'): R & move A (@ex2. move A ex2 s' a T') s' a T'"
-apply (unfold is_simulation_def)
-
-(* Does not perform conditional rewriting on assumptions automatically as
-   usual. Instantiate all variables per hand. Ask Tobias?? *)
-apply (subgoal_tac "? t' ex. (t,t') :R & move A ex s' a t'")
-prefer 2
-apply simp
-apply (erule conjE)
-apply (erule_tac x = "s" in allE)
-apply (erule_tac x = "s'" in allE)
-apply (erule_tac x = "t" in allE)
-apply (erule_tac x = "a" in allE)
-apply simp
-(* Go on as usual *)
-apply (erule exE)
-apply (drule_tac x = "t'" and P = "%t'. ? ex. (t,t') :R & move A ex s' a t'" in someI)
-apply (erule exE)
-apply (erule conjE)
-apply (simp add: Let_def)
-apply (rule_tac x = "ex" in someI)
-apply assumption
-done
-
-declare Let_def [simp]
-
-lemma move_subprop1_sim:
-   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
-    let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
-     is_exec_frag A (s',@x. move A x s' a T')"
-apply (cut_tac move_is_move_sim)
-defer
-apply assumption+
-apply (simp add: move_def)
-done
-
-lemma move_subprop2_sim:
-   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
-    let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
-    Finite (@x. move A x s' a T')"
-apply (cut_tac move_is_move_sim)
-defer
-apply assumption+
-apply (simp add: move_def)
-done
-
-lemma move_subprop3_sim:
-   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
-    let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
-     laststate (s',@x. move A x s' a T') = T'"
-apply (cut_tac move_is_move_sim)
-defer
-apply assumption+
-apply (simp add: move_def)
-done
-
-lemma move_subprop4_sim:
-   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
-    let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
-      mk_trace A$((@x. move A x s' a T')) =
-        (if a:ext A then a\<leadsto>nil else nil)"
-apply (cut_tac move_is_move_sim)
-defer
-apply assumption+
-apply (simp add: move_def)
-done
-
-lemma move_subprop5_sim:
-   "[|is_simulation R C A; reachable C s; s \<midarrow>a\<midarrow>C\<rightarrow> t; (s,s'):R|] ==>
-    let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in
-      (t,T'):R"
-apply (cut_tac move_is_move_sim)
-defer
-apply assumption+
-apply (simp add: move_def)
-done
-
-
-subsection \<open>TRACE INCLUSION Part 1: Traces coincide\<close>
-
-subsubsection "Lemmata for <=="
-
-(* ------------------------------------------------------
-                 Lemma 1 :Traces coincide
-   ------------------------------------------------------- *)
-
-declare split_if [split del]
-lemma traces_coincide_sim [rule_format (no_asm)]:
-  "[|is_simulation R C A; ext C = ext A|] ==>
-         !s s'. reachable C s & is_exec_frag C (s,ex) & (s,s'): R -->
-             mk_trace C$ex = mk_trace A$((corresp_ex_simC A R$ex) s')"
-
-apply (tactic \<open>pair_induct_tac @{context} "ex" [@{thm is_exec_frag_def}] 1\<close>)
-(* cons case *)
-apply auto
-apply (rename_tac ex a t s s')
-apply (simp add: mk_traceConc)
-apply (frule reachable.reachable_n)
-apply assumption
-apply (erule_tac x = "t" in allE)
-apply (erule_tac x = "@t'. ? ex1. (t,t') :R & move A ex1 s' a t'" in allE)
-apply (simp add: move_subprop5_sim [unfolded Let_def]
-  move_subprop4_sim [unfolded Let_def] split add: split_if)
-done
-declare split_if [split]
-
-
-(* ----------------------------------------------------------- *)
-(*               Lemma 2 : corresp_ex_sim is execution         *)
-(* ----------------------------------------------------------- *)
-
-
-lemma correspsim_is_execution [rule_format (no_asm)]:
- "[| is_simulation R C A |] ==>
-  !s s'. reachable C s & is_exec_frag C (s,ex) & (s,s'):R
-  --> is_exec_frag A (s',(corresp_ex_simC A R$ex) s')"
-
-apply (tactic \<open>pair_induct_tac @{context} "ex" [@{thm is_exec_frag_def}] 1\<close>)
-(* main case *)
-apply auto
-apply (rename_tac ex a t s s')
-apply (rule_tac t = "@t'. ? ex1. (t,t') :R & move A ex1 s' a t'" in lemma_2_1)
-
-(* Finite *)
-apply (erule move_subprop2_sim [unfolded Let_def])
-apply assumption+
-apply (rule conjI)
-
-(* is_exec_frag *)
-apply (erule move_subprop1_sim [unfolded Let_def])
-apply assumption+
-apply (rule conjI)
-
-(* Induction hypothesis  *)
-(* reachable_n looping, therefore apply it manually *)
-apply (erule_tac x = "t" in allE)
-apply (erule_tac x = "@t'. ? ex1. (t,t') :R & move A ex1 s' a t'" in allE)
-apply simp
-apply (frule reachable.reachable_n)
-apply assumption
-apply (simp add: move_subprop5_sim [unfolded Let_def])
-(* laststate *)
-apply (erule move_subprop3_sim [unfolded Let_def, symmetric])
-apply assumption+
-done
-
-
-subsection "Main Theorem: TRACE - INCLUSION"
-
-(* -------------------------------------------------------------------------------- *)
-
-  (* generate condition (s,S'):R & S':starts_of A, the first being intereting
-     for the induction cases concerning the two lemmas correpsim_is_execution and
-     traces_coincide_sim, the second for the start state case.
-     S':= @s'. (s,s'):R & s':starts_of A, where s:starts_of C  *)
-
-lemma simulation_starts:
-"[| is_simulation R C A; s:starts_of C |]
-  ==> let S' = @s'. (s,s'):R & s':starts_of A in
-      (s,S'):R & S':starts_of A"
-  apply (simp add: is_simulation_def corresp_ex_sim_def Int_non_empty Image_def)
-  apply (erule conjE)+
-  apply (erule ballE)
-  prefer 2 apply (blast)
-  apply (erule exE)
-  apply (rule someI2)
-  apply assumption
-  apply blast
-  done
-
-lemmas sim_starts1 = simulation_starts [unfolded Let_def, THEN conjunct1]
-lemmas sim_starts2 = simulation_starts [unfolded Let_def, THEN conjunct2]
-
-
-lemma trace_inclusion_for_simulations:
-  "[| ext C = ext A; is_simulation R C A |]
-           ==> traces C <= traces A"
-
-  apply (unfold traces_def)
-
-  apply (simp (no_asm) add: has_trace_def2)
-  apply auto
-
-  (* give execution of abstract automata *)
-  apply (rule_tac x = "corresp_ex_sim A R ex" in bexI)
-
-  (* Traces coincide, Lemma 1 *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (rename_tac s ex)
-  apply (simp (no_asm) add: corresp_ex_sim_def)
-  apply (rule_tac s = "s" in traces_coincide_sim)
-  apply assumption+
-  apply (simp add: executions_def reachable.reachable_0 sim_starts1)
-
-  (* corresp_ex_sim is execution, Lemma 2 *)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (simp add: executions_def)
-  apply (rename_tac s ex)
-
-  (* start state *)
-  apply (rule conjI)
-  apply (simp add: sim_starts2 corresp_ex_sim_def)
-
-  (* is-execution-fragment *)
-  apply (simp add: corresp_ex_sim_def)
-  apply (rule_tac s = s in correspsim_is_execution)
-  apply assumption
-  apply (simp add: reachable.reachable_0 sim_starts1)
-  done
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/Simulations.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,85 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/Simulations.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Simulations in HOLCF/IOA\<close>
-
-theory Simulations
-imports RefCorrectness
-begin
-
-default_sort type
-
-definition
-  is_simulation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
-  "is_simulation R C A =
-   ((!s:starts_of C. R``{s} Int starts_of A ~= {}) &
-   (!s s' t a. reachable C s &
-               s \<midarrow>a\<midarrow>C\<rightarrow> t   &
-               (s,s') : R
-               --> (? t' ex. (t,t'):R & move A ex s' a t')))"
-
-definition
-  is_backward_simulation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
-  "is_backward_simulation R C A =
-   ((!s:starts_of C. R``{s} <= starts_of A) &
-   (!s t t' a. reachable C s &
-               s \<midarrow>a\<midarrow>C\<rightarrow> t   &
-               (t,t') : R
-               --> (? ex s'. (s,s'):R & move A ex s' a t')))"
-
-definition
-  is_forw_back_simulation :: "[('s1 * 's2 set)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
-  "is_forw_back_simulation R C A =
-   ((!s:starts_of C. ? S'. (s,S'):R & S'<= starts_of A) &
-   (!s S' t a. reachable C s &
-               s \<midarrow>a\<midarrow>C\<rightarrow> t   &
-               (s,S') : R
-               --> (? T'. (t,T'):R & (! t':T'. ? s':S'. ? ex. move A ex s' a t'))))"
-
-definition
-  is_back_forw_simulation :: "[('s1 * 's2 set)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
-  "is_back_forw_simulation R C A =
-   ((!s:starts_of C. ! S'. (s,S'):R --> S' Int starts_of A ~={}) &
-   (!s t T' a. reachable C s &
-               s \<midarrow>a\<midarrow>C\<rightarrow> t   &
-               (t,T') : R
-               --> (? S'. (s,S'):R & (! s':S'. ? t':T'. ? ex. move A ex s' a t'))))"
-
-definition
-  is_history_relation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
-  "is_history_relation R C A = (is_simulation R C A &
-                                is_ref_map (%x.(@y. (x,y):(R^-1))) A C)"
-
-definition
-  is_prophecy_relation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
-  "is_prophecy_relation R C A = (is_backward_simulation R C A &
-                                 is_ref_map (%x.(@y. (x,y):(R^-1))) A C)"
-
-
-lemma set_non_empty: "(A~={}) = (? x. x:A)"
-apply auto
-done
-
-lemma Int_non_empty: "(A Int B ~= {}) = (? x. x: A & x:B)"
-apply (simp add: set_non_empty)
-done
-
-
-lemma Sim_start_convert:
-"(R``{x} Int S ~= {}) = (? y. (x,y):R & y:S)"
-apply (unfold Image_def)
-apply (simp add: Int_non_empty)
-done
-
-declare Sim_start_convert [simp]
-
-
-lemma ref_map_is_simulation:
-"!! f. is_ref_map f C A ==> is_simulation {p. (snd p) = f (fst p)} C A"
-
-apply (unfold is_ref_map_def is_simulation_def)
-apply simp
-done
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/TL.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,179 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/TL.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>A General Temporal Logic\<close>
-
-theory TL
-imports Pred Sequence
-begin
-
-default_sort type
-
-type_synonym 'a temporal = "'a Seq predicate"
-
-definition validT :: "'a Seq predicate \<Rightarrow> bool"
-  where "validT P \<longleftrightarrow> (\<forall>s. s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> (s \<Turnstile> P))"
-
-definition unlift :: "'a lift \<Rightarrow> 'a"
-  where "unlift x = (case x of Def y \<Rightarrow> y)"
-
-definition Init :: "'a predicate \<Rightarrow> 'a temporal"  ("\<langle>_\<rangle>" [0] 1000)
-  where "Init P s = P (unlift (HD $ s))"
-    \<comment> \<open>this means that for \<open>nil\<close> and \<open>UU\<close> the effect is unpredictable\<close>
-
-definition Next :: "'a temporal \<Rightarrow> 'a temporal"
-  where "(Next P) s \<longleftrightarrow> (if TL $ s = UU \<or> TL $ s = nil then P s else P (TL $ s))"
-
-definition suffix :: "'a Seq \<Rightarrow> 'a Seq \<Rightarrow> bool"
-  where "suffix s2 s \<longleftrightarrow> (\<exists>s1. Finite s1 \<and> s = s1 @@ s2)"
-
-definition tsuffix :: "'a Seq \<Rightarrow> 'a Seq \<Rightarrow> bool"
-  where "tsuffix s2 s \<longleftrightarrow> s2 \<noteq> nil \<and> s2 \<noteq> UU \<and> suffix s2 s"
-
-definition Box :: "'a temporal \<Rightarrow> 'a temporal"  ("\<box>(_)" [80] 80)
-  where "(\<box>P) s \<longleftrightarrow> (\<forall>s2. tsuffix s2 s \<longrightarrow> P s2)"
-
-definition Diamond :: "'a temporal \<Rightarrow> 'a temporal"  ("\<diamond>(_)" [80] 80)
-  where "\<diamond>P = (\<^bold>\<not> (\<box>(\<^bold>\<not> P)))"
-
-definition Leadsto :: "'a temporal \<Rightarrow> 'a temporal \<Rightarrow> 'a temporal"  (infixr "\<leadsto>" 22)
-  where "(P \<leadsto> Q) = (\<box>(P \<^bold>\<longrightarrow> (\<diamond>Q)))"
-
-
-lemma simple: "\<box>\<diamond>(\<^bold>\<not> P) = (\<^bold>\<not> \<diamond>\<box>P)"
-apply (rule ext)
-apply (simp add: Diamond_def NOT_def Box_def)
-done
-
-lemma Boxnil: "nil \<Turnstile> \<box>P"
-apply (simp add: satisfies_def Box_def tsuffix_def suffix_def nil_is_Conc)
-done
-
-lemma Diamondnil: "~(nil \<Turnstile> \<diamond>P)"
-apply (simp add: Diamond_def satisfies_def NOT_def)
-apply (cut_tac Boxnil)
-apply (simp add: satisfies_def)
-done
-
-lemma Diamond_def2: "(\<diamond>F) s = (? s2. tsuffix s2 s & F s2)"
-apply (simp add: Diamond_def NOT_def Box_def)
-done
-
-
-
-subsection "TLA Axiomatization by Merz"
-
-lemma suffix_refl: "suffix s s"
-apply (simp add: suffix_def)
-apply (rule_tac x = "nil" in exI)
-apply auto
-done
-
-lemma reflT: "s~=UU & s~=nil --> (s \<Turnstile> \<box>F \<^bold>\<longrightarrow> F)"
-apply (simp add: satisfies_def IMPLIES_def Box_def)
-apply (rule impI)+
-apply (erule_tac x = "s" in allE)
-apply (simp add: tsuffix_def suffix_refl)
-done
-
-
-lemma suffix_trans: "[| suffix y x ; suffix z y |]  ==> suffix z x"
-apply (simp add: suffix_def)
-apply auto
-apply (rule_tac x = "s1 @@ s1a" in exI)
-apply auto
-apply (simp (no_asm) add: Conc_assoc)
-done
-
-lemma transT: "s \<Turnstile> \<box>F \<^bold>\<longrightarrow> \<box>\<box>F"
-apply (simp (no_asm) add: satisfies_def IMPLIES_def Box_def tsuffix_def)
-apply auto
-apply (drule suffix_trans)
-apply assumption
-apply (erule_tac x = "s2a" in allE)
-apply auto
-done
-
-
-lemma normalT: "s \<Turnstile> \<box>(F \<^bold>\<longrightarrow> G) \<^bold>\<longrightarrow> \<box>F \<^bold>\<longrightarrow> \<box>G"
-apply (simp (no_asm) add: satisfies_def IMPLIES_def Box_def)
-done
-
-
-subsection "TLA Rules by Lamport"
-
-lemma STL1a: "validT P ==> validT (\<box>P)"
-apply (simp add: validT_def satisfies_def Box_def tsuffix_def)
-done
-
-lemma STL1b: "valid P ==> validT (Init P)"
-apply (simp add: valid_def validT_def satisfies_def Init_def)
-done
-
-lemma STL1: "valid P ==> validT (\<box>(Init P))"
-apply (rule STL1a)
-apply (erule STL1b)
-done
-
-(* Note that unlift and HD is not at all used !!! *)
-lemma STL4: "valid (P \<^bold>\<longrightarrow> Q)  ==> validT (\<box>(Init P) \<^bold>\<longrightarrow> \<box>(Init Q))"
-apply (simp add: valid_def validT_def satisfies_def IMPLIES_def Box_def Init_def)
-done
-
-
-subsection "LTL Axioms by Manna/Pnueli"
-
-lemma tsuffix_TL [rule_format (no_asm)]:
-"s~=UU & s~=nil --> tsuffix s2 (TL$s) --> tsuffix s2 s"
-apply (unfold tsuffix_def suffix_def)
-apply auto
-apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
-apply (rule_tac x = "a\<leadsto>s1" in exI)
-apply auto
-done
-
-lemmas tsuffix_TL2 = conjI [THEN tsuffix_TL]
-
-declare split_if [split del]
-lemma LTL1:
-   "s~=UU & s~=nil --> (s \<Turnstile> \<box>F \<^bold>\<longrightarrow> (F \<^bold>\<and> (Next (\<box>F))))"
-apply (unfold Next_def satisfies_def NOT_def IMPLIES_def AND_def Box_def)
-apply auto
-(* \<box>F \<^bold>\<longrightarrow> F *)
-apply (erule_tac x = "s" in allE)
-apply (simp add: tsuffix_def suffix_refl)
-(* \<box>F \<^bold>\<longrightarrow> Next \<box>F *)
-apply (simp split add: split_if)
-apply auto
-apply (drule tsuffix_TL2)
-apply assumption+
-apply auto
-done
-declare split_if [split]
-
-
-lemma LTL2a:
-    "s \<Turnstile> \<^bold>\<not> (Next F) \<^bold>\<longrightarrow> (Next (\<^bold>\<not> F))"
-apply (unfold Next_def satisfies_def NOT_def IMPLIES_def)
-apply simp
-done
-
-lemma LTL2b:
-    "s \<Turnstile> (Next (\<^bold>\<not> F)) \<^bold>\<longrightarrow> (\<^bold>\<not> (Next F))"
-apply (unfold Next_def satisfies_def NOT_def IMPLIES_def)
-apply simp
-done
-
-lemma LTL3:
-"ex \<Turnstile> (Next (F \<^bold>\<longrightarrow> G)) \<^bold>\<longrightarrow> (Next F) \<^bold>\<longrightarrow> (Next G)"
-apply (unfold Next_def satisfies_def NOT_def IMPLIES_def)
-apply simp
-done
-
-
-lemma ModusPonens: "[| validT (P \<^bold>\<longrightarrow> Q); validT P |] ==> validT Q"
-apply (simp add: validT_def satisfies_def IMPLIES_def)
-done
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/TLS.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,176 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/TLS.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Temporal Logic of Steps -- tailored for I/O automata\<close>
-
-theory TLS
-imports IOA TL
-begin
-
-default_sort type
-
-type_synonym ('a, 's) ioa_temp  = "('a option, 's) transition temporal"
-
-type_synonym ('a, 's) step_pred = "('a option, 's) transition predicate"
-
-type_synonym 's state_pred = "'s predicate"
-
-definition mkfin :: "'a Seq \<Rightarrow> 'a Seq"
-  where "mkfin s = (if Partial s then SOME t. Finite t \<and> s = t @@ UU else s)"
-
-definition option_lift :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a option \<Rightarrow> 'b"
-  where "option_lift f s y = (case y of None \<Rightarrow> s | Some x \<Rightarrow> f x)"
-
-definition plift :: "('a \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> bool"
-(* plift is used to determine that None action is always false in
-   transition predicates *)
-  where "plift P = option_lift P False"
-
-definition xt1 :: "'s predicate \<Rightarrow> ('a, 's) step_pred"
-  where "xt1 P tr = P (fst tr)"
-
-definition xt2 :: "'a option predicate \<Rightarrow> ('a, 's) step_pred"
-  where "xt2 P tr = P (fst (snd tr))"
-
-definition ex2seqC :: "('a, 's) pairs \<rightarrow> ('s \<Rightarrow> ('a option, 's) transition Seq)"
-where
-  "ex2seqC = (fix$(LAM h ex. (%s. case ex of
-      nil =>  (s,None,s)\<leadsto>nil
-    | x##xs => (flift1 (%pr.
-                (s,Some (fst pr), snd pr)\<leadsto> (h$xs) (snd pr))
-                $x)
-      )))"
-
-definition ex2seq :: "('a, 's) execution \<Rightarrow> ('a option, 's) transition Seq"
-  where "ex2seq ex = (ex2seqC $ (mkfin (snd ex))) (fst ex)"
-
-definition temp_sat :: "('a, 's) execution \<Rightarrow> ('a, 's) ioa_temp \<Rightarrow> bool"  (infixr "\<TTurnstile>" 22)
-  where "(ex \<TTurnstile> P) \<longleftrightarrow> ((ex2seq ex) \<Turnstile> P)"
-
-definition validTE :: "('a, 's) ioa_temp \<Rightarrow> bool"
-  where "validTE P \<longleftrightarrow> (\<forall>ex. (ex \<TTurnstile> P))"
-
-definition validIOA :: "('a, 's) ioa \<Rightarrow> ('a, 's) ioa_temp \<Rightarrow> bool"
-  where "validIOA A P \<longleftrightarrow> (\<forall>ex \<in> executions A. (ex \<TTurnstile> P))"
-
-
-axiomatization
-where
-
-mkfin_UU:
-  "mkfin UU = nil" and
-
-mkfin_nil:
-  "mkfin nil =nil" and
-
-mkfin_cons:
-  "(mkfin (a\<leadsto>s)) = (a\<leadsto>(mkfin s))"
-
-
-lemmas [simp del] = HOL.ex_simps HOL.all_simps split_paired_Ex
-
-setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close>
-
-
-subsection \<open>ex2seqC\<close>
-
-lemma ex2seqC_unfold: "ex2seqC  = (LAM ex. (%s. case ex of
-       nil =>  (s,None,s)\<leadsto>nil
-     | x##xs => (flift1 (%pr.
-                 (s,Some (fst pr), snd pr)\<leadsto> (ex2seqC$xs) (snd pr))
-                 $x)
-       ))"
-  apply (rule trans)
-  apply (rule fix_eq4)
-  apply (rule ex2seqC_def)
-  apply (rule beta_cfun)
-  apply (simp add: flift1_def)
-  done
-
-lemma ex2seqC_UU: "(ex2seqC $UU) s=UU"
-  apply (subst ex2seqC_unfold)
-  apply simp
-  done
-
-lemma ex2seqC_nil: "(ex2seqC $nil) s = (s,None,s)\<leadsto>nil"
-  apply (subst ex2seqC_unfold)
-  apply simp
-  done
-
-lemma ex2seqC_cons: "(ex2seqC $((a,t)\<leadsto>xs)) s = (s,Some a,t)\<leadsto> ((ex2seqC$xs) t)"
-  apply (rule trans)
-  apply (subst ex2seqC_unfold)
-  apply (simp add: Consq_def flift1_def)
-  apply (simp add: Consq_def flift1_def)
-  done
-
-declare ex2seqC_UU [simp] ex2seqC_nil [simp] ex2seqC_cons [simp]
-
-
-
-declare mkfin_UU [simp] mkfin_nil [simp] mkfin_cons [simp]
-
-lemma ex2seq_UU: "ex2seq (s, UU) = (s,None,s)\<leadsto>nil"
-  by (simp add: ex2seq_def)
-
-lemma ex2seq_nil: "ex2seq (s, nil) = (s,None,s)\<leadsto>nil"
-  by (simp add: ex2seq_def)
-
-lemma ex2seq_cons: "ex2seq (s, (a,t)\<leadsto>ex) = (s,Some a,t) \<leadsto> ex2seq (t, ex)"
-  by (simp add: ex2seq_def)
-
-declare ex2seqC_UU [simp del] ex2seqC_nil [simp del] ex2seqC_cons [simp del]
-declare ex2seq_UU [simp] ex2seq_nil [simp] ex2seq_cons [simp]
-
-
-lemma ex2seq_nUUnnil: "ex2seq exec ~= UU & ex2seq exec ~= nil"
-  apply (tactic \<open>pair_tac @{context} "exec" 1\<close>)
-  apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
-  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-  done
-
-
-subsection \<open>Interface TL -- TLS\<close>
-
-(* uses the fact that in executions states overlap, which is lost in
-   after the translation via ex2seq !! *)
-
-lemma TL_TLS:
-  "[| ! s a t. (P s) & s \<midarrow>a\<midarrow>A\<rightarrow> t --> (Q t) |]
-    ==> ex \<TTurnstile> (Init (%(s,a,t). P s) \<^bold>\<and> Init (%(s,a,t). s \<midarrow>a\<midarrow>A\<rightarrow> t)
-              \<^bold>\<longrightarrow> (Next (Init (%(s,a,t).Q s))))"
-  apply (unfold Init_def Next_def temp_sat_def satisfies_def IMPLIES_def AND_def)
-
-  apply clarify
-  apply (simp split add: split_if)
-  (* TL = UU *)
-  apply (rule conjI)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
-  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-  apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
-  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-  (* TL = nil *)
-  apply (rule conjI)
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (tactic \<open>Seq_case_tac @{context} "x2" 1\<close>)
-  apply (simp add: unlift_def)
-  apply fast
-  apply (simp add: unlift_def)
-  apply fast
-  apply (simp add: unlift_def)
-  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-  apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
-  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-  (* TL =cons *)
-  apply (simp add: unlift_def)
-
-  apply (tactic \<open>pair_tac @{context} "ex" 1\<close>)
-  apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>)
-  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-  apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
-  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-  done
-
-end
--- a/src/HOL/HOLCF/IOA/meta_theory/Traces.thy	Thu Dec 31 12:37:16 2015 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,356 +0,0 @@
-(*  Title:      HOL/HOLCF/IOA/meta_theory/Traces.thy
-    Author:     Olaf Müller
-*)
-
-section \<open>Executions and Traces of I/O automata in HOLCF\<close>
-
-theory Traces
-imports Sequence Automata
-begin
-
-default_sort type
-
-type_synonym ('a, 's) pairs = "('a * 's) Seq"
-type_synonym ('a, 's) execution = "'s * ('a, 's) pairs"
-type_synonym 'a trace = "'a Seq"
-type_synonym ('a, 's) execution_module = "('a, 's) execution set * 'a signature"
-type_synonym 'a schedule_module = "'a trace set * 'a signature"
-type_synonym 'a trace_module = "'a trace set * 'a signature"
-
-
-(*  ------------------- Executions ------------------------------ *)
-
-definition is_exec_fragC :: "('a, 's) ioa \<Rightarrow> ('a, 's) pairs \<rightarrow> 's \<Rightarrow> tr"
-where
-  "is_exec_fragC A = (fix $ (LAM h ex. (%s. case ex of
-      nil => TT
-    | x##xs => (flift1
-            (%p. Def ((s,p):trans_of A) andalso (h$xs) (snd p))
-             $x)
-   )))"
-
-definition is_exec_frag :: "[('a,'s)ioa, ('a,'s)execution] \<Rightarrow> bool"
-  where "is_exec_frag A ex = ((is_exec_fragC A $ (snd ex)) (fst ex) ~= FF)"
-
-definition executions :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution set"
-  where "executions ioa = {e. ((fst e) \<in> starts_of(ioa)) \<and> is_exec_frag ioa e}"
-
-
-(*  ------------------- Schedules ------------------------------ *)
-
-definition filter_act :: "('a, 's) pairs \<rightarrow> 'a trace"
-  where "filter_act = Map fst"
-
-definition has_schedule :: "[('a, 's) ioa, 'a trace] \<Rightarrow> bool"
-  where "has_schedule ioa sch \<longleftrightarrow> (\<exists>ex \<in> executions ioa. sch = filter_act $ (snd ex))"
-
-definition schedules :: "('a, 's) ioa \<Rightarrow> 'a trace set"
-  where "schedules ioa = {sch. has_schedule ioa sch}"
-
-
-(*  ------------------- Traces ------------------------------ *)
-
-definition has_trace :: "[('a, 's) ioa, 'a trace] \<Rightarrow> bool"
-  where "has_trace ioa tr = (\<exists>sch \<in> schedules ioa. tr = Filter (\<lambda>a. a \<in> ext ioa) $ sch)"
-
-definition traces :: "('a, 's) ioa \<Rightarrow> 'a trace set"
-  where "traces ioa \<equiv> {tr. has_trace ioa tr}"
-
-definition mk_trace :: "('a, 's) ioa \<Rightarrow> ('a, 's) pairs \<rightarrow> 'a trace"
-  where "mk_trace ioa = (LAM tr. Filter (\<lambda>a. a \<in> ext ioa) $ (filter_act $ tr))"
-
-
-(*  ------------------- Fair Traces ------------------------------ *)
-
-definition laststate :: "('a, 's) execution \<Rightarrow> 's"
-where
-  "laststate ex = (case Last $ (snd ex) of
-                    UU  => fst ex
-                  | Def at => snd at)"
-
-(* A predicate holds infinitely (finitely) often in a sequence *)
-
-definition inf_often :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<Rightarrow> bool"
-  where "inf_often P s \<longleftrightarrow> Infinite (Filter P $ s)"
-
-(*  filtering P yields a finite or partial sequence *)
-definition fin_often :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<Rightarrow> bool"
-  where "fin_often P s \<longleftrightarrow> \<not> inf_often P s"
-
-
-(* fairness of executions *)
-
-(* Note that partial execs cannot be wfair as the inf_often predicate in the
-   else branch prohibits it. However they can be sfair in the case when all W
-   are only finitely often enabled: Is this the right model?
-   See LiveIOA for solution conforming with the literature and superseding this one *)
-definition is_wfair :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> ('a, 's) execution \<Rightarrow> bool"
-where
-  "is_wfair A W ex \<longleftrightarrow>
-    (inf_often (\<lambda>x. fst x \<in> W) (snd ex) \<or> inf_often (\<lambda>x. \<not> Enabled A W (snd x)) (snd ex))"
-definition wfair_ex :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution \<Rightarrow> bool"
-where
-  "wfair_ex A ex \<longleftrightarrow> (\<forall>W \<in> wfair_of A.
-                      if Finite (snd ex)
-                      then \<not> Enabled A W (laststate ex)
-                      else is_wfair A W ex)"
-
-definition is_sfair :: "('a, 's) ioa \<Rightarrow> 'a set \<Rightarrow> ('a, 's) execution \<Rightarrow> bool"
-where
-  "is_sfair A W ex \<longleftrightarrow>
-    (inf_often (\<lambda>x. fst x:W) (snd ex) \<or> fin_often (\<lambda>x. Enabled A W (snd x)) (snd ex))"
-definition sfair_ex :: "('a, 's)ioa \<Rightarrow> ('a, 's) execution \<Rightarrow> bool"
-where
-  "sfair_ex A ex \<longleftrightarrow> (\<forall>W \<in> sfair_of A.
-                      if   Finite (snd ex)
-                      then ~Enabled A W (laststate ex)
-                      else is_sfair A W ex)"
-
-definition fair_ex :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution \<Rightarrow> bool"
-  where "fair_ex A ex \<longleftrightarrow> wfair_ex A ex \<and> sfair_ex A ex"
-
-
-(* fair behavior sets *)
-
-definition fairexecutions :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution set"
-  where "fairexecutions A = {ex. ex \<in> executions A \<and> fair_ex A ex}"
-
-definition fairtraces :: "('a, 's) ioa \<Rightarrow> 'a trace set"
-  where "fairtraces A = {mk_trace A $ (snd ex) | ex. ex \<in> fairexecutions A}"
-
-
-(*  ------------------- Implementation ------------------------------ *)
-
-(* Notions of implementation *)
-
-definition ioa_implements :: "[('a, 's1) ioa, ('a, 's2) ioa] \<Rightarrow> bool"  (infixr "=<|" 12)
-where
-  "(ioa1 =<| ioa2) \<longleftrightarrow>
-    (((inputs(asig_of(ioa1)) = inputs(asig_of(ioa2))) \<and>
-     (outputs(asig_of(ioa1)) = outputs(asig_of(ioa2)))) \<and>
-      traces(ioa1) \<subseteq> traces(ioa2))"
-
-definition fair_implements :: "('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool"
-where
-  "fair_implements C A \<longleftrightarrow> inp C = inp A \<and> out C = out A \<and> fairtraces C \<subseteq> fairtraces A"
-
-
-(*  ------------------- Modules ------------------------------ *)
-
-(* Execution, schedule and trace modules *)
-
-definition Execs :: "('a, 's) ioa \<Rightarrow> ('a, 's) execution_module"
-  where "Execs A = (executions A, asig_of A)"
-
-definition Scheds :: "('a, 's) ioa \<Rightarrow> 'a schedule_module"
-  where "Scheds A = (schedules A, asig_of A)"
-
-definition Traces :: "('a, 's) ioa \<Rightarrow> 'a trace_module"
-  where "Traces A = (traces A, asig_of A)"
-
-
-lemmas [simp del] = HOL.ex_simps HOL.all_simps split_paired_Ex
-declare Let_def [simp]
-setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close>
-
-lemmas exec_rws = executions_def is_exec_frag_def
-
-
-
-subsection "recursive equations of operators"
-
-(* ---------------------------------------------------------------- *)
-(*                               filter_act                         *)
-(* ---------------------------------------------------------------- *)
-
-
-lemma filter_act_UU: "filter_act$UU = UU"
-  by (simp add: filter_act_def)
-
-lemma filter_act_nil: "filter_act$nil = nil"
-  by (simp add: filter_act_def)
-
-lemma filter_act_cons: "filter_act$(x\<leadsto>xs) = (fst x) \<leadsto> filter_act$xs"
-  by (simp add: filter_act_def)
-
-declare filter_act_UU [simp] filter_act_nil [simp] filter_act_cons [simp]
-
-
-(* ---------------------------------------------------------------- *)
-(*                             mk_trace                             *)
-(* ---------------------------------------------------------------- *)
-
-lemma mk_trace_UU: "mk_trace A$UU=UU"
-  by (simp add: mk_trace_def)
-
-lemma mk_trace_nil: "mk_trace A$nil=nil"
-  by (simp add: mk_trace_def)
-
-lemma mk_trace_cons: "mk_trace A$(at \<leadsto> xs) =
-             (if ((fst at):ext A)
-                  then (fst at) \<leadsto> (mk_trace A$xs)
-                  else mk_trace A$xs)"
-
-  by (simp add: mk_trace_def)
-
-declare mk_trace_UU [simp] mk_trace_nil [simp] mk_trace_cons [simp]
-
-(* ---------------------------------------------------------------- *)
-(*                             is_exec_fragC                             *)
-(* ---------------------------------------------------------------- *)
-
-
-lemma is_exec_fragC_unfold: "is_exec_fragC A = (LAM ex. (%s. case ex of
-       nil => TT
-     | x##xs => (flift1
-             (%p. Def ((s,p):trans_of A) andalso (is_exec_fragC A$xs) (snd p))
-              $x)
-    ))"
-  apply (rule trans)
-  apply (rule fix_eq4)
-  apply (rule is_exec_fragC_def)
-  apply (rule beta_cfun)
-  apply (simp add: flift1_def)
-  done
-
-lemma is_exec_fragC_UU: "(is_exec_fragC A$UU) s=UU"
-  apply (subst is_exec_fragC_unfold)
-  apply simp
-  done
-
-lemma is_exec_fragC_nil: "(is_exec_fragC A$nil) s = TT"
-  apply (subst is_exec_fragC_unfold)
-  apply simp
-  done
-
-lemma is_exec_fragC_cons: "(is_exec_fragC A$(pr\<leadsto>xs)) s =
-                         (Def ((s,pr):trans_of A)
-                 andalso (is_exec_fragC A$xs)(snd pr))"
-  apply (rule trans)
-  apply (subst is_exec_fragC_unfold)
-  apply (simp add: Consq_def flift1_def)
-  apply simp
-  done
-
-
-declare is_exec_fragC_UU [simp] is_exec_fragC_nil [simp] is_exec_fragC_cons [simp]
-
-
-(* ---------------------------------------------------------------- *)
-(*                        is_exec_frag                              *)
-(* ---------------------------------------------------------------- *)
-
-lemma is_exec_frag_UU: "is_exec_frag A (s, UU)"
-  by (simp add: is_exec_frag_def)
-
-lemma is_exec_frag_nil: "is_exec_frag A (s, nil)"
-  by (simp add: is_exec_frag_def)
-
-lemma is_exec_frag_cons: "is_exec_frag A (s, (a,t)\<leadsto>ex) =
-                                (((s,a,t):trans_of A) &
-                                is_exec_frag A (t, ex))"
-  by (simp add: is_exec_frag_def)
-
-
-(* Delsimps [is_exec_fragC_UU,is_exec_fragC_nil,is_exec_fragC_cons]; *)
-declare is_exec_frag_UU [simp] is_exec_frag_nil [simp] is_exec_frag_cons [simp]
-
-(* ---------------------------------------------------------------------------- *)
-                           section "laststate"
-(* ---------------------------------------------------------------------------- *)
-
-lemma laststate_UU: "laststate (s,UU) = s"
-  by (simp add: laststate_def)
-
-lemma laststate_nil: "laststate (s,nil) = s"
-  by (simp add: laststate_def)
-
-lemma laststate_cons: "!! ex. Finite ex ==> laststate (s,at\<leadsto>ex) = laststate (snd at,ex)"
-  apply (simp (no_asm) add: laststate_def)
-  apply (case_tac "ex=nil")
-  apply (simp (no_asm_simp))
-  apply (simp (no_asm_simp))
-  apply (drule Finite_Last1 [THEN mp])
-  apply assumption
-  apply defined
-  done
-
-declare laststate_UU [simp] laststate_nil [simp] laststate_cons [simp]
-
-lemma exists_laststate: "!!ex. Finite ex ==> (! s. ? u. laststate (s,ex)=u)"
-  apply (tactic "Seq_Finite_induct_tac @{context} 1")
-  done
-
-
-subsection "has_trace, mk_trace"
-
-(* alternative definition of has_trace tailored for the refinement proof, as it does not
-   take the detour of schedules *)
-
-lemma has_trace_def2:
-  "has_trace A b = (? ex:executions A. b = mk_trace A$(snd ex))"
-  apply (unfold executions_def mk_trace_def has_trace_def schedules_def has_schedule_def [abs_def])
-  apply auto
-  done
-
-
-subsection "signatures and executions, schedules"
-
-(* All executions of A have only actions of A. This is only true because of the
-   predicate state_trans (part of the predicate IOA): We have no dependent types.
-   For executions of parallel automata this assumption is not needed, as in par_def
-   this condition is included once more. (see Lemmas 1.1.1c in CompoExecs for example) *)
-
-lemma execfrag_in_sig:
-  "!! A. is_trans_of A ==>
-  ! s. is_exec_frag A (s,xs) --> Forall (%a. a:act A) (filter_act$xs)"
-  apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def},
-    @{thm Forall_def}, @{thm sforall_def}] 1\<close>)
-  (* main case *)
-  apply (auto simp add: is_trans_of_def)
-  done
-
-lemma exec_in_sig:
-  "!! A.[|  is_trans_of A; x:executions A |] ==>
-  Forall (%a. a:act A) (filter_act$(snd x))"
-  apply (simp add: executions_def)
-  apply (tactic \<open>pair_tac @{context} "x" 1\<close>)
-  apply (rule execfrag_in_sig [THEN spec, THEN mp])
-  apply auto
-  done
-
-lemma scheds_in_sig:
-  "!! A.[|  is_trans_of A; x:schedules A |] ==>
-    Forall (%a. a:act A) x"
-  apply (unfold schedules_def has_schedule_def [abs_def])
-  apply (fast intro!: exec_in_sig)
-  done
-
-
-subsection "executions are prefix closed"
-
-(* only admissible in y, not if done in x !! *)
-lemma execfrag_prefixclosed: "!x s. is_exec_frag A (s,x) & y<<x  --> is_exec_frag A (s,y)"
-  apply (tactic \<open>pair_induct_tac @{context} "y" [@{thm is_exec_frag_def}] 1\<close>)
-  apply (intro strip)
-  apply (tactic \<open>Seq_case_simp_tac @{context} "x" 1\<close>)
-  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-  apply auto
-  done
-
-lemmas exec_prefixclosed =
-  conjI [THEN execfrag_prefixclosed [THEN spec, THEN spec, THEN mp]]
-
-
-(* second prefix notion for Finite x *)
-
-lemma exec_prefix2closed [rule_format]:
-  "! y s. is_exec_frag A (s,x@@y) --> is_exec_frag A (s,x)"
-  apply (tactic \<open>pair_induct_tac @{context} "x" [@{thm is_exec_frag_def}] 1\<close>)
-  apply (intro strip)
-  apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
-  apply (tactic \<open>pair_tac @{context} "a" 1\<close>)
-  apply auto
-  done
-
-end
--- a/src/HOL/ROOT	Thu Dec 31 12:37:16 2015 +0100
+++ b/src/HOL/ROOT	Thu Dec 31 12:43:09 2015 +0100
@@ -1100,7 +1100,7 @@
     finite and infinite sequences.
   *}
   options [document = false]
-  theories "meta_theory/Abstraction"
+  theories "Abstraction"
 
 session "IOA-ABP" in "HOLCF/IOA/ABP" = IOA +
   description {*