src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy
author hoelzl
Wed, 08 Oct 2014 10:22:00 +0200
changeset 58627 1329679abb2d
child 58881 b9556a055632
permissions -rw-r--r--
add Linear Temporal Logic on Streams
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
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
     6
header {* Linear Temporal Logic on Streams *}
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
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
     9
  imports Stream Sublist
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
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    12
section{* Preliminaries *}
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"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    16
shows "prefixeq xl yl"
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"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    24
shows "prefixeq xl yl \<or> prefixeq yl xl"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    25
using shift_prefix[OF assms] apply(cases "length xl \<le> length yl")
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    26
by (metis, metis assms nat_le_linear shift_prefix)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    27
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    28
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    29
section{* Linear temporal logic *}
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    30
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    31
(* Propositional connectives: *)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    32
abbreviation (input) IMPL (infix "impl" 60)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    33
where "\<phi> impl \<psi> \<equiv> \<lambda> xs. \<phi> xs \<longrightarrow> \<psi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    34
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    35
abbreviation (input) OR (infix "or" 60)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    36
where "\<phi> or \<psi> \<equiv> \<lambda> xs. \<phi> xs \<or> \<psi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    37
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    38
abbreviation (input) AND (infix "aand" 60)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    39
where "\<phi> aand \<psi> \<equiv> \<lambda> xs. \<phi> xs \<and> \<psi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    40
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    41
abbreviation (input) "not \<phi> \<equiv> \<lambda> xs. \<not> \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    42
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    43
abbreviation (input) "true \<equiv> \<lambda> xs. True"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    44
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    45
abbreviation (input) "false \<equiv> \<lambda> xs. False"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    46
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    47
lemma impl_not_or: "\<phi> impl \<psi> = (not \<phi>) or \<psi>"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    48
by blast
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    49
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    50
lemma not_or: "not (\<phi> or \<psi>) = (not \<phi>) aand (not \<psi>)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    51
by blast
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    52
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    53
lemma not_aand: "not (\<phi> aand \<psi>) = (not \<phi>) or (not \<psi>)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    54
by blast
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    55
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    56
lemma non_not[simp]: "not (not \<phi>) = \<phi>" by simp
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    57
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    58
(* Temporal (LTL) connectives: *)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    59
fun holds where "holds P xs \<longleftrightarrow> P (shd xs)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    60
fun nxt where "nxt \<phi> xs = \<phi> (stl xs)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    61
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    62
inductive ev for \<phi> where
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    63
base: "\<phi> xs \<Longrightarrow> ev \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    64
|
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    65
step: "ev \<phi> (stl xs) \<Longrightarrow> ev \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    66
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    67
coinductive alw for \<phi> where
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    68
alw: "\<lbrakk>\<phi> xs; alw \<phi> (stl xs)\<rbrakk> \<Longrightarrow> alw \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    69
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    70
(* weak until: *)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    71
coinductive UNTIL (infix "until" 60) for \<phi> \<psi> where
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    72
base: "\<psi> xs \<Longrightarrow> (\<phi> until \<psi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    73
|
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    74
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
    75
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    76
lemma holds_mono:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    77
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
    78
shows "holds Q xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    79
using assms by auto
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    80
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    81
lemma holds_aand:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    82
"(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
    83
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    84
lemma nxt_mono:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    85
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
    86
shows "nxt \<psi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    87
using assms by auto
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    88
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    89
lemma ev_mono:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    90
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
    91
shows "ev \<psi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    92
using ev by induct (auto intro: ev.intros simp: 0)
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 alw_mono:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    95
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
    96
shows "alw \<psi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    97
using alw by coinduct (auto elim: alw.cases simp: 0)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    98
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
    99
lemma until_monoL:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   100
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
   101
shows "(\<phi>2 until \<psi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   102
using until by coinduct (auto elim: UNTIL.cases simp: 0)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   103
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   104
lemma until_monoR:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   105
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
   106
shows "(\<phi> until \<psi>2) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   107
using until by coinduct (auto elim: UNTIL.cases simp: 0)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   108
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   109
lemma until_mono:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   110
assumes until: "(\<phi>1 until \<psi>1) xs" and
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   111
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
   112
