# HG changeset patch # User traytel # Date 1443090079 -7200 # Node ID dd949db0ade8e043fee067f87b69a725c0f4e58f # Parent e3d8a313a64916bf4b5a348e012899e1241f96ab tuned proofs (less warnings) diff -r e3d8a313a649 -r dd949db0ade8 src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy --- a/src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy Wed Sep 23 14:11:35 2015 +0100 +++ b/src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy Thu Sep 24 12:21:19 2015 +0200 @@ -22,8 +22,8 @@ lemma shift_prefix_cases: assumes "xl @- xs = yl @- ys" shows "prefixeq xl yl \ prefixeq yl xl" -using shift_prefix[OF assms] apply(cases "length xl \ length yl") -by (metis, metis assms nat_le_linear shift_prefix) +using shift_prefix[OF assms] +by (cases "length xl \ length yl") (metis, metis assms nat_le_linear shift_prefix) section\Linear temporal logic\ @@ -111,12 +111,12 @@ lemma ev_mono: assumes ev: "ev \ xs" and 0: "\ xs. \ xs \ \ xs" shows "ev \ xs" -using ev by induct (auto intro: ev.intros simp: 0) +using ev by induct (auto simp: 0) lemma alw_mono: assumes alw: "alw \ xs" and 0: "\ xs. \ xs \ \ xs" shows "alw \ xs" -using alw by coinduct (auto elim: alw.cases simp: 0) +using alw by coinduct (auto simp: 0) lemma until_monoL: assumes until: "(\1 until \) xs" and 0: "\ xs. \1 xs \ \2 xs" @@ -137,55 +137,55 @@ lemma until_false: "\ until false = alw \" proof- {fix xs assume "(\ until false) xs" hence "alw \ xs" - apply coinduct by (auto elim: UNTIL.cases) + by coinduct (auto elim: UNTIL.cases) } moreover {fix xs assume "alw \ xs" hence "(\ until false) xs" - apply coinduct by (auto elim: alw.cases) + by coinduct auto } ultimately show ?thesis by blast qed lemma ev_nxt: "ev \ = (\ or nxt (ev \))" -apply(rule ext) by (metis ev.simps nxt.simps) +by (rule ext) (metis ev.simps nxt.simps) lemma alw_nxt: "alw \ = (\ aand nxt (alw \))" -apply(rule ext) by (metis alw.simps nxt.simps) +by (rule ext) (metis alw.simps nxt.simps) lemma ev_ev[simp]: "ev (ev \) = ev \" proof- {fix xs assume "ev (ev \) xs" hence "ev \ xs" - by induct (auto intro: ev.intros) + by induct auto } - thus ?thesis by (auto intro: ev.intros) + thus ?thesis by auto qed lemma alw_alw[simp]: "alw (alw \) = alw \" proof- {fix xs assume "alw \ xs" hence "alw (alw \) xs" - by coinduct (auto elim: alw.cases) + by coinduct auto } - thus ?thesis by (auto elim: alw.cases) + thus ?thesis by auto qed lemma ev_shift: assumes "ev \ xs" shows "ev \ (xl @- xs)" -using assms by (induct xl) (auto intro: ev.intros) +using assms by (induct xl) auto lemma ev_imp_shift: assumes "ev \ xs" shows "\ xl xs2. xs = xl @- xs2 \ \ xs2" using assms by induct (metis shift.simps(1), metis shift.simps(2) stream.collapse)+ lemma alw_ev_shift: "alw \ xs1 \ ev (alw \) (xl @- xs1)" -by (auto intro: ev_shift ev.intros) +by (auto intro: ev_shift) lemma alw_shift: assumes "alw \ (xl @- xs)" shows "alw \ xs" -using assms by (induct xl) (auto elim: alw.cases) +using assms by (induct xl) auto lemma ev_ex_nxt: assumes "ev \ xs" @@ -224,15 +224,15 @@ using assms nxt_wait_least unfolding nxt_sdrop by auto lemma nxt_ev: "(nxt ^^ n) \ xs \ ev \ xs" -by (induct n arbitrary: xs) (auto intro: ev.intros) +by (induct n arbitrary: xs) auto lemma not_ev: "not (ev \) = alw (not \)" proof(rule ext, safe) fix xs assume "not (ev \) xs" thus "alw (not \) xs" - by (coinduct) (auto intro: ev.intros) + by (coinduct) auto next fix xs assume "ev \ xs" and "alw (not \) xs" thus False - by (induct) (auto elim: alw.cases) + by (induct) auto qed lemma not_alw: "not (alw \) = ev (not \)" @@ -256,9 +256,7 @@ lemma ev_alw_imp_alw_ev: assumes "ev (alw \) xs" shows "alw (ev \) xs" -using assms apply induct - apply (metis (full_types) alw_mono ev.base) - by (metis alw alw_nxt ev.step) +using assms by induct (metis (full_types) alw_mono ev.base, metis alw alw_nxt ev.step) lemma alw_aand: "alw (\ aand \) = alw \ aand alw \" proof- @@ -267,7 +265,7 @@ } moreover {fix xs assume "(alw \ aand alw \) xs" hence "alw (\ aand \) xs" - by coinduct (auto elim: alw.cases) + by coinduct auto } ultimately show ?thesis by blast qed @@ -279,7 +277,7 @@ } moreover {fix xs assume "ev (\ or \) xs" hence "(ev \ or ev \) xs" - by induct (auto intro: ev.intros) + by induct auto } ultimately show ?thesis by blast qed @@ -314,7 +312,7 @@ lemma ev_alw_alw_impl: assumes "ev (alw \) xs" and "alw (alw \ impl ev \) xs" shows "ev \ xs" -using assms by induct (auto intro: ev.intros elim: alw.cases) +using assms by induct auto lemma ev_alw_stl[simp]: "ev (alw \) (stl x) \ ev (alw \) x" by (metis (full_types) alw_nxt ev_nxt nxt.simps) @@ -323,18 +321,18 @@ "alw (alw \ impl ev \) = (ev (alw \) impl alw (ev \))" (is "?A = ?B") proof- {fix xs assume "?A xs \ ev (alw \) xs" hence "alw (ev \) xs" - apply coinduct using ev_nxt by (auto elim: ev_alw_alw_impl alw.cases intro: ev.intros) + by coinduct (auto elim: ev_alw_alw_impl) } moreover {fix xs assume "?B xs" hence "?A xs" - apply coinduct by (auto elim: alw.cases intro: ev.intros) + by coinduct auto } ultimately show ?thesis by blast qed lemma ev_alw_impl: assumes "ev \ xs" and "alw (\ impl \) xs" shows "ev \ xs" -using assms by induct (auto intro: ev.intros elim: alw.cases) +using assms by induct auto lemma ev_alw_impl_ev: assumes "ev \ xs" and "alw (\ impl ev \) xs" shows "ev \ xs" @@ -345,7 +343,7 @@ shows "alw \ xs" proof- {assume "alw \ xs \ alw (\ impl \) xs" hence ?thesis - apply coinduct by (auto elim: alw.cases) + by coinduct auto } thus ?thesis using assms by auto qed @@ -362,7 +360,7 @@ lemma alw_impl_ev_alw: assumes "alw (\ impl ev \) xs" shows "alw (ev \ impl ev \) xs" -using assms by coinduct (auto elim: alw.cases dest: ev_alw_impl intro: ev.intros) +using assms by coinduct (auto dest: ev_alw_impl) lemma ev_holds_sset: "ev (holds P) xs \ (\ x \ sset xs. P x)" (is "?L \ ?R") @@ -379,7 +377,7 @@ shows "alw \ xs" proof- {assume "\ xs \ alw (\ impl nxt \) xs" hence ?thesis - apply coinduct by(auto elim: alw.cases) + by coinduct auto } thus ?thesis using assms by auto qed @@ -390,7 +388,7 @@ proof- {assume "\ ev \ xs" hence "alw (not \) xs" unfolding not_ev[symmetric] . moreover have "alw (not \ impl (\ impl nxt \)) xs" - using 2 by coinduct (auto elim: alw.cases) + using 2 by coinduct auto ultimately have "alw (\ impl nxt \) xs" by(auto dest: alw_mp) with 1 have "alw \ xs" by(rule alw_invar) } @@ -404,7 +402,7 @@ obtain xl xs1 where xs: "xs = xl @- xs1" and \: "\ xs1" using e by (metis ev_imp_shift) have "\ xs1 \ alw (\ impl (nxt \)) xs1" using a \ unfolding xs by (metis alw_shift) - hence "alw \ xs1" by(coinduct xs1 rule: alw.coinduct) (auto elim: alw.cases) + hence "alw \ xs1" by(coinduct xs1 rule: alw.coinduct) auto thus ?thesis unfolding xs by (auto intro: alw_ev_shift) qed @@ -602,7 +600,7 @@ using suntil.induct[of \ \ x P] by blast lemma ev_suntil: "(\ suntil \) \ \ ev \ \" - by (induct rule: suntil.induct) (auto intro: ev.intros) + by (induct rule: suntil.induct) auto lemma suntil_inv: assumes stl: "\s. f (stl s) = stl (f s)"