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