author | wenzelm |
Tue, 01 Oct 2024 20:39:16 +0200 | |
changeset 81091 | c007e6d9941d |
parent 81089 | 8042869c2072 |
child 81095 | 49c04500c5f9 |
permissions | -rw-r--r-- |
62008 | 1 |
(* Title: HOL/HOLCF/IOA/Sequence.thy |
40945 | 2 |
Author: Olaf Müller |
62116 | 3 |
*) |
3071 | 4 |
|
62116 | 5 |
section \<open>Sequences over flat domains with lifted elements\<close> |
3071 | 6 |
|
17233 | 7 |
theory Sequence |
8 |
imports Seq |
|
9 |
begin |
|
3071 | 10 |
|
36452 | 11 |
default_sort type |
17233 | 12 |
|
41476 | 13 |
type_synonym 'a Seq = "'a lift seq" |
3071 | 14 |
|
62005 | 15 |
definition Consq :: "'a \<Rightarrow> 'a Seq \<rightarrow> 'a Seq" |
16 |
where "Consq a = (LAM s. Def a ## s)" |
|
17 |
||
18 |
definition Filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq" |
|
63549 | 19 |
where "Filter P = sfilter \<cdot> (flift2 P)" |
3071 | 20 |
|
62005 | 21 |
definition Map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a Seq \<rightarrow> 'b Seq" |
63549 | 22 |
where "Map f = smap \<cdot> (flift2 f)" |
3071 | 23 |
|
62005 | 24 |
definition Forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<Rightarrow> bool" |
25 |
where "Forall P = sforall (flift2 P)" |
|
3071 | 26 |
|
62005 | 27 |
definition Last :: "'a Seq \<rightarrow> 'a lift" |
28 |
where "Last = slast" |
|
3071 | 29 |
|
62005 | 30 |
definition Dropwhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq" |
63549 | 31 |
where "Dropwhile P = sdropwhile \<cdot> (flift2 P)" |
3071 | 32 |
|
62005 | 33 |
definition Takewhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq" |
63549 | 34 |
where "Takewhile P = stakewhile \<cdot> (flift2 P)" |
3071 | 35 |
|
62005 | 36 |
definition Zip :: "'a Seq \<rightarrow> 'b Seq \<rightarrow> ('a * 'b) Seq" |
62116 | 37 |
where "Zip = |
63549 | 38 |
(fix \<cdot> (LAM h t1 t2. |
62116 | 39 |
case t1 of |
40 |
nil \<Rightarrow> nil |
|
41 |
| x ## xs \<Rightarrow> |
|
42 |
(case t2 of |
|
43 |
nil \<Rightarrow> UU |
|
44 |
| y ## ys \<Rightarrow> |
|
45 |
(case x of |
|
46 |
UU \<Rightarrow> UU |
|
47 |
| Def a \<Rightarrow> |
|
48 |
(case y of |
|
49 |
UU \<Rightarrow> UU |
|
63549 | 50 |
| Def b \<Rightarrow> Def (a, b) ## (h \<cdot> xs \<cdot> ys))))))" |
3071 | 51 |
|
62005 | 52 |
definition Flat :: "'a Seq seq \<rightarrow> 'a Seq" |
53 |
where "Flat = sflat" |
|
54 |
||
55 |
definition Filter2 :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Seq \<rightarrow> 'a Seq" |
|
62116 | 56 |
where "Filter2 P = |
63549 | 57 |
(fix \<cdot> |
62116 | 58 |
(LAM h t. |
59 |
case t of |
|
60 |
nil \<Rightarrow> nil |
|
61 |
| x ## xs \<Rightarrow> |
|
62 |
(case x of |
|
63 |
UU \<Rightarrow> UU |
|
63549 | 64 |
| Def y \<Rightarrow> (if P y then x ## (h \<cdot> xs) else h \<cdot> xs))))" |
3071 | 65 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
66 |
abbreviation Consq_syn (\<open>(_/\<leadsto>_)\<close> [66, 65] 65) |
63549 | 67 |
where "a \<leadsto> s \<equiv> Consq a \<cdot> s" |
62005 | 68 |
|
69 |
||
62116 | 70 |
subsection \<open>List enumeration\<close> |
71 |
||
62005 | 72 |
syntax |
81089
8042869c2072
clarified syntax: prefer nonterminal "args", use outer block (with indent);
wenzelm
parents:
81019
diff
changeset
|
73 |
"_totlist" :: "args \<Rightarrow> 'a Seq" (\<open>(\<open>indent=1 notation=\<open>mixfix total list enumeration\<close>\<close>[_!])\<close>) |
8042869c2072
clarified syntax: prefer nonterminal "args", use outer block (with indent);
wenzelm
parents:
81019
diff
changeset
|
74 |
"_partlist" :: "args \<Rightarrow> 'a Seq" (\<open>(\<open>indent=1 notation=\<open>mixfix partial list enumeration\<close>\<close>[_?])\<close>) |
62005 | 75 |
translations |
62116 | 76 |
"[x, xs!]" \<rightleftharpoons> "x \<leadsto> [xs!]" |
77 |
"[x!]" \<rightleftharpoons> "x\<leadsto>nil" |
|
78 |
"[x, xs?]" \<rightleftharpoons> "x \<leadsto> [xs?]" |
|
79 |
"[x?]" \<rightleftharpoons> "x \<leadsto> CONST bottom" |
|
62005 | 80 |
|
81 |
||
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
82 |
declare andalso_and [simp] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
83 |
declare andalso_or [simp] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
84 |
|
62005 | 85 |
|
62116 | 86 |
subsection \<open>Recursive equations of operators\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
87 |
|
62116 | 88 |
subsubsection \<open>Map\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
89 |
|
63549 | 90 |
lemma Map_UU: "Map f \<cdot> UU = UU" |
62005 | 91 |
by (simp add: Map_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
92 |
|
63549 | 93 |
lemma Map_nil: "Map f \<cdot> nil = nil" |
62005 | 94 |
by (simp add: Map_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
95 |
|
63549 | 96 |
lemma Map_cons: "Map f \<cdot> (x \<leadsto> xs) = (f x) \<leadsto> Map f \<cdot> xs" |
62005 | 97 |
by (simp add: Map_def Consq_def flift2_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
98 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
99 |
|
62002 | 100 |
subsubsection \<open>Filter\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
101 |
|
63549 | 102 |
lemma Filter_UU: "Filter P \<cdot> UU = UU" |
62005 | 103 |
by (simp add: Filter_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
104 |
|
63549 | 105 |
lemma Filter_nil: "Filter P \<cdot> nil = nil" |
62005 | 106 |
by (simp add: Filter_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
107 |
|
63549 | 108 |
lemma Filter_cons: "Filter P \<cdot> (x \<leadsto> xs) = (if P x then x \<leadsto> (Filter P \<cdot> xs) else Filter P \<cdot> xs)" |
62005 | 109 |
by (simp add: Filter_def Consq_def flift2_def If_and_if) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
110 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
111 |
|
62002 | 112 |
subsubsection \<open>Forall\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
113 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
114 |
lemma Forall_UU: "Forall P UU" |
62005 | 115 |
by (simp add: Forall_def sforall_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
116 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
117 |
lemma Forall_nil: "Forall P nil" |
62005 | 118 |
by (simp add: Forall_def sforall_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
119 |
|
62116 | 120 |
lemma Forall_cons: "Forall P (x \<leadsto> xs) = (P x \<and> Forall P xs)" |
62005 | 121 |
by (simp add: Forall_def sforall_def Consq_def flift2_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
122 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
123 |
|
62002 | 124 |
subsubsection \<open>Conc\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
125 |
|
62116 | 126 |
lemma Conc_cons: "(x \<leadsto> xs) @@ y = x \<leadsto> (xs @@ y)" |
62005 | 127 |
by (simp add: Consq_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
128 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
129 |
|
62002 | 130 |
subsubsection \<open>Takewhile\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
131 |
|
63549 | 132 |
lemma Takewhile_UU: "Takewhile P \<cdot> UU = UU" |
62005 | 133 |
by (simp add: Takewhile_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
134 |
|
63549 | 135 |
lemma Takewhile_nil: "Takewhile P \<cdot> nil = nil" |
62005 | 136 |
by (simp add: Takewhile_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
137 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
138 |
lemma Takewhile_cons: |
63549 | 139 |
"Takewhile P \<cdot> (x \<leadsto> xs) = (if P x then x \<leadsto> (Takewhile P \<cdot> xs) else nil)" |
62005 | 140 |
by (simp add: Takewhile_def Consq_def flift2_def If_and_if) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
141 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
142 |
|
62002 | 143 |
subsubsection \<open>DropWhile\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
144 |
|
63549 | 145 |
lemma Dropwhile_UU: "Dropwhile P \<cdot> UU = UU" |
62005 | 146 |
by (simp add: Dropwhile_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
147 |
|
63549 | 148 |
lemma Dropwhile_nil: "Dropwhile P \<cdot> nil = nil" |
62005 | 149 |
by (simp add: Dropwhile_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
150 |
|
63549 | 151 |
lemma Dropwhile_cons: "Dropwhile P \<cdot> (x \<leadsto> xs) = (if P x then Dropwhile P \<cdot> xs else x \<leadsto> xs)" |
62005 | 152 |
by (simp add: Dropwhile_def Consq_def flift2_def If_and_if) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
153 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
154 |
|
62002 | 155 |
subsubsection \<open>Last\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
156 |
|
63549 | 157 |
lemma Last_UU: "Last \<cdot> UU = UU" |
62005 | 158 |
by (simp add: Last_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
159 |
|
63549 | 160 |
lemma Last_nil: "Last \<cdot> nil = UU" |
62116 | 161 |
by (simp add: Last_def) |
162 |
||
63549 | 163 |
lemma Last_cons: "Last \<cdot> (x \<leadsto> xs) = (if xs = nil then Def x else Last \<cdot> xs)" |
62116 | 164 |
by (cases xs) (simp_all add: Last_def Consq_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
165 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
166 |
|
62002 | 167 |
subsubsection \<open>Flat\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
168 |
|
63549 | 169 |
lemma Flat_UU: "Flat \<cdot> UU = UU" |
62005 | 170 |
by (simp add: Flat_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
171 |
|
63549 | 172 |
lemma Flat_nil: "Flat \<cdot> nil = nil" |
62005 | 173 |
by (simp add: Flat_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
174 |
|
63549 | 175 |
lemma Flat_cons: "Flat \<cdot> (x ## xs) = x @@ (Flat \<cdot> xs)" |
62005 | 176 |
by (simp add: Flat_def Consq_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
177 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
178 |
|
62002 | 179 |
subsubsection \<open>Zip\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
180 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
181 |
lemma Zip_unfold: |
62116 | 182 |
"Zip = |
183 |
(LAM t1 t2. |
|
184 |
case t1 of |
|
185 |
nil \<Rightarrow> nil |
|
186 |
| x ## xs \<Rightarrow> |
|
187 |
(case t2 of |
|
188 |
nil \<Rightarrow> UU |
|
189 |
| y ## ys \<Rightarrow> |
|
190 |
(case x of |
|
191 |
UU \<Rightarrow> UU |
|
192 |
| Def a \<Rightarrow> |
|
193 |
(case y of |
|
194 |
UU \<Rightarrow> UU |
|
63549 | 195 |
| Def b \<Rightarrow> Def (a, b) ## (Zip \<cdot> xs \<cdot> ys)))))" |
62005 | 196 |
apply (rule trans) |
197 |
apply (rule fix_eq4) |
|
198 |
apply (rule Zip_def) |
|
199 |
apply (rule beta_cfun) |
|
200 |
apply simp |
|
201 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
202 |
|
63549 | 203 |
lemma Zip_UU1: "Zip \<cdot> UU \<cdot> y = UU" |
62005 | 204 |
apply (subst Zip_unfold) |
205 |
apply simp |
|
206 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
207 |
|
63549 | 208 |
lemma Zip_UU2: "x \<noteq> nil \<Longrightarrow> Zip \<cdot> x \<cdot> UU = UU" |
62005 | 209 |
apply (subst Zip_unfold) |
210 |
apply simp |
|
211 |
apply (cases x) |
|
212 |
apply simp_all |
|
213 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
214 |
|
63549 | 215 |
lemma Zip_nil: "Zip \<cdot> nil \<cdot> y = nil" |
62005 | 216 |
apply (subst Zip_unfold) |
217 |
apply simp |
|
218 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
219 |
|
63549 | 220 |
lemma Zip_cons_nil: "Zip \<cdot> (x \<leadsto> xs) \<cdot> nil = UU" |
62005 | 221 |
apply (subst Zip_unfold) |
222 |
apply (simp add: Consq_def) |
|
223 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
224 |
|
63549 | 225 |
lemma Zip_cons: "Zip \<cdot> (x \<leadsto> xs) \<cdot> (y \<leadsto> ys) = (x, y) \<leadsto> Zip \<cdot> xs \<cdot> ys" |
62005 | 226 |
apply (rule trans) |
227 |
apply (subst Zip_unfold) |
|
228 |
apply simp |
|
229 |
apply (simp add: Consq_def) |
|
230 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
231 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
232 |
lemmas [simp del] = |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
233 |
sfilter_UU sfilter_nil sfilter_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
234 |
smap_UU smap_nil smap_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
235 |
sforall2_UU sforall2_nil sforall2_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
236 |
slast_UU slast_nil slast_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
237 |
stakewhile_UU stakewhile_nil stakewhile_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
238 |
sdropwhile_UU sdropwhile_nil sdropwhile_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
239 |
sflat_UU sflat_nil sflat_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
240 |
szip_UU1 szip_UU2 szip_nil szip_cons_nil szip_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
241 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
242 |
lemmas [simp] = |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
243 |
Filter_UU Filter_nil Filter_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
244 |
Map_UU Map_nil Map_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
245 |
Forall_UU Forall_nil Forall_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
246 |
Last_UU Last_nil Last_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
247 |
Conc_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
248 |
Takewhile_UU Takewhile_nil Takewhile_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
249 |
Dropwhile_UU Dropwhile_nil Dropwhile_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
250 |
Zip_UU1 Zip_UU2 Zip_nil Zip_cons_nil Zip_cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
251 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
252 |
|
62116 | 253 |
subsection \<open>Cons\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
254 |
|
62116 | 255 |
lemma Consq_def2: "a \<leadsto> s = Def a ## s" |
62005 | 256 |
by (simp add: Consq_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
257 |
|
62116 | 258 |
lemma Seq_exhaust: "x = UU \<or> x = nil \<or> (\<exists>a s. x = a \<leadsto> s)" |
62005 | 259 |
apply (simp add: Consq_def2) |
260 |
apply (cut_tac seq.nchotomy) |
|
261 |
apply (fast dest: not_Undef_is_Def [THEN iffD1]) |
|
262 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
263 |
|
62116 | 264 |
lemma Seq_cases: obtains "x = UU" | "x = nil" | a s where "x = a \<leadsto> s" |
62005 | 265 |
apply (cut_tac x="x" in Seq_exhaust) |
266 |
apply (erule disjE) |
|
267 |
apply simp |
|
268 |
apply (erule disjE) |
|
269 |
apply simp |
|
270 |
apply (erule exE)+ |
|
271 |
apply simp |
|
272 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
273 |
|
62116 | 274 |
lemma Cons_not_UU: "a \<leadsto> s \<noteq> UU" |
62005 | 275 |
apply (subst Consq_def2) |
276 |
apply simp |
|
277 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
278 |
|
62116 | 279 |
lemma Cons_not_less_UU: "\<not> (a \<leadsto> x) \<sqsubseteq> UU" |
62005 | 280 |
apply (rule notI) |
281 |
apply (drule below_antisym) |
|
282 |
apply simp |
|
283 |
apply (simp add: Cons_not_UU) |
|
284 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
285 |
|
62116 | 286 |
lemma Cons_not_less_nil: "\<not> a \<leadsto> s \<sqsubseteq> nil" |
62005 | 287 |
by (simp add: Consq_def2) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
288 |
|
62116 | 289 |
lemma Cons_not_nil: "a \<leadsto> s \<noteq> nil" |
62005 | 290 |
by (simp add: Consq_def2) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
291 |
|
62116 | 292 |
lemma Cons_not_nil2: "nil \<noteq> a \<leadsto> s" |
62005 | 293 |
by (simp add: Consq_def2) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
294 |
|
62116 | 295 |
lemma Cons_inject_eq: "a \<leadsto> s = b \<leadsto> t \<longleftrightarrow> a = b \<and> s = t" |
296 |
by (simp add: Consq_def2 scons_inject_eq) |
|
297 |
||
298 |
lemma Cons_inject_less_eq: "a \<leadsto> s \<sqsubseteq> b \<leadsto> t \<longleftrightarrow> a = b \<and> s \<sqsubseteq> t" |
|
299 |
by (simp add: Consq_def2) |
|
300 |
||
63549 | 301 |
lemma seq_take_Cons: "seq_take (Suc n) \<cdot> (a \<leadsto> x) = a \<leadsto> (seq_take n \<cdot> x)" |
62005 | 302 |
by (simp add: Consq_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
303 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
304 |
lemmas [simp] = |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
305 |
Cons_not_nil2 Cons_inject_eq Cons_inject_less_eq seq_take_Cons |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
306 |
Cons_not_UU Cons_not_less_UU Cons_not_less_nil Cons_not_nil |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
307 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
308 |
|
62116 | 309 |
subsection \<open>Induction\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
310 |
|
62116 | 311 |
lemma Seq_induct: |
312 |
assumes "adm P" |
|
313 |
and "P UU" |
|
314 |
and "P nil" |
|
315 |
and "\<And>a s. P s \<Longrightarrow> P (a \<leadsto> s)" |
|
316 |
shows "P x" |
|
317 |
apply (insert assms) |
|
62005 | 318 |
apply (erule (2) seq.induct) |
319 |
apply defined |
|
320 |
apply (simp add: Consq_def) |
|
321 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
322 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
323 |
lemma Seq_FinitePartial_ind: |
62116 | 324 |
assumes "P UU" |
325 |
and "P nil" |
|
326 |
and "\<And>a s. P s \<Longrightarrow> P (a \<leadsto> s)" |
|
327 |
shows "seq_finite x \<longrightarrow> P x" |
|
328 |
apply (insert assms) |
|
62005 | 329 |
apply (erule (1) seq_finite_ind) |
330 |
apply defined |
|
331 |
apply (simp add: Consq_def) |
|
332 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
333 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
334 |
lemma Seq_Finite_ind: |
62116 | 335 |
assumes "Finite x" |
336 |
and "P nil" |
|
337 |
and "\<And>a s. Finite s \<Longrightarrow> P s \<Longrightarrow> P (a \<leadsto> s)" |
|
338 |
shows "P x" |
|
339 |
apply (insert assms) |
|
62005 | 340 |
apply (erule (1) Finite.induct) |
341 |
apply defined |
|
342 |
apply (simp add: Consq_def) |
|
343 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
344 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
345 |
|
62116 | 346 |
subsection \<open>\<open>HD\<close> and \<open>TL\<close>\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
347 |
|
63549 | 348 |
lemma HD_Cons [simp]: "HD \<cdot> (x \<leadsto> y) = Def x" |
62116 | 349 |
by (simp add: Consq_def) |
350 |
||
63549 | 351 |
lemma TL_Cons [simp]: "TL \<cdot> (x \<leadsto> y) = y" |
62116 | 352 |
by (simp add: Consq_def) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
353 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
354 |
|
62116 | 355 |
subsection \<open>\<open>Finite\<close>, \<open>Partial\<close>, \<open>Infinite\<close>\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
356 |
|
62116 | 357 |
lemma Finite_Cons [simp]: "Finite (a \<leadsto> xs) = Finite xs" |
62005 | 358 |
by (simp add: Consq_def2 Finite_cons) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
359 |
|
62116 | 360 |
lemma FiniteConc_1: "Finite (x::'a Seq) \<Longrightarrow> Finite y \<longrightarrow> Finite (x @@ y)" |
62005 | 361 |
apply (erule Seq_Finite_ind) |
362 |
apply simp_all |
|
363 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
364 |
|
62116 | 365 |
lemma FiniteConc_2: "Finite (z::'a Seq) \<Longrightarrow> \<forall>x y. z = x @@ y \<longrightarrow> Finite x \<and> Finite y" |
62005 | 366 |
apply (erule Seq_Finite_ind) |
62116 | 367 |
text \<open>\<open>nil\<close>\<close> |
62005 | 368 |
apply (intro strip) |
369 |
apply (rule_tac x="x" in Seq_cases, simp_all) |
|
62116 | 370 |
text \<open>\<open>cons\<close>\<close> |
62005 | 371 |
apply (intro strip) |
372 |
apply (rule_tac x="x" in Seq_cases, simp_all) |
|
373 |
apply (rule_tac x="y" in Seq_cases, simp_all) |
|
374 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
375 |
|
62116 | 376 |
lemma FiniteConc [simp]: "Finite (x @@ y) \<longleftrightarrow> Finite (x::'a Seq) \<and> Finite y" |
62005 | 377 |
apply (rule iffI) |
378 |
apply (erule FiniteConc_2 [rule_format]) |
|
379 |
apply (rule refl) |
|
380 |
apply (rule FiniteConc_1 [rule_format]) |
|
381 |
apply auto |
|
382 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
383 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
384 |
|
63549 | 385 |
lemma FiniteMap1: "Finite s \<Longrightarrow> Finite (Map f \<cdot> s)" |
62005 | 386 |
apply (erule Seq_Finite_ind) |
387 |
apply simp_all |
|
388 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
389 |
|
63549 | 390 |
lemma FiniteMap2: "Finite s \<Longrightarrow> \<forall>t. s = Map f \<cdot> t \<longrightarrow> Finite t" |
62005 | 391 |
apply (erule Seq_Finite_ind) |
392 |
apply (intro strip) |
|
393 |
apply (rule_tac x="t" in Seq_cases, simp_all) |
|
62116 | 394 |
text \<open>\<open>main case\<close>\<close> |
62005 | 395 |
apply auto |
396 |
apply (rule_tac x="t" in Seq_cases, simp_all) |
|
397 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
398 |
|
63549 | 399 |
lemma Map2Finite: "Finite (Map f \<cdot> s) = Finite s" |
62005 | 400 |
apply auto |
401 |
apply (erule FiniteMap2 [rule_format]) |
|
402 |
apply (rule refl) |
|
403 |
apply (erule FiniteMap1) |
|
404 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
405 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
406 |
|
63549 | 407 |
lemma FiniteFilter: "Finite s \<Longrightarrow> Finite (Filter P \<cdot> s)" |
62005 | 408 |
apply (erule Seq_Finite_ind) |
409 |
apply simp_all |
|
410 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
411 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
412 |
|
62116 | 413 |
subsection \<open>\<open>Conc\<close>\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
414 |
|
62116 | 415 |
lemma Conc_cong: "\<And>x::'a Seq. Finite x \<Longrightarrow> (x @@ y) = (x @@ z) \<longleftrightarrow> y = z" |
62005 | 416 |
apply (erule Seq_Finite_ind) |
417 |
apply simp_all |
|
418 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
419 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
420 |
lemma Conc_assoc: "(x @@ y) @@ z = (x::'a Seq) @@ y @@ z" |
62005 | 421 |
apply (rule_tac x="x" in Seq_induct) |
422 |
apply simp_all |
|
423 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
424 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
425 |
lemma nilConc [simp]: "s@@ nil = s" |
62005 | 426 |
apply (induct s) |
427 |
apply simp |
|
428 |
apply simp |
|
429 |
apply simp |
|
430 |
apply simp |
|
431 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
432 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
433 |
|
62116 | 434 |
(*Should be same as nil_is_Conc2 when all nils are turned to right side!*) |
435 |
lemma nil_is_Conc: "nil = x @@ y \<longleftrightarrow> (x::'a Seq) = nil \<and> y = nil" |
|
62005 | 436 |
apply (rule_tac x="x" in Seq_cases) |
437 |
apply auto |
|
438 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
439 |
|
62116 | 440 |
lemma nil_is_Conc2: "x @@ y = nil \<longleftrightarrow> (x::'a Seq) = nil \<and> y = nil" |
62005 | 441 |
apply (rule_tac x="x" in Seq_cases) |
442 |
apply auto |
|
443 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
444 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
445 |
|
62116 | 446 |
subsection \<open>Last\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
447 |
|
63549 | 448 |
lemma Finite_Last1: "Finite s \<Longrightarrow> s \<noteq> nil \<longrightarrow> Last \<cdot> s \<noteq> UU" |
62116 | 449 |
by (erule Seq_Finite_ind) simp_all |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
450 |
|
63549 | 451 |
lemma Finite_Last2: "Finite s \<Longrightarrow> Last \<cdot> s = UU \<longrightarrow> s = nil" |
62116 | 452 |
by (erule Seq_Finite_ind) auto |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
453 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
454 |
|
62116 | 455 |
subsection \<open>Filter, Conc\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
456 |
|
63549 | 457 |
lemma FilterPQ: "Filter P \<cdot> (Filter Q \<cdot> s) = Filter (\<lambda>x. P x \<and> Q x) \<cdot> s" |
62116 | 458 |
by (rule_tac x="s" in Seq_induct) simp_all |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
459 |
|
63549 | 460 |
lemma FilterConc: "Filter P \<cdot> (x @@ y) = (Filter P \<cdot> x @@ Filter P \<cdot> y)" |
62116 | 461 |
by (simp add: Filter_def sfiltersconc) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
462 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
463 |
|
62116 | 464 |
subsection \<open>Map\<close> |
465 |
||
63549 | 466 |
lemma MapMap: "Map f \<cdot> (Map g \<cdot> s) = Map (f \<circ> g) \<cdot> s" |
62116 | 467 |
by (rule_tac x="s" in Seq_induct) simp_all |
468 |
||
63549 | 469 |
lemma MapConc: "Map f \<cdot> (x @@ y) = (Map f \<cdot> x) @@ (Map f \<cdot> y)" |
62116 | 470 |
by (rule_tac x="x" in Seq_induct) simp_all |
471 |
||
63549 | 472 |
lemma MapFilter: "Filter P \<cdot> (Map f \<cdot> x) = Map f \<cdot> (Filter (P \<circ> f) \<cdot> x)" |
62116 | 473 |
by (rule_tac x="x" in Seq_induct) simp_all |
474 |
||
63549 | 475 |
lemma nilMap: "nil = (Map f \<cdot> s) \<longrightarrow> s = nil" |
62116 | 476 |
by (rule_tac x="s" in Seq_cases) simp_all |
477 |
||
63549 | 478 |
lemma ForallMap: "Forall P (Map f \<cdot> s) = Forall (P \<circ> f) s" |
62005 | 479 |
apply (rule_tac x="s" in Seq_induct) |
480 |
apply (simp add: Forall_def sforall_def) |
|
481 |
apply simp_all |
|
482 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
483 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
484 |
|
62116 | 485 |
subsection \<open>Forall\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
486 |
|
62116 | 487 |
lemma ForallPForallQ1: "Forall P ys \<and> (\<forall>x. P x \<longrightarrow> Q x) \<longrightarrow> Forall Q ys" |
62005 | 488 |
apply (rule_tac x="ys" in Seq_induct) |
489 |
apply (simp add: Forall_def sforall_def) |
|
490 |
apply simp_all |
|
491 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
492 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
493 |
lemmas ForallPForallQ = |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
494 |
ForallPForallQ1 [THEN mp, OF conjI, OF _ allI, OF _ impI] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
495 |
|
62116 | 496 |
lemma Forall_Conc_impl: "Forall P x \<and> Forall P y \<longrightarrow> Forall P (x @@ y)" |
62005 | 497 |
apply (rule_tac x="x" in Seq_induct) |
498 |
apply (simp add: Forall_def sforall_def) |
|
499 |
apply simp_all |
|
500 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
501 |
|
62116 | 502 |
lemma Forall_Conc [simp]: "Finite x \<Longrightarrow> Forall P (x @@ y) \<longleftrightarrow> Forall P x \<and> Forall P y" |
503 |
by (erule Seq_Finite_ind) simp_all |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
504 |
|
63549 | 505 |
lemma ForallTL1: "Forall P s \<longrightarrow> Forall P (TL \<cdot> s)" |
62005 | 506 |
apply (rule_tac x="s" in Seq_induct) |
507 |
apply (simp add: Forall_def sforall_def) |
|
508 |
apply simp_all |
|
509 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
510 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
511 |
lemmas ForallTL = ForallTL1 [THEN mp] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
512 |
|
63549 | 513 |
lemma ForallDropwhile1: "Forall P s \<longrightarrow> Forall P (Dropwhile Q \<cdot> s)" |
62005 | 514 |
apply (rule_tac x="s" in Seq_induct) |
515 |
apply (simp add: Forall_def sforall_def) |
|
516 |
apply simp_all |
|
517 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
518 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
519 |
lemmas ForallDropwhile = ForallDropwhile1 [THEN mp] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
520 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
521 |
|
62116 | 522 |
(*only admissible in t, not if done in s*) |
523 |
lemma Forall_prefix: "\<forall>s. Forall P s \<longrightarrow> t \<sqsubseteq> s \<longrightarrow> Forall P t" |
|
62005 | 524 |
apply (rule_tac x="t" in Seq_induct) |
525 |
apply (simp add: Forall_def sforall_def) |
|
526 |
apply simp_all |
|
527 |
apply (intro strip) |
|
528 |
apply (rule_tac x="sa" in Seq_cases) |
|
529 |
apply simp |
|
530 |
apply auto |
|
531 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
532 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
533 |
lemmas Forall_prefixclosed = Forall_prefix [rule_format] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
534 |
|
62116 | 535 |
lemma Forall_postfixclosed: "Finite h \<Longrightarrow> Forall P s \<Longrightarrow> s= h @@ t \<Longrightarrow> Forall P t" |
62005 | 536 |
by auto |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
537 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
538 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
539 |
lemma ForallPFilterQR1: |
63549 | 540 |
"(\<forall>x. P x \<longrightarrow> Q x = R x) \<and> Forall P tr \<longrightarrow> Filter Q \<cdot> tr = Filter R \<cdot> tr" |
62005 | 541 |
apply (rule_tac x="tr" in Seq_induct) |
542 |
apply (simp add: Forall_def sforall_def) |
|
543 |
apply simp_all |
|
544 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
545 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
546 |
lemmas ForallPFilterQR = ForallPFilterQR1 [THEN mp, OF conjI, OF allI] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
547 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
548 |
|
62116 | 549 |
subsection \<open>Forall, Filter\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
550 |
|
63549 | 551 |
lemma ForallPFilterP: "Forall P (Filter P \<cdot> x)" |
62005 | 552 |
by (simp add: Filter_def Forall_def forallPsfilterP) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
553 |
|
62116 | 554 |
(*holds also in other direction, then equal to forallPfilterP*) |
63549 | 555 |
lemma ForallPFilterPid1: "Forall P x \<longrightarrow> Filter P \<cdot> x = x" |
62005 | 556 |
apply (rule_tac x="x" in Seq_induct) |
557 |
apply (simp add: Forall_def sforall_def Filter_def) |
|
558 |
apply simp_all |
|
559 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
560 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
561 |
lemmas ForallPFilterPid = ForallPFilterPid1 [THEN mp] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
562 |
|
62116 | 563 |
(*holds also in other direction*) |
63549 | 564 |
lemma ForallnPFilterPnil1: "Finite ys \<Longrightarrow> Forall (\<lambda>x. \<not> P x) ys \<longrightarrow> Filter P \<cdot> ys = nil" |
62116 | 565 |
by (erule Seq_Finite_ind) simp_all |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
566 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
567 |
lemmas ForallnPFilterPnil = ForallnPFilterPnil1 [THEN mp] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
568 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
569 |
|
62116 | 570 |
(*holds also in other direction*) |
63549 | 571 |
lemma ForallnPFilterPUU1: "\<not> Finite ys \<and> Forall (\<lambda>x. \<not> P x) ys \<longrightarrow> Filter P \<cdot> ys = UU" |
62005 | 572 |
apply (rule_tac x="ys" in Seq_induct) |
573 |
apply (simp add: Forall_def sforall_def) |
|
574 |
apply simp_all |
|
575 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
576 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
577 |
lemmas ForallnPFilterPUU = ForallnPFilterPUU1 [THEN mp, OF conjI] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
578 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
579 |
|
62116 | 580 |
(*inverse of ForallnPFilterPnil*) |
63549 | 581 |
lemma FilternPnilForallP [rule_format]: "Filter P \<cdot> ys = nil \<longrightarrow> Forall (\<lambda>x. \<not> P x) ys \<and> Finite ys" |
62005 | 582 |
apply (rule_tac x="ys" in Seq_induct) |
62116 | 583 |
text \<open>adm\<close> |
62005 | 584 |
apply (simp add: Forall_def sforall_def) |
62116 | 585 |
text \<open>base cases\<close> |
62005 | 586 |
apply simp |
587 |
apply simp |
|
62116 | 588 |
text \<open>main case\<close> |
62005 | 589 |
apply simp |
590 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
591 |
|
62116 | 592 |
(*inverse of ForallnPFilterPUU*) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
593 |
lemma FilternPUUForallP: |
63549 | 594 |
assumes "Filter P \<cdot> ys = UU" |
62116 | 595 |
shows "Forall (\<lambda>x. \<not> P x) ys \<and> \<not> Finite ys" |
48194
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
596 |
proof |
62116 | 597 |
show "Forall (\<lambda>x. \<not> P x) ys" |
48194
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
598 |
proof (rule classical) |
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
599 |
assume "\<not> ?thesis" |
63549 | 600 |
then have "Filter P \<cdot> ys \<noteq> UU" |
48194
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
601 |
apply (rule rev_mp) |
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
602 |
apply (induct ys rule: Seq_induct) |
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
603 |
apply (simp add: Forall_def sforall_def) |
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
604 |
apply simp_all |
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
605 |
done |
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
606 |
with assms show ?thesis by contradiction |
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
607 |
qed |
62116 | 608 |
show "\<not> Finite ys" |
48194
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
609 |
proof |
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
610 |
assume "Finite ys" |
63549 | 611 |
then have "Filter P\<cdot>ys \<noteq> UU" |
48194
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
612 |
by (rule Seq_Finite_ind) simp_all |
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
613 |
with assms show False by contradiction |
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
614 |
qed |
1440a3103af0
tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents:
44890
diff
changeset
|
615 |
qed |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
616 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
617 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
618 |
lemma ForallQFilterPnil: |
63549 | 619 |
"Forall Q ys \<Longrightarrow> Finite ys \<Longrightarrow> (\<And>x. Q x \<Longrightarrow> \<not> P x) \<Longrightarrow> Filter P \<cdot> ys = nil" |
62005 | 620 |
apply (erule ForallnPFilterPnil) |
621 |
apply (erule ForallPForallQ) |
|
622 |
apply auto |
|
623 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
624 |
|
63549 | 625 |
lemma ForallQFilterPUU: "\<not> Finite ys \<Longrightarrow> Forall Q ys \<Longrightarrow> (\<And>x. Q x \<Longrightarrow> \<not> P x) \<Longrightarrow> Filter P \<cdot> ys = UU" |
62005 | 626 |
apply (erule ForallnPFilterPUU) |
627 |
apply (erule ForallPForallQ) |
|
628 |
apply auto |
|
629 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
630 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
631 |
|
62116 | 632 |
subsection \<open>Takewhile, Forall, Filter\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
633 |
|
63549 | 634 |
lemma ForallPTakewhileP [simp]: "Forall P (Takewhile P \<cdot> x)" |
62005 | 635 |
by (simp add: Forall_def Takewhile_def sforallPstakewhileP) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
636 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
637 |
|
63549 | 638 |
lemma ForallPTakewhileQ [simp]: "(\<And>x. Q x \<Longrightarrow> P x) \<Longrightarrow> Forall P (Takewhile Q \<cdot> x)" |
62005 | 639 |
apply (rule ForallPForallQ) |
640 |
apply (rule ForallPTakewhileP) |
|
641 |
apply auto |
|
642 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
643 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
644 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
645 |
lemma FilterPTakewhileQnil [simp]: |
63549 | 646 |
"Finite (Takewhile Q \<cdot> ys) \<Longrightarrow> (\<And>x. Q x \<Longrightarrow> \<not> P x) \<Longrightarrow> Filter P \<cdot> (Takewhile Q \<cdot> ys) = nil" |
62005 | 647 |
apply (erule ForallnPFilterPnil) |
648 |
apply (rule ForallPForallQ) |
|
649 |
apply (rule ForallPTakewhileP) |
|
650 |
apply auto |
|
651 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
652 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
653 |
lemma FilterPTakewhileQid [simp]: |
63549 | 654 |
"(\<And>x. Q x \<Longrightarrow> P x) \<Longrightarrow> Filter P \<cdot> (Takewhile Q \<cdot> ys) = Takewhile Q \<cdot> ys" |
62005 | 655 |
apply (rule ForallPFilterPid) |
656 |
apply (rule ForallPForallQ) |
|
657 |
apply (rule ForallPTakewhileP) |
|
658 |
apply auto |
|
659 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
660 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
661 |
|
63549 | 662 |
lemma Takewhile_idempotent: "Takewhile P \<cdot> (Takewhile P \<cdot> s) = Takewhile P \<cdot> s" |
62005 | 663 |
apply (rule_tac x="s" in Seq_induct) |
664 |
apply (simp add: Forall_def sforall_def) |
|
665 |
apply simp_all |
|
666 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
667 |
|
62116 | 668 |
lemma ForallPTakewhileQnP [simp]: |
63549 | 669 |
"Forall P s \<longrightarrow> Takewhile (\<lambda>x. Q x \<or> (\<not> P x)) \<cdot> s = Takewhile Q \<cdot> s" |
62005 | 670 |
apply (rule_tac x="s" in Seq_induct) |
671 |
apply (simp add: Forall_def sforall_def) |
|
672 |
apply simp_all |
|
673 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
674 |
|
62116 | 675 |
lemma ForallPDropwhileQnP [simp]: |
63549 | 676 |
"Forall P s \<longrightarrow> Dropwhile (\<lambda>x. Q x \<or> (\<not> P x)) \<cdot> s = Dropwhile Q \<cdot> s" |
62005 | 677 |
apply (rule_tac x="s" in Seq_induct) |
678 |
apply (simp add: Forall_def sforall_def) |
|
679 |
apply simp_all |
|
680 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
681 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
682 |
|
63549 | 683 |
lemma TakewhileConc1: "Forall P s \<longrightarrow> Takewhile P \<cdot> (s @@ t) = s @@ (Takewhile P \<cdot> t)" |
62005 | 684 |
apply (rule_tac x="s" in Seq_induct) |
685 |
apply (simp add: Forall_def sforall_def) |
|
686 |
apply simp_all |
|
687 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
688 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
689 |
lemmas TakewhileConc = TakewhileConc1 [THEN mp] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
690 |
|
63549 | 691 |
lemma DropwhileConc1: "Finite s \<Longrightarrow> Forall P s \<longrightarrow> Dropwhile P \<cdot> (s @@ t) = Dropwhile P \<cdot> t" |
62116 | 692 |
by (erule Seq_Finite_ind) simp_all |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
693 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
694 |
lemmas DropwhileConc = DropwhileConc1 [THEN mp] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
695 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
696 |
|
62116 | 697 |
subsection \<open>Coinductive characterizations of Filter\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
698 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
699 |
lemma divide_Seq_lemma: |
63549 | 700 |
"HD \<cdot> (Filter P \<cdot> y) = Def x \<longrightarrow> |
701 |
y = (Takewhile (\<lambda>x. \<not> P x) \<cdot> y) @@ (x \<leadsto> TL \<cdot> (Dropwhile (\<lambda>a. \<not> P a) \<cdot> y)) \<and> |
|
702 |
Finite (Takewhile (\<lambda>x. \<not> P x) \<cdot> y) \<and> P x" |
|
62005 | 703 |
(* FIX: pay attention: is only admissible with chain-finite package to be added to |
704 |
adm test and Finite f x admissibility *) |
|
705 |
apply (rule_tac x="y" in Seq_induct) |
|
706 |
apply (simp add: adm_subst [OF _ adm_Finite]) |
|
707 |
apply simp |
|
708 |
apply simp |
|
709 |
apply (case_tac "P a") |
|
710 |
apply simp |
|
711 |
apply blast |
|
62116 | 712 |
text \<open>\<open>\<not> P a\<close>\<close> |
62005 | 713 |
apply simp |
714 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
715 |
|
63549 | 716 |
lemma divide_Seq: "(x \<leadsto> xs) \<sqsubseteq> Filter P \<cdot> y \<Longrightarrow> |
717 |
y = ((Takewhile (\<lambda>a. \<not> P a) \<cdot> y) @@ (x \<leadsto> TL \<cdot> (Dropwhile (\<lambda>a. \<not> P a) \<cdot> y))) \<and> |
|
718 |
Finite (Takewhile (\<lambda>a. \<not> P a) \<cdot> y) \<and> P x" |
|
62005 | 719 |
apply (rule divide_Seq_lemma [THEN mp]) |
62116 | 720 |
apply (drule_tac f="HD" and x="x \<leadsto> xs" in monofun_cfun_arg) |
62005 | 721 |
apply simp |
722 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
723 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
724 |
|
63549 | 725 |
lemma nForall_HDFilter: "\<not> Forall P y \<longrightarrow> (\<exists>x. HD \<cdot> (Filter (\<lambda>a. \<not> P a) \<cdot> y) = Def x)" |
62005 | 726 |
unfolding not_Undef_is_Def [symmetric] |
727 |
apply (induct y rule: Seq_induct) |
|
728 |
apply (simp add: Forall_def sforall_def) |
|
729 |
apply simp_all |
|
730 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
731 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
732 |
|
62116 | 733 |
lemma divide_Seq2: |
734 |
"\<not> Forall P y \<Longrightarrow> |
|
63549 | 735 |
\<exists>x. y = Takewhile P\<cdot>y @@ (x \<leadsto> TL \<cdot> (Dropwhile P \<cdot> y)) \<and> Finite (Takewhile P \<cdot> y) \<and> \<not> P x" |
62005 | 736 |
apply (drule nForall_HDFilter [THEN mp]) |
737 |
apply safe |
|
738 |
apply (rule_tac x="x" in exI) |
|
62116 | 739 |
apply (cut_tac P1="\<lambda>x. \<not> P x" in divide_Seq_lemma [THEN mp]) |
62005 | 740 |
apply auto |
741 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
742 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
743 |
|
62116 | 744 |
lemma divide_Seq3: |
745 |
"\<not> Forall P y \<Longrightarrow> \<exists>x bs rs. y = (bs @@ (x\<leadsto>rs)) \<and> Finite bs \<and> Forall P bs \<and> \<not> P x" |
|
62005 | 746 |
apply (drule divide_Seq2) |
747 |
apply fastforce |
|
748 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
749 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
750 |
lemmas [simp] = FilterPQ FilterConc Conc_cong |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
751 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
752 |
|
62116 | 753 |
subsection \<open>Take-lemma\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
754 |
|
63549 | 755 |
lemma seq_take_lemma: "(\<forall>n. seq_take n \<cdot> x = seq_take n \<cdot> x') \<longleftrightarrow> x = x'" |
62005 | 756 |
apply (rule iffI) |
757 |
apply (rule seq.take_lemma) |
|
758 |
apply auto |
|
759 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
760 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
761 |
lemma take_reduction1: |
63549 | 762 |
"\<forall>n. ((\<forall>k. k < n \<longrightarrow> seq_take k \<cdot> y1 = seq_take k \<cdot> y2) \<longrightarrow> |
763 |
seq_take n \<cdot> (x @@ (t \<leadsto> y1)) = seq_take n \<cdot> (x @@ (t \<leadsto> y2)))" |
|
62005 | 764 |
apply (rule_tac x="x" in Seq_induct) |
765 |
apply simp_all |
|
766 |
apply (intro strip) |
|
767 |
apply (case_tac "n") |
|
768 |
apply auto |
|
769 |
apply (case_tac "n") |
|
770 |
apply auto |
|
771 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
772 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
773 |
lemma take_reduction: |
63549 | 774 |
"x = y \<Longrightarrow> s = t \<Longrightarrow> (\<And>k. k < n \<Longrightarrow> seq_take k \<cdot> y1 = seq_take k \<cdot> y2) |
775 |
\<Longrightarrow> seq_take n \<cdot> (x @@ (s \<leadsto> y1)) = seq_take n \<cdot> (y @@ (t \<leadsto> y2))" |
|
62005 | 776 |
by (auto intro!: take_reduction1 [rule_format]) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
777 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
778 |
|
62116 | 779 |
text \<open> |
780 |
Take-lemma and take-reduction for \<open>\<sqsubseteq>\<close> instead of \<open>=\<close>. |
|
781 |
\<close> |
|
782 |
||
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
783 |
lemma take_reduction_less1: |
63549 | 784 |
"\<forall>n. ((\<forall>k. k < n \<longrightarrow> seq_take k \<cdot> y1 \<sqsubseteq> seq_take k\<cdot>y2) \<longrightarrow> |
785 |
seq_take n \<cdot> (x @@ (t \<leadsto> y1)) \<sqsubseteq> seq_take n \<cdot> (x @@ (t \<leadsto> y2)))" |
|
62005 | 786 |
apply (rule_tac x="x" in Seq_induct) |
787 |
apply simp_all |
|
788 |
apply (intro strip) |
|
789 |
apply (case_tac "n") |
|
790 |
apply auto |
|
791 |
apply (case_tac "n") |
|
792 |
apply auto |
|
793 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
794 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
795 |
lemma take_reduction_less: |
63549 | 796 |
"x = y \<Longrightarrow> s = t \<Longrightarrow> (\<And>k. k < n \<Longrightarrow> seq_take k \<cdot> y1 \<sqsubseteq> seq_take k \<cdot> y2) \<Longrightarrow> |
797 |
seq_take n \<cdot> (x @@ (s \<leadsto> y1)) \<sqsubseteq> seq_take n \<cdot> (y @@ (t \<leadsto> y2))" |
|
62005 | 798 |
by (auto intro!: take_reduction_less1 [rule_format]) |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
799 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
800 |
lemma take_lemma_less1: |
63549 | 801 |
assumes "\<And>n. seq_take n \<cdot> s1 \<sqsubseteq> seq_take n \<cdot> s2" |
62116 | 802 |
shows "s1 \<sqsubseteq> s2" |
62005 | 803 |
apply (rule_tac t="s1" in seq.reach [THEN subst]) |
804 |
apply (rule_tac t="s2" in seq.reach [THEN subst]) |
|
805 |
apply (rule lub_mono) |
|
806 |
apply (rule seq.chain_take [THEN ch2ch_Rep_cfunL]) |
|
807 |
apply (rule seq.chain_take [THEN ch2ch_Rep_cfunL]) |
|
808 |
apply (rule assms) |
|
809 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
810 |
|
63549 | 811 |
lemma take_lemma_less: "(\<forall>n. seq_take n \<cdot> x \<sqsubseteq> seq_take n \<cdot> x') \<longleftrightarrow> x \<sqsubseteq> x'" |
62005 | 812 |
apply (rule iffI) |
813 |
apply (rule take_lemma_less1) |
|
814 |
apply auto |
|
815 |
apply (erule monofun_cfun_arg) |
|
816 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
817 |
|
62116 | 818 |
|
819 |
text \<open>Take-lemma proof principles.\<close> |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
820 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
821 |
lemma take_lemma_principle1: |
62116 | 822 |
assumes "\<And>s. Forall Q s \<Longrightarrow> A s \<Longrightarrow> f s = g s" |
823 |
and "\<And>s1 s2 y. Forall Q s1 \<Longrightarrow> Finite s1 \<Longrightarrow> |
|
824 |
\<not> Q y \<Longrightarrow> A (s1 @@ y \<leadsto> s2) \<Longrightarrow> f (s1 @@ y \<leadsto> s2) = g (s1 @@ y \<leadsto> s2)" |
|
825 |
shows "A x \<longrightarrow> f x = g x" |
|
826 |
using assms by (cases "Forall Q x") (auto dest!: divide_Seq3) |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
827 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
828 |
lemma take_lemma_principle2: |
62116 | 829 |
assumes "\<And>s. Forall Q s \<Longrightarrow> A s \<Longrightarrow> f s = g s" |
830 |
and "\<And>s1 s2 y. Forall Q s1 \<Longrightarrow> Finite s1 \<Longrightarrow> \<not> Q y \<Longrightarrow> A (s1 @@ y \<leadsto> s2) \<Longrightarrow> |
|
63549 | 831 |
\<forall>n. seq_take n \<cdot> (f (s1 @@ y \<leadsto> s2)) = seq_take n \<cdot> (g (s1 @@ y \<leadsto> s2))" |
62116 | 832 |
shows "A x \<longrightarrow> f x = g x" |
833 |
using assms |
|
834 |
apply (cases "Forall Q x") |
|
62005 | 835 |
apply (auto dest!: divide_Seq3) |
836 |
apply (rule seq.take_lemma) |
|
837 |
apply auto |
|
838 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
839 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
840 |
|
62116 | 841 |
text \<open> |
842 |
Note: in the following proofs the ordering of proof steps is very important, |
|
843 |
as otherwise either \<open>Forall Q s1\<close> would be in the IH as assumption (then |
|
844 |
rule useless) or it is not possible to strengthen the IH apply doing a |
|
845 |
forall closure of the sequence \<open>t\<close> (then rule also useless). This is also |
|
846 |
the reason why the induction rule (\<open>nat_less_induct\<close> or \<open>nat_induct\<close>) has to |
|
847 |
to be imbuilt into the rule, as induction has to be done early and the take |
|
848 |
lemma has to be used in the trivial direction afterwards for the |
|
849 |
\<open>Forall Q x\<close> case. |
|
850 |
\<close> |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
851 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
852 |
lemma take_lemma_induct: |
62116 | 853 |
assumes "\<And>s. Forall Q s \<Longrightarrow> A s \<Longrightarrow> f s = g s" |
854 |
and "\<And>s1 s2 y n. |
|
63549 | 855 |
\<forall>t. A t \<longrightarrow> seq_take n \<cdot> (f t) = seq_take n \<cdot> (g t) \<Longrightarrow> |
62116 | 856 |
Forall Q s1 \<Longrightarrow> Finite s1 \<Longrightarrow> \<not> Q y \<Longrightarrow> A (s1 @@ y \<leadsto> s2) \<Longrightarrow> |
63549 | 857 |
seq_take (Suc n) \<cdot> (f (s1 @@ y \<leadsto> s2)) = |
858 |
seq_take (Suc n) \<cdot> (g (s1 @@ y \<leadsto> s2))" |
|
62116 | 859 |
shows "A x \<longrightarrow> f x = g x" |
860 |
apply (insert assms) |
|
62005 | 861 |
apply (rule impI) |
862 |
apply (rule seq.take_lemma) |
|
863 |
apply (rule mp) |
|
864 |
prefer 2 apply assumption |
|
865 |
apply (rule_tac x="x" in spec) |
|
866 |
apply (rule nat.induct) |
|
867 |
apply simp |
|
868 |
apply (rule allI) |
|
869 |
apply (case_tac "Forall Q xa") |
|
870 |
apply (force intro!: seq_take_lemma [THEN iffD2, THEN spec]) |
|
871 |
apply (auto dest!: divide_Seq3) |
|
872 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
873 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
874 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
875 |
lemma take_lemma_less_induct: |
62116 | 876 |
assumes "\<And>s. Forall Q s \<Longrightarrow> A s \<Longrightarrow> f s = g s" |
877 |
and "\<And>s1 s2 y n. |
|
63549 | 878 |
\<forall>t m. m < n \<longrightarrow> A t \<longrightarrow> seq_take m \<cdot> (f t) = seq_take m \<cdot> (g t) \<Longrightarrow> |
62116 | 879 |
Forall Q s1 \<Longrightarrow> Finite s1 \<Longrightarrow> \<not> Q y \<Longrightarrow> A (s1 @@ y \<leadsto> s2) \<Longrightarrow> |
63549 | 880 |
seq_take n \<cdot> (f (s1 @@ y \<leadsto> s2)) = |
881 |
seq_take n \<cdot> (g (s1 @@ y \<leadsto> s2))" |
|
62116 | 882 |
shows "A x \<longrightarrow> f x = g x" |
883 |
apply (insert assms) |
|
62005 | 884 |
apply (rule impI) |
885 |
apply (rule seq.take_lemma) |
|
886 |
apply (rule mp) |
|
887 |
prefer 2 apply assumption |
|
888 |
apply (rule_tac x="x" in spec) |
|
889 |
apply (rule nat_less_induct) |
|
890 |
apply (rule allI) |
|
891 |
apply (case_tac "Forall Q xa") |
|
892 |
apply (force intro!: seq_take_lemma [THEN iffD2, THEN spec]) |
|
893 |
apply (auto dest!: divide_Seq3) |
|
894 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
895 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
896 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
897 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
898 |
lemma take_lemma_in_eq_out: |
62116 | 899 |
assumes "A UU \<Longrightarrow> f UU = g UU" |
900 |
and "A nil \<Longrightarrow> f nil = g nil" |
|
901 |
and "\<And>s y n. |
|
63549 | 902 |
\<forall>t. A t \<longrightarrow> seq_take n \<cdot> (f t) = seq_take n \<cdot> (g t) \<Longrightarrow> A (y \<leadsto> s) \<Longrightarrow> |
903 |
seq_take (Suc n) \<cdot> (f (y \<leadsto> s)) = |
|
904 |
seq_take (Suc n) \<cdot> (g (y \<leadsto> s))" |
|
62116 | 905 |
shows "A x \<longrightarrow> f x = g x" |
906 |
apply (insert assms) |
|
62005 | 907 |
apply (rule impI) |
908 |
apply (rule seq.take_lemma) |
|
909 |
apply (rule mp) |
|
910 |
prefer 2 apply assumption |
|
911 |
apply (rule_tac x="x" in spec) |
|
912 |
apply (rule nat.induct) |
|
913 |
apply simp |
|
914 |
apply (rule allI) |
|
915 |
apply (rule_tac x="xa" in Seq_cases) |
|
916 |
apply simp_all |
|
917 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
918 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
919 |
|
62116 | 920 |
subsection \<open>Alternative take_lemma proofs\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
921 |
|
62116 | 922 |
subsubsection \<open>Alternative Proof of FilterPQ\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
923 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
924 |
declare FilterPQ [simp del] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
925 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
926 |
|
62116 | 927 |
(*In general: How to do this case without the same adm problems |
928 |
as for the entire proof?*) |
|
929 |
lemma Filter_lemma1: |
|
930 |
"Forall (\<lambda>x. \<not> (P x \<and> Q x)) s \<longrightarrow> |
|
63549 | 931 |
Filter P \<cdot> (Filter Q \<cdot> s) = Filter (\<lambda>x. P x \<and> Q x) \<cdot> s" |
62005 | 932 |
apply (rule_tac x="s" in Seq_induct) |
933 |
apply (simp add: Forall_def sforall_def) |
|
934 |
apply simp_all |
|
935 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
936 |
|
62116 | 937 |
lemma Filter_lemma2: "Finite s \<Longrightarrow> |
63549 | 938 |
Forall (\<lambda>x. \<not> P x \<or> \<not> Q x) s \<longrightarrow> Filter P \<cdot> (Filter Q \<cdot> s) = nil" |
62116 | 939 |
by (erule Seq_Finite_ind) simp_all |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
940 |
|
62116 | 941 |
lemma Filter_lemma3: "Finite s \<Longrightarrow> |
63549 | 942 |
Forall (\<lambda>x. \<not> P x \<or> \<not> Q x) s \<longrightarrow> Filter (\<lambda>x. P x \<and> Q x) \<cdot> s = nil" |
62116 | 943 |
by (erule Seq_Finite_ind) simp_all |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
944 |
|
63549 | 945 |
lemma FilterPQ_takelemma: "Filter P \<cdot> (Filter Q \<cdot> s) = Filter (\<lambda>x. P x \<and> Q x) \<cdot> s" |
62116 | 946 |
apply (rule_tac A1="\<lambda>x. True" and Q1="\<lambda>x. \<not> (P x \<and> Q x)" and x1="s" in |
947 |
take_lemma_induct [THEN mp]) |
|
948 |
(*better support for A = \<lambda>x. True*) |
|
62005 | 949 |
apply (simp add: Filter_lemma1) |
950 |
apply (simp add: Filter_lemma2 Filter_lemma3) |
|
951 |
apply simp |
|
952 |
done |
|
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
953 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
954 |
declare FilterPQ [simp] |
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
955 |
|
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
956 |
|
62116 | 957 |
subsubsection \<open>Alternative Proof of \<open>MapConc\<close>\<close> |
19551
4103954f3668
converted to isar theory; removed unsound adm_all axiom
huffman
parents:
17233
diff
changeset
|
958 |
|
63549 | 959 |
lemma MapConc_takelemma: "Map f \<cdot> (x @@ y) = (Map f \<cdot> x) @@ (Map f \<cdot> y)" |
62116 | 960 |
apply (rule_tac A1="\<lambda>x. True" and x1="x" in take_lemma_in_eq_out [THEN mp]) |
62005 | 961 |
apply auto |
962 |
done |
|
3071 | 963 |
|
62002 | 964 |
ML \<open> |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27105
diff
changeset
|
965 |
fun Seq_case_tac ctxt s i = |
59780 | 966 |
Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_cases} i |
62116 | 967 |
THEN hyp_subst_tac ctxt i THEN hyp_subst_tac ctxt (i + 1) THEN hyp_subst_tac ctxt (i + 2); |
19741 | 968 |
|
62001 | 969 |
(* on a\<leadsto>s only simp_tac, as full_simp_tac is uncomplete and often causes errors *) |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27105
diff
changeset
|
970 |
fun Seq_case_simp_tac ctxt s i = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48194
diff
changeset
|
971 |
Seq_case_tac ctxt s i |
62116 | 972 |
THEN asm_simp_tac ctxt (i + 2) |
973 |
THEN asm_full_simp_tac ctxt (i + 1) |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48194
diff
changeset
|
974 |
THEN asm_full_simp_tac ctxt i; |
19741 | 975 |
|
976 |
(* rws are definitions to be unfolded for admissibility check *) |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27105
diff
changeset
|
977 |
fun Seq_induct_tac ctxt s rws i = |
59780 | 978 |
Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_induct} i |
62116 | 979 |
THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ctxt (i + 1)))) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48194
diff
changeset
|
980 |
THEN simp_tac (ctxt addsimps rws) i; |
19741 | 981 |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27105
diff
changeset
|
982 |
fun Seq_Finite_induct_tac ctxt i = |
60754 | 983 |
eresolve_tac ctxt @{thms Seq_Finite_ind} i |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48194
diff
changeset
|
984 |
THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ctxt i))); |
19741 | 985 |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27105
diff
changeset
|
986 |
fun pair_tac ctxt s = |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61032
diff
changeset
|
987 |
Rule_Insts.res_inst_tac ctxt [((("y", 0), Position.none), s)] [] @{thm prod.exhaust} |
51798 | 988 |
THEN' hyp_subst_tac ctxt THEN' asm_full_simp_tac ctxt; |
19741 | 989 |
|
990 |
(* induction on a sequence of pairs with pairsplitting and simplification *) |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27105
diff
changeset
|
991 |
fun pair_induct_tac ctxt s rws i = |
59780 | 992 |
Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_induct} i |
62195 | 993 |
THEN pair_tac ctxt "a" (i + 3) |
994 |
THEN (REPEAT_DETERM (CHANGED (simp_tac ctxt (i + 1)))) |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48194
diff
changeset
|
995 |
THEN simp_tac (ctxt addsimps rws) i; |
62002 | 996 |
\<close> |
19741 | 997 |
|
62195 | 998 |
method_setup Seq_case = |
74563 | 999 |
\<open>Scan.lift Parse.embedded >> (fn s => fn ctxt => SIMPLE_METHOD' (Seq_case_tac ctxt s))\<close> |
62195 | 1000 |
|
1001 |
method_setup Seq_case_simp = |
|
74563 | 1002 |
\<open>Scan.lift Parse.embedded >> (fn s => fn ctxt => SIMPLE_METHOD' (Seq_case_simp_tac ctxt s))\<close> |
62195 | 1003 |
|
1004 |
method_setup Seq_induct = |
|
74563 | 1005 |
\<open>Scan.lift Parse.embedded -- |
62195 | 1006 |
Scan.optional ((Scan.lift (Args.$$$ "simp" -- Args.colon) |-- Attrib.thms)) [] |
1007 |
>> (fn (s, rws) => fn ctxt => SIMPLE_METHOD' (Seq_induct_tac ctxt s rws))\<close> |
|
1008 |
||
1009 |
method_setup Seq_Finite_induct = |
|
1010 |
\<open>Scan.succeed (SIMPLE_METHOD' o Seq_Finite_induct_tac)\<close> |
|
1011 |
||
1012 |
method_setup pair = |
|
74563 | 1013 |
\<open>Scan.lift Parse.embedded >> (fn s => fn ctxt => SIMPLE_METHOD' (pair_tac ctxt s))\<close> |
62195 | 1014 |
|
1015 |
method_setup pair_induct = |
|
74563 | 1016 |
\<open>Scan.lift Parse.embedded -- |
62195 | 1017 |
Scan.optional ((Scan.lift (Args.$$$ "simp" -- Args.colon) |-- Attrib.thms)) [] |
1018 |
>> (fn (s, rws) => fn ctxt => SIMPLE_METHOD' (pair_induct_tac ctxt s rws))\<close> |
|
1019 |
||
63549 | 1020 |
lemma Mapnil: "Map f \<cdot> s = nil \<longleftrightarrow> s = nil" |
62195 | 1021 |
by (Seq_case_simp s) |
62193 | 1022 |
|
63549 | 1023 |
lemma MapUU: "Map f \<cdot> s = UU \<longleftrightarrow> s = UU" |
62195 | 1024 |
by (Seq_case_simp s) |
62193 | 1025 |
|
63549 | 1026 |
lemma MapTL: "Map f \<cdot> (TL \<cdot> s) = TL \<cdot> (Map f \<cdot> s)" |
62195 | 1027 |
by (Seq_induct s) |
62193 | 1028 |
|
19741 | 1029 |
end |