src/HOL/HOLCF/Library/Stream.thy
author huffman
Sat, 27 Nov 2010 16:08:10 -0800
changeset 41022 0437dbc127b3
parent 40567 src/HOLCF/Library/Stream.thy@73f2b99b549d
child 41661 64cd30d6b0b8
permissions -rw-r--r--
moved directory src/HOLCF to src/HOL/HOLCF;
added HOLCF theories to src/HOL/IsaMakefile;
wenzelm@17291
     1
(*  Title:      HOLCF/ex/Stream.thy
wenzelm@17291
     2
    Author:     Franz Regensburger, David von Oheimb, Borislav Gajanovic
oheimb@2570
     3
*)
oheimb@2570
     4
wenzelm@17291
     5
header {* General Stream domain *}
wenzelm@17291
     6
wenzelm@17291
     7
theory Stream
wenzelm@17291
     8
imports HOLCF Nat_Infinity
wenzelm@17291
     9
begin
oheimb@2570
    10
huffman@40567
    11
default_sort pcpo
huffman@40567
    12
huffman@40567
    13
domain (unsafe) 'a stream = scons (ft::'a) (lazy rt::"'a stream") (infixr "&&" 65)
oheimb@2570
    14
wenzelm@19763
    15
definition
wenzelm@21404
    16
  smap :: "('a \<rightarrow> 'b) \<rightarrow> 'a stream \<rightarrow> 'b stream" where
wenzelm@19763
    17
  "smap = fix\<cdot>(\<Lambda> h f s. case s of x && xs \<Rightarrow> f\<cdot>x && h\<cdot>f\<cdot>xs)"
oheimb@11348
    18
wenzelm@21404
    19
definition
wenzelm@21404
    20
  sfilter :: "('a \<rightarrow> tr) \<rightarrow> 'a stream \<rightarrow> 'a stream" where
wenzelm@19763
    21
  "sfilter = fix\<cdot>(\<Lambda> h p s. case s of x && xs \<Rightarrow>
huffman@40560
    22
                                     If p\<cdot>x then x && h\<cdot>p\<cdot>xs else h\<cdot>p\<cdot>xs)"
wenzelm@19763
    23
wenzelm@21404
    24
definition
wenzelm@21404
    25
  slen :: "'a stream \<Rightarrow> inat"  ("#_" [1000] 1000) where
wenzelm@19763
    26
  "#s = (if stream_finite s then Fin (LEAST n. stream_take n\<cdot>s = s) else \<infinity>)"
oheimb@11348
    27
oheimb@15188
    28
oheimb@15188
    29
(* concatenation *)
oheimb@15188
    30
wenzelm@19763
    31
definition
wenzelm@21404
    32
  i_rt :: "nat => 'a stream => 'a stream" where (* chops the first i elements *)
wenzelm@19763
    33
  "i_rt = (%i s. iterate i$rt$s)"
wenzelm@17291
    34
wenzelm@21404
    35
definition
wenzelm@21404
    36
  i_th :: "nat => 'a stream => 'a" where (* the i-th element *)
wenzelm@19763
    37
  "i_th = (%i s. ft$(i_rt i s))"
oheimb@15188
    38
wenzelm@21404
    39
definition
wenzelm@21404
    40
  sconc :: "'a stream => 'a stream => 'a stream"  (infixr "ooo" 65) where
wenzelm@19763
    41
  "s1 ooo s2 = (case #s1 of
wenzelm@19763
    42
                  Fin n \<Rightarrow> (SOME s. (stream_take n$s=s1) & (i_rt n s = s2))
wenzelm@19763
    43
               | \<infinity>     \<Rightarrow> s1)"
wenzelm@19763
    44
wenzelm@27361
    45
primrec constr_sconc' :: "nat => 'a stream => 'a stream => 'a stream"
wenzelm@27361
    46
where
oheimb@15188
    47
  constr_sconc'_0:   "constr_sconc' 0 s1 s2 = s2"
wenzelm@27361
    48
| constr_sconc'_Suc: "constr_sconc' (Suc n) s1 s2 = ft$s1 &&
oheimb@15188
    49
                                                    constr_sconc' n (rt$s1) s2"
oheimb@15188
    50
wenzelm@19763
    51
definition
wenzelm@21404
    52
  constr_sconc  :: "'a stream => 'a stream => 'a stream" where (* constructive *)
wenzelm@19763
    53
  "constr_sconc s1 s2 = (case #s1 of
wenzelm@19763
    54
                          Fin n \<Rightarrow> constr_sconc' n s1 s2
wenzelm@19763
    55
                        | \<infinity>    \<Rightarrow> s1)"
wenzelm@19763
    56
oheimb@15188
    57
oheimb@15188
    58
(* ----------------------------------------------------------------------- *)
oheimb@15188
    59
(* theorems about scons                                                    *)
oheimb@15188
    60
(* ----------------------------------------------------------------------- *)
oheimb@15188
    61
oheimb@15188
    62
oheimb@15188
    63
section "scons"
oheimb@15188
    64
oheimb@15188
    65
lemma scons_eq_UU: "(a && s = UU) = (a = UU)"
huffman@30913
    66
by simp
oheimb@15188
    67
oheimb@15188
    68
lemma scons_not_empty: "[| a && x = UU; a ~= UU |] ==> R"
huffman@30913
    69
by simp
oheimb@15188
    70
oheimb@15188
    71
lemma stream_exhaust_eq: "(x ~= UU) = (EX a y. a ~= UU &  x = a && y)"
huffman@35781
    72
by (cases x, auto)
oheimb@15188
    73
wenzelm@18109
    74
lemma stream_neq_UU: "x~=UU ==> EX a a_s. x=a&&a_s & a~=UU"
oheimb@15188
    75
by (simp add: stream_exhaust_eq,auto)
oheimb@15188
    76
wenzelm@17291
    77
lemma stream_prefix:
oheimb@15188
    78
  "[| a && s << t; a ~= UU  |] ==> EX b tt. t = b && tt &  b ~= UU &  s << tt"
huffman@35781
    79
by (cases t, auto)
oheimb@15188
    80
wenzelm@17291
    81
lemma stream_prefix':
wenzelm@17291
    82
  "b ~= UU ==> x << b && z =
oheimb@15188
    83
   (x = UU |  (EX a y. x = a && y &  a ~= UU &  a << b &  y << z))"
huffman@35781
    84
by (cases x, auto)
huffman@19550
    85
oheimb@15188
    86
oheimb@15188
    87
