# HG changeset patch # User hoelzl # Date 1412756520 -7200 # Node ID 1329679abb2d1c771a895f89a0aebfed54e3d08a # Parent 6c473ed0ac70213e6c0173550bdf3f6894f85e63 add Linear Temporal Logic on Streams diff -r 6c473ed0ac70 -r 1329679abb2d src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Wed Oct 08 09:09:12 2014 +0200 +++ b/src/HOL/Library/Library.thy Wed Oct 08 10:22:00 2014 +0200 @@ -36,6 +36,7 @@ Lattice_Algebras Lattice_Syntax Lattice_Constructions + Linear_Temporal_Logic_on_Streams ListVector Lubs_Glbs Mapping diff -r 6c473ed0ac70 -r 1329679abb2d src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy Wed Oct 08 10:22:00 2014 +0200 @@ -0,0 +1,391 @@ +(* Title: HOL/Library/Linear_Temporal_Logic_on_Streams.thy + Author: Andrei Popescu, TU Muenchen + Author: Dmitriy Traytel, TU Muenchen +*) + +header {* Linear Temporal Logic on Streams *} + +theory Linear_Temporal_Logic_on_Streams + imports Stream Sublist +begin + +section{* Preliminaries *} + +lemma shift_prefix: +assumes "xl @- xs = yl @- ys" and "length xl \ length yl" +shows "prefixeq xl yl" +using assms proof(induct xl arbitrary: yl xs ys) + case (Cons x xl yl xs ys) + thus ?case by (cases yl) auto +qed auto + +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) + + +section{* Linear temporal logic *} + +(* Propositional connectives: *) +abbreviation (input) IMPL (infix "impl" 60) +where "\ impl \ \ \ xs. \ xs \ \ xs" + +abbreviation (input) OR (infix "or" 60) +where "\ or \ \ \ xs. \ xs \ \ xs" + +abbreviation (input) AND (infix "aand" 60) +where "\ aand \ \ \ xs. \ xs \ \ xs" + +abbreviation (input) "not \ \ \ xs. \ \ xs" + +abbreviation (input) "true \ \ xs. True" + +abbreviation (input) "false \ \ xs. False" + +lemma impl_not_or: "\ impl \ = (not \) or \" +by blast + +lemma not_or: "not (\ or \) = (not \) aand (not \)" +by blast + +lemma not_aand: "not (\ aand \) = (not \) or (not \)" +by blast + +lemma non_not[simp]: "not (not \) = \" by simp + +(* Temporal (LTL) connectives: *) +fun holds where "holds P xs \ P (shd xs)" +fun nxt where "nxt \ xs = \ (stl xs)" + +inductive ev for \ where +base: "\ xs \ ev \ xs" +| +step: "ev \ (stl xs) \ ev \ xs" + +coinductive alw for \ where +alw: "\\ xs; alw \ (stl xs)\ \ alw \ xs" + +(* weak until: *) +coinductive UNTIL (infix "until" 60) for \ \ where +base: "\ xs \ (\ until \) xs" +| +step: "\\ xs; (\ until \) (stl xs)\ \ (\ until \) xs" + +lemma holds_mono: +assumes holds: "holds P xs" and 0: "\ x. P x \ Q x" +shows "holds Q xs" +using assms by auto + +lemma holds_aand: +"(holds P aand holds Q) steps \ holds (\ step. P step \ Q step) steps" by auto + +lemma nxt_mono: +assumes nxt: "nxt \ xs" and 0: "\ xs. \ xs \ \ xs" +shows "nxt \ xs" +using assms by auto + +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) + +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) + +lemma until_monoL: +assumes until: "(\1 until \) xs" and 0: "\ xs. \1 xs \ \2 xs" +shows "(\2 until \) xs" +using until by coinduct (auto elim: UNTIL.cases simp: 0) + +lemma until_monoR: +assumes until: "(\ until \1) xs" and 0: "\ xs. \1 xs \ \2 xs" +shows "(\ until \2) xs" +using until by coinduct (auto elim: UNTIL.cases simp: 0) + +lemma until_mono: +assumes until: "(\1 until \1) xs" and +0: "\ xs. \1 xs \ \2 xs" "\ xs. \1 xs \ \2 xs" +shows "(\2 until \2) xs" +using until by coinduct (auto elim: UNTIL.cases simp: 0) + +lemma until_false: "\ until false = alw \" +proof- + {fix xs assume "(\ until false) xs" hence "alw \ xs" + apply coinduct by (auto elim: UNTIL.cases) + } + moreover + {fix xs assume "alw \ xs" hence "(\ until false) xs" + apply coinduct by (auto elim: alw.cases) + } + ultimately show ?thesis by blast +qed + +lemma ev_nxt: "ev \ = (\ or nxt (ev \))" +apply(rule ext) by (metis ev.simps nxt.simps) + +lemma alw_nxt: "alw \ = (\ aand nxt (alw \))" +apply(rule ext) by (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) + } + thus ?thesis by (auto intro: ev.intros) +qed + +lemma alw_alw[simp]: "alw (alw \) = alw \" +proof- + {fix xs + assume "alw \ xs" hence "alw (alw \) xs" + by coinduct (auto elim: alw.cases) + } + thus ?thesis by (auto elim: alw.cases) +qed + +lemma ev_shift: +assumes "ev \ xs" +shows "ev \ (xl @- xs)" +using assms by (induct xl) (auto intro: ev.intros) + +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) + +lemma alw_shift: +assumes "alw \ (xl @- xs)" +shows "alw \ xs" +using assms by (induct xl) (auto elim: alw.cases) + +lemma ev_ex_nxt: +assumes "ev \ xs" +shows "\ n. (nxt ^^ n) \ xs" +using assms proof induct + case (base xs) thus ?case by (intro exI[of _ 0]) auto +next + case (step xs) + then obtain n where "(nxt ^^ n) \ (stl xs)" by blast + thus ?case by (intro exI[of _ "Suc n"]) (metis funpow.simps(2) nxt.simps o_def) +qed + +lemma alw_sdrop: +assumes "alw \ xs" shows "alw \ (sdrop n xs)" +by (metis alw_shift assms stake_sdrop) + +lemma nxt_sdrop: "(nxt ^^ n) \ xs \ \ (sdrop n xs)" +by (induct n arbitrary: xs) auto + +definition "wait \ xs \ LEAST n. (nxt ^^ n) \ xs" + +lemma nxt_wait: +assumes "ev \ xs" shows "(nxt ^^ (wait \ xs)) \ xs" +unfolding wait_def using ev_ex_nxt[OF assms] by(rule LeastI_ex) + +lemma nxt_wait_least: +assumes ev: "ev \ xs" and nxt: "(nxt ^^ n) \ xs" shows "wait \ xs \ n" +unfolding wait_def using ev_ex_nxt[OF ev] by (metis Least_le nxt) + +lemma sdrop_wait: +assumes "ev \ xs" shows "\ (sdrop (wait \ xs) xs)" +using nxt_wait[OF assms] unfolding nxt_sdrop . + +lemma sdrop_wait_least: +assumes ev: "ev \ xs" and nxt: "\ (sdrop n xs)" shows "wait \ xs \ n" +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) + +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) +next + fix xs assume "ev \ xs" and "alw (not \) xs" thus False + by (induct) (auto elim: alw.cases) +qed + +lemma not_alw: "not (alw \) = ev (not \)" +proof- + have "not (alw \) = not (alw (not (not \)))" by simp + also have "... = ev (not \)" unfolding not_ev[symmetric] by simp + finally show ?thesis . +qed + +lemma not_ev_not[simp]: "not (ev (not \)) = alw \" +unfolding not_ev by simp + +lemma not_alw_not[simp]: "not (alw (not \)) = ev \" +unfolding not_alw by simp + +lemma alw_ev_sdrop: +assumes "alw (ev \) (sdrop m xs)" +shows "alw (ev \) xs" +using assms +by coinduct (metis alw_nxt ev_shift funpow_swap1 nxt.simps nxt_sdrop stake_sdrop) + +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) + +lemma alw_aand: "alw (\ aand \) = alw \ aand alw \" +proof- + {fix xs assume "alw (\ aand \) xs" hence "(alw \ aand alw \) xs" + by (auto elim: alw_mono) + } + moreover + {fix xs assume "(alw \ aand alw \) xs" hence "alw (\ aand \) xs" + by coinduct (auto elim: alw.cases) + } + ultimately show ?thesis by blast +qed + +lemma ev_or: "ev (\ or \) = ev \ or ev \" +proof- + {fix xs assume "(ev \ or ev \) xs" hence "ev (\ or \) xs" + by (auto elim: ev_mono) + } + moreover + {fix xs assume "ev (\ or \) xs" hence "(ev \ or ev \) xs" + by induct (auto intro: ev.intros) + } + ultimately show ?thesis by blast +qed + +lemma ev_alw_aand: +assumes \: "ev (alw \) xs" and \: "ev (alw \) xs" +shows "ev (alw (\ aand \)) xs" +proof- + obtain xl xs1 where xs1: "xs = xl @- xs1" and \\: "alw \ xs1" + using \ by (metis ev_imp_shift) + moreover obtain yl ys1 where xs2: "xs = yl @- ys1" and \\: "alw \ ys1" + using \ by (metis ev_imp_shift) + ultimately have 0: "xl @- xs1 = yl @- ys1" by auto + hence "prefixeq xl yl \ prefixeq yl xl" using shift_prefix_cases by auto + thus ?thesis proof + assume "prefixeq xl yl" + then obtain yl1 where yl: "yl = xl @ yl1" by (elim prefixeqE) + have xs1': "xs1 = yl1 @- ys1" using 0 unfolding yl by simp + have "alw \ ys1" using \\ unfolding xs1' by (metis alw_shift) + hence "alw (\ aand \) ys1" using \\ unfolding alw_aand by auto + thus ?thesis unfolding xs2 by (auto intro: alw_ev_shift) + next + assume "prefixeq yl xl" + then obtain xl1 where xl: "xl = yl @ xl1" by (elim prefixeqE) + have ys1': "ys1 = xl1 @- xs1" using 0 unfolding xl by simp + have "alw \ xs1" using \\ unfolding ys1' by (metis alw_shift) + hence "alw (\ aand \) xs1" using \\ unfolding alw_aand by auto + thus ?thesis unfolding xs1 by (auto intro: alw_ev_shift) + qed +qed + +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) + +lemma ev_alw_stl[simp]: "ev (alw \) (stl x) \ ev (alw \) x" +by (metis (full_types) alw_nxt ev_nxt nxt.simps) + +lemma alw_alw_impl_ev: +"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) + } + moreover + {fix xs assume "?B xs" hence "?A xs" + apply coinduct by (auto elim: alw.cases intro: ev.intros) + } + 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) + +lemma ev_alw_impl_ev: +assumes "ev \ xs" and "alw (\ impl ev \) xs" shows "ev \ xs" +using ev_alw_impl[OF assms] by simp + +lemma alw_mp: +assumes "alw \ xs" and "alw (\ impl \) xs" +shows "alw \ xs" +proof- + {assume "alw \ xs \ alw (\ impl \) xs" hence ?thesis + apply coinduct by (auto elim: alw.cases) + } + thus ?thesis using assms by auto +qed + +lemma all_imp_alw: +assumes "\ xs. \ xs" shows "alw \ xs" +proof- + {assume "\ xs. \ xs" + hence ?thesis by coinduct auto + } + thus ?thesis using assms by auto +qed + +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) + +lemma ev_holds_sset: +"ev (holds P) xs \ (\ x \ sset xs. P x)" (is "?L \ ?R") +proof safe + assume ?L thus ?R by induct (metis holds.simps stream.set_sel(1), metis stl_sset) +next + fix x assume "x \ sset xs" "P x" + thus ?L by (induct rule: sset_induct) (simp_all add: ev.base ev.step) +qed + +(* LTL as a program logic: *) +lemma alw_invar: +assumes "\ xs" and "alw (\ impl nxt \) xs" +shows "alw \ xs" +proof- + {assume "\ xs \ alw (\ impl nxt \) xs" hence ?thesis + apply coinduct by(auto elim: alw.cases) + } + thus ?thesis using assms by auto +qed + +lemma variance: +assumes 1: "\ xs" and 2: "alw (\ impl (\ or nxt \)) xs" +shows "(alw \ or ev \) xs" +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) + ultimately have "alw (\ impl nxt \) xs" by(auto dest: alw_mp) + with 1 have "alw \ xs" by(rule alw_invar) + } + thus ?thesis by blast +qed + +lemma ev_alw_imp_nxt: +assumes e: "ev \ xs" and a: "alw (\ impl (nxt \)) xs" +shows "ev (alw \) xs" +proof- + 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) + thus ?thesis unfolding xs by (auto intro: alw_ev_shift) +qed + + + +end \ No newline at end of file