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