(*
oheimb@15188
    88
lemma stream_prefix1: "[| x<<y; xs<<ys |] ==> x&&xs << y&&ys"
oheimb@15188
    89
by (insert stream_prefix' [of y "x&&xs" ys],force)
oheimb@15188
    90
*)
oheimb@15188
    91
wenzelm@17291
    92
lemma stream_flat_prefix:
oheimb@15188
    93
  "[| x && xs << y && ys; (x::'a::flat) ~= UU|] ==> x = y & xs << ys"
oheimb@15188
    94
apply (case_tac "y=UU",auto)
huffman@25920
    95
by (drule ax_flat,simp)
huffman@19550
    96
oheimb@15188
    97
oheimb@15188
    98
oheimb@15188
    99
oheimb@15188
   100
(* ----------------------------------------------------------------------- *)
huffman@40454
   101
(* theorems about stream_case                                              *)
oheimb@15188
   102
(* ----------------------------------------------------------------------- *)
oheimb@15188
   103
huffman@40454
   104
section "stream_case"
oheimb@15188
   105
oheimb@15188
   106
huffman@40454
   107
lemma stream_case_strictf: "stream_case$UU$s=UU"
huffman@35781
   108
by (cases s, auto)
oheimb@15188
   109
oheimb@15188
   110
oheimb@15188
   111
oheimb@15188
   112
(* ----------------------------------------------------------------------- *)
oheimb@15188
   113
(* theorems about ft and rt                                                *)
oheimb@15188
   114
(* ----------------------------------------------------------------------- *)
oheimb@15188
   115
oheimb@15188
   116
oheimb@15188
   117
section "ft & rt"
oheimb@15188
   118
oheimb@15188
   119
oheimb@15188
   120
lemma ft_defin: "s~=UU ==> ft$s~=UU"
huffman@35781
   121
by simp
oheimb@15188
   122
oheimb@15188
   123
lemma rt_strict_rev: "rt$s~=UU ==> s~=UU"
oheimb@15188
   124
by auto
oheimb@15188
   125
oheimb@15188
   126
lemma surjectiv_scons: "(ft$s)&&(rt$s)=s"
huffman@35781
   127
by (cases s, auto)
oheimb@15188
   128
huffman@18075
   129
lemma monofun_rt_mult: "x << s ==> iterate i$rt$x << iterate i$rt$s"
huffman@18075
   130
by (rule monofun_cfun_arg)
oheimb@15188
   131
oheimb@15188
   132
oheimb@15188
   133
oheimb@15188
   134
(* ----------------------------------------------------------------------- *)
oheimb@15188
   135
(* theorems about stream_take                                              *)
oheimb@15188
   136
(* ----------------------------------------------------------------------- *)
oheimb@15188
   137
oheimb@15188
   138
wenzelm@17291
   139
section "stream_take"
oheimb@15188
   140
oheimb@15188
   141
oheimb@15188
   142
lemma stream_reach2: "(LUB i. stream_take i$s) = s"
huffman@35480
   143
by (rule stream.reach)
oheimb@15188
   144
oheimb@15188
   145
lemma chain_stream_take: "chain (%i. stream_take i$s)"
huffman@35781
   146
by simp
oheimb@15188
   147
oheimb@15188
   148
lemma stream_take_prefix [simp]: "stream_take n$s << s"
oheimb@15188
   149
apply (insert stream_reach2 [of s])
oheimb@15188
   150
apply (erule subst) back
oheimb@15188
   151
apply (rule is_ub_thelub)
oheimb@15188
   152
by (simp only: chain_stream_take)
oheimb@15188
   153
wenzelm@17291
   154
lemma stream_take_more [rule_format]:
oheimb@15188
   155
  "ALL x. stream_take n$x = x --> stream_take (Suc n)$x = x"
oheimb@15188
   156
apply (induct_tac n,auto)
oheimb@15188
   157
apply (case_tac "x=UU",auto)
oheimb@15188
   158
by (drule stream_exhaust_eq [THEN iffD1],auto)
oheimb@15188
   159
wenzelm@17291
   160
lemma stream_take_lemma3 [rule_format]:
oheimb@15188
   161
  "ALL x xs. x~=UU --> stream_take n$(x && xs) = x && xs --> stream_take n$xs=xs"
oheimb@15188
   162
apply (induct_tac n,clarsimp)
huffman@16745
   163
(*apply (drule sym, erule scons_not_empty, simp)*)
oheimb@15188
   164
apply (clarify, rule stream_take_more)
oheimb@15188
   165
apply (erule_tac x="x" in allE)
oheimb@15188
   166
by (erule_tac x="xs" in allE,simp)
oheimb@15188
   167
wenzelm@17291
   168
lemma stream_take_lemma4:
oheimb@15188
   169
  "ALL x xs. stream_take n$xs=xs --> stream_take (Suc n)$(x && xs) = x && xs"
oheimb@15188
   170
by auto
oheimb@15188
   171
wenzelm@17291
   172
lemma stream_take_idempotent [rule_format, simp]:
oheimb@15188
   173
 "ALL s. stream_take n$(stream_take n$s) = stream_take n$s"
oheimb@15188
   174
apply (induct_tac n, auto)
oheimb@15188
   175
apply (case_tac "s=UU", auto)
oheimb@15188
   176
by (drule stream_exhaust_eq [THEN iffD1], auto)
oheimb@15188
   177
wenzelm@17291
   178
lemma stream_take_take_Suc [rule_format, simp]:
wenzelm@17291
   179
  "ALL s. stream_take n$(stream_take (Suc n)$s) =
oheimb@15188
   180
                                    stream_take n$s"
oheimb@15188
   181
apply (induct_tac n, auto)
oheimb@15188
   182
apply (case_tac "s=UU", auto)
oheimb@15188
   183
by (drule stream_exhaust_eq [THEN iffD1], auto)
oheimb@15188
   184
wenzelm@17291
   185
lemma mono_stream_take_pred:
oheimb@15188
   186
  "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==>
oheimb@15188
   187
                       stream_take n$s1 << stream_take n$s2"
wenzelm@17291
   188
by (insert monofun_cfun_arg [of "stream_take (Suc n)$s1"
oheimb@15188
   189
  "stream_take (Suc n)$s2" "stream_take n"], auto)
oheimb@15188
   190
(*
wenzelm@17291
   191
lemma mono_stream_take_pred:
oheimb@15188
   192
  "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==>
oheimb@15188
   193
                       stream_take n$s1 << stream_take n$s2"
oheimb@15188
   194
by (drule mono_stream_take [of _ _ n],simp)
oheimb@15188
   195
*)
oheimb@15188
   196
oheimb@15188
   197
lemma stream_take_lemma10 [rule_format]:
wenzelm@17291
   198
  "ALL k<=n. stream_take n$s1 << stream_take n$s2
oheimb@15188
   199
                             --> stream_take k$s1 << stream_take k$s2"
oheimb@15188
   200
apply (induct_tac n,simp,clarsimp)
oheimb@15188
   201
apply (case_tac "k=Suc n",blast)
oheimb@15188
   202
apply (erule_tac x="k" in allE)
oheimb@15188
   203
by (drule mono_stream_take_pred,simp)
oheimb@15188
   204
oheimb@15188
   205
lemma stream_take_le_mono : "k<=n ==> stream_take k$s1 << stream_take n$s1"
oheimb@15188
   206
apply (insert chain_stream_take [of s1])
huffman@25922
   207
by (drule chain_mono,auto)
oheimb@15188
   208
oheimb@15188
   209
lemma mono_stream_take: "s1 << s2 ==> stream_take n$s1 << stream_take n$s2"
oheimb@15188
   210
by (simp add: monofun_cfun_arg)
oheimb@15188
   211
oheimb@15188
   212
(*
oheimb@15188
   213
lemma stream_take_prefix [simp]: "stream_take n$s << s"
oheimb@15188
   214
apply (subgoal_tac "s=(LUB n. stream_take n$s)")
oheimb@15188
   215
 apply (erule ssubst, rule is_ub_thelub)
oheimb@15188
   216
 apply (simp only: chain_stream_take)
oheimb@15188
   217
by (simp only: stream_reach2)
oheimb@15188
   218
*)
oheimb@15188
   219
oheimb@15188
   220
lemma stream_take_take_less:"stream_take k$(stream_take n$s) << stream_take k$s"
oheimb@15188
   221
by (rule monofun_cfun_arg,auto)
oheimb@15188
   222
oheimb@15188
   223
oheimb@15188
   224
(* ------------------------------------------------------------------------- *)
oheimb@15188
   225
(* special induction rules                                                   *)
oheimb@15188
   226
(* ------------------------------------------------------------------------- *)
oheimb@15188
   227
oheimb@15188
   228
oheimb@15188
   229
section "induction"
oheimb@15188
   230
wenzelm@17291
   231
lemma stream_finite_ind:
oheimb@15188
   232
 "[| stream_finite x; P UU; !!a s. [| a ~= UU; P s |] ==> P (a && s) |] ==> P x"
oheimb@15188
   233
apply (simp add: stream.finite_def,auto)
oheimb@15188
   234
apply (erule subst)
huffman@35781
   235
by (drule stream.finite_induct [of P _ x], auto)
oheimb@15188
   236
wenzelm@17291
   237
lemma stream_finite_ind2:
wenzelm@17291
   238
"[| P UU; !! x. x ~= UU ==> P (x && UU); !! y z s. [| y ~= UU; z ~= UU; P s |] ==> P (y && z && s )|] ==>
oheimb@15188
   239
                                 !s. P (stream_take n$s)"
paulson@29792
   240
apply (rule nat_less_induct [of _ n],auto)
paulson@29792
   241
apply (case_tac n, auto) 
paulson@29792
   242
apply (case_tac nat, auto) 
oheimb@15188
   243
apply (case_tac "s=UU",clarsimp)
oheimb@15188
   244
apply (drule stream_exhaust_eq [THEN iffD1],clarsimp)
oheimb@15188
   245
apply (case_tac "s=UU",clarsimp)
oheimb@15188
   246
apply (drule stream_exhaust_eq [THEN iffD1],clarsimp)
oheimb@15188
   247
apply (case_tac "y=UU",clarsimp)
oheimb@15188
   248
by (drule stream_exhaust_eq [THEN iffD1],clarsimp)
oheimb@15188
   249
wenzelm@17291
   250
lemma stream_ind2:
oheimb@15188
   251
"[| adm P; P UU; !!a. a ~= UU ==> P (a && UU); !!a b s. [| a ~= UU; b ~= UU; P s |] ==> P (a && b && s) |] ==> P x"
oheimb@15188
   252
apply (insert stream.reach [of x],erule subst)
huffman@35480
   253
apply (erule admD, rule chain_stream_take)
oheimb@15188
   254
apply (insert stream_finite_ind2 [of P])
huffman@35480
   255
by simp
oheimb@15188
   256
oheimb@15188
   257
oheimb@15188
   258
oheimb@15188
   259
(* ----------------------------------------------------------------------- *)
oheimb@15188
   260
(* simplify use of coinduction                                             *)
oheimb@15188
   261
(* ----------------------------------------------------------------------- *)
oheimb@15188
   262
oheimb@15188
   263
oheimb@15188
   264
section "coinduction"
oheimb@15188
   265
oheimb@15188
   266
lemma stream_coind_lemma2: "!s1 s2. R s1 s2 --> ft$s1 = ft$s2 &  R (rt$s1) (rt$s2) ==> stream_bisim R"
huffman@30803
   267
 apply (simp add: stream.bisim_def,clarsimp)
huffman@35483
   268
 apply (drule spec, drule spec, drule (1) mp)
huffman@35483
   269
 apply (case_tac "x", simp)
huffman@40206
   270
 apply (case_tac "y", simp)
oheimb@15188
   271
by auto
oheimb@15188
   272
oheimb@15188
   273
oheimb@15188
   274
oheimb@15188
   275
(* ----------------------------------------------------------------------- *)
oheimb@15188
   276
(* theorems about stream_finite                                            *)
oheimb@15188
   277
(* ----------------------------------------------------------------------- *)
oheimb@15188
   278
oheimb@15188
   279
oheimb@15188
   280
section "stream_finite"
oheimb@15188
   281
oheimb@15188
   282
lemma stream_finite_UU [simp]: "stream_finite UU"
oheimb@15188
   283
by (simp add: stream.finite_def)
oheimb@15188
   284
oheimb@15188
   285
lemma stream_finite_UU_rev: "~  stream_finite s ==> s ~= UU"
oheimb@15188
   286
by (auto simp add: stream.finite_def)
oheimb@15188
   287
oheimb@15188
   288
lemma stream_finite_lemma1: "stream_finite xs ==> stream_finite (x && xs)"
oheimb@15188
   289
apply (simp add: stream.finite_def,auto)
huffman@35557
   290
apply (rule_tac x="Suc n" in exI)
oheimb@15188
   291
by (simp add: stream_take_lemma4)
oheimb@15188
   292
oheimb@15188
   293
lemma stream_finite_lemma2: "[| x ~= UU; stream_finite (x && xs) |] ==> stream_finite xs"
oheimb@15188
   294
apply (simp add: stream.finite_def, auto)
huffman@35557
   295
apply (rule_tac x="n" in exI)
oheimb@15188
   296
by (erule stream_take_lemma3,simp)
oheimb@15188
   297
oheimb@15188
   298
lemma stream_finite_rt_eq: "stream_finite (rt$s) = stream_finite s"
huffman@35781
   299
apply (cases s, auto)
oheimb@15188
   300
apply (rule stream_finite_lemma1, simp)
oheimb@15188
   301
by (rule stream_finite_lemma2,simp)
oheimb@15188
   302
oheimb@15188
   303
lemma stream_finite_less: "stream_finite s ==> !t. t<<s --> stream_finite t"
huffman@19440
   304
apply (erule stream_finite_ind [of s], auto)
oheimb@15188
   305
apply (case_tac "t=UU", auto)
oheimb@15188
   306
apply (drule stream_exhaust_eq [THEN iffD1],auto)
oheimb@15188
   307
apply (erule_tac x="y" in allE, simp)
oheimb@15188
   308
by (rule stream_finite_lemma1, simp)
oheimb@15188
   309
oheimb@15188
   310
lemma stream_take_finite [simp]: "stream_finite (stream_take n$s)"
oheimb@15188
   311
apply (simp add: stream.finite_def)
oheimb@15188
   312
by (rule_tac x="n" in exI,simp)
oheimb@15188
   313
oheimb@15188
   314
lemma adm_not_stream_finite: "adm (%x. ~ stream_finite x)"
huffman@25833
   315
apply (rule adm_upward)
huffman@25833
   316
apply (erule contrapos_nn)
huffman@25833
   317
apply (erule (1) stream_finite_less [rule_format])
huffman@25833
   318
done
oheimb@15188
   319
oheimb@15188
   320
oheimb@15188
   321
oheimb@15188
   322
(* ----------------------------------------------------------------------- *)
oheimb@15188
   323
(* theorems about stream length                                            *)
oheimb@15188
   324
(* ----------------------------------------------------------------------- *)
oheimb@15188
   325
oheimb@15188
   326
oheimb@15188
   327
section "slen"
oheimb@15188
   328
oheimb@15188
   329
lemma slen_empty [simp]: "#\<bottom> = 0"
haftmann@27111
   330
by (simp add: slen_def stream.finite_def zero_inat_def Least_equality)
oheimb@15188
   331
oheimb@15188
   332
lemma slen_scons [simp]: "x ~= \<bottom> ==> #(x&&xs) = iSuc (#xs)"
oheimb@15188
   333
apply (case_tac "stream_finite (x && xs)")
oheimb@15188
   334
apply (simp add: slen_def, auto)
haftmann@27111
   335
apply (simp add: stream.finite_def, auto simp add: iSuc_Fin)
haftmann@27111
   336
apply (rule Least_Suc2, auto)
huffman@16745
   337
(*apply (drule sym)*)
huffman@16745
   338
(*apply (drule sym scons_eq_UU [THEN iffD1],simp)*)
oheimb@15188
   339
apply (erule stream_finite_lemma2, simp)
oheimb@15188
   340
apply (simp add: slen_def, auto)
oheimb@15188
   341
by (drule stream_finite_lemma1,auto)
oheimb@15188
   342
oheimb@15188
   343
lemma slen_less_1_eq: "(#x < Fin (Suc 0)) = (x = \<bottom>)"
huffman@35781
   344
by (cases x, auto simp add: Fin_0 iSuc_Fin[THEN sym])
oheimb@15188
   345
oheimb@15188
   346
lemma slen_empty_eq: "(#x = 0) = (x = \<bottom>)"
huffman@35781
   347
by (cases x, auto)
oheimb@15188
   348
oheimb@15188
   349
lemma slen_scons_eq: "(Fin (Suc n) < #x) = (? a y. x = a && y &  a ~= \<bottom> &  Fin n < #y)"
oheimb@15188
   350
apply (auto, case_tac "x=UU",auto)
oheimb@15188
   351
apply (drule stream_exhaust_eq [THEN iffD1], auto)
haftmann@27111
   352
apply (case_tac "#y") apply simp_all
haftmann@27111
   353
apply (case_tac "#y") apply simp_all
haftmann@27111
   354
done
oheimb@15188
   355
oheimb@15188
   356
lemma slen_iSuc: "#x = iSuc n --> (? a y. x = a&&y &  a ~= \<bottom> &  #y = n)"
huffman@35781
   357
by (cases x, auto)
oheimb@15188
   358
oheimb@15188
   359
lemma slen_stream_take_finite [simp]: "#(stream_take n$s) ~= \<infinity>"
oheimb@15188
   360
by (simp add: slen_def)
oheimb@15188
   361
oheimb@15188
   362
lemma slen_scons_eq_rev: "(#x < Fin (Suc (Suc n))) = (!a y. x ~= a && y |  a = \<bottom> |  #y < Fin (Suc n))"
huffman@35781
   363
 apply (cases x, auto)
huffman@30803
   364
   apply (simp add: zero_inat_def)
huffman@35429
   365
  apply (case_tac "#stream") apply (simp_all add: iSuc_Fin)
huffman@35429
   366
 apply (case_tac "#stream") apply (simp_all add: iSuc_Fin)
haftmann@27111
   367
done
oheimb@15188
   368
wenzelm@17291
   369
lemma slen_take_lemma4 [rule_format]:
oheimb@15188
   370
  "!s. stream_take n$s ~= s --> #(stream_take n$s) = Fin n"
haftmann@27111
   371
apply (induct n, auto simp add: Fin_0)
haftmann@27111
   372
apply (case_tac "s=UU", simp)
haftmann@27111
   373
by (drule stream_exhaust_eq [THEN iffD1], auto simp add: iSuc_Fin)
oheimb@15188
   374
oheimb@15188
   375
(*
wenzelm@17291
   376
lemma stream_take_idempotent [simp]:
oheimb@15188
   377
 "stream_take n$(stream_take n$s) = stream_take n$s"
oheimb@15188
   378
apply (case_tac "stream_take n$s = s")
oheimb@15188
   379
apply (auto,insert slen_take_lemma4 [of n s]);
oheimb@15188
   380
by (auto,insert slen_take_lemma1 [of "stream_take n$s" n],simp)
oheimb@15188
   381
wenzelm@17291
   382
lemma stream_take_take_Suc [simp]: "stream_take n$(stream_take (Suc n)$s) =
oheimb@15188
   383
                                    stream_take n$s"
oheimb@15188
   384
apply (simp add: po_eq_conv,auto)
oheimb@15188
   385
 apply (simp add: stream_take_take_less)
oheimb@15188
   386
apply (subgoal_tac "stream_take n$s = stream_take n$(stream_take n$s)")
oheimb@15188
   387
 apply (erule ssubst)
oheimb@15188
   388
 apply (rule_tac monofun_cfun_arg)
oheimb@15188
   389
 apply (insert chain_stream_take [of s])
oheimb@15188
   390
by (simp add: chain_def,simp)
oheimb@15188
   391
*)
oheimb@15188
   392
oheimb@15188
   393
lemma slen_take_eq: "ALL x. (Fin n < #x) = (stream_take n\<cdot>x ~= x)"
oheimb@15188
   394
apply (induct_tac n, auto)
oheimb@15188
   395
apply (simp add: Fin_0, clarsimp)
oheimb@15188
   396
apply (drule not_sym)
oheimb@15188
   397
apply (drule slen_empty_eq [THEN iffD1], simp)
oheimb@15188
   398
apply (case_tac "x=UU", simp)
oheimb@15188
   399
apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)
oheimb@15188
   400
