diff -r 94f018b2a4fb -r 78d95ff11ade src/HOL/HOLCF/IOA/TL.thy --- a/src/HOL/HOLCF/IOA/TL.thy Sun Jan 05 21:17:36 2025 +0100 +++ b/src/HOL/HOLCF/IOA/TL.thy Sun Jan 05 22:28:05 2025 +0100 @@ -57,34 +57,53 @@ 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 +proof - + have "Finite nil \ s = nil @@ s" by simp + then show ?thesis unfolding suffix_def .. +qed + +lemma suffix_trans: + assumes "suffix y x" + and "suffix z y" + shows "suffix z x" +proof - + from assms obtain s1 s2 + where "Finite s1 \ x = s1 @@ y" + and "Finite s2 \ y = s2 @@ z" + unfolding suffix_def by iprover + then have "Finite (s1 @@ s2) \ x = (s1 @@ s2) @@ z" + by (simp add: Conc_assoc) + then show ?thesis unfolding suffix_def .. +qed lemma reflT: "s \ UU \ s \ nil \ (s \ \F \<^bold>\ 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 add: Conc_assoc) - done +proof + assume s: "s \ UU \ s \ nil" + have "F s" if box_F: "\s2. tsuffix s2 s \ F s2" + proof - + from s and suffix_refl [of s] have "tsuffix s s" + by (simp add: tsuffix_def) + also from box_F have "?this \ F s" .. + finally show ?thesis . + qed + then show "s \ \F \<^bold>\ F" + by (simp add: satisfies_def IMPLIES_def Box_def) +qed lemma transT: "s \ \F \<^bold>\ \\F" - apply (simp 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 +proof - + have "F s2" if *: "tsuffix s1 s" "tsuffix s2 s1" + and **: "\s'. tsuffix s' s \ F s'" + for s1 s2 + proof - + from * have "tsuffix s2 s" + by (auto simp: tsuffix_def elim: suffix_trans) + also from ** have "?this \ F s2" .. + finally show ?thesis . + qed + then show ?thesis + by (simp add: satisfies_def IMPLIES_def Box_def) +qed lemma normalT: "s \ \(F \<^bold>\ G) \<^bold>\ \F \<^bold>\ \G" by (simp add: satisfies_def IMPLIES_def Box_def) @@ -98,10 +117,11 @@ lemma STL1b: "\ P \ \<^bold>\ (Init P)" by (simp add: valid_def validT_def satisfies_def Init_def) -lemma STL1: "\ P \ \<^bold>\ (\(Init P))" - apply (rule STL1a) - apply (erule STL1b) - done +lemma STL1: assumes "\ P" shows "\<^bold>\ (\(Init P))" +proof - + from assms have "\<^bold>\ (Init P)" by (rule STL1b) + then show ?thesis by (rule STL1a) +qed (*Note that unlift and HD is not at all used!*) lemma STL4: "\ (P \<^bold>\ Q) \ \<^bold>\ (\(Init P) \<^bold>\ \(Init Q))"