diff -r bdaa0eb0fc74 -r 0826f6b6ba09 src/HOL/HOLCF/IOA/Abstraction.thy --- a/src/HOL/HOLCF/IOA/Abstraction.thy Sat Jan 16 23:24:50 2016 +0100 +++ b/src/HOL/HOLCF/IOA/Abstraction.thy Sat Jan 16 23:31:28 2016 +0100 @@ -114,21 +114,6 @@ done -(* FIXME: to TLS.thy *) - -lemma IMPLIES_temp_sat [simp]: "(ex \ P \<^bold>\ Q) \ ((ex \ P) \ (ex \ Q))" - by (simp add: IMPLIES_def temp_sat_def satisfies_def) - -lemma AND_temp_sat [simp]: "(ex \ P \<^bold>\ Q) \ ((ex \ P) \ (ex \ Q))" - by (simp add: AND_def temp_sat_def satisfies_def) - -lemma OR_temp_sat [simp]: "(ex \ P \<^bold>\ Q) \ ((ex \ P) \ (ex \ Q))" - by (simp add: OR_def temp_sat_def satisfies_def) - -lemma NOT_temp_sat [simp]: "(ex \ \<^bold>\ P) \ (\ (ex \ P))" - by (simp add: NOT_def temp_sat_def satisfies_def) - - lemma AbsRuleT2: "is_live_abstraction h (C, L) (A, M) \ validLIOA (A, M) Q \ temp_strengthening Q P h \ validLIOA (C, L) P" @@ -217,11 +202,6 @@ apply (tactic \pair_tac @{context} "ex" 1\) done -(* FIXME: to Traces.thy *) - -lemma implements_trans: "A =<| B \ B =<| C \ A =<| C" - by (auto simp add: ioa_implements_def) - subsection \Abstraction Rules for Automata\ @@ -328,15 +308,6 @@ done -(* FIXME: to Sequence.thy *) - -lemma Mapnil: "Map f $ s = nil \ s = nil" - by (tactic \Seq_case_simp_tac @{context} "s" 1\) - -lemma MapUU: "Map f $ s = UU \ s = UU" - by (tactic \Seq_case_simp_tac @{context} "s" 1\) - - (*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)" @@ -390,10 +361,6 @@ apply (tactic \pair_tac @{context} "a" 1\) done -(* FIXME: put to Sequence Lemmas *) -lemma MapTL: "Map f $ (TL $ s) = TL $ (Map f $ s)" - by (tactic \Seq_induct_tac @{context} "s" [] 1\) - (*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)"