apply (erule_tac x="y" in allE, auto)
haftmann@27111
   401
apply (simp_all add: not_less iSuc_Fin)
haftmann@27111
   402
apply (case_tac "#y") apply simp_all
oheimb@15188
   403
apply (case_tac "x=UU", simp)
oheimb@15188
   404
apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)
oheimb@15188
   405
apply (erule_tac x="y" in allE, simp)
haftmann@27111
   406
apply (case_tac "#y") by simp_all
oheimb@15188
   407
oheimb@15188
   408
lemma slen_take_eq_rev: "(#x <= Fin n) = (stream_take n\<cdot>x = x)"
huffman@26102
   409
by (simp add: linorder_not_less [symmetric] slen_take_eq)
oheimb@15188
   410
oheimb@15188
   411
lemma slen_take_lemma1: "#x = Fin n ==> stream_take n\<cdot>x = x"
oheimb@15188
   412
by (rule slen_take_eq_rev [THEN iffD1], auto)
oheimb@15188
   413
oheimb@15188
   414
lemma slen_rt_mono: "#s2 <= #s1 ==> #(rt$s2) <= #(rt$s1)"
huffman@35781
   415
apply (cases s1)
huffman@35781
   416
 by (cases s2, simp+)+
oheimb@15188
   417
wenzelm@17291
   418
