| 58627 |      1 | (*  Title:      HOL/Library/Linear_Temporal_Logic_on_Streams.thy
 | 
|  |      2 |     Author:     Andrei Popescu, TU Muenchen
 | 
|  |      3 |     Author:     Dmitriy Traytel, TU Muenchen
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
| 60500 |      6 | section \<open>Linear Temporal Logic on Streams\<close>
 | 
| 58627 |      7 | 
 | 
|  |      8 | theory Linear_Temporal_Logic_on_Streams
 | 
| 59000 |      9 |   imports Stream Sublist Extended_Nat Infinite_Set
 | 
| 58627 |     10 | begin
 | 
|  |     11 | 
 | 
| 60500 |     12 | section\<open>Preliminaries\<close>
 | 
| 58627 |     13 | 
 | 
|  |     14 | lemma shift_prefix:
 | 
|  |     15 | assumes "xl @- xs = yl @- ys" and "length xl \<le> length yl"
 | 
|  |     16 | shows "prefixeq xl yl"
 | 
|  |     17 | using assms proof(induct xl arbitrary: yl xs ys)
 | 
|  |     18 |   case (Cons x xl yl xs ys)
 | 
|  |     19 |   thus ?case by (cases yl) auto
 | 
|  |     20 | qed auto
 | 
|  |     21 | 
 | 
|  |     22 | lemma shift_prefix_cases:
 | 
|  |     23 | assumes "xl @- xs = yl @- ys"
 | 
|  |     24 | shows "prefixeq xl yl \<or> prefixeq yl xl"
 | 
|  |     25 | using shift_prefix[OF assms] apply(cases "length xl \<le> length yl")
 | 
|  |     26 | by (metis, metis assms nat_le_linear shift_prefix)
 | 
|  |     27 | 
 | 
|  |     28 | 
 | 
| 60500 |     29 | section\<open>Linear temporal logic\<close>
 | 
| 58627 |     30 | 
 | 
|  |     31 | (* Propositional connectives: *)
 | 
|  |     32 | abbreviation (input) IMPL (infix "impl" 60)
 | 
|  |     33 | where "\<phi> impl \<psi> \<equiv> \<lambda> xs. \<phi> xs \<longrightarrow> \<psi> xs"
 | 
|  |     34 | 
 | 
|  |     35 | abbreviation (input) OR (infix "or" 60)
 | 
|  |     36 | where "\<phi> or \<psi> \<equiv> \<lambda> xs. \<phi> xs \<or> \<psi> xs"
 | 
|  |     37 | 
 | 
|  |     38 | abbreviation (input) AND (infix "aand" 60)
 | 
|  |     39 | where "\<phi> aand \<psi> \<equiv> \<lambda> xs. \<phi> xs \<and> \<psi> xs"
 | 
|  |     40 | 
 | 
|  |     41 | abbreviation (input) "not \<phi> \<equiv> \<lambda> xs. \<not> \<phi> xs"
 | 
|  |     42 | 
 | 
|  |     43 | abbreviation (input) "true \<equiv> \<lambda> xs. True"
 | 
|  |     44 | 
 | 
|  |     45 | abbreviation (input) "false \<equiv> \<lambda> xs. False"
 | 
|  |     46 | 
 | 
|  |     47 | lemma impl_not_or: "\<phi> impl \<psi> = (not \<phi>) or \<psi>"
 | 
|  |     48 | by blast
 | 
|  |     49 | 
 | 
|  |     50 | lemma not_or: "not (\<phi> or \<psi>) = (not \<phi>) aand (not \<psi>)"
 | 
|  |     51 | by blast
 | 
|  |     52 | 
 | 
|  |     53 | lemma not_aand: "not (\<phi> aand \<psi>) = (not \<phi>) or (not \<psi>)"
 | 
|  |     54 | by blast
 | 
|  |     55 | 
 | 
|  |     56 | lemma non_not[simp]: "not (not \<phi>) = \<phi>" by simp
 | 
|  |     57 | 
 | 
|  |     58 | (* Temporal (LTL) connectives: *)
 | 
|  |     59 | fun holds where "holds P xs \<longleftrightarrow> P (shd xs)"
 | 
|  |     60 | fun nxt where "nxt \<phi> xs = \<phi> (stl xs)"
 | 
|  |     61 | 
 | 
| 59000 |     62 | definition "HLD s = holds (\<lambda>x. x \<in> s)"
 | 
|  |     63 | 
 | 
|  |     64 | abbreviation HLD_nxt (infixr "\<cdot>" 65) where
 | 
|  |     65 |   "s \<cdot> P \<equiv> HLD s aand nxt P"
 | 
|  |     66 | 
 | 
| 58627 |     67 | inductive ev for \<phi> where
 | 
|  |     68 | base: "\<phi> xs \<Longrightarrow> ev \<phi> xs"
 | 
|  |     69 | |
 | 
|  |     70 | step: "ev \<phi> (stl xs) \<Longrightarrow> ev \<phi> xs"
 | 
|  |     71 | 
 | 
|  |     72 | coinductive alw for \<phi> where
 | 
|  |     73 | alw: "\<lbrakk>\<phi> xs; alw \<phi> (stl xs)\<rbrakk> \<Longrightarrow> alw \<phi> xs"
 | 
|  |     74 | 
 | 
|  |     75 | (* weak until: *)
 | 
|  |     76 | coinductive UNTIL (infix "until" 60) for \<phi> \<psi> where
 | 
|  |     77 | base: "\<psi> xs \<Longrightarrow> (\<phi> until \<psi>) xs"
 | 
|  |     78 | |
 | 
|  |     79 | step: "\<lbrakk>\<phi> xs; (\<phi> until \<psi>) (stl xs)\<rbrakk> \<Longrightarrow> (\<phi> until \<psi>) xs"
 | 
|  |     80 | 
 | 
|  |     81 | lemma holds_mono:
 | 
|  |     82 | assumes holds: "holds P xs" and 0: "\<And> x. P x \<Longrightarrow> Q x"
 | 
|  |     83 | shows "holds Q xs"
 | 
|  |     84 | using assms by auto
 | 
|  |     85 | 
 | 
|  |     86 | lemma holds_aand:
 | 
|  |     87 | "(holds P aand holds Q) steps \<longleftrightarrow> holds (\<lambda> step. P step \<and> Q step) steps" by auto
 | 
|  |     88 | 
 | 
| 59000 |     89 | lemma HLD_iff: "HLD s \<omega> \<longleftrightarrow> shd \<omega> \<in> s"
 | 
|  |     90 |   by (simp add: HLD_def)
 | 
|  |     91 | 
 | 
|  |     92 | lemma HLD_Stream[simp]: "HLD X (x ## \<omega>) \<longleftrightarrow> x \<in> X"
 | 
|  |     93 |   by (simp add: HLD_iff)
 | 
|  |     94 | 
 | 
| 58627 |     95 | lemma nxt_mono:
 | 
|  |     96 | assumes nxt: "nxt \<phi> xs" and 0: "\<And> xs. \<phi> xs \<Longrightarrow> \<psi> xs"
 | 
|  |     97 | shows "nxt \<psi> xs"
 | 
|  |     98 | using assms by auto
 | 
|  |     99 | 
 | 
| 59000 |    100 | declare ev.intros[intro]
 | 
|  |    101 | declare alw.cases[elim]
 | 
|  |    102 | 
 | 
|  |    103 | lemma ev_induct_strong[consumes 1, case_names base step]:
 | 
|  |    104 |   "ev \<phi> x \<Longrightarrow> (\<And>xs. \<phi> xs \<Longrightarrow> P xs) \<Longrightarrow> (\<And>xs. ev \<phi> (stl xs) \<Longrightarrow> \<not> \<phi> xs \<Longrightarrow> P (stl xs) \<Longrightarrow> P xs) \<Longrightarrow> P x"
 | 
|  |    105 |   by (induct rule: ev.induct) auto
 | 
|  |    106 | 
 | 
|  |    107 | lemma alw_coinduct[consumes 1, case_names alw stl]:
 | 
|  |    108 |   "X x \<Longrightarrow> (\<And>x. X x \<Longrightarrow> \<phi> x) \<Longrightarrow> (\<And>x. X x \<Longrightarrow> \<not> alw \<phi> (stl x) \<Longrightarrow> X (stl x)) \<Longrightarrow> alw \<phi> x"
 | 
|  |    109 |   using alw.coinduct[of X x \<phi>] by auto
 | 
|  |    110 | 
 | 
| 58627 |    111 | lemma ev_mono:
 | 
|  |    112 | assumes ev: "ev \<phi> xs" and 0: "\<And> xs. \<phi> xs \<Longrightarrow> \<psi> xs"
 | 
|  |    113 | shows "ev \<psi> xs"
 | 
|  |    114 | using ev by induct (auto intro: ev.intros simp: 0)
 | 
|  |    115 | 
 | 
|  |    116 | lemma alw_mono:
 | 
|  |    117 | assumes alw: "alw \<phi> xs" and 0: "\<And> xs. \<phi> xs \<Longrightarrow> \<psi> xs"
 | 
|  |    118 | shows "alw \<psi> xs"
 | 
|  |    119 | using alw by coinduct (auto elim: alw.cases simp: 0)
 | 
|  |    120 | 
 | 
|  |    121 | lemma until_monoL:
 | 
|  |    122 | assumes until: "(\<phi>1 until \<psi>) xs" and 0: "\<And> xs. \<phi>1 xs \<Longrightarrow> \<phi>2 xs"
 | 
|  |    123 | shows "(\<phi>2 until \<psi>) xs"
 | 
|  |    124 | using until by coinduct (auto elim: UNTIL.cases simp: 0)
 | 
|  |    125 | 
 | 
|  |    126 | lemma until_monoR:
 | 
|  |    127 | assumes until: "(\<phi> until \<psi>1) xs" and 0: "\<And> xs. \<psi>1 xs \<Longrightarrow> \<psi>2 xs"
 | 
|  |    128 | shows "(\<phi> until \<psi>2) xs"
 | 
|  |    129 | using until by coinduct (auto elim: UNTIL.cases simp: 0)
 | 
|  |    130 | 
 | 
|  |    131 | lemma until_mono:
 | 
|  |    132 | assumes until: "(\<phi>1 until \<psi>1) xs" and
 | 
|  |    133 | 0: "\<And> xs. \<phi>1 xs \<Longrightarrow> \<phi>2 xs" "\<And> xs. \<psi>1 xs \<Longrightarrow> \<psi>2 xs"
 | 
|  |    134 | shows "(\<phi>2 until \<psi>2) xs"
 | 
|  |    135 | using until by coinduct (auto elim: UNTIL.cases simp: 0)
 | 
|  |    136 | 
 | 
|  |    137 | lemma until_false: "\<phi> until false = alw \<phi>"
 | 
|  |    138 | proof-
 | 
|  |    139 |   {fix xs assume "(\<phi> until false) xs" hence "alw \<phi> xs"
 | 
|  |    140 |    apply coinduct by (auto elim: UNTIL.cases)
 | 
|  |    141 |   }
 | 
|  |    142 |   moreover
 | 
|  |    143 |   {fix xs assume "alw \<phi> xs" hence "(\<phi> until false) xs"
 | 
|  |    144 |    apply coinduct by (auto elim: alw.cases)
 | 
|  |    145 |   }
 | 
|  |    146 |   ultimately show ?thesis by blast
 | 
|  |    147 | qed
 | 
|  |    148 | 
 | 
|  |    149 | lemma ev_nxt: "ev \<phi> = (\<phi> or nxt (ev \<phi>))"
 | 
|  |    150 | apply(rule ext) by (metis ev.simps nxt.simps)
 | 
|  |    151 | 
 | 
|  |    152 | lemma alw_nxt: "alw \<phi> = (\<phi> aand nxt (alw \<phi>))"
 | 
|  |    153 | apply(rule ext) by (metis alw.simps nxt.simps)
 | 
|  |    154 | 
 | 
|  |    155 | lemma ev_ev[simp]: "ev (ev \<phi>) = ev \<phi>"
 | 
|  |    156 | proof-
 | 
|  |    157 |   {fix xs
 | 
|  |    158 |    assume "ev (ev \<phi>) xs" hence "ev \<phi> xs"
 | 
|  |    159 |    by induct (auto intro: ev.intros)
 | 
|  |    160 |   }
 | 
|  |    161 |   thus ?thesis by (auto intro: ev.intros)
 | 
|  |    162 | qed
 | 
|  |    163 | 
 | 
|  |    164 | lemma alw_alw[simp]: "alw (alw \<phi>) = alw \<phi>"
 | 
|  |    165 | proof-
 | 
|  |    166 |   {fix xs
 | 
|  |    167 |    assume "alw \<phi> xs" hence "alw (alw \<phi>) xs"
 | 
|  |    168 |    by coinduct (auto elim: alw.cases)
 | 
|  |    169 |   }
 | 
|  |    170 |   thus ?thesis by (auto elim: alw.cases)
 | 
|  |    171 | qed
 | 
|  |    172 | 
 | 
|  |    173 | lemma ev_shift:
 | 
|  |    174 | assumes "ev \<phi> xs"
 | 
|  |    175 | shows "ev \<phi> (xl @- xs)"
 | 
|  |    176 | using assms by (induct xl) (auto intro: ev.intros)
 | 
|  |    177 | 
 | 
|  |    178 | lemma ev_imp_shift:
 | 
|  |    179 | assumes "ev \<phi> xs"  shows "\<exists> xl xs2. xs = xl @- xs2 \<and> \<phi> xs2"
 | 
|  |    180 | using assms by induct (metis shift.simps(1), metis shift.simps(2) stream.collapse)+
 | 
|  |    181 | 
 | 
|  |    182 | lemma alw_ev_shift: "alw \<phi> xs1 \<Longrightarrow> ev (alw \<phi>) (xl @- xs1)"
 | 
|  |    183 | by (auto intro: ev_shift ev.intros)
 | 
|  |    184 | 
 | 
|  |    185 | lemma alw_shift:
 | 
|  |    186 | assumes "alw \<phi> (xl @- xs)"
 | 
|  |    187 | shows "alw \<phi> xs"
 | 
|  |    188 | using assms by (induct xl) (auto elim: alw.cases)
 | 
|  |    189 | 
 | 
|  |    190 | lemma ev_ex_nxt:
 | 
|  |    191 | assumes "ev \<phi> xs"
 | 
|  |    192 | shows "\<exists> n. (nxt ^^ n) \<phi> xs"
 | 
|  |    193 | using assms proof induct
 | 
|  |    194 |   case (base xs) thus ?case by (intro exI[of _ 0]) auto
 | 
|  |    195 | next
 | 
|  |    196 |   case (step xs)
 | 
|  |    197 |   then obtain n where "(nxt ^^ n) \<phi> (stl xs)" by blast
 | 
|  |    198 |   thus ?case by (intro exI[of _ "Suc n"]) (metis funpow.simps(2) nxt.simps o_def)
 | 
|  |    199 | qed
 | 
|  |    200 | 
 | 
|  |    201 | lemma alw_sdrop:
 | 
|  |    202 | assumes "alw \<phi> xs"  shows "alw \<phi> (sdrop n xs)"
 | 
|  |    203 | by (metis alw_shift assms stake_sdrop)
 | 
|  |    204 | 
 | 
|  |    205 | lemma nxt_sdrop: "(nxt ^^ n) \<phi> xs \<longleftrightarrow> \<phi> (sdrop n xs)"
 | 
|  |    206 | by (induct n arbitrary: xs) auto
 | 
|  |    207 | 
 | 
|  |    208 | definition "wait \<phi> xs \<equiv> LEAST n. (nxt ^^ n) \<phi> xs"
 | 
|  |    209 | 
 | 
|  |    210 | lemma nxt_wait:
 | 
|  |    211 | assumes "ev \<phi> xs"  shows "(nxt ^^ (wait \<phi> xs)) \<phi> xs"
 | 
|  |    212 | unfolding wait_def using ev_ex_nxt[OF assms] by(rule LeastI_ex)
 | 
|  |    213 | 
 | 
|  |    214 | lemma nxt_wait_least:
 | 
|  |    215 | assumes ev: "ev \<phi> xs" and nxt: "(nxt ^^ n) \<phi> xs"  shows "wait \<phi> xs \<le> n"
 | 
|  |    216 | unfolding wait_def using ev_ex_nxt[OF ev] by (metis Least_le nxt)
 | 
|  |    217 | 
 | 
|  |    218 | lemma sdrop_wait:
 | 
|  |    219 | assumes "ev \<phi> xs"  shows "\<phi> (sdrop (wait \<phi> xs) xs)"
 | 
|  |    220 | using nxt_wait[OF assms] unfolding nxt_sdrop .
 | 
|  |    221 | 
 | 
|  |    222 | lemma sdrop_wait_least:
 | 
|  |    223 | assumes ev: "ev \<phi> xs" and nxt: "\<phi> (sdrop n xs)"  shows "wait \<phi> xs \<le> n"
 | 
|  |    224 | using assms nxt_wait_least unfolding nxt_sdrop by auto
 | 
|  |    225 | 
 | 
|  |    226 | lemma nxt_ev: "(nxt ^^ n) \<phi> xs \<Longrightarrow> ev \<phi> xs"
 | 
|  |    227 | by (induct n arbitrary: xs) (auto intro: ev.intros)
 | 
|  |    228 | 
 | 
|  |    229 | lemma not_ev: "not (ev \<phi>) = alw (not \<phi>)"
 | 
|  |    230 | proof(rule ext, safe)
 | 
|  |    231 |   fix xs assume "not (ev \<phi>) xs" thus "alw (not \<phi>) xs"
 | 
|  |    232 |   by (coinduct) (auto intro: ev.intros)
 | 
|  |    233 | next
 | 
|  |    234 |   fix xs assume "ev \<phi> xs" and "alw (not \<phi>) xs" thus False
 | 
|  |    235 |   by (induct) (auto elim: alw.cases)
 | 
|  |    236 | qed
 | 
|  |    237 | 
 | 
|  |    238 | lemma not_alw: "not (alw \<phi>) = ev (not \<phi>)"
 | 
|  |    239 | proof-
 | 
|  |    240 |   have "not (alw \<phi>) = not (alw (not (not \<phi>)))" by simp
 | 
|  |    241 |   also have "... = ev (not \<phi>)" unfolding not_ev[symmetric] by simp
 | 
|  |    242 |   finally show ?thesis .
 | 
|  |    243 | qed
 | 
|  |    244 | 
 | 
|  |    245 | lemma not_ev_not[simp]: "not (ev (not \<phi>)) = alw \<phi>"
 | 
|  |    246 | unfolding not_ev by simp
 | 
|  |    247 | 
 | 
|  |    248 | lemma not_alw_not[simp]: "not (alw (not \<phi>)) = ev \<phi>"
 | 
|  |    249 | unfolding not_alw by simp
 | 
|  |    250 | 
 | 
|  |    251 | lemma alw_ev_sdrop:
 | 
|  |    252 | assumes "alw (ev \<phi>) (sdrop m xs)"
 | 
|  |    253 | shows "alw (ev \<phi>) xs"
 | 
|  |    254 | using assms
 | 
|  |    255 | by coinduct (metis alw_nxt ev_shift funpow_swap1 nxt.simps nxt_sdrop stake_sdrop)
 | 
|  |    256 | 
 | 
|  |    257 | lemma ev_alw_imp_alw_ev:
 | 
|  |    258 | assumes "ev (alw \<phi>) xs"  shows "alw (ev \<phi>) xs"
 | 
|  |    259 | using assms apply induct
 | 
|  |    260 |   apply (metis (full_types) alw_mono ev.base)
 | 
|  |    261 |   by (metis alw alw_nxt ev.step)
 | 
|  |    262 | 
 | 
|  |    263 | lemma alw_aand: "alw (\<phi> aand \<psi>) = alw \<phi> aand alw \<psi>"
 | 
|  |    264 | proof-
 | 
|  |    265 |   {fix xs assume "alw (\<phi> aand \<psi>) xs" hence "(alw \<phi> aand alw \<psi>) xs"
 | 
|  |    266 |    by (auto elim: alw_mono)
 | 
|  |    267 |   }
 | 
|  |    268 |   moreover
 | 
|  |    269 |   {fix xs assume "(alw \<phi> aand alw \<psi>) xs" hence "alw (\<phi> aand \<psi>) xs"
 | 
|  |    270 |    by coinduct (auto elim: alw.cases)
 | 
|  |    271 |   }
 | 
|  |    272 |   ultimately show ?thesis by blast
 | 
|  |    273 | qed
 | 
|  |    274 | 
 | 
|  |    275 | lemma ev_or: "ev (\<phi> or \<psi>) = ev \<phi> or ev \<psi>"
 | 
|  |    276 | proof-
 | 
|  |    277 |   {fix xs assume "(ev \<phi> or ev \<psi>) xs" hence "ev (\<phi> or \<psi>) xs"
 | 
|  |    278 |    by (auto elim: ev_mono)
 | 
|  |    279 |   }
 | 
|  |    280 |   moreover
 | 
|  |    281 |   {fix xs assume "ev (\<phi> or \<psi>) xs" hence "(ev \<phi> or ev \<psi>) xs"
 | 
|  |    282 |    by induct (auto intro: ev.intros)
 | 
|  |    283 |   }
 | 
|  |    284 |   ultimately show ?thesis by blast
 | 
|  |    285 | qed
 | 
|  |    286 | 
 | 
|  |    287 | lemma ev_alw_aand:
 | 
|  |    288 | assumes \<phi>: "ev (alw \<phi>) xs" and \<psi>: "ev (alw \<psi>) xs"
 | 
|  |    289 | shows "ev (alw (\<phi> aand \<psi>)) xs"
 | 
|  |    290 | proof-
 | 
|  |    291 |   obtain xl xs1 where xs1: "xs = xl @- xs1" and \<phi>\<phi>: "alw \<phi> xs1"
 | 
|  |    292 |   using \<phi> by (metis ev_imp_shift)
 | 
|  |    293 |   moreover obtain yl ys1 where xs2: "xs = yl @- ys1" and \<psi>\<psi>: "alw \<psi> ys1"
 | 
|  |    294 |   using \<psi> by (metis ev_imp_shift)
 | 
|  |    295 |   ultimately have 0: "xl @- xs1 = yl @- ys1" by auto
 | 
|  |    296 |   hence "prefixeq xl yl \<or> prefixeq yl xl" using shift_prefix_cases by auto
 | 
|  |    297 |   thus ?thesis proof
 | 
|  |    298 |     assume "prefixeq xl yl"
 | 
|  |    299 |     then obtain yl1 where yl: "yl = xl @ yl1" by (elim prefixeqE)
 | 
|  |    300 |     have xs1': "xs1 = yl1 @- ys1" using 0 unfolding yl by simp
 | 
|  |    301 |     have "alw \<phi> ys1" using \<phi>\<phi> unfolding xs1' by (metis alw_shift)
 | 
|  |    302 |     hence "alw (\<phi> aand \<psi>) ys1" using \<psi>\<psi> unfolding alw_aand by auto
 | 
|  |    303 |     thus ?thesis unfolding xs2 by (auto intro: alw_ev_shift)
 | 
|  |    304 |   next
 | 
|  |    305 |     assume "prefixeq yl xl"
 | 
|  |    306 |     then obtain xl1 where xl: "xl = yl @ xl1" by (elim prefixeqE)
 | 
|  |    307 |     have ys1': "ys1 = xl1 @- xs1" using 0 unfolding xl by simp
 | 
|  |    308 |     have "alw \<psi> xs1" using \<psi>\<psi> unfolding ys1' by (metis alw_shift)
 | 
|  |    309 |     hence "alw (\<phi> aand \<psi>) xs1" using \<phi>\<phi> unfolding alw_aand by auto
 | 
|  |    310 |     thus ?thesis unfolding xs1 by (auto intro: alw_ev_shift)
 | 
|  |    311 |   qed
 | 
|  |    312 | qed
 | 
|  |    313 | 
 | 
|  |    314 | lemma ev_alw_alw_impl:
 | 
|  |    315 | assumes "ev (alw \<phi>) xs" and "alw (alw \<phi> impl ev \<psi>) xs"
 | 
|  |    316 | shows "ev \<psi> xs"
 | 
|  |    317 | using assms by induct (auto intro: ev.intros elim: alw.cases)
 | 
|  |    318 | 
 | 
|  |    319 | lemma ev_alw_stl[simp]: "ev (alw \<phi>) (stl x) \<longleftrightarrow> ev (alw \<phi>) x"
 | 
|  |    320 | by (metis (full_types) alw_nxt ev_nxt nxt.simps)
 | 
|  |    321 | 
 | 
|  |    322 | lemma alw_alw_impl_ev:
 | 
|  |    323 | "alw (alw \<phi> impl ev \<psi>) = (ev (alw \<phi>) impl alw (ev \<psi>))" (is "?A = ?B")
 | 
|  |    324 | proof-
 | 
|  |    325 |   {fix xs assume "?A xs \<and> ev (alw \<phi>) xs" hence "alw (ev \<psi>) xs"
 | 
|  |    326 |    apply coinduct using ev_nxt by (auto elim: ev_alw_alw_impl alw.cases intro: ev.intros)
 | 
|  |    327 |   }
 | 
|  |    328 |   moreover
 | 
|  |    329 |   {fix xs assume "?B xs" hence "?A xs"
 | 
|  |    330 |    apply coinduct by (auto elim: alw.cases intro: ev.intros)
 | 
|  |    331 |   }
 | 
|  |    332 |   ultimately show ?thesis by blast
 | 
|  |    333 | qed
 | 
|  |    334 | 
 | 
|  |    335 | lemma ev_alw_impl:
 | 
|  |    336 | assumes "ev \<phi> xs" and "alw (\<phi> impl \<psi>) xs"  shows "ev \<psi> xs"
 | 
|  |    337 | using assms by induct (auto intro: ev.intros elim: alw.cases)
 | 
|  |    338 | 
 | 
|  |    339 | lemma ev_alw_impl_ev:
 | 
|  |    340 | assumes "ev \<phi> xs" and "alw (\<phi> impl ev \<psi>) xs"  shows "ev \<psi> xs"
 | 
|  |    341 | using ev_alw_impl[OF assms] by simp
 | 
|  |    342 | 
 | 
|  |    343 | lemma alw_mp:
 | 
|  |    344 | assumes "alw \<phi> xs" and "alw (\<phi> impl \<psi>) xs"
 | 
|  |    345 | shows "alw \<psi> xs"
 | 
|  |    346 | proof-
 | 
|  |    347 |   {assume "alw \<phi> xs \<and> alw (\<phi> impl \<psi>) xs" hence ?thesis
 | 
|  |    348 |    apply coinduct by (auto elim: alw.cases)
 | 
|  |    349 |   }
 | 
|  |    350 |   thus ?thesis using assms by auto
 | 
|  |    351 | qed
 | 
|  |    352 | 
 | 
|  |    353 | lemma all_imp_alw:
 | 
|  |    354 | assumes "\<And> xs. \<phi> xs"  shows "alw \<phi> xs"
 | 
|  |    355 | proof-
 | 
|  |    356 |   {assume "\<forall> xs. \<phi> xs"
 | 
|  |    357 |    hence ?thesis by coinduct auto
 | 
|  |    358 |   }
 | 
|  |    359 |   thus ?thesis using assms by auto
 | 
|  |    360 | qed
 | 
|  |    361 | 
 | 
|  |    362 | lemma alw_impl_ev_alw:
 | 
|  |    363 | assumes "alw (\<phi> impl ev \<psi>) xs"
 | 
|  |    364 | shows "alw (ev \<phi> impl ev \<psi>) xs"
 | 
|  |    365 | using assms by coinduct (auto elim: alw.cases dest: ev_alw_impl intro: ev.intros)
 | 
|  |    366 | 
 | 
|  |    367 | lemma ev_holds_sset:
 | 
|  |    368 | "ev (holds P) xs \<longleftrightarrow> (\<exists> x \<in> sset xs. P x)" (is "?L \<longleftrightarrow> ?R")
 | 
|  |    369 | proof safe
 | 
|  |    370 |   assume ?L thus ?R by induct (metis holds.simps stream.set_sel(1), metis stl_sset)
 | 
|  |    371 | next
 | 
|  |    372 |   fix x assume "x \<in> sset xs" "P x"
 | 
|  |    373 |   thus ?L by (induct rule: sset_induct) (simp_all add: ev.base ev.step)
 | 
|  |    374 | qed
 | 
|  |    375 | 
 | 
|  |    376 | (* LTL as a program logic: *)
 | 
|  |    377 | lemma alw_invar:
 | 
|  |    378 | assumes "\<phi> xs" and "alw (\<phi> impl nxt \<phi>) xs"
 | 
|  |    379 | shows "alw \<phi> xs"
 | 
|  |    380 | proof-
 | 
|  |    381 |   {assume "\<phi> xs \<and> alw (\<phi> impl nxt \<phi>) xs" hence ?thesis
 | 
|  |    382 |    apply coinduct by(auto elim: alw.cases)
 | 
|  |    383 |   }
 | 
|  |    384 |   thus ?thesis using assms by auto
 | 
|  |    385 | qed
 | 
|  |    386 | 
 | 
|  |    387 | lemma variance:
 | 
|  |    388 | assumes 1: "\<phi> xs" and 2: "alw (\<phi> impl (\<psi> or nxt \<phi>)) xs"
 | 
|  |    389 | shows "(alw \<phi> or ev \<psi>) xs"
 | 
|  |    390 | proof-
 | 
|  |    391 |   {assume "\<not> ev \<psi> xs" hence "alw (not \<psi>) xs" unfolding not_ev[symmetric] .
 | 
|  |    392 |    moreover have "alw (not \<psi> impl (\<phi> impl nxt \<phi>)) xs"
 | 
|  |    393 |    using 2 by coinduct (auto elim: alw.cases)
 | 
|  |    394 |    ultimately have "alw (\<phi> impl nxt \<phi>) xs" by(auto dest: alw_mp)
 | 
|  |    395 |    with 1 have "alw \<phi> xs" by(rule alw_invar)
 | 
|  |    396 |   }
 | 
|  |    397 |   thus ?thesis by blast
 | 
|  |    398 | qed
 | 
|  |    399 | 
 | 
|  |    400 | lemma ev_alw_imp_nxt:
 | 
|  |    401 | assumes e: "ev \<phi> xs" and a: "alw (\<phi> impl (nxt \<phi>)) xs"
 | 
|  |    402 | shows "ev (alw \<phi>) xs"
 | 
|  |    403 | proof-
 | 
|  |    404 |   obtain xl xs1 where xs: "xs = xl @- xs1" and \<phi>: "\<phi> xs1"
 | 
|  |    405 |   using e by (metis ev_imp_shift)
 | 
|  |    406 |   have "\<phi> xs1 \<and> alw (\<phi> impl (nxt \<phi>)) xs1" using a \<phi> unfolding xs by (metis alw_shift)
 | 
|  |    407 |   hence "alw \<phi> xs1" by(coinduct xs1 rule: alw.coinduct) (auto elim: alw.cases)
 | 
|  |    408 |   thus ?thesis unfolding xs by (auto intro: alw_ev_shift)
 | 
|  |    409 | qed
 | 
|  |    410 | 
 | 
|  |    411 | 
 | 
| 59000 |    412 | inductive ev_at :: "('a stream \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> 'a stream \<Rightarrow> bool" for P :: "'a stream \<Rightarrow> bool" where
 | 
|  |    413 |   base: "P \<omega> \<Longrightarrow> ev_at P 0 \<omega>"
 | 
|  |    414 | | step:"\<not> P \<omega> \<Longrightarrow> ev_at P n (stl \<omega>) \<Longrightarrow> ev_at P (Suc n) \<omega>"
 | 
|  |    415 | 
 | 
|  |    416 | inductive_simps ev_at_0[simp]: "ev_at P 0 \<omega>"
 | 
|  |    417 | inductive_simps ev_at_Suc[simp]: "ev_at P (Suc n) \<omega>"
 | 
|  |    418 | 
 | 
|  |    419 | lemma ev_at_imp_snth: "ev_at P n \<omega> \<Longrightarrow> P (sdrop n \<omega>)"
 | 
|  |    420 |   by (induction n arbitrary: \<omega>) auto
 | 
|  |    421 | 
 | 
|  |    422 | lemma ev_at_HLD_imp_snth: "ev_at (HLD X) n \<omega> \<Longrightarrow> \<omega> !! n \<in> X"
 | 
|  |    423 |   by (auto dest!: ev_at_imp_snth simp: HLD_iff)
 | 
|  |    424 | 
 | 
|  |    425 | lemma ev_at_HLD_single_imp_snth: "ev_at (HLD {x}) n \<omega> \<Longrightarrow> \<omega> !! n = x"
 | 
|  |    426 |   by (drule ev_at_HLD_imp_snth) simp
 | 
|  |    427 | 
 | 
|  |    428 | lemma ev_at_unique: "ev_at P n \<omega> \<Longrightarrow> ev_at P m \<omega> \<Longrightarrow> n = m"
 | 
|  |    429 | proof (induction arbitrary: m rule: ev_at.induct)
 | 
|  |    430 |   case (base \<omega>) then show ?case
 | 
|  |    431 |     by (simp add: ev_at.simps[of _ _ \<omega>])
 | 
|  |    432 | next
 | 
|  |    433 |   case (step \<omega> n) from step.prems step.hyps step.IH[of "m - 1"] show ?case
 | 
|  |    434 |     by (auto simp add: ev_at.simps[of _ _ \<omega>])
 | 
|  |    435 | qed
 | 
|  |    436 | 
 | 
|  |    437 | lemma ev_iff_ev_at: "ev P \<omega> \<longleftrightarrow> (\<exists>n. ev_at P n \<omega>)"
 | 
|  |    438 | proof
 | 
|  |    439 |   assume "ev P \<omega>" then show "\<exists>n. ev_at P n \<omega>"
 | 
|  |    440 |     by (induction rule: ev_induct_strong) (auto intro: ev_at.intros)
 | 
|  |    441 | next
 | 
|  |    442 |   assume "\<exists>n. ev_at P n \<omega>"
 | 
|  |    443 |   then obtain n where "ev_at P n \<omega>"
 | 
|  |    444 |     by auto
 | 
|  |    445 |   then show "ev P \<omega>"
 | 
|  |    446 |     by induction auto
 | 
|  |    447 | qed
 | 
|  |    448 | 
 | 
|  |    449 | lemma ev_at_shift: "ev_at (HLD X) i (stake (Suc i) \<omega> @- \<omega>' :: 's stream) \<longleftrightarrow> ev_at (HLD X) i \<omega>"
 | 
|  |    450 |   by (induction i arbitrary: \<omega>) (auto simp: HLD_iff)
 | 
|  |    451 | 
 | 
|  |    452 | lemma ev_iff_ev_at_unqiue: "ev P \<omega> \<longleftrightarrow> (\<exists>!n. ev_at P n \<omega>)"
 | 
|  |    453 |   by (auto intro: ev_at_unique simp: ev_iff_ev_at)
 | 
|  |    454 | 
 | 
|  |    455 | lemma alw_HLD_iff_streams: "alw (HLD X) \<omega> \<longleftrightarrow> \<omega> \<in> streams X"
 | 
|  |    456 | proof
 | 
|  |    457 |   assume "alw (HLD X) \<omega>" then show "\<omega> \<in> streams X"
 | 
|  |    458 |   proof (coinduction arbitrary: \<omega>)
 | 
|  |    459 |     case (streams \<omega>) then show ?case by (cases \<omega>) auto
 | 
|  |    460 |   qed
 | 
|  |    461 | next
 | 
|  |    462 |   assume "\<omega> \<in> streams X" then show "alw (HLD X) \<omega>"
 | 
|  |    463 |   proof (coinduction arbitrary: \<omega>)
 | 
|  |    464 |     case (alw \<omega>) then show ?case by (cases \<omega>) auto
 | 
|  |    465 |   qed
 | 
|  |    466 | qed
 | 
|  |    467 | 
 | 
|  |    468 | lemma not_HLD: "not (HLD X) = HLD (- X)"
 | 
|  |    469 |   by (auto simp: HLD_iff)
 | 
|  |    470 | 
 | 
|  |    471 | lemma not_alw_iff: "\<not> (alw P \<omega>) \<longleftrightarrow> ev (not P) \<omega>"
 | 
|  |    472 |   using not_alw[of P] by (simp add: fun_eq_iff)
 | 
|  |    473 | 
 | 
|  |    474 | lemma not_ev_iff: "\<not> (ev P \<omega>) \<longleftrightarrow> alw (not P) \<omega>"
 | 
|  |    475 |   using not_alw_iff[of "not P" \<omega>, symmetric]  by simp
 | 
|  |    476 | 
 | 
|  |    477 | lemma ev_Stream: "ev P (x ## s) \<longleftrightarrow> P (x ## s) \<or> ev P s"
 | 
|  |    478 |   by (auto elim: ev.cases)
 | 
|  |    479 | 
 | 
|  |    480 | lemma alw_ev_imp_ev_alw:
 | 
|  |    481 |   assumes "alw (ev P) \<omega>" shows "ev (P aand alw (ev P)) \<omega>"
 | 
|  |    482 | proof -
 | 
|  |    483 |   have "ev P \<omega>" using assms by auto
 | 
|  |    484 |   from this assms show ?thesis
 | 
|  |    485 |     by induct auto
 | 
|  |    486 | qed
 | 
|  |    487 | 
 | 
|  |    488 | lemma ev_False: "ev (\<lambda>x. False) \<omega> \<longleftrightarrow> False"
 | 
|  |    489 | proof
 | 
|  |    490 |   assume "ev (\<lambda>x. False) \<omega>" then show False
 | 
|  |    491 |     by induct auto
 | 
|  |    492 | qed auto
 | 
|  |    493 | 
 | 
|  |    494 | lemma alw_False: "alw (\<lambda>x. False) \<omega> \<longleftrightarrow> False"
 | 
|  |    495 |   by auto
 | 
|  |    496 | 
 | 
|  |    497 | lemma ev_iff_sdrop: "ev P \<omega> \<longleftrightarrow> (\<exists>m. P (sdrop m \<omega>))"
 | 
|  |    498 | proof safe
 | 
|  |    499 |   assume "ev P \<omega>" then show "\<exists>m. P (sdrop m \<omega>)"
 | 
|  |    500 |     by (induct rule: ev_induct_strong) (auto intro: exI[of _ 0] exI[of _ "Suc n" for n])
 | 
|  |    501 | next
 | 
|  |    502 |   fix m assume "P (sdrop m \<omega>)" then show "ev P \<omega>"
 | 
|  |    503 |     by (induct m arbitrary: \<omega>) auto
 | 
|  |    504 | qed
 | 
|  |    505 | 
 | 
|  |    506 | lemma alw_iff_sdrop: "alw P \<omega> \<longleftrightarrow> (\<forall>m. P (sdrop m \<omega>))"
 | 
|  |    507 | proof safe
 | 
|  |    508 |   fix m assume "alw P \<omega>" then show "P (sdrop m \<omega>)"
 | 
|  |    509 |     by (induct m arbitrary: \<omega>) auto
 | 
|  |    510 | next
 | 
|  |    511 |   assume "\<forall>m. P (sdrop m \<omega>)" then show "alw P \<omega>"
 | 
|  |    512 |     by (coinduction arbitrary: \<omega>) (auto elim: allE[of _ 0] allE[of _ "Suc n" for n])
 | 
|  |    513 | qed
 | 
|  |    514 | 
 | 
|  |    515 | lemma infinite_iff_alw_ev: "infinite {m. P (sdrop m \<omega>)} \<longleftrightarrow> alw (ev P) \<omega>"
 | 
|  |    516 |   unfolding infinite_nat_iff_unbounded_le alw_iff_sdrop ev_iff_sdrop
 | 
|  |    517 |   by simp (metis le_Suc_ex le_add1)
 | 
|  |    518 | 
 | 
|  |    519 | lemma alw_inv:
 | 
|  |    520 |   assumes stl: "\<And>s. f (stl s) = stl (f s)"
 | 
|  |    521 |   shows "alw P (f s) \<longleftrightarrow> alw (\<lambda>x. P (f x)) s"
 | 
|  |    522 | proof
 | 
|  |    523 |   assume "alw P (f s)" then show "alw (\<lambda>x. P (f x)) s"
 | 
|  |    524 |     by (coinduction arbitrary: s rule: alw_coinduct)
 | 
|  |    525 |        (auto simp: stl)
 | 
|  |    526 | next
 | 
|  |    527 |   assume "alw (\<lambda>x. P (f x)) s" then show "alw P (f s)"
 | 
|  |    528 |     by (coinduction arbitrary: s rule: alw_coinduct) (auto simp: stl[symmetric])
 | 
|  |    529 | qed
 | 
|  |    530 | 
 | 
|  |    531 | lemma ev_inv:
 | 
|  |    532 |   assumes stl: "\<And>s. f (stl s) = stl (f s)"
 | 
|  |    533 |   shows "ev P (f s) \<longleftrightarrow> ev (\<lambda>x. P (f x)) s"
 | 
|  |    534 | proof
 | 
|  |    535 |   assume "ev P (f s)" then show "ev (\<lambda>x. P (f x)) s"
 | 
|  |    536 |     by (induction "f s" arbitrary: s) (auto simp: stl)
 | 
|  |    537 | next
 | 
|  |    538 |   assume "ev (\<lambda>x. P (f x)) s" then show "ev P (f s)"
 | 
|  |    539 |     by induction (auto simp: stl[symmetric])
 | 
|  |    540 | qed
 | 
|  |    541 | 
 | 
|  |    542 | lemma alw_smap: "alw P (smap f s) \<longleftrightarrow> alw (\<lambda>x. P (smap f x)) s"
 | 
|  |    543 |   by (rule alw_inv) simp
 | 
|  |    544 | 
 | 
|  |    545 | lemma ev_smap: "ev P (smap f s) \<longleftrightarrow> ev (\<lambda>x. P (smap f x)) s"
 | 
|  |    546 |   by (rule ev_inv) simp
 | 
|  |    547 | 
 | 
|  |    548 | lemma alw_cong:
 | 
|  |    549 |   assumes P: "alw P \<omega>" and eq: "\<And>\<omega>. P \<omega> \<Longrightarrow> Q1 \<omega> \<longleftrightarrow> Q2 \<omega>"
 | 
|  |    550 |   shows "alw Q1 \<omega> \<longleftrightarrow> alw Q2 \<omega>"
 | 
|  |    551 | proof -
 | 
|  |    552 |   from eq have "(alw P aand Q1) = (alw P aand Q2)" by auto
 | 
|  |    553 |   then have "alw (alw P aand Q1) \<omega> = alw (alw P aand Q2) \<omega>" by auto
 | 
|  |    554 |   with P show "alw Q1 \<omega> \<longleftrightarrow> alw Q2 \<omega>"
 | 
|  |    555 |     by (simp add: alw_aand)
 | 
|  |    556 | qed
 | 
|  |    557 | 
 | 
|  |    558 | lemma ev_cong:
 | 
|  |    559 |   assumes P: "alw P \<omega>" and eq: "\<And>\<omega>. P \<omega> \<Longrightarrow> Q1 \<omega> \<longleftrightarrow> Q2 \<omega>"
 | 
|  |    560 |   shows "ev Q1 \<omega> \<longleftrightarrow> ev Q2 \<omega>"
 | 
|  |    561 | proof -
 | 
|  |    562 |   from P have "alw (\<lambda>xs. Q1 xs \<longrightarrow> Q2 xs) \<omega>" by (rule alw_mono) (simp add: eq)
 | 
|  |    563 |   moreover from P have "alw (\<lambda>xs. Q2 xs \<longrightarrow> Q1 xs) \<omega>" by (rule alw_mono) (simp add: eq)
 | 
|  |    564 |   moreover note ev_alw_impl[of Q1 \<omega> Q2] ev_alw_impl[of Q2 \<omega> Q1]
 | 
|  |    565 |   ultimately show "ev Q1 \<omega> \<longleftrightarrow> ev Q2 \<omega>"
 | 
|  |    566 |     by auto
 | 
|  |    567 | qed
 | 
|  |    568 | 
 | 
|  |    569 | lemma alwD: "alw P x \<Longrightarrow> P x"
 | 
|  |    570 |   by auto
 | 
|  |    571 | 
 | 
|  |    572 | lemma alw_alwD: "alw P \<omega> \<Longrightarrow> alw (alw P) \<omega>"
 | 
|  |    573 |   by simp
 | 
|  |    574 | 
 | 
|  |    575 | lemma alw_ev_stl: "alw (ev P) (stl \<omega>) \<longleftrightarrow> alw (ev P) \<omega>"
 | 
|  |    576 |   by (auto intro: alw.intros)
 | 
|  |    577 | 
 | 
|  |    578 | lemma holds_Stream: "holds P (x ## s) \<longleftrightarrow> P x"
 | 
|  |    579 |   by simp
 | 
|  |    580 | 
 | 
|  |    581 | lemma holds_eq1[simp]: "holds (op = x) = HLD {x}"
 | 
|  |    582 |   by rule (auto simp: HLD_iff)
 | 
|  |    583 | 
 | 
|  |    584 | lemma holds_eq2[simp]: "holds (\<lambda>y. y = x) = HLD {x}"
 | 
|  |    585 |   by rule (auto simp: HLD_iff)
 | 
|  |    586 | 
 | 
|  |    587 | lemma not_holds_eq[simp]: "holds (- op = x) = not (HLD {x})"
 | 
|  |    588 |   by rule (auto simp: HLD_iff)
 | 
|  |    589 | 
 | 
| 60500 |    590 | text \<open>Strong until\<close>
 | 
| 59000 |    591 | 
 | 
|  |    592 | inductive suntil (infix "suntil" 60) for \<phi> \<psi> where
 | 
|  |    593 |   base: "\<psi> \<omega> \<Longrightarrow> (\<phi> suntil \<psi>) \<omega>"
 | 
|  |    594 | | step: "\<phi> \<omega> \<Longrightarrow> (\<phi> suntil \<psi>) (stl \<omega>) \<Longrightarrow> (\<phi> suntil \<psi>) \<omega>"
 | 
|  |    595 | 
 | 
|  |    596 | inductive_simps suntil_Stream: "(\<phi> suntil \<psi>) (x ## s)"
 | 
|  |    597 | 
 | 
|  |    598 | lemma suntil_induct_strong[consumes 1, case_names base step]:
 | 
|  |    599 |   "(\<phi> suntil \<psi>) x \<Longrightarrow>
 | 
|  |    600 |     (\<And>\<omega>. \<psi> \<omega> \<Longrightarrow> P \<omega>) \<Longrightarrow>
 | 
|  |    601 |     (\<And>\<omega>. \<phi> \<omega> \<Longrightarrow> \<not> \<psi> \<omega> \<Longrightarrow> (\<phi> suntil \<psi>) (stl \<omega>) \<Longrightarrow> P (stl \<omega>) \<Longrightarrow> P \<omega>) \<Longrightarrow> P x"
 | 
|  |    602 |   using suntil.induct[of \<phi> \<psi> x P] by blast
 | 
|  |    603 | 
 | 
|  |    604 | lemma ev_suntil: "(\<phi> suntil \<psi>) \<omega> \<Longrightarrow> ev \<psi> \<omega>"
 | 
|  |    605 |   by (induct rule: suntil.induct) (auto intro: ev.intros)
 | 
|  |    606 | 
 | 
|  |    607 | lemma suntil_inv:
 | 
|  |    608 |   assumes stl: "\<And>s. f (stl s) = stl (f s)"
 | 
|  |    609 |   shows "(P suntil Q) (f s) \<longleftrightarrow> ((\<lambda>x. P (f x)) suntil (\<lambda>x. Q (f x))) s"
 | 
|  |    610 | proof
 | 
|  |    611 |   assume "(P suntil Q) (f s)" then show "((\<lambda>x. P (f x)) suntil (\<lambda>x. Q (f x))) s"
 | 
|  |    612 |     by (induction "f s" arbitrary: s) (auto simp: stl intro: suntil.intros)
 | 
|  |    613 | next
 | 
|  |    614 |   assume "((\<lambda>x. P (f x)) suntil (\<lambda>x. Q (f x))) s" then show "(P suntil Q) (f s)"
 | 
|  |    615 |     by induction (auto simp: stl[symmetric] intro: suntil.intros)
 | 
|  |    616 | qed
 | 
|  |    617 | 
 | 
|  |    618 | lemma suntil_smap: "(P suntil Q) (smap f s) \<longleftrightarrow> ((\<lambda>x. P (smap f x)) suntil (\<lambda>x. Q (smap f x))) s"
 | 
|  |    619 |   by (rule suntil_inv) simp
 | 
|  |    620 | 
 | 
|  |    621 | lemma hld_smap: "HLD x (smap f s) = holds (\<lambda>y. f y \<in> x) s"
 | 
|  |    622 |   by (simp add: HLD_def)
 | 
|  |    623 | 
 | 
|  |    624 | lemma suntil_mono:
 | 
|  |    625 |   assumes eq: "\<And>\<omega>. P \<omega> \<Longrightarrow> Q1 \<omega> \<Longrightarrow> Q2 \<omega>" "\<And>\<omega>. P \<omega> \<Longrightarrow> R1 \<omega> \<Longrightarrow> R2 \<omega>"
 | 
|  |    626 |   assumes *: "(Q1 suntil R1) \<omega>" "alw P \<omega>" shows "(Q2 suntil R2) \<omega>"
 | 
|  |    627 |   using * by induct (auto intro: eq suntil.intros)
 | 
|  |    628 | 
 | 
|  |    629 | lemma suntil_cong:
 | 
|  |    630 |   "alw P \<omega> \<Longrightarrow> (\<And>\<omega>. P \<omega> \<Longrightarrow> Q1 \<omega> \<longleftrightarrow> Q2 \<omega>) \<Longrightarrow> (\<And>\<omega>. P \<omega> \<Longrightarrow> R1 \<omega> \<longleftrightarrow> R2 \<omega>) \<Longrightarrow>
 | 
|  |    631 |     (Q1 suntil R1) \<omega> \<longleftrightarrow> (Q2 suntil R2) \<omega>"
 | 
|  |    632 |   using suntil_mono[of P Q1 Q2 R1 R2 \<omega>] suntil_mono[of P Q2 Q1 R2 R1 \<omega>] by auto
 | 
|  |    633 | 
 | 
|  |    634 | lemma ev_suntil_iff: "ev (P suntil Q) \<omega> \<longleftrightarrow> ev Q \<omega>"
 | 
|  |    635 | proof
 | 
|  |    636 |   assume "ev (P suntil Q) \<omega>" then show "ev Q \<omega>"
 | 
|  |    637 |    by induct (auto dest: ev_suntil)
 | 
|  |    638 | next
 | 
|  |    639 |   assume "ev Q \<omega>" then show "ev (P suntil Q) \<omega>"
 | 
|  |    640 |     by induct (auto intro: suntil.intros)
 | 
|  |    641 | qed
 | 
|  |    642 | 
 | 
|  |    643 | lemma true_suntil: "((\<lambda>_. True) suntil P) = ev P"
 | 
|  |    644 |   by (simp add: suntil_def ev_def)
 | 
|  |    645 | 
 | 
|  |    646 | lemma suntil_lfp: "(\<phi> suntil \<psi>) = lfp (\<lambda>P s. \<psi> s \<or> (\<phi> s \<and> P (stl s)))"
 | 
|  |    647 |   by (simp add: suntil_def)
 | 
|  |    648 | 
 | 
|  |    649 | lemma sfilter_P[simp]: "P (shd s) \<Longrightarrow> sfilter P s = shd s ## sfilter P (stl s)"
 | 
|  |    650 |   using sfilter_Stream[of P "shd s" "stl s"] by simp
 | 
|  |    651 | 
 | 
|  |    652 | lemma sfilter_not_P[simp]: "\<not> P (shd s) \<Longrightarrow> sfilter P s = sfilter P (stl s)"
 | 
|  |    653 |   using sfilter_Stream[of P "shd s" "stl s"] by simp
 | 
|  |    654 | 
 | 
|  |    655 | lemma sfilter_eq: 
 | 
|  |    656 |   assumes "ev (holds P) s"
 | 
|  |    657 |   shows "sfilter P s = x ## s' \<longleftrightarrow>
 | 
|  |    658 |     P x \<and> (not (holds P) suntil (HLD {x} aand nxt (\<lambda>s. sfilter P s = s'))) s"
 | 
|  |    659 |   using assms
 | 
|  |    660 |   by (induct rule: ev_induct_strong)
 | 
|  |    661 |      (auto simp add: HLD_iff intro: suntil.intros elim: suntil.cases)
 | 
|  |    662 | 
 | 
|  |    663 | lemma sfilter_streams:
 | 
|  |    664 |   "alw (ev (holds P)) \<omega> \<Longrightarrow> \<omega> \<in> streams A \<Longrightarrow> sfilter P \<omega> \<in> streams {x\<in>A. P x}"
 | 
|  |    665 | proof (coinduction arbitrary: \<omega>)
 | 
|  |    666 |   case (streams \<omega>)
 | 
|  |    667 |   then have "ev (holds P) \<omega>" by blast
 | 
|  |    668 |   from this streams show ?case
 | 
|  |    669 |     by (induct rule: ev_induct_strong) (auto elim: streamsE)
 | 
|  |    670 | qed
 | 
|  |    671 | 
 | 
|  |    672 | lemma alw_sfilter:
 | 
|  |    673 |   assumes *: "alw (ev (holds P)) s"
 | 
|  |    674 |   shows "alw Q (sfilter P s) \<longleftrightarrow> alw (\<lambda>x. Q (sfilter P x)) s"
 | 
|  |    675 | proof
 | 
|  |    676 |   assume "alw Q (sfilter P s)" with * show "alw (\<lambda>x. Q (sfilter P x)) s"
 | 
|  |    677 |   proof (coinduction arbitrary: s rule: alw_coinduct)
 | 
|  |    678 |     case (stl s) 
 | 
|  |    679 |     then have "ev (holds P) s"
 | 
|  |    680 |       by blast
 | 
|  |    681 |     from this stl show ?case
 | 
|  |    682 |       by (induct rule: ev_induct_strong) auto
 | 
|  |    683 |   qed auto
 | 
|  |    684 | next
 | 
|  |    685 |   assume "alw (\<lambda>x. Q (sfilter P x)) s" with * show "alw Q (sfilter P s)"
 | 
|  |    686 |   proof (coinduction arbitrary: s rule: alw_coinduct)
 | 
|  |    687 |     case (stl s) 
 | 
|  |    688 |     then have "ev (holds P) s"
 | 
|  |    689 |       by blast
 | 
|  |    690 |     from this stl show ?case
 | 
|  |    691 |       by (induct rule: ev_induct_strong) auto
 | 
|  |    692 |   qed auto
 | 
|  |    693 | qed
 | 
|  |    694 | 
 | 
|  |    695 | lemma ev_sfilter:
 | 
|  |    696 |   assumes *: "alw (ev (holds P)) s"
 | 
|  |    697 |   shows "ev Q (sfilter P s) \<longleftrightarrow> ev (\<lambda>x. Q (sfilter P x)) s"
 | 
|  |    698 | proof
 | 
|  |    699 |   assume "ev Q (sfilter P s)" from this * show "ev (\<lambda>x. Q (sfilter P x)) s"
 | 
|  |    700 |   proof (induction "sfilter P s" arbitrary: s rule: ev_induct_strong)
 | 
|  |    701 |     case (step s)
 | 
|  |    702 |     then have "ev (holds P) s"
 | 
|  |    703 |       by blast
 | 
|  |    704 |     from this step show ?case
 | 
|  |    705 |       by (induct rule: ev_induct_strong) auto
 | 
|  |    706 |   qed auto
 | 
|  |    707 | next
 | 
|  |    708 |   assume "ev (\<lambda>x. Q (sfilter P x)) s" then show "ev Q (sfilter P s)"
 | 
|  |    709 |   proof (induction rule: ev_induct_strong)
 | 
|  |    710 |     case (step s) then show ?case
 | 
|  |    711 |       by (cases "P (shd s)") auto
 | 
|  |    712 |   qed auto
 | 
|  |    713 | qed
 | 
|  |    714 | 
 | 
|  |    715 | lemma holds_sfilter:
 | 
|  |    716 |   assumes "ev (holds Q) s" shows "holds P (sfilter Q s) \<longleftrightarrow> (not (holds Q) suntil (holds (Q aand P))) s"
 | 
|  |    717 | proof
 | 
|  |    718 |   assume "holds P (sfilter Q s)" with assms show "(not (holds Q) suntil (holds (Q aand P))) s"
 | 
|  |    719 |     by (induct rule: ev_induct_strong) (auto intro: suntil.intros)
 | 
|  |    720 | next
 | 
|  |    721 |   assume "(not (holds Q) suntil (holds (Q aand P))) s" then show "holds P (sfilter Q s)"
 | 
|  |    722 |     by induct auto
 | 
|  |    723 | qed
 | 
|  |    724 | 
 | 
|  |    725 | lemma suntil_aand_nxt:
 | 
|  |    726 |   "(\<phi> suntil (\<phi> aand nxt \<psi>)) \<omega> \<longleftrightarrow> (\<phi> aand nxt (\<phi> suntil \<psi>)) \<omega>"
 | 
|  |    727 | proof
 | 
|  |    728 |   assume "(\<phi> suntil (\<phi> aand nxt \<psi>)) \<omega>" then show "(\<phi> aand nxt (\<phi> suntil \<psi>)) \<omega>"
 | 
|  |    729 |     by induction (auto intro: suntil.intros)
 | 
|  |    730 | next
 | 
|  |    731 |   assume "(\<phi> aand nxt (\<phi> suntil \<psi>)) \<omega>"
 | 
|  |    732 |   then have "(\<phi> suntil \<psi>) (stl \<omega>)" "\<phi> \<omega>"
 | 
|  |    733 |     by auto
 | 
|  |    734 |   then show "(\<phi> suntil (\<phi> aand nxt \<psi>)) \<omega>"
 | 
|  |    735 |     by (induction "stl \<omega>" arbitrary: \<omega>)
 | 
|  |    736 |        (auto elim: suntil.cases intro: suntil.intros)
 | 
|  |    737 | qed
 | 
|  |    738 | 
 | 
|  |    739 | lemma alw_sconst: "alw P (sconst x) \<longleftrightarrow> P (sconst x)"
 | 
|  |    740 | proof
 | 
|  |    741 |   assume "P (sconst x)" then show "alw P (sconst x)"
 | 
|  |    742 |     by coinduction auto
 | 
|  |    743 | qed auto
 | 
|  |    744 | 
 | 
|  |    745 | lemma ev_sconst: "ev P (sconst x) \<longleftrightarrow> P (sconst x)"
 | 
|  |    746 | proof
 | 
|  |    747 |   assume "ev P (sconst x)" then show "P (sconst x)"
 | 
|  |    748 |     by (induction "sconst x") auto
 | 
|  |    749 | qed auto
 | 
|  |    750 | 
 | 
|  |    751 | lemma suntil_sconst: "(\<phi> suntil \<psi>) (sconst x) \<longleftrightarrow> \<psi> (sconst x)"
 | 
|  |    752 | proof
 | 
|  |    753 |   assume "(\<phi> suntil \<psi>) (sconst x)" then show "\<psi> (sconst x)"
 | 
|  |    754 |     by (induction "sconst x") auto
 | 
|  |    755 | qed (auto intro: suntil.intros)
 | 
|  |    756 | 
 | 
|  |    757 | lemma hld_smap': "HLD x (smap f s) = HLD (f -` x) s"
 | 
|  |    758 |   by (simp add: HLD_def)
 | 
| 58627 |    759 | 
 | 
|  |    760 | end |