--- a/src/HOLCF/IOA/Modelcheck/MuIOA.ML Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/Modelcheck/MuIOA.ML Sun May 28 19:54:20 2006 +0200
@@ -5,12 +5,13 @@
exception SimFailureExn of string;
-val ioa_simps = [asig_of_def,starts_of_def,trans_of_def];
-val asig_simps = [asig_inputs_def,asig_outputs_def,asig_internals_def,actions_def];
-val comp_simps = [par_def,asig_comp_def];
-val restrict_simps = [restrict_def,restrict_asig_def];
-val hide_simps = [hide_def,hide_asig_def];
-val rename_simps = [rename_def,rename_set_def];
+val ioa_simps = [thm "asig_of_def", thm "starts_of_def", thm "trans_of_def"];
+val asig_simps = [thm "asig_inputs_def", thm "asig_outputs_def",
+ thm "asig_internals_def", thm "actions_def"];
+val comp_simps = [thm "par_def", thm "asig_comp_def"];
+val restrict_simps = [thm "restrict_def", thm "restrict_asig_def"];
+val hide_simps = [thm "hide_def", thm "hide_asig_def"];
+val rename_simps = [thm "rename_def", thm "rename_set_def"];
local
--- a/src/HOLCF/IOA/Modelcheck/MuIOAOracle.ML Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/Modelcheck/MuIOAOracle.ML Sun May 28 19:54:20 2006 +0200
@@ -6,28 +6,24 @@
let val sign = #sign (rep_thm state);
val (subgoal::_) = Library.drop(i-1,prems_of state);
val OraAss = sim_oracle sign (subgoal,thm_list);
-in
-(cut_facts_tac [OraAss] i) state
-end;
+in cut_facts_tac [OraAss] i state end;
+
+
+val ioa_implements_def = thm "ioa_implements_def";
(* is_sim_tac makes additionally to call_sim_tac some simplifications,
which are suitable for implementation realtion formulas *)
-fun is_sim_tac thm_list i state =
-let val sign = #sign (rep_thm state);
-in
-case Library.drop(i-1,prems_of state) of
- [] => no_tac state |
- subgoal::_ => EVERY[REPEAT(etac thin_rl i),
- simp_tac (simpset() addsimps [ioa_implements_def]) i,
- rtac conjI i,
- rtac conjI (i+1),
- TRY(call_sim_tac thm_list (i+2)),
- TRY(atac (i+2)),
- REPEAT(rtac refl (i+2)),
- simp_tac (simpset() addsimps (thm_list @
- comp_simps @ restrict_simps @ hide_simps @
- rename_simps @ ioa_simps @ asig_simps)) (i+1),
- simp_tac (simpset() addsimps (thm_list @
- comp_simps @ restrict_simps @ hide_simps @
- rename_simps @ ioa_simps @ asig_simps)) (i)] state
-end;
+fun is_sim_tac thm_list = SUBGOAL (fn (goal, i) =>
+ EVERY [REPEAT (etac thin_rl i),
+ simp_tac (simpset() addsimps [ioa_implements_def]) i,
+ rtac conjI i,
+ rtac conjI (i+1),
+ TRY(call_sim_tac thm_list (i+2)),
+ TRY(atac (i+2)),
+ REPEAT(rtac refl (i+2)),
+ simp_tac (simpset() addsimps (thm_list @
+ comp_simps @ restrict_simps @ hide_simps @
+ rename_simps @ ioa_simps @ asig_simps)) (i+1),
+ simp_tac (simpset() addsimps (thm_list @
+ comp_simps @ restrict_simps @ hide_simps @
+ rename_simps @ ioa_simps @ asig_simps)) (i)]);
--- a/src/HOLCF/IOA/NTP/Impl.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/NTP/Impl.thy Sun May 28 19:54:20 2006 +0200
@@ -214,19 +214,19 @@
apply (tactic "EVERY1 [tac2,tac2,tac2,tac2]")
txt {* 6 *}
apply (tactic {* forward_tac [rewrite_rule [thm "Impl.inv1_def"]
- (thm "inv1" RS invariantE) RS conjunct1] 1 *})
+ (thm "inv1" RS thm "invariantE") RS conjunct1] 1 *})
txt {* 6 - 5 *}
apply (tactic "EVERY1 [tac2,tac2]")
txt {* 4 *}
apply (tactic {* forward_tac [rewrite_rule [thm "Impl.inv1_def"]
- (thm "inv1" RS invariantE) RS conjunct1] 1 *})
+ (thm "inv1" RS thm "invariantE") RS conjunct1] 1 *})
apply (tactic "tac2 1")
txt {* 3 *}
apply (tactic {* forward_tac [rewrite_rule [thm "Impl.inv1_def"]
- (thm "inv1" RS invariantE)] 1 *})
+ (thm "inv1" RS thm "invariantE")] 1 *})
apply (tactic "tac2 1")
apply (tactic {* fold_tac [rewrite_rule [thm "Packet.hdr_def"] (thm "Impl.hdr_sum_def")] *})
@@ -235,7 +235,7 @@
txt {* 2 *}
apply (tactic "tac2 1")
apply (tactic {* forward_tac [rewrite_rule [thm "Impl.inv1_def"]
- (thm "inv1" RS invariantE) RS conjunct1] 1 *})
+ (thm "inv1" RS thm "invariantE") RS conjunct1] 1 *})
apply (intro strip)
apply (erule conjE)+
apply simp
@@ -243,7 +243,7 @@
txt {* 1 *}
apply (tactic "tac2 1")
apply (tactic {* forward_tac [rewrite_rule [thm "Impl.inv1_def"]
- (thm "inv1" RS invariantE) RS conjunct2] 1 *})
+ (thm "inv1" RS thm "invariantE") RS conjunct2] 1 *})
apply (intro strip)
apply (erule conjE)+
apply (tactic {* fold_tac [rewrite_rule[thm "Packet.hdr_def"] (thm "Impl.hdr_sum_def")] *})
@@ -291,13 +291,13 @@
apply (rule imp_disjL [THEN iffD1])
apply (rule impI)
apply (tactic {* forward_tac [rewrite_rule [thm "Impl.inv2_def"]
- (thm "inv2" RS invariantE)] 1 *})
+ (thm "inv2" RS thm "invariantE")] 1 *})
apply simp
apply (erule conjE)+
apply (rule_tac j = "count (ssent (sen s)) (~sbit (sen s))" and
k = "count (rsent (rec s)) (sbit (sen s))" in le_trans)
apply (tactic {* forward_tac [rewrite_rule [thm "inv1_def"]
- (thm "inv1" RS invariantE) RS conjunct2] 1 *})
+ (thm "inv1" RS thm "invariantE") RS conjunct2] 1 *})
apply (simp add: hdr_sum_def Multiset.count_def)
apply (rule add_le_mono)
apply (rule countm_props)
@@ -312,7 +312,7 @@
apply (rule imp_disjL [THEN iffD1])
apply (rule impI)
apply (tactic {* forward_tac [rewrite_rule [thm "Impl.inv2_def"]
- (thm "inv2" RS invariantE)] 1 *})
+ (thm "inv2" RS thm "invariantE")] 1 *})
apply simp
done
@@ -338,7 +338,7 @@
apply (intro strip, (erule conjE)+)
apply (tactic {* forward_tac [rewrite_rule [thm "Impl.inv2_def"]
- (thm "inv2" RS invariantE)] 1 *})
+ (thm "inv2" RS thm "invariantE")] 1 *})
apply simp
txt {* 1 *}
@@ -346,9 +346,9 @@
apply (intro strip, (erule conjE)+)
apply (rule ccontr)
apply (tactic {* forward_tac [rewrite_rule [thm "Impl.inv2_def"]
- (thm "inv2" RS invariantE)] 1 *})
+ (thm "inv2" RS thm "invariantE")] 1 *})
apply (tactic {* forward_tac [rewrite_rule [thm "Impl.inv3_def"]
- (thm "inv3" RS invariantE)] 1 *})
+ (thm "inv3" RS thm "invariantE")] 1 *})
apply simp
apply (erule_tac x = "m" in allE)
apply simp
--- a/src/HOLCF/IOA/meta_theory/Abstraction.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,602 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/Abstraction.thy
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-section "cex_abs";
-
-
-(* ---------------------------------------------------------------- *)
-(* cex_abs *)
-(* ---------------------------------------------------------------- *)
-
-Goal "cex_abs f (s,UU) = (f s, UU)";
-by (simp_tac (simpset() addsimps [cex_abs_def]) 1);
-qed"cex_abs_UU";
-
-Goal "cex_abs f (s,nil) = (f s, nil)";
-by (simp_tac (simpset() addsimps [cex_abs_def]) 1);
-qed"cex_abs_nil";
-
-Goal "cex_abs f (s,(a,t)>>ex) = (f s, (a,f t) >> (snd (cex_abs f (t,ex))))";
-by (simp_tac (simpset() addsimps [cex_abs_def]) 1);
-qed"cex_abs_cons";
-
-Addsimps [cex_abs_UU, cex_abs_nil, cex_abs_cons];
-
-
-
-section "lemmas";
-
-(* ---------------------------------------------------------------- *)
-(* Lemmas *)
-(* ---------------------------------------------------------------- *)
-
-Goal "temp_weakening Q P h = (! ex. (ex |== P) --> (cex_abs h ex |== Q))";
-by (simp_tac (simpset() addsimps [temp_weakening_def,temp_strengthening_def,
- NOT_def,temp_sat_def,satisfies_def]) 1);
-by Auto_tac;
-qed"temp_weakening_def2";
-
-Goal "state_weakening Q P h = (! s t a. P (s,a,t) --> Q (h(s),a,h(t)))";
-by (simp_tac (simpset() addsimps [state_weakening_def,state_strengthening_def,
- NOT_def]) 1);
-by Auto_tac;
-qed"state_weakening_def2";
-
-
-section "Abstraction Rules for Properties";
-
-(* ---------------------------------------------------------------- *)
-(* Abstraction Rules for Properties *)
-(* ---------------------------------------------------------------- *)
-
-
-Goalw [cex_abs_def]
- "[| is_abstraction h C A |] ==>\
-\ !s. reachable C s & is_exec_frag C (s,xs) \
-\ --> is_exec_frag A (cex_abs h (s,xs))";
-
-by (Asm_full_simp_tac 1);
-by (pair_induct_tac "xs" [is_exec_frag_def] 1);
-(* main case *)
-by (safe_tac set_cs);
-by (asm_full_simp_tac (simpset() addsimps [is_abstraction_def])1);
-by (forward_tac [reachable.reachable_n] 1);
-by (assume_tac 1);
-by (Asm_full_simp_tac 1);
-qed_spec_mp"exec_frag_abstraction";
-
-
-Goal "is_abstraction h C A ==> weakeningIOA A C h";
-by (asm_full_simp_tac (simpset() addsimps [weakeningIOA_def])1);
-by Auto_tac;
-by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
-(* start state *)
-by (rtac conjI 1);
-by (asm_full_simp_tac (simpset() addsimps [is_abstraction_def,cex_abs_def]) 1);
-(* is-execution-fragment *)
-by (etac exec_frag_abstraction 1);
-by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1);
-qed"abs_is_weakening";
-
-
-Goal "[|is_abstraction h C A; validIOA A Q; temp_strengthening Q P h |] \
-\ ==> validIOA C P";
-by (dtac abs_is_weakening 1);
-by (asm_full_simp_tac (simpset() addsimps [weakeningIOA_def,
- validIOA_def, temp_strengthening_def])1);
-by (safe_tac set_cs);
-by (pair_tac "ex" 1);
-qed"AbsRuleT1";
-
-
-(* FIX: Nach TLS.ML *)
-
-Goal "(ex |== P .--> Q) = ((ex |== P) --> (ex |== Q))";
-by (simp_tac (simpset() addsimps [IMPLIES_def,temp_sat_def, satisfies_def])1);
-qed"IMPLIES_temp_sat";
-
-Goal "(ex |== P .& Q) = ((ex |== P) & (ex |== Q))";
-by (simp_tac (simpset() addsimps [AND_def,temp_sat_def, satisfies_def])1);
-qed"AND_temp_sat";
-
-Goal "(ex |== P .| Q) = ((ex |== P) | (ex |== Q))";
-by (simp_tac (simpset() addsimps [OR_def,temp_sat_def, satisfies_def])1);
-qed"OR_temp_sat";
-
-Goal "(ex |== .~ P) = (~ (ex |== P))";
-by (simp_tac (simpset() addsimps [NOT_def,temp_sat_def, satisfies_def])1);
-qed"NOT_temp_sat";
-
-Addsimps [IMPLIES_temp_sat,AND_temp_sat,OR_temp_sat,NOT_temp_sat];
-
-
-Goalw [is_live_abstraction_def]
- "[|is_live_abstraction h (C,L) (A,M); \
-\ validLIOA (A,M) Q; temp_strengthening Q P h |] \
-\ ==> validLIOA (C,L) P";
-by Auto_tac;
-by (dtac abs_is_weakening 1);
-by (asm_full_simp_tac (simpset() addsimps [weakeningIOA_def, temp_weakening_def2,
- validLIOA_def, validIOA_def, temp_strengthening_def])1);
-by (safe_tac set_cs);
-by (pair_tac "ex" 1);
-qed"AbsRuleT2";
-
-
-Goalw [is_live_abstraction_def]
- "[|is_live_abstraction h (C,L) (A,M); \
-\ validLIOA (A,M) (H1 .--> Q); temp_strengthening Q P h; \
-\ temp_weakening H1 H2 h; validLIOA (C,L) H2 |] \
-\ ==> validLIOA (C,L) P";
-by Auto_tac;
-by (dtac abs_is_weakening 1);
-by (asm_full_simp_tac (simpset() addsimps [weakeningIOA_def, temp_weakening_def2,
- validLIOA_def, validIOA_def, temp_strengthening_def])1);
-by (safe_tac set_cs);
-by (pair_tac "ex" 1);
-qed"AbsRuleTImprove";
-
-
-section "Correctness of safe abstraction";
-
-(* ---------------------------------------------------------------- *)
-(* Correctness of safe abstraction *)
-(* ---------------------------------------------------------------- *)
-
-
-Goalw [is_abstraction_def,is_ref_map_def]
-"is_abstraction h C A ==> is_ref_map h C A";
-by (safe_tac set_cs);
-by (res_inst_tac[("x","(a,h t)>>nil")] exI 1);
-by (asm_full_simp_tac (simpset() addsimps [move_def])1);
-qed"abstraction_is_ref_map";
-
-
-Goal "[| inp(C)=inp(A); out(C)=out(A); \
-\ is_abstraction h C A |] \
-\ ==> C =<| A";
-by (asm_full_simp_tac (simpset() addsimps [ioa_implements_def]) 1);
-by (rtac trace_inclusion 1);
-by (simp_tac (simpset() addsimps [externals_def])1);
-by (SELECT_GOAL (auto_tac (claset(),simpset()))1);
-by (etac abstraction_is_ref_map 1);
-qed"abs_safety";
-
-
-section "Correctness of life abstraction";
-
-(* ---------------------------------------------------------------- *)
-(* 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 *)
-Goalw [cex_abs_def,mk_trace_def,filter_act_def]
- "ext C = ext A \
-\ ==> mk_trace C$xs = mk_trace A$(snd (cex_abs f (s,xs)))";
-by (Asm_full_simp_tac 1);
-by (pair_induct_tac "xs" [] 1);
-qed"traces_coincide_abs";
-
-
-(* 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 *)
-Goal "[| inp(C)=inp(A); out(C)=out(A); \
-\ is_live_abstraction h (C,M) (A,L) |] \
-\ ==> live_implements (C,M) (A,L)";
-
-by (asm_full_simp_tac (simpset() addsimps [is_live_abstraction_def, live_implements_def,
-livetraces_def,liveexecutions_def]) 1);
-by (safe_tac set_cs);
-by (res_inst_tac[("x","cex_abs h ex")] exI 1);
-by (safe_tac set_cs);
- (* Traces coincide *)
- by (pair_tac "ex" 1);
- by (rtac traces_coincide_abs 1);
- by (simp_tac (simpset() addsimps [externals_def])1);
- by (SELECT_GOAL (auto_tac (claset(),simpset()))1);
-
- (* cex_abs is execution *)
- by (pair_tac "ex" 1);
- by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
- (* start state *)
- by (rtac conjI 1);
- by (asm_full_simp_tac (simpset() addsimps [is_abstraction_def,cex_abs_def]) 1);
- (* is-execution-fragment *)
- by (etac exec_frag_abstraction 1);
- by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1);
-
- (* Liveness *)
-by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2]) 1);
- by (pair_tac "ex" 1);
-qed"abs_liveness";
-
-(* FIX: NAch Traces.ML bringen *)
-
-Goalw [ioa_implements_def]
-"[| A =<| B; B =<| C|] ==> A =<| C";
-by Auto_tac;
-qed"implements_trans";
-
-
-section "Abstraction Rules for Automata";
-
-(* ---------------------------------------------------------------- *)
-(* Abstraction Rules for Automata *)
-(* ---------------------------------------------------------------- *)
-
-Goal "[| 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";
-by (dtac abs_safety 1);
-by (REPEAT (atac 1));
-by (dtac abs_safety 1);
-by (REPEAT (atac 1));
-by (etac implements_trans 1);
-by (etac implements_trans 1);
-by (assume_tac 1);
-qed"AbsRuleA1";
-
-
-Goal "!!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)";
-by (dtac abs_liveness 1);
-by (REPEAT (atac 1));
-by (dtac abs_liveness 1);
-by (REPEAT (atac 1));
-by (etac live_implements_trans 1);
-by (etac live_implements_trans 1);
-by (assume_tac 1);
-qed"AbsRuleA2";
-
-
-Delsimps [split_paired_All];
-
-
-section "Localizing Temporal Strengthenings and Weakenings";
-
-(* ---------------------------------------------------------------- *)
-(* Localizing Temproal Strengthenings - 1 *)
-(* ---------------------------------------------------------------- *)
-
-Goalw [temp_strengthening_def]
-"[| temp_strengthening P1 Q1 h; \
-\ temp_strengthening P2 Q2 h |] \
-\ ==> temp_strengthening (P1 .& P2) (Q1 .& Q2) h";
-by Auto_tac;
-qed"strength_AND";
-
-Goalw [temp_strengthening_def]
-"[| temp_strengthening P1 Q1 h; \
-\ temp_strengthening P2 Q2 h |] \
-\ ==> temp_strengthening (P1 .| P2) (Q1 .| Q2) h";
-by Auto_tac;
-qed"strength_OR";
-
-Goalw [temp_strengthening_def]
-"[| temp_weakening P Q h |] \
-\ ==> temp_strengthening (.~ P) (.~ Q) h";
-by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1);
-by Auto_tac;
-qed"strength_NOT";
-
-Goalw [temp_strengthening_def]
-"[| temp_weakening P1 Q1 h; \
-\ temp_strengthening P2 Q2 h |] \
-\ ==> temp_strengthening (P1 .--> P2) (Q1 .--> Q2) h";
-by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1);
-qed"strength_IMPLIES";
-
-
-
-(* ---------------------------------------------------------------- *)
-(* Localizing Temproal Weakenings - Part 1 *)
-(* ---------------------------------------------------------------- *)
-
-Goal
-"[| temp_weakening P1 Q1 h; \
-\ temp_weakening P2 Q2 h |] \
-\ ==> temp_weakening (P1 .& P2) (Q1 .& Q2) h";
-by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1);
-qed"weak_AND";
-
-Goal
-"[| temp_weakening P1 Q1 h; \
-\ temp_weakening P2 Q2 h |] \
-\ ==> temp_weakening (P1 .| P2) (Q1 .| Q2) h";
-by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1);
-qed"weak_OR";
-
-Goalw [temp_strengthening_def]
-"[| temp_strengthening P Q h |] \
-\ ==> temp_weakening (.~ P) (.~ Q) h";
-by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1);
-by Auto_tac;
-qed"weak_NOT";
-
-Goalw [temp_strengthening_def]
-"[| temp_strengthening P1 Q1 h; \
-\ temp_weakening P2 Q2 h |] \
-\ ==> temp_weakening (P1 .--> P2) (Q1 .--> Q2) h";
-by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2])1);
-qed"weak_IMPLIES";
-
-
-(* ---------------------------------------------------------------- *)
-(* Localizing Temproal Strengthenings - 2 *)
-(* ---------------------------------------------------------------- *)
-
-
-(* ------------------ Box ----------------------------*)
-
-(* FIX: should be same as nil_is_Conc2 when all nils are turned to right side !! *)
-Goal "(UU = x @@ y) = (((x::'a Seq)= UU) | (x=nil & y=UU))";
-by (Seq_case_simp_tac "x" 1);
-by Auto_tac;
-qed"UU_is_Conc";
-
-Goal
-"Finite s1 --> \
-\ (! ex. (s~=nil & s~=UU & ex2seq ex = s1 @@ s) --> (? ex'. s = ex2seq ex'))";
-by (rtac impI 1);
-by (Seq_Finite_induct_tac 1);
-by (Blast_tac 1);
-(* main case *)
-by (clarify_tac set_cs 1);
-by (pair_tac "ex" 1);
-by (Seq_case_simp_tac "y" 1);
-(* UU case *)
-by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc])1);
-(* nil case *)
-by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc])1);
-(* cons case *)
-by (pair_tac "aa" 1);
-by Auto_tac;
-qed_spec_mp"ex2seqConc";
-
-(* important property of ex2seq: can be shiftet, as defined "pointwise" *)
-
-Goalw [tsuffix_def,suffix_def]
-"tsuffix s (ex2seq ex) ==> ? ex'. s = (ex2seq ex')";
-by Auto_tac;
-by (dtac ex2seqConc 1);
-by Auto_tac;
-qed"ex2seq_tsuffix";
-
-
-(* FIX: NAch Sequence.ML bringen *)
-
-Goal "(Map f$s = nil) = (s=nil)";
-by (Seq_case_simp_tac "s" 1);
-qed"Mapnil";
-
-Goal "(Map f$s = UU) = (s=UU)";
-by (Seq_case_simp_tac "s" 1);
-qed"MapUU";
-
-
-(* important property of cex_absSeq: As it is a 1to1 correspondence,
- properties carry over *)
-
-Goalw [tsuffix_def,suffix_def,cex_absSeq_def]
-"tsuffix s t ==> tsuffix (cex_absSeq h s) (cex_absSeq h t)";
-by Auto_tac;
-by (asm_full_simp_tac (simpset() addsimps [Mapnil])1);
-by (asm_full_simp_tac (simpset() addsimps [MapUU])1);
-by (res_inst_tac [("x","Map (%(s,a,t). (h s,a, h t))$s1")] exI 1);
-by (asm_full_simp_tac (simpset() addsimps [Map2Finite,MapConc])1);
-qed"cex_absSeq_tsuffix";
-
-
-Goalw [temp_strengthening_def,state_strengthening_def, temp_sat_def,
-satisfies_def,Box_def]
-"[| temp_strengthening P Q h |]\
-\ ==> temp_strengthening ([] P) ([] Q) h";
-by (clarify_tac set_cs 1);
-by (ftac ex2seq_tsuffix 1);
-by (clarify_tac set_cs 1);
-by (dres_inst_tac [("h","h")] cex_absSeq_tsuffix 1);
-by (asm_full_simp_tac (simpset() addsimps [ex2seq_abs_cex])1);
-qed"strength_Box";
-
-
-(* ------------------ Init ----------------------------*)
-
-Goalw [temp_strengthening_def,state_strengthening_def,
-temp_sat_def,satisfies_def,Init_def,unlift_def]
-"[| state_strengthening P Q h |]\
-\ ==> temp_strengthening (Init P) (Init Q) h";
-by (safe_tac set_cs);
-by (pair_tac "ex" 1);
-by (Seq_case_simp_tac "y" 1);
-by (pair_tac "a" 1);
-qed"strength_Init";
-
-
-(* ------------------ Next ----------------------------*)
-
-Goal
-"(TL$(ex2seq (cex_abs h ex))=UU) = (TL$(ex2seq ex)=UU)";
-by (pair_tac "ex" 1);
-by (Seq_case_simp_tac "y" 1);
-by (pair_tac "a" 1);
-by (Seq_case_simp_tac "s" 1);
-by (pair_tac "a" 1);
-qed"TL_ex2seq_UU";
-
-Goal
-"(TL$(ex2seq (cex_abs h ex))=nil) = (TL$(ex2seq ex)=nil)";
-by (pair_tac "ex" 1);
-by (Seq_case_simp_tac "y" 1);
-by (pair_tac "a" 1);
-by (Seq_case_simp_tac "s" 1);
-by (pair_tac "a" 1);
-qed"TL_ex2seq_nil";
-
-(* FIX: put to Sequence Lemmas *)
-Goal "Map f$(TL$s) = TL$(Map f$s)";
-by (Seq_induct_tac "s" [] 1);
-qed"MapTL";
-
-(* important property of cex_absSeq: As it is a 1to1 correspondence,
- properties carry over *)
-
-Goalw [cex_absSeq_def]
-"cex_absSeq h (TL$s) = (TL$(cex_absSeq h s))";
-by (simp_tac (simpset() addsimps [MapTL]) 1);
-qed"cex_absSeq_TL";
-
-(* important property of ex2seq: can be shiftet, as defined "pointwise" *)
-
-Goal "[| (snd ex)~=UU ; (snd ex)~=nil |] ==> (? ex'. TL$(ex2seq ex) = ex2seq ex')";
-by (pair_tac "ex" 1);
-by (Seq_case_simp_tac "y" 1);
-by (pair_tac "a" 1);
-by Auto_tac;
-qed"TLex2seq";
-
-
-Goal "(TL$(ex2seq ex)~=nil) = ((snd ex)~=nil & (snd ex)~=UU)";
-by (pair_tac "ex" 1);
-by (Seq_case_simp_tac "y" 1);
-by (pair_tac "a" 1);
-by (Seq_case_simp_tac "s" 1);
-by (pair_tac "a" 1);
-qed"ex2seqnilTL";
-
-
-Goalw [temp_strengthening_def,state_strengthening_def,
-temp_sat_def, satisfies_def,Next_def]
-"[| temp_strengthening P Q h |]\
-\ ==> temp_strengthening (Next P) (Next Q) h";
-by (Asm_full_simp_tac 1);
-by (safe_tac set_cs);
-by (asm_full_simp_tac (simpset() addsimps [TL_ex2seq_nil,TL_ex2seq_UU]) 1);
-by (asm_full_simp_tac (simpset() addsimps [TL_ex2seq_nil,TL_ex2seq_UU]) 1);
-by (asm_full_simp_tac (simpset() addsimps [TL_ex2seq_nil,TL_ex2seq_UU]) 1);
-by (asm_full_simp_tac (simpset() addsimps [TL_ex2seq_nil,TL_ex2seq_UU]) 1);
-(* cons case *)
-by (asm_full_simp_tac (simpset() addsimps [TL_ex2seq_nil,TL_ex2seq_UU,
- ex2seq_abs_cex,cex_absSeq_TL RS sym, ex2seqnilTL])1);
-by (etac conjE 1);
-by (dtac TLex2seq 1);
-by (assume_tac 1);
-by Auto_tac;
-qed"strength_Next";
-
-
-
-(* ---------------------------------------------------------------- *)
-(* Localizing Temporal Weakenings - 2 *)
-(* ---------------------------------------------------------------- *)
-
-
-Goal
-"[| state_weakening P Q h |]\
-\ ==> temp_weakening (Init P) (Init Q) h";
-by (asm_full_simp_tac (simpset() addsimps [temp_weakening_def2,
- state_weakening_def2, temp_sat_def,satisfies_def,Init_def,unlift_def])1);
-by (safe_tac set_cs);
-by (pair_tac "ex" 1);
-by (Seq_case_simp_tac "y" 1);
-by (pair_tac "a" 1);
-qed"weak_Init";
-
-
-(* ---------------------------------------------------------------- *)
-(* Localizing Temproal Strengthenings - 3 *)
-(* ---------------------------------------------------------------- *)
-
-
-Goalw [Diamond_def]
-"[| temp_strengthening P Q h |]\
-\ ==> temp_strengthening (<> P) (<> Q) h";
-by (rtac strength_NOT 1);
-by (rtac weak_Box 1);
-by (etac weak_NOT 1);
-qed"strength_Diamond";
-
-Goalw [Leadsto_def]
-"[| temp_weakening P1 P2 h;\
-\ temp_strengthening Q1 Q2 h |]\
-\ ==> temp_strengthening (P1 ~> Q1) (P2 ~> Q2) h";
-by (rtac strength_Box 1);
-by (etac strength_IMPLIES 1);
-by (etac strength_Diamond 1);
-qed"strength_Leadsto";
-
-
-(* ---------------------------------------------------------------- *)
-(* Localizing Temporal Weakenings - 3 *)
-(* ---------------------------------------------------------------- *)
-
-
-Goalw [Diamond_def]
-"[| temp_weakening P Q h |]\
-\ ==> temp_weakening (<> P) (<> Q) h";
-by (rtac weak_NOT 1);
-by (rtac strength_Box 1);
-by (etac strength_NOT 1);
-qed"weak_Diamond";
-
-Goalw [Leadsto_def]
-"[| temp_strengthening P1 P2 h;\
-\ temp_weakening Q1 Q2 h |]\
-\ ==> temp_weakening (P1 ~> Q1) (P2 ~> Q2) h";
-by (rtac weak_Box 1);
-by (etac weak_IMPLIES 1);
-by (etac weak_Diamond 1);
-qed"weak_Leadsto";
-
-Goalw [WF_def]
- " !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|] \
-\ ==> temp_weakening (WF A acts) (WF C acts) h";
-by (rtac weak_IMPLIES 1);
-by (rtac strength_Diamond 1);
-by (rtac strength_Box 1);
-by (rtac strength_Init 1);
-by (rtac weak_Box 2);
-by (rtac weak_Diamond 2);
-by (rtac weak_Init 2);
-by (auto_tac (claset(),
- simpset() addsimps [state_weakening_def,state_strengthening_def,
- xt2_def,plift_def,option_lift_def,NOT_def]));
-qed"weak_WF";
-
-Goalw [SF_def]
- " !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|] \
-\ ==> temp_weakening (SF A acts) (SF C acts) h";
-by (rtac weak_IMPLIES 1);
-by (rtac strength_Box 1);
-by (rtac strength_Diamond 1);
-by (rtac strength_Init 1);
-by (rtac weak_Box 2);
-by (rtac weak_Diamond 2);
-by (rtac weak_Init 2);
-by (auto_tac (claset(),
- simpset() addsimps [state_weakening_def,state_strengthening_def,
- xt2_def,plift_def,option_lift_def,NOT_def]));
-qed"weak_SF";
-
-
-val 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];
-
-fun abstraction_tac i =
- SELECT_GOAL (auto_tac (claset() addSIs weak_strength_lemmas,
- simpset() addsimps [state_strengthening_def,state_weakening_def])) i;
--- a/src/HOLCF/IOA/meta_theory/Abstraction.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/Abstraction.thy Sun May 28 19:54:20 2006 +0200
@@ -79,6 +79,564 @@
"temp_weakening P Q h
==> temp_weakening (Next P) (Next Q) h"
-ML {* use_legacy_bindings (the_context ()) *}
+
+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)>>ex) = (f s, (a,f t) >> (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 |== P) --> (cex_abs h ex |== 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 {* pair_induct_tac "xs" [thm "is_exec_frag_def"] 1 *})
+txt {* main case *}
+apply (tactic "safe_tac set_cs")
+apply (simp add: is_abstraction_def)
+apply (frule reachable.reachable_n)
+apply assumption
+apply simp
+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 {* start state *}
+apply (rule conjI)
+apply (simp add: is_abstraction_def cex_abs_def)
+txt {* is-execution-fragment *}
+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 (tactic "safe_tac set_cs")
+apply (tactic {* pair_tac "ex" 1 *})
+done
+
+
+(* FIX: Nach TLS.ML *)
+
+lemma IMPLIES_temp_sat: "(ex |== P .--> Q) = ((ex |== P) --> (ex |== Q))"
+ by (simp add: IMPLIES_def temp_sat_def satisfies_def)
+
+lemma AND_temp_sat: "(ex |== P .& Q) = ((ex |== P) & (ex |== Q))"
+ by (simp add: AND_def temp_sat_def satisfies_def)
+
+lemma OR_temp_sat: "(ex |== P .| Q) = ((ex |== P) | (ex |== Q))"
+ by (simp add: OR_def temp_sat_def satisfies_def)
+
+lemma NOT_temp_sat: "(ex |== .~ P) = (~ (ex |== 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 (tactic "safe_tac set_cs")
+apply (tactic {* pair_tac "ex" 1 *})
+done
+
+
+lemma AbsRuleTImprove:
+ "[|is_live_abstraction h (C,L) (A,M);
+ validLIOA (A,M) (H1 .--> 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 (tactic "safe_tac set_cs")
+apply (tactic {* pair_tac "ex" 1 *})
+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 (tactic "safe_tac set_cs")
+apply (rule_tac x = "(a,h t) >>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 {* pair_induct_tac "xs" [] 1 *})
+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 (tactic "safe_tac set_cs")
+apply (rule_tac x = "cex_abs h ex" in exI)
+apply (tactic "safe_tac set_cs")
+ (* Traces coincide *)
+ apply (tactic {* pair_tac "ex" 1 *})
+ apply (rule traces_coincide_abs)
+ apply (simp (no_asm) add: externals_def)
+ apply (auto)[1]
+
+ (* cex_abs is execution *)
+ apply (tactic {* pair_tac "ex" 1 *})
+ 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 {* pair_tac "ex" 1 *})
+done
+
+(* FIX: NAch Traces.ML bringen *)
+
+lemma implements_trans:
+"[| A =<| B; B =<| C|] ==> A =<| C"
+apply (unfold ioa_implements_def)
+apply auto
+done
+
+
+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 .& P2) (Q1 .& 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 .| P2) (Q1 .| Q2) h"
+apply (unfold temp_strengthening_def)
+apply auto
+done
+
+lemma strength_NOT:
+"[| temp_weakening P Q h |]
+ ==> temp_strengthening (.~ P) (.~ 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 .--> P2) (Q1 .--> 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 .& P2) (Q1 .& 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 .| P2) (Q1 .| Q2) h"
+apply (simp add: temp_weakening_def2)
+done
+
+lemma weak_NOT:
+"[| temp_strengthening P Q h |]
+ ==> temp_weakening (.~ P) (.~ 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 .--> P2) (Q1 .--> Q2) h"
+apply (unfold temp_strengthening_def)
+apply (simp add: temp_weakening_def2)
+done
+
+
+subsubsection {* Box *}
+
+(* 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 {* Seq_case_simp_tac "x" 1 *})
+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 {* Seq_Finite_induct_tac 1 *})
+apply blast
+(* main case *)
+apply (tactic "clarify_tac set_cs 1")
+apply (tactic {* pair_tac "ex" 1 *})
+apply (tactic {* Seq_case_simp_tac "y" 1 *})
+(* UU case *)
+apply (simp add: nil_is_Conc)
+(* nil case *)
+apply (simp add: nil_is_Conc)
+(* cons case *)
+apply (tactic {* pair_tac "aa" 1 *})
+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 {* Seq_case_simp_tac "s" 1 *})
+done
+
+lemma MapUU: "(Map f$s = UU) = (s=UU)"
+apply (tactic {* Seq_case_simp_tac "s" 1 *})
+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 ([] P) ([] Q) h"
+apply (unfold temp_strengthening_def state_strengthening_def temp_sat_def satisfies_def Box_def)
+apply (tactic "clarify_tac set_cs 1")
+apply (frule ex2seq_tsuffix)
+apply (tactic "clarify_tac set_cs 1")
+apply (drule_tac h = "h" in cex_absSeq_tsuffix)
+apply (simp add: ex2seq_abs_cex)
+done
+
+
+subsubsection {* Init *}
+
+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 (tactic "safe_tac set_cs")
+apply (tactic {* pair_tac "ex" 1 *})
+apply (tactic {* Seq_case_simp_tac "y" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+done
+
+
+subsubsection {* Next *}
+
+lemma TL_ex2seq_UU:
+"(TL$(ex2seq (cex_abs h ex))=UU) = (TL$(ex2seq ex)=UU)"
+apply (tactic {* pair_tac "ex" 1 *})
+apply (tactic {* Seq_case_simp_tac "y" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+apply (tactic {* Seq_case_simp_tac "s" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+done
+
+lemma TL_ex2seq_nil:
+"(TL$(ex2seq (cex_abs h ex))=nil) = (TL$(ex2seq ex)=nil)"
+apply (tactic {* pair_tac "ex" 1 *})
+apply (tactic {* Seq_case_simp_tac "y" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+apply (tactic {* Seq_case_simp_tac "s" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+done
+
+(* FIX: put to Sequence Lemmas *)
+lemma MapTL: "Map f$(TL$s) = TL$(Map f$s)"
+apply (tactic {* Seq_induct_tac "s" [] 1 *})
+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 {* pair_tac "ex" 1 *})
+apply (tactic {* Seq_case_simp_tac "y" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+apply auto
+done
+
+
+lemma ex2seqnilTL: "(TL$(ex2seq ex)~=nil) = ((snd ex)~=nil & (snd ex)~=UU)"
+apply (tactic {* pair_tac "ex" 1 *})
+apply (tactic {* Seq_case_simp_tac "y" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+apply (tactic {* Seq_case_simp_tac "s" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+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 (tactic "safe_tac set_cs")
+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 {* Localizing Temporal Weakenings - 2 *}
+
+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 (tactic "safe_tac set_cs")
+apply (tactic {* pair_tac "ex" 1 *})
+apply (tactic {* Seq_case_simp_tac "y" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+done
+
+
+text {* Localizing Temproal Strengthenings - 3 *}
+
+lemma strength_Diamond:
+"[| temp_strengthening P Q h |]
+ ==> temp_strengthening (<> P) (<> 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 ~> Q1) (P2 ~> Q2) h"
+apply (unfold Leadsto_def)
+apply (rule strength_Box)
+apply (erule strength_IMPLIES)
+apply (erule strength_Diamond)
+done
+
+
+text {* Localizing Temporal Weakenings - 3 *}
+
+lemma weak_Diamond:
+"[| temp_weakening P Q h |]
+ ==> temp_weakening (<> P) (<> 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 ~> Q1) (P2 ~> 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 {*
+
+local
+ val weak_strength_lemmas = thms "weak_strength_lemmas"
+ val state_strengthening_def = thm "state_strengthening_def"
+ val state_weakening_def = thm "state_weakening_def"
+in
+
+fun abstraction_tac i =
+ SELECT_GOAL (CLASIMPSET (fn (cs, ss) =>
+ auto_tac (cs addSIs weak_strength_lemmas,
+ ss addsimps [state_strengthening_def, state_weakening_def]))) i
+end
+*}
end
--- a/src/HOLCF/IOA/meta_theory/Asig.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,53 +0,0 @@
-(* Title: HOL/IOA/meta_theory/Asig.ML
- ID: $Id$
- Author: Olaf Mueller, Tobias Nipkow & Konrad Slind
-*)
-
-bind_thms ("asig_projections", [asig_inputs_def, asig_outputs_def, asig_internals_def]);
-
-Goal
- "(outputs (a,b,c) = b) & \
-\ (inputs (a,b,c) = a) & \
-\ (internals (a,b,c) = c)";
- by (simp_tac (simpset() addsimps asig_projections) 1);
-qed "asig_triple_proj";
-
-Goal "[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)";
-by (asm_full_simp_tac (simpset() addsimps [externals_def,actions_def]) 1);
-qed"int_and_ext_is_act";
-
-Goal "[|a:externals(S)|] ==> a:actions(S)";
-by (asm_full_simp_tac (simpset() addsimps [externals_def,actions_def]) 1);
-qed"ext_is_act";
-
-Goal "[|a:internals S|] ==> a:actions S";
-by (asm_full_simp_tac (simpset() addsimps [asig_internals_def,actions_def]) 1);
-qed"int_is_act";
-
-Goal "[|a:inputs S|] ==> a:actions S";
-by (asm_full_simp_tac (simpset() addsimps [asig_inputs_def,actions_def]) 1);
-qed"inp_is_act";
-
-Goal "[|a:outputs S|] ==> a:actions S";
-by (asm_full_simp_tac (simpset() addsimps [asig_outputs_def,actions_def]) 1);
-qed"out_is_act";
-
-Goal "(x: actions S & x : externals S) = (x: externals S)";
-by (fast_tac (claset() addSIs [ext_is_act]) 1 );
-qed"ext_and_act";
-
-Goal "[|is_asig S;x: actions S|] ==> (x~:externals S) = (x: internals S)";
-by (asm_full_simp_tac (simpset() addsimps [actions_def,is_asig_def,externals_def]) 1);
-by (Blast_tac 1);
-qed"not_ext_is_int";
-
-Goal "is_asig S ==> (x~:externals S) = (x: internals S | x~:actions S)";
-by (asm_full_simp_tac (simpset() addsimps [actions_def,is_asig_def,externals_def]) 1);
-by (Blast_tac 1);
-qed"not_ext_is_int_or_not_act";
-
-Goalw [externals_def,actions_def,is_asig_def]
- "[| is_asig (S); x:internals S |] ==> x~:externals S";
-by (Asm_full_simp_tac 1);
-by (Blast_tac 1);
-qed"int_is_not_ext";
--- a/src/HOLCF/IOA/meta_theory/Asig.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/Asig.thy Sun May 28 19:54:20 2006 +0200
@@ -47,6 +47,56 @@
mk_ext_asig_def:
"mk_ext_asig(triple) == (inputs(triple), outputs(triple), {})"
-ML {* use_legacy_bindings (the_context ()) *}
+
+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/HOLCF/IOA/meta_theory/Automata.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,473 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/Automata.ML
- ID: $Id$
- Author: Olaf Mueller, Tobias Nipkow, Konrad Slind
-*)
-
-(* this modification of the simpset is local to this file *)
-Delsimps [split_paired_Ex];
-
-val ioa_projections = [asig_of_def, starts_of_def, trans_of_def,wfair_of_def,sfair_of_def];
-
-(* ------------------------------------------------------------------------- *)
-
-section "asig_of, starts_of, trans_of";
-
-
-Goal
- "((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)";
- by (simp_tac (simpset() addsimps ioa_projections) 1);
-qed "ioa_triple_proj";
-
-Goalw [is_trans_of_def,actions_def, is_asig_def]
- "[| is_trans_of A; (s1,a,s2):trans_of(A) |] ==> a:act A";
- by (REPEAT(etac conjE 1));
- by (EVERY1[etac allE, etac impE, atac]);
- by (Asm_full_simp_tac 1);
-qed "trans_in_actions";
-
-Goal
-"starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}";
- by (simp_tac (simpset() addsimps (par_def::ioa_projections)) 1);
-qed "starts_of_par";
-
-Goal
-"trans_of(A || 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_tac (simpset() addsimps (par_def::ioa_projections)) 1);
-qed "trans_of_par";
-
-
-(* ------------------------------------------------------------------------- *)
-
-section "actions and par";
-
-
-Goal
-"actions(asig_comp a b) = actions(a) Un actions(b)";
- by (simp_tac (simpset() addsimps
- ([actions_def,asig_comp_def]@asig_projections)) 1);
- by (fast_tac (set_cs addSIs [equalityI]) 1);
-qed "actions_asig_comp";
-
-
-Goal "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)";
- by (simp_tac (simpset() addsimps (par_def::ioa_projections)) 1);
-qed "asig_of_par";
-
-
-Goal "ext (A1||A2) = \
-\ (ext A1) Un (ext A2)";
-by (asm_full_simp_tac (simpset() addsimps [externals_def,asig_of_par,asig_comp_def,
- asig_inputs_def,asig_outputs_def,Un_def,set_diff_def]) 1);
-by (rtac set_ext 1);
-by (fast_tac set_cs 1);
-qed"externals_of_par";
-
-Goal "act (A1||A2) = \
-\ (act A1) Un (act A2)";
-by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def,
- asig_inputs_def,asig_outputs_def,asig_internals_def,Un_def,set_diff_def]) 1);
-by (rtac set_ext 1);
-by (fast_tac set_cs 1);
-qed"actions_of_par";
-
-Goal "inp (A1||A2) =\
-\ ((inp A1) Un (inp A2)) - ((out A1) Un (out A2))";
-by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def,
- asig_inputs_def,asig_outputs_def,Un_def,set_diff_def]) 1);
-qed"inputs_of_par";
-
-Goal "out (A1||A2) =\
-\ (out A1) Un (out A2)";
-by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def,
- asig_outputs_def,Un_def,set_diff_def]) 1);
-qed"outputs_of_par";
-
-Goal "int (A1||A2) =\
-\ (int A1) Un (int A2)";
-by (asm_full_simp_tac (simpset() addsimps [actions_def,asig_of_par,asig_comp_def,
- asig_inputs_def,asig_outputs_def,asig_internals_def,Un_def,set_diff_def]) 1);
-qed"internals_of_par";
-
-(* ------------------------------------------------------------------------ *)
-
-section "actions and compatibility";
-
-Goal"compatible A B = compatible B A";
-by (asm_full_simp_tac (simpset() addsimps [compatible_def,Int_commute]) 1);
-by Auto_tac;
-qed"compat_commute";
-
-Goalw [externals_def,actions_def,compatible_def]
- "[| compatible A1 A2; a:ext A1|] ==> a~:int A2";
-by (Asm_full_simp_tac 1);
-by (Blast_tac 1);
-qed"ext1_is_not_int2";
-
-(* just commuting the previous one: better commute compatible *)
-Goalw [externals_def,actions_def,compatible_def]
- "[| compatible A2 A1 ; a:ext A1|] ==> a~:int A2";
-by (Asm_full_simp_tac 1);
-by (Blast_tac 1);
-qed"ext2_is_not_int1";
-
-bind_thm("ext1_ext2_is_not_act2",ext1_is_not_int2 RS int_and_ext_is_act);
-bind_thm("ext1_ext2_is_not_act1",ext2_is_not_int1 RS int_and_ext_is_act);
-
-Goalw [externals_def,actions_def,compatible_def]
- "[| compatible A B; x:int A |] ==> x~:ext B";
-by (Asm_full_simp_tac 1);
-by (Blast_tac 1);
-qed"intA_is_not_extB";
-
-Goalw [externals_def,actions_def,compatible_def,is_asig_def,asig_of_def]
-"[| compatible A B; a:int A |] ==> a ~: act B";
-by (Asm_full_simp_tac 1);
-by (Blast_tac 1);
-qed"intA_is_not_actB";
-
-(* the only one that needs disjointness of outputs and of internals and _all_ acts *)
-Goalw [asig_outputs_def,asig_internals_def,actions_def,asig_inputs_def,
- compatible_def,is_asig_def,asig_of_def]
-"[| compatible A B; a:out A ;a:act B|] ==> a : inp B";
-by (Asm_full_simp_tac 1);
-by (Blast_tac 1);
-qed"outAactB_is_inpB";
-
-(* needed for propagation of input_enabledness from A,B to A||B *)
-Goalw [asig_outputs_def,asig_internals_def,actions_def,asig_inputs_def,
- compatible_def,is_asig_def,asig_of_def]
-"[| compatible A B; a:inp A ;a:act B|] ==> a : inp B | a: out B";
-by (Asm_full_simp_tac 1);
-by (Blast_tac 1);
-qed"inpAAactB_is_inpBoroutB";
-
-(* ------------------------------------------------------------------------- *)
-
-section "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 *)
-Goalw [input_enabled_def]
-"[| compatible A B; input_enabled A; input_enabled B|] \
-\ ==> input_enabled (A||B)";
-by (asm_full_simp_tac (simpset()addsimps[Let_def,inputs_of_par,trans_of_par])1);
-by (safe_tac set_cs);
-by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1);
-by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 2);
-(* a: inp A *)
-by (case_tac "a:act B" 1);
-(* a:act B *)
-by (eres_inst_tac [("x","a")] allE 1);
-by (Asm_full_simp_tac 1);
-by (dtac inpAAactB_is_inpBoroutB 1);
-by (assume_tac 1);
-by (assume_tac 1);
-by (eres_inst_tac [("x","a")] allE 1);
-by (Asm_full_simp_tac 1);
-by (eres_inst_tac [("x","aa")] allE 1);
-by (eres_inst_tac [("x","b")] allE 1);
-by (etac exE 1);
-by (etac exE 1);
-by (res_inst_tac [("x","(s2,s2a)")] exI 1);
-by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1);
-(* a~: act B*)
-by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1);
-by (eres_inst_tac [("x","a")] allE 1);
-by (Asm_full_simp_tac 1);
-by (eres_inst_tac [("x","aa")] allE 1);
-by (etac exE 1);
-by (res_inst_tac [("x","(s2,b)")] exI 1);
-by (Asm_full_simp_tac 1);
-
-(* a:inp B *)
-by (case_tac "a:act A" 1);
-(* a:act A *)
-by (eres_inst_tac [("x","a")] allE 1);
-by (eres_inst_tac [("x","a")] allE 1);
-by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1);
-by (forw_inst_tac [("A1","A")] (compat_commute RS iffD1) 1);
-by (dtac inpAAactB_is_inpBoroutB 1);
-back();
-by (assume_tac 1);
-by (assume_tac 1);
-by (Asm_full_simp_tac 1);
-by (rotate_tac ~1 1);
-by (Asm_full_simp_tac 1);
-by (eres_inst_tac [("x","aa")] allE 1);
-by (eres_inst_tac [("x","b")] allE 1);
-by (etac exE 1);
-by (etac exE 1);
-by (res_inst_tac [("x","(s2,s2a)")] exI 1);
-by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1);
-(* a~: act B*)
-by (asm_full_simp_tac (simpset() addsimps [inp_is_act]) 1);
-by (eres_inst_tac [("x","a")] allE 1);
-by (Asm_full_simp_tac 1);
-by (eres_inst_tac [("x","a")] allE 1);
-by (Asm_full_simp_tac 1);
-by (eres_inst_tac [("x","b")] allE 1);
-by (etac exE 1);
-by (res_inst_tac [("x","(aa,s2)")] exI 1);
-by (Asm_full_simp_tac 1);
-qed"input_enabled_par";
-
-(* ------------------------------------------------------------------------- *)
-
-section "invariants";
-
-val [p1,p2] = goalw (the_context ()) [invariant_def]
- "[| !!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";
-by (rtac allI 1);
-by (rtac impI 1);
-by (res_inst_tac [("xa","s")] reachable.induct 1);
-by (atac 1);
-by (etac p1 1);
-by (eres_inst_tac [("s1","sa")] (p2 RS mp) 1);
-by (REPEAT (atac 1));
-qed"invariantI";
-
-
-val [p1,p2] = goal (the_context ())
- "[| !!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";
- by (fast_tac (HOL_cs addSIs [invariantI] addSDs [p1,p2]) 1);
-qed "invariantI1";
-
-Goalw [invariant_def] "[| invariant A P; reachable A s |] ==> P(s)";
-by (Blast_tac 1);
-qed "invariantE";
-
-(* ------------------------------------------------------------------------- *)
-
-section "restrict";
-
-
-val reachable_0 = reachable.reachable_0
-and reachable_n = reachable.reachable_n;
-
-Goal "starts_of(restrict ioa acts) = starts_of(ioa) & \
-\ trans_of(restrict ioa acts) = trans_of(ioa)";
-by (simp_tac (simpset() addsimps ([restrict_def]@ioa_projections)) 1);
-qed "cancel_restrict_a";
-
-Goal "reachable (restrict ioa acts) s = reachable ioa s";
-by (rtac iffI 1);
-by (etac reachable.induct 1);
-by (asm_full_simp_tac (simpset() addsimps [cancel_restrict_a,reachable_0]) 1);
-by (etac reachable_n 1);
-by (asm_full_simp_tac (simpset() addsimps [cancel_restrict_a]) 1);
-(* <-- *)
-by (etac reachable.induct 1);
-by (rtac reachable_0 1);
-by (asm_full_simp_tac (simpset() addsimps [cancel_restrict_a]) 1);
-by (etac reachable_n 1);
-by (asm_full_simp_tac (simpset() addsimps [cancel_restrict_a]) 1);
-qed "cancel_restrict_b";
-
-Goal "act (restrict A acts) = act A";
-by (simp_tac (simpset() addsimps [actions_def,asig_internals_def,
- asig_outputs_def,asig_inputs_def,externals_def,asig_of_def,restrict_def,
- restrict_asig_def]) 1);
-by Auto_tac;
-qed"acts_restrict";
-
-Goal "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_tac (simpset() addsimps [cancel_restrict_a,cancel_restrict_b,acts_restrict]) 1);
-qed"cancel_restrict";
-
-(* ------------------------------------------------------------------------- *)
-
-section "rename";
-
-
-
-Goal "s -a--(rename C f)-> t ==> (? x. Some(x) = f(a) & s -x--C-> t)";
-by (asm_full_simp_tac (simpset() addsimps [Let_def,rename_def,trans_of_def]) 1);
-qed"trans_rename";
-
-
-Goal "[| reachable (rename C g) s |] ==> reachable C s";
-by (etac reachable.induct 1);
-by (rtac reachable_0 1);
-by (asm_full_simp_tac (simpset() addsimps [rename_def]@ioa_projections) 1);
-by (dtac trans_rename 1);
-by (etac exE 1);
-by (etac conjE 1);
-by (etac reachable_n 1);
-by (assume_tac 1);
-qed"reachable_rename";
-
-
-
-(* ------------------------------------------------------------------------- *)
-
-section "trans_of(A||B)";
-
-
-Goal "[|(s,a,t):trans_of (A||B); a:act A|] \
-\ ==> (fst s,a,fst t):trans_of A";
-by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
-qed"trans_A_proj";
-
-Goal "[|(s,a,t):trans_of (A||B); a:act B|] \
-\ ==> (snd s,a,snd t):trans_of B";
-by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
-qed"trans_B_proj";
-
-Goal "[|(s,a,t):trans_of (A||B); a~:act A|]\
-\ ==> fst s = fst t";
-by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
-qed"trans_A_proj2";
-
-Goal "[|(s,a,t):trans_of (A||B); a~:act B|]\
-\ ==> snd s = snd t";
-by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
-qed"trans_B_proj2";
-
-Goal "(s,a,t):trans_of (A||B) \
-\ ==> a :act A | a :act B";
-by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
-qed"trans_AB_proj";
-
-Goal "[|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||B)";
-by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
-qed"trans_AB";
-
-Goal "[|a:act A;a~:act B;\
-\ (fst s,a,fst t):trans_of A;snd s=snd t|]\
-\ ==> (s,a,t):trans_of (A||B)";
-by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
-qed"trans_A_notB";
-
-Goal "[|a~:act A;a:act B;\
-\ (snd s,a,snd t):trans_of B;fst s=fst t|]\
-\ ==> (s,a,t):trans_of (A||B)";
-by (asm_full_simp_tac (simpset() addsimps [Let_def,par_def,trans_of_def]) 1);
-qed"trans_notA_B";
-
-val trans_of_defs1 = [trans_AB,trans_A_notB,trans_notA_B];
-val trans_of_defs2 = [trans_A_proj,trans_B_proj,trans_A_proj2,
- trans_B_proj2,trans_AB_proj];
-
-
-Goal
-"((s,a,t) : trans_of(A || B || C || 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_tac (simpset() addsimps ([par_def,actions_asig_comp,Pair_fst_snd_eq,Let_def]@
- ioa_projections)) 1);
-qed "trans_of_par4";
-
-
-(* ------------------------------------------------------------------------- *)
-
-section "proof obligation generator for IOA requirements";
-
-(* without assumptions on A and B because is_trans_of is also incorporated in ||def *)
-Goalw [is_trans_of_def] "is_trans_of (A||B)";
-by (simp_tac (simpset() addsimps [Let_def,actions_of_par,trans_of_par]) 1);
-qed"is_trans_of_par";
-
-Goalw [is_trans_of_def]
-"is_trans_of A ==> is_trans_of (restrict A acts)";
-by (asm_simp_tac (simpset() addsimps [cancel_restrict,acts_restrict])1);
-qed"is_trans_of_restrict";
-
-Goalw [is_trans_of_def,restrict_def,restrict_asig_def]
-"is_trans_of A ==> is_trans_of (rename A f)";
-by (asm_full_simp_tac
- (simpset() addsimps [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]) 1);
-by (Blast_tac 1);
-qed"is_trans_of_rename";
-
-Goal "[| is_asig_of A; is_asig_of B; compatible A B|] \
-\ ==> is_asig_of (A||B)";
-by (asm_full_simp_tac (simpset() addsimps [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]) 1);
-by (asm_full_simp_tac (simpset() addsimps [asig_of_def]) 1);
-by Auto_tac;
-qed"is_asig_of_par";
-
-Goalw [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]
-"is_asig_of A ==> is_asig_of (restrict A f)";
-by (Asm_full_simp_tac 1);
-by Auto_tac;
-qed"is_asig_of_restrict";
-
-Goal "is_asig_of A ==> is_asig_of (rename A f)";
-by (asm_full_simp_tac (simpset() addsimps [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]) 1);
-by Auto_tac;
-by (ALLGOALS(dres_inst_tac [("s","Some ?x")] sym THEN' Asm_full_simp_tac));
-by (ALLGOALS(Blast_tac));
-qed"is_asig_of_rename";
-
-
-Addsimps [is_asig_of_par,is_asig_of_restrict,is_asig_of_rename,
- is_trans_of_par,is_trans_of_restrict,is_trans_of_rename];
-
-
-Goalw [compatible_def]
-"[|compatible A B; compatible A C |]==> compatible A (B||C)";
-by (asm_full_simp_tac (simpset() addsimps [internals_of_par,
- outputs_of_par,actions_of_par]) 1);
-by Auto_tac;
-qed"compatible_par";
-
-(* better derive by previous one and compat_commute *)
-Goalw [compatible_def]
-"[|compatible A C; compatible B C |]==> compatible (A||B) C";
-by (asm_full_simp_tac (simpset() addsimps [internals_of_par,
- outputs_of_par,actions_of_par]) 1);
-by Auto_tac;
-qed"compatible_par2";
-
-Goalw [compatible_def]
-"[| compatible A B; (ext B - S) Int ext A = {}|] \
-\ ==> compatible A (restrict B S)";
-by (asm_full_simp_tac (simpset() addsimps [ioa_triple_proj,asig_triple_proj,
- externals_def,restrict_def,restrict_asig_def,actions_def]) 1);
-by Auto_tac;
-qed"compatible_restrict";
-
-
-Addsimps [split_paired_Ex];
--- a/src/HOLCF/IOA/meta_theory/Automata.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/Automata.thy Sun May 28 19:54:20 2006 +0200
@@ -252,6 +252,453 @@
set_was_enabled_def:
"set_was_enabled A W t == ? w:W. was_enabled A w t"
-ML {* use_legacy_bindings (the_context ()) *}
+
+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 || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}"
+ apply (simp add: par_def ioa_projections)
+done
+
+lemma trans_of_par:
+"trans_of(A || 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))}"
+
+apply (simp add: par_def ioa_projections)
+done
+
+
+subsection "actions and par"
+
+lemma actions_asig_comp:
+ "actions(asig_comp a b) = actions(a) Un actions(b)"
+ apply (simp (no_asm) add: actions_def asig_comp_def asig_projections)
+ apply blast
+ done
+
+lemma asig_of_par: "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)"
+ apply (simp add: par_def ioa_projections)
+ done
+
+
+lemma externals_of_par: "ext (A1||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_def)
+apply blast
+done
+
+lemma actions_of_par: "act (A1||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_def)
+apply blast
+done
+
+lemma inputs_of_par: "inp (A1||A2) =
+ ((inp A1) Un (inp A2)) - ((out A1) Un (out A2))"
+apply (simp add: actions_def asig_of_par asig_comp_def
+ asig_inputs_def asig_outputs_def Un_def set_diff_def)
+done
+
+lemma outputs_of_par: "out (A1||A2) =
+ (out A1) Un (out A2)"
+apply (simp add: actions_def asig_of_par asig_comp_def
+ asig_outputs_def Un_def set_diff_def)
+done
+
+lemma internals_of_par: "int (A1||A2) =
+ (int A1) Un (int 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_def)
+done
+
+
+subsection "actions and compatibility"
+
+lemma compat_commute: "compatible A B = compatible B A"
+apply (simp add: compatible_def Int_commute)
+apply auto
+done
+
+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, standard]
+lemmas ext1_ext2_is_not_act1 = ext2_is_not_int1 [THEN int_and_ext_is_act, standard]
+
+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||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||B)"
+apply (unfold input_enabled_def)
+apply (simp add: Let_def inputs_of_par trans_of_par)
+apply (tactic "safe_tac set_cs")
+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 xa = "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)"
+apply (simp add: restrict_def ioa_projections)
+done
+
+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"
+ apply (simp (no_asm) add: cancel_restrict_a cancel_restrict_b acts_restrict)
+ done
+
+
+subsection "rename"
+
+lemma trans_rename: "s -a--(rename C f)-> t ==> (? x. Some(x) = f(a) & s -x--C-> t)"
+apply (simp add: Let_def rename_def trans_of_def)
+done
+
+
+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||B)"
+
+
+lemma trans_A_proj: "[|(s,a,t):trans_of (A||B); a:act A|]
+ ==> (fst s,a,fst t):trans_of A"
+apply (simp add: Let_def par_def trans_of_def)
+done
+
+lemma trans_B_proj: "[|(s,a,t):trans_of (A||B); a:act B|]
+ ==> (snd s,a,snd t):trans_of B"
+apply (simp add: Let_def par_def trans_of_def)
+done
+
+lemma trans_A_proj2: "[|(s,a,t):trans_of (A||B); a~:act A|]
+ ==> fst s = fst t"
+apply (simp add: Let_def par_def trans_of_def)
+done
+
+lemma trans_B_proj2: "[|(s,a,t):trans_of (A||B); a~:act B|]
+ ==> snd s = snd t"
+apply (simp add: Let_def par_def trans_of_def)
+done
+
+lemma trans_AB_proj: "(s,a,t):trans_of (A||B)
+ ==> a :act A | a :act B"
+apply (simp add: Let_def par_def trans_of_def)
+done
+
+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||B)"
+apply (simp add: Let_def par_def trans_of_def)
+done
+
+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||B)"
+apply (simp add: Let_def par_def trans_of_def)
+done
+
+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||B)"
+apply (simp add: Let_def par_def trans_of_def)
+done
+
+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 || B || C || 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)))))"
+ apply (simp (no_asm) add: par_def actions_asig_comp Pair_fst_snd_eq Let_def ioa_projections)
+ done
+
+
+subsection "proof obligation generator for IOA requirements"
+
+(* without assumptions on A and B because is_trans_of is also incorporated in ||def *)
+lemma is_trans_of_par: "is_trans_of (A||B)"
+apply (unfold is_trans_of_def)
+apply (simp add: Let_def actions_of_par trans_of_par)
+done
+
+lemma is_trans_of_restrict:
+"is_trans_of A ==> is_trans_of (restrict A acts)"
+apply (unfold is_trans_of_def)
+apply (simp add: cancel_restrict acts_restrict)
+done
+
+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||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 ?x" 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||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||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/HOLCF/IOA/meta_theory/CompoExecs.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,263 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/CompoExecs.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-Delsimps (ex_simps @ all_simps);
-Delsimps [split_paired_All];
-
-(* ----------------------------------------------------------------------------------- *)
-
-
-section "recursive equations of operators";
-
-
-(* ---------------------------------------------------------------- *)
-(* ProjA2 *)
-(* ---------------------------------------------------------------- *)
-
-
-Goal "ProjA2$UU = UU";
-by (simp_tac (simpset() addsimps [ProjA2_def]) 1);
-qed"ProjA2_UU";
-
-Goal "ProjA2$nil = nil";
-by (simp_tac (simpset() addsimps [ProjA2_def]) 1);
-qed"ProjA2_nil";
-
-Goal "ProjA2$((a,t)>>xs) = (a,fst t) >> ProjA2$xs";
-by (simp_tac (simpset() addsimps [ProjA2_def]) 1);
-qed"ProjA2_cons";
-
-
-(* ---------------------------------------------------------------- *)
-(* ProjB2 *)
-(* ---------------------------------------------------------------- *)
-
-
-Goal "ProjB2$UU = UU";
-by (simp_tac (simpset() addsimps [ProjB2_def]) 1);
-qed"ProjB2_UU";
-
-Goal "ProjB2$nil = nil";
-by (simp_tac (simpset() addsimps [ProjB2_def]) 1);
-qed"ProjB2_nil";
-
-Goal "ProjB2$((a,t)>>xs) = (a,snd t) >> ProjB2$xs";
-by (simp_tac (simpset() addsimps [ProjB2_def]) 1);
-qed"ProjB2_cons";
-
-
-
-(* ---------------------------------------------------------------- *)
-(* Filter_ex2 *)
-(* ---------------------------------------------------------------- *)
-
-
-Goal "Filter_ex2 sig$UU=UU";
-by (simp_tac (simpset() addsimps [Filter_ex2_def]) 1);
-qed"Filter_ex2_UU";
-
-Goal "Filter_ex2 sig$nil=nil";
-by (simp_tac (simpset() addsimps [Filter_ex2_def]) 1);
-qed"Filter_ex2_nil";
-
-Goal "Filter_ex2 sig$(at >> xs) = \
-\ (if (fst at:actions sig) \
-\ then at >> (Filter_ex2 sig$xs) \
-\ else Filter_ex2 sig$xs)";
-
-by (asm_full_simp_tac (simpset() addsimps [Filter_ex2_def]) 1);
-qed"Filter_ex2_cons";
-
-
-(* ---------------------------------------------------------------- *)
-(* stutter2 *)
-(* ---------------------------------------------------------------- *)
-
-
-Goal "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 fi) \
-\ andalso (stutter2 sig$xs) (snd p)) \
-\ $x) \
-\ ))";
-by (rtac trans 1);
-by (rtac fix_eq2 1);
-by (rtac stutter2_def 1);
-by (rtac beta_cfun 1);
-by (simp_tac (simpset() addsimps [flift1_def]) 1);
-qed"stutter2_unfold";
-
-Goal "(stutter2 sig$UU) s=UU";
-by (stac stutter2_unfold 1);
-by (Simp_tac 1);
-qed"stutter2_UU";
-
-Goal "(stutter2 sig$nil) s = TT";
-by (stac stutter2_unfold 1);
-by (Simp_tac 1);
-qed"stutter2_nil";
-
-Goal "(stutter2 sig$(at>>xs)) s = \
-\ ((if (fst at)~:actions sig then Def (s=snd at) else TT) \
-\ andalso (stutter2 sig$xs) (snd at))";
-by (rtac trans 1);
-by (stac stutter2_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,flift1_def,If_and_if]) 1);
-by (Simp_tac 1);
-qed"stutter2_cons";
-
-
-Addsimps [stutter2_UU,stutter2_nil,stutter2_cons];
-
-
-(* ---------------------------------------------------------------- *)
-(* stutter *)
-(* ---------------------------------------------------------------- *)
-
-Goal "stutter sig (s, UU)";
-by (simp_tac (simpset() addsimps [stutter_def]) 1);
-qed"stutter_UU";
-
-Goal "stutter sig (s, nil)";
-by (simp_tac (simpset() addsimps [stutter_def]) 1);
-qed"stutter_nil";
-
-Goal "stutter sig (s, (a,t)>>ex) = \
-\ ((a~:actions sig --> (s=t)) & stutter sig (t,ex))";
-by (simp_tac (simpset() addsimps [stutter_def]) 1);
-qed"stutter_cons";
-
-(* ----------------------------------------------------------------------------------- *)
-
-Delsimps [stutter2_UU,stutter2_nil,stutter2_cons];
-
-val 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];
-
-Addsimps compoex_simps;
-
-
-
-(* ------------------------------------------------------------------ *)
-(* The following lemmata aim for *)
-(* COMPOSITIONALITY on EXECUTION Level *)
-(* ------------------------------------------------------------------ *)
-
-
-(* --------------------------------------------------------------------- *)
-(* Lemma_1_1a : is_ex_fr propagates from A||B to Projections A and B *)
-(* --------------------------------------------------------------------- *)
-
-Goal "!s. is_exec_frag (A||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))";
-
-by (pair_induct_tac "xs" [is_exec_frag_def] 1);
-(* main case *)
-by (rename_tac "ss a t" 1);
-by (safe_tac set_cs);
-by (REPEAT (asm_full_simp_tac (simpset() addsimps trans_of_defs2) 1));
-qed"lemma_1_1a";
-
-
-(* --------------------------------------------------------------------- *)
-(* Lemma_1_1b : is_ex_fr (A||B) implies stuttering on Projections *)
-(* --------------------------------------------------------------------- *)
-
-Goal "!s. is_exec_frag (A||B) (s,xs) \
-\ --> stutter (asig_of A) (fst s,ProjA2$xs) &\
-\ stutter (asig_of B) (snd s,ProjB2$xs)";
-
-by (pair_induct_tac "xs" [stutter_def,is_exec_frag_def] 1);
-(* main case *)
-by (safe_tac set_cs);
-by (REPEAT (asm_full_simp_tac (simpset() addsimps trans_of_defs2) 1));
-qed"lemma_1_1b";
-
-
-(* --------------------------------------------------------------------- *)
-(* Lemma_1_1c : Executions of A||B have only A- or B-actions *)
-(* --------------------------------------------------------------------- *)
-
-Goal "!s. (is_exec_frag (A||B) (s,xs) \
-\ --> Forall (%x. fst x:act (A||B)) xs)";
-
-by (pair_induct_tac "xs" [Forall_def,sforall_def,is_exec_frag_def] 1);
-(* main case *)
-by (safe_tac set_cs);
-by (REPEAT (asm_full_simp_tac (simpset() addsimps trans_of_defs2 @
- [actions_asig_comp,asig_of_par]) 1));
-qed"lemma_1_1c";
-
-
-(* ----------------------------------------------------------------------- *)
-(* Lemma_1_2 : ex A, exB, stuttering and forall a:A|B implies ex (A||B) *)
-(* ----------------------------------------------------------------------- *)
-
-Goal
-"!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||B)) xs \
-\ --> is_exec_frag (A||B) (s,xs)";
-
-by (pair_induct_tac "xs" [Forall_def,sforall_def,
- is_exec_frag_def, stutter_def] 1);
-by (asm_full_simp_tac (simpset() addsimps trans_of_defs1 @ [actions_asig_comp,asig_of_par]) 1);
-by (safe_tac set_cs);
-(* problem with asm_full_simp_tac: (fst s) = (fst t) is too high in assumption order ! *)
-by (rotate_tac ~4 1);
-by (Asm_full_simp_tac 1);
-by (rotate_tac ~4 1);
-by (Asm_full_simp_tac 1);
-qed"lemma_1_2";
-
-
-(* ------------------------------------------------------------------ *)
-(* COMPOSITIONALITY on EXECUTION Level *)
-(* Main Theorem *)
-(* ------------------------------------------------------------------ *)
-
-
-Goal
-"(ex:executions(A||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||B)) (snd ex))";
-
-by (simp_tac (simpset() addsimps [executions_def,ProjB_def,
- Filter_ex_def,ProjA_def,starts_of_par]) 1);
-by (pair_tac "ex" 1);
-by (rtac iffI 1);
-(* ==> *)
-by (REPEAT (etac conjE 1));
-by (asm_full_simp_tac (simpset() addsimps [lemma_1_1a RS spec RS mp,
- lemma_1_1b RS spec RS mp,lemma_1_1c RS spec RS mp]) 1);
-(* <== *)
-by (REPEAT (etac conjE 1));
-by (asm_full_simp_tac (simpset() addsimps [lemma_1_2 RS spec RS mp]) 1);
-qed"compositionality_ex";
-
-
-(* ------------------------------------------------------------------ *)
-(* COMPOSITIONALITY on EXECUTION Level *)
-(* For Modules *)
-(* ------------------------------------------------------------------ *)
-
-Goalw [Execs_def,par_execs_def]
-
-"Execs (A||B) = par_execs (Execs A) (Execs B)";
-
-by (asm_full_simp_tac (simpset() addsimps [asig_of_par]) 1);
-by (rtac set_ext 1);
-by (asm_full_simp_tac (simpset() addsimps [compositionality_ex,actions_of_par]) 1);
-qed"compositionality_ex_modules";
--- a/src/HOLCF/IOA/meta_theory/CompoExecs.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/CompoExecs.thy Sun May 28 19:54:20 2006 +0200
@@ -73,6 +73,249 @@
Int {ex. Forall (%x. fst x:(actions sigA Un actions sigB)) (snd ex)},
asig_comp sigA sigB)"
-ML {* use_legacy_bindings (the_context ()) *}
+
+
+lemmas [simp del] = ex_simps all_simps 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)>>xs) = (a,fst t) >> 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)>>xs) = (a,snd t) >> 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 >> xs) =
+ (if (fst at:actions sig)
+ then at >> (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 fi)
+ andalso (stutter2 sig$xs) (snd p))
+ $x)
+ ))"
+apply (rule trans)
+apply (rule fix_eq2)
+apply (rule 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>>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)>>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||B to Projections A and B *)
+(* --------------------------------------------------------------------- *)
+
+lemma lemma_1_1a: "!s. is_exec_frag (A||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 {* pair_induct_tac "xs" [thm "is_exec_frag_def"] 1 *})
+(* main case *)
+apply (rename_tac ss a t)
+apply (tactic "safe_tac set_cs")
+apply (simp_all add: trans_of_defs2)
+done
+
+
+(* --------------------------------------------------------------------- *)
+(* Lemma_1_1b : is_ex_fr (A||B) implies stuttering on Projections *)
+(* --------------------------------------------------------------------- *)
+
+lemma lemma_1_1b: "!s. is_exec_frag (A||B) (s,xs)
+ --> stutter (asig_of A) (fst s,ProjA2$xs) &
+ stutter (asig_of B) (snd s,ProjB2$xs)"
+
+apply (tactic {* pair_induct_tac "xs" [thm "stutter_def", thm "is_exec_frag_def"] 1 *})
+(* main case *)
+apply (tactic "safe_tac set_cs")
+apply (simp_all add: trans_of_defs2)
+done
+
+
+(* --------------------------------------------------------------------- *)
+(* Lemma_1_1c : Executions of A||B have only A- or B-actions *)
+(* --------------------------------------------------------------------- *)
+
+lemma lemma_1_1c: "!s. (is_exec_frag (A||B) (s,xs)
+ --> Forall (%x. fst x:act (A||B)) xs)"
+
+apply (tactic {* pair_induct_tac "xs" [thm "Forall_def", thm "sforall_def",
+ thm "is_exec_frag_def"] 1 *})
+(* main case *)
+apply (tactic "safe_tac set_cs")
+apply (simp_all 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||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||B)) xs
+ --> is_exec_frag (A||B) (s,xs)"
+
+apply (tactic {* pair_induct_tac "xs" [thm "Forall_def", thm "sforall_def",
+ thm "is_exec_frag_def", thm "stutter_def"] 1 *})
+apply (simp add: trans_of_defs1 actions_asig_comp asig_of_par)
+apply (tactic "safe_tac set_cs")
+apply simp
+apply simp
+done
+
+
+subsection {* COMPOSITIONALITY on EXECUTION Level -- Main Theorem *}
+
+lemma compositionality_ex:
+"(ex:executions(A||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||B)) (snd ex))"
+
+apply (simp (no_asm) add: executions_def ProjB_def Filter_ex_def ProjA_def starts_of_par)
+apply (tactic {* pair_tac "ex" 1 *})
+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 {* COMPOSITIONALITY on EXECUTION Level -- for Modules *}
+
+lemma compositionality_ex_modules:
+ "Execs (A||B) = par_execs (Execs A) (Execs B)"
+apply (unfold Execs_def par_execs_def)
+apply (simp add: asig_of_par)
+apply (rule set_ext)
+apply (simp add: compositionality_ex actions_of_par)
+done
end
--- a/src/HOLCF/IOA/meta_theory/CompoScheds.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,509 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/CompoScheds.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-Addsimps [surjective_pairing RS sym];
-
-
-
-(* ------------------------------------------------------------------------------- *)
-
-section "mkex rewrite rules";
-
-(* ---------------------------------------------------------------- *)
-(* mkex2 *)
-(* ---------------------------------------------------------------- *)
-
-
-bind_thm ("mkex2_unfold", fix_prover2 (the_context ()) mkex2_def
-"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))>> \
-\ (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))>>(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))>>(mkex2 A B$xs$exA$(TL$exB)) s (snd b)) \
-\ else \
-\ UU \
-\ ) \
-\ ) \
-\ )))");
-
-
-Goal "(mkex2 A B$UU$exA$exB) s t = UU";
-by (stac mkex2_unfold 1);
-by (Simp_tac 1);
-qed"mkex2_UU";
-
-Goal "(mkex2 A B$nil$exA$exB) s t= nil";
-by (stac mkex2_unfold 1);
-by (Simp_tac 1);
-qed"mkex2_nil";
-
-Goal "[| x:act A; x~:act B; HD$exA=Def a|] \
-\ ==> (mkex2 A B$(x>>sch)$exA$exB) s t = \
-\ (x,snd a,t) >> (mkex2 A B$sch$(TL$exA)$exB) (snd a) t";
-by (rtac trans 1);
-by (stac mkex2_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,If_and_if]) 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def]) 1);
-qed"mkex2_cons_1";
-
-Goal "[| x~:act A; x:act B; HD$exB=Def b|] \
-\ ==> (mkex2 A B$(x>>sch)$exA$exB) s t = \
-\ (x,s,snd b) >> (mkex2 A B$sch$exA$(TL$exB)) s (snd b)";
-by (rtac trans 1);
-by (stac mkex2_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,If_and_if]) 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def]) 1);
-qed"mkex2_cons_2";
-
-Goal "[| x:act A; x:act B; HD$exA=Def a;HD$exB=Def b|] \
-\ ==> (mkex2 A B$(x>>sch)$exA$exB) s t = \
-\ (x,snd a,snd b) >> \
-\ (mkex2 A B$sch$(TL$exA)$(TL$exB)) (snd a) (snd b)";
-by (rtac trans 1);
-by (stac mkex2_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,If_and_if]) 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def]) 1);
-qed"mkex2_cons_3";
-
-Addsimps [mkex2_UU,mkex2_nil,mkex2_cons_1,mkex2_cons_2,mkex2_cons_3];
-
-
-(* ---------------------------------------------------------------- *)
-(* mkex *)
-(* ---------------------------------------------------------------- *)
-
-Goal "mkex A B UU (s,exA) (t,exB) = ((s,t),UU)";
-by (simp_tac (simpset() addsimps [mkex_def]) 1);
-qed"mkex_UU";
-
-Goal "mkex A B nil (s,exA) (t,exB) = ((s,t),nil)";
-by (simp_tac (simpset() addsimps [mkex_def]) 1);
-qed"mkex_nil";
-
-Goal "[| x:act A; x~:act B |] \
-\ ==> mkex A B (x>>sch) (s,a>>exA) (t,exB) = \
-\ ((s,t), (x,snd a,t) >> snd (mkex A B sch (snd a,exA) (t,exB)))";
-by (simp_tac (simpset() addsimps [mkex_def]) 1);
-by (cut_inst_tac [("exA","a>>exA")] mkex2_cons_1 1);
-by Auto_tac;
-qed"mkex_cons_1";
-
-Goal "[| x~:act A; x:act B |] \
-\ ==> mkex A B (x>>sch) (s,exA) (t,b>>exB) = \
-\ ((s,t), (x,s,snd b) >> snd (mkex A B sch (s,exA) (snd b,exB)))";
-by (simp_tac (simpset() addsimps [mkex_def]) 1);
-by (cut_inst_tac [("exB","b>>exB")] mkex2_cons_2 1);
-by Auto_tac;
-qed"mkex_cons_2";
-
-Goal "[| x:act A; x:act B |] \
-\ ==> mkex A B (x>>sch) (s,a>>exA) (t,b>>exB) = \
-\ ((s,t), (x,snd a,snd b) >> snd (mkex A B sch (snd a,exA) (snd b,exB)))";
-by (simp_tac (simpset() addsimps [mkex_def]) 1);
-by (cut_inst_tac [("exB","b>>exB"),("exA","a>>exA")] mkex2_cons_3 1);
-by Auto_tac;
-qed"mkex_cons_3";
-
-Delsimps [mkex2_UU,mkex2_nil,mkex2_cons_1,mkex2_cons_2,mkex2_cons_3];
-
-val composch_simps = [mkex_UU,mkex_nil,
- mkex_cons_1,mkex_cons_2,mkex_cons_3];
-
-Addsimps composch_simps;
-
-
-
-(* ------------------------------------------------------------------ *)
-(* The following lemmata aim for *)
-(* COMPOSITIONALITY on SCHEDULE Level *)
-(* ------------------------------------------------------------------ *)
-
-(* ---------------------------------------------------------------------- *)
- section "Lemmas for ==>";
-(* ----------------------------------------------------------------------*)
-
-(* --------------------------------------------------------------------- *)
-(* Lemma_2_1 : tfilter(ex) and filter_act are commutative *)
-(* --------------------------------------------------------------------- *)
-
-Goalw [filter_act_def,Filter_ex2_def]
- "filter_act$(Filter_ex2 (asig_of A)$xs)=\
-\ Filter (%a. a:act A)$(filter_act$xs)";
-
-by (simp_tac (simpset() addsimps [MapFilter,o_def]) 1);
-qed"lemma_2_1a";
-
-
-(* --------------------------------------------------------------------- *)
-(* Lemma_2_2 : State-projections do not affect filter_act *)
-(* --------------------------------------------------------------------- *)
-
-Goal
- "filter_act$(ProjA2$xs) =filter_act$xs &\
-\ filter_act$(ProjB2$xs) =filter_act$xs";
-
-by (pair_induct_tac "xs" [] 1);
-qed"lemma_2_1b";
-
-
-(* --------------------------------------------------------------------- *)
-(* Schedules of A||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 *)
-
-Goal "!s. is_exec_frag (A||B) (s,xs) \
-\ --> Forall (%x. x:act (A||B)) (filter_act$xs)";
-
-by (pair_induct_tac "xs" [is_exec_frag_def,Forall_def,sforall_def] 1);
-(* main case *)
-by (safe_tac set_cs);
-by (REPEAT (asm_full_simp_tac (simpset() addsimps trans_of_defs2 @
- [actions_asig_comp,asig_of_par]) 1));
-qed"sch_actions_in_AorB";
-
-
-(* --------------------------------------------------------------------------*)
- section "Lemmas for <==";
-(* ---------------------------------------------------------------------------*)
-
-(*---------------------------------------------------------------------------
- Filtering actions out of mkex(sch,exA,exB) yields the oracle sch
- structural induction
- --------------------------------------------------------------------------- *)
-
-Goal "! exA exB s t. \
-\ Forall (%x. x:act (A||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";
-
-by (Seq_induct_tac "sch" [Filter_def,Forall_def,sforall_def,mkex_def] 1);
-
-(* main case *)
-(* splitting into 4 cases according to a:A, a:B *)
-by (Asm_full_simp_tac 1);
-by (safe_tac set_cs);
-
-(* Case y:A, y:B *)
-by (Seq_case_simp_tac "exA" 1);
-(* 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>>ss<< UU(nil), using theorems
- Cons_not_less_UU and Cons_not_less_nil *)
-by (Seq_case_simp_tac "exB" 1);
-(* Case exA=a>>x, exB=b>>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 *)
-by (Asm_full_simp_tac 1);
-
-(* Case y:A, y~:B *)
-by (Seq_case_simp_tac "exA" 1);
-by (Asm_full_simp_tac 1);
-
-(* Case y~:A, y:B *)
-by (Seq_case_simp_tac "exB" 1);
-by (Asm_full_simp_tac 1);
-
-(* Case y~:A, y~:B *)
-by (asm_full_simp_tac (simpset() addsimps [asig_of_par,actions_asig_comp]) 1);
-qed"Mapfst_mkex_is_sch";
-
-
-(* generalizing the proof above to a tactic *)
-
-fun mkex_induct_tac sch exA exB =
- EVERY1[Seq_induct_tac sch [Filter_def,Forall_def,sforall_def,mkex_def,stutter_def],
- Asm_full_simp_tac,
- SELECT_GOAL (safe_tac set_cs),
- Seq_case_simp_tac exA,
- Seq_case_simp_tac exB,
- Asm_full_simp_tac,
- Seq_case_simp_tac exA,
- Asm_full_simp_tac,
- Seq_case_simp_tac exB,
- Asm_full_simp_tac,
- asm_full_simp_tac (simpset() addsimps [asig_of_par,actions_asig_comp])
- ];
-
-
-
-(*---------------------------------------------------------------------------
- Projection of mkex(sch,exA,exB) onto A stutters on A
- structural induction
- --------------------------------------------------------------------------- *)
-
-
-Goal "! exA exB s t. \
-\ Forall (%x. x:act (A||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_tac "sch" "exA" "exB");
-
-qed"stutterA_mkex";
-
-
-Goal "[| \
-\ Forall (%x. x:act (A||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))";
-
-by (cut_facts_tac [stutterA_mkex] 1);
-by (asm_full_simp_tac (simpset() addsimps [stutter_def,ProjA_def,mkex_def]) 1);
-by (REPEAT (etac allE 1));
-by (dtac mp 1);
-by (assume_tac 2);
-by (Asm_full_simp_tac 1);
-qed"stutter_mkex_on_A";
-
-
-(*---------------------------------------------------------------------------
- Projection of mkex(sch,exA,exB) onto B stutters on B
- structural induction
- --------------------------------------------------------------------------- *)
-
-Goal "! exA exB s t. \
-\ Forall (%x. x:act (A||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_tac "sch" "exA" "exB");
-
-qed"stutterB_mkex";
-
-
-Goal "[| \
-\ Forall (%x. x:act (A||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))";
-
-by (cut_facts_tac [stutterB_mkex] 1);
-by (asm_full_simp_tac (simpset() addsimps [stutter_def,ProjB_def,mkex_def]) 1);
-by (REPEAT (etac allE 1));
-by (dtac mp 1);
-by (assume_tac 2);
-by (Asm_full_simp_tac 1);
-qed"stutter_mkex_on_B";
-
-
-(*---------------------------------------------------------------------------
- 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
- --------------------------------------------------------------------------- *)
-
-Goal "! exA exB s t. \
-\ Forall (%x. x:act (A||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_tac "sch" "exB" "exA");
-
-qed"filter_mkex_is_exA_tmp";
-
-(*---------------------------------------------------------------------------
- zip$(proj1$y)$(proj2$y) = y (using the lift operations)
- lemma for admissibility problems
- --------------------------------------------------------------------------- *)
-
-Goal "Zip$(Map fst$y)$(Map snd$y) = y";
-by (Seq_induct_tac "y" [] 1);
-qed"Zip_Map_fst_snd";
-
-
-(*---------------------------------------------------------------------------
- filter A$sch = proj1$ex --> zip$(filter A$sch)$(proj2$ex) = ex
- lemma for eliminating non admissible equations in assumptions
- --------------------------------------------------------------------------- *)
-
-Goal "!! sch ex. \
-\ Filter (%a. a:act AB)$sch = filter_act$ex \
-\ ==> ex = Zip$(Filter (%a. a:act AB)$sch)$(Map snd$ex)";
-by (asm_full_simp_tac (simpset() addsimps [filter_act_def]) 1);
-by (rtac (Zip_Map_fst_snd RS sym) 1);
-qed"trick_against_eq_in_ass";
-
-(*---------------------------------------------------------------------------
- Filter of mkex(sch,exA,exB) to A after projection onto A is exA
- using the above trick
- --------------------------------------------------------------------------- *)
-
-
-Goal "!!sch exA exB.\
-\ [| Forall (%a. a:act (A||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";
-by (asm_full_simp_tac (simpset() addsimps [ProjA_def,Filter_ex_def]) 1);
-by (pair_tac "exA" 1);
-by (pair_tac "exB" 1);
-by (rtac conjI 1);
-by (simp_tac (simpset() addsimps [mkex_def]) 1);
-by (stac trick_against_eq_in_ass 1);
-back();
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [filter_mkex_is_exA_tmp]) 1);
-qed"filter_mkex_is_exA";
-
-
-(*---------------------------------------------------------------------------
- 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
- --------------------------------------------------------------------------- *)
-
-
-Goal "! exA exB s t. \
-\ Forall (%x. x:act (A||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_tac "sch" "exA" "exB");
-
-qed"filter_mkex_is_exB_tmp";
-
-
-(*---------------------------------------------------------------------------
- Filter of mkex(sch,exA,exB) to A after projection onto B is exB
- using the above trick
- --------------------------------------------------------------------------- *)
-
-
-Goal "!!sch exA exB.\
-\ [| Forall (%a. a:act (A||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";
-by (asm_full_simp_tac (simpset() addsimps [ProjB_def,Filter_ex_def]) 1);
-by (pair_tac "exA" 1);
-by (pair_tac "exB" 1);
-by (rtac conjI 1);
-by (simp_tac (simpset() addsimps [mkex_def]) 1);
-by (stac trick_against_eq_in_ass 1);
-back();
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [filter_mkex_is_exB_tmp]) 1);
-qed"filter_mkex_is_exB";
-
-(* --------------------------------------------------------------------- *)
-(* mkex has only A- or B-actions *)
-(* --------------------------------------------------------------------- *)
-
-
-Goal "!s t exA exB. \
-\ Forall (%x. x : act (A || 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 ||B)) \
-\ (snd (mkex A B sch (s,exA) (t,exB)))";
-
-by (mkex_induct_tac "sch" "exA" "exB");
-
-qed"mkex_actions_in_AorB";
-
-
-(* ------------------------------------------------------------------ *)
-(* COMPOSITIONALITY on SCHEDULE Level *)
-(* Main Theorem *)
-(* ------------------------------------------------------------------ *)
-
-Goal
-"(sch : schedules (A||B)) = \
-\ (Filter (%a. a:act A)$sch : schedules A &\
-\ Filter (%a. a:act B)$sch : schedules B &\
-\ Forall (%x. x:act (A||B)) sch)";
-
-by (simp_tac (simpset() addsimps [schedules_def, has_schedule_def]) 1);
-by (safe_tac set_cs);
-(* ==> *)
-by (res_inst_tac [("x","Filter_ex (asig_of A) (ProjA ex)")] bexI 1);
-by (asm_full_simp_tac (simpset() addsimps [compositionality_ex]) 2);
-by (simp_tac (simpset() addsimps [Filter_ex_def,ProjA_def,
- lemma_2_1a,lemma_2_1b]) 1);
-by (res_inst_tac [("x","Filter_ex (asig_of B) (ProjB ex)")] bexI 1);
-by (asm_full_simp_tac (simpset() addsimps [compositionality_ex]) 2);
-by (simp_tac (simpset() addsimps [Filter_ex_def,ProjB_def,
- lemma_2_1a,lemma_2_1b]) 1);
-by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
-by (pair_tac "ex" 1);
-by (etac conjE 1);
-by (asm_full_simp_tac (simpset() addsimps [sch_actions_in_AorB]) 1);
-
-(* <== *)
-
-(* mkex is exactly the construction of exA||B out of exA, exB, and the oracle sch,
- we need here *)
-by (rename_tac "exA exB" 1);
-by (res_inst_tac [("x","mkex A B sch exA exB")] bexI 1);
-(* mkex actions are just the oracle *)
-by (pair_tac "exA" 1);
-by (pair_tac "exB" 1);
-by (asm_full_simp_tac (simpset() addsimps [Mapfst_mkex_is_sch]) 1);
-
-(* mkex is an execution -- use compositionality on ex-level *)
-by (asm_full_simp_tac (simpset() addsimps [compositionality_ex]) 1);
-by (asm_full_simp_tac (simpset() addsimps
- [stutter_mkex_on_A, stutter_mkex_on_B,
- filter_mkex_is_exB,filter_mkex_is_exA]) 1);
-by (pair_tac "exA" 1);
-by (pair_tac "exB" 1);
-by (asm_full_simp_tac (simpset() addsimps [mkex_actions_in_AorB]) 1);
-qed"compositionality_sch";
-
-
-(* ------------------------------------------------------------------ *)
-(* COMPOSITIONALITY on SCHEDULE Level *)
-(* For Modules *)
-(* ------------------------------------------------------------------ *)
-
-Goalw [Scheds_def,par_scheds_def]
-
-"Scheds (A||B) = par_scheds (Scheds A) (Scheds B)";
-
-by (asm_full_simp_tac (simpset() addsimps [asig_of_par]) 1);
-by (rtac set_ext 1);
-by (asm_full_simp_tac (simpset() addsimps [compositionality_sch,actions_of_par]) 1);
-qed"compositionality_sch_modules";
-
-
-Delsimps compoex_simps;
-Delsimps composch_simps;
--- a/src/HOLCF/IOA/meta_theory/CompoScheds.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/CompoScheds.thy Sun May 28 19:54:20 2006 +0200
@@ -70,6 +70,492 @@
Int {sch. Forall (%x. x:(actions sigA Un actions sigB)) sch},
asig_comp sigA sigB)"
-ML {* use_legacy_bindings (the_context ()) *}
+
+declare surjective_pairing [symmetric, simp]
+
+
+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))>>
+ (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))>>(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))>>(mkex2 A B$xs$exA$(TL$exB)) s (snd b))
+ else
+ UU
+ )
+ )
+ )))"
+apply (rule trans)
+apply (rule fix_eq2)
+apply (rule 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>>sch)$exA$exB) s t =
+ (x,snd a,t) >> (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>>sch)$exA$exB) s t =
+ (x,s,snd b) >> (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>>sch)$exA$exB) s t =
+ (x,snd a,snd b) >>
+ (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 {* mkex *}
+
+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>>sch) (s,a>>exA) (t,exB) =
+ ((s,t), (x,snd a,t) >> snd (mkex A B sch (snd a,exA) (t,exB)))"
+apply (simp (no_asm) add: mkex_def)
+apply (cut_tac exA = "a>>exA" in mkex2_cons_1)
+apply auto
+done
+
+lemma mkex_cons_2: "[| x~:act A; x:act B |]
+ ==> mkex A B (x>>sch) (s,exA) (t,b>>exB) =
+ ((s,t), (x,s,snd b) >> snd (mkex A B sch (s,exA) (snd b,exB)))"
+apply (simp (no_asm) add: mkex_def)
+apply (cut_tac exB = "b>>exB" in mkex2_cons_2)
+apply auto
+done
+
+lemma mkex_cons_3: "[| x:act A; x:act B |]
+ ==> mkex A B (x>>sch) (s,a>>exA) (t,b>>exB) =
+ ((s,t), (x,snd a,snd b) >> snd (mkex A B sch (snd a,exA) (snd b,exB)))"
+apply (simp (no_asm) add: mkex_def)
+apply (cut_tac exB = "b>>exB" and exA = "a>>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 {* COMPOSITIONALITY on SCHEDULE Level *}
+
+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 {* pair_induct_tac "xs" [] 1 *})
+done
+
+
+(* --------------------------------------------------------------------- *)
+(* Schedules of A||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||B) (s,xs)
+ --> Forall (%x. x:act (A||B)) (filter_act$xs)"
+
+apply (tactic {* pair_induct_tac "xs" [thm "is_exec_frag_def", thm "Forall_def",
+ thm "sforall_def"] 1 *})
+(* main case *)
+apply (tactic "safe_tac set_cs")
+apply (auto 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||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 {* Seq_induct_tac "sch" [thm "Filter_def", thm "Forall_def",
+ thm "sforall_def", thm "mkex_def"] 1 *})
+
+(* main case *)
+(* splitting into 4 cases according to a:A, a:B *)
+apply (tactic "safe_tac set_cs")
+
+(* Case y:A, y:B *)
+apply (tactic {* Seq_case_simp_tac "exA" 1 *})
+(* 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>>ss<< UU(nil), using theorems
+ Cons_not_less_UU and Cons_not_less_nil *)
+apply (tactic {* Seq_case_simp_tac "exB" 1 *})
+(* Case exA=a>>x, exB=b>>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 {* Seq_case_simp_tac "exA" 1 *})
+apply simp
+
+(* Case y~:A, y:B *)
+apply (tactic {* Seq_case_simp_tac "exB" 1 *})
+apply simp
+
+(* Case y~:A, y~:B *)
+apply (simp add: asig_of_par actions_asig_comp)
+done
+
+
+(* generalizing the proof above to a tactic *)
+
+ML {*
+
+local
+ val defs = [thm "Filter_def", thm "Forall_def", thm "sforall_def", thm "mkex_def",
+ thm "stutter_def"]
+ val asigs = [thm "asig_of_par", thm "actions_asig_comp"]
+in
+
+fun mkex_induct_tac sch exA exB =
+ EVERY1[Seq_induct_tac sch defs,
+ SIMPSET' asm_full_simp_tac,
+ SELECT_GOAL (safe_tac set_cs),
+ Seq_case_simp_tac exA,
+ Seq_case_simp_tac exB,
+ SIMPSET' asm_full_simp_tac,
+ Seq_case_simp_tac exA,
+ SIMPSET' asm_full_simp_tac,
+ Seq_case_simp_tac exB,
+ SIMPSET' asm_full_simp_tac,
+ SIMPSET' (fn ss => asm_full_simp_tac (ss addsimps asigs))
+ ]
end
+*}
+
+
+(*---------------------------------------------------------------------------
+ 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||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))))"
+
+apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *})
+done
+
+
+lemma stutter_mkex_on_A: "[|
+ Forall (%x. x:act (A||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||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))))"
+apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *})
+done
+
+
+lemma stutter_mkex_on_B: "[|
+ Forall (%x. x:act (A||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||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)"
+apply (tactic {* mkex_induct_tac "sch" "exB" "exA" *})
+done
+
+(*---------------------------------------------------------------------------
+ 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 {* Seq_induct_tac "y" [] 1 *})
+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||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 {* pair_tac "exA" 1 *})
+apply (tactic {* pair_tac "exB" 1 *})
+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||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 *)
+apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *})
+done
+
+
+(*---------------------------------------------------------------------------
+ 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||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 {* pair_tac "exA" 1 *})
+apply (tactic {* pair_tac "exB" 1 *})
+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 || 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 ||B))
+ (snd (mkex A B sch (s,exA) (t,exB)))"
+apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *})
+done
+
+
+(* ------------------------------------------------------------------ *)
+(* COMPOSITIONALITY on SCHEDULE Level *)
+(* Main Theorem *)
+(* ------------------------------------------------------------------ *)
+
+lemma compositionality_sch:
+"(sch : schedules (A||B)) =
+ (Filter (%a. a:act A)$sch : schedules A &
+ Filter (%a. a:act B)$sch : schedules B &
+ Forall (%x. x:act (A||B)) sch)"
+apply (simp (no_asm) add: schedules_def has_schedule_def)
+apply (tactic "safe_tac set_cs")
+(* ==> *)
+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 {* pair_tac "ex" 1 *})
+apply (erule conjE)
+apply (simp add: sch_actions_in_AorB)
+
+(* <== *)
+
+(* mkex is exactly the construction of exA||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 {* pair_tac "exA" 1 *})
+apply (tactic {* pair_tac "exB" 1 *})
+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 {* pair_tac "exA" 1 *})
+apply (tactic {* pair_tac "exB" 1 *})
+apply (simp add: mkex_actions_in_AorB)
+done
+
+
+subsection {* COMPOSITIONALITY on SCHEDULE Level -- for Modules *}
+
+lemma compositionality_sch_modules:
+ "Scheds (A||B) = par_scheds (Scheds A) (Scheds B)"
+
+apply (unfold Scheds_def par_scheds_def)
+apply (simp add: asig_of_par)
+apply (rule set_ext)
+apply (simp add: compositionality_sch actions_of_par)
+done
+
+
+declare compoex_simps [simp del]
+declare composch_simps [simp del]
+
+end
--- a/src/HOLCF/IOA/meta_theory/CompoTraces.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1226 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/CompoTraces.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-change_simpset (fn ss => ss setmksym (K NONE));
-
-(* ---------------------------------------------------------------- *)
- section "mksch rewrite rules";
-(* ---------------------------------------------------------------- *)
-
-bind_thm ("mksch_unfold", fix_prover2 (the_context ()) mksch_def
-"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>>(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>>(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>>(mksch A B$xs \
-\ $schA \
-\ $(TL$(Dropwhile (%a. a:int B)$schB)) \
-\ ))) \
-\ else \
-\ UU \
-\ ) \
-\ ) \
-\ ))");
-
-
-Goal "mksch A B$UU$schA$schB = UU";
-by (stac mksch_unfold 1);
-by (Simp_tac 1);
-qed"mksch_UU";
-
-Goal "mksch A B$nil$schA$schB = nil";
-by (stac mksch_unfold 1);
-by (Simp_tac 1);
-qed"mksch_nil";
-
-Goal "[|x:act A;x~:act B|] \
-\ ==> mksch A B$(x>>tr)$schA$schB = \
-\ (Takewhile (%a. a:int A)$schA) \
-\ @@ (x>>(mksch A B$tr$(TL$(Dropwhile (%a. a:int A)$schA)) \
-\ $schB))";
-by (rtac trans 1);
-by (stac mksch_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,If_and_if]) 1);
-by (simp_tac (simpset() addsimps [Consq_def]) 1);
-qed"mksch_cons1";
-
-Goal "[|x~:act A;x:act B|] \
-\ ==> mksch A B$(x>>tr)$schA$schB = \
-\ (Takewhile (%a. a:int B)$schB) \
-\ @@ (x>>(mksch A B$tr$schA$(TL$(Dropwhile (%a. a:int B)$schB)) \
-\ ))";
-by (rtac trans 1);
-by (stac mksch_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,If_and_if]) 1);
-by (simp_tac (simpset() addsimps [Consq_def]) 1);
-qed"mksch_cons2";
-
-Goal "[|x:act A;x:act B|] \
-\ ==> mksch A B$(x>>tr)$schA$schB = \
-\ (Takewhile (%a. a:int A)$schA) \
-\ @@ ((Takewhile (%a. a:int B)$schB) \
-\ @@ (x>>(mksch A B$tr$(TL$(Dropwhile (%a. a:int A)$schA)) \
-\ $(TL$(Dropwhile (%a. a:int B)$schB)))) \
-\ )";
-by (rtac trans 1);
-by (stac mksch_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,If_and_if]) 1);
-by (simp_tac (simpset() addsimps [Consq_def]) 1);
-qed"mksch_cons3";
-
-val compotr_simps =[mksch_UU,mksch_nil, mksch_cons1,mksch_cons2,mksch_cons3];
-
-Addsimps compotr_simps;
-
-
-(* ------------------------------------------------------------------ *)
-(* The following lemmata aim for *)
-(* COMPOSITIONALITY on TRACE Level *)
-(* ------------------------------------------------------------------ *)
-
-
-(* ---------------------------------------------------------------------------- *)
- section"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. *)
-
-Goal "(eB & ~eA --> ~A) --> \
-\ (A & (eA | eB)) = (eA & A)";
-by (Fast_tac 1);
-qed"compatibility_consequence1";
-
-
-(* very similar to above, only the commutativity of | is used to make a slight change *)
-
-Goal "(eB & ~eA --> ~A) --> \
-\ (A & (eB | eA)) = (eA & A)";
-by (Fast_tac 1);
-qed"compatibility_consequence2";
-
-
-(* ---------------------------------------------------------------------------- *)
- section"Lemmata for <==";
-(* ---------------------------------------------------------------------------- *)
-
-
-(* Lemma for substitution of looping assumption in another specific assumption *)
-val [p1,p2] = goal (the_context ()) "[| f << (g x) ; x=(h x) |] ==> f << g (h x)";
-by (cut_facts_tac [p1] 1);
-by (etac (p2 RS subst) 1);
-qed"subst_lemma1";
-
-(* Lemma for substitution of looping assumption in another specific assumption *)
-val [p1,p2] = goal (the_context ()) "[| (f x) = y >> g; x=(h x) |] ==> (f (h x)) = y >> g";
-by (cut_facts_tac [p1] 1);
-by (etac (p2 RS subst) 1);
-qed"subst_lemma2";
-
-
-Goal "!!A B. compatible A B ==> \
-\ ! schA schB. Forall (%x. x:act (A||B)) tr \
-\ --> Forall (%x. x:act (A||B)) (mksch A B$tr$schA$schB)";
-by (Seq_induct_tac "tr" [Forall_def,sforall_def,mksch_def] 1);
-by (safe_tac set_cs);
-by (asm_full_simp_tac (simpset() addsimps [actions_of_par]) 1);
-by (case_tac "a:act A" 1);
-by (case_tac "a:act B" 1);
-(* a:A, a:B *)
-by (Asm_full_simp_tac 1);
-by (rtac (Forall_Conc_impl RS mp) 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ,intA_is_not_actB,int_is_act]) 1);
-by (rtac (Forall_Conc_impl RS mp) 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ,intA_is_not_actB,int_is_act]) 1);
-(* a:A,a~:B *)
-by (Asm_full_simp_tac 1);
-by (rtac (Forall_Conc_impl RS mp) 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ,intA_is_not_actB,int_is_act]) 1);
-by (case_tac "a:act B" 1);
-(* a~:A, a:B *)
-by (Asm_full_simp_tac 1);
-by (rtac (Forall_Conc_impl RS mp) 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ,intA_is_not_actB,int_is_act]) 1);
-(* a~:A,a~:B *)
-by Auto_tac;
-qed_spec_mp"ForallAorB_mksch";
-
-Goal "!!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))";
-
-by (Seq_induct_tac "tr" [Forall_def,sforall_def,mksch_def] 1);
-by (safe_tac set_cs);
-by (rtac (Forall_Conc_impl RS mp) 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ,
- intA_is_not_actB,int_is_act]) 1);
-qed_spec_mp "ForallBnAmksch";
-
-Goal "!!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))";
-
-by (Seq_induct_tac "tr" [Forall_def,sforall_def,mksch_def] 1);
-by (safe_tac set_cs);
-by (Asm_full_simp_tac 1);
-by (rtac (Forall_Conc_impl RS mp) 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ,
- intA_is_not_actB,int_is_act]) 1);
-qed_spec_mp"ForallAnBmksch";
-
-(* safe-tac makes too many case distinctions with this lemma in the next proof *)
-Delsimps [FiniteConc];
-
-Goal "[| 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||B)) tr \
-\ --> Finite (mksch A B$tr$x$y)";
-
-by (etac Seq_Finite_ind 1);
-by (Asm_full_simp_tac 1);
-(* main case *)
-by (Asm_full_simp_tac 1);
-by (safe_tac set_cs);
-
-(* a: act A; a: act B *)
-by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1);
-by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1);
-back();
-by (REPEAT (etac conjE 1));
-(* Finite (tw iA x) and Finite (tw iB y) *)
-by (asm_full_simp_tac (simpset() addsimps
- [not_ext_is_int_or_not_act,FiniteConc]) 1);
-(* now for conclusion IH applicable, but assumptions have to be transformed *)
-by (dres_inst_tac [("x","x"),
- ("g","Filter (%a. a:act A)$s")] subst_lemma2 1);
-by (assume_tac 1);
-by (dres_inst_tac [("x","y"),
- ("g","Filter (%a. a:act B)$s")] subst_lemma2 1);
-by (assume_tac 1);
-(* IH *)
-by (asm_full_simp_tac (simpset()
- addsimps [not_ext_is_int_or_not_act,ForallTL,ForallDropwhile]) 1);
-
-(* a: act B; a~: act A *)
-by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1);
-
-by (REPEAT (etac conjE 1));
-(* Finite (tw iB y) *)
-by (asm_full_simp_tac (simpset() addsimps
- [not_ext_is_int_or_not_act,FiniteConc]) 1);
-(* now for conclusion IH applicable, but assumptions have to be transformed *)
-by (dres_inst_tac [("x","y"),
- ("g","Filter (%a. a:act B)$s")] subst_lemma2 1);
-by (assume_tac 1);
-(* IH *)
-by (asm_full_simp_tac (simpset()
- addsimps [not_ext_is_int_or_not_act,ForallTL,ForallDropwhile]) 1);
-
-(* a~: act B; a: act A *)
-by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1);
-
-by (REPEAT (etac conjE 1));
-(* Finite (tw iA x) *)
-by (asm_full_simp_tac (simpset() addsimps
- [not_ext_is_int_or_not_act,FiniteConc]) 1);
-(* now for conclusion IH applicable, but assumptions have to be transformed *)
-by (dres_inst_tac [("x","x"),
- ("g","Filter (%a. a:act A)$s")] subst_lemma2 1);
-by (assume_tac 1);
-(* IH *)
-by (asm_full_simp_tac (simpset()
- addsimps [not_ext_is_int_or_not_act,ForallTL,ForallDropwhile]) 1);
-
-(* a~: act B; a~: act A *)
-by (fast_tac (claset() addSIs [ext_is_act]
- addss (simpset() addsimps [externals_of_par]) ) 1);
-qed_spec_mp"FiniteL_mksch";
-
-Addsimps [FiniteConc];
-
-
-(* otherwise subst_lemma2 does not fit exactly, just to avoid a subst_lemma3 *)
-Delsimps [FilterConc];
-
-Goal " [| 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)";
-by (forw_inst_tac [("A1","A")] (compat_commute RS iffD1) 1);
-by (etac Seq_Finite_ind 1);
-by (REPEAT (rtac allI 1));
-by (rtac impI 1);
-by (res_inst_tac [("x","nil")] exI 1);
-by (res_inst_tac [("x","y")] exI 1);
-by (Asm_full_simp_tac 1);
-(* main case *)
-by (REPEAT (rtac allI 1));
-by (rtac impI 1);
-by (Asm_full_simp_tac 1);
-by (REPEAT (etac conjE 1));
-by (Asm_full_simp_tac 1);
-(* divide_Seq on s *)
-by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1);
-by (REPEAT (etac conjE 1));
-(* transform assumption f eB y = f B (s@z) *)
-by (dres_inst_tac [("x","y"),
- ("g","Filter (%a. a:act B)$(s@@z)")] subst_lemma2 1);
-by (assume_tac 1);
-Addsimps [FilterConc];
-by (asm_full_simp_tac (simpset() addsimps [not_ext_is_int_or_not_act]) 1);
-(* apply IH *)
-by (eres_inst_tac [("x","TL$(Dropwhile (%a. a:int B)$y)")] allE 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallTL,ForallDropwhile])1);
-by (REPEAT (etac exE 1));
-by (REPEAT (etac conjE 1));
-by (Asm_full_simp_tac 1);
-(* for replacing IH in conclusion *)
-by (rotate_tac ~2 1);
-by (Asm_full_simp_tac 1);
-(* instantiate y1a and y2a *)
-by (res_inst_tac [("x","Takewhile (%a. a:int B)$y @@ a>>y1")] exI 1);
-by (res_inst_tac [("x","y2")] exI 1);
-(* elminate all obligations up to two depending on Conc_assoc *)
-by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ, intA_is_not_actB,
- int_is_act,int_is_not_ext]) 1);
-by (simp_tac (simpset() addsimps [Conc_assoc]) 1);
-qed_spec_mp "reduceA_mksch1";
-
-bind_thm("reduceA_mksch",conjI RSN (6,conjI RSN (5,reduceA_mksch1)));
-
-
-(* otherwise subst_lemma2 does not fit exactly, just to avoid a subst_lemma3 *)
-Delsimps [FilterConc];
-
-
-Goal " [| 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)";
-by (forw_inst_tac [("A1","A")] (compat_commute RS iffD1) 1);
-by (etac Seq_Finite_ind 1);
-by (REPEAT (rtac allI 1));
-by (rtac impI 1);
-by (res_inst_tac [("x","nil")] exI 1);
-by (res_inst_tac [("x","x")] exI 1);
-by (Asm_full_simp_tac 1);
-(* main case *)
-by (REPEAT (rtac allI 1));
-by (rtac impI 1);
-by (Asm_full_simp_tac 1);
-by (REPEAT (etac conjE 1));
-by (Asm_full_simp_tac 1);
-(* divide_Seq on s *)
-by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1);
-by (REPEAT (etac conjE 1));
-(* transform assumption f eA x = f A (s@z) *)
-by (dres_inst_tac [("x","x"),
- ("g","Filter (%a. a:act A)$(s@@z)")] subst_lemma2 1);
-by (assume_tac 1);
-Addsimps [FilterConc];
-by (asm_full_simp_tac (simpset() addsimps [not_ext_is_int_or_not_act]) 1);
-(* apply IH *)
-by (eres_inst_tac [("x","TL$(Dropwhile (%a. a:int A)$x)")] allE 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallTL,ForallDropwhile])1);
-by (REPEAT (etac exE 1));
-by (REPEAT (etac conjE 1));
-by (Asm_full_simp_tac 1);
-(* for replacing IH in conclusion *)
-by (rotate_tac ~2 1);
-by (Asm_full_simp_tac 1);
-(* instantiate y1a and y2a *)
-by (res_inst_tac [("x","Takewhile (%a. a:int A)$x @@ a>>x1")] exI 1);
-by (res_inst_tac [("x","x2")] exI 1);
-(* elminate all obligations up to two depending on Conc_assoc *)
-by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ, intA_is_not_actB,
- int_is_act,int_is_not_ext]) 1);
-by (simp_tac (simpset() addsimps [Conc_assoc]) 1);
-qed_spec_mp"reduceB_mksch1";
-
-bind_thm("reduceB_mksch",conjI RSN (6,conjI RSN (5,reduceB_mksch1)));
-
-
-(*
-
-
-Goal "Finite y ==> ! z tr. Forall (%a.a:ext (A||B)) tr & \
-\ y = z @@ mksch A B$tr$a$b\
-\ --> Finite tr";
-by (etac Seq_Finite_ind 1);
-by Auto_tac;
-by (Seq_case_simp_tac "tr" 1);
-(* tr = UU *)
-by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc]) 1);
-(* tr = nil *)
-by Auto_tac;
-(* tr = Conc *)
-ren "s ss" 1;
-
-by (case_tac "s:act A" 1);
-by (case_tac "s:act B" 1);
-by (rotate_tac ~2 1);
-by (rotate_tac ~2 2);
-by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc,nil_is_Conc2]) 1);
-by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc,nil_is_Conc2]) 1);
-by (case_tac "s:act B" 1);
-by (rotate_tac ~2 1);
-by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc,nil_is_Conc2]) 1);
-by (fast_tac (claset() addSIs [ext_is_act]
- addss (simpset() addsimps [externals_of_par]) ) 1);
-(* main case *)
-by (Seq_case_simp_tac "tr" 1);
-(* tr = UU *)
-by (asm_full_simp_tac (simpset() addsimps [Conc_Conc_eq]) 1);
-by Auto_tac;
-(* tr = nil ok *)
-(* tr = Conc *)
-by (Seq_case_simp_tac "z" 1);
-(* z = UU ok *)
-(* z = nil *)
-
-(* z= Cons *)
-
-
-by (case_tac "aaa:act A" 2);
-by (case_tac "aaa:act B" 2);
-by (rotate_tac ~2 2);
-by (rotate_tac ~2 3);
-by (asm_full_simp_tac (HOL_basic_ss addsimps [mksch_cons3]) 2);
-by (eres_inst_tac [("x","sb@@Takewhile (%a. a: int A)$a @@ Takewhile (%a. a:int B)$b@@(aaa>>nil)")] allE 2);
-by (eres_inst_tac [("x","sa")] allE 2);
-by (asm_full_simp_tac (simpset() addsimps [Conc_assoc])2);
-
-
-
-by (eres_inst_tac [("x","sa")] allE 1);
-by (Asm_full_simp_tac 1);
-by (case_tac "aaa:act A" 1);
-by (case_tac "aaa:act B" 1);
-by (rotate_tac ~2 1);
-by (rotate_tac ~2 2);
-by (asm_full_simp_tac (simpset() addsimps [Conc_Conc_eq]) 1);
-
-
-
-Goal "(x>>xs = y @@ z) = ((y=nil & x>>xs=z) | (? ys. y=x>>ys & xs=ys@@z))";
-by (Seq_case_simp_tac "y" 1);
-by Auto_tac;
-qed"Conc_Conc_eq";
-
-Goal "!! (y::'a Seq).Finite y ==> ~ y= x@@UU";
-by (etac Seq_Finite_ind 1);
-by (Seq_case_simp_tac "x" 1);
-by (Seq_case_simp_tac "x" 1);
-by Auto_tac;
-qed"FiniteConcUU1";
-
-Goal "~ Finite ((x::'a Seq)@@UU)";
-by (auto_tac (claset() addDs [FiniteConcUU1], simpset()));
-qed"FiniteConcUU";
-
-finiteR_mksch
- "Finite (mksch A B$tr$x$y) --> Finite tr"
-*)
-
-
-
-(*------------------------------------------------------------------------------------- *)
- section"Filtering external actions out of mksch(tr,schA,schB) yields the oracle tr";
-(* structural induction
- ------------------------------------------------------------------------------------- *)
-
-Goal
-"!! 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||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||B))$(mksch A B$tr$schA$schB) = tr";
-
-by (Seq_induct_tac "tr" [Forall_def,sforall_def,mksch_def] 1);
-
-(* main case *)
-(* splitting into 4 cases according to a:A, a:B *)
-by (Asm_full_simp_tac 1);
-by (safe_tac set_cs);
-
-(* Case a:A, a:B *)
-by (ftac divide_Seq 1);
-by (ftac divide_Seq 1);
-back();
-by (REPEAT (etac conjE 1));
-(* filtering internals of A in schA and of B in schB is nil *)
-by (asm_full_simp_tac (simpset() addsimps
- [not_ext_is_int_or_not_act,
- externals_of_par, intA_is_not_extB,int_is_not_ext]) 1);
-(* conclusion of IH ok, but assumptions of IH have to be transformed *)
-by (dres_inst_tac [("x","schA")] subst_lemma1 1);
-by (assume_tac 1);
-by (dres_inst_tac [("x","schB")] subst_lemma1 1);
-by (assume_tac 1);
-(* IH *)
-by (asm_full_simp_tac (simpset() addsimps [not_ext_is_int_or_not_act,
- ForallTL,ForallDropwhile]) 1);
-
-(* Case a:A, a~:B *)
-by (ftac divide_Seq 1);
-by (REPEAT (etac conjE 1));
-(* filtering internals of A is nil *)
-by (asm_full_simp_tac (simpset() addsimps
- [not_ext_is_int_or_not_act,
- externals_of_par, intA_is_not_extB,int_is_not_ext]) 1);
-by (dres_inst_tac [("x","schA")] subst_lemma1 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [not_ext_is_int_or_not_act,
- ForallTL,ForallDropwhile]) 1);
-
-(* Case a:B, a~:A *)
-by (ftac divide_Seq 1);
-by (REPEAT (etac conjE 1));
-(* filtering internals of A is nil *)
-by (asm_full_simp_tac (simpset() addsimps
- [not_ext_is_int_or_not_act,
- externals_of_par, intA_is_not_extB,int_is_not_ext]) 1);
-by (dres_inst_tac [("x","schB")] subst_lemma1 1);
-back();
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [not_ext_is_int_or_not_act,
- ForallTL,ForallDropwhile]) 1);
-
-(* Case a~:A, a~:B *)
-by (fast_tac (claset() addSIs [ext_is_act]
- addss (simpset() addsimps [externals_of_par]) ) 1);
-qed"FilterA_mksch_is_tr";
-
-
-
-(*
-
-***************************************************************8
-With uncomplete take lemma rule should be reused afterwards !!!!!!!!!!!!!!!!!
-**********************************************************************
-
-(*---------------------------------------------------------------------------
- Filter of mksch(tr,schA,schB) to A is schA
- take lemma
- --------------------------------------------------------------------------- *)
-
-Goal "!! A B. [| compatible A B; compatible B A; \
-\ is_asig(asig_of A); is_asig(asig_of B) |] ==> \
-\ Forall (%x. x:ext (A||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";
-
-by (res_inst_tac [("Q","%x. x:act B & x~:act A"),("x","tr")] take_lemma_less_induct 1);
-by (REPEAT (etac conjE 1));
-
-by (case_tac "Finite s" 1);
-
-(* both sides of this equation are nil *)
-by (subgoal_tac "schA=nil" 1);
-by (Asm_simp_tac 1);
-(* first side: mksch = nil *)
-by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPnil,ForallBnAmksch,FiniteL_mksch],
- simpset())) 1);
-(* second side: schA = nil *)
-by (eres_inst_tac [("A","A")] LastActExtimplnil 1);
-by (Asm_simp_tac 1);
-by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPnil],
- simpset())) 1);
-
-(* case ~ Finite s *)
-
-(* both sides of this equation are UU *)
-by (subgoal_tac "schA=UU" 1);
-by (Asm_simp_tac 1);
-(* first side: mksch = UU *)
-by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPUU,
- (finiteR_mksch RS mp COMP rev_contrapos),
- ForallBnAmksch],
- simpset())) 1);
-(* schA = UU *)
-by (eres_inst_tac [("A","A")] LastActExtimplUU 1);
-by (Asm_simp_tac 1);
-by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPUU],
- simpset())) 1);
-
-(* case" ~ Forall (%x.x:act B & x~:act A) s" *)
-
-by (REPEAT (etac conjE 1));
-
-(* bring in lemma reduceA_mksch *)
-by (forw_inst_tac [("y","schB"),("x","schA")] reduceA_mksch 1);
-by (REPEAT (atac 1));
-by (REPEAT (etac exE 1));
-by (REPEAT (etac conjE 1));
-
-(* use reduceA_mksch to rewrite conclusion *)
-by (hyp_subst_tac 1);
-by (Asm_full_simp_tac 1);
-
-(* eliminate the B-only prefix *)
-
-by (subgoal_tac "(Filter (%a. a :act A)$y1) = nil" 1);
-by (etac ForallQFilterPnil 2);
-by (assume_tac 2);
-by (Fast_tac 2);
-
-(* Now real recursive step follows (in y) *)
-
-by (Asm_full_simp_tac 1);
-by (case_tac "y:act A" 1);
-by (case_tac "y~:act B" 1);
-by (rotate_tac ~2 1);
-by (Asm_full_simp_tac 1);
-
-by (subgoal_tac "Filter (%a. a:act A & a:ext B)$y1=nil" 1);
-by (rotate_tac ~1 1);
-by (Asm_full_simp_tac 1);
-(* eliminate introduced subgoal 2 *)
-by (etac ForallQFilterPnil 2);
-by (assume_tac 2);
-by (Fast_tac 2);
-
-(* bring in divide Seq for s *)
-by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1);
-by (REPEAT (etac conjE 1));
-
-(* subst divide_Seq in conclusion, but only at the righest occurence *)
-by (res_inst_tac [("t","schA")] ssubst 1);
-back();
-back();
-back();
-by (assume_tac 1);
-
-(* reduce trace_takes from n to strictly smaller k *)
-by (rtac take_reduction 1);
-
-(* f A (tw iA) = tw ~eA *)
-by (asm_full_simp_tac (simpset() addsimps [FilterPTakewhileQid,int_is_act,
- not_ext_is_int_or_not_act]) 1);
-by (rtac refl 1);
-
-(* now conclusion fulfills induction hypothesis, but assumptions are not ready *)
-(*
-
-nacessary anymore ????????????????
-by (rotate_tac ~10 1);
-
-*)
-(* assumption schB *)
-by (asm_full_simp_tac (simpset() addsimps [ext_and_act]) 1);
-(* assumption schA *)
-by (dres_inst_tac [("x","schA"),
- ("g","Filter (%a. a:act A)$s2")] subst_lemma2 1);
-by (assume_tac 1);
-by (Asm_full_simp_tac 1);
-(* assumptions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping *)
-by (dres_inst_tac [("sch","schA"),("P","%a. a:int A")] LastActExtsmall1 1);
-by (dres_inst_tac [("sch1.0","y1")] LastActExtsmall2 1);
-by (assume_tac 1);
-
-FIX: problem: schA and schB are not quantified in the new take lemma version !!!
-
-by (Asm_full_simp_tac 1);
-
-**********************************************************************************************
-*)
-
-
-
-(*--------------------------------------------------------------------------- *)
-
- section" Filter of mksch(tr,schA,schB) to A is schA -- take lemma proof";
-
-(* --------------------------------------------------------------------------- *)
-
-
-
-Goal "!! A B. [| compatible A B; compatible B A; \
-\ is_asig(asig_of A); is_asig(asig_of B) |] ==> \
-\ Forall (%x. x:ext (A||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";
-
-by (strip_tac 1);
-by (resolve_tac [seq.take_lemmas] 1);
-by (rtac mp 1);
-by (assume_tac 2);
-back();back();back();back();
-by (res_inst_tac [("x","schA")] spec 1);
-by (res_inst_tac [("x","schB")] spec 1);
-by (res_inst_tac [("x","tr")] spec 1);
-by (thin_tac' 5 1);
-by (rtac nat_less_induct 1);
-by (REPEAT (rtac allI 1));
-by (rename_tac "tr schB schA" 1);
-by (strip_tac 1);
-by (REPEAT (etac conjE 1));
-
-by (case_tac "Forall (%x. x:act B & x~:act A) tr" 1);
-
-by (rtac (seq_take_lemma RS iffD2 RS spec) 1);
-by (thin_tac' 5 1);
-
-
-by (case_tac "Finite tr" 1);
-
-(* both sides of this equation are nil *)
-by (subgoal_tac "schA=nil" 1);
-by (Asm_simp_tac 1);
-(* first side: mksch = nil *)
-by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPnil,ForallBnAmksch,FiniteL_mksch],
- simpset())) 1);
-(* second side: schA = nil *)
-by (eres_inst_tac [("A","A")] LastActExtimplnil 1);
-by (Asm_simp_tac 1);
-by (eres_inst_tac [("Q","%x. x:act B & x~:act A")] ForallQFilterPnil 1);
-by (assume_tac 1);
-by (Fast_tac 1);
-
-(* case ~ Finite s *)
-
-(* both sides of this equation are UU *)
-by (subgoal_tac "schA=UU" 1);
-by (Asm_simp_tac 1);
-(* first side: mksch = UU *)
-by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPUU,
- (finiteR_mksch RS mp COMP rev_contrapos),ForallBnAmksch],simpset())) 1);
-(* schA = UU *)
-by (eres_inst_tac [("A","A")] LastActExtimplUU 1);
-by (Asm_simp_tac 1);
-by (eres_inst_tac [("Q","%x. x:act B & x~:act A")] ForallQFilterPUU 1);
-by (assume_tac 1);
-by (Fast_tac 1);
-
-(* case" ~ Forall (%x.x:act B & x~:act A) s" *)
-
-by (dtac divide_Seq3 1);
-
-by (REPEAT (etac exE 1));
-by (REPEAT (etac conjE 1));
-by (hyp_subst_tac 1);
-
-(* bring in lemma reduceA_mksch *)
-by (forw_inst_tac [("x","schA"),("y","schB"),("A","A"),("B","B")] reduceA_mksch 1);
-by (REPEAT (atac 1));
-by (REPEAT (etac exE 1));
-by (REPEAT (etac conjE 1));
-
-(* use reduceA_mksch to rewrite conclusion *)
-by (hyp_subst_tac 1);
-by (Asm_full_simp_tac 1);
-
-(* eliminate the B-only prefix *)
-
-by (subgoal_tac "(Filter (%a. a :act A)$y1) = nil" 1);
-by (etac ForallQFilterPnil 2);
-by (assume_tac 2);
-by (Fast_tac 2);
-
-(* Now real recursive step follows (in y) *)
-
-by (Asm_full_simp_tac 1);
-by (case_tac "x:act A" 1);
-by (case_tac "x~:act B" 1);
-by (rotate_tac ~2 1);
-by (Asm_full_simp_tac 1);
-
-by (subgoal_tac "Filter (%a. a:act A & a:ext B)$y1=nil" 1);
-by (rotate_tac ~1 1);
-by (Asm_full_simp_tac 1);
-(* eliminate introduced subgoal 2 *)
-by (etac ForallQFilterPnil 2);
-by (assume_tac 2);
-by (Fast_tac 2);
-
-(* bring in divide Seq for s *)
-by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1);
-by (REPEAT (etac conjE 1));
-
-(* subst divide_Seq in conclusion, but only at the righest occurence *)
-by (res_inst_tac [("t","schA")] ssubst 1);
-back();
-back();
-back();
-by (assume_tac 1);
-
-(* reduce trace_takes from n to strictly smaller k *)
-by (rtac take_reduction 1);
-
-(* f A (tw iA) = tw ~eA *)
-by (asm_full_simp_tac (simpset() addsimps [FilterPTakewhileQid,int_is_act,
- not_ext_is_int_or_not_act]) 1);
-by (rtac refl 1);
-by (asm_full_simp_tac (simpset() addsimps [int_is_act,
- not_ext_is_int_or_not_act]) 1);
-by (rotate_tac ~11 1);
-
-(* now conclusion fulfills induction hypothesis, but assumptions are not ready *)
-
-(* assumption Forall tr *)
-by (asm_full_simp_tac (simpset() addsimps [Forall_Conc]) 1);
-(* assumption schB *)
-by (asm_full_simp_tac (simpset() addsimps [Forall_Conc,ext_and_act]) 1);
-by (REPEAT (etac conjE 1));
-(* assumption schA *)
-by (dres_inst_tac [("x","schA"),
- ("g","Filter (%a. a:act A)$rs")] subst_lemma2 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext]) 1);
-(* assumptions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping *)
-by (dres_inst_tac [("sch","schA"),("P","%a. a:int A")] LastActExtsmall1 1);
-by (forw_inst_tac [("sch1.0","y1")] LastActExtsmall2 1);
-by (assume_tac 1);
-
-(* assumption Forall schA *)
-by (dres_inst_tac [("s","schA"),
- ("P","Forall (%x. x:act A)")] subst 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ, int_is_act]) 1);
-
-(* case x:actions(asig_of A) & x: actions(asig_of B) *)
-
-
-by (rotate_tac ~2 1);
-by (Asm_full_simp_tac 1);
-
-by (subgoal_tac "Filter (%a. a:act A & a:ext B)$y1=nil" 1);
-by (rotate_tac ~1 1);
-by (Asm_full_simp_tac 1);
-(* eliminate introduced subgoal 2 *)
-by (etac ForallQFilterPnil 2);
-by (assume_tac 2);
-by (Fast_tac 2);
-
-(* bring in divide Seq for s *)
-by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1);
-by (REPEAT (etac conjE 1));
-
-(* subst divide_Seq in conclusion, but only at the righest occurence *)
-by (res_inst_tac [("t","schA")] ssubst 1);
-back();
-back();
-back();
-by (assume_tac 1);
-
-(* f A (tw iA) = tw ~eA *)
-by (asm_full_simp_tac (simpset() addsimps [FilterPTakewhileQid,int_is_act,
- not_ext_is_int_or_not_act]) 1);
-
-(* rewrite assumption forall and schB *)
-by (rotate_tac 13 1);
-by (asm_full_simp_tac (simpset() addsimps [ext_and_act]) 1);
-
-(* divide_Seq for schB2 *)
-by (forw_inst_tac [("y","y2")] (sym RS antisym_less_inverse RS conjunct1 RS divide_Seq) 1);
-by (REPEAT (etac conjE 1));
-(* assumption schA *)
-by (dres_inst_tac [("x","schA"),
- ("g","Filter (%a. a:act A)$rs")] subst_lemma2 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext]) 1);
-
-(* f A (tw iB schB2) = nil *)
-by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext,not_ext_is_int_or_not_act,
- intA_is_not_actB]) 1);
-
-
-(* reduce trace_takes from n to strictly smaller k *)
-by (rtac take_reduction 1);
-by (rtac refl 1);
-by (rtac refl 1);
-
-(* now conclusion fulfills induction hypothesis, but assumptions are not all ready *)
-
-(* assumption schB *)
-by (dres_inst_tac [("x","y2"),
- ("g","Filter (%a. a:act B)$rs")] subst_lemma2 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [intA_is_not_actB,int_is_not_ext]) 1);
-
-(* conclusions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping *)
-by (dres_inst_tac [("sch","schA"),("P","%a. a:int A")] LastActExtsmall1 1);
-by (forw_inst_tac [("sch1.0","y1")] LastActExtsmall2 1);
-by (assume_tac 1);
-by (dres_inst_tac [("sch","y2"),("P","%a. a:int B")] LastActExtsmall1 1);
-
-(* assumption Forall schA, and Forall schA are performed by ForallTL,ForallDropwhile *)
-by (asm_full_simp_tac (simpset() addsimps [ForallTL,ForallDropwhile]) 1);
-
-(* case x~:A & x:B *)
-(* cannot occur, as just this case has been scheduled out before as the B-only prefix *)
-by (case_tac "x:act B" 1);
-by (Fast_tac 1);
-
-(* case x~:A & x~:B *)
-(* cannot occur because of assumption: Forall (a:ext A | a:ext B) *)
-by (rotate_tac ~9 1);
-(* reduce forall assumption from tr to (x>>rs) *)
-by (asm_full_simp_tac (simpset() addsimps [externals_of_par]) 1);
-by (REPEAT (etac conjE 1));
-by (fast_tac (claset() addSIs [ext_is_act]) 1);
-
-qed"FilterAmksch_is_schA";
-
-
-
-(*--------------------------------------------------------------------------- *)
-
- section" Filter of mksch(tr,schA,schB) to B is schB -- take lemma proof";
-
-(* --------------------------------------------------------------------------- *)
-
-
-
-Goal "!! A B. [| compatible A B; compatible B A; \
-\ is_asig(asig_of A); is_asig(asig_of B) |] ==> \
-\ Forall (%x. x:ext (A||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";
-
-by (strip_tac 1);
-by (resolve_tac [seq.take_lemmas] 1);
-by (rtac mp 1);
-by (assume_tac 2);
-back();back();back();back();
-by (res_inst_tac [("x","schA")] spec 1);
-by (res_inst_tac [("x","schB")] spec 1);
-by (res_inst_tac [("x","tr")] spec 1);
-by (thin_tac' 5 1);
-by (rtac nat_less_induct 1);
-by (REPEAT (rtac allI 1));
-by (rename_tac "tr schB schA" 1);
-by (strip_tac 1);
-by (REPEAT (etac conjE 1));
-
-by (case_tac "Forall (%x. x:act A & x~:act B) tr" 1);
-
-by (rtac (seq_take_lemma RS iffD2 RS spec) 1);
-by (thin_tac' 5 1);
-
-
-by (case_tac "Finite tr" 1);
-
-(* both sides of this equation are nil *)
-by (subgoal_tac "schB=nil" 1);
-by (Asm_simp_tac 1);
-(* first side: mksch = nil *)
-by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPnil,ForallAnBmksch,FiniteL_mksch],
- simpset())) 1);
-(* second side: schA = nil *)
-by (eres_inst_tac [("A","B")] LastActExtimplnil 1);
-by (Asm_simp_tac 1);
-by (eres_inst_tac [("Q","%x. x:act A & x~:act B")] ForallQFilterPnil 1);
-by (assume_tac 1);
-by (Fast_tac 1);
-
-(* case ~ Finite tr *)
-
-(* both sides of this equation are UU *)
-by (subgoal_tac "schB=UU" 1);
-by (Asm_simp_tac 1);
-(* first side: mksch = UU *)
-by (force_tac (claset() addSIs [ForallQFilterPUU,
- (finiteR_mksch RS mp COMP rev_contrapos),
- ForallAnBmksch],
- simpset()) 1);
-(* schA = UU *)
-by (eres_inst_tac [("A","B")] LastActExtimplUU 1);
-by (Asm_simp_tac 1);
-by (eres_inst_tac [("Q","%x. x:act A & x~:act B")] ForallQFilterPUU 1);
-by (assume_tac 1);
-by (Fast_tac 1);
-
-(* case" ~ Forall (%x.x:act B & x~:act A) s" *)
-
-by (dtac divide_Seq3 1);
-
-by (REPEAT (etac exE 1));
-by (REPEAT (etac conjE 1));
-by (hyp_subst_tac 1);
-
-(* bring in lemma reduceB_mksch *)
-by (forw_inst_tac [("y","schB"),("x","schA"),("A","A"),("B","B")] reduceB_mksch 1);
-by (REPEAT (atac 1));
-by (REPEAT (etac exE 1));
-by (REPEAT (etac conjE 1));
-
-(* use reduceB_mksch to rewrite conclusion *)
-by (hyp_subst_tac 1);
-by (Asm_full_simp_tac 1);
-
-(* eliminate the A-only prefix *)
-
-by (subgoal_tac "(Filter (%a. a :act B)$x1) = nil" 1);
-by (etac ForallQFilterPnil 2);
-by (assume_tac 2);
-by (Fast_tac 2);
-
-(* Now real recursive step follows (in x) *)
-
-by (Asm_full_simp_tac 1);
-by (case_tac "x:act B" 1);
-by (case_tac "x~:act A" 1);
-by (rotate_tac ~2 1);
-by (Asm_full_simp_tac 1);
-
-by (subgoal_tac "Filter (%a. a:act B & a:ext A)$x1=nil" 1);
-by (rotate_tac ~1 1);
-by (Asm_full_simp_tac 1);
-(* eliminate introduced subgoal 2 *)
-by (etac ForallQFilterPnil 2);
-by (assume_tac 2);
-by (Fast_tac 2);
-
-(* bring in divide Seq for s *)
-by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1);
-by (REPEAT (etac conjE 1));
-
-(* subst divide_Seq in conclusion, but only at the righest occurence *)
-by (res_inst_tac [("t","schB")] ssubst 1);
-back();
-back();
-back();
-by (assume_tac 1);
-
-(* reduce trace_takes from n to strictly smaller k *)
-by (rtac take_reduction 1);
-
-(* f B (tw iB) = tw ~eB *)
-by (asm_full_simp_tac (simpset() addsimps [FilterPTakewhileQid,int_is_act,
- not_ext_is_int_or_not_act]) 1);
-by (rtac refl 1);
-by (asm_full_simp_tac (simpset() addsimps [int_is_act,
- not_ext_is_int_or_not_act]) 1);
-by (rotate_tac ~11 1);
-
-(* now conclusion fulfills induction hypothesis, but assumptions are not ready *)
-
-(* assumption Forall tr *)
-by (asm_full_simp_tac (simpset() addsimps [Forall_Conc]) 1);
-(* assumption schA *)
-by (asm_full_simp_tac (simpset() addsimps [Forall_Conc,ext_and_act]) 1);
-by (REPEAT (etac conjE 1));
-(* assumption schB *)
-by (dres_inst_tac [("x","schB"),
- ("g","Filter (%a. a:act B)$rs")] subst_lemma2 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext]) 1);
-(* assumptions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping *)
-by (dres_inst_tac [("sch","schB"),("P","%a. a:int B")] LastActExtsmall1 1);
-by (forw_inst_tac [("sch1.0","x1")] LastActExtsmall2 1);
-by (assume_tac 1);
-
-(* assumption Forall schB *)
-by (dres_inst_tac [("s","schB"),
- ("P","Forall (%x. x:act B)")] subst 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ, int_is_act]) 1);
-
-(* case x:actions(asig_of A) & x: actions(asig_of B) *)
-
-
-by (rotate_tac ~2 1);
-by (Asm_full_simp_tac 1);
-
-by (subgoal_tac "Filter (%a. a:act B & a:ext A)$x1=nil" 1);
-by (rotate_tac ~1 1);
-by (Asm_full_simp_tac 1);
-(* eliminate introduced subgoal 2 *)
-by (etac ForallQFilterPnil 2);
-by (assume_tac 2);
-by (Fast_tac 2);
-
-(* bring in divide Seq for s *)
-by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1);
-by (REPEAT (etac conjE 1));
-
-(* subst divide_Seq in conclusion, but only at the righest occurence *)
-by (res_inst_tac [("t","schB")] ssubst 1);
-back();
-back();
-back();
-by (assume_tac 1);
-
-(* f B (tw iB) = tw ~eB *)
-by (asm_full_simp_tac (simpset() addsimps [FilterPTakewhileQid,int_is_act,
- not_ext_is_int_or_not_act]) 1);
-
-(* rewrite assumption forall and schB *)
-by (rotate_tac 13 1);
-by (asm_full_simp_tac (simpset() addsimps [ext_and_act]) 1);
-
-(* divide_Seq for schB2 *)
-by (forw_inst_tac [("y","x2")] (sym RS antisym_less_inverse RS conjunct1 RS divide_Seq) 1);
-by (REPEAT (etac conjE 1));
-(* assumption schA *)
-by (dres_inst_tac [("x","schB"),
- ("g","Filter (%a. a:act B)$rs")] subst_lemma2 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext]) 1);
-
-(* f B (tw iA schA2) = nil *)
-by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext,not_ext_is_int_or_not_act,
- intA_is_not_actB]) 1);
-
-
-(* reduce trace_takes from n to strictly smaller k *)
-by (rtac take_reduction 1);
-by (rtac refl 1);
-by (rtac refl 1);
-
-(* now conclusion fulfills induction hypothesis, but assumptions are not all ready *)
-
-(* assumption schA *)
-by (dres_inst_tac [("x","x2"),
- ("g","Filter (%a. a:act A)$rs")] subst_lemma2 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [intA_is_not_actB,int_is_not_ext]) 1);
-
-(* conclusions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping *)
-by (dres_inst_tac [("sch","schB"),("P","%a. a:int B")] LastActExtsmall1 1);
-by (forw_inst_tac [("sch1.0","x1")] LastActExtsmall2 1);
-by (assume_tac 1);
-by (dres_inst_tac [("sch","x2"),("P","%a. a:int A")] LastActExtsmall1 1);
-
-(* assumption Forall schA, and Forall schA are performed by ForallTL,ForallDropwhile *)
-by (asm_full_simp_tac (simpset() addsimps [ForallTL,ForallDropwhile]) 1);
-
-(* case x~:B & x:A *)
-(* cannot occur, as just this case has been scheduled out before as the B-only prefix *)
-by (case_tac "x:act A" 1);
-by (Fast_tac 1);
-
-(* case x~:B & x~:A *)
-(* cannot occur because of assumption: Forall (a:ext A | a:ext B) *)
-by (rotate_tac ~9 1);
-(* reduce forall assumption from tr to (x>>rs) *)
-by (asm_full_simp_tac (simpset() addsimps [externals_of_par]) 1);
-by (REPEAT (etac conjE 1));
-by (fast_tac (claset() addSIs [ext_is_act]) 1);
-
-qed"FilterBmksch_is_schB";
-
-
-
-(* ------------------------------------------------------------------ *)
- section"COMPOSITIONALITY on TRACE Level -- Main Theorem";
-(* ------------------------------------------------------------------ *)
-
-Goal
-"!! 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||B)) = \
-\ (Filter (%a. a:act A)$tr : traces A &\
-\ Filter (%a. a:act B)$tr : traces B &\
-\ Forall (%x. x:ext(A||B)) tr)";
-
-by (simp_tac (simpset() addsimps [traces_def,has_trace_def]) 1);
-by (safe_tac set_cs);
-
-(* ==> *)
-(* There is a schedule of A *)
-by (res_inst_tac [("x","Filter (%a. a:act A)$sch")] bexI 1);
-by (asm_full_simp_tac (simpset() addsimps [compositionality_sch]) 2);
-by (asm_full_simp_tac (simpset() addsimps [compatibility_consequence1,
- externals_of_par,ext1_ext2_is_not_act1]) 1);
-(* There is a schedule of B *)
-by (res_inst_tac [("x","Filter (%a. a:act B)$sch")] bexI 1);
-by (asm_full_simp_tac (simpset() addsimps [compositionality_sch]) 2);
-by (asm_full_simp_tac (simpset() addsimps [compatibility_consequence2,
- externals_of_par,ext1_ext2_is_not_act2]) 1);
-(* Traces of A||B have only external actions from A or B *)
-by (rtac ForallPFilterP 1);
-
-(* <== *)
-
-(* replace schA and schB by Cut(schA) and Cut(schB) *)
-by (dtac exists_LastActExtsch 1);
-by (assume_tac 1);
-by (dtac exists_LastActExtsch 1);
-by (assume_tac 1);
-by (REPEAT (etac exE 1));
-by (REPEAT (etac conjE 1));
-(* Schedules of A(B) have only actions of A(B) *)
-by (dtac scheds_in_sig 1);
-by (assume_tac 1);
-by (dtac scheds_in_sig 1);
-by (assume_tac 1);
-
-by (rename_tac "h1 h2 schA schB" 1);
-(* mksch is exactly the construction of trA||B out of schA, schB, and the oracle tr,
- we need here *)
-by (res_inst_tac [("x","mksch A B$tr$schA$schB")] bexI 1);
-
-(* External actions of mksch are just the oracle *)
-by (asm_full_simp_tac (simpset() addsimps [FilterA_mksch_is_tr RS spec RS spec RS mp]) 1);
-
-(* mksch is a schedule -- use compositionality on sch-level *)
-by (asm_full_simp_tac (simpset() addsimps [compositionality_sch]) 1);
-by (asm_full_simp_tac (simpset() addsimps [FilterAmksch_is_schA,FilterBmksch_is_schB]) 1);
-by (etac ForallAorB_mksch 1);
-by (etac ForallPForallQ 1);
-by (etac ext_is_act 1);
-qed"compositionality_tr";
-
-
-
-(* ------------------------------------------------------------------ *)
-(* COMPOSITIONALITY on TRACE Level *)
-(* For Modules *)
-(* ------------------------------------------------------------------ *)
-
-Goalw [Traces_def,par_traces_def]
-
-"!! 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||B) = par_traces (Traces A) (Traces B)";
-
-by (asm_full_simp_tac (simpset() addsimps [asig_of_par]) 1);
-by (rtac set_ext 1);
-by (asm_full_simp_tac (simpset() addsimps [compositionality_tr,externals_of_par]) 1);
-qed"compositionality_tr_modules";
-
-
-change_simpset (fn ss => ss setmksym (SOME o symmetric_fun));
--- a/src/HOLCF/IOA/meta_theory/CompoTraces.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/CompoTraces.thy Sun May 28 19:54:20 2006 +0200
@@ -67,7 +67,907 @@
finiteR_mksch:
"Finite (mksch A B$tr$x$y) --> Finite tr"
-ML {* use_legacy_bindings (the_context ()) *}
+
+ML_setup {* change_simpset (fn ss => ss setmksym (K NONE)) *}
+
+
+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>>(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>>(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>>(mksch A B$xs
+ $schA
+ $(TL$(Dropwhile (%a. a:int B)$schB))
+ )))
+ else
+ UU
+ )
+ )
+ ))"
+apply (rule trans)
+apply (rule fix_eq2)
+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>>tr)$schA$schB =
+ (Takewhile (%a. a:int A)$schA)
+ @@ (x>>(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>>tr)$schA$schB =
+ (Takewhile (%a. a:int B)$schB)
+ @@ (x>>(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>>tr)$schA$schB =
+ (Takewhile (%a. a:int A)$schA)
+ @@ ((Takewhile (%a. a:int B)$schB)
+ @@ (x>>(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 {* COMPOSITIONALITY on TRACE Level *}
+
+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 >> g; x=(h x) |] ==> (f (h x)) = y >> g"
+by (erule subst)
+
+lemma ForallAorB_mksch [rule_format]:
+ "!!A B. compatible A B ==>
+ ! schA schB. Forall (%x. x:act (A||B)) tr
+ --> Forall (%x. x:act (A||B)) (mksch A B$tr$schA$schB)"
+apply (tactic {* Seq_induct_tac "tr" [thm "Forall_def", thm "sforall_def", thm "mksch_def"] 1 *})
+apply (tactic "safe_tac set_cs")
+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: ForallPTakewhileQ intA_is_not_actB int_is_act)
+apply (rule Forall_Conc_impl [THEN mp])
+apply (simp add: ForallPTakewhileQ intA_is_not_actB int_is_act)
+(* a:A,a~:B *)
+apply simp
+apply (rule Forall_Conc_impl [THEN mp])
+apply (simp add: ForallPTakewhileQ 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: ForallPTakewhileQ 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 {* Seq_induct_tac "tr" [thm "Forall_def", thm "sforall_def", thm "mksch_def"] 1 *})
+apply (tactic "safe_tac set_cs")
+apply (rule Forall_Conc_impl [THEN mp])
+apply (simp add: ForallPTakewhileQ 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 {* Seq_induct_tac "tr" [thm "Forall_def", thm "sforall_def", thm "mksch_def"] 1 *})
+apply (tactic "safe_tac set_cs")
+apply simp
+apply (rule Forall_Conc_impl [THEN mp])
+apply (simp add: ForallPTakewhileQ 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||B)) tr
+ --> Finite (mksch A B$tr$x$y)"
+
+apply (erule Seq_Finite_ind)
+apply simp
+(* main case *)
+apply simp
+apply (tactic "safe_tac set_cs")
+
+(* a: act A; a: act B *)
+apply (frule sym [THEN antisym_less_inverse, THEN conjunct1, THEN divide_Seq])
+apply (frule sym [THEN antisym_less_inverse, THEN conjunct1, 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 antisym_less_inverse, THEN conjunct1, 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 antisym_less_inverse, THEN conjunct1, 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 (fastsimp 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 antisym_less_inverse, THEN conjunct1, 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>>y1" in exI)
+apply (rule_tac x = "y2" in exI)
+(* elminate all obligations up to two depending on Conc_assoc *)
+apply (simp add: ForallPTakewhileQ 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]]
+
+declare FilterConc [simp del]
+
+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 antisym_less_inverse, THEN conjunct1, 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>>x1" in exI)
+apply (rule_tac x = "x2" in exI)
+(* elminate all obligations up to two depending on Conc_assoc *)
+apply (simp add: ForallPTakewhileQ 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||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||B))$(mksch A B$tr$schA$schB) = tr"
+
+apply (tactic {* Seq_induct_tac "tr" [thm "Forall_def", thm "sforall_def", thm "mksch_def"] 1 *})
+(* main case *)
+(* splitting into 4 cases according to a:A, a:B *)
+apply (tactic "safe_tac set_cs")
+
+(* 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 (fastsimp 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||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_lemmas)
+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' 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' 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 [THEN mp, COMP rev_contrapos] 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 antisym_less_inverse, THEN conjunct1, THEN divide_Seq])
+apply (erule conjE)+
+
+(* subst divide_Seq in conclusion, but only at the righest occurence *)
+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: FilterPTakewhileQid 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: Forall_Conc 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: ForallPTakewhileQ 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 antisym_less_inverse, THEN conjunct1, THEN divide_Seq])
+apply (erule conjE)+
+
+(* subst divide_Seq in conclusion, but only at the righest occurence *)
+apply (rule_tac t = "schA" in ssubst)
+back
+back
+back
+apply assumption
+
+(* f A (tw iA) = tw ~eA *)
+apply (simp add: FilterPTakewhileQid 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 antisym_less_inverse, THEN conjunct1, 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>>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||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_lemmas)
+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' 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' 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 [THEN mp, COMP rev_contrapos] 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 antisym_less_inverse, THEN conjunct1, THEN divide_Seq])
+apply (erule conjE)+
+
+(* subst divide_Seq in conclusion, but only at the righest occurence *)
+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: FilterPTakewhileQid 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: Forall_Conc 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: ForallPTakewhileQ 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 antisym_less_inverse, THEN conjunct1, THEN divide_Seq])
+apply (erule conjE)+
+
+(* subst divide_Seq in conclusion, but only at the righest occurence *)
+apply (rule_tac t = "schB" in ssubst)
+back
+back
+back
+apply assumption
+
+(* f B (tw iB) = tw ~eB *)
+apply (simp add: FilterPTakewhileQid 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 antisym_less_inverse, THEN conjunct1, 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>>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||B)) =
+ (Filter (%a. a:act A)$tr : traces A &
+ Filter (%a. a:act B)$tr : traces B &
+ Forall (%x. x:ext(A||B)) tr)"
+
+apply (simp (no_asm) add: traces_def has_trace_def)
+apply (tactic "safe_tac set_cs")
+
+(* ==> *)
+(* 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||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||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 {* COMPOSITIONALITY on TRACE Level -- for Modules *}
+
+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||B) = par_traces (Traces A) (Traces B)"
+
+apply (unfold Traces_def par_traces_def)
+apply (simp add: asig_of_par)
+apply (rule set_ext)
+apply (simp add: compositionality_tr externals_of_par)
+done
+
+
+ML_setup {* change_simpset (fn ss => ss setmksym (SOME o symmetric_fun)) *}
end
--- a/src/HOLCF/IOA/meta_theory/Compositionality.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,76 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/Compositionality.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-Goal "[|eA --> A ; eB & ~eA --> ~A|] ==> (eA | eB) --> A=eA";
-by Auto_tac;
-qed"compatibility_consequence3";
-
-
-Goal
-"!! 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";
-by (rtac ForallPFilterQR 1);
-(* i.e.: [| (! x. P x --> (Q x = R x)) ; Forall P tr |] ==> Filter Q$tr = Filter R$tr *)
-by (assume_tac 2);
-by (rtac compatibility_consequence3 1);
-by (REPEAT (asm_full_simp_tac (simpset()
- addsimps [ext_is_act,ext1_ext2_is_not_act1]) 1));
-qed"Filter_actAisFilter_extA";
-
-
-(* the next two theorems are only necessary, as there is no theorem ext (A||B) = ext (B||A) *)
-
-Goal "[|eA --> A ; eB & ~eA --> ~A|] ==> (eB | eA) --> A=eA";
-by Auto_tac;
-qed"compatibility_consequence4";
-
-Goal "[| compatible A B; Forall (%a. a: ext B | a: ext A) tr |] ==> \
-\ Filter (%a. a: act A)$tr= Filter (%a. a: ext A)$tr";
-by (rtac ForallPFilterQR 1);
-by (assume_tac 2);
-by (rtac compatibility_consequence4 1);
-by (REPEAT (asm_full_simp_tac (simpset()
- addsimps [ext_is_act,ext1_ext2_is_not_act1]) 1));
-qed"Filter_actAisFilter_extA2";
-
-
-(* -------------------------------------------------------------------------- *)
- section " Main Compositionality Theorem ";
-(* -------------------------------------------------------------------------- *)
-
-
-
-Goal "[| 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 || B1) =<| (A2 || B2)";
-
-by (asm_full_simp_tac (simpset() addsimps [is_asig_of_def]) 1);
-by (forw_inst_tac [("A1","A1")] (compat_commute RS iffD1) 1);
-by (forw_inst_tac [("A1","A2")] (compat_commute RS iffD1) 1);
-by (asm_full_simp_tac (simpset() addsimps [ioa_implements_def,
- inputs_of_par,outputs_of_par,externals_of_par]) 1);
-by (safe_tac set_cs);
-by (asm_full_simp_tac (simpset() addsimps [compositionality_tr]) 1);
-by (subgoal_tac "ext A1 = ext A2 & ext B1 = ext B2" 1);
-by (asm_full_simp_tac (simpset() addsimps [externals_def]) 2);
-by (REPEAT (etac conjE 1));
-(* rewrite with proven subgoal *)
-by (asm_full_simp_tac (simpset() addsimps [externals_of_par]) 1);
-by (safe_tac set_cs);
-
-(* 2 goals, the 3rd has been solved automatically *)
-(* 1: Filter A2 x : traces A2 *)
-by (dres_inst_tac [("A","traces A1")] subsetD 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [Filter_actAisFilter_extA]) 1);
-(* 2: Filter B2 x : traces B2 *)
-by (dres_inst_tac [("A","traces B1")] subsetD 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [Filter_actAisFilter_extA2]) 1);
-qed"compositionality";
--- a/src/HOLCF/IOA/meta_theory/Compositionality.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/Compositionality.thy Sun May 28 19:54:20 2006 +0200
@@ -8,4 +8,69 @@
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||B) = ext (B||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 || B1) =<| (A2 || 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 (tactic "safe_tac set_cs")
+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 (tactic "safe_tac set_cs")
+
+(* 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/HOLCF/IOA/meta_theory/Deadlock.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,83 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/Deadlock.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-(********************************************************************************
- input actions may always be added to a schedule
-**********************************************************************************)
-
-Goal "[| Filter (%x. x:act A)$sch : schedules A; a:inp A; input_enabled A; Finite sch|] \
-\ ==> Filter (%x. x:act A)$sch @@ a>>nil : schedules A";
-by (asm_full_simp_tac (simpset() addsimps [schedules_def,has_schedule_def]) 1);
-by (safe_tac set_cs);
-by (ftac inp_is_act 1);
-by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
-by (pair_tac "ex" 1);
-by (rename_tac "s ex" 1);
-by (subgoal_tac "Finite ex" 1);
-by (asm_full_simp_tac (simpset() addsimps [filter_act_def]) 2);
-by (rtac (Map2Finite RS iffD1) 2);
-by (res_inst_tac [("t","Map fst$ex")] subst 2);
-by (assume_tac 2);
-by (etac FiniteFilter 2);
-(* subgoal 1 *)
-by (ftac exists_laststate 1);
-by (etac allE 1);
-by (etac exE 1);
-(* using input-enabledness *)
-by (asm_full_simp_tac (simpset() addsimps [input_enabled_def]) 1);
-by (REPEAT (etac conjE 1));
-by (eres_inst_tac [("x","a")] allE 1);
-by (Asm_full_simp_tac 1);
-by (eres_inst_tac [("x","u")] allE 1);
-by (etac exE 1);
-(* instantiate execution *)
-by (res_inst_tac [("x","(s,ex @@ (a,s2)>>nil)")] exI 1);
-by (asm_full_simp_tac (simpset() addsimps [filter_act_def,MapConc]) 1);
-by (eres_inst_tac [("t","u")] lemma_2_1 1);
-by (Asm_full_simp_tac 1);
-by (rtac sym 1);
-by (assume_tac 1);
-qed"scheds_input_enabled";
-
-(********************************************************************************
- 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 @@
-**********************************************************************************)
-Delsplits [split_if];
-Goal "[| a : local A; Finite sch; sch : schedules (A||B); \
-\ Filter (%x. x:act A)$(sch @@ a>>nil) : schedules A; compatible A B; input_enabled B |] \
-\ ==> (sch @@ a>>nil) : schedules (A||B)";
-
-by (asm_full_simp_tac (simpset() addsimps [compositionality_sch,locals_def]) 1);
-by (rtac conjI 1);
-(* a : act (A||B) *)
-by (asm_full_simp_tac (simpset() addsimps [actions_of_par]) 2);
-by (blast_tac (claset() addDs [int_is_act,out_is_act]) 2);
-
-(* Filter B (sch@@[a]) : schedules B *)
-
-by (case_tac "a:int A" 1);
-by (dtac intA_is_not_actB 1);
-by (assume_tac 1); (* --> a~:act B *)
-by (Asm_full_simp_tac 1);
-
-(* case a~:int A , i.e. a:out A *)
-by (case_tac "a~:act B" 1);
-by (Asm_full_simp_tac 1);
-(* case a:act B *)
-by (Asm_full_simp_tac 1);
-by (subgoal_tac "a:out A" 1);
-by (Blast_tac 2);
-by (dtac outAactB_is_inpB 1);
-by (assume_tac 1);
-by (assume_tac 1);
-by (rtac scheds_input_enabled 1);
-by (Asm_full_simp_tac 1);
-by (REPEAT (atac 1));
-qed"IOA_deadlock_free";
-
-Addsplits [split_if];
--- a/src/HOLCF/IOA/meta_theory/Deadlock.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/Deadlock.thy Sun May 28 19:54:20 2006 +0200
@@ -9,4 +9,85 @@
imports RefCorrectness CompoScheds
begin
+text {* input actions may always be added to a schedule *}
+
+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>>nil : schedules A"
+apply (simp add: schedules_def has_schedule_def)
+apply (tactic "safe_tac set_cs")
+apply (frule inp_is_act)
+apply (simp add: executions_def)
+apply (tactic {* pair_tac "ex" 1 *})
+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) >>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 {*
+ 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 @@
+*}
+
+declare split_if [split del]
+lemma IOA_deadlock_free: "[| a : local A; Finite sch; sch : schedules (A||B);
+ Filter (%x. x:act A)$(sch @@ a>>nil) : schedules A; compatible A B; input_enabled B |]
+ ==> (sch @@ a>>nil) : schedules (A||B)"
+apply (simp add: compositionality_sch locals_def)
+apply (rule conjI)
+(* a : act (A||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/HOLCF/IOA/meta_theory/LiveIOA.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,51 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/LiveIOA.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-Delsimps [split_paired_Ex];
-
-Goalw [live_implements_def]
-"!!LC. [| live_implements (A,LA) (B,LB); live_implements (B,LB) (C,LC) |] \
-\ ==> live_implements (A,LA) (C,LC)";
-by Auto_tac;
-qed"live_implements_trans";
-
-
-section "Correctness of live refmap";
-
-
-(* ---------------------------------------------------------------- *)
-(* Correctness of live refmap *)
-(* ---------------------------------------------------------------- *)
-
-
-Goal "[| inp(C)=inp(A); out(C)=out(A); \
-\ is_live_ref_map f (C,M) (A,L) |] \
-\ ==> live_implements (C,M) (A,L)";
-
-by (asm_full_simp_tac (simpset() addsimps [is_live_ref_map_def, live_implements_def,
-livetraces_def,liveexecutions_def]) 1);
-by (safe_tac set_cs);
-by (res_inst_tac[("x","corresp_ex A f ex")] exI 1);
-by (safe_tac set_cs);
- (* Traces coincide, Lemma 1 *)
- by (pair_tac "ex" 1);
- by (etac (lemma_1 RS spec RS mp) 1);
- by (simp_tac (simpset() addsimps [externals_def])1);
- by (SELECT_GOAL (auto_tac (claset(),simpset()))1);
- by (asm_full_simp_tac (simpset() addsimps [executions_def,reachable.reachable_0]) 1);
-
- (* corresp_ex is execution, Lemma 2 *)
- by (pair_tac "ex" 1);
- by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
- (* start state *)
- by (rtac conjI 1);
- by (asm_full_simp_tac (simpset() addsimps [is_ref_map_def,corresp_ex_def]) 1);
- (* is-execution-fragment *)
- by (etac (lemma_2 RS spec RS mp) 1);
- by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1);
-
- (* Liveness *)
-by Auto_tac;
-qed"live_implements";
--- a/src/HOLCF/IOA/meta_theory/LiveIOA.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/LiveIOA.thy Sun May 28 19:54:20 2006 +0200
@@ -57,6 +57,45 @@
(! exec : executions (fst CL). (exec |== (snd CL)) -->
((corresp_ex (fst AM) f exec) |== (snd AM)))"
-ML {* use_legacy_bindings (the_context ()) *}
+
+declare split_paired_Ex [simp del]
+
+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 (tactic "safe_tac set_cs")
+apply (rule_tac x = "corresp_ex A f ex" in exI)
+apply (tactic "safe_tac set_cs")
+ (* Traces coincide, Lemma 1 *)
+ apply (tactic {* pair_tac "ex" 1 *})
+ 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 {* pair_tac "ex" 1 *})
+ 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)
+
+ (* Liveness *)
+apply auto
+done
end
--- a/src/HOLCF/IOA/meta_theory/Pred.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/Pred.thy Sun May 28 19:54:20 2006 +0200
@@ -67,6 +67,4 @@
IMPLIES_def:
"(P .--> Q) s == (P s) --> (Q s)"
-ML {* use_legacy_bindings (the_context ()) *}
-
end
--- a/src/HOLCF/IOA/meta_theory/RefCorrectness.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,333 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/RefCorrectness.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-
-(* -------------------------------------------------------------------------------- *)
-
-section "corresp_ex";
-
-
-(* ---------------------------------------------------------------- *)
-(* corresp_exC *)
-(* ---------------------------------------------------------------- *)
-
-
-Goal "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) ))";
-by (rtac trans 1);
-by (rtac fix_eq2 1);
-by (rtac corresp_exC_def 1);
-by (rtac beta_cfun 1);
-by (simp_tac (simpset() addsimps [flift1_def]) 1);
-qed"corresp_exC_unfold";
-
-Goal "(corresp_exC A f$UU) s=UU";
-by (stac corresp_exC_unfold 1);
-by (Simp_tac 1);
-qed"corresp_exC_UU";
-
-Goal "(corresp_exC A f$nil) s = nil";
-by (stac corresp_exC_unfold 1);
-by (Simp_tac 1);
-qed"corresp_exC_nil";
-
-Goal "(corresp_exC A f$(at>>xs)) s = \
-\ (@cex. move A cex (f s) (fst at) (f (snd at))) \
-\ @@ ((corresp_exC A f$xs) (snd at))";
-by (rtac trans 1);
-by (stac corresp_exC_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,flift1_def]) 1);
-by (Simp_tac 1);
-qed"corresp_exC_cons";
-
-
-Addsimps [corresp_exC_UU,corresp_exC_nil,corresp_exC_cons];
-
-
-
-(* ------------------------------------------------------------------ *)
-(* The following lemmata describe the definition *)
-(* of move in more detail *)
-(* ------------------------------------------------------------------ *)
-
-section"properties of move";
-
-Goalw [is_ref_map_def]
- "[|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)";
-
-by (subgoal_tac "? ex. move A ex (f s) a (f t)" 1);
-by (Asm_full_simp_tac 2);
-by (etac exE 1);
-by (rtac someI 1);
-by (assume_tac 1);
-qed"move_is_move";
-
-Goal
- "[|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))";
-by (cut_inst_tac [] move_is_move 1);
-by (REPEAT (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
-qed"move_subprop1";
-
-Goal
- "[|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)))";
-by (cut_inst_tac [] move_is_move 1);
-by (REPEAT (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
-qed"move_subprop2";
-
-Goal
- "[|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)";
-by (cut_inst_tac [] move_is_move 1);
-by (REPEAT (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
-qed"move_subprop3";
-
-Goal
- "[|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>>nil else nil)";
-
-by (cut_inst_tac [] move_is_move 1);
-by (REPEAT (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
-qed"move_subprop4";
-
-
-(* ------------------------------------------------------------------ *)
-(* The following lemmata contribute to *)
-(* TRACE INCLUSION Part 1: Traces coincide *)
-(* ------------------------------------------------------------------ *)
-
-section "Lemmata for <==";
-
-(* --------------------------------------------------- *)
-(* Lemma 1.1: Distribution of mk_trace and @@ *)
-(* --------------------------------------------------- *)
-
-
-Goal "mk_trace C$(ex1 @@ ex2)= (mk_trace C$ex1) @@ (mk_trace C$ex2)";
-by (simp_tac (simpset() addsimps [mk_trace_def,filter_act_def,
- FilterConc,MapConc]) 1);
-qed"mk_traceConc";
-
-
-
-(* ------------------------------------------------------
- Lemma 1 :Traces coincide
- ------------------------------------------------------- *)
-Delsplits[split_if];
-Goalw [corresp_ex_def]
- "[|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)))";
-
-by (pair_induct_tac "xs" [is_exec_frag_def] 1);
-(* cons case *)
-by (safe_tac set_cs);
-by (asm_full_simp_tac (simpset() addsimps [mk_traceConc]) 1);
-by (forward_tac [reachable.reachable_n] 1);
-by (assume_tac 1);
-by (eres_inst_tac [("x","y")] allE 1);
-by (Asm_full_simp_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [move_subprop4]
- addsplits [split_if]) 1);
-qed"lemma_1";
-Addsplits[split_if];
-
-(* ------------------------------------------------------------------ *)
-(* The following lemmata contribute to *)
-(* TRACE INCLUSION Part 2: corresp_ex is execution *)
-(* ------------------------------------------------------------------ *)
-
-section "Lemmata for ==>";
-
-(* -------------------------------------------------- *)
-(* Lemma 2.1 *)
-(* -------------------------------------------------- *)
-
-Goal
-"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))";
-
-by (rtac impI 1);
-by (Seq_Finite_induct_tac 1);
-(* main case *)
-by (safe_tac set_cs);
-by (pair_tac "a" 1);
-qed_spec_mp"lemma_2_1";
-
-
-(* ----------------------------------------------------------- *)
-(* Lemma 2 : corresp_ex is execution *)
-(* ----------------------------------------------------------- *)
-
-
-
-Goalw [corresp_ex_def]
- "[| 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))";
-
-by (Asm_full_simp_tac 1);
-by (pair_induct_tac "xs" [is_exec_frag_def] 1);
-(* main case *)
-by (safe_tac set_cs);
-by (res_inst_tac [("t","f y")] lemma_2_1 1);
-
-(* Finite *)
-by (etac move_subprop2 1);
-by (REPEAT (atac 1));
-by (rtac conjI 1);
-
-(* is_exec_frag *)
-by (etac move_subprop1 1);
-by (REPEAT (atac 1));
-by (rtac conjI 1);
-
-(* Induction hypothesis *)
-(* reachable_n looping, therefore apply it manually *)
-by (eres_inst_tac [("x","y")] allE 1);
-by (Asm_full_simp_tac 1);
-by (forward_tac [reachable.reachable_n] 1);
-by (assume_tac 1);
-by (Asm_full_simp_tac 1);
-(* laststate *)
-by (etac (move_subprop3 RS sym) 1);
-by (REPEAT (atac 1));
-qed"lemma_2";
-
-
-(* -------------------------------------------------------------------------------- *)
-
-section "Main Theorem: T R A C E - I N C L U S I O N";
-
-(* -------------------------------------------------------------------------------- *)
-
-
-Goalw [traces_def]
- "[| ext C = ext A; is_ref_map f C A |] \
-\ ==> traces C <= traces A";
-
- by (simp_tac(simpset() addsimps [has_trace_def2])1);
- by (safe_tac set_cs);
-
- (* give execution of abstract automata *)
- by (res_inst_tac[("x","corresp_ex A f ex")] bexI 1);
-
- (* Traces coincide, Lemma 1 *)
- by (pair_tac "ex" 1);
- by (etac (lemma_1 RS spec RS mp) 1);
- by (REPEAT (atac 1));
- by (asm_full_simp_tac (simpset() addsimps [executions_def,reachable.reachable_0]) 1);
-
- (* corresp_ex is execution, Lemma 2 *)
- by (pair_tac "ex" 1);
- by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
- (* start state *)
- by (rtac conjI 1);
- by (asm_full_simp_tac (simpset() addsimps [is_ref_map_def,corresp_ex_def]) 1);
- (* is-execution-fragment *)
- by (etac (lemma_2 RS spec RS mp) 1);
- by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1);
-qed"trace_inclusion";
-
-
-(* -------------------------------------------------------------------------------- *)
-
-section "Corollary: F A I R T R A C E - I N C L U S I O N";
-
-(* -------------------------------------------------------------------------------- *)
-
-
-Goalw [fin_often_def] "(~inf_often P s) = fin_often P s";
-by Auto_tac;
-qed"fininf";
-
-
-Goal "is_wfair A W (s,ex) = \
-\ (fin_often (%x. ~Enabled A W (snd x)) ex --> inf_often (%x. fst x :W) ex)";
-by (asm_full_simp_tac (simpset() addsimps [is_wfair_def,fin_often_def])1);
-by Auto_tac;
-qed"WF_alt";
-
-Goal "[|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";
-by (dtac persistent 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [WF_alt])1);
-by Auto_tac;
-qed"WF_persistent";
-
-
-Goal "!! 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";
-by (simp_tac (simpset() addsimps [fairtraces_def,
- fairexecutions_def]) 1);
-by (safe_tac set_cs);
-by (res_inst_tac[("x","corresp_ex A f ex")] exI 1);
-by (safe_tac set_cs);
-
- (* Traces coincide, Lemma 1 *)
- by (pair_tac "ex" 1);
- by (etac (lemma_1 RS spec RS mp) 1);
- by (REPEAT (atac 1));
- by (asm_full_simp_tac (simpset() addsimps [executions_def,reachable.reachable_0]) 1);
-
- (* corresp_ex is execution, Lemma 2 *)
- by (pair_tac "ex" 1);
- by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
- (* start state *)
- by (rtac conjI 1);
- by (asm_full_simp_tac (simpset() addsimps [is_ref_map_def,corresp_ex_def]) 1);
- (* is-execution-fragment *)
- by (etac (lemma_2 RS spec RS mp) 1);
- by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1);
-
- (* Fairness *)
-by Auto_tac;
-qed"fair_trace_inclusion";
-
-Goal "!! C A. \
-\ [| inp(C) = inp(A); out(C)=out(A); \
-\ is_fair_ref_map f C A |] \
-\ ==> fair_implements C A";
-by (asm_full_simp_tac (simpset() addsimps [is_fair_ref_map_def, fair_implements_def,
- fairtraces_def, fairexecutions_def]) 1);
-by (safe_tac set_cs);
-by (res_inst_tac[("x","corresp_ex A f ex")] exI 1);
-by (safe_tac set_cs);
- (* Traces coincide, Lemma 1 *)
- by (pair_tac "ex" 1);
- by (etac (lemma_1 RS spec RS mp) 1);
- by (simp_tac (simpset() addsimps [externals_def])1);
- by (SELECT_GOAL (auto_tac (claset(),simpset()))1);
- by (asm_full_simp_tac (simpset() addsimps [executions_def,reachable.reachable_0]) 1);
-
- (* corresp_ex is execution, Lemma 2 *)
- by (pair_tac "ex" 1);
- by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
- (* start state *)
- by (rtac conjI 1);
- by (asm_full_simp_tac (simpset() addsimps [is_ref_map_def,corresp_ex_def]) 1);
- (* is-execution-fragment *)
- by (etac (lemma_2 RS spec RS mp) 1);
- by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1);
-
- (* Fairness *)
-by Auto_tac;
-qed"fair_trace_inclusion2";
--- a/src/HOLCF/IOA/meta_theory/RefCorrectness.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/RefCorrectness.thy Sun May 28 19:54:20 2006 +0200
@@ -67,6 +67,317 @@
"[| is_exec_frag A (s,ex); inf_often (%x. fst x:W) ex|]
==> inf_often (% x. set_was_enabled A W (snd x)) ex"
-ML {* use_legacy_bindings (the_context ()) *}
+
+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 (rule 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>>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>>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 FilterConc 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 {* pair_induct_tac "xs" [thm "is_exec_frag_def"] 1 *})
+(* cons case *)
+apply (tactic "safe_tac set_cs")
+apply (simp add: mk_traceConc)
+apply (frule reachable.reachable_n)
+apply assumption
+apply (erule_tac x = "y" in allE)
+apply simp
+apply (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 {* Seq_Finite_induct_tac 1 *})
+(* main case *)
+apply (tactic "safe_tac set_cs")
+apply (tactic {* pair_tac "a" 1 *})
+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 {* pair_induct_tac "xs" [thm "is_exec_frag_def"] 1 *})
+(* main case *)
+apply (tactic "safe_tac set_cs")
+apply (rule_tac t = "f y" 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 = "y" 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 (tactic "safe_tac set_cs")
+
+ (* give execution of abstract automata *)
+ apply (rule_tac x = "corresp_ex A f ex" in bexI)
+
+ (* Traces coincide, Lemma 1 *)
+ apply (tactic {* pair_tac "ex" 1 *})
+ 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 {* pair_tac "ex" 1 *})
+ 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 (tactic "safe_tac set_cs")
+apply (rule_tac x = "corresp_ex A f ex" in exI)
+apply (tactic "safe_tac set_cs")
+
+ (* Traces coincide, Lemma 1 *)
+ apply (tactic {* pair_tac "ex" 1 *})
+ 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 {* pair_tac "ex" 1 *})
+ 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)
+
+ (* Fairness *)
+apply auto
+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 (tactic "safe_tac set_cs")
+apply (rule_tac x = "corresp_ex A f ex" in exI)
+apply (tactic "safe_tac set_cs")
+ (* Traces coincide, Lemma 1 *)
+ apply (tactic {* pair_tac "ex" 1 *})
+ 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 {* pair_tac "ex" 1 *})
+ 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)
+
+ (* Fairness *)
+apply auto
+done
+
end
--- a/src/HOLCF/IOA/meta_theory/RefMappings.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,105 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/RefMappings.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-
-(* ---------------------------------------------------------------------------- *)
-
-section "transitions and moves";
-
-
-Goal "s -a--A-> t ==> ? ex. move A ex s a t";
-
-by (res_inst_tac [("x","(a,t)>>nil")] exI 1);
-by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
-qed"transition_is_ex";
-
-
-Goal "(~a:ext A) & s=t ==> ? ex. move A ex s a t";
-
-by (res_inst_tac [("x","nil")] exI 1);
-by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
-qed"nothing_is_ex";
-
-
-Goal "(s -a--A-> s') & (s' -a'--A-> s'') & (~a':ext A) \
-\ ==> ? ex. move A ex s a s''";
-
-by (res_inst_tac [("x","(a,s')>>(a',s'')>>nil")] exI 1);
-by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
-qed"ei_transitions_are_ex";
-
-
-Goal "(s1 -a1--A-> s2) & (s2 -a2--A-> s3) & (s3 -a3--A-> s4) &\
-\ (~a2:ext A) & (~a3:ext A) ==> \
-\ ? ex. move A ex s1 a1 s4";
-
-by (res_inst_tac [("x","(a1,s2)>>(a2,s3)>>(a3,s4)>>nil")] exI 1);
-by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);
-qed"eii_transitions_are_ex";
-
-
-(* ---------------------------------------------------------------------------- *)
-
-section "weak_ref_map and ref_map";
-
-
-Goalw [is_weak_ref_map_def,is_ref_map_def]
- "[| ext C = ext A; \
-\ is_weak_ref_map f C A |] ==> is_ref_map f C A";
-by (safe_tac set_cs);
-by (case_tac "a:ext A" 1);
-by (rtac transition_is_ex 1);
-by (Asm_simp_tac 1);
-by (rtac nothing_is_ex 1);
-by (Asm_simp_tac 1);
-qed"weak_ref_map2ref_map";
-
-
-val prems = Goal "(P ==> Q-->R) ==> P&Q --> R";
- by (fast_tac (claset() addDs prems) 1);
-qed "imp_conj_lemma";
-
-Delsplits [split_if]; Delcongs [if_weak_cong];
-
-Goal "[| is_weak_ref_map f C A |] \
-\ ==> (is_weak_ref_map f (rename C g) (rename A g))";
-by (asm_full_simp_tac (simpset() addsimps [is_weak_ref_map_def]) 1);
-by (rtac conjI 1);
-(* 1: start states *)
-by (asm_full_simp_tac (simpset() addsimps [rename_def,rename_set_def,starts_of_def]) 1);
-(* 2: reachable transitions *)
-by (REPEAT (rtac allI 1));
-by (rtac imp_conj_lemma 1);
-by (simp_tac (simpset() addsimps [rename_def,rename_set_def]) 1);
-by (asm_full_simp_tac (simpset() addsimps [externals_def,asig_inputs_def,
-asig_outputs_def,asig_of_def,trans_of_def]) 1);
-by Safe_tac;
-by (stac split_if 1);
- by (rtac conjI 1);
- by (rtac impI 1);
- by (etac disjE 1);
- by (etac exE 1);
-by (etac conjE 1);
-(* x is input *)
- by (dtac sym 1);
- by (dtac sym 1);
-by (Asm_full_simp_tac 1);
-by (REPEAT (hyp_subst_tac 1));
-by (ftac reachable_rename 1);
-by (Asm_full_simp_tac 1);
-(* x is output *)
- by (etac exE 1);
-by (etac conjE 1);
- by (dtac sym 1);
- by (dtac sym 1);
-by (Asm_full_simp_tac 1);
-by (REPEAT (hyp_subst_tac 1));
-by (ftac reachable_rename 1);
-by (Asm_full_simp_tac 1);
-(* x is internal *)
-by (ftac reachable_rename 1);
-by Auto_tac;
-qed"rename_through_pmap";
-Addsplits [split_if]; Addcongs [if_weak_cong];
--- a/src/HOLCF/IOA/meta_theory/RefMappings.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/RefMappings.thy Sun May 28 19:54:20 2006 +0200
@@ -41,6 +41,98 @@
then (f s) -a--A-> (f t)
else (f s)=(f t)))"
-ML {* use_legacy_bindings (the_context ()) *}
+
+subsection "transitions and moves"
+
+
+lemma transition_is_ex: "s -a--A-> t ==> ? ex. move A ex s a t"
+apply (rule_tac x = " (a,t) >>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 -a--A-> s') & (s' -a'--A-> s'') & (~a':ext A)
+ ==> ? ex. move A ex s a s''"
+apply (rule_tac x = " (a,s') >> (a',s'') >>nil" in exI)
+apply (simp add: move_def)
+done
+
+
+lemma eii_transitions_are_ex: "(s1 -a1--A-> s2) & (s2 -a2--A-> s3) & (s3 -a3--A-> s4) &
+ (~a2:ext A) & (~a3:ext A) ==>
+ ? ex. move A ex s1 a1 s4"
+apply (rule_tac x = " (a1,s2) >> (a2,s3) >> (a3,s4) >>nil" in exI)
+apply (simp add: move_def)
+done
+
+
+subsection "weak_ref_map and ref_map"
+
+lemma imp_conj_lemma:
+ "[| 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 (tactic "safe_tac set_cs")
+apply (case_tac "a:ext A")
+apply (rule transition_is_ex)
+apply (simp (no_asm_simp))
+apply (rule nothing_is_ex)
+apply simp
+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/HOLCF/IOA/meta_theory/Sequence.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,32 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/Sequence.ML
- ID: $Id$
- Author: Olaf Mller
-*)
-
-(* ----------------------------------------------------------------------------------- *)
-
-fun Seq_case_tac s i = res_inst_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>>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;
-
-(* rws are definitions to be unfolded for admissibility check *)
-fun Seq_induct_tac s rws i = res_inst_tac [("x",s)] Seq_induct i
- THEN (REPEAT_DETERM (CHANGED (Asm_simp_tac (i+1))))
- THEN simp_tac (simpset() addsimps rws) i;
-
-fun Seq_Finite_induct_tac i = etac Seq_Finite_ind i
- THEN (REPEAT_DETERM (CHANGED (Asm_simp_tac i)));
-
-fun pair_tac s = res_inst_tac [("p",s)] PairE
- THEN' hyp_subst_tac THEN' Asm_full_simp_tac;
-
-(* induction on a sequence of pairs with pairsplitting and simplification *)
-fun pair_induct_tac s rws i =
- res_inst_tac [("x",s)] Seq_induct i
- THEN pair_tac "a" (i+3)
- THEN (REPEAT_DETERM (CHANGED (Simp_tac (i+1))))
- THEN simp_tac (simpset() addsimps rws) i;
--- a/src/HOLCF/IOA/meta_theory/Sequence.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/Sequence.thy Sun May 28 19:54:20 2006 +0200
@@ -1124,6 +1124,45 @@
apply auto
done
-ML {* use_legacy_bindings (the_context ()) *}
+
+
+ML {*
+
+local
+ val Seq_cases = thm "Seq_cases"
+ val Seq_induct = thm "Seq_induct"
+ val Seq_Finite_ind = thm "Seq_Finite_ind"
+in
+
+fun Seq_case_tac s i = res_inst_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>>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 SIMPSET' asm_simp_tac (i+2)
+ THEN SIMPSET' asm_full_simp_tac (i+1)
+ THEN SIMPSET' asm_full_simp_tac i;
+
+(* rws are definitions to be unfolded for admissibility check *)
+fun Seq_induct_tac s rws i = res_inst_tac [("x",s)] Seq_induct i
+ THEN (REPEAT_DETERM (CHANGED (SIMPSET' asm_simp_tac (i+1))))
+ THEN SIMPSET' (fn ss => simp_tac (ss addsimps rws)) i;
+
+fun Seq_Finite_induct_tac i = etac Seq_Finite_ind i
+ THEN (REPEAT_DETERM (CHANGED (SIMPSET' asm_simp_tac i)));
+
+fun pair_tac s = res_inst_tac [("p",s)] PairE
+ THEN' hyp_subst_tac THEN' SIMPSET' asm_full_simp_tac;
+
+(* induction on a sequence of pairs with pairsplitting and simplification *)
+fun pair_induct_tac s rws i =
+ res_inst_tac [("x",s)] Seq_induct i
+ THEN pair_tac "a" (i+3)
+ THEN (REPEAT_DETERM (CHANGED (SIMPSET' simp_tac (i+1))))
+ THEN SIMPSET' (fn ss => simp_tac (ss addsimps rws)) i;
end
+
+*}
+
+
+end
--- a/src/HOLCF/IOA/meta_theory/ShortExecutions.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,252 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/ShortExecutions.thy
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-
-fun thin_tac' j =
- rotate_tac (j - 1) THEN'
- etac thin_rl THEN'
- rotate_tac (~ (j - 1));
-
-
-
-(* ---------------------------------------------------------------- *)
- section "oraclebuild rewrite rules";
-(* ---------------------------------------------------------------- *)
-
-
-bind_thm ("oraclebuild_unfold", fix_prover2 (the_context ()) oraclebuild_def
-"oraclebuild P = (LAM s t. case t of \
-\ nil => nil\
-\ | x##xs => \
-\ (case x of\
-\ UU => UU\
-\ | Def y => (Takewhile (%a.~ P a)$s)\
-\ @@ (y>>(oraclebuild P$(TL$(Dropwhile (%a.~ P a)$s))$xs))\
-\ )\
-\ )");
-
-Goal "oraclebuild P$sch$UU = UU";
-by (stac oraclebuild_unfold 1);
-by (Simp_tac 1);
-qed"oraclebuild_UU";
-
-Goal "oraclebuild P$sch$nil = nil";
-by (stac oraclebuild_unfold 1);
-by (Simp_tac 1);
-qed"oraclebuild_nil";
-
-Goal "oraclebuild P$s$(x>>t) = \
-\ (Takewhile (%a.~ P a)$s) \
-\ @@ (x>>(oraclebuild P$(TL$(Dropwhile (%a.~ P a)$s))$t))";
-by (rtac trans 1);
-by (stac oraclebuild_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def]) 1);
-by (simp_tac (simpset() addsimps [Consq_def]) 1);
-qed"oraclebuild_cons";
-
-
-
-
-(* ---------------------------------------------------------------- *)
- section "Cut rewrite rules";
-(* ---------------------------------------------------------------- *)
-
-Goalw [Cut_def]
-"[| Forall (%a.~ P a) s; Finite s|] \
-\ ==> Cut P s =nil";
-by (subgoal_tac "Filter P$s = nil" 1);
-by (asm_simp_tac (simpset() addsimps [oraclebuild_nil]) 1);
-by (rtac ForallQFilterPnil 1);
-by (REPEAT (atac 1));
-qed"Cut_nil";
-
-Goalw [Cut_def]
-"[| Forall (%a.~ P a) s; ~Finite s|] \
-\ ==> Cut P s =UU";
-by (subgoal_tac "Filter P$s= UU" 1);
-by (asm_simp_tac (simpset() addsimps [oraclebuild_UU]) 1);
-by (rtac ForallQFilterPUU 1);
-by (REPEAT (atac 1));
-qed"Cut_UU";
-
-Goalw [Cut_def]
-"[| P t; Forall (%x.~ P x) ss; Finite ss|] \
-\ ==> Cut P (ss @@ (t>> rs)) \
-\ = ss @@ (t >> Cut P rs)";
-
-by (asm_full_simp_tac (simpset() addsimps [ForallQFilterPnil,oraclebuild_cons,
- TakewhileConc,DropwhileConc]) 1);
-qed"Cut_Cons";
-
-
-(* ---------------------------------------------------------------- *)
- section "Cut lemmas for main theorem";
-(* ---------------------------------------------------------------- *)
-
-
-Goal "Filter P$s = Filter P$(Cut P s)";
-
-by (res_inst_tac [("A1","%x. True")
- ,("Q1","%x.~ P x"), ("x1","s")]
- (take_lemma_induct RS mp) 1);
-by (Fast_tac 3);
-by (case_tac "Finite s" 1);
-by (asm_full_simp_tac (simpset() addsimps [Cut_nil,
- ForallQFilterPnil]) 1);
-by (asm_full_simp_tac (simpset() addsimps [Cut_UU,
- ForallQFilterPUU]) 1);
-(* main case *)
-by (asm_full_simp_tac (simpset() addsimps [Cut_Cons,ForallQFilterPnil]) 1);
-by Auto_tac;
-qed"FilterCut";
-
-
-Goal "Cut P (Cut P s) = (Cut P s)";
-
-by (res_inst_tac [("A1","%x. True")
- ,("Q1","%x.~ P x"), ("x1","s")]
- (take_lemma_less_induct RS mp) 1);
-by (Fast_tac 3);
-by (case_tac "Finite s" 1);
-by (asm_full_simp_tac (simpset() addsimps [Cut_nil,
- ForallQFilterPnil]) 1);
-by (asm_full_simp_tac (simpset() addsimps [Cut_UU,
- ForallQFilterPUU]) 1);
-(* main case *)
-by (asm_full_simp_tac (simpset() addsimps [Cut_Cons,ForallQFilterPnil]) 1);
-by (rtac take_reduction 1);
-by Auto_tac;
-qed"Cut_idemp";
-
-
-Goal "Map f$(Cut (P o f) s) = Cut P (Map f$s)";
-
-by (res_inst_tac [("A1","%x. True")
- ,("Q1","%x.~ P (f x)"), ("x1","s")]
- (take_lemma_less_induct RS mp) 1);
-by (Fast_tac 3);
-by (case_tac "Finite s" 1);
-by (asm_full_simp_tac (simpset() addsimps [Cut_nil]) 1);
-by (rtac (Cut_nil RS sym) 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallMap,o_def]) 1);
-by (asm_full_simp_tac (simpset() addsimps [Map2Finite]) 1);
-(* csae ~ Finite s *)
-by (asm_full_simp_tac (simpset() addsimps [Cut_UU]) 1);
-by (rtac (Cut_UU (*RS sym*)) 1);
-by (asm_full_simp_tac (simpset() addsimps [ForallMap,o_def]) 1);
-by (asm_full_simp_tac (simpset() addsimps [Map2Finite]) 1);
-(* main case *)
-by (asm_full_simp_tac (simpset() addsimps [Cut_Cons,MapConc,
- ForallMap,FiniteMap1,o_def]) 1);
-by (rtac take_reduction 1);
-by Auto_tac;
-qed"MapCut";
-
-
-Goal "~Finite s --> Cut P s << s";
-by (strip_tac 1);
-by (rtac (take_lemma_less RS iffD1) 1);
-by (strip_tac 1);
-by (rtac mp 1);
-by (assume_tac 2);
-by (thin_tac' 1 1);
-by (res_inst_tac [("x","s")] spec 1);
-by (rtac nat_less_induct 1);
-by (strip_tac 1);
-by (rename_tac "na n s" 1);
-by (case_tac "Forall (%x. ~ P x) s" 1);
-by (rtac (take_lemma_less RS iffD2 RS spec) 1);
-by (asm_full_simp_tac (simpset() addsimps [Cut_UU]) 1);
-(* main case *)
-by (dtac divide_Seq3 1);
-by (REPEAT (etac exE 1));
-by (REPEAT (etac conjE 1));
-by (hyp_subst_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [Cut_Cons]) 1);
-by (rtac take_reduction_less 1);
-(* auto makes also reasoning about Finiteness of parts of s ! *)
-by Auto_tac;
-qed_spec_mp"Cut_prefixcl_nFinite";
-
-
-
-Goal "!!ex .is_exec_frag A (s,ex) ==> is_exec_frag A (s,Cut P ex)";
-by (case_tac "Finite ex" 1);
-by (cut_inst_tac [("s","ex"),("P","P")] Cut_prefixcl_Finite 1);
-by (assume_tac 1);
-by (etac exE 1);
-by (rtac exec_prefix2closed 1);
-by (eres_inst_tac [("s","ex"),("t","Cut P ex @@ y")] subst 1);
-by (assume_tac 1);
-by (etac exec_prefixclosed 1);
-by (etac Cut_prefixcl_nFinite 1);
-qed"execThruCut";
-
-
-
-(* ---------------------------------------------------------------- *)
- section "Main Cut Theorem";
-(* ---------------------------------------------------------------- *)
-
-
-
-Goalw [schedules_def,has_schedule_def]
- "[|sch : schedules A ; tr = Filter (%a. a:ext A)$sch|] \
-\ ==> ? sch. sch : schedules A & \
-\ tr = Filter (%a. a:ext A)$sch &\
-\ LastActExtsch A sch";
-
-by (safe_tac set_cs);
-by (res_inst_tac [("x","filter_act$(Cut (%a. fst a:ext A) (snd ex))")] exI 1);
-by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
-by (pair_tac "ex" 1);
-by (safe_tac set_cs);
-by (res_inst_tac [("x","(x,Cut (%a. fst a:ext A) y)")] exI 1);
-by (Asm_simp_tac 1);
-
-(* Subgoal 1: Lemma: propagation of execution through Cut *)
-
-by (asm_full_simp_tac (simpset() addsimps [execThruCut]) 1);
-
-(* Subgoal 2: Lemma: Filter P s = Filter P (Cut P s) *)
-
-by (simp_tac (simpset() addsimps [filter_act_def]) 1);
-by (subgoal_tac "Map fst$(Cut (%a. fst a: ext A) y)= Cut (%a. a:ext A) (Map fst$y)" 1);
-
-by (rtac (rewrite_rule [o_def] MapCut) 2);
-by (asm_full_simp_tac (simpset() addsimps [FilterCut RS sym]) 1);
-
-(* Subgoal 3: Lemma: Cut idempotent *)
-
-by (simp_tac (simpset() addsimps [LastActExtsch_def,filter_act_def]) 1);
-by (subgoal_tac "Map fst$(Cut (%a. fst a: ext A) y)= Cut (%a. a:ext A) (Map fst$y)" 1);
-by (rtac (rewrite_rule [o_def] MapCut) 2);
-by (asm_full_simp_tac (simpset() addsimps [Cut_idemp]) 1);
-qed"exists_LastActExtsch";
-
-
-(* ---------------------------------------------------------------- *)
- section "Further Cut lemmas";
-(* ---------------------------------------------------------------- *)
-
-Goalw [LastActExtsch_def]
- "[| LastActExtsch A sch; Filter (%x. x:ext A)$sch = nil |] \
-\ ==> sch=nil";
-by (dtac FilternPnilForallP 1);
-by (etac conjE 1);
-by (dtac Cut_nil 1);
-by (assume_tac 1);
-by (Asm_full_simp_tac 1);
-qed"LastActExtimplnil";
-
-Goalw [LastActExtsch_def]
- "[| LastActExtsch A sch; Filter (%x. x:ext A)$sch = UU |] \
-\ ==> sch=UU";
-by (dtac FilternPUUForallP 1);
-by (etac conjE 1);
-by (dtac Cut_UU 1);
-by (assume_tac 1);
-by (Asm_full_simp_tac 1);
-qed"LastActExtimplUU";
--- a/src/HOLCF/IOA/meta_theory/ShortExecutions.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/ShortExecutions.thy Sun May 28 19:54:20 2006 +0200
@@ -57,6 +57,234 @@
LastActExtsmall2:
"[| Finite sch1; LastActExtsch A (sch1 @@ sch2) |] ==> LastActExtsch A sch2"
-ML {* use_legacy_bindings (the_context ()) *}
+
+
+ML {*
+fun thin_tac' j =
+ rotate_tac (j - 1) THEN'
+ etac thin_rl THEN'
+ rotate_tac (~ (j - 1))
+*}
+
+
+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.~ P a)$s)
+ @@ (y>>(oraclebuild P$(TL$(Dropwhile (%a.~ P a)$s))$xs))
+ )
+ )"
+apply (rule trans)
+apply (rule fix_eq2)
+apply (rule 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>>t) =
+ (Takewhile (%a.~ P a)$s)
+ @@ (x>>(oraclebuild P$(TL$(Dropwhile (%a.~ 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.~ 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.~ 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.~ P x) ss; Finite ss|]
+ ==> Cut P (ss @@ (t>> rs))
+ = ss @@ (t >> 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.~ 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.~ 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.~ 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' 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)
+apply (tactic "safe_tac set_cs")
+apply (rule_tac x = "filter_act$ (Cut (%a. fst a:ext A) (snd ex))" in exI)
+apply (simp add: executions_def)
+apply (tactic {* pair_tac "ex" 1 *})
+apply (tactic "safe_tac set_cs")
+apply (rule_tac x = " (x,Cut (%a. fst a:ext A) y) " 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) y) = Cut (%a. a:ext A) (Map fst$y) ")
+
+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) y) = Cut (%a. a:ext A) (Map fst$y) ")
+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/HOLCF/IOA/meta_theory/SimCorrectness.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,290 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/SimCorrectness.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-(* -------------------------------------------------------------------------------- *)
-
-section "corresp_ex_sim";
-
-
-(* ---------------------------------------------------------------- *)
-(* corresp_ex_simC *)
-(* ---------------------------------------------------------------- *)
-
-
-Goal "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) ))";
-by (rtac trans 1);
-by (rtac fix_eq2 1);
-by (rtac corresp_ex_simC_def 1);
-by (rtac beta_cfun 1);
-by (simp_tac (simpset() addsimps [flift1_def]) 1);
-qed"corresp_ex_simC_unfold";
-
-Goal "(corresp_ex_simC A R$UU) s=UU";
-by (stac corresp_ex_simC_unfold 1);
-by (Simp_tac 1);
-qed"corresp_ex_simC_UU";
-
-Goal "(corresp_ex_simC A R$nil) s = nil";
-by (stac corresp_ex_simC_unfold 1);
-by (Simp_tac 1);
-qed"corresp_ex_simC_nil";
-
-Goal "(corresp_ex_simC A R$((a,t)>>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'))";
-by (rtac trans 1);
-by (stac corresp_ex_simC_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,flift1_def]) 1);
-by (Simp_tac 1);
-qed"corresp_ex_simC_cons";
-
-
-Addsimps [corresp_ex_simC_UU,corresp_ex_simC_nil,corresp_ex_simC_cons];
-
-
-
-(* ------------------------------------------------------------------ *)
-(* The following lemmata describe the definition *)
-(* of move in more detail *)
-(* ------------------------------------------------------------------ *)
-
-section"properties of move";
-
-Delsimps [Let_def];
-
-Goalw [is_simulation_def]
- "[|is_simulation R C A; reachable C s; s -a--C-> 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'";
-
-(* Does not perform conditional rewriting on assumptions automatically as
- usual. Instantiate all variables per hand. Ask Tobias?? *)
-by (subgoal_tac "? t' ex. (t,t'):R & move A ex s' a t'" 1);
-by (Asm_full_simp_tac 2);
-by (etac conjE 2);
-by (eres_inst_tac [("x","s")] allE 2);
-by (eres_inst_tac [("x","s'")] allE 2);
-by (eres_inst_tac [("x","t")] allE 2);
-by (eres_inst_tac [("x","a")] allE 2);
-by (Asm_full_simp_tac 2);
-(* Go on as usual *)
-by (etac exE 1);
-by (dres_inst_tac [("x","t'"),
- ("P","%t'. ? ex.(t,t'):R & move A ex s' a t'")] someI 1);
-by (etac exE 1);
-by (etac conjE 1);
-by (asm_full_simp_tac (simpset() addsimps [Let_def]) 1);
-by (res_inst_tac [("x","ex")] someI 1);
-by (etac conjE 1);
-by (assume_tac 1);
-qed"move_is_move_sim";
-
-
-Addsimps [Let_def];
-
-Goal
- "[|is_simulation R C A; reachable C s; s-a--C-> 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')";
-by (cut_inst_tac [] move_is_move_sim 1);
-by (REPEAT (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [move_def,Let_def]) 1);
-qed"move_subprop1_sim";
-
-Goal
- "[|is_simulation R C A; reachable C s; s-a--C-> 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')";
-by (cut_inst_tac [] move_is_move_sim 1);
-by (REPEAT (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [move_def,Let_def]) 1);
-qed"move_subprop2_sim";
-
-Goal
- "[|is_simulation R C A; reachable C s; s-a--C-> 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'";
-by (cut_inst_tac [] move_is_move_sim 1);
-by (REPEAT (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [move_def,Let_def]) 1);
-qed"move_subprop3_sim";
-
-Goal
- "[|is_simulation R C A; reachable C s; s-a--C-> 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>>nil else nil)";
-by (cut_inst_tac [] move_is_move_sim 1);
-by (REPEAT (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [move_def,Let_def]) 1);
-qed"move_subprop4_sim";
-
-Goal
- "[|is_simulation R C A; reachable C s; s-a--C-> t; (s,s'):R|] ==>\
-\ let T' = @t'. ? ex1. (t,t'):R & move A ex1 s' a t' in \
-\ (t,T'):R";
-by (cut_inst_tac [] move_is_move_sim 1);
-by (REPEAT (assume_tac 1));
-by (asm_full_simp_tac (simpset() addsimps [move_def,Let_def]) 1);
-qed"move_subprop5_sim";
-
-(* ------------------------------------------------------------------ *)
-(* The following lemmata contribute to *)
-(* TRACE INCLUSION Part 1: Traces coincide *)
-(* ------------------------------------------------------------------ *)
-
-section "Lemmata for <==";
-
-
-(* ------------------------------------------------------
- Lemma 1 :Traces coincide
- ------------------------------------------------------- *)
-
-Delsplits[split_if];
-Goal
- "[|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')";
-
-by (pair_induct_tac "ex" [is_exec_frag_def] 1);
-(* cons case *)
-by (safe_tac set_cs);
-by (rename_tac "ex a t s s'" 1);
-by (asm_full_simp_tac (simpset() addsimps [mk_traceConc]) 1);
-by (forward_tac [reachable.reachable_n] 1);
-by (assume_tac 1);
-by (eres_inst_tac [("x","t")] allE 1);
-by (eres_inst_tac [("x",
- "@t'. ? ex1. (t,t'):R & move A ex1 s' a t'")]
- allE 1);
-by (Asm_full_simp_tac 1);
-by (asm_full_simp_tac (simpset() addsimps
- [rewrite_rule [Let_def] move_subprop5_sim,
- rewrite_rule [Let_def] move_subprop4_sim]
- addsplits [split_if]) 1);
-qed_spec_mp"traces_coincide_sim";
-Addsplits[split_if];
-
-
-(* ----------------------------------------------------------- *)
-(* Lemma 2 : corresp_ex_sim is execution *)
-(* ----------------------------------------------------------- *)
-
-
-Goal
- "[| 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')";
-
-by (Asm_full_simp_tac 1);
-by (pair_induct_tac "ex" [is_exec_frag_def] 1);
-(* main case *)
-by (safe_tac set_cs);
-by (rename_tac "ex a t s s'" 1);
-by (res_inst_tac [("t",
- "@t'. ? ex1. (t,t'):R & move A ex1 s' a t'")]
- lemma_2_1 1);
-
-(* Finite *)
-by (etac (rewrite_rule [Let_def] move_subprop2_sim) 1);
-by (REPEAT (atac 1));
-by (rtac conjI 1);
-
-(* is_exec_frag *)
-by (etac (rewrite_rule [Let_def] move_subprop1_sim) 1);
-by (REPEAT (atac 1));
-by (rtac conjI 1);
-
-(* Induction hypothesis *)
-(* reachable_n looping, therefore apply it manually *)
-by (eres_inst_tac [("x","t")] allE 1);
-by (eres_inst_tac [("x",
- "@t'. ? ex1. (t,t'):R & move A ex1 s' a t'")]
- allE 1);
-by (Asm_full_simp_tac 1);
-by (forward_tac [reachable.reachable_n] 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (simpset() addsimps
- [rewrite_rule [Let_def] move_subprop5_sim]) 1);
-(* laststate *)
-by (etac ((rewrite_rule [Let_def] move_subprop3_sim) RS sym) 1);
-by (REPEAT (atac 1));
-qed_spec_mp"correspsim_is_execution";
-
-
-(* -------------------------------------------------------------------------------- *)
-
-section "Main Theorem: T R A C E - I N C L U S I O N";
-
-(* -------------------------------------------------------------------------------- *)
-
- (* 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 *)
-
-Goal
-"[| 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";
- by (asm_full_simp_tac (simpset() addsimps [is_simulation_def,
- corresp_ex_sim_def, Int_non_empty,Image_def]) 1);
- by (REPEAT (etac conjE 1));
- by (etac ballE 1);
- by (Blast_tac 2);
- by (etac exE 1);
- by (rtac someI2 1);
- by (assume_tac 1);
- by (Blast_tac 1);
-qed"simulation_starts";
-
-bind_thm("sim_starts1",(rewrite_rule [Let_def] simulation_starts) RS conjunct1);
-bind_thm("sim_starts2",(rewrite_rule [Let_def] simulation_starts) RS conjunct2);
-
-
-Goalw [traces_def]
- "[| ext C = ext A; is_simulation R C A |] \
-\ ==> traces C <= traces A";
-
- by (simp_tac(simpset() addsimps [has_trace_def2])1);
- by (safe_tac set_cs);
-
- (* give execution of abstract automata *)
- by (res_inst_tac[("x","corresp_ex_sim A R ex")] bexI 1);
-
- (* Traces coincide, Lemma 1 *)
- by (pair_tac "ex" 1);
- by (rename_tac "s ex" 1);
- by (simp_tac (simpset() addsimps [corresp_ex_sim_def]) 1);
- by (res_inst_tac [("s","s")] traces_coincide_sim 1);
- by (REPEAT (atac 1));
- by (asm_full_simp_tac (simpset() addsimps [executions_def,
- reachable.reachable_0,sim_starts1]) 1);
-
- (* corresp_ex_sim is execution, Lemma 2 *)
- by (pair_tac "ex" 1);
- by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
- by (rename_tac "s ex" 1);
-
- (* start state *)
- by (rtac conjI 1);
- by (asm_full_simp_tac (simpset() addsimps [sim_starts2,
- corresp_ex_sim_def]) 1);
-
- (* is-execution-fragment *)
- by (asm_full_simp_tac (simpset() addsimps [corresp_ex_sim_def]) 1);
- by (res_inst_tac [("s","s")] correspsim_is_execution 1);
- by (assume_tac 1);
- by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0,sim_starts1]) 1);
-qed"trace_inclusion_for_simulations";
--- a/src/HOLCF/IOA/meta_theory/SimCorrectness.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/SimCorrectness.thy Sun May 28 19:54:20 2006 +0200
@@ -36,6 +36,266 @@
@@ ((h$xs) T'))
$x) )))"
-ML {* use_legacy_bindings (the_context ()) *}
+
+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 (rule 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)>>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 -a--C-> 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 (erule conjE)
+apply assumption
+done
+
+declare Let_def [simp]
+
+lemma move_subprop1_sim:
+ "[|is_simulation R C A; reachable C s; s-a--C-> 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-a--C-> 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-a--C-> 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-a--C-> 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>>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-a--C-> 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 {* TRACE INCLUSION Part 1: Traces coincide *}
+
+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 {* pair_induct_tac "ex" [thm "is_exec_frag_def"] 1 *})
+(* cons case *)
+apply (tactic "safe_tac set_cs")
+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
+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 {* pair_induct_tac "ex" [thm "is_exec_frag_def"] 1 *})
+(* main case *)
+apply (tactic "safe_tac set_cs")
+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, standard]
+lemmas sim_starts2 = simulation_starts [unfolded Let_def, THEN conjunct2, standard]
+
+
+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 (tactic "safe_tac set_cs")
+
+ (* give execution of abstract automata *)
+ apply (rule_tac x = "corresp_ex_sim A R ex" in bexI)
+
+ (* Traces coincide, Lemma 1 *)
+ apply (tactic {* pair_tac "ex" 1 *})
+ 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 {* pair_tac "ex" 1 *})
+ 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/HOLCF/IOA/meta_theory/Simulations.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,27 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/Simulations.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-Goal "(A~={}) = (? x. x:A)";
-by (safe_tac set_cs);
-by Auto_tac;
-qed"set_non_empty";
-
-Goal "(A Int B ~= {}) = (? x. x: A & x:B)";
-by (simp_tac (simpset() addsimps [set_non_empty]) 1);
-qed"Int_non_empty";
-
-
-Goalw [Image_def]
-"(R``{x} Int S ~= {}) = (? y. (x,y):R & y:S)";
-by (simp_tac (simpset() addsimps [Int_non_empty]) 1);
-qed"Sim_start_convert";
-
-Addsimps [Sim_start_convert];
-
-
-Goalw [is_ref_map_def,is_simulation_def]
-"!! f. is_ref_map f C A ==> is_simulation {p. (snd p) = f (fst p)} C A";
-by (Asm_full_simp_tac 1);
-qed"ref_map_is_simulation";
--- a/src/HOLCF/IOA/meta_theory/Simulations.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/Simulations.thy Sun May 28 19:54:20 2006 +0200
@@ -62,6 +62,30 @@
"is_prophecy_relation R C A == is_backward_simulation R C A &
is_ref_map (%x.(@y. (x,y):(R^-1))) A C"
-ML {* use_legacy_bindings (the_context ()) *}
+
+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/HOLCF/IOA/meta_theory/TL.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,165 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/TL.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-
-Goal "[] <> (.~ P) = (.~ <> [] P)";
-by (rtac ext 1);
-by (simp_tac (simpset() addsimps [Diamond_def,NOT_def,Box_def])1);
-by Auto_tac;
-qed"simple";
-
-Goal "nil |= [] P";
-by (asm_full_simp_tac (simpset() addsimps [satisfies_def,
- Box_def,tsuffix_def,suffix_def,nil_is_Conc])1);
-qed"Boxnil";
-
-Goal "~(nil |= <> P)";
-by (simp_tac (simpset() addsimps [Diamond_def,satisfies_def,NOT_def])1);
-by (cut_inst_tac [] Boxnil 1);
-by (asm_full_simp_tac (simpset() addsimps [satisfies_def])1);
-qed"Diamondnil";
-
-Goal "(<> F) s = (? s2. tsuffix s2 s & F s2)";
-by (simp_tac (simpset() addsimps [Diamond_def,NOT_def,Box_def])1);
-qed"Diamond_def2";
-
-
-
-section "TLA Axiomatization by Merz";
-
-(* ---------------------------------------------------------------- *)
-(* TLA Axiomatization by Merz *)
-(* ---------------------------------------------------------------- *)
-
-Goal "suffix s s";
-by (simp_tac (simpset() addsimps [suffix_def])1);
-by (res_inst_tac [("x","nil")] exI 1);
-by Auto_tac;
-qed"suffix_refl";
-
-Goal "s~=UU & s~=nil --> (s |= [] F .--> F)";
-by (simp_tac (simpset() addsimps [satisfies_def,IMPLIES_def,Box_def])1);
-by (REPEAT (rtac impI 1));
-by (eres_inst_tac [("x","s")] allE 1);
-by (asm_full_simp_tac (simpset() addsimps [tsuffix_def,suffix_refl])1);
-qed"reflT";
-
-
-Goal "[| suffix y x ; suffix z y |] ==> suffix z x";
-by (asm_full_simp_tac (simpset() addsimps [suffix_def])1);
-by Auto_tac;
-by (res_inst_tac [("x","s1 @@ s1a")] exI 1);
-by Auto_tac;
-by (simp_tac (simpset() addsimps [Conc_assoc]) 1);
-qed"suffix_trans";
-
-Goal "s |= [] F .--> [] [] F";
-by (simp_tac (simpset() addsimps [satisfies_def,IMPLIES_def,Box_def,tsuffix_def])1);
-by Auto_tac;
-by (dtac suffix_trans 1);
-by (assume_tac 1);
-by (eres_inst_tac [("x","s2a")] allE 1);
-by Auto_tac;
-qed"transT";
-
-
-Goal "s |= [] (F .--> G) .--> [] F .--> [] G";
-by (simp_tac (simpset() addsimps [satisfies_def,IMPLIES_def,Box_def])1);
-qed"normalT";
-
-
-section "TLA Rules by Lamport";
-
-(* ---------------------------------------------------------------- *)
-(* TLA Rules by Lamport *)
-(* ---------------------------------------------------------------- *)
-
-Goal "validT P ==> validT ([] P)";
-by (asm_full_simp_tac (simpset() addsimps [validT_def,satisfies_def,Box_def,tsuffix_def])1);
-qed"STL1a";
-
-Goal "valid P ==> validT (Init P)";
-by (asm_full_simp_tac (simpset() addsimps [valid_def,validT_def,satisfies_def,Init_def])1);
-qed"STL1b";
-
-Goal "valid P ==> validT ([] (Init P))";
-by (rtac STL1a 1);
-by (etac STL1b 1);
-qed"STL1";
-
-(* Note that unlift and HD is not at all used !!! *)
-Goal "valid (P .--> Q) ==> validT ([] (Init P) .--> [] (Init Q))";
-by (asm_full_simp_tac (simpset() addsimps [valid_def,validT_def,satisfies_def,IMPLIES_def,Box_def,Init_def])1);
-qed"STL4";
-
-
-section "LTL Axioms by Manna/Pnueli";
-
-(* ---------------------------------------------------------------- *)
-(* LTL Axioms by Manna/Pnueli *)
-(* ---------------------------------------------------------------- *)
-
-
-Goalw [tsuffix_def,suffix_def]
-"s~=UU & s~=nil --> tsuffix s2 (TL$s) --> tsuffix s2 s";
-by Auto_tac;
-by (Seq_case_simp_tac "s" 1);
-by (res_inst_tac [("x","a>>s1")] exI 1);
-by Auto_tac;
-qed_spec_mp"tsuffix_TL";
-
-val tsuffix_TL2 = conjI RS tsuffix_TL;
-
-Delsplits[split_if];
-Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def,AND_def,Box_def]
- "s~=UU & s~=nil --> (s |= [] F .--> (F .& (Next ([] F))))";
-by Auto_tac;
-(* []F .--> F *)
-by (eres_inst_tac [("x","s")] allE 1);
-by (asm_full_simp_tac (simpset() addsimps [tsuffix_def,suffix_refl])1);
-(* []F .--> Next [] F *)
-by (asm_full_simp_tac (simpset() addsplits [split_if]) 1);
-by Auto_tac;
-by (dtac tsuffix_TL2 1);
-by (REPEAT (atac 1));
-by Auto_tac;
-qed"LTL1";
-Addsplits[split_if];
-
-
-Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def]
- "s |= .~ (Next F) .--> (Next (.~ F))";
-by (Asm_full_simp_tac 1);
-qed"LTL2a";
-
-Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def]
- "s |= (Next (.~ F)) .--> (.~ (Next F))";
-by (Asm_full_simp_tac 1);
-qed"LTL2b";
-
-Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def]
-"ex |= (Next (F .--> G)) .--> (Next F) .--> (Next G)";
-by (Asm_full_simp_tac 1);
-qed"LTL3";
-
-Goalw [Next_def,satisfies_def,Box_def,NOT_def,IMPLIES_def]
- "s |= [] (F .--> Next F) .--> F .--> []F";
-by (Asm_full_simp_tac 1);
-by Auto_tac;
-
-by (asm_full_simp_tac (simpset() addsimps [tsuffix_def,suffix_def])1);
-by Auto_tac;
-
-
-
-
-Goal "[| validT (P .--> Q); validT P |] ==> validT Q";
-by (asm_full_simp_tac (simpset() addsimps [validT_def,satisfies_def,IMPLIES_def])1);
-qed"ModusPonens";
-
-(* works only if validT is defined without restriction to s~=UU, s~=nil *)
-Goal "validT P ==> validT (Next P)";
-by (asm_full_simp_tac (simpset() addsimps [validT_def,satisfies_def,Next_def])1);
-(* qed"NextTauto"; *)
--- a/src/HOLCF/IOA/meta_theory/TL.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/TL.thy Sun May 28 19:54:20 2006 +0200
@@ -67,6 +67,140 @@
validT_def:
"validT P == ! s. s~=UU & s~=nil --> (s |= P)"
-ML {* use_legacy_bindings (the_context ()) *}
+
+lemma simple: "[] <> (.~ P) = (.~ <> [] P)"
+apply (rule ext)
+apply (simp add: Diamond_def NOT_def Box_def)
+done
+
+lemma Boxnil: "nil |= [] P"
+apply (simp add: satisfies_def Box_def tsuffix_def suffix_def nil_is_Conc)
+done
+
+lemma Diamondnil: "~(nil |= <> P)"
+apply (simp add: Diamond_def satisfies_def NOT_def)
+apply (cut_tac Boxnil)
+apply (simp add: satisfies_def)
+done
+
+lemma Diamond_def2: "(<> 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 |= [] F .--> 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 |= [] F .--> [] [] 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 |= [] (F .--> G) .--> [] F .--> [] G"
+apply (simp (no_asm) add: satisfies_def IMPLIES_def Box_def)
+done
+
+
+subsection "TLA Rules by Lamport"
+
+lemma STL1a: "validT P ==> validT ([] 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 ([] (Init P))"
+apply (rule STL1a)
+apply (erule STL1b)
+done
+
+(* Note that unlift and HD is not at all used !!! *)
+lemma STL4: "valid (P .--> Q) ==> validT ([] (Init P) .--> [] (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 {* Seq_case_simp_tac "s" 1 *})
+apply (rule_tac x = "a>>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 |= [] F .--> (F .& (Next ([] F))))"
+apply (unfold Next_def satisfies_def NOT_def IMPLIES_def AND_def Box_def)
+apply auto
+(* []F .--> F *)
+apply (erule_tac x = "s" in allE)
+apply (simp add: tsuffix_def suffix_refl)
+(* []F .--> Next [] 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 |= .~ (Next F) .--> (Next (.~ F))"
+apply (unfold Next_def satisfies_def NOT_def IMPLIES_def)
+apply simp
+done
+
+lemma LTL2b:
+ "s |= (Next (.~ F)) .--> (.~ (Next F))"
+apply (unfold Next_def satisfies_def NOT_def IMPLIES_def)
+apply simp
+done
+
+lemma LTL3:
+"ex |= (Next (F .--> G)) .--> (Next F) .--> (Next G)"
+apply (unfold Next_def satisfies_def NOT_def IMPLIES_def)
+apply simp
+done
+
+
+lemma ModusPonens: "[| validT (P .--> Q); validT P |] ==> validT Q"
+apply (simp add: validT_def satisfies_def IMPLIES_def)
+done
end
--- a/src/HOLCF/IOA/meta_theory/TLS.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,122 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/TLS.ML
- ID: $Id$
- Author: Olaf Mueller
-*)
-
-(* global changes to simpset() and claset(), repeated from Traces.ML *)
-Delsimps (ex_simps @ all_simps);
-Delsimps [split_paired_Ex];
-Addsimps [Let_def];
-change_claset (fn cs => cs delSWrapper "split_all_tac");
-
-
-(* ---------------------------------------------------------------- *)
-(* ex2seqC *)
-(* ---------------------------------------------------------------- *)
-
-Goal "ex2seqC = (LAM ex. (%s. case ex of \
-\ nil => (s,None,s)>>nil \
-\ | x##xs => (flift1 (%pr. \
-\ (s,Some (fst pr), snd pr)>> (ex2seqC$xs) (snd pr)) \
-\ $x) \
-\ ))";
-by (rtac trans 1);
-by (rtac fix_eq2 1);
-by (rtac ex2seqC_def 1);
-by (rtac beta_cfun 1);
-by (simp_tac (simpset() addsimps [flift1_def]) 1);
-qed"ex2seqC_unfold";
-
-Goal "(ex2seqC $UU) s=UU";
-by (stac ex2seqC_unfold 1);
-by (Simp_tac 1);
-qed"ex2seqC_UU";
-
-Goal "(ex2seqC $nil) s = (s,None,s)>>nil";
-by (stac ex2seqC_unfold 1);
-by (Simp_tac 1);
-qed"ex2seqC_nil";
-
-Goal "(ex2seqC $((a,t)>>xs)) s = \
-\ (s,Some a,t)>> ((ex2seqC$xs) t)";
-by (rtac trans 1);
-by (stac ex2seqC_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,flift1_def]) 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,flift1_def]) 1);
-qed"ex2seqC_cons";
-
-Addsimps [ex2seqC_UU,ex2seqC_nil,ex2seqC_cons];
-
-
-
-Addsimps [mkfin_UU,mkfin_nil,mkfin_cons];
-
-Goal "ex2seq (s, UU) = (s,None,s)>>nil";
-by (simp_tac (simpset() addsimps [ex2seq_def]) 1);
-qed"ex2seq_UU";
-
-Goal "ex2seq (s, nil) = (s,None,s)>>nil";
-by (simp_tac (simpset() addsimps [ex2seq_def]) 1);
-qed"ex2seq_nil";
-
-Goal "ex2seq (s, (a,t)>>ex) = (s,Some a,t) >> ex2seq (t, ex)";
-by (simp_tac (simpset() addsimps [ex2seq_def]) 1);
-qed"ex2seq_cons";
-
-Delsimps [ex2seqC_UU,ex2seqC_nil,ex2seqC_cons];
-Addsimps [ex2seq_UU,ex2seq_nil, ex2seq_cons];
-
-
-Goal "ex2seq exec ~= UU & ex2seq exec ~= nil";
-by (pair_tac "exec" 1);
-by (Seq_case_simp_tac "y" 1);
-by (pair_tac "a" 1);
-qed"ex2seq_nUUnnil";
-
-
-(* ----------------------------------------------------------- *)
-(* Interface TL -- TLS *)
-(* ---------------------------------------------------------- *)
-
-
-(* uses the fact that in executions states overlap, which is lost in
- after the translation via ex2seq !! *)
-
-Goalw [Init_def,Next_def,temp_sat_def,satisfies_def,IMPLIES_def,AND_def]
- "[| ! s a t. (P s) & s-a--A-> t --> (Q t) |]\
-\ ==> ex |== (Init (%(s,a,t). P s) .& Init (%(s,a,t). s -a--A-> t) \
-\ .--> (Next (Init (%(s,a,t).Q s))))";
-
-by (clarify_tac set_cs 1);
-by (asm_full_simp_tac (simpset() addsplits [split_if]) 1);
-(* TL = UU *)
-by (rtac conjI 1);
-by (pair_tac "ex" 1);
-by (Seq_case_simp_tac "y" 1);
-by (pair_tac "a" 1);
-by (Seq_case_simp_tac "s" 1);
-by (pair_tac "a" 1);
-(* TL = nil *)
-by (rtac conjI 1);
-by (pair_tac "ex" 1);
-by (Seq_case_simp_tac "y" 1);
-by (asm_full_simp_tac (simpset() addsimps [unlift_def])1);
-by (Fast_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [unlift_def])1);
-by (Fast_tac 1);
-by (asm_full_simp_tac (simpset() addsimps [unlift_def])1);
-by (pair_tac "a" 1);
-by (Seq_case_simp_tac "s" 1);
-by (pair_tac "a" 1);
-(* TL =cons *)
-by (asm_full_simp_tac (simpset() addsimps [unlift_def])1);
-
-by (pair_tac "ex" 1);
-by (Seq_case_simp_tac "y" 1);
-by (pair_tac "a" 1);
-by (Seq_case_simp_tac "s" 1);
- by (Fast_tac 1);
- by (Fast_tac 1);
-by (pair_tac "a" 1);
- by (Fast_tac 1);
-qed"TL_TLS";
--- a/src/HOLCF/IOA/meta_theory/TLS.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/TLS.thy Sun May 28 19:54:20 2006 +0200
@@ -86,6 +86,118 @@
mkfin_cons:
"(mkfin (a>>s)) = (a>>(mkfin s))"
-ML {* use_legacy_bindings (the_context ()) *}
+
+lemmas [simp del] = ex_simps all_simps split_paired_Ex
+declare Let_def [simp]
+
+ML_setup {* change_claset (fn cs => cs delSWrapper "split_all_tac") *}
+
+
+subsection {* ex2seqC *}
+
+lemma ex2seqC_unfold: "ex2seqC = (LAM ex. (%s. case ex of
+ nil => (s,None,s)>>nil
+ | x##xs => (flift1 (%pr.
+ (s,Some (fst pr), snd pr)>> (ex2seqC$xs) (snd pr))
+ $x)
+ ))"
+apply (rule trans)
+apply (rule fix_eq2)
+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)>>nil"
+apply (subst ex2seqC_unfold)
+apply simp
+done
+
+lemma ex2seqC_cons: "(ex2seqC $((a,t)>>xs)) s =
+ (s,Some a,t)>> ((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)>>nil"
+apply (simp add: ex2seq_def)
+done
+
+lemma ex2seq_nil: "ex2seq (s, nil) = (s,None,s)>>nil"
+apply (simp add: ex2seq_def)
+done
+
+lemma ex2seq_cons: "ex2seq (s, (a,t)>>ex) = (s,Some a,t) >> ex2seq (t, ex)"
+apply (simp add: ex2seq_def)
+done
+
+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 {* pair_tac "exec" 1 *})
+apply (tactic {* Seq_case_simp_tac "y" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+done
+
+
+subsection {* Interface TL -- TLS *}
+
+(* 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-a--A-> t --> (Q t) |]
+ ==> ex |== (Init (%(s,a,t). P s) .& Init (%(s,a,t). s -a--A-> t)
+ .--> (Next (Init (%(s,a,t).Q s))))"
+apply (unfold Init_def Next_def temp_sat_def satisfies_def IMPLIES_def AND_def)
+
+apply (tactic "clarify_tac set_cs 1")
+apply (simp split add: split_if)
+(* TL = UU *)
+apply (rule conjI)
+apply (tactic {* pair_tac "ex" 1 *})
+apply (tactic {* Seq_case_simp_tac "y" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+apply (tactic {* Seq_case_simp_tac "s" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+(* TL = nil *)
+apply (rule conjI)
+apply (tactic {* pair_tac "ex" 1 *})
+apply (tactic {* Seq_case_tac "y" 1 *})
+apply (simp add: unlift_def)
+apply fast
+apply (simp add: unlift_def)
+apply fast
+apply (simp add: unlift_def)
+apply (tactic {* pair_tac "a" 1 *})
+apply (tactic {* Seq_case_simp_tac "s" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+(* TL =cons *)
+apply (simp add: unlift_def)
+
+apply (tactic {* pair_tac "ex" 1 *})
+apply (tactic {* Seq_case_simp_tac "y" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+apply (tactic {* Seq_case_simp_tac "s" 1 *})
+apply blast
+apply fastsimp
+apply (tactic {* pair_tac "a" 1 *})
+ apply fastsimp
+done
end
--- a/src/HOLCF/IOA/meta_theory/Traces.ML Sat May 27 21:18:51 2006 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,262 +0,0 @@
-(* Title: HOLCF/IOA/meta_theory/Traces.ML
- ID: $Id$
- Author: Olaf Mueller
-
-Theorems about Executions and Traces of I/O automata in HOLCF.
-*)
-
-(* global changes to simpset() and claset(), see also TLS.ML *)
-Delsimps (ex_simps @ all_simps);
-Delsimps [split_paired_Ex];
-Addsimps [Let_def];
-change_claset (fn cs => cs delSWrapper "split_all_tac");
-
-val exec_rws = [executions_def,is_exec_frag_def];
-
-
-
-(* ----------------------------------------------------------------------------------- *)
-
-section "recursive equations of operators";
-
-
-(* ---------------------------------------------------------------- *)
-(* filter_act *)
-(* ---------------------------------------------------------------- *)
-
-
-Goal "filter_act$UU = UU";
-by (simp_tac (simpset() addsimps [filter_act_def]) 1);
-qed"filter_act_UU";
-
-Goal "filter_act$nil = nil";
-by (simp_tac (simpset() addsimps [filter_act_def]) 1);
-qed"filter_act_nil";
-
-Goal "filter_act$(x>>xs) = (fst x) >> filter_act$xs";
-by (simp_tac (simpset() addsimps [filter_act_def]) 1);
-qed"filter_act_cons";
-
-Addsimps [filter_act_UU,filter_act_nil,filter_act_cons];
-
-
-(* ---------------------------------------------------------------- *)
-(* mk_trace *)
-(* ---------------------------------------------------------------- *)
-
-Goal "mk_trace A$UU=UU";
-by (simp_tac (simpset() addsimps [mk_trace_def]) 1);
-qed"mk_trace_UU";
-
-Goal "mk_trace A$nil=nil";
-by (simp_tac (simpset() addsimps [mk_trace_def]) 1);
-qed"mk_trace_nil";
-
-Goal "mk_trace A$(at >> xs) = \
-\ (if ((fst at):ext A) \
-\ then (fst at) >> (mk_trace A$xs) \
-\ else mk_trace A$xs)";
-
-by (asm_full_simp_tac (simpset() addsimps [mk_trace_def]) 1);
-qed"mk_trace_cons";
-
-Addsimps [mk_trace_UU,mk_trace_nil,mk_trace_cons];
-
-(* ---------------------------------------------------------------- *)
-(* is_exec_fragC *)
-(* ---------------------------------------------------------------- *)
-
-
-Goal "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) \
-\ ))";
-by (rtac trans 1);
-by (rtac fix_eq2 1);
-by (rtac is_exec_fragC_def 1);
-by (rtac beta_cfun 1);
-by (simp_tac (simpset() addsimps [flift1_def]) 1);
-qed"is_exec_fragC_unfold";
-
-Goal "(is_exec_fragC A$UU) s=UU";
-by (stac is_exec_fragC_unfold 1);
-by (Simp_tac 1);
-qed"is_exec_fragC_UU";
-
-Goal "(is_exec_fragC A$nil) s = TT";
-by (stac is_exec_fragC_unfold 1);
-by (Simp_tac 1);
-qed"is_exec_fragC_nil";
-
-Goal "(is_exec_fragC A$(pr>>xs)) s = \
-\ (Def ((s,pr):trans_of A) \
-\ andalso (is_exec_fragC A$xs)(snd pr))";
-by (rtac trans 1);
-by (stac is_exec_fragC_unfold 1);
-by (asm_full_simp_tac (simpset() addsimps [Consq_def,flift1_def]) 1);
-by (Simp_tac 1);
-qed"is_exec_fragC_cons";
-
-
-Addsimps [is_exec_fragC_UU,is_exec_fragC_nil,is_exec_fragC_cons];
-
-
-(* ---------------------------------------------------------------- *)
-(* is_exec_frag *)
-(* ---------------------------------------------------------------- *)
-
-Goal "is_exec_frag A (s, UU)";
-by (simp_tac (simpset() addsimps [is_exec_frag_def]) 1);
-qed"is_exec_frag_UU";
-
-Goal "is_exec_frag A (s, nil)";
-by (simp_tac (simpset() addsimps [is_exec_frag_def]) 1);
-qed"is_exec_frag_nil";
-
-Goal "is_exec_frag A (s, (a,t)>>ex) = \
-\ (((s,a,t):trans_of A) & \
-\ is_exec_frag A (t, ex))";
-by (simp_tac (simpset() addsimps [is_exec_frag_def]) 1);
-qed"is_exec_frag_cons";
-
-
-(* Delsimps [is_exec_fragC_UU,is_exec_fragC_nil,is_exec_fragC_cons]; *)
-Addsimps [is_exec_frag_UU,is_exec_frag_nil, is_exec_frag_cons];
-
-(* ---------------------------------------------------------------------------- *)
- section "laststate";
-(* ---------------------------------------------------------------------------- *)
-
-Goal "laststate (s,UU) = s";
-by (simp_tac (simpset() addsimps [laststate_def]) 1);
-qed"laststate_UU";
-
-Goal "laststate (s,nil) = s";
-by (simp_tac (simpset() addsimps [laststate_def]) 1);
-qed"laststate_nil";
-
-Goal "!! ex. Finite ex ==> laststate (s,at>>ex) = laststate (snd at,ex)";
-by (simp_tac (simpset() addsimps [laststate_def]) 1);
-by (case_tac "ex=nil" 1);
-by (Asm_simp_tac 1);
-by (Asm_simp_tac 1);
-by (dtac (Finite_Last1 RS mp) 1);
-by (assume_tac 1);
-by (def_tac 1);
-qed"laststate_cons";
-
-Addsimps [laststate_UU,laststate_nil,laststate_cons];
-
-Goal "!!ex. Finite ex ==> (! s. ? u. laststate (s,ex)=u)";
-by (Seq_Finite_induct_tac 1);
-qed"exists_laststate";
-
-
-(* -------------------------------------------------------------------------------- *)
-
-section "has_trace, mk_trace";
-
-(* alternative definition of has_trace tailored for the refinement proof, as it does not
- take the detour of schedules *)
-
-Goalw [executions_def,mk_trace_def,has_trace_def,schedules_def,has_schedule_def]
-"has_trace A b = (? ex:executions A. b = mk_trace A$(snd ex))";
-
-by (safe_tac set_cs);
-(* 1 *)
-by (res_inst_tac[("x","ex")] bexI 1);
-by (stac beta_cfun 1);
-by (cont_tacR 1);
-by (Simp_tac 1);
-by (Asm_simp_tac 1);
-(* 2 *)
-by (res_inst_tac[("x","filter_act$(snd ex)")] bexI 1);
-by (stac beta_cfun 1);
-by (cont_tacR 1);
-by (Simp_tac 1);
-by (safe_tac set_cs);
-by (res_inst_tac[("x","ex")] bexI 1);
-by (REPEAT (Asm_simp_tac 1));
-qed"has_trace_def2";
-
-
-(* -------------------------------------------------------------------------------- *)
-
-section "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) *)
-
-Goal
- "!! A. is_trans_of A ==> \
-\ ! s. is_exec_frag A (s,xs) --> Forall (%a. a:act A) (filter_act$xs)";
-
-by (pair_induct_tac "xs" [is_exec_frag_def,Forall_def,sforall_def] 1);
-(* main case *)
-by (rename_tac "ss a t" 1);
-by (safe_tac set_cs);
-by (REPEAT (asm_full_simp_tac (simpset() addsimps [is_trans_of_def]) 1));
-qed"execfrag_in_sig";
-
-Goal
- "!! A.[| is_trans_of A; x:executions A |] ==> \
-\ Forall (%a. a:act A) (filter_act$(snd x))";
-
-by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
-by (pair_tac "x" 1);
-by (rtac (execfrag_in_sig RS spec RS mp) 1);
-by Auto_tac;
-qed"exec_in_sig";
-
-Goalw [schedules_def,has_schedule_def]
- "!! A.[| is_trans_of A; x:schedules A |] ==> \
-\ Forall (%a. a:act A) x";
-
-by (fast_tac (claset() addSIs [exec_in_sig]) 1);
-qed"scheds_in_sig";
-
-(*
-
-is ok but needs ForallQFilterP which has to been proven first (is trivial also)
-
-Goalw [traces_def,has_trace_def]
- "!! A.[| x:traces A |] ==> \
-\ Forall (%a. a:act A) x";
- by (safe_tac set_cs );
-by (rtac ForallQFilterP 1);
-by (fast_tac (!claset addSIs [ext_is_act]) 1);
-qed"traces_in_sig";
-*)
-
-
-(* -------------------------------------------------------------------------------- *)
-
-section "executions are prefix closed";
-
-(* only admissible in y, not if done in x !! *)
-Goal "!x s. is_exec_frag A (s,x) & y<<x --> is_exec_frag A (s,y)";
-by (pair_induct_tac "y" [is_exec_frag_def] 1);
-by (strip_tac 1);
-by (Seq_case_simp_tac "xa" 1);
-by (pair_tac "a" 1);
-by Auto_tac;
-qed"execfrag_prefixclosed";
-
-bind_thm ("exec_prefixclosed",conjI RS (execfrag_prefixclosed RS spec RS spec RS mp));
-
-
-(* second prefix notion for Finite x *)
-
-Goal "! y s. is_exec_frag A (s,x@@y) --> is_exec_frag A (s,x)";
-by (pair_induct_tac "x" [is_exec_frag_def] 1);
-by (strip_tac 1);
-by (Seq_case_simp_tac "s" 1);
-by (pair_tac "a" 1);
-by Auto_tac;
-qed_spec_mp"exec_prefix2closed";
-
--- a/src/HOLCF/IOA/meta_theory/Traces.thy Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IOA/meta_theory/Traces.thy Sun May 28 19:54:20 2006 +0200
@@ -192,6 +192,237 @@
Traces_def:
"Traces A == (traces A,asig_of A)"
-ML {* use_legacy_bindings (the_context ()) *}
+
+lemmas [simp del] = ex_simps all_simps split_paired_Ex
+declare Let_def [simp]
+ML_setup {* change_claset (fn cs => cs delSWrapper "split_all_tac") *}
+
+lemmas exec_rws = executions_def is_exec_frag_def
+
+
+
+subsection "recursive equations of operators"
+
+(* ---------------------------------------------------------------- *)
+(* filter_act *)
+(* ---------------------------------------------------------------- *)
+
+
+lemma filter_act_UU: "filter_act$UU = UU"
+apply (simp add: filter_act_def)
+done
+
+lemma filter_act_nil: "filter_act$nil = nil"
+apply (simp add: filter_act_def)
+done
+
+lemma filter_act_cons: "filter_act$(x>>xs) = (fst x) >> filter_act$xs"
+apply (simp add: filter_act_def)
+done
+
+declare filter_act_UU [simp] filter_act_nil [simp] filter_act_cons [simp]
+
+
+(* ---------------------------------------------------------------- *)
+(* mk_trace *)
+(* ---------------------------------------------------------------- *)
+
+lemma mk_trace_UU: "mk_trace A$UU=UU"
+apply (simp add: mk_trace_def)
+done
+
+lemma mk_trace_nil: "mk_trace A$nil=nil"
+apply (simp add: mk_trace_def)
+done
+
+lemma mk_trace_cons: "mk_trace A$(at >> xs) =
+ (if ((fst at):ext A)
+ then (fst at) >> (mk_trace A$xs)
+ else mk_trace A$xs)"
+
+apply (simp add: mk_trace_def)
+done
+
+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_eq2)
+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>>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)"
+apply (simp add: is_exec_frag_def)
+done
+
+lemma is_exec_frag_nil: "is_exec_frag A (s, nil)"
+apply (simp add: is_exec_frag_def)
+done
+
+lemma is_exec_frag_cons: "is_exec_frag A (s, (a,t)>>ex) =
+ (((s,a,t):trans_of A) &
+ is_exec_frag A (t, ex))"
+apply (simp add: is_exec_frag_def)
+done
+
+
+(* 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"
+apply (simp add: laststate_def)
+done
+
+lemma laststate_nil: "laststate (s,nil) = s"
+apply (simp add: laststate_def)
+done
+
+lemma laststate_cons: "!! ex. Finite ex ==> laststate (s,at>>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 (tactic "def_tac 1")
+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 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)
+apply (tactic "safe_tac set_cs")
+(* 1 *)
+apply (rule_tac x = "ex" in bexI)
+apply (simplesubst beta_cfun)
+apply (tactic "cont_tacR 1")
+apply (simp (no_asm))
+apply (simp (no_asm_simp))
+(* 2 *)
+apply (rule_tac x = "filter_act$ (snd ex) " in bexI)
+apply (simplesubst beta_cfun)
+apply (tactic "cont_tacR 1")
+apply (simp (no_asm))
+apply (tactic "safe_tac set_cs")
+apply (rule_tac x = "ex" in bexI)
+apply simp_all
+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 {* pair_induct_tac "xs" [thm "is_exec_frag_def",
+ thm "Forall_def", thm "sforall_def"] 1 *})
+(* main case *)
+apply (rename_tac ss a t)
+apply (tactic "safe_tac set_cs")
+apply (simp_all 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 {* pair_tac "x" 1 *})
+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)
+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 {* pair_induct_tac "y" [thm "is_exec_frag_def"] 1 *})
+apply (intro strip)
+apply (tactic {* Seq_case_simp_tac "xa" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+apply auto
+done
+
+lemmas exec_prefixclosed =
+ conjI [THEN execfrag_prefixclosed [THEN spec, THEN spec, THEN mp], standard]
+
+
+(* 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 {* pair_induct_tac "x" [thm "is_exec_frag_def"] 1 *})
+apply (intro strip)
+apply (tactic {* Seq_case_simp_tac "s" 1 *})
+apply (tactic {* pair_tac "a" 1 *})
+apply auto
+done
+
end
--- a/src/HOLCF/IsaMakefile Sat May 27 21:18:51 2006 +0200
+++ b/src/HOLCF/IsaMakefile Sun May 28 19:54:20 2006 +0200
@@ -75,27 +75,17 @@
IOA: HOLCF $(OUT)/IOA
-$(OUT)/IOA: $(OUT)/HOLCF IOA/ROOT.ML IOA/meta_theory/Traces.thy \
- IOA/meta_theory/Asig.ML IOA/meta_theory/Asig.thy \
- IOA/meta_theory/CompoScheds.thy IOA/meta_theory/CompoExecs.ML \
- IOA/meta_theory/CompoTraces.thy IOA/meta_theory/CompoScheds.ML \
- IOA/meta_theory/CompoTraces.ML IOA/meta_theory/Sequence.ML \
- IOA/meta_theory/Seq.thy IOA/meta_theory/RefCorrectness.thy \
- IOA/meta_theory/Automata.thy IOA/meta_theory/Traces.ML \
- IOA/meta_theory/RefMappings.ML \
- IOA/meta_theory/ShortExecutions.thy IOA/meta_theory/ShortExecutions.ML \
- IOA/meta_theory/IOA.thy \
- IOA/meta_theory/Sequence.thy IOA/meta_theory/Automata.ML \
- IOA/meta_theory/CompoExecs.thy IOA/meta_theory/RefMappings.thy \
- IOA/meta_theory/RefCorrectness.ML IOA/meta_theory/Compositionality.ML \
- IOA/meta_theory/Compositionality.thy \
- IOA/meta_theory/TL.thy IOA/meta_theory/TL.ML IOA/meta_theory/TLS.thy \
- IOA/meta_theory/TLS.ML IOA/meta_theory/LiveIOA.thy IOA/meta_theory/LiveIOA.ML \
- IOA/meta_theory/Pred.thy IOA/meta_theory/Abstraction.thy \
- IOA/meta_theory/Abstraction.ML \
- IOA/meta_theory/Simulations.thy IOA/meta_theory/Simulations.ML \
- IOA/meta_theory/SimCorrectness.thy IOA/meta_theory/SimCorrectness.ML \
- IOA/meta_theory/ioa_package.ML
+$(OUT)/IOA: $(OUT)/HOLCF IOA/ROOT.ML IOA/meta_theory/Traces.thy \
+ IOA/meta_theory/Asig.thy IOA/meta_theory/CompoScheds.thy \
+ IOA/meta_theory/CompoTraces.thy IOA/meta_theory/Seq.thy \
+ IOA/meta_theory/RefCorrectness.thy IOA/meta_theory/Automata.thy \
+ IOA/meta_theory/ShortExecutions.thy IOA/meta_theory/IOA.thy \
+ IOA/meta_theory/Sequence.thy IOA/meta_theory/CompoExecs.thy \
+ IOA/meta_theory/RefMappings.thy IOA/meta_theory/Compositionality.thy \
+ IOA/meta_theory/TL.thy IOA/meta_theory/TLS.thy \
+ IOA/meta_theory/LiveIOA.thy IOA/meta_theory/Pred.thy \
+ IOA/meta_theory/Abstraction.thy IOA/meta_theory/Simulations.thy \
+ IOA/meta_theory/SimCorrectness.thy IOA/meta_theory/ioa_package.ML
@cd IOA; $(ISATOOL) usedir -b $(OUT)/HOLCF IOA