lemma slen_take_lemma5: "#(stream_take n$s) <= Fin n"
oheimb@15188
   419
apply (case_tac "stream_take n$s = s")
oheimb@15188
   420
 apply (simp add: slen_take_eq_rev)
oheimb@15188
   421
by (simp add: slen_take_lemma4)
oheimb@15188
   422
oheimb@15188
   423
lemma slen_take_lemma2: "!x. ~stream_finite x --> #(stream_take i\<cdot>x) = Fin i"
oheimb@15188
   424
apply (simp add: stream.finite_def, auto)
oheimb@15188
   425
by (simp add: slen_take_lemma4)
oheimb@15188
   426
oheimb@15188
   427
lemma slen_infinite: "stream_finite x = (#x ~= Infty)"
oheimb@15188
   428
by (simp add: slen_def)
oheimb@15188
   429
oheimb@15188
   430
lemma slen_mono_lemma: "stream_finite s ==> ALL t. s << t --> #s <= #t"
oheimb@15188
   431
apply (erule stream_finite_ind [of s], auto)
oheimb@15188
   432
apply (case_tac "t=UU", auto)
oheimb@15188
   433
apply (drule stream_exhaust_eq [THEN iffD1], auto)
huffman@30803
   434
done
oheimb@15188
   435
oheimb@15188
   436
lemma slen_mono: "s << t ==> #s <= #t"
oheimb@15188
   437
apply (case_tac "stream_finite t")
wenzelm@17291
   438
apply (frule stream_finite_less)
oheimb@15188
   439
apply (erule_tac x="s" in allE, simp)
oheimb@15188
   440
apply (drule slen_mono_lemma, auto)
oheimb@15188
   441
by (simp add: slen_def)
oheimb@15188
   442
huffman@18075
   443
lemma iterate_lemma: "F$(iterate n$F$x) = iterate n$F$(F$x)"
oheimb@15188
   444
by (insert iterate_Suc2 [of n F x], auto)
oheimb@15188
   445
huffman@18075
   446
lemma slen_rt_mult [rule_format]: "!x. Fin (i + j) <= #x --> Fin j <= #(iterate i$rt$x)"
haftmann@27111
   447
apply (induct i, auto)
haftmann@27111
   448
apply (case_tac "x=UU", auto simp add: zero_inat_def)
oheimb@15188
   449
apply (drule stream_exhaust_eq [THEN iffD1], auto)
oheimb@15188
   450
apply (erule_tac x="y" in allE, auto)
haftmann@27111
   451
apply (simp add: not_le) apply (case_tac "#y") apply (simp_all add: iSuc_Fin)
oheimb@15188
   452
by (simp add: iterate_lemma)
oheimb@15188
   453
wenzelm@17291
   454
lemma slen_take_lemma3 [rule_format]:
oheimb@15188
   455
  "!(x::'a::flat stream) y. Fin n <= #x --> x << y --> stream_take n\<cdot>x = stream_take n\<cdot>y"
oheimb@15188
   456
apply (induct_tac n, auto)
oheimb@15188
   457
apply (case_tac "x=UU", auto)
haftmann@27111
   458
apply (simp add: zero_inat_def)
oheimb@15188
   459
apply (simp add: Suc_ile_eq)
oheimb@15188
   460
apply (case_tac "y=UU", clarsimp)
oheimb@15188
   461
apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)+
oheimb@15188
   462
apply (erule_tac x="ya" in allE, simp)
huffman@25920
   463
by (drule ax_flat, simp)
oheimb@15188
   464
wenzelm@17291
   465
lemma slen_strict_mono_lemma:
oheimb@15188
   466
  "stream_finite t ==> !s. #(s::'a::flat stream) = #t &  s << t --> s = t"
oheimb@15188
   467
apply (erule stream_finite_ind, auto)
oheimb@15188
   468
apply (case_tac "sa=UU", auto)
oheimb@15188
   469
apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)
huffman@25920
   470
by (drule ax_flat, simp)
oheimb@15188
   471
oheimb@15188
   472
lemma slen_strict_mono: "[|stream_finite t; s ~= t; s << (t::'a::flat stream) |] ==> #s < #t"
haftmann@27111
   473
by (auto simp add: slen_mono less_le dest: slen_strict_mono_lemma)
oheimb@15188
   474
wenzelm@17291
   475
lemma stream_take_Suc_neq: "stream_take (Suc n)$s ~=s ==>
oheimb@15188
   476
                     stream_take n$s ~= stream_take (Suc n)$s"
oheimb@15188
   477
apply auto
oheimb@15188
   478
apply (subgoal_tac "stream_take n$s ~=s")
oheimb@15188
   479
 apply (insert slen_take_lemma4 [of n s],auto)
huffman@35781
   480
apply (cases s, simp)
haftmann@27111
   481
by (simp add: slen_take_lemma4 iSuc_Fin)
oheimb@15188
   482
oheimb@15188
   483
(* ----------------------------------------------------------------------- *)
oheimb@15188
   484
(* theorems about smap                                                     *)
oheimb@15188
   485
(* ----------------------------------------------------------------------- *)
oheimb@15188
   486
oheimb@15188
   487
oheimb@15188
   488
section "smap"
oheimb@15188
   489
oheimb@15188
   490
lemma smap_unfold: "smap = (\<Lambda> f t. case t of x&&xs \<Rightarrow> f$x && smap$f$xs)"
huffman@29530
   491
by (insert smap_def [where 'a='a and 'b='b, THEN eq_reflection, THEN fix_eq2], auto)
oheimb@15188
   492
oheimb@15188
   493
lemma smap_empty [simp]: "smap\<cdot>f\<cdot>\<bottom> = \<bottom>"
oheimb@15188
   494
by (subst smap_unfold, simp)
oheimb@15188
   495
oheimb@15188
   496
