author | haftmann |
Fri, 03 Nov 2006 14:22:44 +0100 | |
changeset 21158 | b379fdc3a3bd |
parent 19763 | ec18656a2c10 |
child 21210 | c17fd2df4e9e |
permissions | -rw-r--r-- |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
1 |
(* Title: HOLCF/FOCUS/Fstreams.thy |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
2 |
ID: $Id$ |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
3 |
Author: Borislav Gajanovic |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
4 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
5 |
FOCUS flat streams (with lifted elements) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
6 |
###TODO: integrate this with Fstream.* |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
7 |
*) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
8 |
|
16417 | 9 |
theory Fstreams imports Stream begin |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
10 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
11 |
defaultsort type |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
12 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
13 |
types 'a fstream = "('a lift) stream" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
14 |
|
19763 | 15 |
definition |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
16 |
fsingleton :: "'a => 'a fstream" ("<_>" [1000] 999) |
19763 | 17 |
fsingleton_def2: "fsingleton = (%a. Def a && UU)" |
18 |
||
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
19 |
fsfilter :: "'a set \<Rightarrow> 'a fstream \<rightarrow> 'a fstream" |
19763 | 20 |
"fsfilter A = sfilter\<cdot>(flift2 (\<lambda>x. x\<in>A))" |
21 |
||
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
22 |
fsmap :: "('a => 'b) => 'a fstream -> 'b fstream" |
19763 | 23 |
"fsmap f = smap$(flift2 f)" |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
24 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
25 |
jth :: "nat => 'a fstream => 'a" |
19763 | 26 |
"jth = (%n s. if Fin n < #s then THE a. i_th n s = Def a else arbitrary)" |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
27 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
28 |
first :: "'a fstream => 'a" |
19763 | 29 |
"first = (%s. jth 0 s)" |
30 |
||
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
31 |
last :: "'a fstream => 'a" |
19763 | 32 |
"last = (%s. case #s of Fin n => (if n~=0 then jth (THE k. Suc k = n) s else arbitrary) |
33 |
| Infty => arbitrary)" |
|
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
34 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
35 |
|
19763 | 36 |
abbreviation |
37 |
emptystream :: "'a fstream" ("<>") |
|
38 |
"<> == \<bottom>" |
|
39 |
fsfilter' :: "'a set \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" ("(_'(C')_)" [64,63] 63) |
|
40 |
"A(C)s == fsfilter A\<cdot>s" |
|
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
41 |
|
19763 | 42 |
const_syntax (xsymbols) |
43 |
fsfilter' ("(_\<copyright>_)" [64,63] 63) |
|
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
44 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
45 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
46 |
lemma ft_fsingleton[simp]: "ft$(<a>) = Def a" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
47 |
by (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
48 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
49 |
lemma slen_fsingleton[simp]: "#(<a>) = Fin 1" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
50 |
by (simp add: fsingleton_def2 inat_defs) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
51 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
52 |
lemma slen_fstreams[simp]: "#(<a> ooo s) = iSuc (#s)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
53 |
by (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
54 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
55 |
lemma slen_fstreams2[simp]: "#(s ooo <a>) = iSuc (#s)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
56 |
apply (case_tac "#s", auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
57 |
apply (insert slen_sconc [of _ s "Suc 0" "<a>"], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
58 |
by (simp add: sconc_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
59 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
60 |
lemma j_th_0_fsingleton[simp]:"jth 0 (<a>) = a" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
61 |
apply (simp add: fsingleton_def2 jth_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
62 |
by (simp add: i_th_def Fin_0) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
63 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
64 |
lemma jth_0[simp]: "jth 0 (<a> ooo s) = a" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
65 |
apply (simp add: fsingleton_def2 jth_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
66 |
by (simp add: i_th_def Fin_0) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
67 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
68 |
lemma first_sconc[simp]: "first (<a> ooo s) = a" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
69 |
by (simp add: first_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
70 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
71 |
lemma first_fsingleton[simp]: "first (<a>) = a" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
72 |
by (simp add: first_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
73 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
74 |
lemma jth_n[simp]: "Fin n = #s ==> jth n (s ooo <a>) = a" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
75 |
apply (simp add: jth_def, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
76 |
apply (simp add: i_th_def rt_sconc1) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
77 |
by (simp add: inat_defs split: inat_splits) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
78 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
79 |
lemma last_sconc[simp]: "Fin n = #s ==> last (s ooo <a>) = a" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
80 |
apply (simp add: last_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
81 |
apply (simp add: inat_defs split:inat_splits) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
82 |
by (drule sym, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
83 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
84 |
lemma last_fsingleton[simp]: "last (<a>) = a" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
85 |
by (simp add: last_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
86 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
87 |
lemma first_UU[simp]: "first UU = arbitrary" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
88 |
by (simp add: first_def jth_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
89 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
90 |
lemma last_UU[simp]:"last UU = arbitrary" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
91 |
by (simp add: last_def jth_def inat_defs) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
92 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
93 |
lemma last_infinite[simp]:"#s = Infty ==> last s = arbitrary" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
94 |
by (simp add: last_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
95 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
96 |
lemma jth_slen_lemma1:"n <= k & Fin n = #s ==> jth k s = arbitrary" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
97 |
by (simp add: jth_def inat_defs split:inat_splits, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
98 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
99 |
lemma jth_UU[simp]:"jth n UU = arbitrary" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
100 |
by (simp add: jth_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
101 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
102 |
lemma ext_last:"[|s ~= UU; Fin (Suc n) = #s|] ==> (stream_take n$s) ooo <(last s)> = s" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
103 |
apply (simp add: last_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
104 |
apply (case_tac "#s", auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
105 |
apply (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
106 |
apply (subgoal_tac "Def (jth n s) = i_th n s") |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
107 |
apply (auto simp add: i_th_last) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
108 |
apply (drule slen_take_lemma1, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
109 |
apply (simp add: jth_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
110 |
apply (case_tac "i_th n s = UU") |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
111 |
apply auto |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
112 |
apply (simp add: i_th_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
113 |
apply (case_tac "i_rt n s = UU", auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
114 |
apply (drule i_rt_slen [THEN iffD1]) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
115 |
apply (drule slen_take_eq_rev [rule_format, THEN iffD2],auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
116 |
by (drule not_Undef_is_Def [THEN iffD1], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
117 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
118 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
119 |
lemma fsingleton_lemma1[simp]: "(<a> = <b>) = (a=b)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
120 |
by (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
121 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
122 |
lemma fsingleton_lemma2[simp]: "<a> ~= <>" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
123 |
by (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
124 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
125 |
lemma fsingleton_sconc:"<a> ooo s = Def a && s" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
126 |
by (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
127 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
128 |
lemma fstreams_ind: |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
129 |
"[| adm P; P <>; !!a s. P s ==> P (<a> ooo s) |] ==> P x" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
130 |
apply (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
131 |
apply (rule stream.ind, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
132 |
by (drule not_Undef_is_Def [THEN iffD1], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
133 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
134 |
lemma fstreams_ind2: |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
135 |
"[| adm P; P <>; !!a. P (<a>); !!a b s. P s ==> P (<a> ooo <b> ooo s) |] ==> P x" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
136 |
apply (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
137 |
apply (rule stream_ind2, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
138 |
by (drule not_Undef_is_Def [THEN iffD1], auto)+ |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
139 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
140 |
lemma fstreams_take_Suc[simp]: "stream_take (Suc n)$(<a> ooo s) = <a> ooo stream_take n$s" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
141 |
by (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
142 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
143 |
lemma fstreams_not_empty[simp]: "<a> ooo s ~= <>" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
144 |
by (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
145 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
146 |
lemma fstreams_not_empty2[simp]: "s ooo <a> ~= <>" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
147 |
by (case_tac "s=UU", auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
148 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
149 |
lemma fstreams_exhaust: "x = UU | (EX a s. x = <a> ooo s)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
150 |
apply (simp add: fsingleton_def2, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
151 |
apply (erule contrapos_pp, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
152 |
apply (drule stream_exhaust_eq [THEN iffD1], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
153 |
by (drule not_Undef_is_Def [THEN iffD1], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
154 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
155 |
lemma fstreams_cases: "[| x = UU ==> P; !!a y. x = <a> ooo y ==> P |] ==> P" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
156 |
by (insert fstreams_exhaust [of x], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
157 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
158 |
lemma fstreams_exhaust_eq: "(x ~= UU) = (? a y. x = <a> ooo y)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
159 |
apply (simp add: fsingleton_def2, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
160 |
apply (drule stream_exhaust_eq [THEN iffD1], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
161 |
by (drule not_Undef_is_Def [THEN iffD1], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
162 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
163 |
lemma fstreams_inject: "(<a> ooo s = <b> ooo t) = (a=b & s=t)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
164 |
by (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
165 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
166 |
lemma fstreams_prefix: "<a> ooo s << t ==> EX tt. t = <a> ooo tt & s << tt" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
167 |
apply (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
168 |
apply (insert stream_prefix [of "Def a" s t], auto) |
19550
ae77a20f6995
update to reflect changes in inverts/injects lemmas
huffman
parents:
16746
diff
changeset
|
169 |
by (auto simp add: stream.inverts) |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
170 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
171 |
lemma fstreams_prefix': "x << <a> ooo z = (x = <> | (EX y. x = <a> ooo y & y << z))" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
172 |
apply (auto, case_tac "x=UU", auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
173 |
apply (drule stream_exhaust_eq [THEN iffD1], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
174 |
apply (simp add: fsingleton_def2, auto) |
19550
ae77a20f6995
update to reflect changes in inverts/injects lemmas
huffman
parents:
16746
diff
changeset
|
175 |
apply (auto simp add: stream.inverts) |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
176 |
apply (drule ax_flat [rule_format], simp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
177 |
by (erule sconc_mono) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
178 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
179 |
lemma ft_fstreams[simp]: "ft$(<a> ooo s) = Def a" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
180 |
by (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
181 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
182 |
lemma rt_fstreams[simp]: "rt$(<a> ooo s) = s" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
183 |
by (simp add: fsingleton_def2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
184 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
185 |
lemma ft_eq[simp]: "(ft$s = Def a) = (EX t. s = <a> ooo t)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
186 |
apply (rule stream.casedist [of s],auto) |
16746 | 187 |
by ((*drule sym,*) auto simp add: fsingleton_def2) |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
188 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
189 |
lemma surjective_fstreams: "(<d> ooo y = x) = (ft$x = Def d & rt$x = y)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
190 |
by auto |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
191 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
192 |
lemma fstreams_mono: "<a> ooo b << <a> ooo c ==> b << c" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
193 |
apply (simp add: fsingleton_def2) |
19550
ae77a20f6995
update to reflect changes in inverts/injects lemmas
huffman
parents:
16746
diff
changeset
|
194 |
by (auto simp add: stream.inverts) |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
195 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
196 |
lemma fsmap_UU[simp]: "fsmap f$UU = UU" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
197 |
by (simp add: fsmap_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
198 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
199 |
lemma fsmap_fsingleton_sconc: "fsmap f$(<x> ooo xs) = <(f x)> ooo (fsmap f$xs)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
200 |
by (simp add: fsmap_def fsingleton_def2 flift2_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
201 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
202 |
lemma fsmap_fsingleton[simp]: "fsmap f$(<x>) = <(f x)>" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
203 |
by (simp add: fsmap_def fsingleton_def2 flift2_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
204 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
205 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
206 |
declare range_composition[simp del] |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
207 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
208 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
209 |
lemma fstreams_chain_lemma[rule_format]: |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
210 |
"ALL s x y. stream_take n$(s::'a fstream) << x & x << y & y << s & x ~= y --> stream_take (Suc n)$s << y" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
211 |
apply (induct_tac n, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
212 |
apply (case_tac "s=UU", auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
213 |
apply (drule stream_exhaust_eq [THEN iffD1], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
214 |
apply (case_tac "y=UU", auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
215 |
apply (drule stream_exhaust_eq [THEN iffD1], auto) |
19550
ae77a20f6995
update to reflect changes in inverts/injects lemmas
huffman
parents:
16746
diff
changeset
|
216 |
apply (auto simp add: stream.inverts) |
ae77a20f6995
update to reflect changes in inverts/injects lemmas
huffman
parents:
16746
diff
changeset
|
217 |
apply (simp add: less_lift) |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
218 |
apply (case_tac "s=UU", auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
219 |
apply (drule stream_exhaust_eq [THEN iffD1], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
220 |
apply (erule_tac x="ya" in allE) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
221 |
apply (drule stream_prefix, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
222 |
apply (case_tac "y=UU",auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
223 |
apply (drule stream_exhaust_eq [THEN iffD1], clarsimp) |
19550
ae77a20f6995
update to reflect changes in inverts/injects lemmas
huffman
parents:
16746
diff
changeset
|
224 |
apply (auto simp add: stream.inverts) |
16746 | 225 |
apply (simp add: less_lift) |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
226 |
apply (erule_tac x="tt" in allE) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
227 |
apply (erule_tac x="yb" in allE, auto) |
16746 | 228 |
apply (simp add: less_lift) |
19550
ae77a20f6995
update to reflect changes in inverts/injects lemmas
huffman
parents:
16746
diff
changeset
|
229 |
by (simp add: less_lift) |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
230 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
231 |
lemma fstreams_lub_lemma1: "[| chain Y; lub (range Y) = <a> ooo s |] ==> EX j t. Y j = <a> ooo t" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
232 |
apply (subgoal_tac "lub(range Y) ~= UU") |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
233 |
apply (drule chain_UU_I_inverse2, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
234 |
apply (drule_tac x="i" in is_ub_thelub, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
235 |
by (drule fstreams_prefix' [THEN iffD1], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
236 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
237 |
lemma fstreams_lub1: |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
238 |
"[| chain Y; lub (range Y) = <a> ooo s |] |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
239 |
==> (EX j t. Y j = <a> ooo t) & (EX X. chain X & (ALL i. EX j. <a> ooo X i << Y j) & lub (range X) = s)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
240 |
apply (auto simp add: fstreams_lub_lemma1) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
241 |
apply (rule_tac x="%n. stream_take n$s" in exI, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
242 |
apply (simp add: chain_stream_take) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
243 |
apply (induct_tac i, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
244 |
apply (drule fstreams_lub_lemma1, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
245 |
apply (rule_tac x="j" in exI, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
246 |
apply (case_tac "max_in_chain j Y") |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
247 |
apply (frule lub_finch1 [THEN thelubI], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
248 |
apply (rule_tac x="j" in exI) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
249 |
apply (erule subst) back back |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
250 |
apply (simp add: less_cprod_def sconc_mono) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
251 |
apply (simp add: max_in_chain_def, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
252 |
apply (rule_tac x="ja" in exI) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
253 |
apply (subgoal_tac "Y j << Y ja") |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
254 |
apply (drule fstreams_prefix, auto)+ |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
255 |
apply (rule sconc_mono) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
256 |
apply (rule fstreams_chain_lemma, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
257 |
apply (subgoal_tac "Y ja << (LUB i. (Y i))", clarsimp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
258 |
apply (drule fstreams_mono, simp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
259 |
apply (rule is_ub_thelub, simp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
260 |
apply (blast intro: chain_mono3) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
261 |
by (rule stream_reach2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
262 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
263 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
264 |
lemma lub_Pair_not_UU_lemma: |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
265 |
"[| chain Y; lub (range Y) = ((a::'a::flat), b); a ~= UU; b ~= UU |] |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
266 |
==> EX j c d. Y j = (c, d) & c ~= UU & d ~= UU" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
267 |
apply (frule thelub_cprod, clarsimp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
268 |
apply (drule chain_UU_I_inverse2, clarsimp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
269 |
apply (case_tac "Y i", clarsimp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
270 |
apply (case_tac "max_in_chain i Y") |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
271 |
apply (drule maxinch_is_thelub, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
272 |
apply (rule_tac x="i" in exI, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
273 |
apply (simp add: max_in_chain_def, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
274 |
apply (subgoal_tac "Y i << Y j",auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
275 |
apply (simp add: less_cprod_def, clarsimp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
276 |
apply (drule ax_flat [rule_format], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
277 |
apply (case_tac "snd (Y j) = UU",auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
278 |
apply (case_tac "Y j", auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
279 |
apply (rule_tac x="j" in exI) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
280 |
apply (case_tac "Y j",auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
281 |
by (drule chain_mono3, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
282 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
283 |
lemma fstreams_lub_lemma2: |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
284 |
"[| chain Y; lub (range Y) = (a, <m> ooo ms); (a::'a::flat) ~= UU |] ==> EX j t. Y j = (a, <m> ooo t)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
285 |
apply (frule lub_Pair_not_UU_lemma, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
286 |
apply (drule_tac x="j" in is_ub_thelub, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
287 |
apply (simp add: less_cprod_def, clarsimp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
288 |
apply (drule ax_flat [rule_format], clarsimp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
289 |
by (drule fstreams_prefix' [THEN iffD1], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
290 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
291 |
lemma fstreams_lub2: |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
292 |
"[| chain Y; lub (range Y) = (a, <m> ooo ms); (a::'a::flat) ~= UU |] |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
293 |
==> (EX j t. Y j = (a, <m> ooo t)) & (EX X. chain X & (ALL i. EX j. (a, <m> ooo X i) << Y j) & lub (range X) = ms)" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
294 |
apply (auto simp add: fstreams_lub_lemma2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
295 |
apply (rule_tac x="%n. stream_take n$ms" in exI, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
296 |
apply (simp add: chain_stream_take) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
297 |
apply (induct_tac i, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
298 |
apply (drule fstreams_lub_lemma2, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
299 |
apply (rule_tac x="j" in exI, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
300 |
apply (simp add: less_cprod_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
301 |
apply (case_tac "max_in_chain j Y") |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
302 |
apply (frule lub_finch1 [THEN thelubI], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
303 |
apply (rule_tac x="j" in exI) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
304 |
apply (erule subst) back back |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
305 |
apply (simp add: less_cprod_def sconc_mono) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
306 |
apply (simp add: max_in_chain_def, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
307 |
apply (rule_tac x="ja" in exI) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
308 |
apply (subgoal_tac "Y j << Y ja") |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
309 |
apply (simp add: less_cprod_def, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
310 |
apply (drule trans_less) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
311 |
apply (simp add: ax_flat, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
312 |
apply (drule fstreams_prefix, auto)+ |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
313 |
apply (rule sconc_mono) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
314 |
apply (subgoal_tac "tt ~= tta" "tta << ms") |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
315 |
apply (blast intro: fstreams_chain_lemma) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
316 |
apply (frule lub_cprod [THEN thelubI], auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
317 |
apply (subgoal_tac "snd (Y ja) << (LUB i. snd (Y i))", clarsimp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
318 |
apply (drule fstreams_mono, simp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
319 |
apply (rule is_ub_thelub chainI) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
320 |
apply (simp add: chain_def less_cprod_def) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
321 |
apply (subgoal_tac "fst (Y j) ~= fst (Y ja) | snd (Y j) ~= snd (Y ja)", simp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
322 |
apply (drule ax_flat[rule_format], simp)+ |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
323 |
apply (drule prod_eqI, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
324 |
apply (simp add: chain_mono3) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
325 |
by (rule stream_reach2) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
326 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
327 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
328 |
lemma cpo_cont_lemma: |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
329 |
"[| monofun (f::'a::cpo => 'b::cpo); (!Y. chain Y --> f (lub(range Y)) << (LUB i. f (Y i))) |] ==> cont f" |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
330 |
apply (rule monocontlub2cont, auto) |
16219 | 331 |
apply (simp add: contlub_def, auto) |
15188
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
332 |
apply (erule_tac x="Y" in allE, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
333 |
apply (simp add: po_eq_conv) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
334 |
apply (frule cpo,auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
335 |
apply (frule is_lubD1) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
336 |
apply (frule ub2ub_monofun, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
337 |
apply (drule thelubI, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
338 |
apply (rule is_lub_thelub, auto) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
339 |
by (erule ch2ch_monofun, simp) |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
340 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
341 |
end |
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
342 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
343 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
344 |
|
9d57263faf9e
integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy
oheimb
parents:
diff
changeset
|
345 |