| author | nipkow | 
| Fri, 08 Mar 2013 11:28:04 +0100 | |
| changeset 51372 | d315e9a9ee72 | 
| parent 49521 | 06cb12198b92 | 
| child 57492 | 74bf65a1910a | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/Library/Stream.thy | 
| 17291 | 2 | Author: Franz Regensburger, David von Oheimb, Borislav Gajanovic | 
| 2570 | 3 | *) | 
| 4 | ||
| 17291 | 5 | header {* General Stream domain *}
 | 
| 6 | ||
| 7 | theory Stream | |
| 43919 | 8 | imports HOLCF "~~/src/HOL/Library/Extended_Nat" | 
| 17291 | 9 | begin | 
| 2570 | 10 | |
| 40329 
73f2b99b549d
change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
 huffman parents: 
40327diff
changeset | 11 | default_sort pcpo | 
| 
73f2b99b549d
change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
 huffman parents: 
40327diff
changeset | 12 | |
| 
73f2b99b549d
change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
 huffman parents: 
40327diff
changeset | 13 | domain (unsafe) 'a stream = scons (ft::'a) (lazy rt::"'a stream") (infixr "&&" 65) | 
| 2570 | 14 | |
| 19763 | 15 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 16 |   smap :: "('a \<rightarrow> 'b) \<rightarrow> 'a stream \<rightarrow> 'b stream" where
 | 
| 19763 | 17 | "smap = fix\<cdot>(\<Lambda> h f s. case s of x && xs \<Rightarrow> f\<cdot>x && h\<cdot>f\<cdot>xs)" | 
| 11348 | 18 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 19 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 20 |   sfilter :: "('a \<rightarrow> tr) \<rightarrow> 'a stream \<rightarrow> 'a stream" where
 | 
| 19763 | 21 | "sfilter = fix\<cdot>(\<Lambda> h p s. case s of x && xs \<Rightarrow> | 
| 40322 
707eb30e8a53
make syntax of continuous if-then-else consistent with HOL if-then-else
 huffman parents: 
40213diff
changeset | 22 | If p\<cdot>x then x && h\<cdot>p\<cdot>xs else h\<cdot>p\<cdot>xs)" | 
| 11348 | 23 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 24 | definition | 
| 43919 | 25 |   slen :: "'a stream \<Rightarrow> enat"  ("#_" [1000] 1000) where
 | 
| 43924 | 26 | "#s = (if stream_finite s then enat (LEAST n. stream_take n\<cdot>s = s) else \<infinity>)" | 
| 19763 | 27 | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 28 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 29 | (* concatenation *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 30 | |
| 19763 | 31 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 32 | i_rt :: "nat => 'a stream => 'a stream" where (* chops the first i elements *) | 
| 19763 | 33 | "i_rt = (%i s. iterate i$rt$s)" | 
| 17291 | 34 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 35 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 36 | i_th :: "nat => 'a stream => 'a" where (* the i-th element *) | 
| 19763 | 37 | "i_th = (%i s. ft$(i_rt i s))" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 38 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 39 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 40 | sconc :: "'a stream => 'a stream => 'a stream" (infixr "ooo" 65) where | 
| 19763 | 41 | "s1 ooo s2 = (case #s1 of | 
| 43924 | 42 | enat n \<Rightarrow> (SOME s. (stream_take n$s=s1) & (i_rt n s = s2)) | 
| 19763 | 43 | | \<infinity> \<Rightarrow> s1)" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 44 | |
| 27361 | 45 | primrec constr_sconc' :: "nat => 'a stream => 'a stream => 'a stream" | 
| 46 | where | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 47 | constr_sconc'_0: "constr_sconc' 0 s1 s2 = s2" | 
| 27361 | 48 | | 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 | 49 | constr_sconc' n (rt$s1) s2" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 50 | |
| 19763 | 51 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 52 | constr_sconc :: "'a stream => 'a stream => 'a stream" where (* constructive *) | 
| 19763 | 53 | "constr_sconc s1 s2 = (case #s1 of | 
| 43924 | 54 | enat n \<Rightarrow> constr_sconc' n s1 s2 | 
| 19763 | 55 | | \<infinity> \<Rightarrow> s1)" | 
| 56 | ||
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 57 | |
| 
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 | (* theorems about scons *) | 
| 
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 | |
| 
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 | section "scons" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 64 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 65 | lemma scons_eq_UU: "(a && s = UU) = (a = UU)" | 
| 30913 
10b26965a08f
domain package now generates iff rules for definedness of constructors
 huffman parents: 
30807diff
changeset | 66 | by simp | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 67 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 68 | lemma scons_not_empty: "[| a && x = UU; a ~= UU |] ==> R" | 
| 30913 
10b26965a08f
domain package now generates iff rules for definedness of constructors
 huffman parents: 