lemma smap_scons [simp]: "x~=\<bottom> ==> smap\<cdot>f\<cdot>(x&&xs) = (f\<cdot>x)&&(smap\<cdot>f\<cdot>xs)"
oheimb@15188
   497
by (subst smap_unfold, force)
oheimb@15188
   498
oheimb@15188
   499
oheimb@15188
   500
oheimb@15188
   501
(* ----------------------------------------------------------------------- *)
oheimb@15188
   502
(* theorems about sfilter                                                  *)
oheimb@15188
   503
(* ----------------------------------------------------------------------- *)
oheimb@15188
   504
oheimb@15188
   505
section "sfilter"
oheimb@15188
   506
wenzelm@17291
   507
lemma sfilter_unfold:
oheimb@15188
   508
 "sfilter = (\<Lambda> p s. case s of x && xs \<Rightarrow>
huffman@40560
   509
  If p\<cdot>x then x && sfilter\<cdot>p\<cdot>xs else sfilter\<cdot>p\<cdot>xs)"
huffman@29530
   510
by (insert sfilter_def [where 'a='a, THEN eq_reflection, THEN fix_eq2], auto)
oheimb@15188
   511
oheimb@15188
   512
lemma strict_sfilter: "sfilter\<cdot>\<bottom> = \<bottom>"
huffman@40183
   513
apply (rule cfun_eqI)
oheimb@15188
   514
apply (subst sfilter_unfold, auto)
oheimb@15188
   515
apply (case_tac "x=UU", auto)
oheimb@15188
   516
by (drule stream_exhaust_eq [THEN iffD1], auto)
oheimb@15188
   517
oheimb@15188
   518
lemma sfilter_empty [simp]: "sfilter\<cdot>f\<cdot>\<bottom> = \<bottom>"
oheimb@15188
   519
by (subst sfilter_unfold, force)
oheimb@15188
   520
wenzelm@17291
   521
lemma sfilter_scons [simp]:
wenzelm@17291
   522
  "x ~= \<bottom> ==> sfilter\<cdot>f\<cdot>(x && xs) =
huffman@40560
   523
                           If f\<cdot>x then x && sfilter\<cdot>f\<cdot>xs else sfilter\<cdot>f\<cdot>xs"
oheimb@15188
   524
by (subst sfilter_unfold, force)
oheimb@15188
   525
oheimb@15188
   526
oheimb@15188
   527
(* ----------------------------------------------------------------------- *)
oheimb@15188
   528
   section "i_rt"
oheimb@15188
   529
(* ----------------------------------------------------------------------- *)
oheimb@15188
   530
oheimb@15188
   531
lemma i_rt_UU [simp]: "i_rt n UU = UU"
haftmann@34928
   532
  by (induct n) (simp_all add: i_rt_def)
oheimb@15188
   533
oheimb@15188
   534
lemma i_rt_0 [simp]: "i_rt 0 s = s"
oheimb@15188
   535
by (simp add: i_rt_def)
oheimb@15188
   536
oheimb@15188
   537
lemma i_rt_Suc [simp]: "a ~= UU ==> i_rt (Suc n) (a&&s) = i_rt n s"
oheimb@15188
   538
by (simp add: i_rt_def iterate_Suc2 del: iterate_Suc)
oheimb@15188
   539
oheimb@15188
   540
lemma i_rt_Suc_forw: "i_rt (Suc n) s = i_rt n (rt$s)"
oheimb@15188
   541
by (simp only: i_rt_def iterate_Suc2)
oheimb@15188
   542
oheimb@15188
   543
lemma i_rt_Suc_back:"i_rt (Suc n) s = rt$(i_rt n s)"
oheimb@15188
   544
by (simp only: i_rt_def,auto)
oheimb@15188
   545
oheimb@15188
   546
lemma i_rt_mono: "x << s ==> i_rt n x  << i_rt n s"
oheimb@15188
   547
by (simp add: i_rt_def monofun_rt_mult)
oheimb@15188
   548
oheimb@15188
   549
lemma i_rt_ij_lemma: "Fin (i + j) <= #x ==> Fin j <= #(i_rt i x)"
oheimb@15188
   550
by (simp add: i_rt_def slen_rt_mult)
oheimb@15188
   551
oheimb@15188
   552
lemma slen_i_rt_mono: "#s2 <= #s1 ==> #(i_rt n s2) <= #(i_rt n s1)"
oheimb@15188
   553
apply (induct_tac n,auto)
oheimb@15188
   554
apply (simp add: i_rt_Suc_back)
oheimb@15188
   555
by (drule slen_rt_mono,simp)
oheimb@15188
   556
oheimb@15188
   557
lemma i_rt_take_lemma1 [rule_format]: "ALL s. i_rt n (stream_take n$s) = UU"
wenzelm@17291
   558
apply (induct_tac n)
oheimb@15188
   559
 apply (simp add: i_rt_Suc_back,auto)
oheimb@15188
   560
apply (case_tac "s=UU",auto)
oheimb@15188
   561
by (drule stream_exhaust_eq [THEN iffD1],auto)
oheimb@15188
   562
oheimb@15188
   563
lemma i_rt_slen: "(i_rt n s = UU) = (stream_take n$s = s)"
oheimb@15188
   564
apply auto
wenzelm@17291
   565
 apply (insert i_rt_ij_lemma [of n "Suc 0" s])
oheimb@15188
   566
 apply (subgoal_tac "#(i_rt n s)=0")
oheimb@15188
   567
  apply (case_tac "stream_take n$s = s",simp+)
oheimb@15188
   568
  apply (insert slen_take_eq [rule_format,of n s],simp)
haftmann@27111
   569
  apply (cases "#s") apply (simp_all add: zero_inat_def)
haftmann@27111
   570
  apply (simp add: slen_take_eq)
haftmann@27111
   571
  apply (cases "#s")
haftmann@27111
   572
  using i_rt_take_lemma1 [of n s]
haftmann@27111
   573
  apply (simp_all add: zero_inat_def)
haftmann@27111
   574
  done
oheimb@15188
   575
oheimb@15188
   576
lemma i_rt_lemma_slen: "#s=Fin n ==> i_rt n s = UU"
oheimb@15188
   577
by (simp add: i_rt_slen slen_take_lemma1)
oheimb@15188
   578
oheimb@15188
   579
lemma stream_finite_i_rt [simp]: "stream_finite (i_rt n s) = stream_finite s"
oheimb@15188
   580
apply (induct_tac n, auto)
huffman@35781
   581
 apply (cases s, auto simp del: i_rt_Suc)
oheimb@15188
   582
by (simp add: i_rt_Suc_back stream_finite_rt_eq)+
oheimb@15188
   583
oheimb@15188
   584
lemma take_i_rt_len_lemma: "ALL sl x j t. Fin sl = #x & n <= sl &
wenzelm@17291
   585
                            #(stream_take n$x) = Fin t & #(i_rt n x)= Fin j
oheimb@15188
   586
                                              --> Fin (j + t) = #x"
haftmann@27111
   587
apply (induct n, auto)
haftmann@27111
   588
 apply (simp add: zero_inat_def)
oheimb@15188
   589
apply (case_tac "x=UU",auto)
haftmann@27111
   590
 apply (simp add: zero_inat_def)
oheimb@15188
   591
apply (drule stream_exhaust_eq [THEN iffD1],clarsimp)
oheimb@15188
   592
apply (subgoal_tac "EX k. Fin k = #y",clarify)
oheimb@15188
   593
 apply (erule_tac x="k" in allE)
oheimb@15188
   594
 apply (erule_tac x="y" in allE,auto)
oheimb@15188
   595
 apply (erule_tac x="THE p. Suc p = t" in allE,auto)
haftmann@27111
   596
   apply (simp add: iSuc_def split: inat.splits)
haftmann@27111
   597
  apply (simp add: iSuc_def split: inat.splits)
oheimb@15188
   598
  apply (simp only: the_equality)
haftmann@27111
   599
 apply (simp add: iSuc_def split: inat.splits)
oheimb@15188
   600
 apply force
haftmann@27111
   601
apply (simp add: iSuc_def split: inat.splits)
haftmann@27111
   602
done
oheimb@15188
   603
wenzelm@17291
   604
lemma take_i_rt_len:
oheimb@15188
   605
"[| Fin sl = #x; n <= sl; #(stream_take n$x) = Fin t; #(i_rt n x) = Fin j |] ==>
oheimb@15188
   606
    Fin (j + t) = #x"
oheimb@15188
   607
by (blast intro: take_i_rt_len_lemma [rule_format])
oheimb@15188
   608
oheimb@15188
   609
oheimb@15188
   610
