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