30807diff
changeset | 69 | by simp | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 70 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 71 | lemma stream_exhaust_eq: "(x ~= UU) = (EX a y. a ~= UU & x = a && y)" | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 72 | by (cases x, auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 73 | |
| 18109 | 74 | 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 | 75 | by (simp add: stream_exhaust_eq,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 76 | |
| 17291 | 77 | lemma stream_prefix: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 78 | "[| a && s << t; a ~= UU |] ==> EX b tt. t = b && tt & b ~= UU & s << tt" | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 79 | by (cases t, auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 80 | |
| 17291 | 81 | lemma stream_prefix': | 
| 82 | "b ~= UU ==> x << b && z = | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 83 | (x = UU | (EX a y. x = a && y & a ~= UU & a << b & y << z))" | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 84 | by (cases x, auto) | 
| 19550 
ae77a20f6995
update to reflect changes in inverts/injects lemmas
 huffman parents: 
19440diff
changeset | 85 | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 86 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 87 | (* | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 88 | 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 | 89 | by (insert stream_prefix' [of y "x&&xs" ys],force) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 90 | *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 91 | |
| 17291 | 92 | lemma stream_flat_prefix: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 93 | "[| 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 | 94 | apply (case_tac "y=UU",auto) | 
| 49521 | 95 | apply (drule ax_flat,simp) | 
| 96 | done | |
| 19550 
ae77a20f6995
update to reflect changes in inverts/injects lemmas
 huffman parents: 
19440diff
changeset | 97 | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 98 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 99 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 100 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 101 | (* ----------------------------------------------------------------------- *) | 
| 40213 
b63e966564da
rename case combinators generated by domain package to 'foo_case' instead of 'foo_when'
 huffman parents: 
40025diff
changeset | 102 | (* theorems about stream_case *) | 
| 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 | |
| 40213 
b63e966564da
rename case combinators generated by domain package to 'foo_case' instead of 'foo_when'
 huffman parents: 
40025diff
changeset | 105 | section "stream_case" | 
| 15188 
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 | |
| 40213 
b63e966564da
rename case combinators generated by domain package to 'foo_case' instead of 'foo_when'
 huffman parents: 
40025diff
changeset | 108 | lemma stream_case_strictf: "stream_case$UU$s=UU" | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 109 | by (cases s, auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 110 | |
| 
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 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 114 | (* theorems about ft and rt *) | 
| 
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 | section "ft & rt" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 119 | |
| 
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 | lemma ft_defin: "s~=UU ==> ft$s~=UU" | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 122 | by simp | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 123 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 124 | lemma rt_strict_rev: "rt$s~=UU ==> s~=UU" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 125 | by auto | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 126 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 127 | lemma surjectiv_scons: "(ft$s)&&(rt$s)=s" | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 128 | by (cases s, auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 129 | |
| 18075 | 130 | lemma monofun_rt_mult: "x << s ==> iterate i$rt$x << iterate i$rt$s" | 
| 131 | by (rule monofun_cfun_arg) | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 132 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 133 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 134 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 135 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 136 | (* theorems about stream_take *) | 
| 
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 | |
| 17291 | 140 | section "stream_take" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 141 | |
| 
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 | lemma stream_reach2: "(LUB i. stream_take i$s) = s" | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35444diff
changeset | 144 | by (rule stream.reach) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 145 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 146 | lemma chain_stream_take: "chain (%i. stream_take i$s)" | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 147 | by simp | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 148 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 149 | lemma stream_take_prefix [simp]: "stream_take n$s << s" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 150 | apply (insert stream_reach2 [of s]) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 151 | apply (erule subst) back | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 152 | apply (rule is_ub_thelub) | 
| 49521 | 153 | apply (simp only: chain_stream_take) | 
| 154 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 155 | |
| 17291 | 156 | lemma stream_take_more [rule_format]: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 157 | "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 | 158 | apply (induct_tac n,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 159 | apply (case_tac "x=UU",auto) | 
| 49521 | 160 | apply (drule stream_exhaust_eq [THEN iffD1],auto) | 
| 161 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 162 | |
| 17291 | 163 | lemma stream_take_lemma3 [rule_format]: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 164 | "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 | 165 | apply (induct_tac n,clarsimp) | 
| 16745 | 166 | (*apply (drule sym, erule scons_not_empty, simp)*) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 167 | apply (clarify, rule stream_take_more) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 168 | apply (erule_tac x="x" in allE) | 
| 49521 | 169 | apply (erule_tac x="xs" in allE,simp) | 
| 170 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 171 | |
| 17291 | 172 | lemma stream_take_lemma4: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 173 | "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 | 174 | by auto | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 175 | |
| 17291 | 176 | lemma stream_take_idempotent [rule_format, simp]: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 177 | "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 | 178 | apply (induct_tac n, auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 179 | apply (case_tac "s=UU", auto) | 
| 49521 | 180 | apply (drule stream_exhaust_eq [THEN iffD1], auto) | 
| 181 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 182 | |
| 17291 | 183 | lemma stream_take_take_Suc [rule_format, simp]: | 
| 184 | "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 | 185 | 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) | 
| 49521 | 188 | apply (drule stream_exhaust_eq [THEN iffD1], auto) | 
| 189 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 190 | |
| 17291 | 191 | lemma mono_stream_take_pred: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 192 | "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==> | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 193 | stream_take n$s1 << stream_take n$s2" | 
| 17291 | 194 | 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 | 195 | "stream_take (Suc n)$s2" "stream_take n"], 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" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 200 | by (drule mono_stream_take [of _ _ n],simp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 201 | *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 202 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 203 | lemma stream_take_lemma10 [rule_format]: | 
| 17291 | 204 | "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 | 205 | --> stream_take k$s1 << stream_take k$s2" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 206 | apply (induct_tac n,simp,clarsimp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 207 | apply (case_tac "k=Suc n",blast) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 208 | apply (erule_tac x="k" in allE) | 
| 49521 | 209 | apply (drule mono_stream_take_pred,simp) | 
| 210 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 211 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 212 | 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 | 213 | apply (insert chain_stream_take [of s1]) | 
| 49521 | 214 | apply (drule chain_mono,auto) | 
| 215 | done | |
| 15188 
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 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 | 218 | by (simp add: monofun_cfun_arg) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 219 | |
| 
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 stream_take_prefix [simp]: "stream_take n$s << s" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 222 | apply (subgoal_tac "s=(LUB n. stream_take n$s)") | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 223 | apply (erule ssubst, rule is_ub_thelub) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 224 | apply (simp only: chain_stream_take) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 225 | by (simp only: stream_reach2) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 226 | *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 227 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 228 | 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 | 229 | by (rule monofun_cfun_arg,auto) | 
| 
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 | (* ------------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 233 | (* special induction rules *) | 
| 
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 | section "induction" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 238 | |
| 17291 | 239 | lemma stream_finite_ind: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 240 | "[| 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 | 241 | apply (simp add: stream.finite_def,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 242 | apply (erule subst) | 
| 49521 | 243 | apply (drule stream.finite_induct [of P _ x], auto) | 
| 244 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 245 | |
| 17291 | 246 | lemma stream_finite_ind2: | 
| 247 | "[| 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 | 248 | !s. P (stream_take n$s)" | 
| 29855 
e0ab6cf95539
Repaired a proof that did, after all, refer to the theorem nat_induct2.
 paulson parents: 
29530diff
changeset | 249 | apply (rule nat_less_induct [of _ n],auto) | 
| 
e0ab6cf95539
Repaired a proof that did, after all, refer to the theorem nat_induct2.
 paulson parents: 
29530diff
changeset | 250 | apply (case_tac n, auto) | 
| 
e0ab6cf95539
Repaired a proof that did, after all, refer to the theorem nat_induct2.
 paulson parents: 
29530diff
changeset | 251 | apply (case_tac nat, auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 252 | apply (case_tac "s=UU",clarsimp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 253 | apply (drule stream_exhaust_eq [THEN iffD1],clarsimp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 254 | apply (case_tac "s=UU",clarsimp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 255 | apply (drule stream_exhaust_eq [THEN iffD1],clarsimp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 256 | apply (case_tac "y=UU",clarsimp) | 
| 49521 | 257 | apply (drule stream_exhaust_eq [THEN iffD1],clarsimp) | 
| 258 | done | |
| 15188 
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) | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35444diff
changeset | 263 | apply (erule admD, rule chain_stream_take) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 264 | apply (insert stream_finite_ind2 [of P]) | 
| 35494 
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
 huffman parents: 
35444diff
changeset | 265 | by simp | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 266 | |
| 
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 | (* simplify use of coinduction *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 271 | (* ----------------------------------------------------------------------- *) | 
| 
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 | section "coinduction" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 275 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 276 | lemma stream_coind_lemma2: "!s1 s2. R s1 s2 --> ft$s1 = ft$s2 & R (rt$s1) (rt$s2) ==> stream_bisim R" | 
| 30807 | 277 | apply (simp add: stream.bisim_def,clarsimp) | 
| 35497 | 278 | apply (drule spec, drule spec, drule (1) mp) | 
| 279 | apply (case_tac "x", simp) | |
| 40025 | 280 | apply (case_tac "y", simp) | 
| 49521 | 281 | apply auto | 
| 282 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 283 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 284 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 285 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 286 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 287 | (* theorems about stream_finite *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 288 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 289 | |
| 
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 | section "stream_finite" | 
| 
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 | lemma stream_finite_UU [simp]: "stream_finite UU" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 294 | by (simp add: stream.finite_def) | 
| 
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 | lemma stream_finite_UU_rev: "~ stream_finite s ==> s ~= UU" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 297 | by (auto simp add: stream.finite_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 298 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 299 | 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 | 300 | apply (simp add: stream.finite_def,auto) | 
| 35557 
5da670d57118
uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
 huffman parents: 
35524diff
changeset | 301 | apply (rule_tac x="Suc n" in exI) | 
| 49521 | 302 | apply (simp add: stream_take_lemma4) | 
| 303 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 304 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 305 | 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 | 306 | apply (simp add: stream.finite_def, auto) | 
| 35557 
5da670d57118
uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
 huffman parents: 
35524diff
changeset | 307 | apply (rule_tac x="n" in exI) | 
| 49521 | 308 | apply (erule stream_take_lemma3,simp) | 
| 309 | done | |
| 15188 
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_rt_eq: "stream_finite (rt$s) = stream_finite s" | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 312 | apply (cases s, auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 313 | apply (rule stream_finite_lemma1, simp) | 
| 49521 | 314 | apply (rule stream_finite_lemma2,simp) | 
| 315 | apply assumption | |
| 316 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 317 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 318 | lemma stream_finite_less: "stream_finite s ==> !t. t<<s --> stream_finite t" | 
| 19440 | 319 | apply (erule stream_finite_ind [of s], auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 320 | apply (case_tac "t=UU", auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 321 | apply (drule stream_exhaust_eq [THEN iffD1],auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 322 | apply (erule_tac x="y" in allE, simp) | 
| 49521 | 323 | apply (rule stream_finite_lemma1, simp) | 
| 324 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 325 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 326 | 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 | 327 | apply (simp add: stream.finite_def) | 
| 49521 | 328 | apply (rule_tac x="n" in exI,simp) | 
| 329 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 330 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 331 | lemma adm_not_stream_finite: "adm (%x. ~ stream_finite x)" | 
| 25833 | 332 | apply (rule adm_upward) | 
| 333 | apply (erule contrapos_nn) | |
| 334 | apply (erule (1) stream_finite_less [rule_format]) | |
| 335 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 336 | |
| 
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 | (* theorems about stream length *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 341 | (* ----------------------------------------------------------------------- *) | 
| 
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 | section "slen" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 345 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 346 | lemma slen_empty [simp]: "#\<bottom> = 0" | 
| 43919 | 347 | by (simp add: slen_def stream.finite_def zero_enat_def Least_equality) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 348 | |
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 349 | lemma slen_scons [simp]: "x ~= \<bottom> ==> #(x&&xs) = eSuc (#xs)" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 350 | apply (case_tac "stream_finite (x && xs)") | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 351 | apply (simp add: slen_def, auto) | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 352 | apply (simp add: stream.finite_def, auto simp add: eSuc_enat) | 
| 27111 | 353 | apply (rule Least_Suc2, auto) | 
| 16745 | 354 | (*apply (drule sym)*) | 
| 355 | (*apply (drule sym scons_eq_UU [THEN iffD1],simp)*) | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 356 | apply (erule stream_finite_lemma2, simp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 357 | apply (simp add: slen_def, auto) | 
| 49521 | 358 | apply (drule stream_finite_lemma1,auto) | 
| 359 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 360 | |
| 43924 | 361 | lemma slen_less_1_eq: "(#x < enat (Suc 0)) = (x = \<bottom>)" | 
| 49521 | 362 | by (cases x) (auto simp add: enat_0 eSuc_enat[THEN sym]) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 363 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 364 | lemma slen_empty_eq: "(#x = 0) = (x = \<bottom>)" | 
| 49521 | 365 | by (cases x) auto | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 366 | |
| 43924 | 367 | lemma slen_scons_eq: "(enat (Suc n) < #x) = (? a y. x = a && y & a ~= \<bottom> & enat n < #y)" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 368 | apply (auto, case_tac "x=UU",auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 369 | apply (drule stream_exhaust_eq [THEN iffD1], auto) | 
| 27111 | 370 | apply (case_tac "#y") apply simp_all | 
| 371 | apply (case_tac "#y") apply simp_all | |
| 372 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 373 | |
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 374 | lemma slen_eSuc: "#x = eSuc n --> (? a y. x = a&&y & a ~= \<bottom> & #y = n)" | 
| 49521 | 375 | by (cases x) auto | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 376 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 377 | 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 | 378 | by (simp add: slen_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 379 | |
| 43924 | 380 | lemma slen_scons_eq_rev: "(#x < enat (Suc (Suc n))) = (!a y. x ~= a && y | a = \<bottom> | #y < enat (Suc n))" | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 381 | apply (cases x, auto) | 
| 43919 | 382 | apply (simp add: zero_enat_def) | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 383 | apply (case_tac "#stream") apply (simp_all add: eSuc_enat) | 
| 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 384 | apply (case_tac "#stream") apply (simp_all add: eSuc_enat) | 
| 27111 | 385 | done | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 386 | |
| 17291 | 387 | lemma slen_take_lemma4 [rule_format]: | 
| 43924 | 388 | "!s. stream_take n$s ~= s --> #(stream_take n$s) = enat n" | 
| 389 | apply (induct n, auto simp add: enat_0) | |
| 27111 | 390 | apply (case_tac "s=UU", simp) | 
| 49521 | 391 | apply (drule stream_exhaust_eq [THEN iffD1], auto simp add: eSuc_enat) | 
| 392 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 393 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 394 | (* | 
| 17291 | 395 | lemma stream_take_idempotent [simp]: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 396 | "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 | 397 | apply (case_tac "stream_take n$s = s") | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 398 | apply (auto,insert slen_take_lemma4 [of n s]); | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 399 | 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 | 400 | |
| 17291 | 401 | 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 | 402 | stream_take n$s" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 403 | apply (simp add: po_eq_conv,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 404 | apply (simp add: stream_take_take_less) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 405 | 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 | 406 | apply (erule ssubst) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 407 | apply (rule_tac monofun_cfun_arg) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 408 | apply (insert chain_stream_take [of s]) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 409 | by (simp add: chain_def,simp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 410 | *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 411 | |
| 43924 | 412 | lemma slen_take_eq: "ALL x. (enat n < #x) = (stream_take n\<cdot>x ~= x)" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 413 | apply (induct_tac n, auto) | 
| 43924 | 414 | apply (simp add: enat_0, clarsimp) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 415 | apply (drule not_sym) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 416 | apply (drule slen_empty_eq [THEN iffD1], simp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 417 | apply (case_tac "x=UU", simp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 418 | apply (drule stream_exhaust_eq [THEN iffD1], clarsimp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 419 | apply (erule_tac x="y" in allE, auto) | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 420 | apply (simp_all add: not_less eSuc_enat) | 
| 27111 | 421 | apply (case_tac "#y") apply simp_all | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 422 | apply (case_tac "x=UU", simp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 423 | apply (drule stream_exhaust_eq [THEN iffD1], clarsimp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 424 | apply (erule_tac x="y" in allE, simp) | 
| 49521 | 425 | apply (case_tac "#y") | 
| 426 | apply simp_all | |
| 427 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 428 | |
| 43924 | 429 | lemma slen_take_eq_rev: "(#x <= enat n) = (stream_take n\<cdot>x = x)" | 
| 26102 | 430 | by (simp add: linorder_not_less [symmetric] slen_take_eq) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 431 | |
| 43924 | 432 | lemma slen_take_lemma1: "#x = enat n ==> stream_take n\<cdot>x = x" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 433 | by (rule slen_take_eq_rev [THEN iffD1], auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 434 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 435 | lemma slen_rt_mono: "#s2 <= #s1 ==> #(rt$s2) <= #(rt$s1)" | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 436 | apply (cases s1) | 
| 49521 | 437 | apply (cases s2, simp+)+ | 
| 438 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 439 | |
| 43924 | 440 | lemma slen_take_lemma5: "#(stream_take n$s) <= enat 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) | 
| 49521 | 443 | apply (simp add: slen_take_lemma4) | 
| 444 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 445 | |
| 43924 | 446 | lemma slen_take_lemma2: "!x. ~stream_finite x --> #(stream_take i\<cdot>x) = enat i" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 447 | apply (simp add: stream.finite_def, auto) | 
| 49521 | 448 | apply (simp add: slen_take_lemma4) | 
| 449 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 450 | |
| 43921 | 451 | lemma slen_infinite: "stream_finite x = (#x ~= \<infinity>)" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 452 | by (simp add: slen_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 453 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 454 | 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 | 455 | apply (erule stream_finite_ind [of s], auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 456 | apply (case_tac "t=UU", auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 457 | apply (drule stream_exhaust_eq [THEN iffD1], auto) | 
| 30807 | 458 | done | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 459 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 460 | lemma slen_mono: "s << t ==> #s <= #t" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 461 | apply (case_tac "stream_finite t") | 
| 17291 | 462 | apply (frule stream_finite_less) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 463 | apply (erule_tac x="s" in allE, simp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 464 | apply (drule slen_mono_lemma, auto) | 
| 49521 | 465 | apply (simp add: slen_def) | 
| 466 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 467 | |
| 18075 | 468 | 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 | 469 | by (insert iterate_Suc2 [of n F x], auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 470 | |
| 43924 | 471 | lemma slen_rt_mult [rule_format]: "!x. enat (i + j) <= #x --> enat j <= #(iterate i$rt$x)" | 
| 27111 | 472 | apply (induct i, auto) | 
| 43919 | 473 | apply (case_tac "x=UU", auto simp add: zero_enat_def) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 474 | apply (drule stream_exhaust_eq [THEN iffD1], auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 475 | apply (erule_tac x="y" in allE, auto) | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 476 | apply (simp add: not_le) apply (case_tac "#y") apply (simp_all add: eSuc_enat) | 
| 49521 | 477 | apply (simp add: iterate_lemma) | 
| 478 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 479 | |
| 17291 | 480 | lemma slen_take_lemma3 [rule_format]: | 
| 43924 | 481 | "!(x::'a::flat stream) y. enat n <= #x --> x << y --> stream_take n\<cdot>x = stream_take n\<cdot>y" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 482 | apply (induct_tac n, auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 483 | apply (case_tac "x=UU", auto) | 
| 43919 | 484 | apply (simp add: zero_enat_def) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 485 | apply (simp add: Suc_ile_eq) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 486 | apply (case_tac "y=UU", clarsimp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 487 | apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)+ | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 488 | apply (erule_tac x="ya" in allE, simp) | 
| 25920 | 489 | by (drule ax_flat, simp) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 490 | |
| 17291 | 491 | lemma slen_strict_mono_lemma: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 492 | "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 | 493 | apply (erule stream_finite_ind, auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 494 | apply (case_tac "sa=UU", auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 495 | apply (drule stream_exhaust_eq [THEN iffD1], clarsimp) | 
| 49521 | 496 | apply (drule ax_flat, simp) | 
| 497 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 498 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 499 | lemma slen_strict_mono: "[|stream_finite t; s ~= t; s << (t::'a::flat stream) |] ==> #s < #t" | 
| 27111 | 500 | by (auto simp add: slen_mono less_le dest: slen_strict_mono_lemma) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 501 | |
| 17291 | 502 | 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 | 503 | stream_take n$s ~= stream_take (Suc n)$s" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 504 | apply auto | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 505 | apply (subgoal_tac "stream_take n$s ~=s") | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 506 | apply (insert slen_take_lemma4 [of n s],auto) | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 507 | apply (cases s, simp) | 
| 49521 | 508 | apply (simp add: slen_take_lemma4 eSuc_enat) | 
| 509 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 510 | |
| 
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 | (* theorems about smap *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 513 | (* ----------------------------------------------------------------------- *) | 
| 
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 | section "smap" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 517 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 518 | lemma smap_unfold: "smap = (\<Lambda> f t. case t of x&&xs \<Rightarrow> f$x && smap$f$xs)" | 
| 29530 
9905b660612b
change to simpler, more extensible continuity simproc
 huffman parents: 
27361diff
changeset | 519 | by (insert smap_def [where 'a='a and 'b='b, THEN eq_reflection, THEN fix_eq2], auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 520 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 521 | lemma smap_empty [simp]: "smap\<cdot>f\<cdot>\<bottom> = \<bottom>" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 522 | by (subst smap_unfold, simp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 523 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 524 | 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 | 525 | by (subst smap_unfold, force) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 526 | |
| 
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 | (* theorems about sfilter *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 531 | (* ----------------------------------------------------------------------- *) | 
| 
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 | section "sfilter" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 534 | |
| 17291 | 535 | lemma sfilter_unfold: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 536 | "sfilter = (\<Lambda> p s. case s of x && xs \<Rightarrow> | 
| 40322 
707eb30e8a53
make syntax of continuous if-then-else consistent with HOL if-then-else
 huffman parents: 
40213diff
changeset | 537 | If p\<cdot>x then x && sfilter\<cdot>p\<cdot>xs else sfilter\<cdot>p\<cdot>xs)" | 
| 29530 
9905b660612b
change to simpler, more extensible continuity simproc
 huffman parents: 
27361diff
changeset | 538 | by (insert sfilter_def [where 'a='a, THEN eq_reflection, THEN fix_eq2], auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 539 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 540 | lemma strict_sfilter: "sfilter\<cdot>\<bottom> = \<bottom>" | 
| 40002 
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
 huffman parents: 
37110diff
changeset | 541 | apply (rule cfun_eqI) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 542 | apply (subst sfilter_unfold, auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 543 | apply (case_tac "x=UU", auto) | 
| 49521 | 544 | apply (drule stream_exhaust_eq [THEN iffD1], auto) | 
| 545 | done | |
| 15188 
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) = | |
| 40322 
707eb30e8a53
make syntax of continuous if-then-else consistent with HOL if-then-else
 huffman parents: 
40213diff
changeset | 552 | If f\<cdot>x then x && sfilter\<cdot>f\<cdot>xs else sfilter\<cdot>f\<cdot>xs" | 
| 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" | 
| 34941 | 561 | by (induct n) (simp_all add: i_rt_def) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 562 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 563 | lemma i_rt_0 [simp]: "i_rt 0 s = s" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 564 | by (simp add: i_rt_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 565 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 566 | 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 | 567 | 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 | 568 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 569 | 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 | 570 | by (simp only: i_rt_def iterate_Suc2) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 571 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 572 | 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 | 573 | by (simp only: i_rt_def,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 574 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 575 | 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 | 576 | by (simp add: i_rt_def monofun_rt_mult) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 577 | |
| 43924 | 578 | lemma i_rt_ij_lemma: "enat (i + j) <= #x ==> enat j <= #(i_rt i x)" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 579 | by (simp add: i_rt_def slen_rt_mult) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 580 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 581 | 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 | 582 | apply (induct_tac n,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 583 | apply (simp add: i_rt_Suc_back) | 
| 49521 | 584 | apply (drule slen_rt_mono,simp) | 
| 585 | done | |
| 15188 
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) | 
| 49521 | 591 | apply (drule stream_exhaust_eq [THEN iffD1],auto) | 
| 592 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 593 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 594 | 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 | 595 | apply auto | 
| 17291 | 596 | 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 | 597 | apply (subgoal_tac "#(i_rt n s)=0") | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 598 | apply (case_tac "stream_take n$s = s",simp+) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 599 | apply (insert slen_take_eq [rule_format,of n s],simp) | 
| 43919 | 600 | apply (cases "#s") apply (simp_all add: zero_enat_def) | 
| 27111 | 601 | apply (simp add: slen_take_eq) | 
| 602 | apply (cases "#s") | |
| 603 | using i_rt_take_lemma1 [of n s] | |
| 43919 | 604 | apply (simp_all add: zero_enat_def) | 
| 27111 | 605 | done | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 606 | |
| 43924 | 607 | lemma i_rt_lemma_slen: "#s=enat n ==> i_rt n s = UU" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 608 | by (simp add: i_rt_slen slen_take_lemma1) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 609 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 610 | 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 | 611 | apply (induct_tac n, auto) | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 612 | apply (cases s, auto simp del: i_rt_Suc) | 
| 49521 | 613 | apply (simp add: i_rt_Suc_back stream_finite_rt_eq)+ | 
| 614 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 615 | |
| 43924 | 616 | lemma take_i_rt_len_lemma: "ALL sl x j t. enat sl = #x & n <= sl & | 
| 617 | #(stream_take n$x) = enat t & #(i_rt n x)= enat j | |
| 618 | --> enat (j + t) = #x" | |
| 27111 | 619 | apply (induct n, auto) | 
| 43919 | 620 | apply (simp add: zero_enat_def) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 621 | apply (case_tac "x=UU",auto) | 
| 43919 | 622 | apply (simp add: zero_enat_def) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 623 | apply (drule stream_exhaust_eq [THEN iffD1],clarsimp) | 
| 43924 | 624 | apply (subgoal_tac "EX k. enat k = #y",clarify) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 625 | apply (erule_tac x="k" in allE) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 626 | apply (erule_tac x="y" in allE,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 627 | apply (erule_tac x="THE p. Suc p = t" in allE,auto) | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 628 | apply (simp add: eSuc_def split: enat.splits) | 
| 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 629 | apply (simp add: eSuc_def split: enat.splits) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 630 | apply (simp only: the_equality) | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 631 | apply (simp add: eSuc_def split: enat.splits) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 632 | apply force | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 633 | apply (simp add: eSuc_def split: enat.splits) | 
| 27111 | 634 | done | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 635 | |
| 17291 | 636 | lemma take_i_rt_len: | 
| 43924 | 637 | "[| enat sl = #x; n <= sl; #(stream_take n$x) = enat t; #(i_rt n x) = enat j |] ==> | 
| 638 | enat (j + t) = #x" | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 639 | by (blast intro: take_i_rt_len_lemma [rule_format]) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 640 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 641 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 642 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 643 | section "i_th" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 644 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 645 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 646 | lemma i_th_i_rt_step: | 
| 17291 | 647 | "[| 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 | 648 | i_rt n s1 << i_rt n s2" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 649 | apply (simp add: i_th_def i_rt_Suc_back) | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 650 | apply (cases "i_rt n s1", simp) | 
| 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 651 | apply (cases "i_rt n s2", auto) | 
| 30807 | 652 | done | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 653 | |
| 17291 | 654 | lemma i_th_stream_take_Suc [rule_format]: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 655 | "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 | 656 | apply (induct_tac n,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 657 | apply (simp add: i_th_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 658 | apply (case_tac "s=UU",auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 659 | apply (drule stream_exhaust_eq [THEN iffD1],auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 660 | apply (case_tac "s=UU",simp add: i_th_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 661 | apply (drule stream_exhaust_eq [THEN iffD1],auto) | 
| 49521 | 662 | apply (simp add: i_th_def i_rt_Suc_forw) | 
| 663 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 664 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 665 | 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 | 666 | 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 | 667 | apply (rule i_th_stream_take_Suc [THEN subst]) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 668 | 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 | 669 | by (simp add: i_rt_take_lemma1) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 670 | |
| 17291 | 671 | lemma i_th_last_eq: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 672 | "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 | 673 | apply (insert i_th_last [of n s1]) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 674 | apply (insert i_th_last [of n s2]) | 
| 49521 | 675 | apply auto | 
| 676 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 677 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 678 | lemma i_th_prefix_lemma: | 
| 17291 | 679 | "[| 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 | 680 | i_th k s1 << i_th k s2" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 681 | apply (insert i_th_stream_take_Suc [of k s1, THEN sym]) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 682 | apply (insert i_th_stream_take_Suc [of k s2, THEN sym],auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 683 | apply (simp add: i_th_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 684 | apply (rule monofun_cfun, auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 685 | apply (rule i_rt_mono) | 
| 49521 | 686 | apply (blast intro: stream_take_lemma10) | 
| 687 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 688 | |
| 17291 | 689 | lemma take_i_rt_prefix_lemma1: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 690 | "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==> | 
| 17291 | 691 | 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 | 692 | 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 | 693 | apply auto | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 694 | 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 | 695 | apply (rule i_th_i_rt_step,auto) | 
| 49521 | 696 | apply (drule mono_stream_take_pred,simp) | 
| 697 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 698 | |
| 17291 | 699 | lemma take_i_rt_prefix_lemma: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 700 | "[| 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 | 701 | apply (case_tac "n=0",simp) | 
| 25161 | 702 | apply (auto) | 
| 17291 | 703 | 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 | 704 | i_rt 0 s1 << i_rt 0 s2") | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 705 | defer 1 | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 706 | apply (rule zero_induct,blast) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 707 | apply (blast dest: take_i_rt_prefix_lemma1) | 
| 49521 | 708 | apply simp | 
| 709 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 710 | |
| 17291 | 711 | lemma streams_prefix_lemma: "(s1 << s2) = | 
| 712 | (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 | 713 | apply auto | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 714 | apply (simp add: monofun_cfun_arg) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 715 | apply (simp add: i_rt_mono) | 
| 49521 | 716 | apply (erule take_i_rt_prefix_lemma,simp) | 
| 717 | done | |
| 15188 
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 streams_prefix_lemma1: | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 720 | "[| 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 | 721 | apply (simp add: po_eq_conv,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 722 | apply (insert streams_prefix_lemma) | 
| 49521 | 723 | apply blast+ | 
| 724 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 725 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 726 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 727 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 728 | section "sconc" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 729 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 730 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 731 | lemma UU_sconc [simp]: " UU ooo s = s " | 
| 43919 | 732 | by (simp add: sconc_def zero_enat_def) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 733 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 734 | lemma scons_neq_UU: "a~=UU ==> a && s ~=UU" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 735 | by auto | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 736 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 737 | lemma singleton_sconc [rule_format, simp]: "x~=UU --> (x && UU) ooo y = x && y" | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 738 | apply (simp add: sconc_def zero_enat_def eSuc_def split: enat.splits, auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 739 | apply (rule someI2_ex,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 740 | apply (rule_tac x="x && y" in exI,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 741 | apply (simp add: i_rt_Suc_forw) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 742 | apply (case_tac "xa=UU",simp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 743 | by (drule stream_exhaust_eq [THEN iffD1],auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 744 | |
| 17291 | 745 | lemma ex_sconc [rule_format]: | 
| 43924 | 746 | "ALL k y. #x = enat k --> (EX w. stream_take k$w = x & i_rt k w = y)" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 747 | apply (case_tac "#x") | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 748 | apply (rule stream_finite_ind [of x],auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 749 | apply (simp add: stream.finite_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 750 | apply (drule slen_take_lemma1,blast) | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 751 | apply (simp_all add: zero_enat_def eSuc_def split: enat.splits) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 752 | apply (erule_tac x="y" in allE,auto) | 
| 49521 | 753 | apply (rule_tac x="a && w" in exI,auto) | 
| 754 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 755 | |
| 43924 | 756 | lemma rt_sconc1: "enat n = #x ==> i_rt n (x ooo y) = y" | 
| 43919 | 757 | apply (simp add: sconc_def split: enat.splits, arith?,auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 758 | apply (rule someI2_ex,auto) | 
| 49521 | 759 | apply (drule ex_sconc,simp) | 
| 760 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 761 | |
| 43924 | 762 | lemma sconc_inj2: "\<lbrakk>enat n = #x; x ooo y = x ooo z\<rbrakk> \<Longrightarrow> y = z" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 763 | apply (frule_tac y=y in rt_sconc1) | 
| 49521 | 764 | apply (auto elim: rt_sconc1) | 
| 765 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 766 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 767 | lemma sconc_UU [simp]:"s ooo UU = s" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 768 | apply (case_tac "#s") | 
| 27111 | 769 | apply (simp add: sconc_def) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 770 | apply (rule someI2_ex) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 771 | apply (rule_tac x="s" in exI) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 772 | apply auto | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 773 | apply (drule slen_take_lemma1,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 774 | apply (simp add: i_rt_lemma_slen) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 775 | apply (drule slen_take_lemma1,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 776 | apply (simp add: i_rt_slen) | 
| 49521 | 777 | apply (simp add: sconc_def) | 
| 778 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 779 | |
| 43924 | 780 | lemma stream_take_sconc [simp]: "enat n = #x ==> stream_take n$(x ooo y) = x" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 781 | apply (simp add: sconc_def) | 
| 27111 | 782 | apply (cases "#x") | 
| 783 | apply auto | |
| 784 | apply (rule someI2_ex, auto) | |
| 49521 | 785 | apply (drule ex_sconc,simp) | 
| 786 | done | |
| 15188 
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 scons_sconc [rule_format,simp]: "a~=UU --> (a && x) ooo y = a && x ooo y" | 
| 27111 | 789 | apply (cases "#x",auto) | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 790 | apply (simp add: sconc_def eSuc_enat) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 791 | apply (rule someI2_ex) | 
| 27111 | 792 | apply (drule ex_sconc, simp) | 
| 793 | apply (rule someI2_ex, auto) | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 794 | apply (simp add: i_rt_Suc_forw) | 
| 27111 | 795 | apply (rule_tac x="a && x" in exI, auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 796 | apply (case_tac "xa=UU",auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 797 | apply (drule stream_exhaust_eq [THEN iffD1],auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 798 | apply (drule streams_prefix_lemma1,simp+) | 
| 49521 | 799 | apply (simp add: sconc_def) | 
| 800 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 801 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 802 | lemma ft_sconc: "x ~= UU ==> ft$(x ooo y) = ft$x" | 
| 49521 | 803 | by (cases x) auto | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 804 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 805 | 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 | 806 | apply (case_tac "#x") | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 807 | 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 | 808 | apply (simp add: stream.finite_def del: scons_sconc) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 809 | apply (drule slen_take_lemma1,auto simp del: scons_sconc) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 810 | apply (case_tac "a = UU", auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 811 | by (simp add: sconc_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 812 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 813 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 814 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 815 | |
| 25833 | 816 | lemma cont_sconc_lemma1: "stream_finite x \<Longrightarrow> cont (\<lambda>y. x ooo y)" | 
| 817 | by (erule stream_finite_ind, simp_all) | |
| 818 | ||
| 819 | lemma cont_sconc_lemma2: "\<not> stream_finite x \<Longrightarrow> cont (\<lambda>y. x ooo y)" | |
| 820 | by (simp add: sconc_def slen_def) | |
| 821 | ||
| 822 | lemma cont_sconc: "cont (\<lambda>y. x ooo y)" | |
| 823 | apply (cases "stream_finite x") | |
| 824 | apply (erule cont_sconc_lemma1) | |
| 825 | apply (erule cont_sconc_lemma2) | |
| 826 | done | |
| 827 | ||
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 828 | lemma sconc_mono: "y << y' ==> x ooo y << x ooo y'" | 
| 25833 | 829 | by (rule cont_sconc [THEN cont2mono, THEN monofunE]) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 830 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 831 | lemma sconc_mono1 [simp]: "x << x ooo y" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 832 | by (rule sconc_mono [of UU, simplified]) | 
| 
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 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 | 837 | apply (case_tac "#x",auto) | 
| 17291 | 838 | apply (insert sconc_mono1 [of x y]) | 
| 49521 | 839 | apply auto | 
| 840 | done | |
| 15188 
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 | (* ----------------------------------------------------------------------- *) | 
| 
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 | lemma rt_sconc [rule_format, simp]: "s~=UU --> rt$(s ooo x) = rt$s ooo x" | 
| 35781 
b7738ab762b1
renamed some lemmas generated by the domain package
 huffman parents: 
35642diff
changeset | 845 | by (cases s, auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 846 | |
| 17291 | 847 | lemma i_th_sconc_lemma [rule_format]: | 
| 43924 | 848 | "ALL x y. enat n < #x --> i_th n (x ooo y) = i_th n x" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 849 | apply (induct_tac n, auto) | 
| 43924 | 850 | apply (simp add: enat_0 i_th_def) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 851 | apply (simp add: slen_empty_eq ft_sconc) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 852 | apply (simp add: i_th_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 853 | apply (case_tac "x=UU",auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 854 | apply (drule stream_exhaust_eq [THEN iffD1], auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 855 | apply (erule_tac x="ya" in allE) | 
| 49521 | 856 | apply (case_tac "#ya") | 
| 857 | apply simp_all | |
| 858 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 859 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 860 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 861 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 862 | (* ----------------------------------------------------------------------- *) | 
| 
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 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 | 865 | apply (induct_tac n,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 866 | apply (case_tac "s=UU",auto) | 
| 49521 | 867 | apply (drule stream_exhaust_eq [THEN iffD1],auto) | 
| 868 | done | |
| 15188 
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 | subsection "pointwise equality" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 872 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 873 | |
| 17291 | 874 | 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 | 875 | 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 | 876 | 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 | 877 | |
| 17291 | 878 | lemma i_th_stream_take_eq: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 879 | "!!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 | 880 | apply (induct_tac n,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 881 | apply (subgoal_tac "stream_take (Suc na)$s1 = | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 882 | stream_take na$s1 ooo i_rt na (stream_take (Suc na)$s1)") | 
| 17291 | 883 | 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 | 884 | i_rt na (stream_take (Suc na)$s2)") | 
| 17291 | 885 | apply (subgoal_tac "stream_take (Suc na)$s2 = | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 886 | 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 | 887 | apply (insert ex_last_stream_take_scons,simp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 888 | apply blast | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 889 | apply (erule_tac x="na" in allE) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 890 | apply (insert i_th_last_eq [of _ s1 s2]) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 891 | by blast+ | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 892 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 893 | lemma pointwise_eq_lemma[rule_format]: "ALL n. i_th n s1 = i_th n s2 ==> s1 = s2" | 
| 35642 
f478d5a9d238
generate separate qualified theorem name for each type's reach and take_lemma
 huffman parents: 
35557diff
changeset | 894 | by (insert i_th_stream_take_eq [THEN stream.take_lemma],blast) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 895 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 896 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 897 | subsection "finiteness" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 898 | (* ----------------------------------------------------------------------- *) | 
| 
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 | lemma slen_sconc_finite1: | 
| 43924 | 901 | "[| #(x ooo y) = \<infinity>; enat n = #x |] ==> #y = \<infinity>" | 
| 43921 | 902 | apply (case_tac "#y ~= \<infinity>",auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 903 | apply (drule_tac y=y in rt_sconc1) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 904 | apply (insert stream_finite_i_rt [of n "x ooo y"]) | 
| 49521 | 905 | apply (simp add: slen_infinite) | 
| 906 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 907 | |
| 43921 | 908 | lemma slen_sconc_infinite1: "#x=\<infinity> ==> #(x ooo y) = \<infinity>" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 909 | by (simp add: sconc_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 910 | |
| 43921 | 911 | lemma slen_sconc_infinite2: "#y=\<infinity> ==> #(x ooo y) = \<infinity>" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 912 | apply (case_tac "#x") | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 913 | apply (simp add: sconc_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 914 | apply (rule someI2_ex) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 915 | apply (drule ex_sconc,auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 916 | apply (erule contrapos_pp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 917 | apply (insert stream_finite_i_rt) | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44019diff
changeset | 918 | apply (fastforce simp add: slen_infinite,auto) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 919 | by (simp add: sconc_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 920 | |
| 43921 | 921 | lemma sconc_finite: "(#x~=\<infinity> & #y~=\<infinity>) = (#(x ooo y)~=\<infinity>)" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 922 | apply auto | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 923 | apply (metis not_infinity_eq slen_sconc_finite1) | 
| 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 924 | apply (metis not_infinity_eq slen_sconc_infinite1) | 
| 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 925 | apply (metis not_infinity_eq slen_sconc_infinite2) | 
| 31084 | 926 | done | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 927 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 928 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 929 | |
| 43924 | 930 | lemma slen_sconc_mono3: "[| enat n = #x; enat k = #(x ooo y) |] ==> n <= k" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 931 | apply (insert slen_mono [of "x" "x ooo y"]) | 
| 27111 | 932 | apply (cases "#x") apply simp_all | 
| 933 | apply (cases "#(x ooo y)") apply simp_all | |
| 934 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 935 | |
| 
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 | subsection "finite slen" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 938 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 939 | |
| 43924 | 940 | lemma slen_sconc: "[| enat n = #x; enat m = #y |] ==> #(x ooo y) = enat (n + m)" | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 941 | apply (case_tac "#(x ooo y)") | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 942 | apply (frule_tac y=y in rt_sconc1) | 
| 43924 | 943 | apply (insert take_i_rt_len [of "THE j. enat j = #(x ooo y)" "x ooo y" n n m],simp) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 944 | apply (insert slen_sconc_mono3 [of n x _ y],simp) | 
| 49521 | 945 | apply (insert sconc_finite [of x y],auto) | 
| 946 | done | |
| 15188 
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 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 949 | subsection "flat prefix" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 950 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 951 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 952 | 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 | 953 | apply (case_tac "#s1") | 
| 17291 | 954 | 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 | 955 | apply (rule_tac x="i_rt nat s2" in exI) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 956 | apply (simp add: sconc_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 957 | apply (rule someI2_ex) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 958 | apply (drule ex_sconc) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 959 | apply (simp,clarsimp,drule streams_prefix_lemma1) | 
| 17291 | 960 | apply (simp+,rule slen_take_lemma3 [of _ s1 s2]) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 961 | apply (simp+,rule_tac x="UU" in exI) | 
| 17291 | 962 | apply (insert slen_take_lemma3 [of _ s1 s2]) | 
| 49521 | 963 | apply (rule stream.take_lemma,simp) | 
| 964 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 965 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 966 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 967 | subsection "continuity" | 
| 
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 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 970 | 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 | 971 | by (simp add: chain_def,auto simp add: sconc_mono) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 972 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 973 | lemma chain_scons: "chain S ==> chain (%i. a && S i)" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 974 | apply (simp add: chain_def,auto) | 
| 49521 | 975 | apply (rule monofun_cfun_arg,simp) | 
| 976 | done | |
| 15188 
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 | lemma contlub_scons_lemma: "chain S ==> (LUB i. a && S i) = a && (LUB i. S i)" | 
| 40327 | 979 | by (rule cont2contlubE [OF cont_Rep_cfun2, symmetric]) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 980 | |
| 17291 | 981 | lemma finite_lub_sconc: "chain Y ==> (stream_finite x) ==> | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 982 | (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 | 983 | apply (rule stream_finite_ind [of x]) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 984 | apply (auto) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 985 | apply (subgoal_tac "(LUB i. a && (s ooo Y i)) = a && (LUB i. s ooo Y i)") | 
| 49521 | 986 | apply (force,blast dest: contlub_scons_lemma chain_sconc) | 
| 987 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 988 | |
| 17291 | 989 | lemma contlub_sconc_lemma: | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 990 | "chain Y ==> (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))" | 
| 43921 | 991 | apply (case_tac "#x=\<infinity>") | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 992 | apply (simp add: sconc_def) | 
| 18075 | 993 | apply (drule finite_lub_sconc,auto simp add: slen_infinite) | 
| 994 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 995 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 996 | lemma monofun_sconc: "monofun (%y. x ooo y)" | 
| 16218 | 997 | by (simp add: monofun_def sconc_mono) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 998 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 999 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1000 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1001 | section "constr_sconc" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1002 | (* ----------------------------------------------------------------------- *) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1003 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1004 | lemma constr_sconc_UUs [simp]: "constr_sconc UU s = s" | 
| 43919 | 1005 | by (simp add: constr_sconc_def zero_enat_def) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1006 | |
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1007 | lemma "x ooo y = constr_sconc x y" | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1008 | apply (case_tac "#x") | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1009 | 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 | 1010 | defer 1 | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1011 | apply (simp add: constr_sconc_def del: scons_sconc) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1012 | apply (case_tac "#s") | 
| 44019 
ee784502aed5
Extended_Nat.thy: renamed iSuc to eSuc, standardized theorem names
 huffman parents: 
43924diff
changeset | 1013 | apply (simp add: eSuc_enat) | 
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1014 | apply (case_tac "a=UU",auto simp del: scons_sconc) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1015 | apply (simp) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1016 | apply (simp add: sconc_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1017 | apply (simp add: constr_sconc_def) | 
| 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1018 | apply (simp add: stream.finite_def) | 
| 49521 | 1019 | apply (drule slen_take_lemma1,auto) | 
| 1020 | done | |
| 15188 
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
 oheimb parents: 
14981diff
changeset | 1021 | |
| 2570 | 1022 | end |