(* ----------------------------------------------------------------------- *)
oheimb@15188
   611
   section "i_th"
oheimb@15188
   612
(* ----------------------------------------------------------------------- *)
oheimb@15188
   613
oheimb@15188
   614
lemma i_th_i_rt_step:
wenzelm@17291
   615
"[| i_th n s1 << i_th n s2; i_rt (Suc n) s1 << i_rt (Suc n) s2 |] ==>
oheimb@15188
   616
   i_rt n s1 << i_rt n s2"
oheimb@15188
   617
apply (simp add: i_th_def i_rt_Suc_back)
huffman@35781
   618
apply (cases "i_rt n s1", simp)
huffman@35781
   619
apply (cases "i_rt n s2", auto)
huffman@30803
   620
done
oheimb@15188
   621
wenzelm@17291
   622
lemma i_th_stream_take_Suc [rule_format]:
oheimb@15188
   623
 "ALL s. i_th n (stream_take (Suc n)$s) = i_th n s"
oheimb@15188
   624
apply (induct_tac n,auto)
oheimb@15188
   625
 apply (simp add: i_th_def)
oheimb@15188
   626
 apply (case_tac "s=UU",auto)
oheimb@15188
   627
 apply (drule stream_exhaust_eq [THEN iffD1],auto)
oheimb@15188
   628
apply (case_tac "s=UU",simp add: i_th_def)
oheimb@15188
   629
apply (drule stream_exhaust_eq [THEN iffD1],auto)
oheimb@15188
   630
by (simp add: i_th_def i_rt_Suc_forw)
oheimb@15188
   631
oheimb@15188
   632
lemma i_th_last: "i_th n s && UU = i_rt n (stream_take (Suc n)$s)"
oheimb@15188
   633
apply (insert surjectiv_scons [of "i_rt n (stream_take (Suc n)$s)"])
oheimb@15188
   634
apply (rule i_th_stream_take_Suc [THEN subst])
oheimb@15188
   635
apply (simp add: i_th_def  i_rt_Suc_back [symmetric])
oheimb@15188
   636
by (simp add: i_rt_take_lemma1)
oheimb@15188
   637
wenzelm@17291
   638
lemma i_th_last_eq:
oheimb@15188
   639
"i_th n s1 = i_th n s2 ==> i_rt n (stream_take (Suc n)$s1) = i_rt n (stream_take (Suc n)$s2)"
oheimb@15188
   640
apply (insert i_th_last [of n s1])
oheimb@15188
   641
apply (insert i_th_last [of n s2])
oheimb@15188
   642
by auto
oheimb@15188
   643
oheimb@15188
   644
lemma i_th_prefix_lemma:
wenzelm@17291
   645
"[| k <= n; stream_take (Suc n)$s1 << stream_take (Suc n)$s2 |] ==>
oheimb@15188
   646
    i_th k s1 << i_th k s2"
oheimb@15188
   647
apply (insert i_th_stream_take_Suc [of k s1, THEN sym])
oheimb@15188
   648
apply (insert i_th_stream_take_Suc [of k s2, THEN sym],auto)
oheimb@15188
   649
apply (simp add: i_th_def)
oheimb@15188
   650
apply (rule monofun_cfun, auto)
oheimb@15188
   651
apply (rule i_rt_mono)
oheimb@15188
   652
by (blast intro: stream_take_lemma10)
oheimb@15188
   653
wenzelm@17291
   654
lemma take_i_rt_prefix_lemma1:
oheimb@15188
   655
  "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==>
wenzelm@17291
   656
   i_rt (Suc n) s1 << i_rt (Suc n) s2 ==>
oheimb@15188
   657
   i_rt n s1 << i_rt n s2 & stream_take n$s1 << stream_take n$s2"
oheimb@15188
   658
apply auto
oheimb@15188
   659
 apply (insert i_th_prefix_lemma [of n n s1 s2])
oheimb@15188
   660
 apply (rule i_th_i_rt_step,auto)
oheimb@15188
   661
by (drule mono_stream_take_pred,simp)
oheimb@15188
   662
wenzelm@17291
   663
lemma take_i_rt_prefix_lemma:
oheimb@15188
   664
"[| stream_take n$s1 << stream_take n$s2; i_rt n s1 << i_rt n s2 |] ==> s1 << s2"
oheimb@15188
   665
apply (case_tac "n=0",simp)
nipkow@25161
   666
apply (auto)
wenzelm@17291
   667