shows "(\<phi>2 until \<psi>2) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   113
using until by coinduct (auto elim: UNTIL.cases simp: 0)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   114
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   115
lemma until_false: "\<phi> until false = alw \<phi>"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   116
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   117
  {fix xs assume "(\<phi> until false) xs" hence "alw \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   118
   apply coinduct by (auto elim: UNTIL.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   119
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   120
  moreover
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   121
  {fix xs assume "alw \<phi> xs" hence "(\<phi> until false) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   122
   apply coinduct by (auto elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   123
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   124
  ultimately show ?thesis by blast
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   125
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   126
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   127
lemma ev_nxt: "ev \<phi> = (\<phi> or nxt (ev \<phi>))"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   128
apply(rule ext) by (metis ev.simps nxt.simps)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   129
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   130
lemma alw_nxt: "alw \<phi> = (\<phi> aand nxt (alw \<phi>))"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   131
apply(rule ext) by (metis alw.simps nxt.simps)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   132
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   133
lemma ev_ev[simp]: "ev (ev \<phi>) = ev \<phi>"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   134
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   135
  {fix xs
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   136
   assume "ev (ev \<phi>) xs" hence "ev \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   137
   by induct (auto intro: ev.intros)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   138
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   139
  thus ?thesis by (auto intro: ev.intros)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   140
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   141
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   142
lemma alw_alw[simp]: "alw (alw \<phi>) = alw \<phi>"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   143
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   144
  {fix xs
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   145
   assume "alw \<phi> xs" hence "alw (alw \<phi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   146
   by coinduct (auto elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   147
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   148
  thus ?thesis by (auto elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   149
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   150
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   151
lemma ev_shift:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   152
assumes "ev \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   153
shows "ev \<phi> (xl @- xs)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   154
using assms by (induct xl) (auto intro: ev.intros)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   155
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   156
lemma ev_imp_shift:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   157
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
   158
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
   159
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   160
lemma alw_ev_shift: "alw \<phi> xs1 \<Longrightarrow> ev (alw \<phi>) (xl @- xs1)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   161
by (auto intro: ev_shift ev.intros)
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 alw_shift:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   164
assumes "alw \<phi> (xl @- xs)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   165
shows "alw \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   166
using assms by (induct xl) (auto elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   167
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   168
lemma ev_ex_nxt:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   169
assumes "ev \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   170
shows "\<exists> n. (nxt ^^ n) \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   171
using assms proof induct
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   172
  case (base xs) thus ?case by (intro exI[of _ 0]) auto
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   173
next
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   174
  case (step xs)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   175
  then obtain n where "(nxt ^^ n) \<phi> (stl xs)" by blast
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   176
  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
   177
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   178
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   179
lemma alw_sdrop:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   180
assumes "alw \<phi> xs"  shows "alw \<phi> (sdrop n xs)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   181
by (metis alw_shift assms stake_sdrop)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   182
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   183
lemma nxt_sdrop: "(nxt ^^ n) \<phi> xs \<longleftrightarrow> \<phi> (sdrop n xs)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   184
by (induct n arbitrary: xs) auto
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   185
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   186
definition "wait \<phi> xs \<equiv> LEAST n. (nxt ^^ n) \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   187
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   188
lemma nxt_wait:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   189
assumes "ev \<phi> xs"  shows "(nxt ^^ (wait \<phi> xs)) \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   190
unfolding wait_def using ev_ex_nxt[OF assms] by(rule LeastI_ex)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   191
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   192
lemma nxt_wait_least:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   193
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
   194
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
   195
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   196
lemma sdrop_wait:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   197
assumes "ev \<phi> xs"  shows "\<phi> (sdrop (wait \<phi> xs) xs)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   198
using nxt_wait[OF assms] unfolding nxt_sdrop .
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   199
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   200
lemma sdrop_wait_least:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   201
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
   202
using assms nxt_wait_least unfolding nxt_sdrop by auto
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   203
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   204
lemma nxt_ev: "(nxt ^^ n) \<phi> xs \<Longrightarrow> ev \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   205
by (induct n arbitrary: xs) (auto intro: ev.intros)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   206
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   207
lemma not_ev: "not (ev \<phi>) = alw (not \<phi>)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   208
proof(rule ext, safe)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   209
  fix xs assume "not (ev \<phi>) xs" thus "alw (not \<phi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   210
  by (coinduct) (auto intro: ev.intros)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   211
next
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   212
  fix xs assume "ev \<phi> xs" and "alw (not \<phi>) xs" thus False
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   213
  by (induct) (auto elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   214
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   215
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   216
lemma not_alw: "not (alw \<phi>) = ev (not \<phi>)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   217
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   218
  have "not (alw \<phi>) = not (alw (not (not \<phi>)))" by simp
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   219
  also have "... = ev (not \<phi>)" unfolding not_ev[symmetric] by simp
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   220
  finally show ?thesis .
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   221
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   222
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   223
lemma not_ev_not[simp]: "not (ev (not \<phi>)) = alw \<phi>"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   224
unfolding not_ev by simp
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 not_alw_not[simp]: "not (alw (not \<phi>)) = ev \<phi>"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   227
unfolding not_alw by simp
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   228
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   229
lemma alw_ev_sdrop:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   230
assumes "alw (ev \<phi>) (sdrop m xs)"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   231
shows "alw (ev \<phi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   232
using assms
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   233
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
   234
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   235
lemma ev_alw_imp_alw_ev:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   236
assumes "ev (alw \<phi>) xs"  shows "alw (ev \<phi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   237
using assms apply induct
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   238
  apply (metis (full_types) alw_mono ev.base)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   239
  by (metis alw alw_nxt ev.step)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   240
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   241
lemma alw_aand: "alw (\<phi> aand \<psi>) = alw \<phi> aand alw \<psi>"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   242
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   243
  {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
   244
   by (auto elim: alw_mono)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   245
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   246
  moreover
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   247
  {fix xs assume "(alw \<phi> aand alw \<psi>) xs" hence "alw (\<phi> aand \<psi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   248
   by coinduct (auto elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   249
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   250
  ultimately show ?thesis by blast
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 ev_or: "ev (\<phi> or \<psi>) = ev \<phi> or ev \<psi>"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   254
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   255
  {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
   256
   by (auto elim: ev_mono)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   257
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   258
  moreover
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   259
  {fix xs assume "ev (\<phi> or \<psi>) xs" hence "(ev \<phi> or ev \<psi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   260
   by induct (auto intro: ev.intros)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   261
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   262
  ultimately show ?thesis by blast
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   263
qed
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_aand:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   266
assumes \<phi>: "ev (alw \<phi>) xs" and \<psi>: "ev (alw \<psi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   267
shows "ev (alw (\<phi> aand \<psi>)) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   268
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   269
  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
   270
  using \<phi> by (metis ev_imp_shift)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   271
  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
   272
  using \<psi> by (metis ev_imp_shift)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   273
  ultimately have 0: "xl @- xs1 = yl @- ys1" by auto
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   274
  hence "prefixeq xl yl \<or> prefixeq yl xl" using shift_prefix_cases by auto
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   275
  thus ?thesis proof
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   276
    assume "prefixeq xl yl"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   277
    then obtain yl1 where yl: "yl = xl @ yl1" by (elim prefixeqE)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   278
    have xs1': "xs1 = yl1 @- ys1" using 0 unfolding yl by simp
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   279
    have "alw \<phi> ys1" using \<phi>\<phi> unfolding xs1' by (metis alw_shift)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   280
    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
   281
    thus ?thesis unfolding xs2 by (auto intro: alw_ev_shift)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   282
  next
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   283
    assume "prefixeq yl xl"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   284
    then obtain xl1 where xl: "xl = yl @ xl1" by (elim prefixeqE)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   285
    have ys1': "ys1 = xl1 @- xs1" using 0 unfolding xl by simp
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   286
    have "alw \<psi> xs1" using \<psi>\<psi> unfolding ys1' by (metis alw_shift)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   287
    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
   288
    thus ?thesis unfolding xs1 by (auto intro: alw_ev_shift)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   289
  qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   290
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   291
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   292
lemma ev_alw_alw_impl:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   293
assumes "ev (alw \<phi>) xs" and "alw (alw \<phi> impl ev \<psi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   294
shows "ev \<psi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   295
using assms by induct (auto intro: ev.intros elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   296
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   297
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
   298
by (metis (full_types) alw_nxt ev_nxt nxt.simps)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   299
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   300
lemma alw_alw_impl_ev:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   301
"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
   302
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   303
  {fix xs assume "?A xs \<and> ev (alw \<phi>) xs" hence "alw (ev \<psi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   304
   apply coinduct using ev_nxt by (auto elim: ev_alw_alw_impl alw.cases intro: ev.intros)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   305
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   306
  moreover
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   307
  {fix xs assume "?B xs" hence "?A xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   308
   apply coinduct by (auto elim: alw.cases intro: ev.intros)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   309
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   310
  ultimately show ?thesis by blast
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   311
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   312
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   313
lemma ev_alw_impl:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   314
assumes "ev \<phi> xs" and "alw (\<phi> impl \<psi>) xs"  shows "ev \<psi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   315
using assms by induct (auto intro: ev.intros elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   316
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   317
lemma ev_alw_impl_ev:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   318
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
   319
using ev_alw_impl[OF assms] by simp
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   320
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   321
lemma alw_mp:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   322
assumes "alw \<phi> xs" and "alw (\<phi> impl \<psi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   323
shows "alw \<psi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   324
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   325
  {assume "alw \<phi> xs \<and> alw (\<phi> impl \<psi>) xs" hence ?thesis
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   326
   apply coinduct by (auto elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   327
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   328
  thus ?thesis using assms by auto
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   329
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   330
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   331
lemma all_imp_alw:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   332
assumes "\<And> xs. \<phi> xs"  shows "alw \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   333
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   334
  {assume "\<forall> xs. \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   335
   hence ?thesis by coinduct auto
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   336
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   337
  thus ?thesis using assms by auto
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   338
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   339
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   340
lemma alw_impl_ev_alw:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   341
assumes "alw (\<phi> impl ev \<psi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   342
shows "alw (ev \<phi> impl ev \<psi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   343
using assms by coinduct (auto elim: alw.cases dest: ev_alw_impl intro: ev.intros)
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_holds_sset:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   346
"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
   347
proof safe
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   348
  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
   349
next
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   350
  fix x assume "x \<in> sset xs" "P x"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   351
  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
   352
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   353
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   354
(* LTL as a program logic: *)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   355
lemma alw_invar:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   356
assumes "\<phi> xs" and "alw (\<phi> impl nxt \<phi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   357
shows "alw \<phi> xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   358
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   359
  {assume "\<phi> xs \<and> alw (\<phi> impl nxt \<phi>) xs" hence ?thesis
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   360
   apply coinduct by(auto elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   361
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   362
  thus ?thesis using assms by auto
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   363
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   364
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   365
lemma variance:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   366
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
   367
shows "(alw \<phi> or ev \<psi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   368
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   369
  {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
   370
   moreover have "alw (not \<psi> impl (\<phi> impl nxt \<phi>)) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   371
   using 2 by coinduct (auto elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   372
   ultimately have "alw (\<phi> impl nxt \<phi>) xs" by(auto dest: alw_mp)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   373
   with 1 have "alw \<phi> xs" by(rule alw_invar)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   374
  }
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   375
  thus ?thesis by blast
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   376
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   377
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   378
lemma ev_alw_imp_nxt:
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   379
assumes e: "ev \<phi> xs" and a: "alw (\<phi> impl (nxt \<phi>)) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   380
shows "ev (alw \<phi>) xs"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   381
proof-
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   382
  obtain xl xs1 where xs: "xs = xl @- xs1" and \<phi>: "\<phi> xs1"
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   383
  using e by (metis ev_imp_shift)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   384
  have "\<phi> xs1 \<and> alw (\<phi> impl (nxt \<phi>)) xs1" using a \<phi> unfolding xs by (metis alw_shift)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   385
  hence "alw \<phi> xs1" by(coinduct xs1 rule: alw.coinduct) (auto elim: alw.cases)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   386
  thus ?thesis unfolding xs by (auto intro: alw_ev_shift)
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   387
qed
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   388
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   389
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   390
1329679abb2d add Linear Temporal Logic on Streams
hoelzl
parents:
diff changeset
   391
end