apply (subgoal_tac "stream_take 0$s1 << stream_take 0$s2 &
oheimb@15188
   668
                    i_rt 0 s1 << i_rt 0 s2")
oheimb@15188
   669
 defer 1
oheimb@15188
   670
 apply (rule zero_induct,blast)
oheimb@15188
   671
 apply (blast dest: take_i_rt_prefix_lemma1)
oheimb@15188
   672
by simp
oheimb@15188
   673
wenzelm@17291
   674
lemma streams_prefix_lemma: "(s1 << s2) =
wenzelm@17291
   675
  (stream_take n$s1 << stream_take n$s2 & i_rt n s1 << i_rt n s2)"
oheimb@15188
   676
apply auto
oheimb@15188
   677
  apply (simp add: monofun_cfun_arg)
oheimb@15188
   678
 apply (simp add: i_rt_mono)
oheimb@15188
   679
by (erule take_i_rt_prefix_lemma,simp)
oheimb@15188
   680
oheimb@15188
   681
lemma streams_prefix_lemma1:
oheimb@15188
   682
 "[| stream_take n$s1 = stream_take n$s2; i_rt n s1 = i_rt n s2 |] ==> s1 = s2"
oheimb@15188
   683
apply (simp add: po_eq_conv,auto)
oheimb@15188
   684
 apply (insert streams_prefix_lemma)
oheimb@15188
   685
 by blast+
oheimb@15188
   686
oheimb@15188
   687
oheimb@15188
   688
(* ----------------------------------------------------------------------- *)
oheimb@15188
   689
   section "sconc"
oheimb@15188
   690
(* ----------------------------------------------------------------------- *)
oheimb@15188
   691
oheimb@15188
   692
lemma UU_sconc [simp]: " UU ooo s = s "
haftmann@27111
   693
by (simp add: sconc_def zero_inat_def)
oheimb@15188
   694
oheimb@15188
   695
lemma scons_neq_UU: "a~=UU ==> a && s ~=UU"
oheimb@15188
   696
by auto
oheimb@15188
   697
oheimb@15188
   698
lemma singleton_sconc [rule_format, simp]: "x~=UU --> (x && UU) ooo y = x && y"
haftmann@27111
   699
apply (simp add: sconc_def zero_inat_def iSuc_def split: inat.splits, auto)
oheimb@15188
   700
apply (rule someI2_ex,auto)
oheimb@15188
   701
 apply (rule_tac x="x && y" in exI,auto)
oheimb@15188
   702
apply (simp add: i_rt_Suc_forw)
oheimb@15188
   703
apply (case_tac "xa=UU",simp)
oheimb@15188
   704
by (drule stream_exhaust_eq [THEN iffD1],auto)
oheimb@15188
   705
wenzelm@17291
   706
lemma ex_sconc [rule_format]:
oheimb@15188
   707
  "ALL k y. #x = Fin k --> (EX w. stream_take k$w = x & i_rt k w = y)"
oheimb@15188
   708
apply (case_tac "#x")
oheimb@15188
   709
 apply (rule stream_finite_ind [of x],auto)
oheimb@15188
   710
  apply (simp add: stream.finite_def)
oheimb@15188
   711
  apply (drule slen_take_lemma1,blast)
haftmann@27111
   712
 apply (simp_all add: zero_inat_def iSuc_def split: inat.splits)
oheimb@15188
   713
apply (erule_tac x="y" in allE,auto)
oheimb@15188
   714
by (rule_tac x="a && w" in exI,auto)
oheimb@15188
   715
wenzelm@17291
   716
lemma rt_sconc1: "Fin n = #x ==> i_rt n (x ooo y) = y"
haftmann@27111
   717
apply (simp add: sconc_def split: inat.splits, arith?,auto)
oheimb@15188
   718
apply (rule someI2_ex,auto)
oheimb@15188
   719
by (drule ex_sconc,simp)
oheimb@15188
   720
oheimb@15188
   721
lemma sconc_inj2: "\<lbrakk>Fin n = #x; x ooo y = x ooo z\<rbrakk> \<Longrightarrow> y = z"
oheimb@15188
   722
apply (frule_tac y=y in rt_sconc1)
oheimb@15188
   723
by (auto elim: rt_sconc1)
oheimb@15188
   724
oheimb@15188
   725
lemma sconc_UU [simp]:"s ooo UU = s"
oheimb@15188
   726
apply (case_tac "#s")
haftmann@27111
   727
 apply (simp add: sconc_def)
oheimb@15188
   728
 apply (rule someI2_ex)
oheimb@15188
   729
  apply (rule_tac x="s" in exI)
oheimb@15188
   730
  apply auto
oheimb@15188
   731
   apply (drule slen_take_lemma1,auto)
oheimb@15188
   732
  apply (simp add: i_rt_lemma_slen)
oheimb@15188
   733
 apply (drule slen_take_lemma1,auto)
oheimb@15188
   734
 apply (simp add: i_rt_slen)
haftmann@27111
   735
by (simp add: sconc_def)
oheimb@15188
   736
oheimb@15188
   737
lemma stream_take_sconc [simp]: "Fin n = #x ==> stream_take n$(x ooo y) = x"
oheimb@15188
   738
apply (simp add: sconc_def)
haftmann@27111
   739
apply (cases "#x")
haftmann@27111
   740
apply auto
haftmann@27111
   741
apply (rule someI2_ex, auto)
oheimb@15188
   742
by (drule ex_sconc,simp)
oheimb@15188
   743
oheimb@15188
   744
lemma scons_sconc [rule_format,simp]: "a~=UU --> (a && x) ooo y = a && x ooo y"
haftmann@27111
   745
apply (cases "#x",auto)
haftmann@27111
   746
 apply (simp add: sconc_def iSuc_Fin)
oheimb@15188
   747
 apply (rule someI2_ex)
haftmann@27111
   748
  apply (drule ex_sconc, simp)
haftmann@27111
   749
 apply (rule someI2_ex, auto)
oheimb@15188
   750
  apply (simp add: i_rt_Suc_forw)
haftmann@27111
   751
  apply (rule_tac x="a && x" in exI, auto)
oheimb@15188
   752
 apply (case_tac "xa=UU",auto)
oheimb@15188
   753
 apply (drule stream_exhaust_eq [THEN iffD1],auto)
oheimb@15188
   754
 apply (drule streams_prefix_lemma1,simp+)
oheimb@15188
   755
by (simp add: sconc_def)
oheimb@15188
   756
oheimb@15188
   757
lemma ft_sconc: "x ~= UU ==> ft$(x ooo y) = ft$x"
huffman@35781
   758
by (cases x, auto)
oheimb@15188
   759
oheimb@15188
   760
lemma sconc_assoc: "(x ooo y) ooo z = x ooo y ooo z"
oheimb@15188
   761
apply (case_tac "#x")
oheimb@15188
   762
 apply (rule stream_finite_ind [of x],auto simp del: scons_sconc)
oheimb@15188
   763
  apply (simp add: stream.finite_def del: scons_sconc)
oheimb@15188
   764
  apply (drule slen_take_lemma1,auto simp del: scons_sconc)
oheimb@15188
   765
 apply (case_tac "a = UU", auto)
oheimb@15188
   766
by (simp add: sconc_def)
oheimb@15188
   767
oheimb@15188
   768
oheimb@15188
   769
(* ----------------------------------------------------------------------- *)
oheimb@15188
   770
huffman@25833
   771
lemma cont_sconc_lemma1: "stream_finite x \<Longrightarrow> cont (\<lambda>y. x ooo y)"
huffman@25833
   772
by (erule stream_finite_ind, simp_all)
huffman@25833
   773
huffman@25833
   774
lemma cont_sconc_lemma2: "\<not> stream_finite x \<Longrightarrow> cont (\<lambda>y. x ooo y)"
huffman@25833
   775
by (simp add: sconc_def slen_def)
huffman@25833
   776
huffman@25833
   777
lemma cont_sconc: "cont (\<lambda>y. x ooo y)"
huffman@25833
   778
apply (cases "stream_finite x")
huffman@25833
   779
apply (erule cont_sconc_lemma1)
huffman@25833
   780
apply (erule cont_sconc_lemma2)
huffman@25833
   781
done
huffman@25833
   782
oheimb@15188
   783
lemma sconc_mono: "y << y' ==> x ooo y << x ooo y'"
huffman@25833
   784
by (rule cont_sconc [THEN cont2mono, THEN monofunE])
oheimb@15188
   785
oheimb@15188
   786
lemma sconc_mono1 [simp]: "x << x ooo y"
oheimb@15188
   787
by (rule sconc_mono [of UU, simplified])
oheimb@15188
   788
oheimb@15188
   789
(* ----------------------------------------------------------------------- *)
oheimb@15188
   790
oheimb@15188
   791
lemma empty_sconc [simp]: "(x ooo y = UU) = (x = UU & y = UU)"
oheimb@15188
   792
apply (case_tac "#x",auto)
wenzelm@17291
   793
   apply (insert sconc_mono1 [of x y])
huffman@19440
   794
   by auto
oheimb@15188
   795
oheimb@15188
   796
(* ----------------------------------------------------------------------- *)
oheimb@15188
   797
oheimb@15188
   798
lemma rt_sconc [rule_format, simp]: "s~=UU --> rt$(s ooo x) = rt$s ooo x"
huffman@35781
   799
by (cases s, auto)
oheimb@15188
   800
wenzelm@17291
   801
lemma i_th_sconc_lemma [rule_format]:
oheimb@15188
   802
  "ALL x y. Fin n < #x --> i_th n (x ooo y) = i_th n x"
oheimb@15188
   803
apply (induct_tac n, auto)
oheimb@15188
   804
apply (simp add: Fin_0 i_th_def)
oheimb@15188
   805
apply (simp add: slen_empty_eq ft_sconc)
oheimb@15188
   806
apply (simp add: i_th_def)
oheimb@15188
   807
apply (case_tac "x=UU",auto)
oheimb@15188
   808
apply (drule stream_exhaust_eq [THEN iffD1], auto)
oheimb@15188
   809
apply (erule_tac x="ya" in allE)
haftmann@27111
   810
apply (case_tac "#ya") by simp_all
oheimb@15188
   811
oheimb@15188
   812
oheimb@15188
   813
oheimb@15188
   814
(* ----------------------------------------------------------------------- *)
oheimb@15188
   815
oheimb@15188
   816
lemma sconc_lemma [rule_format, simp]: "ALL s. stream_take n$s ooo i_rt n s = s"
oheimb@15188
   817
apply (induct_tac n,auto)
oheimb@15188
   818
apply (case_tac "s=UU",auto)
oheimb@15188
   819
by (drule stream_exhaust_eq [THEN iffD1],auto)
oheimb@15188
   820
oheimb@15188
   821
(* ----------------------------------------------------------------------- *)
oheimb@15188
   822
   subsection "pointwise equality"
oheimb@15188
   823
(* ----------------------------------------------------------------------- *)
oheimb@15188
   824
wenzelm@17291
   825
lemma ex_last_stream_take_scons: "stream_take (Suc n)$s =
oheimb@15188
   826
                     stream_take n$s ooo i_rt n (stream_take (Suc n)$s)"
oheimb@15188
   827
by (insert sconc_lemma [of n "stream_take (Suc n)$s"],simp)
oheimb@15188
   828
wenzelm@17291
   829
lemma i_th_stream_take_eq:
oheimb@15188
   830
"!!n. ALL n. i_th n s1 = i_th n s2 ==> stream_take n$s1 = stream_take n$s2"
oheimb@15188
   831
apply (induct_tac n,auto)
oheimb@15188
   832
apply (subgoal_tac "stream_take (Suc na)$s1 =
oheimb@15188
   833
                    stream_take na$s1 ooo i_rt na (stream_take (Suc na)$s1)")
wenzelm@17291
   834
 apply (subgoal_tac "i_rt na (stream_take (Suc na)$s1) =
oheimb@15188
   835
                    i_rt na (stream_take (Suc na)$s2)")
wenzelm@17291
   836
  apply (subgoal_tac "stream_take (Suc na)$s2 =
oheimb@15188
   837
                    stream_take na$s2 ooo i_rt na (stream_take (Suc na)$s2)")
oheimb@15188
   838
   apply (insert ex_last_stream_take_scons,simp)
oheimb@15188
   839
  apply blast
oheimb@15188
   840
 apply (erule_tac x="na" in allE)
oheimb@15188
   841
 apply (insert i_th_last_eq [of _ s1 s2])
oheimb@15188
   842
by blast+
oheimb@15188
   843
oheimb@15188
   844
lemma pointwise_eq_lemma[rule_format]: "ALL n. i_th n s1 = i_th n s2 ==> s1 = s2"
huffman@35642
   845
by (insert i_th_stream_take_eq [THEN stream.take_lemma],blast)
oheimb@15188
   846
oheimb@15188
   847
(* ----------------------------------------------------------------------- *)
oheimb@15188
   848
   subsection "finiteness"
oheimb@15188
   849
(* ----------------------------------------------------------------------- *)
oheimb@15188
   850
oheimb@15188
   851
lemma slen_sconc_finite1:
oheimb@15188
   852
  "[| #(x ooo y) = Infty; Fin n = #x |] ==> #y = Infty"
oheimb@15188
   853
apply (case_tac "#y ~= Infty",auto)
oheimb@15188
   854
apply (drule_tac y=y in rt_sconc1)
oheimb@15188
   855
apply (insert stream_finite_i_rt [of n "x ooo y"])
oheimb@15188
   856
by (simp add: slen_infinite)
oheimb@15188
   857
oheimb@15188
   858
lemma slen_sconc_infinite1: "#x=Infty ==> #(x ooo y) = Infty"
oheimb@15188
   859
by (simp add: sconc_def)
oheimb@15188
   860
oheimb@15188
   861
lemma slen_sconc_infinite2: "#y=Infty ==> #(x ooo y) = Infty"
oheimb@15188
   862
apply (case_tac "#x")
oheimb@15188
   863
 apply (simp add: sconc_def)
oheimb@15188
   864
 apply (rule someI2_ex)
oheimb@15188
   865
  apply (drule ex_sconc,auto)
oheimb@15188
   866
 apply (erule contrapos_pp)
oheimb@15188
   867
 apply (insert stream_finite_i_rt)
nipkow@31084
   868
 apply (fastsimp simp add: slen_infinite,auto)
oheimb@15188
   869
by (simp add: sconc_def)
oheimb@15188
   870
oheimb@15188
   871
lemma sconc_finite: "(#x~=Infty & #y~=Infty) = (#(x ooo y)~=Infty)"
oheimb@15188
   872
apply auto
nipkow@31084
   873
  apply (metis not_Infty_eq slen_sconc_finite1)
nipkow@31084
   874
 apply (metis not_Infty_eq slen_sconc_infinite1)
nipkow@31084
   875
apply (metis not_Infty_eq slen_sconc_infinite2)
nipkow@31084
   876
done
oheimb@15188
   877
oheimb@15188
   878
(* ----------------------------------------------------------------------- *)
oheimb@15188
   879
oheimb@15188
   880
lemma slen_sconc_mono3: "[| Fin n = #x; Fin k = #(x ooo y) |] ==> n <= k"
oheimb@15188
   881
apply (insert slen_mono [of "x" "x ooo y"])
haftmann@27111
   882
apply (cases "#x") apply simp_all
haftmann@27111
   883
apply (cases "#(x ooo y)") apply simp_all
haftmann@27111
   884
done
oheimb@15188
   885
oheimb@15188
   886
(* ----------------------------------------------------------------------- *)
oheimb@15188
   887
   subsection "finite slen"
oheimb@15188
   888
(* ----------------------------------------------------------------------- *)
oheimb@15188
   889
oheimb@15188
   890
lemma slen_sconc: "[| Fin n = #x; Fin m = #y |] ==> #(x ooo y) = Fin (n + m)"
oheimb@15188
   891
apply (case_tac "#(x ooo y)")
oheimb@15188
   892
 apply (frule_tac y=y in rt_sconc1)
oheimb@15188
   893
 apply (insert take_i_rt_len [of "THE j. Fin j = #(x ooo y)" "x ooo y" n n m],simp)
oheimb@15188
   894
 apply (insert slen_sconc_mono3 [of n x _ y],simp)
oheimb@15188
   895
by (insert sconc_finite [of x y],auto)
oheimb@15188
   896
oheimb@15188
   897
(* ----------------------------------------------------------------------- *)
oheimb@15188
   898
   subsection "flat prefix"
oheimb@15188
   899
(* ----------------------------------------------------------------------- *)
oheimb@15188
   900
oheimb@15188
   901
lemma sconc_prefix: "(s1::'a::flat stream) << s2 ==> EX t. s1 ooo t = s2"
oheimb@15188
   902
apply (case_tac "#s1")
wenzelm@17291
   903
 apply (subgoal_tac "stream_take nat$s1 = stream_take nat$s2")
oheimb@15188
   904
  apply (rule_tac x="i_rt nat s2" in exI)
oheimb@15188
   905
  apply (simp add: sconc_def)
oheimb@15188
   906
  apply (rule someI2_ex)
oheimb@15188
   907
   apply (drule ex_sconc)
oheimb@15188
   908
   apply (simp,clarsimp,drule streams_prefix_lemma1)
wenzelm@17291
   909
   apply (simp+,rule slen_take_lemma3 [of _ s1 s2])
oheimb@15188
   910
  apply (simp+,rule_tac x="UU" in exI)
wenzelm@17291
   911
apply (insert slen_take_lemma3 [of _ s1 s2])
huffman@35642
   912
by (rule stream.take_lemma,simp)
oheimb@15188
   913
oheimb@15188
   914
(* ----------------------------------------------------------------------- *)
oheimb@15188
   915
   subsection "continuity"
oheimb@15188
   916
(* ----------------------------------------------------------------------- *)
oheimb@15188
   917
oheimb@15188
   918
lemma chain_sconc: "chain S ==> chain (%i. (x ooo S i))"
oheimb@15188
   919
by (simp add: chain_def,auto simp add: sconc_mono)
oheimb@15188
   920
oheimb@15188
   921
lemma chain_scons: "chain S ==> chain (%i. a && S i)"
oheimb@15188
   922
apply (simp add: chain_def,auto)
oheimb@15188
   923
by (rule monofun_cfun_arg,simp)
oheimb@15188
   924
oheimb@15188
   925
lemma contlub_scons_lemma: "chain S ==> (LUB i. a && S i) = a && (LUB i. S i)"
huffman@40565
   926
by (rule cont2contlubE [OF cont_Rep_cfun2, symmetric])
oheimb@15188
   927
wenzelm@17291
   928
lemma finite_lub_sconc: "chain Y ==> (stream_finite x) ==>
oheimb@15188
   929
                        (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))"
oheimb@15188
   930
apply (rule stream_finite_ind [of x])
oheimb@15188
   931
 apply (auto)
oheimb@15188
   932
apply (subgoal_tac "(LUB i. a && (s ooo Y i)) = a && (LUB i. s ooo Y i)")
oheimb@15188
   933
 by (force,blast dest: contlub_scons_lemma chain_sconc)
oheimb@15188
   934
wenzelm@17291
   935
lemma contlub_sconc_lemma:
oheimb@15188
   936
  "chain Y ==> (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))"
oheimb@15188
   937
apply (case_tac "#x=Infty")
oheimb@15188
   938
 apply (simp add: sconc_def)
huffman@18075
   939
apply (drule finite_lub_sconc,auto simp add: slen_infinite)
huffman@18075
   940
done
oheimb@15188
   941
oheimb@15188
   942
lemma monofun_sconc: "monofun (%y. x ooo y)"
huffman@16218
   943
by (simp add: monofun_def sconc_mono)
oheimb@15188
   944
oheimb@15188
   945
oheimb@15188
   946
(* ----------------------------------------------------------------------- *)
oheimb@15188
   947
   section "constr_sconc"
oheimb@15188
   948
(* ----------------------------------------------------------------------- *)
oheimb@15188
   949
oheimb@15188
   950
lemma constr_sconc_UUs [simp]: "constr_sconc UU s = s"
haftmann@27111
   951
by (simp add: constr_sconc_def zero_inat_def)
oheimb@15188
   952
oheimb@15188
   953
lemma "x ooo y = constr_sconc x y"
oheimb@15188
   954
apply (case_tac "#x")
oheimb@15188
   955
 apply (rule stream_finite_ind [of x],auto simp del: scons_sconc)
oheimb@15188
   956
  defer 1
oheimb@15188
   957
  apply (simp add: constr_sconc_def del: scons_sconc)
oheimb@15188
   958
  apply (case_tac "#s")
haftmann@27111
   959
   apply (simp add: iSuc_Fin)
oheimb@15188
   960
   apply (case_tac "a=UU",auto simp del: scons_sconc)
oheimb@15188
   961
   apply (simp)
oheimb@15188
   962
  apply (simp add: sconc_def)
oheimb@15188
   963
 apply (simp add: constr_sconc_def)
oheimb@15188
   964
apply (simp add: stream.finite_def)
oheimb@15188
   965
by (drule slen_take_lemma1,auto)
oheimb@15188
   966
oheimb@2570
   967
end