transformed 'option' and 'list' into new-style datatypes (but register them as old-style as well)
* * *
compile
* * *
tuned imports to prevent merge issues in 'Main'
5 header {* The datatype of finite lists *}
8 imports Presburger Code_Numeral Quotient Lifting_Set Lifting_Option Lifting_Product
11 datatype_new 'a list =
13 | Cons 'a "'a list" (infixr "#" 65)
15 datatype_new_compat list
17 -- {* Compatibility *}
18 setup {* Sign.mandatory_path "list" *}
20 lemmas inducts = list.induct
21 lemmas recs = list.rec
22 lemmas cases = list.case
24 setup {* Sign.parent_path *}
27 -- {* list Enumeration *}
28 "_list" :: "args => 'a list" ("[(_)]")
35 subsection {* Basic list processing functions *}
37 primrec hd :: "'a list \<Rightarrow> 'a" where
40 primrec tl :: "'a list \<Rightarrow> 'a list" where
44 primrec last :: "'a list \<Rightarrow> 'a" where
45 "last (x # xs) = (if xs = [] then x else last xs)"
47 primrec butlast :: "'a list \<Rightarrow> 'a list" where
49 "butlast (x # xs) = (if xs = [] then [] else x # butlast xs)"
51 primrec set :: "'a list \<Rightarrow> 'a set" where
53 "set (x # xs) = insert x (set xs)"
55 definition coset :: "'a list \<Rightarrow> 'a set" where
56 [simp]: "coset xs = - set xs"
58 primrec map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
60 "map f (x # xs) = f x # map f xs"
62 primrec append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "@" 65) where
63 append_Nil: "[] @ ys = ys" |
64 append_Cons: "(x#xs) @ ys = x # xs @ ys"
66 primrec rev :: "'a list \<Rightarrow> 'a list" where
68 "rev (x # xs) = rev xs @ [x]"
70 primrec filter:: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
72 "filter P (x # xs) = (if P x then x # filter P xs else filter P xs)"
75 -- {* Special syntax for filter *}
76 "_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_<-_./ _])")
79 "[x<-xs . P]"== "CONST filter (%x. P) xs"
82 "_filter" :: "[pttrn, 'a list, bool] => 'a list"("(1[_\<leftarrow>_ ./ _])")
84 "_filter" :: "[pttrn, 'a list, bool] => 'a list"("(1[_\<leftarrow>_ ./ _])")
86 primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
87 fold_Nil: "fold f [] = id" |
88 fold_Cons: "fold f (x # xs) = fold f xs \<circ> f x"
90 primrec foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
91 foldr_Nil: "foldr f [] = id" |
92 foldr_Cons: "foldr f (x # xs) = f x \<circ> foldr f xs"
94 primrec foldl :: "('b \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b" where
95 foldl_Nil: "foldl f a [] = a" |
96 foldl_Cons: "foldl f a (x # xs) = foldl f (f a x) xs"
98 primrec concat:: "'a list list \<Rightarrow> 'a list" where
100 "concat (x # xs) = x @ concat xs"
102 definition (in monoid_add) listsum :: "'a list \<Rightarrow> 'a" where
103 "listsum xs = foldr plus xs 0"
105 primrec drop:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
106 drop_Nil: "drop n [] = []" |
107 drop_Cons: "drop n (x # xs) = (case n of 0 \<Rightarrow> x # xs | Suc m \<Rightarrow> drop m xs)"
108 -- {*Warning: simpset does not contain this definition, but separate
109 theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
111 primrec take:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
112 take_Nil:"take n [] = []" |
113 take_Cons: "take n (x # xs) = (case n of 0 \<Rightarrow> [] | Suc m \<Rightarrow> x # take m xs)"
114 -- {*Warning: simpset does not contain this definition, but separate
115 theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
117 primrec nth :: "'a list => nat => 'a" (infixl "!" 100) where
118 nth_Cons: "(x # xs) ! n = (case n of 0 \<Rightarrow> x | Suc k \<Rightarrow> xs ! k)"
119 -- {*Warning: simpset does not contain this definition, but separate
120 theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
122 primrec list_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
123 "list_update [] i v = []" |
124 "list_update (x # xs) i v =
125 (case i of 0 \<Rightarrow> v # xs | Suc j \<Rightarrow> x # list_update xs j v)"
127 nonterminal lupdbinds and lupdbind
130 "_lupdbind":: "['a, 'a] => lupdbind" ("(2_ :=/ _)")
131 "" :: "lupdbind => lupdbinds" ("_")
132 "_lupdbinds" :: "[lupdbind, lupdbinds] => lupdbinds" ("_,/ _")
133 "_LUpdate" :: "['a, lupdbinds] => 'a" ("_/[(_)]" [900,0] 900)
136 "_LUpdate xs (_lupdbinds b bs)" == "_LUpdate (_LUpdate xs b) bs"
137 "xs[i:=x]" == "CONST list_update xs i x"
139 primrec takeWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
140 "takeWhile P [] = []" |
141 "takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])"
143 primrec dropWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
144 "dropWhile P [] = []" |
145 "dropWhile P (x # xs) = (if P x then dropWhile P xs else x # xs)"
147 primrec zip :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
149 zip_Cons: "zip xs (y # ys) =
150 (case xs of [] => [] | z # zs => (z, y) # zip zs ys)"
151 -- {*Warning: simpset does not contain this definition, but separate
152 theorems for @{text "xs = []"} and @{text "xs = z # zs"} *}
154 primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
155 "product [] _ = []" |
156 "product (x#xs) ys = map (Pair x) ys @ product xs ys"
158 hide_const (open) product
160 primrec product_lists :: "'a list list \<Rightarrow> 'a list list" where
161 "product_lists [] = [[]]" |
162 "product_lists (xs # xss) = concat (map (\<lambda>x. map (Cons x) (product_lists xss)) xs)"
164 primrec upt :: "nat \<Rightarrow> nat \<Rightarrow> nat list" ("(1[_..</_'])") where
165 upt_0: "[i..<0] = []" |
166 upt_Suc: "[i..<(Suc j)] = (if i <= j then [i..<j] @ [j] else [])"
168 definition insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
169 "insert x xs = (if x \<in> set xs then xs else x # xs)"
171 hide_const (open) insert
172 hide_fact (open) insert_def
174 primrec find :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a option" where
176 "find P (x#xs) = (if P x then Some x else find P xs)"
178 hide_const (open) find
180 primrec those :: "'a option list \<Rightarrow> 'a list option"
182 "those [] = Some []" |
183 "those (x # xs) = (case x of
184 None \<Rightarrow> None
185 | Some y \<Rightarrow> Option.map (Cons y) (those xs))"
187 primrec remove1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
188 "remove1 x [] = []" |
189 "remove1 x (y # xs) = (if x = y then xs else y # remove1 x xs)"
191 primrec removeAll :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
192 "removeAll x [] = []" |
193 "removeAll x (y # xs) = (if x = y then removeAll x xs else y # removeAll x xs)"
195 primrec distinct :: "'a list \<Rightarrow> bool" where
196 "distinct [] \<longleftrightarrow> True" |
197 "distinct (x # xs) \<longleftrightarrow> x \<notin> set xs \<and> distinct xs"
199 primrec remdups :: "'a list \<Rightarrow> 'a list" where
201 "remdups (x # xs) = (if x \<in> set xs then remdups xs else x # remdups xs)"
203 fun remdups_adj :: "'a list \<Rightarrow> 'a list" where
204 "remdups_adj [] = []" |
205 "remdups_adj [x] = [x]" |
206 "remdups_adj (x # y # xs) = (if x = y then remdups_adj (x # xs) else x # remdups_adj (y # xs))"
208 primrec replicate :: "nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
209 replicate_0: "replicate 0 x = []" |
210 replicate_Suc: "replicate (Suc n) x = x # replicate n x"
213 Function @{text size} is overloaded for all datatypes. Users may
214 refer to the list version as @{text length}. *}
216 abbreviation length :: "'a list \<Rightarrow> nat" where
217 "length \<equiv> size"
219 definition enumerate :: "nat \<Rightarrow> 'a list \<Rightarrow> (nat \<times> 'a) list" where
220 enumerate_eq_zip: "enumerate n xs = zip [n..<n + length xs] xs"
222 primrec rotate1 :: "'a list \<Rightarrow> 'a list" where
224 "rotate1 (x # xs) = xs @ [x]"
226 definition rotate :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
227 "rotate n = rotate1 ^^ n"
229 definition list_all2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> bool" where
231 (length xs = length ys \<and> (\<forall>(x, y) \<in> set (zip xs ys). P x y))"
233 definition sublist :: "'a list => nat set => 'a list" where
234 "sublist xs A = map fst (filter (\<lambda>p. snd p \<in> A) (zip xs [0..<size xs]))"
236 primrec sublists :: "'a list \<Rightarrow> 'a list list" where
237 "sublists [] = [[]]" |
238 "sublists (x#xs) = (let xss = sublists xs in map (Cons x) xss @ xss)"
240 primrec n_lists :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list list" where
241 "n_lists 0 xs = [[]]" |
242 "n_lists (Suc n) xs = concat (map (\<lambda>ys. map (\<lambda>y. y # ys) xs) (n_lists n xs))"
244 hide_const (open) n_lists
246 fun splice :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
247 "splice [] ys = ys" |
248 "splice xs [] = xs" |
249 "splice (x#xs) (y#ys) = x # y # splice xs ys"
255 @{lemma "[a,b]@[c,d] = [a,b,c,d]" by simp}\\
256 @{lemma "length [a,b,c] = 3" by simp}\\
257 @{lemma "set [a,b,c] = {a,b,c}" by simp}\\
258 @{lemma "map f [a,b,c] = [f a, f b, f c]" by simp}\\
259 @{lemma "rev [a,b,c] = [c,b,a]" by simp}\\
260 @{lemma "hd [a,b,c,d] = a" by simp}\\
261 @{lemma "tl [a,b,c,d] = [b,c,d]" by simp}\\
262 @{lemma "last [a,b,c,d] = d" by simp}\\
263 @{lemma "butlast [a,b,c,d] = [a,b,c]" by simp}\\
264 @{lemma[source] "filter (\<lambda>n::nat. n<2) [0,2,1] = [0,1]" by simp}\\
265 @{lemma "concat [[a,b],[c,d,e],[],[f]] = [a,b,c,d,e,f]" by simp}\\
266 @{lemma "fold f [a,b,c] x = f c (f b (f a x))" by simp}\\
267 @{lemma "foldr f [a,b,c] x = f a (f b (f c x))" by simp}\\
268 @{lemma "foldl f x [a,b,c] = f (f (f x a) b) c" by simp}\\
269 @{lemma "zip [a,b,c] [x,y,z] = [(a,x),(b,y),(c,z)]" by simp}\\
270 @{lemma "zip [a,b] [x,y,z] = [(a,x),(b,y)]" by simp}\\
271 @{lemma "enumerate 3 [a,b,c] = [(3,a),(4,b),(5,c)]" by normalization}\\
272 @{lemma "List.product [a,b] [c,d] = [(a, c), (a, d), (b, c), (b, d)]" by simp}\\
273 @{lemma "product_lists [[a,b], [c], [d,e]] = [[a,c,d], [a,c,e], [b,c,d], [b,c,e]]" by simp}\\
274 @{lemma "splice [a,b,c] [x,y,z] = [a,x,b,y,c,z]" by simp}\\
275 @{lemma "splice [a,b,c,d] [x,y] = [a,x,b,y,c,d]" by simp}\\
276 @{lemma "take 2 [a,b,c,d] = [a,b]" by simp}\\
277 @{lemma "take 6 [a,b,c,d] = [a,b,c,d]" by simp}\\
278 @{lemma "drop 2 [a,b,c,d] = [c,d]" by simp}\\
279 @{lemma "drop 6 [a,b,c,d] = []" by simp}\\
280 @{lemma "takeWhile (%n::nat. n<3) [1,2,3,0] = [1,2]" by simp}\\
281 @{lemma "dropWhile (%n::nat. n<3) [1,2,3,0] = [3,0]" by simp}\\
282 @{lemma "distinct [2,0,1::nat]" by simp}\\
283 @{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\
284 @{lemma "remdups_adj [2,2,3,1,1::nat,2,1] = [2,3,1,2,1]" by simp}\\
285 @{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\
286 @{lemma "List.insert 3 [0::nat,1,2] = [3,0,1,2]" by (simp add: List.insert_def)}\\
287 @{lemma "List.find (%i::int. i>0) [0,0] = None" by simp}\\
288 @{lemma "List.find (%i::int. i>0) [0,1,0,2] = Some 1" by simp}\\
289 @{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\
290 @{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\
291 @{lemma "nth [a,b,c,d] 2 = c" by simp}\\
292 @{lemma "[a,b,c,d][2 := x] = [a,b,x,d]" by simp}\\
293 @{lemma "sublist [a,b,c,d,e] {0,2,3} = [a,c,d]" by (simp add:sublist_def)}\\
294 @{lemma "sublists [a,b] = [[a, b], [a], [b], []]" by simp}\\
295 @{lemma "List.n_lists 2 [a,b,c] = [[a, a], [b, a], [c, a], [a, b], [b, b], [c, b], [a, c], [b, c], [c, c]]" by (simp add: eval_nat_numeral)}\\
296 @{lemma "rotate1 [a,b,c,d] = [b,c,d,a]" by simp}\\
297 @{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate_def eval_nat_numeral)}\\
298 @{lemma "replicate 4 a = [a,a,a,a]" by (simp add:eval_nat_numeral)}\\
299 @{lemma "[2..<5] = [2,3,4]" by (simp add:eval_nat_numeral)}\\
300 @{lemma "listsum [1,2,3::nat] = 6" by (simp add: listsum_def)}
302 \caption{Characteristic examples}
303 \label{fig:Characteristic}
305 Figure~\ref{fig:Characteristic} shows characteristic examples
306 that should give an intuitive understanding of the above functions.
309 text{* The following simple sort functions are intended for proofs,
310 not for efficient implementations. *}
315 inductive sorted :: "'a list \<Rightarrow> bool" where
316 Nil [iff]: "sorted []"
317 | Cons: "\<forall>y\<in>set xs. x \<le> y \<Longrightarrow> sorted xs \<Longrightarrow> sorted (x # xs)"
319 lemma sorted_single [iff]:
321 by (rule sorted.Cons) auto
324 "x \<le> y \<Longrightarrow> sorted (y # zs) \<Longrightarrow> sorted (x # y # zs)"
325 by (rule sorted.Cons) (cases "y # zs" rule: sorted.cases, auto)
327 lemma sorted_many_eq [simp, code]:
328 "sorted (x # y # zs) \<longleftrightarrow> x \<le> y \<and> sorted (y # zs)"
329 by (auto intro: sorted_many elim: sorted.cases)
332 "sorted [] \<longleftrightarrow> True"
333 "sorted [x] \<longleftrightarrow> True"
336 primrec insort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where
337 "insort_key f x [] = [x]" |
338 "insort_key f x (y#ys) =
339 (if f x \<le> f y then (x#y#ys) else y#(insort_key f x ys))"
341 definition sort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b list \<Rightarrow> 'b list" where
342 "sort_key f xs = foldr (insort_key f) xs []"
344 definition insort_insert_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where
345 "insort_insert_key f x xs =
346 (if f x \<in> f ` set xs then xs else insort_key f x xs)"
348 abbreviation "sort \<equiv> sort_key (\<lambda>x. x)"
349 abbreviation "insort \<equiv> insort_key (\<lambda>x. x)"
350 abbreviation "insort_insert \<equiv> insort_insert_key (\<lambda>x. x)"
355 subsubsection {* List comprehension *}
357 text{* Input syntax for Haskell-like list comprehension notation.
358 Typical example: @{text"[(x,y). x \<leftarrow> xs, y \<leftarrow> ys, x \<noteq> y]"},
359 the list of all pairs of distinct elements from @{text xs} and @{text ys}.
360 The syntax is as in Haskell, except that @{text"|"} becomes a dot
361 (like in Isabelle's set comprehension): @{text"[e. x \<leftarrow> xs, \<dots>]"} rather than
362 \verb![e| x <- xs, ...]!.
364 The qualifiers after the dot are
366 \item[generators] @{text"p \<leftarrow> xs"},
367 where @{text p} is a pattern and @{text xs} an expression of list type, or
368 \item[guards] @{text"b"}, where @{text b} is a boolean expression.
369 %\item[local bindings] @ {text"let x = e"}.
372 Just like in Haskell, list comprehension is just a shorthand. To avoid
373 misunderstandings, the translation into desugared form is not reversed
374 upon output. Note that the translation of @{text"[e. x \<leftarrow> xs]"} is
375 optmized to @{term"map (%x. e) xs"}.
377 It is easy to write short list comprehensions which stand for complex
378 expressions. During proofs, they may become unreadable (and
379 mangled). In such cases it can be advisable to introduce separate
380 definitions for the list comprehensions in question. *}
382 nonterminal lc_qual and lc_quals
385 "_listcompr" :: "'a \<Rightarrow> lc_qual \<Rightarrow> lc_quals \<Rightarrow> 'a list" ("[_ . __")
386 "_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ <- _")
387 "_lc_test" :: "bool \<Rightarrow> lc_qual" ("_")
388 (*"_lc_let" :: "letbinds => lc_qual" ("let _")*)
389 "_lc_end" :: "lc_quals" ("]")
390 "_lc_quals" :: "lc_qual \<Rightarrow> lc_quals \<Rightarrow> lc_quals" (", __")
391 "_lc_abs" :: "'a => 'b list => 'b list"
393 (* These are easier than ML code but cannot express the optimized
394 translation of [e. p<-xs]
396 "[e. p<-xs]" => "concat(map (_lc_abs p [e]) xs)"
397 "_listcompr e (_lc_gen p xs) (_lc_quals Q Qs)"
398 => "concat (map (_lc_abs p (_listcompr e Q Qs)) xs)"
399 "[e. P]" => "if P then [e] else []"
400 "_listcompr e (_lc_test P) (_lc_quals Q Qs)"
401 => "if P then (_listcompr e Q Qs) else []"
402 "_listcompr e (_lc_let b) (_lc_quals Q Qs)"
403 => "_Let b (_listcompr e Q Qs)"
407 "_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
409 "_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
413 val NilC = Syntax.const @{const_syntax Nil};
414 val ConsC = Syntax.const @{const_syntax Cons};
415 val mapC = Syntax.const @{const_syntax map};
416 val concatC = Syntax.const @{const_syntax concat};
417 val IfC = Syntax.const @{const_syntax If};
419 fun single x = ConsC $ x $ NilC;
421 fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)
423 (* FIXME proper name context!? *)
425 Free (singleton (Name.variant_list (fold Term.add_free_names [p, e] [])) "x", dummyT);
426 val e = if opti then single e else e;
427 val case1 = Syntax.const @{syntax_const "_case1"} $ p $ e;
429 Syntax.const @{syntax_const "_case1"} $
430 Syntax.const @{const_syntax dummy_pattern} $ NilC;
431 val cs = Syntax.const @{syntax_const "_case2"} $ case1 $ case2;
432 in Syntax_Trans.abs_tr [x, Case_Translation.case_tr false ctxt [x, cs]] end;
434 fun abs_tr ctxt p e opti =
435 (case Term_Position.strip_positions p of
438 val thy = Proof_Context.theory_of ctxt;
439 val s' = Proof_Context.intern_const ctxt s;
441 if Sign.declared_const thy s'
442 then (pat_tr ctxt p e opti, false)
443 else (Syntax_Trans.abs_tr [p, e], true)
445 | _ => (pat_tr ctxt p e opti, false));
447 fun lc_tr ctxt [e, Const (@{syntax_const "_lc_test"}, _) $ b, qs] =
451 Const (@{syntax_const "_lc_end"}, _) => single e
452 | Const (@{syntax_const "_lc_quals"}, _) $ q $ qs => lc_tr ctxt [e, q, qs]);
453 in IfC $ b $ res $ NilC end
455 [e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es,
456 Const(@{syntax_const "_lc_end"}, _)] =
457 (case abs_tr ctxt p e true of
458 (f, true) => mapC $ f $ es
459 | (f, false) => concatC $ (mapC $ f $ es))
461 [e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es,
462 Const (@{syntax_const "_lc_quals"}, _) $ q $ qs] =
463 let val e' = lc_tr ctxt [e, q, qs];
464 in concatC $ (mapC $ (fst (abs_tr ctxt p e' false)) $ es) end;
466 in [(@{syntax_const "_listcompr"}, lc_tr)] end
471 val read = Syntax.read_term @{context};
472 fun check s1 s2 = read s1 aconv read s2 orelse error ("Check failed: " ^ quote s1);
474 check "[(x,y,z). b]" "if b then [(x, y, z)] else []";
475 check "[(x,y,z). x\<leftarrow>xs]" "map (\<lambda>x. (x, y, z)) xs";
476 check "[e x y. x\<leftarrow>xs, y\<leftarrow>ys]" "concat (map (\<lambda>x. map (\<lambda>y. e x y) ys) xs)";
477 check "[(x,y,z). x<a, x>b]" "if x < a then if b < x then [(x, y, z)] else [] else []";
478 check "[(x,y,z). x\<leftarrow>xs, x>b]" "concat (map (\<lambda>x. if b < x then [(x, y, z)] else []) xs)";
479 check "[(x,y,z). x<a, x\<leftarrow>xs]" "if x < a then map (\<lambda>x. (x, y, z)) xs else []";
480 check "[(x,y). Cons True x \<leftarrow> xs]"
481 "concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | True # x \<Rightarrow> [(x, y)] | False # x \<Rightarrow> []) xs)";
482 check "[(x,y,z). Cons x [] \<leftarrow> xs]"
483 "concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | [x] \<Rightarrow> [(x, y, z)] | x # aa # lista \<Rightarrow> []) xs)";
484 check "[(x,y,z). x<a, x>b, x=d]"
485 "if x < a then if b < x then if x = d then [(x, y, z)] else [] else [] else []";
486 check "[(x,y,z). x<a, x>b, y\<leftarrow>ys]"
487 "if x < a then if b < x then map (\<lambda>y. (x, y, z)) ys else [] else []";
488 check "[(x,y,z). x<a, x\<leftarrow>xs,y>b]"
489 "if x < a then concat (map (\<lambda>x. if b < y then [(x, y, z)] else []) xs) else []";
490 check "[(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys]"
491 "if x < a then concat (map (\<lambda>x. map (\<lambda>y. (x, y, z)) ys) xs) else []";
492 check "[(x,y,z). x\<leftarrow>xs, x>b, y<a]"
493 "concat (map (\<lambda>x. if b < x then if y < a then [(x, y, z)] else [] else []) xs)";
494 check "[(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys]"
495 "concat (map (\<lambda>x. if b < x then map (\<lambda>y. (x, y, z)) ys else []) xs)";
496 check "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x]"
497 "concat (map (\<lambda>x. concat (map (\<lambda>y. if x < y then [(x, y, z)] else []) ys)) xs)";
498 check "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs]"
499 "concat (map (\<lambda>x. concat (map (\<lambda>y. map (\<lambda>z. (x, y, z)) zs) ys)) xs)"
504 term "[(x,y). x\<leftarrow>xs, let xx = x+x, y\<leftarrow>ys, y \<noteq> xx]"
509 (* Simproc for rewriting list comprehensions applied to List.set to set
512 signature LIST_TO_SET_COMPREHENSION =
514 val simproc : Proof.context -> cterm -> thm option
517 structure List_to_Set_Comprehension : LIST_TO_SET_COMPREHENSION =
522 fun all_exists_conv cv ctxt ct =
523 (case Thm.term_of ct of
524 Const (@{const_name HOL.Ex}, _) $ Abs _ =>
525 Conv.arg_conv (Conv.abs_conv (all_exists_conv cv o #2) ctxt) ct
528 fun all_but_last_exists_conv cv ctxt ct =
529 (case Thm.term_of ct of
530 Const (@{const_name HOL.Ex}, _) $ Abs (_, _, Const (@{const_name HOL.Ex}, _) $ _) =>
531 Conv.arg_conv (Conv.abs_conv (all_but_last_exists_conv cv o #2) ctxt) ct
534 fun Collect_conv cv ctxt ct =
535 (case Thm.term_of ct of
536 Const (@{const_name Set.Collect}, _) $ Abs _ => Conv.arg_conv (Conv.abs_conv cv ctxt) ct
537 | _ => raise CTERM ("Collect_conv", [ct]))
539 fun rewr_conv' th = Conv.rewr_conv (mk_meta_eq th)
541 fun conjunct_assoc_conv ct =
543 (rewr_conv' @{thm conj_assoc} then_conv HOLogic.conj_conv Conv.all_conv conjunct_assoc_conv) ct
545 fun right_hand_set_comprehension_conv conv ctxt =
546 HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv
547 (Collect_conv (all_exists_conv conv o #2) ctxt))
550 (* term abstraction of list comprehension patterns *)
552 datatype termlets = If | Case of (typ * int)
554 fun simproc ctxt redex =
556 val set_Nil_I = @{thm trans} OF [@{thm set.simps(1)}, @{thm empty_def}]
557 val set_singleton = @{lemma "set [a] = {x. x = a}" by simp}
558 val inst_Collect_mem_eq = @{lemma "set A = {x. x : set A}" by simp}
559 val del_refl_eq = @{lemma "(t = t & P) == P" by simp}
560 fun mk_set T = Const (@{const_name List.set}, HOLogic.listT T --> HOLogic.mk_setT T)
561 fun dest_set (Const (@{const_name List.set}, _) $ xs) = xs
562 fun dest_singleton_list (Const (@{const_name List.Cons}, _)
563 $ t $ (Const (@{const_name List.Nil}, _))) = t
564 | dest_singleton_list t = raise TERM ("dest_singleton_list", [t])
565 (* We check that one case returns a singleton list and all other cases
566 return [], and return the index of the one singleton list case *)
567 fun possible_index_of_singleton_case cases =
569 fun check (i, case_t) s =
570 (case strip_abs_body case_t of
571 (Const (@{const_name List.Nil}, _)) => s
572 | _ => (case s of SOME NONE => SOME (SOME i) | _ => NONE))
574 fold_index check cases (SOME NONE) |> the_default NONE
576 (* returns (case_expr type index chosen_case constr_name) option *)
577 fun dest_case case_term =
579 val (case_const, args) = strip_comb case_term
581 (case try dest_Const case_const of
583 (case Ctr_Sugar.ctr_sugar_of_case ctxt c of
585 (case possible_index_of_singleton_case (fst (split_last args)) of
588 val constr_names = map (fst o dest_Const) ctrs
589 val (Ts, _) = strip_type T
590 val T' = List.last Ts
591 in SOME (List.last args, T', i, nth args i, nth constr_names i) end
596 (* returns condition continuing term option *)
597 fun dest_if (Const (@{const_name If}, _) $ cond $ then_t $ Const (@{const_name Nil}, _)) =
600 fun tac _ [] = rtac set_singleton 1 ORELSE rtac inst_Collect_mem_eq 1
601 | tac ctxt (If :: cont) =
602 Splitter.split_tac [@{thm split_if}] 1
603 THEN rtac @{thm conjI} 1
604 THEN rtac @{thm impI} 1
605 THEN Subgoal.FOCUS (fn {prems, context, ...} =>
606 CONVERSION (right_hand_set_comprehension_conv (K
607 (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_TrueI})) Conv.all_conv
609 rewr_conv' @{lemma "(True & P) = P" by simp})) context) 1) ctxt 1
611 THEN rtac @{thm impI} 1
612 THEN Subgoal.FOCUS (fn {prems, context, ...} =>
613 CONVERSION (right_hand_set_comprehension_conv (K
614 (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_FalseI})) Conv.all_conv
615 then_conv rewr_conv' @{lemma "(False & P) = False" by simp})) context) 1) ctxt 1
616 THEN rtac set_Nil_I 1
617 | tac ctxt (Case (T, i) :: cont) =
619 val SOME {injects, distincts, case_thms, split, ...} =
620 Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type T))
622 (* do case distinction *)
623 Splitter.split_tac [split] 1
624 THEN EVERY (map_index (fn (i', _) =>
625 (if i' < length case_thms - 1 then rtac @{thm conjI} 1 else all_tac)
626 THEN REPEAT_DETERM (rtac @{thm allI} 1)
627 THEN rtac @{thm impI} 1
629 (* continue recursively *)
630 Subgoal.FOCUS (fn {prems, context, ...} =>
631 CONVERSION (Thm.eta_conversion then_conv right_hand_set_comprehension_conv (K
633 (HOLogic.eq_conv Conv.all_conv (rewr_conv' (List.last prems)) then_conv
634 (Conv.try_conv (Conv.rewrs_conv (map mk_meta_eq injects))))
636 then_conv (Conv.try_conv (Conv.rewr_conv del_refl_eq))
637 then_conv conjunct_assoc_conv)) context
638 then_conv (HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt) =>
640 (all_but_last_exists_conv
642 @{lemma "(EX x. x = t & P x) = P t" by simp})) ctxt)) context)))) 1) ctxt 1
645 Subgoal.FOCUS (fn {prems, context, ...} =>
647 (right_hand_set_comprehension_conv (K
649 ((HOLogic.eq_conv Conv.all_conv
650 (rewr_conv' (List.last prems))) then_conv
651 (Conv.rewrs_conv (map (fn th => th RS @{thm Eq_FalseI}) distincts)))
652 Conv.all_conv then_conv
653 (rewr_conv' @{lemma "(False & P) = False" by simp}))) context then_conv
654 HOLogic.Trueprop_conv
655 (HOLogic.eq_conv Conv.all_conv
656 (Collect_conv (fn (_, ctxt) =>
660 @{lemma "(EX x. P) = P" by simp})) ctxt)) context))) 1) ctxt 1
661 THEN rtac set_Nil_I 1)) case_thms)
663 fun make_inner_eqs bound_vs Tis eqs t =
665 SOME (x, T, i, cont, constr_name) =>
667 val (vs, body) = strip_abs (Envir.eta_long (map snd bound_vs) cont)
668 val x' = incr_boundvars (length vs) x
669 val eqs' = map (incr_boundvars (length vs)) eqs
672 (Const (constr_name, map snd vs ---> T), map Bound (((length vs) - 1) downto 0))
673 val constr_eq = Const (@{const_name HOL.eq}, T --> T --> @{typ bool}) $ constr_t $ x'
675 make_inner_eqs (rev vs @ bound_vs) (Case (T, i) :: Tis) (constr_eq :: eqs') body
679 SOME (condition, cont) => make_inner_eqs bound_vs (If :: Tis) (condition :: eqs) cont
681 if eqs = [] then NONE (* no rewriting, nothing to be done *)
684 val Type (@{type_name List.list}, [rT]) = fastype_of1 (map snd bound_vs, t)
686 (case try dest_singleton_list t of
688 Const (@{const_name HOL.eq}, rT --> rT --> @{typ bool}) $
689 Bound (length bound_vs) $ t'
691 Const (@{const_name Set.member}, rT --> HOLogic.mk_setT rT --> @{typ bool}) $
692 Bound (length bound_vs) $ (mk_set rT $ t))
693 val reverse_bounds = curry subst_bounds
694 ((map Bound ((length bound_vs - 1) downto 0)) @ [Bound (length bound_vs)])
695 val eqs' = map reverse_bounds eqs
696 val pat_eq' = reverse_bounds pat_eq
698 fold (fn (_, T) => fn t => HOLogic.exists_const T $ absdummy T t)
699 (rev bound_vs) (fold (curry HOLogic.mk_conj) eqs' pat_eq')
700 val lhs = term_of redex
701 val rhs = HOLogic.mk_Collect ("x", rT, inner_t)
702 val rewrite_rule_t = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
705 ((Goal.prove ctxt [] [] rewrite_rule_t
706 (fn {context, ...} => tac context (rev Tis))) RS @{thm eq_reflection})
709 make_inner_eqs [] [] [] (dest_set (term_of redex))
715 simproc_setup list_to_set_comprehension ("set xs") = {* K List_to_Set_Comprehension.simproc *}
717 code_datatype set coset
719 hide_const (open) coset
722 subsubsection {* @{const Nil} and @{const Cons} *}
724 lemma not_Cons_self [simp]:
728 lemma not_Cons_self2 [simp]:
730 by (rule not_Cons_self [symmetric])
732 lemma neq_Nil_conv: "(xs \<noteq> []) = (\<exists>y ys. xs = y # ys)"
735 lemma tl_Nil: "tl xs = [] \<longleftrightarrow> xs = [] \<or> (EX x. xs = [x])"
738 lemma Nil_tl: "[] = tl xs \<longleftrightarrow> xs = [] \<or> (EX x. xs = [x])"
742 "(\<And>xs. \<forall>ys. length ys < length xs \<longrightarrow> P ys \<Longrightarrow> P xs) \<Longrightarrow> P xs"
743 by (fact measure_induct)
745 lemma list_nonempty_induct [consumes 1, case_names single cons]:
746 assumes "xs \<noteq> []"
747 assumes single: "\<And>x. P [x]"
748 assumes cons: "\<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (x # xs)"
750 using `xs \<noteq> []` proof (induct xs)
751 case Nil then show ?case by simp
757 with single show ?thesis by simp
762 from Cons show "xs \<noteq> []" by simp
763 with Cons.hyps show "P xs" .
768 lemma inj_split_Cons: "inj_on (\<lambda>(xs, n). n#xs) X"
769 by (auto intro!: inj_onI)
772 subsubsection {* @{const length} *}
775 Needs to come before @{text "@"} because of theorem @{text
776 append_eq_append_conv}.
779 lemma length_append [simp]: "length (xs @ ys) = length xs + length ys"
782 lemma length_map [simp]: "length (map f xs) = length xs"
785 lemma length_rev [simp]: "length (rev xs) = length xs"
788 lemma length_tl [simp]: "length (tl xs) = length xs - 1"
791 lemma length_0_conv [iff]: "(length xs = 0) = (xs = [])"
794 lemma length_greater_0_conv [iff]: "(0 < length xs) = (xs \<noteq> [])"
797 lemma length_pos_if_in_set: "x : set xs \<Longrightarrow> length xs > 0"
800 lemma length_Suc_conv:
801 "(length xs = Suc n) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
804 lemma Suc_length_conv:
805 "(Suc n = length xs) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
806 apply (induct xs, simp, simp)
810 lemma impossible_Cons: "length xs <= length ys ==> xs = x # ys = False"
813 lemma list_induct2 [consumes 1, case_names Nil Cons]:
814 "length xs = length ys \<Longrightarrow> P [] [] \<Longrightarrow>
815 (\<And>x xs y ys. length xs = length ys \<Longrightarrow> P xs ys \<Longrightarrow> P (x#xs) (y#ys))
816 \<Longrightarrow> P xs ys"
817 proof (induct xs arbitrary: ys)
818 case Nil then show ?case by simp
820 case (Cons x xs ys) then show ?case by (cases ys) simp_all
823 lemma list_induct3 [consumes 2, case_names Nil Cons]:
824 "length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> P [] [] [] \<Longrightarrow>
825 (\<And>x xs y ys z zs. length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> P xs ys zs \<Longrightarrow> P (x#xs) (y#ys) (z#zs))
826 \<Longrightarrow> P xs ys zs"
827 proof (induct xs arbitrary: ys zs)
828 case Nil then show ?case by simp
830 case (Cons x xs ys zs) then show ?case by (cases ys, simp_all)
834 lemma list_induct4 [consumes 3, case_names Nil Cons]:
835 "length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow>
836 P [] [] [] [] \<Longrightarrow> (\<And>x xs y ys z zs w ws. length xs = length ys \<Longrightarrow>
837 length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow> P xs ys zs ws \<Longrightarrow>
838 P (x#xs) (y#ys) (z#zs) (w#ws)) \<Longrightarrow> P xs ys zs ws"
839 proof (induct xs arbitrary: ys zs ws)
840 case Nil then show ?case by simp
842 case (Cons x xs ys zs ws) then show ?case by ((cases ys, simp_all), (cases zs,simp_all)) (cases ws, simp_all)
847 \<And>x xs. P (x#xs) [];
848 \<And>y ys. P [] (y#ys);
849 \<And>x xs y ys. P xs ys \<Longrightarrow> P (x#xs) (y#ys) \<rbrakk>
850 \<Longrightarrow> P xs ys"
851 by (induct xs arbitrary: ys) (case_tac x, auto)+
853 lemma neq_if_length_neq: "length xs \<noteq> length ys \<Longrightarrow> (xs = ys) == False"
854 by (rule Eq_FalseI) auto
856 simproc_setup list_neq ("(xs::'a list) = ys") = {*
858 Reduces xs=ys to False if xs and ys cannot be of the same length.
859 This is the case if the atomic sublists of one are a submultiset
860 of those of the other list and there are fewer Cons's in one than the other.
865 fun len (Const(@{const_name Nil},_)) acc = acc
866 | len (Const(@{const_name Cons},_) $ _ $ xs) (ts,n) = len xs (ts,n+1)
867 | len (Const(@{const_name append},_) $ xs $ ys) acc = len xs (len ys acc)
868 | len (Const(@{const_name rev},_) $ xs) acc = len xs acc
869 | len (Const(@{const_name map},_) $ _ $ xs) acc = len xs acc
870 | len t (ts,n) = (t::ts,n);
872 val ss = simpset_of @{context};
874 fun list_neq ctxt ct =
876 val (Const(_,eqT) $ lhs $ rhs) = Thm.term_of ct;
877 val (ls,m) = len lhs ([],0) and (rs,n) = len rhs ([],0);
880 val Type(_,listT::_) = eqT;
881 val size = HOLogic.size_const listT;
882 val eq_len = HOLogic.mk_eq (size $ lhs, size $ rhs);
883 val neq_len = HOLogic.mk_Trueprop (HOLogic.Not $ eq_len);
884 val thm = Goal.prove ctxt [] [] neq_len
885 (K (simp_tac (put_simpset ss ctxt) 1));
886 in SOME (thm RS @{thm neq_if_length_neq}) end
888 if m < n andalso submultiset (op aconv) (ls,rs) orelse
889 n < m andalso submultiset (op aconv) (rs,ls)
890 then prove_neq() else NONE
896 subsubsection {* @{text "@"} -- append *}
898 lemma append_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)"
901 lemma append_Nil2 [simp]: "xs @ [] = xs"
904 lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \<and> ys = [])"
907 lemma Nil_is_append_conv [iff]: "([] = xs @ ys) = (xs = [] \<and> ys = [])"
910 lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])"
913 lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])"
916 lemma append_eq_append_conv [simp]:
917 "length xs = length ys \<or> length us = length vs
918 ==> (xs@us = ys@vs) = (xs=ys \<and> us=vs)"
919 apply (induct xs arbitrary: ys)
920 apply (case_tac ys, simp, force)
921 apply (case_tac ys, force, simp)
924 lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) =
925 (EX us. xs = zs @ us & us @ ys = ts | xs @ us = zs & ys = us@ ts)"
926 apply (induct xs arbitrary: ys zs ts)
933 lemma same_append_eq [iff, induct_simp]: "(xs @ ys = xs @ zs) = (ys = zs)"
936 lemma append1_eq_conv [iff]: "(xs @ [x] = ys @ [y]) = (xs = ys \<and> x = y)"
939 lemma append_same_eq [iff, induct_simp]: "(ys @ xs = zs @ xs) = (ys = zs)"
942 lemma append_self_conv2 [iff]: "(xs @ ys = ys) = (xs = [])"
943 using append_same_eq [of _ _ "[]"] by auto
945 lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])"
946 using append_same_eq [of "[]"] by auto
948 lemma hd_Cons_tl [simp]: "xs \<noteq> [] ==> hd xs # tl xs = xs"
951 lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)"
954 lemma hd_append2 [simp]: "xs \<noteq> [] ==> hd (xs @ ys) = hd xs"
955 by (simp add: hd_append split: list.split)
957 lemma tl_append: "tl (xs @ ys) = (case xs of [] => tl ys | z#zs => zs @ ys)"
958 by (simp split: list.split)
960 lemma tl_append2 [simp]: "xs \<noteq> [] ==> tl (xs @ ys) = tl xs @ ys"
961 by (simp add: tl_append split: list.split)
964 lemma Cons_eq_append_conv: "x#xs = ys@zs =
965 (ys = [] & x#xs = zs | (EX ys'. x#ys' = ys & xs = ys'@zs))"
968 lemma append_eq_Cons_conv: "(ys@zs = x#xs) =
969 (ys = [] & zs = x#xs | (EX ys'. ys = x#ys' & ys'@zs = xs))"
973 text {* Trivial rules for solving @{text "@"}-equations automatically. *}
975 lemma eq_Nil_appendI: "xs = ys ==> xs = [] @ ys"
978 lemma Cons_eq_appendI:
979 "[| x # xs1 = ys; xs = xs1 @ zs |] ==> x # xs = ys @ zs"
982 lemma append_eq_appendI:
983 "[| xs @ xs1 = zs; ys = xs1 @ us |] ==> xs @ ys = zs @ us"
988 Simplification procedure for all list equalities.
989 Currently only tries to rearrange @{text "@"} to see if
990 - both lists end in a singleton list,
991 - or both lists end in the same list.
994 simproc_setup list_eq ("(xs::'a list) = ys") = {*
996 fun last (cons as Const (@{const_name Cons}, _) $ _ $ xs) =
997 (case xs of Const (@{const_name Nil}, _) => cons | _ => last xs)
998 | last (Const(@{const_name append},_) $ _ $ ys) = last ys
1001 fun list1 (Const(@{const_name Cons},_) $ _ $ Const(@{const_name Nil},_)) = true
1004 fun butlast ((cons as Const(@{const_name Cons},_) $ x) $ xs) =
1005 (case xs of Const (@{const_name Nil}, _) => xs | _ => cons $ butlast xs)
1006 | butlast ((app as Const (@{const_name append}, _) $ xs) $ ys) = app $ butlast ys
1007 | butlast xs = Const(@{const_name Nil}, fastype_of xs);
1010 simpset_of (put_simpset HOL_basic_ss @{context}
1011 addsimps [@{thm append_assoc}, @{thm append_Nil}, @{thm append_Cons}]);
1013 fun list_eq ctxt (F as (eq as Const(_,eqT)) $ lhs $ rhs) =
1015 val lastl = last lhs and lastr = last rhs;
1018 val lhs1 = butlast lhs and rhs1 = butlast rhs;
1019 val Type(_,listT::_) = eqT
1020 val appT = [listT,listT] ---> listT
1021 val app = Const(@{const_name append},appT)
1022 val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr)
1023 val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2));
1024 val thm = Goal.prove ctxt [] [] eq
1025 (K (simp_tac (put_simpset rearr_ss ctxt) 1));
1026 in SOME ((conv RS (thm RS trans)) RS eq_reflection) end;
1028 if list1 lastl andalso list1 lastr then rearr @{thm append1_eq_conv}
1029 else if lastl aconv lastr then rearr @{thm append_same_eq}
1032 in fn _ => fn ctxt => fn ct => list_eq ctxt (term_of ct) end;
1036 subsubsection {* @{const map} *}
1039 "xs \<noteq> [] \<Longrightarrow> hd (map f xs) = f (hd xs)"
1040 by (cases xs) simp_all
1043 "map f (tl xs) = tl (map f xs)"
1044 by (cases xs) simp_all
1046 lemma map_ext: "(!!x. x : set xs --> f x = g x) ==> map f xs = map g xs"
1047 by (induct xs) simp_all
1049 lemma map_ident [simp]: "map (\<lambda>x. x) = (\<lambda>xs. xs)"
1050 by (rule ext, induct_tac xs) auto
1052 lemma map_append [simp]: "map f (xs @ ys) = map f xs @ map f ys"
1055 lemma map_map [simp]: "map f (map g xs) = map (f \<circ> g) xs"
1058 lemma map_comp_map[simp]: "((map f) o (map g)) = map(f o g)"
1063 lemma rev_map: "rev (map f xs) = map f (rev xs)"
1066 lemma map_eq_conv[simp]: "(map f xs = map g xs) = (!x : set xs. f x = g x)"
1069 lemma map_cong [fundef_cong]:
1070 "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> map f xs = map g ys"
1073 lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])"
1076 lemma Nil_is_map_conv [iff]: "([] = map f xs) = (xs = [])"
1079 lemma map_eq_Cons_conv:
1080 "(map f xs = y#ys) = (\<exists>z zs. xs = z#zs \<and> f z = y \<and> map f zs = ys)"
1083 lemma Cons_eq_map_conv:
1084 "(x#xs = map f ys) = (\<exists>z zs. ys = z#zs \<and> x = f z \<and> xs = map f zs)"
1087 lemmas map_eq_Cons_D = map_eq_Cons_conv [THEN iffD1]
1088 lemmas Cons_eq_map_D = Cons_eq_map_conv [THEN iffD1]
1089 declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!]
1092 "(EX xs. ys = map f xs) = (ALL y : set ys. EX x. y = f x)"
1093 by(induct ys, auto simp add: Cons_eq_map_conv)
1095 lemma map_eq_imp_length_eq:
1096 assumes "map f xs = map g ys"
1097 shows "length xs = length ys"
1099 proof (induct ys arbitrary: xs)
1100 case Nil then show ?case by simp
1102 case (Cons y ys) then obtain z zs where xs: "xs = z # zs" by auto
1103 from Cons xs have "map f zs = map g ys" by simp
1104 with Cons have "length zs = length ys" by blast
1105 with xs show ?case by simp
1109 "[| map f xs = map f ys; inj_on f (set xs Un set ys) |]
1111 apply(frule map_eq_imp_length_eq)
1112 apply(rotate_tac -1)
1113 apply(induct rule:list_induct2)
1116 apply (blast intro:sym)
1119 lemma inj_on_map_eq_map:
1120 "inj_on f (set xs Un set ys) \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
1121 by(blast dest:map_inj_on)
1123 lemma map_injective:
1124 "map f xs = map f ys ==> inj f ==> xs = ys"
1125 by (induct ys arbitrary: xs) (auto dest!:injD)
1127 lemma inj_map_eq_map[simp]: "inj f \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
1128 by(blast dest:map_injective)
1130 lemma inj_mapI: "inj f ==> inj (map f)"
1131 by (iprover dest: map_injective injD intro: inj_onI)
1133 lemma inj_mapD: "inj (map f) ==> inj f"
1134 apply (unfold inj_on_def, clarify)
1135 apply (erule_tac x = "[x]" in ballE)
1136 apply (erule_tac x = "[y]" in ballE, simp, blast)
1140 lemma inj_map[iff]: "inj (map f) = inj f"
1141 by (blast dest: inj_mapD intro: inj_mapI)
1143 lemma inj_on_mapI: "inj_on f (\<Union>(set ` A)) \<Longrightarrow> inj_on (map f) A"
1145 apply(erule map_inj_on)
1146 apply(blast intro:inj_onI dest:inj_onD)
1149 lemma map_idI: "(\<And>x. x \<in> set xs \<Longrightarrow> f x = x) \<Longrightarrow> map f xs = xs"
1150 by (induct xs, auto)
1152 lemma map_fun_upd [simp]: "y \<notin> set xs \<Longrightarrow> map (f(y:=v)) xs = map f xs"
1155 lemma map_fst_zip[simp]:
1156 "length xs = length ys \<Longrightarrow> map fst (zip xs ys) = xs"
1157 by (induct rule:list_induct2, simp_all)
1159 lemma map_snd_zip[simp]:
1160 "length xs = length ys \<Longrightarrow> map snd (zip xs ys) = ys"
1161 by (induct rule:list_induct2, simp_all)
1163 enriched_type map: map
1164 by (simp_all add: id_def)
1166 declare map.id [simp]
1169 subsubsection {* @{const rev} *}
1171 lemma rev_append [simp]: "rev (xs @ ys) = rev ys @ rev xs"
1174 lemma rev_rev_ident [simp]: "rev (rev xs) = xs"
1177 lemma rev_swap: "(rev xs = ys) = (xs = rev ys)"
1180 lemma rev_is_Nil_conv [iff]: "(rev xs = []) = (xs = [])"
1183 lemma Nil_is_rev_conv [iff]: "([] = rev xs) = (xs = [])"
1186 lemma rev_singleton_conv [simp]: "(rev xs = [x]) = (xs = [x])"
1189 lemma singleton_rev_conv [simp]: "([x] = rev xs) = (xs = [x])"
1192 lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)"
1193 apply (induct xs arbitrary: ys, force)
1194 apply (case_tac ys, simp, force)
1197 lemma inj_on_rev[iff]: "inj_on rev A"
1198 by(simp add:inj_on_def)
1200 lemma rev_induct [case_names Nil snoc]:
1201 "[| P []; !!x xs. P xs ==> P (xs @ [x]) |] ==> P xs"
1202 apply(simplesubst rev_rev_ident[symmetric])
1203 apply(rule_tac list = "rev xs" in list.induct, simp_all)
1206 lemma rev_exhaust [case_names Nil snoc]:
1207 "(xs = [] ==> P) ==>(!!ys y. xs = ys @ [y] ==> P) ==> P"
1208 by (induct xs rule: rev_induct) auto
1210 lemmas rev_cases = rev_exhaust
1212 lemma rev_eq_Cons_iff[iff]: "(rev xs = y#ys) = (xs = rev ys @ [y])"
1213 by(rule rev_cases[of xs]) auto
1216 subsubsection {* @{const set} *}
1218 declare set.simps [code_post] --"pretty output"
1220 lemma finite_set [iff]: "finite (set xs)"
1223 lemma set_append [simp]: "set (xs @ ys) = (set xs \<union> set ys)"
1226 lemma hd_in_set[simp]: "xs \<noteq> [] \<Longrightarrow> hd xs : set xs"
1229 lemma set_subset_Cons: "set xs \<subseteq> set (x # xs)"
1232 lemma set_ConsD: "y \<in> set (x # xs) \<Longrightarrow> y=x \<or> y \<in> set xs"
1235 lemma set_empty [iff]: "(set xs = {}) = (xs = [])"
1238 lemma set_empty2[iff]: "({} = set xs) = (xs = [])"
1241 lemma set_rev [simp]: "set (rev xs) = set xs"
1244 lemma set_map [simp]: "set (map f xs) = f`(set xs)"
1247 lemma set_filter [simp]: "set (filter P xs) = {x. x : set xs \<and> P x}"
1250 lemma set_upt [simp]: "set[i..<j] = {i..<j}"
1254 lemma split_list: "x : set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs"
1256 case Nil thus ?case by simp
1258 case Cons thus ?case by (auto intro: Cons_eq_appendI)
1261 lemma in_set_conv_decomp: "x \<in> set xs \<longleftrightarrow> (\<exists>ys zs. xs = ys @ x # zs)"
1262 by (auto elim: split_list)
1264 lemma split_list_first: "x : set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys"
1266 case Nil thus ?case by simp
1271 assume "x = a" thus ?case using Cons by fastforce
1273 assume "x \<noteq> a" thus ?case using Cons by(fastforce intro!: Cons_eq_appendI)
1277 lemma in_set_conv_decomp_first:
1278 "(x : set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys)"
1279 by (auto dest!: split_list_first)
1281 lemma split_list_last: "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs"
1282 proof (induct xs rule: rev_induct)
1283 case Nil thus ?case by simp
1288 assume "x = a" thus ?case using snoc by (metis List.set.simps(1) emptyE)
1290 assume "x \<noteq> a" thus ?case using snoc by fastforce
1294 lemma in_set_conv_decomp_last:
1295 "(x : set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs)"
1296 by (auto dest!: split_list_last)
1298 lemma split_list_prop: "\<exists>x \<in> set xs. P x \<Longrightarrow> \<exists>ys x zs. xs = ys @ x # zs & P x"
1300 case Nil thus ?case by simp
1302 case Cons thus ?case
1303 by(simp add:Bex_def)(metis append_Cons append.simps(1))
1306 lemma split_list_propE:
1307 assumes "\<exists>x \<in> set xs. P x"
1308 obtains ys x zs where "xs = ys @ x # zs" and "P x"
1309 using split_list_prop [OF assms] by blast
1311 lemma split_list_first_prop:
1312 "\<exists>x \<in> set xs. P x \<Longrightarrow>
1313 \<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y)"
1315 case Nil thus ?case by simp
1321 thus ?thesis by simp (metis Un_upper1 contra_subsetD in_set_conv_decomp_first self_append_conv2 set_append)
1324 hence "\<exists>x\<in>set xs. P x" using Cons(2) by simp
1325 thus ?thesis using `\<not> P x` Cons(1) by (metis append_Cons set_ConsD)
1329 lemma split_list_first_propE:
1330 assumes "\<exists>x \<in> set xs. P x"
1331 obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>y \<in> set ys. \<not> P y"
1332 using split_list_first_prop [OF assms] by blast
1334 lemma split_list_first_prop_iff:
1335 "(\<exists>x \<in> set xs. P x) \<longleftrightarrow>
1336 (\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y))"
1337 by (rule, erule split_list_first_prop) auto
1339 lemma split_list_last_prop:
1340 "\<exists>x \<in> set xs. P x \<Longrightarrow>
1341 \<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z)"
1342 proof(induct xs rule:rev_induct)
1343 case Nil thus ?case by simp
1348 assume "P x" thus ?thesis by (metis emptyE set_empty)
1351 hence "\<exists>x\<in>set xs. P x" using snoc(2) by simp
1352 thus ?thesis using `\<not> P x` snoc(1) by fastforce
1356 lemma split_list_last_propE:
1357 assumes "\<exists>x \<in> set xs. P x"
1358 obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>z \<in> set zs. \<not> P z"
1359 using split_list_last_prop [OF assms] by blast
1361 lemma split_list_last_prop_iff:
1362 "(\<exists>x \<in> set xs. P x) \<longleftrightarrow>
1363 (\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z))"
1364 by (metis split_list_last_prop [where P=P] in_set_conv_decomp)
1366 lemma finite_list: "finite A ==> EX xs. set xs = A"
1367 by (erule finite_induct)
1368 (auto simp add: set.simps(2) [symmetric] simp del: set.simps(2))
1370 lemma card_length: "card (set xs) \<le> length xs"
1371 by (induct xs) (auto simp add: card_insert_if)
1373 lemma set_minus_filter_out:
1374 "set xs - {y} = set (filter (\<lambda>x. \<not> (x = y)) xs)"
1378 subsubsection {* @{const filter} *}
1380 lemma filter_append [simp]: "filter P (xs @ ys) = filter P xs @ filter P ys"
1383 lemma rev_filter: "rev (filter P xs) = filter P (rev xs)"
1384 by (induct xs) simp_all
1386 lemma filter_filter [simp]: "filter P (filter Q xs) = filter (\<lambda>x. Q x \<and> P x) xs"
1389 lemma length_filter_le [simp]: "length (filter P xs) \<le> length xs"
1390 by (induct xs) (auto simp add: le_SucI)
1392 lemma sum_length_filter_compl:
1393 "length(filter P xs) + length(filter (%x. ~P x) xs) = length xs"
1394 by(induct xs) simp_all
1396 lemma filter_True [simp]: "\<forall>x \<in> set xs. P x ==> filter P xs = xs"
1399 lemma filter_False [simp]: "\<forall>x \<in> set xs. \<not> P x ==> filter P xs = []"
1402 lemma filter_empty_conv: "(filter P xs = []) = (\<forall>x\<in>set xs. \<not> P x)"
1403 by (induct xs) simp_all
1405 lemma filter_id_conv: "(filter P xs = xs) = (\<forall>x\<in>set xs. P x)"
1408 apply(cut_tac P=P and xs=xs in length_filter_le)
1413 "filter P (map f xs) = map f (filter (P o f) xs)"
1414 by (induct xs) simp_all
1416 lemma length_filter_map[simp]:
1417 "length (filter P (map f xs)) = length(filter (P o f) xs)"
1418 by (simp add:filter_map)
1420 lemma filter_is_subset [simp]: "set (filter P xs) \<le> set xs"
1423 lemma length_filter_less:
1424 "\<lbrakk> x : set xs; ~ P x \<rbrakk> \<Longrightarrow> length(filter P xs) < length xs"
1426 case Nil thus ?case by simp
1428 case (Cons x xs) thus ?case
1429 apply (auto split:split_if_asm)
1430 using length_filter_le[of P xs] apply arith
1434 lemma length_filter_conv_card:
1435 "length(filter p xs) = card{i. i < length xs & p(xs!i)}"
1437 case Nil thus ?case by simp
1440 let ?S = "{i. i < length xs & p(xs!i)}"
1441 have fin: "finite ?S" by(fast intro: bounded_nat_set_is_finite)
1442 show ?case (is "?l = card ?S'")
1445 hence eq: "?S' = insert 0 (Suc ` ?S)"
1446 by(auto simp: image_def split:nat.split dest:gr0_implies_Suc)
1447 have "length (filter p (x # xs)) = Suc(card ?S)"
1448 using Cons `p x` by simp
1449 also have "\<dots> = Suc(card(Suc ` ?S))" using fin
1450 by (simp add: card_image)
1451 also have "\<dots> = card ?S'" using eq fin
1452 by (simp add:card_insert_if) (simp add:image_def)
1453 finally show ?thesis .
1456 hence eq: "?S' = Suc ` ?S"
1457 by(auto simp add: image_def split:nat.split elim:lessE)
1458 have "length (filter p (x # xs)) = card ?S"
1459 using Cons `\<not> p x` by simp
1460 also have "\<dots> = card(Suc ` ?S)" using fin
1461 by (simp add: card_image)
1462 also have "\<dots> = card ?S'" using eq fin
1463 by (simp add:card_insert_if)
1464 finally show ?thesis .
1468 lemma Cons_eq_filterD:
1469 "x#xs = filter P ys \<Longrightarrow>
1470 \<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
1471 (is "_ \<Longrightarrow> \<exists>us vs. ?P ys us vs")
1473 case Nil thus ?case by simp
1476 show ?case (is "\<exists>x. ?Q x")
1482 with Py Cons.prems have "?Q []" by simp
1483 then show ?thesis ..
1485 assume "x \<noteq> y"
1486 with Py Cons.prems show ?thesis by simp
1490 with Cons obtain us vs where "?P (y#ys) (y#us) vs" by fastforce
1491 then have "?Q (y#us)" by simp
1492 then show ?thesis ..
1496 lemma filter_eq_ConsD:
1497 "filter P ys = x#xs \<Longrightarrow>
1498 \<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
1499 by(rule Cons_eq_filterD) simp
1501 lemma filter_eq_Cons_iff:
1502 "(filter P ys = x#xs) =
1503 (\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
1504 by(auto dest:filter_eq_ConsD)
1506 lemma Cons_eq_filter_iff:
1507 "(x#xs = filter P ys) =
1508 (\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
1509 by(auto dest:Cons_eq_filterD)
1511 lemma filter_cong[fundef_cong]:
1512 "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> P x = Q x) \<Longrightarrow> filter P xs = filter Q ys"
1514 apply(erule thin_rl)
1515 by (induct ys) simp_all
1518 subsubsection {* List partitioning *}
1520 primrec partition :: "('a \<Rightarrow> bool) \<Rightarrow>'a list \<Rightarrow> 'a list \<times> 'a list" where
1521 "partition P [] = ([], [])" |
1522 "partition P (x # xs) =
1523 (let (yes, no) = partition P xs
1524 in if P x then (x # yes, no) else (yes, x # no))"
1526 lemma partition_filter1:
1527 "fst (partition P xs) = filter P xs"
1528 by (induct xs) (auto simp add: Let_def split_def)
1530 lemma partition_filter2:
1531 "snd (partition P xs) = filter (Not o P) xs"
1532 by (induct xs) (auto simp add: Let_def split_def)
1535 assumes "partition P xs = (yes, no)"
1536 shows "(\<forall>p \<in> set yes. P p) \<and> (\<forall>p \<in> set no. \<not> P p)"
1538 from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)"
1540 then show ?thesis by (simp_all add: partition_filter1 partition_filter2)
1543 lemma partition_set:
1544 assumes "partition P xs = (yes, no)"
1545 shows "set yes \<union> set no = set xs"
1547 from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)"
1549 then show ?thesis by (auto simp add: partition_filter1 partition_filter2)
1552 lemma partition_filter_conv[simp]:
1553 "partition f xs = (filter f xs,filter (Not o f) xs)"
1554 unfolding partition_filter2[symmetric]
1555 unfolding partition_filter1[symmetric] by simp
1557 declare partition.simps[simp del]
1560 subsubsection {* @{const concat} *}
1562 lemma concat_append [simp]: "concat (xs @ ys) = concat xs @ concat ys"
1565 lemma concat_eq_Nil_conv [simp]: "(concat xss = []) = (\<forall>xs \<in> set xss. xs = [])"
1566 by (induct xss) auto
1568 lemma Nil_eq_concat_conv [simp]: "([] = concat xss) = (\<forall>xs \<in> set xss. xs = [])"
1569 by (induct xss) auto
1571 lemma set_concat [simp]: "set (concat xs) = (UN x:set xs. set x)"
1574 lemma concat_map_singleton[simp]: "concat(map (%x. [f x]) xs) = map f xs"
1577 lemma map_concat: "map f (concat xs) = concat (map (map f) xs)"
1580 lemma filter_concat: "filter p (concat xs) = concat (map (filter p) xs)"
1583 lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))"
1586 lemma concat_eq_concat_iff: "\<forall>(x, y) \<in> set (zip xs ys). length x = length y ==> length xs = length ys ==> (concat xs = concat ys) = (xs = ys)"
1587 proof (induct xs arbitrary: ys)
1589 thus ?case by (cases ys) auto
1592 lemma concat_injective: "concat xs = concat ys ==> length xs = length ys ==> \<forall>(x, y) \<in> set (zip xs ys). length x = length y ==> xs = ys"
1593 by (simp add: concat_eq_concat_iff)
1596 subsubsection {* @{const nth} *}
1598 lemma nth_Cons_0 [simp, code]: "(x # xs)!0 = x"
1601 lemma nth_Cons_Suc [simp, code]: "(x # xs)!(Suc n) = xs!n"
1604 declare nth.simps [simp del]
1606 lemma nth_Cons_pos[simp]: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)"
1607 by(auto simp: Nat.gr0_conv_Suc)
1610 "(xs @ ys)!n = (if n < length xs then xs!n else ys!(n - length xs))"
1611 apply (induct xs arbitrary: n, simp)
1612 apply (case_tac n, auto)
1615 lemma nth_append_length [simp]: "(xs @ x # ys) ! length xs = x"
1618 lemma nth_append_length_plus[simp]: "(xs @ ys) ! (length xs + n) = ys ! n"
1621 lemma nth_map [simp]: "n < length xs ==> (map f xs)!n = f(xs!n)"
1622 apply (induct xs arbitrary: n, simp)
1623 apply (case_tac n, auto)
1627 assumes "n < length (tl x)" shows "tl x ! n = x ! Suc n"
1628 using assms by (induct x) auto
1630 lemma hd_conv_nth: "xs \<noteq> [] \<Longrightarrow> hd xs = xs!0"
1631 by(cases xs) simp_all
1634 lemma list_eq_iff_nth_eq:
1635 "(xs = ys) = (length xs = length ys \<and> (ALL i<length xs. xs!i = ys!i))"
1636 apply(induct xs arbitrary: ys)
1640 apply(simp add:nth_Cons split:nat.split)apply blast
1643 lemma set_conv_nth: "set xs = {xs!i | i. i < length xs}"
1644 apply (induct xs, simp, simp)
1646 apply (metis nat_case_0 nth.simps zero_less_Suc)
1647 apply (metis less_Suc_eq_0_disj nth_Cons_Suc)
1648 apply (case_tac i, simp)
1649 apply (metis diff_Suc_Suc nat_case_Suc nth.simps zero_less_diff)
1652 lemma in_set_conv_nth: "(x \<in> set xs) = (\<exists>i < length xs. xs!i = x)"
1653 by(auto simp:set_conv_nth)
1655 lemma nth_equal_first_eq:
1656 assumes "x \<notin> set xs"
1657 assumes "n \<le> length xs"
1658 shows "(x # xs) ! n = x \<longleftrightarrow> n = 0" (is "?lhs \<longleftrightarrow> ?rhs")
1663 assume "n \<noteq> 0"
1664 then have "n > 0" by simp
1665 with `?lhs` have "xs ! (n - 1) = x" by simp
1666 moreover from `n > 0` `n \<le> length xs` have "n - 1 < length xs" by simp
1667 ultimately have "\<exists>i<length xs. xs ! i = x" by auto
1668 with `x \<notin> set xs` in_set_conv_nth [of x xs] show False by simp
1671 assume ?rhs then show ?lhs by simp
1674 lemma nth_non_equal_first_eq:
1675 assumes "x \<noteq> y"
1676 shows "(x # xs) ! n = y \<longleftrightarrow> xs ! (n - 1) = y \<and> n > 0" (is "?lhs \<longleftrightarrow> ?rhs")
1678 assume "?lhs" with assms have "n > 0" by (cases n) simp_all
1679 with `?lhs` show ?rhs by simp
1681 assume "?rhs" then show "?lhs" by simp
1684 lemma list_ball_nth: "[| n < length xs; !x : set xs. P x|] ==> P(xs!n)"
1685 by (auto simp add: set_conv_nth)
1687 lemma nth_mem [simp]: "n < length xs ==> xs!n : set xs"
1688 by (auto simp add: set_conv_nth)
1690 lemma all_nth_imp_all_set:
1691 "[| !i < length xs. P(xs!i); x : set xs|] ==> P x"
1692 by (auto simp add: set_conv_nth)
1694 lemma all_set_conv_all_nth:
1695 "(\<forall>x \<in> set xs. P x) = (\<forall>i. i < length xs --> P (xs ! i))"
1696 by (auto simp add: set_conv_nth)
1699 "n < size xs \<Longrightarrow> rev xs ! n = xs ! (length xs - Suc n)"
1700 proof (induct xs arbitrary: n)
1701 case Nil thus ?case by simp
1704 hence n: "n < Suc (length xs)" by simp
1706 { assume "n < length xs"
1707 with n obtain n' where n': "length xs - n = Suc n'"
1708 by (cases "length xs - n", auto)
1710 from n' have "length xs - Suc n = n'" by simp
1712 have "xs ! (length xs - Suc n) = (x # xs) ! (length xs - n)" by simp
1715 show ?case by (clarsimp simp add: Cons nth_append)
1718 lemma Skolem_list_nth:
1719 "(ALL i<k. EX x. P i x) = (EX xs. size xs = k & (ALL i<k. P i (xs!i)))"
1720 (is "_ = (EX xs. ?P k xs)")
1722 case 0 show ?case by simp
1725 show ?case (is "?L = ?R" is "_ = (EX xs. ?P' xs)")
1727 assume "?R" thus "?L" using Suc by auto
1730 with Suc obtain x xs where "?P k xs & P k x" by (metis less_Suc_eq)
1731 hence "?P'(xs@[x])" by(simp add:nth_append less_Suc_eq)
1737 subsubsection {* @{const list_update} *}
1739 lemma length_list_update [simp]: "length(xs[i:=x]) = length xs"
1740 by (induct xs arbitrary: i) (auto split: nat.split)
1742 lemma nth_list_update:
1743 "i < length xs==> (xs[i:=x])!j = (if i = j then x else xs!j)"
1744 by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
1746 lemma nth_list_update_eq [simp]: "i < length xs ==> (xs[i:=x])!i = x"
1747 by (simp add: nth_list_update)
1749 lemma nth_list_update_neq [simp]: "i \<noteq> j ==> xs[i:=x]!j = xs!j"
1750 by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
1752 lemma list_update_id[simp]: "xs[i := xs!i] = xs"
1753 by (induct xs arbitrary: i) (simp_all split:nat.splits)
1755 lemma list_update_beyond[simp]: "length xs \<le> i \<Longrightarrow> xs[i:=x] = xs"
1756 apply (induct xs arbitrary: i)
1762 lemma list_update_nonempty[simp]: "xs[k:=x] = [] \<longleftrightarrow> xs=[]"
1763 by(metis length_0_conv length_list_update)
1765 lemma list_update_same_conv:
1766 "i < length xs ==> (xs[i := x] = xs) = (xs!i = x)"
1767 by (induct xs arbitrary: i) (auto split: nat.split)
1769 lemma list_update_append1:
1770 "i < size xs \<Longrightarrow> (xs @ ys)[i:=x] = xs[i:=x] @ ys"
1771 apply (induct xs arbitrary: i, simp)
1772 apply(simp split:nat.split)
1775 lemma list_update_append:
1776 "(xs @ ys) [n:= x] =
1777 (if n < length xs then xs[n:= x] @ ys else xs @ (ys [n-length xs:= x]))"
1778 by (induct xs arbitrary: n) (auto split:nat.splits)
1780 lemma list_update_length [simp]:
1781 "(xs @ x # ys)[length xs := y] = (xs @ y # ys)"
1782 by (induct xs, auto)
1784 lemma map_update: "map f (xs[k:= y]) = (map f xs)[k := f y]"
1785 by(induct xs arbitrary: k)(auto split:nat.splits)
1788 "k < length xs \<Longrightarrow> rev (xs[k:= y]) = (rev xs)[length xs - k - 1 := y]"
1789 by (induct xs arbitrary: k) (auto simp: list_update_append split:nat.splits)
1792 "(zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])"
1793 by (induct ys arbitrary: i xy xs) (auto, case_tac xs, auto split: nat.split)
1795 lemma set_update_subset_insert: "set(xs[i:=x]) <= insert x (set xs)"
1796 by (induct xs arbitrary: i) (auto split: nat.split)
1798 lemma set_update_subsetI: "[| set xs <= A; x:A |] ==> set(xs[i := x]) <= A"
1799 by (blast dest!: set_update_subset_insert [THEN subsetD])
1801 lemma set_update_memI: "n < length xs \<Longrightarrow> x \<in> set (xs[n := x])"
1802 by (induct xs arbitrary: n) (auto split:nat.splits)
1804 lemma list_update_overwrite[simp]:
1805 "xs [i := x, i := y] = xs [i := y]"
1806 apply (induct xs arbitrary: i) apply simp
1807 apply (case_tac i, simp_all)
1810 lemma list_update_swap:
1811 "i \<noteq> i' \<Longrightarrow> xs [i := x, i' := x'] = xs [i' := x', i := x]"
1812 apply (induct xs arbitrary: i i')
1814 apply (case_tac i, case_tac i')
1820 lemma list_update_code [code]:
1822 "(x # xs)[0 := y] = y # xs"
1823 "(x # xs)[Suc i := y] = x # xs[i := y]"
1827 subsubsection {* @{const last} and @{const butlast} *}
1829 lemma last_snoc [simp]: "last (xs @ [x]) = x"
1832 lemma butlast_snoc [simp]: "butlast (xs @ [x]) = xs"
1835 lemma last_ConsL: "xs = [] \<Longrightarrow> last(x#xs) = x"
1838 lemma last_ConsR: "xs \<noteq> [] \<Longrightarrow> last(x#xs) = last xs"
1841 lemma last_append: "last(xs @ ys) = (if ys = [] then last xs else last ys)"
1842 by (induct xs) (auto)
1844 lemma last_appendL[simp]: "ys = [] \<Longrightarrow> last(xs @ ys) = last xs"
1845 by(simp add:last_append)
1847 lemma last_appendR[simp]: "ys \<noteq> [] \<Longrightarrow> last(xs @ ys) = last ys"
1848 by(simp add:last_append)
1850 lemma last_tl: "xs = [] \<or> tl xs \<noteq> [] \<Longrightarrow>last (tl xs) = last xs"
1851 by (induct xs) simp_all
1853 lemma butlast_tl: "butlast (tl xs) = tl (butlast xs)"
1854 by (induct xs) simp_all
1856 lemma hd_rev: "xs \<noteq> [] \<Longrightarrow> hd(rev xs) = last xs"
1857 by(rule rev_exhaust[of xs]) simp_all
1859 lemma last_rev: "xs \<noteq> [] \<Longrightarrow> last(rev xs) = hd xs"
1860 by(cases xs) simp_all
1862 lemma last_in_set[simp]: "as \<noteq> [] \<Longrightarrow> last as \<in> set as"
1865 lemma length_butlast [simp]: "length (butlast xs) = length xs - 1"
1866 by (induct xs rule: rev_induct) auto
1868 lemma butlast_append:
1869 "butlast (xs @ ys) = (if ys = [] then butlast xs else xs @ butlast ys)"
1870 by (induct xs arbitrary: ys) auto
1872 lemma append_butlast_last_id [simp]:
1873 "xs \<noteq> [] ==> butlast xs @ [last xs] = xs"
1876 lemma in_set_butlastD: "x : set (butlast xs) ==> x : set xs"
1877 by (induct xs) (auto split: split_if_asm)
1879 lemma in_set_butlast_appendI:
1880 "x : set (butlast xs) | x : set (butlast ys) ==> x : set (butlast (xs @ ys))"
1881 by (auto dest: in_set_butlastD simp add: butlast_append)
1883 lemma last_drop[simp]: "n < length xs \<Longrightarrow> last (drop n xs) = last xs"
1884 apply (induct xs arbitrary: n)
1886 apply (auto split:nat.split)
1890 assumes "n < length (butlast xs)" shows "butlast xs ! n = xs ! n"
1893 moreover from assms have "butlast xs ! n = (butlast xs @ [last xs]) ! n"
1894 by (simp add: nth_append)
1895 ultimately show ?thesis using append_butlast_last_id by simp
1898 lemma last_conv_nth: "xs\<noteq>[] \<Longrightarrow> last xs = xs!(length xs - 1)"
1899 by(induct xs)(auto simp:neq_Nil_conv)
1901 lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs"
1902 by (induct xs, simp, case_tac xs, simp_all)
1904 lemma last_list_update:
1905 "xs \<noteq> [] \<Longrightarrow> last(xs[k:=x]) = (if k = size xs - 1 then x else last xs)"
1906 by (auto simp: last_conv_nth)
1908 lemma butlast_list_update:
1909 "butlast(xs[k:=x]) =
1910 (if k = size xs - 1 then butlast xs else (butlast xs)[k:=x])"
1911 apply(cases xs rule:rev_cases)
1913 apply(simp add:list_update_append split:nat.splits)
1917 "xs \<noteq> [] \<Longrightarrow> last (map f xs) = f (last xs)"
1918 by (cases xs rule: rev_cases) simp_all
1921 "map f (butlast xs) = butlast (map f xs)"
1922 by (induct xs) simp_all
1924 lemma snoc_eq_iff_butlast:
1925 "xs @ [x] = ys \<longleftrightarrow> (ys \<noteq> [] & butlast ys = xs & last ys = x)"
1926 by (metis append_butlast_last_id append_is_Nil_conv butlast_snoc last_snoc not_Cons_self)
1929 subsubsection {* @{const take} and @{const drop} *}
1931 lemma take_0 [simp]: "take 0 xs = []"
1934 lemma drop_0 [simp]: "drop 0 xs = xs"
1937 lemma take_Suc_Cons [simp]: "take (Suc n) (x # xs) = x # take n xs"
1940 lemma drop_Suc_Cons [simp]: "drop (Suc n) (x # xs) = drop n xs"
1943 declare take_Cons [simp del] and drop_Cons [simp del]
1945 lemma take_1_Cons [simp]: "take 1 (x # xs) = [x]"
1946 unfolding One_nat_def by simp
1948 lemma drop_1_Cons [simp]: "drop 1 (x # xs) = xs"
1949 unfolding One_nat_def by simp
1951 lemma take_Suc: "xs ~= [] ==> take (Suc n) xs = hd xs # take n (tl xs)"
1952 by(clarsimp simp add:neq_Nil_conv)
1954 lemma drop_Suc: "drop (Suc n) xs = drop n (tl xs)"
1955 by(cases xs, simp_all)
1957 lemma take_tl: "take n (tl xs) = tl (take (Suc n) xs)"
1958 by (induct xs arbitrary: n) simp_all
1960 lemma drop_tl: "drop n (tl xs) = tl(drop n xs)"
1961 by(induct xs arbitrary: n, simp_all add:drop_Cons drop_Suc split:nat.split)
1963 lemma tl_take: "tl (take n xs) = take (n - 1) (tl xs)"
1964 by (cases n, simp, cases xs, auto)
1966 lemma tl_drop: "tl (drop n xs) = drop n (tl xs)"
1967 by (simp only: drop_tl)
1969 lemma nth_via_drop: "drop n xs = y#ys \<Longrightarrow> xs!n = y"
1970 apply (induct xs arbitrary: n, simp)
1971 apply(simp add:drop_Cons nth_Cons split:nat.splits)
1974 lemma take_Suc_conv_app_nth:
1975 "i < length xs \<Longrightarrow> take (Suc i) xs = take i xs @ [xs!i]"
1976 apply (induct xs arbitrary: i, simp)
1977 apply (case_tac i, auto)
1980 lemma drop_Suc_conv_tl:
1981 "i < length xs \<Longrightarrow> (xs!i) # (drop (Suc i) xs) = drop i xs"
1982 apply (induct xs arbitrary: i, simp)
1983 apply (case_tac i, auto)
1986 lemma length_take [simp]: "length (take n xs) = min (length xs) n"
1987 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
1989 lemma length_drop [simp]: "length (drop n xs) = (length xs - n)"
1990 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
1992 lemma take_all [simp]: "length xs <= n ==> take n xs = xs"
1993 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
1995 lemma drop_all [simp]: "length xs <= n ==> drop n xs = []"
1996 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
1998 lemma take_append [simp]:
1999 "take n (xs @ ys) = (take n xs @ take (n - length xs) ys)"
2000 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
2002 lemma drop_append [simp]:
2003 "drop n (xs @ ys) = drop n xs @ drop (n - length xs) ys"
2004 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
2006 lemma take_take [simp]: "take n (take m xs) = take (min n m) xs"
2007 apply (induct m arbitrary: xs n, auto)
2008 apply (case_tac xs, auto)
2009 apply (case_tac n, auto)
2012 lemma drop_drop [simp]: "drop n (drop m xs) = drop (n + m) xs"
2013 apply (induct m arbitrary: xs, auto)
2014 apply (case_tac xs, auto)
2017 lemma take_drop: "take n (drop m xs) = drop m (take (n + m) xs)"
2018 apply (induct m arbitrary: xs n, auto)
2019 apply (case_tac xs, auto)
2022 lemma drop_take: "drop n (take m xs) = take (m-n) (drop n xs)"
2023 apply(induct xs arbitrary: m n)
2025 apply(simp add: take_Cons drop_Cons split:nat.split)
2028 lemma append_take_drop_id [simp]: "take n xs @ drop n xs = xs"
2029 apply (induct n arbitrary: xs, auto)
2030 apply (case_tac xs, auto)
2033 lemma take_eq_Nil[simp]: "(take n xs = []) = (n = 0 \<or> xs = [])"
2034 apply(induct xs arbitrary: n)
2036 apply(simp add:take_Cons split:nat.split)
2039 lemma drop_eq_Nil[simp]: "(drop n xs = []) = (length xs <= n)"
2040 apply(induct xs arbitrary: n)
2042 apply(simp add:drop_Cons split:nat.split)
2045 lemma take_map: "take n (map f xs) = map f (take n xs)"
2046 apply (induct n arbitrary: xs, auto)
2047 apply (case_tac xs, auto)
2050 lemma drop_map: "drop n (map f xs) = map f (drop n xs)"
2051 apply (induct n arbitrary: xs, auto)
2052 apply (case_tac xs, auto)
2055 lemma rev_take: "rev (take i xs) = drop (length xs - i) (rev xs)"
2056 apply (induct xs arbitrary: i, auto)
2057 apply (case_tac i, auto)
2060 lemma rev_drop: "rev (drop i xs) = take (length xs - i) (rev xs)"
2061 apply (induct xs arbitrary: i, auto)
2062 apply (case_tac i, auto)
2065 lemma nth_take [simp]: "i < n ==> (take n xs)!i = xs!i"
2066 apply (induct xs arbitrary: i n, auto)
2067 apply (case_tac n, blast)
2068 apply (case_tac i, auto)
2071 lemma nth_drop [simp]:
2072 "n + i <= length xs ==> (drop n xs)!i = xs!(n + i)"
2073 apply (induct n arbitrary: xs i, auto)
2074 apply (case_tac xs, auto)
2078 "n <= length xs ==> butlast (take n xs) = take (n - 1) xs"
2079 by (simp add: butlast_conv_take min.absorb1 min.absorb2)
2081 lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)"
2082 by (simp add: butlast_conv_take drop_take add_ac)
2084 lemma take_butlast: "n < length xs ==> take n (butlast xs) = take n xs"
2085 by (simp add: butlast_conv_take min.absorb1)
2087 lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)"
2088 by (simp add: butlast_conv_take drop_take add_ac)
2090 lemma hd_drop_conv_nth: "n < length xs \<Longrightarrow> hd(drop n xs) = xs!n"
2091 by(simp add: hd_conv_nth)
2093 lemma set_take_subset_set_take:
2094 "m <= n \<Longrightarrow> set(take m xs) <= set(take n xs)"
2095 apply (induct xs arbitrary: m n)
2098 apply (auto simp: take_Cons)
2101 lemma set_take_subset: "set(take n xs) \<subseteq> set xs"
2102 by(induct xs arbitrary: n)(auto simp:take_Cons split:nat.split)
2104 lemma set_drop_subset: "set(drop n xs) \<subseteq> set xs"
2105 by(induct xs arbitrary: n)(auto simp:drop_Cons split:nat.split)
2107 lemma set_drop_subset_set_drop:
2108 "m >= n \<Longrightarrow> set(drop m xs) <= set(drop n xs)"
2109 apply(induct xs arbitrary: m n)
2110 apply(auto simp:drop_Cons split:nat.split)
2111 apply (metis set_drop_subset subset_iff)
2114 lemma in_set_takeD: "x : set(take n xs) \<Longrightarrow> x : set xs"
2115 using set_take_subset by fast
2117 lemma in_set_dropD: "x : set(drop n xs) \<Longrightarrow> x : set xs"
2118 using set_drop_subset by fast
2120 lemma append_eq_conv_conj:
2121 "(xs @ ys = zs) = (xs = take (length xs) zs \<and> ys = drop (length xs) zs)"
2122 apply (induct xs arbitrary: zs, simp, clarsimp)
2123 apply (case_tac zs, auto)
2127 "take (i+j) xs = take i xs @ take j (drop i xs)"
2128 apply (induct xs arbitrary: i, auto)
2129 apply (case_tac i, simp_all)
2132 lemma append_eq_append_conv_if:
2133 "(xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>1 @ ys\<^sub>2) =
2134 (if size xs\<^sub>1 \<le> size ys\<^sub>1
2135 then xs\<^sub>1 = take (size xs\<^sub>1) ys\<^sub>1 \<and> xs\<^sub>2 = drop (size xs\<^sub>1) ys\<^sub>1 @ ys\<^sub>2
2136 else take (size ys\<^sub>1) xs\<^sub>1 = ys\<^sub>1 \<and> drop (size ys\<^sub>1) xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>2)"
2137 apply(induct xs\<^sub>1 arbitrary: ys\<^sub>1)
2139 apply(case_tac ys\<^sub>1)
2144 "n < length xs \<Longrightarrow> take n xs @ [hd (drop n xs)] = take (Suc n) xs"
2145 apply(induct xs arbitrary: n)
2147 apply(simp add:drop_Cons split:nat.split)
2150 lemma id_take_nth_drop:
2151 "i < length xs \<Longrightarrow> xs = take i xs @ xs!i # drop (Suc i) xs"
2153 assume si: "i < length xs"
2154 hence "xs = take (Suc i) xs @ drop (Suc i) xs" by auto
2156 from si have "take (Suc i) xs = take i xs @ [xs!i]"
2157 apply (rule_tac take_Suc_conv_app_nth) by arith
2158 ultimately show ?thesis by auto
2161 lemma upd_conv_take_nth_drop:
2162 "i < length xs \<Longrightarrow> xs[i:=a] = take i xs @ a # drop (Suc i) xs"
2164 assume i: "i < length xs"
2165 have "xs[i:=a] = (take i xs @ xs!i # drop (Suc i) xs)[i:=a]"
2166 by(rule arg_cong[OF id_take_nth_drop[OF i]])
2167 also have "\<dots> = take i xs @ a # drop (Suc i) xs"
2168 using i by (simp add: list_update_append)
2169 finally show ?thesis .
2173 "i < length xs \<Longrightarrow> xs ! i # drop (Suc i) xs = drop i xs"
2174 apply (induct i arbitrary: xs)
2175 apply (simp add: neq_Nil_conv)
2183 subsubsection {* @{const takeWhile} and @{const dropWhile} *}
2185 lemma length_takeWhile_le: "length (takeWhile P xs) \<le> length xs"
2188 lemma takeWhile_dropWhile_id [simp]: "takeWhile P xs @ dropWhile P xs = xs"
2191 lemma takeWhile_append1 [simp]:
2192 "[| x:set xs; ~P(x)|] ==> takeWhile P (xs @ ys) = takeWhile P xs"
2195 lemma takeWhile_append2 [simp]:
2196 "(!!x. x : set xs ==> P x) ==> takeWhile P (xs @ ys) = xs @ takeWhile P ys"
2199 lemma takeWhile_tail: "\<not> P x ==> takeWhile P (xs @ (x#l)) = takeWhile P xs"
2202 lemma takeWhile_nth: "j < length (takeWhile P xs) \<Longrightarrow> takeWhile P xs ! j = xs ! j"
2203 apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto
2205 lemma dropWhile_nth: "j < length (dropWhile P xs) \<Longrightarrow> dropWhile P xs ! j = xs ! (j + length (takeWhile P xs))"
2206 apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto
2208 lemma length_dropWhile_le: "length (dropWhile P xs) \<le> length xs"
2211 lemma dropWhile_append1 [simp]:
2212 "[| x : set xs; ~P(x)|] ==> dropWhile P (xs @ ys) = (dropWhile P xs)@ys"
2215 lemma dropWhile_append2 [simp]:
2216 "(!!x. x:set xs ==> P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys"
2219 lemma dropWhile_append3:
2220 "\<not> P y \<Longrightarrow>dropWhile P (xs @ y # ys) = dropWhile P xs @ y # ys"
2223 lemma dropWhile_last:
2224 "x \<in> set xs \<Longrightarrow> \<not> P x \<Longrightarrow> last (dropWhile P xs) = last xs"
2225 by (auto simp add: dropWhile_append3 in_set_conv_decomp)
2227 lemma set_dropWhileD: "x \<in> set (dropWhile P xs) \<Longrightarrow> x \<in> set xs"
2228 by (induct xs) (auto split: split_if_asm)
2230 lemma set_takeWhileD: "x : set (takeWhile P xs) ==> x : set xs \<and> P x"
2231 by (induct xs) (auto split: split_if_asm)
2233 lemma takeWhile_eq_all_conv[simp]:
2234 "(takeWhile P xs = xs) = (\<forall>x \<in> set xs. P x)"
2237 lemma dropWhile_eq_Nil_conv[simp]:
2238 "(dropWhile P xs = []) = (\<forall>x \<in> set xs. P x)"
2241 lemma dropWhile_eq_Cons_conv:
2242 "(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys & \<not> P y)"
2245 lemma distinct_takeWhile[simp]: "distinct xs ==> distinct (takeWhile P xs)"
2246 by (induct xs) (auto dest: set_takeWhileD)
2248 lemma distinct_dropWhile[simp]: "distinct xs ==> distinct (dropWhile P xs)"
2251 lemma takeWhile_map: "takeWhile P (map f xs) = map f (takeWhile (P \<circ> f) xs)"
2254 lemma dropWhile_map: "dropWhile P (map f xs) = map f (dropWhile (P \<circ> f) xs)"
2257 lemma takeWhile_eq_take: "takeWhile P xs = take (length (takeWhile P xs)) xs"
2260 lemma dropWhile_eq_drop: "dropWhile P xs = drop (length (takeWhile P xs)) xs"
2264 "dropWhile P xs \<noteq> [] \<Longrightarrow> \<not> P (hd (dropWhile P xs))"
2265 using assms by (induct xs) auto
2267 lemma takeWhile_eq_filter:
2268 assumes "\<And> x. x \<in> set (dropWhile P xs) \<Longrightarrow> \<not> P x"
2269 shows "takeWhile P xs = filter P xs"
2271 have A: "filter P xs = filter P (takeWhile P xs @ dropWhile P xs)"
2273 have B: "filter P (dropWhile P xs) = []"
2274 unfolding filter_empty_conv using assms by blast
2275 have "filter P xs = takeWhile P xs"
2276 unfolding A filter_append B
2277 by (auto simp add: filter_id_conv dest: set_takeWhileD)
2281 lemma takeWhile_eq_take_P_nth:
2282 "\<lbrakk> \<And> i. \<lbrakk> i < n ; i < length xs \<rbrakk> \<Longrightarrow> P (xs ! i) ; n < length xs \<Longrightarrow> \<not> P (xs ! n) \<rbrakk> \<Longrightarrow>
2283 takeWhile P xs = take n xs"
2284 proof (induct xs arbitrary: n)
2288 case (Suc n') note this[simp]
2289 have "P x" using Cons.prems(1)[of 0] by simp
2290 moreover have "takeWhile P xs = take n' xs"
2291 proof (rule Cons.hyps)
2292 case goal1 thus "P (xs ! i)" using Cons.prems(1)[of "Suc i"] by simp
2293 next case goal2 thus ?case using Cons by auto
2295 ultimately show ?thesis by simp
2299 lemma nth_length_takeWhile:
2300 "length (takeWhile P xs) < length xs \<Longrightarrow> \<not> P (xs ! length (takeWhile P xs))"
2303 lemma length_takeWhile_less_P_nth:
2304 assumes all: "\<And> i. i < j \<Longrightarrow> P (xs ! i)" and "j \<le> length xs"
2305 shows "j \<le> length (takeWhile P xs)"
2306 proof (rule classical)
2307 assume "\<not> ?thesis"
2308 hence "length (takeWhile P xs) < length xs" using assms by simp
2309 thus ?thesis using all `\<not> ?thesis` nth_length_takeWhile[of P xs] by auto
2312 text{* The following two lemmmas could be generalized to an arbitrary
2315 lemma takeWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow>
2316 takeWhile (\<lambda>y. y \<noteq> x) (rev xs) = rev (tl (dropWhile (\<lambda>y. y \<noteq> x) xs))"
2317 by(induct xs) (auto simp: takeWhile_tail[where l="[]"])
2319 lemma dropWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow>
2320 dropWhile (\<lambda>y. y \<noteq> x) (rev xs) = x # rev (takeWhile (\<lambda>y. y \<noteq> x) xs)"
2324 apply(subst dropWhile_append2)
2328 lemma takeWhile_not_last:
2329 "distinct xs \<Longrightarrow> takeWhile (\<lambda>y. y \<noteq> last xs) xs = butlast xs"
2336 lemma takeWhile_cong [fundef_cong]:
2337 "[| l = k; !!x. x : set l ==> P x = Q x |]
2338 ==> takeWhile P l = takeWhile Q k"
2339 by (induct k arbitrary: l) (simp_all)
2341 lemma dropWhile_cong [fundef_cong]:
2342 "[| l = k; !!x. x : set l ==> P x = Q x |]
2343 ==> dropWhile P l = dropWhile Q k"
2344 by (induct k arbitrary: l, simp_all)
2346 lemma takeWhile_idem [simp]:
2347 "takeWhile P (takeWhile P xs) = takeWhile P xs"
2350 lemma dropWhile_idem [simp]:
2351 "dropWhile P (dropWhile P xs) = dropWhile P xs"
2355 subsubsection {* @{const zip} *}
2357 lemma zip_Nil [simp]: "zip [] ys = []"
2360 lemma zip_Cons_Cons [simp]: "zip (x # xs) (y # ys) = (x, y) # zip xs ys"
2363 declare zip_Cons [simp del]
2368 "zip (x # xs) (y # ys) = (x, y) # zip xs ys"
2369 by (fact zip_Nil zip.simps(1) zip_Cons_Cons)+
2372 "zip (x#xs) ys = (case ys of [] \<Rightarrow> [] | y#ys \<Rightarrow> (x,y)#zip xs ys)"
2373 by(auto split:list.split)
2375 lemma length_zip [simp]:
2376 "length (zip xs ys) = min (length xs) (length ys)"
2377 by (induct xs ys rule:list_induct2') auto
2379 lemma zip_obtain_same_length:
2380 assumes "\<And>zs ws n. length zs = length ws \<Longrightarrow> n = min (length xs) (length ys)
2381 \<Longrightarrow> zs = take n xs \<Longrightarrow> ws = take n ys \<Longrightarrow> P (zip zs ws)"
2382 shows "P (zip xs ys)"
2384 let ?n = "min (length xs) (length ys)"
2385 have "P (zip (take ?n xs) (take ?n ys))"
2386 by (rule assms) simp_all
2387 moreover have "zip xs ys = zip (take ?n xs) (take ?n ys)"
2388 proof (induct xs arbitrary: ys)
2389 case Nil then show ?case by simp
2391 case (Cons x xs) then show ?case by (cases ys) simp_all
2393 ultimately show ?thesis by simp
2398 zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)"
2399 by (induct xs zs rule:list_induct2') auto
2403 zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs"
2404 by (induct xs ys rule:list_induct2') auto
2406 lemma zip_append [simp]:
2407 "[| length xs = length us |] ==>
2408 zip (xs@ys) (us@vs) = zip xs us @ zip ys vs"
2409 by (simp add: zip_append1)
2412 "length xs = length ys ==> zip (rev xs) (rev ys) = rev (zip xs ys)"
2413 by (induct rule:list_induct2, simp_all)
2416 "zip (map f xs) (map g ys) = map (\<lambda> (x, y). (f x, g y)) (zip xs ys)"
2417 proof (induct xs arbitrary: ys)
2418 case (Cons x xs) note Cons_x_xs = Cons.hyps
2422 show ?thesis unfolding Cons using Cons_x_xs by simp
2427 "zip (map f xs) ys = map (\<lambda>(x, y). (f x, y)) (zip xs ys)"
2428 using zip_map_map[of f xs "\<lambda>x. x" ys] by simp
2431 "zip xs (map f ys) = map (\<lambda>(x, y). (x, f y)) (zip xs ys)"
2432 using zip_map_map[of "\<lambda>x. x" xs f ys] by simp
2435 "map f (zip (map g xs) ys) = map (%(x,y). f(g x, y)) (zip xs ys)"
2436 unfolding zip_map1 by auto
2439 "map f (zip xs (map g ys)) = map (%(x,y). f(x, g y)) (zip xs ys)"
2440 unfolding zip_map2 by auto
2442 text{* Courtesy of Andreas Lochbihler: *}
2443 lemma zip_same_conv_map: "zip xs xs = map (\<lambda>x. (x, x)) xs"
2446 lemma nth_zip [simp]:
2447 "[| i < length xs; i < length ys|] ==> (zip xs ys)!i = (xs!i, ys!i)"
2448 apply (induct ys arbitrary: i xs, simp)
2450 apply (simp_all add: nth.simps split: nat.split)
2454 "set (zip xs ys) = {(xs!i, ys!i) | i. i < min (length xs) (length ys)}"
2455 by(simp add: set_conv_nth cong: rev_conj_cong)
2457 lemma zip_same: "((a,b) \<in> set (zip xs xs)) = (a \<in> set xs \<and> a = b)"
2461 "zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]"
2462 by(rule sym, simp add: update_zip)
2464 lemma zip_replicate [simp]:
2465 "zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)"
2466 apply (induct i arbitrary: j, auto)
2467 apply (case_tac j, auto)
2471 "take n (zip xs ys) = zip (take n xs) (take n ys)"
2472 apply (induct n arbitrary: xs ys)
2474 apply (case_tac xs, simp)
2475 apply (case_tac ys, simp_all)
2479 "drop n (zip xs ys) = zip (drop n xs) (drop n ys)"
2480 apply (induct n arbitrary: xs ys)
2482 apply (case_tac xs, simp)
2483 apply (case_tac ys, simp_all)
2486 lemma zip_takeWhile_fst: "zip (takeWhile P xs) ys = takeWhile (P \<circ> fst) (zip xs ys)"
2487 proof (induct xs arbitrary: ys)
2488 case (Cons x xs) thus ?case by (cases ys) auto
2491 lemma zip_takeWhile_snd: "zip xs (takeWhile P ys) = takeWhile (P \<circ> snd) (zip xs ys)"
2492 proof (induct xs arbitrary: ys)
2493 case (Cons x xs) thus ?case by (cases ys) auto
2496 lemma set_zip_leftD:
2497 "(x,y)\<in> set (zip xs ys) \<Longrightarrow> x \<in> set xs"
2498 by (induct xs ys rule:list_induct2') auto
2500 lemma set_zip_rightD:
2501 "(x,y)\<in> set (zip xs ys) \<Longrightarrow> y \<in> set ys"
2502 by (induct xs ys rule:list_induct2') auto
2505 "(x,y) : set(zip xs ys) \<Longrightarrow> (\<lbrakk> x : set xs; y : set ys \<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R"
2506 by(blast dest: set_zip_leftD set_zip_rightD)
2508 lemma zip_map_fst_snd:
2509 "zip (map fst zs) (map snd zs) = zs"
2510 by (induct zs) simp_all
2513 "length xs = length ys \<Longrightarrow> zip xs ys = zs \<longleftrightarrow> map fst zs = xs \<and> map snd zs = ys"
2514 by (auto simp add: zip_map_fst_snd)
2517 "p \<in> set (zip xs ys) \<longleftrightarrow> (\<exists>n. xs ! n = fst p \<and> ys ! n = snd p
2518 \<and> n < length xs \<and> n < length ys)"
2519 by (cases p) (auto simp add: set_zip)
2521 lemma pair_list_eqI:
2522 assumes "map fst xs = map fst ys" and "map snd xs = map snd ys"
2525 from assms(1) have "length xs = length ys" by (rule map_eq_imp_length_eq)
2526 from this assms show ?thesis
2527 by (induct xs ys rule: list_induct2) (simp_all add: prod_eqI)
2531 subsubsection {* @{const list_all2} *}
2533 lemma list_all2_lengthD [intro?]:
2534 "list_all2 P xs ys ==> length xs = length ys"
2535 by (simp add: list_all2_def)
2537 lemma list_all2_Nil [iff, code]: "list_all2 P [] ys = (ys = [])"
2538 by (simp add: list_all2_def)
2540 lemma list_all2_Nil2 [iff, code]: "list_all2 P xs [] = (xs = [])"
2541 by (simp add: list_all2_def)
2543 lemma list_all2_Cons [iff, code]:
2544 "list_all2 P (x # xs) (y # ys) = (P x y \<and> list_all2 P xs ys)"
2545 by (auto simp add: list_all2_def)
2547 lemma list_all2_Cons1:
2548 "list_all2 P (x # xs) ys = (\<exists>z zs. ys = z # zs \<and> P x z \<and> list_all2 P xs zs)"
2551 lemma list_all2_Cons2:
2552 "list_all2 P xs (y # ys) = (\<exists>z zs. xs = z # zs \<and> P z y \<and> list_all2 P zs ys)"
2555 lemma list_all2_induct
2556 [consumes 1, case_names Nil Cons, induct set: list_all2]:
2557 assumes P: "list_all2 P xs ys"
2558 assumes Nil: "R [] []"
2559 assumes Cons: "\<And>x xs y ys.
2560 \<lbrakk>P x y; list_all2 P xs ys; R xs ys\<rbrakk> \<Longrightarrow> R (x # xs) (y # ys)"
2563 by (induct xs arbitrary: ys) (auto simp add: list_all2_Cons1 Nil Cons)
2565 lemma list_all2_rev [iff]:
2566 "list_all2 P (rev xs) (rev ys) = list_all2 P xs ys"
2567 by (simp add: list_all2_def zip_rev cong: conj_cong)
2569 lemma list_all2_rev1:
2570 "list_all2 P (rev xs) ys = list_all2 P xs (rev ys)"
2571 by (subst list_all2_rev [symmetric]) simp
2573 lemma list_all2_append1:
2574 "list_all2 P (xs @ ys) zs =
2575 (EX us vs. zs = us @ vs \<and> length us = length xs \<and> length vs = length ys \<and>
2576 list_all2 P xs us \<and> list_all2 P ys vs)"
2577 apply (simp add: list_all2_def zip_append1)
2579 apply (rule_tac x = "take (length xs) zs" in exI)
2580 apply (rule_tac x = "drop (length xs) zs" in exI)
2581 apply (force split: nat_diff_split simp add: min_def, clarify)
2582 apply (simp add: ball_Un)
2585 lemma list_all2_append2:
2586 "list_all2 P xs (ys @ zs) =
2587 (EX us vs. xs = us @ vs \<and> length us = length ys \<and> length vs = length zs \<and>
2588 list_all2 P us ys \<and> list_all2 P vs zs)"
2589 apply (simp add: list_all2_def zip_append2)
2591 apply (rule_tac x = "take (length ys) xs" in exI)
2592 apply (rule_tac x = "drop (length ys) xs" in exI)
2593 apply (force split: nat_diff_split simp add: min_def, clarify)
2594 apply (simp add: ball_Un)
2597 lemma list_all2_append:
2598 "length xs = length ys \<Longrightarrow>
2599 list_all2 P (xs@us) (ys@vs) = (list_all2 P xs ys \<and> list_all2 P us vs)"
2600 by (induct rule:list_induct2, simp_all)
2602 lemma list_all2_appendI [intro?, trans]:
2603 "\<lbrakk> list_all2 P a b; list_all2 P c d \<rbrakk> \<Longrightarrow> list_all2 P (a@c) (b@d)"
2604 by (simp add: list_all2_append list_all2_lengthD)
2606 lemma list_all2_conv_all_nth:
2607 "list_all2 P xs ys =
2608 (length xs = length ys \<and> (\<forall>i < length xs. P (xs!i) (ys!i)))"
2609 by (force simp add: list_all2_def set_zip)
2611 lemma list_all2_trans:
2612 assumes tr: "!!a b c. P1 a b ==> P2 b c ==> P3 a c"
2613 shows "!!bs cs. list_all2 P1 as bs ==> list_all2 P2 bs cs ==> list_all2 P3 as cs"
2614 (is "!!bs cs. PROP ?Q as bs cs")
2616 fix x xs bs assume I1: "!!bs cs. PROP ?Q xs bs cs"
2617 show "!!cs. PROP ?Q (x # xs) bs cs"
2619 fix y ys cs assume I2: "!!cs. PROP ?Q (x # xs) ys cs"
2620 show "PROP ?Q (x # xs) (y # ys) cs"
2621 by (induct cs) (auto intro: tr I1 I2)
2625 lemma list_all2_all_nthI [intro?]:
2626 "length a = length b \<Longrightarrow> (\<And>n. n < length a \<Longrightarrow> P (a!n) (b!n)) \<Longrightarrow> list_all2 P a b"
2627 by (simp add: list_all2_conv_all_nth)
2630 "\<forall>x \<in> set (zip a b). split P x \<Longrightarrow> length a = length b \<Longrightarrow> list_all2 P a b"
2631 by (simp add: list_all2_def)
2633 lemma list_all2_nthD:
2634 "\<lbrakk> list_all2 P xs ys; p < size xs \<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)"
2635 by (simp add: list_all2_conv_all_nth)
2637 lemma list_all2_nthD2:
2638 "\<lbrakk>list_all2 P xs ys; p < size ys\<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)"
2639 by (frule list_all2_lengthD) (auto intro: list_all2_nthD)
2641 lemma list_all2_map1:
2642 "list_all2 P (map f as) bs = list_all2 (\<lambda>x y. P (f x) y) as bs"
2643 by (simp add: list_all2_conv_all_nth)
2645 lemma list_all2_map2:
2646 "list_all2 P as (map f bs) = list_all2 (\<lambda>x y. P x (f y)) as bs"
2647 by (auto simp add: list_all2_conv_all_nth)
2649 lemma list_all2_refl [intro?]:
2650 "(\<And>x. P x x) \<Longrightarrow> list_all2 P xs xs"
2651 by (simp add: list_all2_conv_all_nth)
2653 lemma list_all2_update_cong:
2654 "\<lbrakk> list_all2 P xs ys; P x y \<rbrakk> \<Longrightarrow> list_all2 P (xs[i:=x]) (ys[i:=y])"
2655 by (cases "i < length ys") (auto simp add: list_all2_conv_all_nth nth_list_update)
2657 lemma list_all2_takeI [simp,intro?]:
2658 "list_all2 P xs ys \<Longrightarrow> list_all2 P (take n xs) (take n ys)"
2659 apply (induct xs arbitrary: n ys)
2661 apply (clarsimp simp add: list_all2_Cons1)
2666 lemma list_all2_dropI [simp,intro?]:
2667 "list_all2 P as bs \<Longrightarrow> list_all2 P (drop n as) (drop n bs)"
2668 apply (induct as arbitrary: n bs, simp)
2669 apply (clarsimp simp add: list_all2_Cons1)
2670 apply (case_tac n, simp, simp)
2673 lemma list_all2_mono [intro?]:
2674 "list_all2 P xs ys \<Longrightarrow> (\<And>xs ys. P xs ys \<Longrightarrow> Q xs ys) \<Longrightarrow> list_all2 Q xs ys"
2675 apply (induct xs arbitrary: ys, simp)
2676 apply (case_tac ys, auto)
2680 "xs = ys \<longleftrightarrow> list_all2 (op =) xs ys"
2681 by (induct xs ys rule: list_induct2') auto
2683 lemma list_eq_iff_zip_eq:
2684 "xs = ys \<longleftrightarrow> length xs = length ys \<and> (\<forall>(x,y) \<in> set (zip xs ys). x = y)"
2685 by(auto simp add: set_zip list_all2_eq list_all2_conv_all_nth cong: conj_cong)
2688 subsubsection {* @{const List.product} and @{const product_lists} *}
2690 lemma product_list_set:
2691 "set (List.product xs ys) = set xs \<times> set ys"
2694 lemma length_product [simp]:
2695 "length (List.product xs ys) = length xs * length ys"
2696 by (induct xs) simp_all
2699 assumes "n < length xs * length ys"
2700 shows "List.product xs ys ! n = (xs ! (n div length ys), ys ! (n mod length ys))"
2701 using assms proof (induct xs arbitrary: n)
2702 case Nil then show ?case by simp
2705 then have "length ys > 0" by auto
2706 with Cons show ?case
2707 by (auto simp add: nth_append not_less le_mod_geq le_div_geq)
2710 lemma in_set_product_lists_length:
2711 "xs \<in> set (product_lists xss) \<Longrightarrow> length xs = length xss"
2712 by (induct xss arbitrary: xs) auto
2714 lemma product_lists_set:
2715 "set (product_lists xss) = {xs. list_all2 (\<lambda>x ys. x \<in> set ys) xs xss}" (is "?L = Collect ?R")
2716 proof (intro equalityI subsetI, unfold mem_Collect_eq)
2717 fix xs assume "xs \<in> ?L"
2718 then have "length xs = length xss" by (rule in_set_product_lists_length)
2719 from this `xs \<in> ?L` show "?R xs" by (induct xs xss rule: list_induct2) auto
2721 fix xs assume "?R xs"
2722 then show "xs \<in> ?L" by induct auto
2726 subsubsection {* @{const fold} with natural argument order *}
2728 lemma fold_simps [code]: -- {* eta-expanded variant for generated code -- enables tail-recursion optimisation in Scala *}
2730 "fold f (x # xs) s = fold f xs (f x s)"
2733 lemma fold_remove1_split:
2734 assumes f: "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f x \<circ> f y = f y \<circ> f x"
2735 and x: "x \<in> set xs"
2736 shows "fold f xs = fold f (remove1 x xs) \<circ> f x"
2737 using assms by (induct xs) (auto simp add: comp_assoc)
2739 lemma fold_cong [fundef_cong]:
2740 "a = b \<Longrightarrow> xs = ys \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> f x = g x)
2741 \<Longrightarrow> fold f xs a = fold g ys b"
2742 by (induct ys arbitrary: a b xs) simp_all
2745 assumes "\<And>x. x \<in> set xs \<Longrightarrow> f x = id"
2746 shows "fold f xs = id"
2747 using assms by (induct xs) simp_all
2750 assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<circ> g x = f x \<circ> h"
2751 shows "h \<circ> fold g xs = fold f xs \<circ> h"
2752 using assms by (induct xs) (simp_all add: fun_eq_iff)
2754 lemma fold_commute_apply:
2755 assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<circ> g x = f x \<circ> h"
2756 shows "h (fold g xs s) = fold f xs (h s)"
2758 from assms have "h \<circ> fold g xs = fold f xs \<circ> h" by (rule fold_commute)
2759 then show ?thesis by (simp add: fun_eq_iff)
2762 lemma fold_invariant:
2763 assumes "\<And>x. x \<in> set xs \<Longrightarrow> Q x" and "P s"
2764 and "\<And>x s. Q x \<Longrightarrow> P s \<Longrightarrow> P (f x s)"
2765 shows "P (fold f xs s)"
2766 using assms by (induct xs arbitrary: s) simp_all
2768 lemma fold_append [simp]:
2769 "fold f (xs @ ys) = fold f ys \<circ> fold f xs"
2770 by (induct xs) simp_all
2772 lemma fold_map [code_unfold]:
2773 "fold g (map f xs) = fold (g o f) xs"
2774 by (induct xs) simp_all
2777 assumes "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y"
2778 shows "fold f (rev xs) = fold f xs"
2779 using assms by (induct xs) (simp_all add: fold_commute_apply fun_eq_iff)
2781 lemma fold_Cons_rev:
2782 "fold Cons xs = append (rev xs)"
2783 by (induct xs) simp_all
2785 lemma rev_conv_fold [code]:
2786 "rev xs = fold Cons xs []"
2787 by (simp add: fold_Cons_rev)
2789 lemma fold_append_concat_rev:
2790 "fold append xss = append (concat (rev xss))"
2791 by (induct xss) simp_all
2793 text {* @{const Finite_Set.fold} and @{const fold} *}
2795 lemma (in comp_fun_commute) fold_set_fold_remdups:
2796 "Finite_Set.fold f y (set xs) = fold f (remdups xs) y"
2797 by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm insert_absorb)
2799 lemma (in comp_fun_idem) fold_set_fold:
2800 "Finite_Set.fold f y (set xs) = fold f xs y"
2801 by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm)
2803 lemma union_set_fold [code]:
2804 "set xs \<union> A = fold Set.insert xs A"
2806 interpret comp_fun_idem Set.insert
2807 by (fact comp_fun_idem_insert)
2808 show ?thesis by (simp add: union_fold_insert fold_set_fold)
2811 lemma union_coset_filter [code]:
2812 "List.coset xs \<union> A = List.coset (List.filter (\<lambda>x. x \<notin> A) xs)"
2815 lemma minus_set_fold [code]:
2816 "A - set xs = fold Set.remove xs A"
2818 interpret comp_fun_idem Set.remove
2819 by (fact comp_fun_idem_remove)
2821 by (simp add: minus_fold_remove [of _ A] fold_set_fold)
2824 lemma minus_coset_filter [code]:
2825 "A - List.coset xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
2828 lemma inter_set_filter [code]:
2829 "A \<inter> set xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
2832 lemma inter_coset_fold [code]:
2833 "A \<inter> List.coset xs = fold Set.remove xs A"
2834 by (simp add: Diff_eq [symmetric] minus_set_fold)
2836 lemma (in semilattice_set) set_eq_fold [code]:
2837 "F (set (x # xs)) = fold f xs x"
2839 interpret comp_fun_idem f
2840 by default (simp_all add: fun_eq_iff left_commute)
2841 show ?thesis by (simp add: eq_fold fold_set_fold)
2844 lemma (in complete_lattice) Inf_set_fold:
2845 "Inf (set xs) = fold inf xs top"
2847 interpret comp_fun_idem "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
2848 by (fact comp_fun_idem_inf)
2849 show ?thesis by (simp add: Inf_fold_inf fold_set_fold inf_commute)
2852 declare Inf_set_fold [where 'a = "'a set", code]
2854 lemma (in complete_lattice) Sup_set_fold:
2855 "Sup (set xs) = fold sup xs bot"
2857 interpret comp_fun_idem "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
2858 by (fact comp_fun_idem_sup)
2859 show ?thesis by (simp add: Sup_fold_sup fold_set_fold sup_commute)
2862 declare Sup_set_fold [where 'a = "'a set", code]
2864 lemma (in complete_lattice) INF_set_fold:
2865 "INFI (set xs) f = fold (inf \<circ> f) xs top"
2866 unfolding INF_def set_map [symmetric] Inf_set_fold fold_map ..
2868 declare INF_set_fold [code]
2870 lemma (in complete_lattice) SUP_set_fold:
2871 "SUPR (set xs) f = fold (sup \<circ> f) xs bot"
2872 unfolding SUP_def set_map [symmetric] Sup_set_fold fold_map ..
2874 declare SUP_set_fold [code]
2877 subsubsection {* Fold variants: @{const foldr} and @{const foldl} *}
2879 text {* Correspondence *}
2881 lemma foldr_conv_fold [code_abbrev]:
2882 "foldr f xs = fold f (rev xs)"
2883 by (induct xs) simp_all
2885 lemma foldl_conv_fold:
2886 "foldl f s xs = fold (\<lambda>x s. f s x) xs s"
2887 by (induct xs arbitrary: s) simp_all
2889 lemma foldr_conv_foldl: -- {* The ``Third Duality Theorem'' in Bird \& Wadler: *}
2890 "foldr f xs a = foldl (\<lambda>x y. f y x) a (rev xs)"
2891 by (simp add: foldr_conv_fold foldl_conv_fold)
2893 lemma foldl_conv_foldr:
2894 "foldl f a xs = foldr (\<lambda>x y. f y x) (rev xs) a"
2895 by (simp add: foldr_conv_fold foldl_conv_fold)
2898 assumes "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y"
2899 shows "foldr f xs = fold f xs"
2900 using assms unfolding foldr_conv_fold by (rule fold_rev)
2902 lemma foldr_cong [fundef_cong]:
2903 "a = b \<Longrightarrow> l = k \<Longrightarrow> (\<And>a x. x \<in> set l \<Longrightarrow> f x a = g x a) \<Longrightarrow> foldr f l a = foldr g k b"
2904 by (auto simp add: foldr_conv_fold intro!: fold_cong)
2906 lemma foldl_cong [fundef_cong]:
2907 "a = b \<Longrightarrow> l = k \<Longrightarrow> (\<And>a x. x \<in> set l \<Longrightarrow> f a x = g a x) \<Longrightarrow> foldl f a l = foldl g b k"
2908 by (auto simp add: foldl_conv_fold intro!: fold_cong)
2910 lemma foldr_append [simp]:
2911 "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)"
2912 by (simp add: foldr_conv_fold)
2914 lemma foldl_append [simp]:
2915 "foldl f a (xs @ ys) = foldl f (foldl f a xs) ys"
2916 by (simp add: foldl_conv_fold)
2918 lemma foldr_map [code_unfold]:
2919 "foldr g (map f xs) a = foldr (g o f) xs a"
2920 by (simp add: foldr_conv_fold fold_map rev_map)
2922 lemma foldl_map [code_unfold]:
2923 "foldl g a (map f xs) = foldl (\<lambda>a x. g a (f x)) a xs"
2924 by (simp add: foldl_conv_fold fold_map comp_def)
2926 lemma concat_conv_foldr [code]:
2927 "concat xss = foldr append xss []"
2928 by (simp add: fold_append_concat_rev foldr_conv_fold)
2931 subsubsection {* @{const upt} *}
2933 lemma upt_rec[code]: "[i..<j] = (if i<j then i#[Suc i..<j] else [])"
2934 -- {* simp does not terminate! *}
2937 lemmas upt_rec_numeral[simp] = upt_rec[of "numeral m" "numeral n"] for m n
2939 lemma upt_conv_Nil [simp]: "j <= i ==> [i..<j] = []"
2940 by (subst upt_rec) simp
2942 lemma upt_eq_Nil_conv[simp]: "([i..<j] = []) = (j = 0 \<or> j <= i)"
2943 by(induct j)simp_all
2945 lemma upt_eq_Cons_conv:
2946 "([i..<j] = x#xs) = (i < j & i = x & [i+1..<j] = xs)"
2947 apply(induct j arbitrary: x xs)
2949 apply(clarsimp simp add: append_eq_Cons_conv)
2953 lemma upt_Suc_append: "i <= j ==> [i..<(Suc j)] = [i..<j]@[j]"
2954 -- {* Only needed if @{text upt_Suc} is deleted from the simpset. *}
2957 lemma upt_conv_Cons: "i < j ==> [i..<j] = i # [Suc i..<j]"
2958 by (simp add: upt_rec)
2960 lemma upt_add_eq_append: "i<=j ==> [i..<j+k] = [i..<j]@[j..<j+k]"
2961 -- {* LOOPS as a simprule, since @{text "j <= j"}. *}
2964 lemma length_upt [simp]: "length [i..<j] = j - i"
2965 by (induct j) (auto simp add: Suc_diff_le)
2967 lemma nth_upt [simp]: "i + k < j ==> [i..<j] ! k = i + k"
2969 apply (auto simp add: less_Suc_eq nth_append split: nat_diff_split)
2973 lemma hd_upt[simp]: "i < j \<Longrightarrow> hd[i..<j] = i"
2974 by(simp add:upt_conv_Cons)
2976 lemma last_upt[simp]: "i < j \<Longrightarrow> last[i..<j] = j - 1"
2979 by(simp add:upt_Suc_append)
2981 lemma take_upt [simp]: "i+m <= n ==> take m [i..<n] = [i..<i+m]"
2982 apply (induct m arbitrary: i, simp)
2983 apply (subst upt_rec)
2985 apply (subst upt_rec)
2986 apply (simp del: upt.simps)
2989 lemma drop_upt[simp]: "drop m [i..<j] = [i+m..<j]"
2994 lemma map_Suc_upt: "map Suc [m..<n] = [Suc m..<Suc n]"
2997 lemma map_add_upt: "map (\<lambda>i. i + n) [0..<m] = [n..<m + n]"
2998 by (induct m) simp_all
3000 lemma nth_map_upt: "i < n-m ==> (map f [m..<n]) ! i = f(m+i)"
3001 apply (induct n m arbitrary: i rule: diff_induct)
3002 prefer 3 apply (subst map_Suc_upt[symmetric])
3003 apply (auto simp add: less_diff_conv)
3007 "map (\<lambda>n. n - Suc 0) [Suc m..<Suc n] = [m..<n]"
3008 by (induct n) simp_all
3010 lemma nth_take_lemma:
3011 "k <= length xs ==> k <= length ys ==>
3012 (!!i. i < k --> xs!i = ys!i) ==> take k xs = take k ys"
3013 apply (atomize, induct k arbitrary: xs ys)
3014 apply (simp_all add: less_Suc_eq_0_disj all_conj_distrib, clarify)
3015 txt {* Both lists must be non-empty *}
3016 apply (case_tac xs, simp)
3017 apply (case_tac ys, clarify)
3018 apply (simp (no_asm_use))
3020 txt {* prenexing's needed, not miniscoping *}
3021 apply (simp (no_asm_use) add: all_simps [symmetric] del: all_simps)
3025 lemma nth_equalityI:
3026 "[| length xs = length ys; ALL i < length xs. xs!i = ys!i |] ==> xs = ys"
3027 by (frule nth_take_lemma [OF le_refl eq_imp_le]) simp_all
3030 "map (\<lambda>i. xs ! i) [0..<length xs] = xs"
3031 by (rule nth_equalityI, auto)
3033 (* needs nth_equalityI *)
3034 lemma list_all2_antisym:
3035 "\<lbrakk> (\<And>x y. \<lbrakk>P x y; Q y x\<rbrakk> \<Longrightarrow> x = y); list_all2 P xs ys; list_all2 Q ys xs \<rbrakk>
3036 \<Longrightarrow> xs = ys"
3037 apply (simp add: list_all2_conv_all_nth)
3038 apply (rule nth_equalityI, blast, simp)
3041 lemma take_equalityI: "(\<forall>i. take i xs = take i ys) ==> xs = ys"
3042 -- {* The famous take-lemma. *}
3043 apply (drule_tac x = "max (length xs) (length ys)" in spec)
3044 apply (simp add: le_max_iff_disj)
3049 "take n (x # xs) = (if n = 0 then [] else x # take (n - 1) xs)"
3050 by (cases n) simp_all
3053 "drop n (x # xs) = (if n = 0 then x # xs else drop (n - 1) xs)"
3054 by (cases n) simp_all
3056 lemma nth_Cons': "(x # xs)!n = (if n = 0 then x else xs!(n - 1))"
3057 by (cases n) simp_all
3059 lemma take_Cons_numeral [simp]:
3060 "take (numeral v) (x # xs) = x # take (numeral v - 1) xs"
3061 by (simp add: take_Cons')
3063 lemma drop_Cons_numeral [simp]:
3064 "drop (numeral v) (x # xs) = drop (numeral v - 1) xs"
3065 by (simp add: drop_Cons')
3067 lemma nth_Cons_numeral [simp]:
3068 "(x # xs) ! numeral v = xs ! (numeral v - 1)"
3069 by (simp add: nth_Cons')
3072 subsubsection {* @{text upto}: interval-list on @{typ int} *}
3074 function upto :: "int \<Rightarrow> int \<Rightarrow> int list" ("(1[_../_])") where
3075 "upto i j = (if i \<le> j then i # [i+1..j] else [])"
3078 by(relation "measure(%(i::int,j). nat(j - i + 1))") auto
3080 declare upto.simps[simp del]
3082 lemmas upto_rec_numeral [simp] =
3083 upto.simps[of "numeral m" "numeral n"]
3084 upto.simps[of "numeral m" "- numeral n"]
3085 upto.simps[of "- numeral m" "numeral n"]
3086 upto.simps[of "- numeral m" "- numeral n"] for m n
3088 lemma upto_empty[simp]: "j < i \<Longrightarrow> [i..j] = []"
3089 by(simp add: upto.simps)
3091 lemma upto_rec1: "i \<le> j \<Longrightarrow> [i..j] = i#[i+1..j]"
3092 by(simp add: upto.simps)
3094 lemma upto_rec2: "i \<le> j \<Longrightarrow> [i..j] = [i..j - 1]@[j]"
3095 proof(induct "nat(j-i)" arbitrary: i j)
3096 case 0 thus ?case by(simp add: upto.simps)
3099 hence "n = nat (j - (i + 1))" "i < j" by linarith+
3100 from this(2) Suc.hyps(1)[OF this(1)] Suc(2,3) upto_rec1 show ?case by simp
3103 lemma set_upto[simp]: "set[i..j] = {i..j}"
3104 proof(induct i j rule:upto.induct)
3106 from this show ?case
3107 unfolding upto.simps[of i j] simp_from_to[of i j] by auto
3110 text{* Tail recursive version for code generation: *}
3112 definition upto_aux :: "int \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" where
3113 "upto_aux i j js = [i..j] @ js"
3115 lemma upto_aux_rec [code]:
3116 "upto_aux i j js = (if j<i then js else upto_aux i (j - 1) (j#js))"
3117 by (simp add: upto_aux_def upto_rec2)
3119 lemma upto_code[code]: "[i..j] = upto_aux i j []"
3120 by(simp add: upto_aux_def)
3123 subsubsection {* @{const distinct} and @{const remdups} and @{const remdups_adj} *}
3126 "distinct xs \<Longrightarrow> distinct (tl xs)"
3127 by (cases xs) simp_all
3129 lemma distinct_append [simp]:
3130 "distinct (xs @ ys) = (distinct xs \<and> distinct ys \<and> set xs \<inter> set ys = {})"
3133 lemma distinct_rev[simp]: "distinct(rev xs) = distinct xs"
3136 lemma set_remdups [simp]: "set (remdups xs) = set xs"
3137 by (induct xs) (auto simp add: insert_absorb)
3139 lemma distinct_remdups [iff]: "distinct (remdups xs)"
3142 lemma distinct_remdups_id: "distinct xs ==> remdups xs = xs"
3143 by (induct xs, auto)
3145 lemma remdups_id_iff_distinct [simp]: "remdups xs = xs \<longleftrightarrow> distinct xs"
3146 by (metis distinct_remdups distinct_remdups_id)
3148 lemma finite_distinct_list: "finite A \<Longrightarrow> EX xs. set xs = A & distinct xs"
3149 by (metis distinct_remdups finite_list set_remdups)
3151 lemma remdups_eq_nil_iff [simp]: "(remdups x = []) = (x = [])"
3154 lemma remdups_eq_nil_right_iff [simp]: "([] = remdups x) = (x = [])"
3157 lemma length_remdups_leq[iff]: "length(remdups xs) <= length xs"
3160 lemma length_remdups_eq[iff]:
3161 "(length (remdups xs) = length xs) = (remdups xs = xs)"
3164 apply(subgoal_tac "length (remdups xs) <= length xs")
3166 apply(rule length_remdups_leq)
3169 lemma remdups_filter: "remdups(filter P xs) = filter P (remdups xs)"
3175 "distinct(map f xs) = (distinct xs & inj_on f (set xs))"
3178 lemma distinct_filter [simp]: "distinct xs ==> distinct (filter P xs)"
3181 lemma distinct_upt[simp]: "distinct[i..<j]"
3184 lemma distinct_upto[simp]: "distinct[i..j]"
3185 apply(induct i j rule:upto.induct)
3186 apply(subst upto.simps)
3190 lemma distinct_take[simp]: "distinct xs \<Longrightarrow> distinct (take i xs)"
3191 apply(induct xs arbitrary: i)
3195 apply(blast dest:in_set_takeD)
3198 lemma distinct_drop[simp]: "distinct xs \<Longrightarrow> distinct (drop i xs)"
3199 apply(induct xs arbitrary: i)
3205 lemma distinct_list_update:
3206 assumes d: "distinct xs" and a: "a \<notin> set xs - {xs!i}"
3207 shows "distinct (xs[i:=a])"
3208 proof (cases "i < length xs")
3210 with a have "a \<notin> set (take i xs @ xs ! i # drop (Suc i) xs) - {xs!i}"
3211 apply (drule_tac id_take_nth_drop) by simp
3212 with d True show ?thesis
3213 apply (simp add: upd_conv_take_nth_drop)
3214 apply (drule subst [OF id_take_nth_drop]) apply assumption
3215 apply simp apply (cases "a = xs!i") apply simp by blast
3217 case False with d show ?thesis by auto
3220 lemma distinct_concat:
3221 assumes "distinct xs"
3222 and "\<And> ys. ys \<in> set xs \<Longrightarrow> distinct ys"
3223 and "\<And> ys zs. \<lbrakk> ys \<in> set xs ; zs \<in> set xs ; ys \<noteq> zs \<rbrakk> \<Longrightarrow> set ys \<inter> set zs = {}"
3224 shows "distinct (concat xs)"
3225 using assms by (induct xs) auto
3227 text {* It is best to avoid this indexed version of distinct, but
3228 sometimes it is useful. *}
3230 lemma distinct_conv_nth:
3231 "distinct xs = (\<forall>i < size xs. \<forall>j < size xs. i \<noteq> j --> xs!i \<noteq> xs!j)"
3232 apply (induct xs, simp, simp)
3233 apply (rule iffI, clarsimp)
3235 apply (case_tac j, simp)
3236 apply (simp add: set_conv_nth)
3238 apply (clarsimp simp add: set_conv_nth, simp)
3241 apply (metis Zero_neq_Suc gr0_conv_Suc in_set_conv_nth lessI less_trans_Suc nth_Cons' nth_Cons_Suc)
3243 apply (clarsimp simp add: set_conv_nth)
3244 apply (erule_tac x = 0 in allE, simp)
3245 apply (erule_tac x = "Suc i" in allE, simp, clarsimp)
3247 apply (metis Suc_Suc_eq lessI less_trans_Suc nth_Cons_Suc)
3249 apply (erule_tac x = "Suc i" in allE, simp)
3250 apply (erule_tac x = "Suc j" in allE, simp)
3253 lemma nth_eq_iff_index_eq:
3254 "\<lbrakk> distinct xs; i < length xs; j < length xs \<rbrakk> \<Longrightarrow> (xs!i = xs!j) = (i = j)"
3255 by(auto simp: distinct_conv_nth)
3257 lemma distinct_card: "distinct xs ==> card (set xs) = size xs"
3260 lemma card_distinct: "card (set xs) = size xs ==> distinct xs"
3262 case Nil thus ?case by simp
3266 proof (cases "x \<in> set xs")
3267 case False with Cons show ?thesis by simp
3269 case True with Cons.prems
3270 have "card (set xs) = Suc (length xs)"
3271 by (simp add: card_insert_if split: split_if_asm)
3272 moreover have "card (set xs) \<le> length xs" by (rule card_length)
3273 ultimately have False by simp
3278 lemma distinct_length_filter: "distinct xs \<Longrightarrow> length (filter P xs) = card ({x. P x} Int set xs)"
3279 by (induct xs) (auto)
3281 lemma not_distinct_decomp: "~ distinct ws ==> EX xs ys zs y. ws = xs@[y]@ys@[y]@zs"
3282 apply (induct n == "length ws" arbitrary:ws) apply simp
3283 apply(case_tac ws) apply simp
3284 apply (simp split:split_if_asm)
3285 apply (metis Cons_eq_appendI eq_Nil_appendI split_list)
3288 lemma not_distinct_conv_prefix:
3289 defines "dec as xs y ys \<equiv> y \<in> set xs \<and> distinct xs \<and> as = xs @ y # ys"
3290 shows "\<not>distinct as \<longleftrightarrow> (\<exists>xs y ys. dec as xs y ys)" (is "?L = ?R")
3292 assume "?L" then show "?R"
3293 proof (induct "length as" arbitrary: as rule: less_induct)
3295 obtain xs ys zs y where decomp: "as = (xs @ y # ys) @ y # zs"
3296 using not_distinct_decomp[OF less.prems] by auto
3298 proof (cases "distinct (xs @ y # ys)")
3300 with decomp have "dec as (xs @ y # ys) y zs" by (simp add: dec_def)
3301 then show ?thesis by blast
3304 with less decomp obtain xs' y' ys' where "dec (xs @ y # ys) xs' y' ys'"
3305 by atomize_elim auto
3306 with decomp have "dec as xs' y' (ys' @ y # zs)" by (simp add: dec_def)
3307 then show ?thesis by blast
3310 qed (auto simp: dec_def)
3312 lemma distinct_product:
3313 assumes "distinct xs" and "distinct ys"
3314 shows "distinct (List.product xs ys)"
3315 using assms by (induct xs)
3316 (auto intro: inj_onI simp add: product_list_set distinct_map)
3318 lemma distinct_product_lists:
3319 assumes "\<forall>xs \<in> set xss. distinct xs"
3320 shows "distinct (product_lists xss)"
3321 using assms proof (induction xss)
3322 case (Cons xs xss) note * = this
3324 proof (cases "product_lists xss")
3325 case Nil then show ?thesis by (induct xs) simp_all
3327 case (Cons ps pss) with * show ?thesis
3328 by (auto intro!: inj_onI distinct_concat simp add: distinct_map)
3332 lemma length_remdups_concat:
3333 "length (remdups (concat xss)) = card (\<Union>xs\<in>set xss. set xs)"
3334 by (simp add: distinct_card [symmetric])
3336 lemma length_remdups_card_conv: "length(remdups xs) = card(set xs)"
3338 have xs: "concat[xs] = xs" by simp
3339 from length_remdups_concat[of "[xs]"] show ?thesis unfolding xs by simp
3342 lemma remdups_remdups:
3343 "remdups (remdups xs) = remdups xs"
3344 by (induct xs) simp_all
3346 lemma distinct_butlast:
3347 assumes "distinct xs"
3348 shows "distinct (butlast xs)"
3349 proof (cases "xs = []")
3351 from `xs \<noteq> []` obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto
3352 with `distinct xs` show ?thesis by simp
3355 lemma remdups_map_remdups:
3356 "remdups (map f (remdups xs)) = remdups (map f xs)"
3357 by (induct xs) simp_all
3359 lemma distinct_zipI1:
3360 assumes "distinct xs"
3361 shows "distinct (zip xs ys)"
3362 proof (rule zip_obtain_same_length)
3363 fix xs' :: "'a list" and ys' :: "'b list" and n
3364 assume "length xs' = length ys'"
3365 assume "xs' = take n xs"
3366 with assms have "distinct xs'" by simp
3367 with `length xs' = length ys'` show "distinct (zip xs' ys')"
3368 by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE)
3371 lemma distinct_zipI2:
3372 assumes "distinct ys"
3373 shows "distinct (zip xs ys)"
3374 proof (rule zip_obtain_same_length)
3375 fix xs' :: "'b list" and ys' :: "'a list" and n
3376 assume "length xs' = length ys'"
3377 assume "ys' = take n ys"
3378 with assms have "distinct ys'" by simp
3379 with `length xs' = length ys'` show "distinct (zip xs' ys')"
3380 by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE)
3383 lemma set_take_disj_set_drop_if_distinct:
3384 "distinct vs \<Longrightarrow> i \<le> j \<Longrightarrow> set (take i vs) \<inter> set (drop j vs) = {}"
3385 by (auto simp: in_set_conv_nth distinct_conv_nth)
3387 (* The next two lemmas help Sledgehammer. *)
3389 lemma distinct_singleton: "distinct [x]" by simp
3391 lemma distinct_length_2_or_more:
3392 "distinct (a # b # xs) \<longleftrightarrow> (a \<noteq> b \<and> distinct (a # xs) \<and> distinct (b # xs))"
3393 by (metis distinct.simps(2) hd.simps hd_in_set list.simps(2) set_ConsD set_rev_mp set_subset_Cons)
3395 lemma remdups_adj_Cons: "remdups_adj (x # xs) =
3396 (case remdups_adj xs of [] \<Rightarrow> [x] | y # xs \<Rightarrow> if x = y then y # xs else x # y # xs)"
3397 by (induct xs arbitrary: x) (auto split: list.splits)
3399 lemma remdups_adj_append_two:
3400 "remdups_adj (xs @ [x,y]) = remdups_adj (xs @ [x]) @ (if x = y then [] else [y])"
3401 by (induct xs rule: remdups_adj.induct, simp_all)
3403 lemma remdups_adj_rev[simp]: "remdups_adj (rev xs) = rev (remdups_adj xs)"
3404 by (induct xs rule: remdups_adj.induct, simp_all add: remdups_adj_append_two)
3406 lemma remdups_adj_length[simp]: "length (remdups_adj xs) \<le> length xs"
3407 by (induct xs rule: remdups_adj.induct, auto)
3409 lemma remdups_adj_length_ge1[simp]: "xs \<noteq> [] \<Longrightarrow> length (remdups_adj xs) \<ge> Suc 0"
3410 by (induct xs rule: remdups_adj.induct, simp_all)
3412 lemma remdups_adj_Nil_iff[simp]: "remdups_adj xs = [] \<longleftrightarrow> xs = []"
3413 by (induct xs rule: remdups_adj.induct, simp_all)
3415 lemma remdups_adj_set[simp]: "set (remdups_adj xs) = set xs"
3416 by (induct xs rule: remdups_adj.induct, simp_all)
3418 lemma remdups_adj_Cons_alt[simp]: "x # tl (remdups_adj (x # xs)) = remdups_adj (x # xs)"
3419 by (induct xs rule: remdups_adj.induct, auto)
3421 lemma remdups_adj_distinct: "distinct xs \<Longrightarrow> remdups_adj xs = xs"
3422 by (induct xs rule: remdups_adj.induct, simp_all)
3424 lemma remdups_adj_append:
3425 "remdups_adj (xs\<^sub>1 @ x # xs\<^sub>2) = remdups_adj (xs\<^sub>1 @ [x]) @ tl (remdups_adj (x # xs\<^sub>2))"
3426 by (induct xs\<^sub>1 rule: remdups_adj.induct, simp_all)
3428 lemma remdups_adj_singleton:
3429 "remdups_adj xs = [x] \<Longrightarrow> xs = replicate (length xs) x"
3430 by (induct xs rule: remdups_adj.induct, auto split: split_if_asm)
3432 lemma remdups_adj_map_injective:
3434 shows "remdups_adj (map f xs) = map f (remdups_adj xs)"
3435 by (induct xs rule: remdups_adj.induct,
3436 auto simp add: injD[OF assms])
3439 subsubsection {* List summation: @{const listsum} and @{text"\<Sum>"}*}
3441 lemma (in monoid_add) listsum_simps [simp]:
3443 "listsum (x # xs) = x + listsum xs"
3444 by (simp_all add: listsum_def)
3446 lemma (in monoid_add) listsum_append [simp]:
3447 "listsum (xs @ ys) = listsum xs + listsum ys"
3448 by (induct xs) (simp_all add: add.assoc)
3450 lemma (in comm_monoid_add) listsum_rev [simp]:
3451 "listsum (rev xs) = listsum xs"
3452 by (simp add: listsum_def foldr_fold fold_rev fun_eq_iff add_ac)
3454 lemma (in monoid_add) fold_plus_listsum_rev:
3455 "fold plus xs = plus (listsum (rev xs))"
3458 have "fold plus xs x = fold plus xs (x + 0)" by simp
3459 also have "\<dots> = fold plus (x # xs) 0" by simp
3460 also have "\<dots> = foldr plus (rev xs @ [x]) 0" by (simp add: foldr_conv_fold)
3461 also have "\<dots> = listsum (rev xs @ [x])" by (simp add: listsum_def)
3462 also have "\<dots> = listsum (rev xs) + listsum [x]" by simp
3463 finally show "fold plus xs x = listsum (rev xs) + x" by simp
3466 text{* Some syntactic sugar for summing a function over a list: *}
3469 "_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3SUM _<-_. _)" [0, 51, 10] 10)
3471 "_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3\<Sum>_\<leftarrow>_. _)" [0, 51, 10] 10)
3472 syntax (HTML output)
3473 "_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3\<Sum>_\<leftarrow>_. _)" [0, 51, 10] 10)
3475 translations -- {* Beware of argument permutation! *}
3476 "SUM x<-xs. b" == "CONST listsum (CONST map (%x. b) xs)"
3477 "\<Sum>x\<leftarrow>xs. b" == "CONST listsum (CONST map (%x. b) xs)"
3479 lemma (in comm_monoid_add) listsum_map_remove1:
3480 "x \<in> set xs \<Longrightarrow> listsum (map f xs) = f x + listsum (map f (remove1 x xs))"
3481 by (induct xs) (auto simp add: ac_simps)
3483 lemma (in monoid_add) list_size_conv_listsum:
3484 "list_size f xs = listsum (map f xs) + size xs"
3487 lemma (in monoid_add) length_concat:
3488 "length (concat xss) = listsum (map length xss)"
3489 by (induct xss) simp_all
3491 lemma (in monoid_add) length_product_lists:
3492 "length (product_lists xss) = foldr op * (map length xss) 1"
3494 case (Cons xs xss) then show ?case by (induct xs) (auto simp: length_concat o_def)
3497 lemma (in monoid_add) listsum_map_filter:
3498 assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<not> P x \<Longrightarrow> f x = 0"
3499 shows "listsum (map f (filter P xs)) = listsum (map f xs)"
3500 using assms by (induct xs) auto
3502 lemma (in comm_monoid_add) distinct_listsum_conv_Setsum:
3503 "distinct xs \<Longrightarrow> listsum xs = Setsum (set xs)"
3504 by (induct xs) simp_all
3506 lemma listsum_eq_0_nat_iff_nat [simp]:
3507 "listsum ns = (0::nat) \<longleftrightarrow> (\<forall>n \<in> set ns. n = 0)"
3508 by (induct ns) simp_all
3510 lemma member_le_listsum_nat:
3511 "(n :: nat) \<in> set ns \<Longrightarrow> n \<le> listsum ns"
3514 lemma elem_le_listsum_nat:
3515 "k < size ns \<Longrightarrow> ns ! k \<le> listsum (ns::nat list)"
3516 by (rule member_le_listsum_nat) simp
3518 lemma listsum_update_nat:
3519 "k < size ns \<Longrightarrow> listsum (ns[k := (n::nat)]) = listsum ns + n - ns ! k"
3520 apply(induct ns arbitrary:k)
3521 apply (auto split:nat.split)
3522 apply(drule elem_le_listsum_nat)
3526 lemma (in monoid_add) listsum_triv:
3527 "(\<Sum>x\<leftarrow>xs. r) = of_nat (length xs) * r"
3528 by (induct xs) (simp_all add: distrib_right)
3530 lemma (in monoid_add) listsum_0 [simp]:
3531 "(\<Sum>x\<leftarrow>xs. 0) = 0"
3532 by (induct xs) (simp_all add: distrib_right)
3534 text{* For non-Abelian groups @{text xs} needs to be reversed on one side: *}
3535 lemma (in ab_group_add) uminus_listsum_map:
3536 "- listsum (map f xs) = listsum (map (uminus \<circ> f) xs)"
3537 by (induct xs) simp_all
3539 lemma (in comm_monoid_add) listsum_addf:
3540 "(\<Sum>x\<leftarrow>xs. f x + g x) = listsum (map f xs) + listsum (map g xs)"
3541 by (induct xs) (simp_all add: algebra_simps)
3543 lemma (in ab_group_add) listsum_subtractf:
3544 "(\<Sum>x\<leftarrow>xs. f x - g x) = listsum (map f xs) - listsum (map g xs)"
3545 by (induct xs) (simp_all add: algebra_simps)
3547 lemma (in semiring_0) listsum_const_mult:
3548 "(\<Sum>x\<leftarrow>xs. c * f x) = c * (\<Sum>x\<leftarrow>xs. f x)"
3549 by (induct xs) (simp_all add: algebra_simps)
3551 lemma (in semiring_0) listsum_mult_const:
3552 "(\<Sum>x\<leftarrow>xs. f x * c) = (\<Sum>x\<leftarrow>xs. f x) * c"
3553 by (induct xs) (simp_all add: algebra_simps)
3555 lemma (in ordered_ab_group_add_abs) listsum_abs:
3556 "\<bar>listsum xs\<bar> \<le> listsum (map abs xs)"
3557 by (induct xs) (simp_all add: order_trans [OF abs_triangle_ineq])
3560 fixes f g :: "'a \<Rightarrow> 'b::{monoid_add, ordered_ab_semigroup_add}"
3561 shows "(\<And>x. x \<in> set xs \<Longrightarrow> f x \<le> g x) \<Longrightarrow> (\<Sum>x\<leftarrow>xs. f x) \<le> (\<Sum>x\<leftarrow>xs. g x)"
3562 by (induct xs) (simp, simp add: add_mono)
3564 lemma (in monoid_add) listsum_distinct_conv_setsum_set:
3565 "distinct xs \<Longrightarrow> listsum (map f xs) = setsum f (set xs)"
3566 by (induct xs) simp_all
3568 lemma (in monoid_add) interv_listsum_conv_setsum_set_nat:
3569 "listsum (map f [m..<n]) = setsum f (set [m..<n])"
3570 by (simp add: listsum_distinct_conv_setsum_set)
3572 lemma (in monoid_add) interv_listsum_conv_setsum_set_int:
3573 "listsum (map f [k..l]) = setsum f (set [k..l])"
3574 by (simp add: listsum_distinct_conv_setsum_set)
3576 text {* General equivalence between @{const listsum} and @{const setsum} *}
3577 lemma (in monoid_add) listsum_setsum_nth:
3578 "listsum xs = (\<Sum> i = 0 ..< length xs. xs ! i)"
3579 using interv_listsum_conv_setsum_set_nat [of "op ! xs" 0 "length xs"] by (simp add: map_nth)
3582 subsubsection {* @{const insert} *}
3584 lemma in_set_insert [simp]:
3585 "x \<in> set xs \<Longrightarrow> List.insert x xs = xs"
3586 by (simp add: List.insert_def)
3588 lemma not_in_set_insert [simp]:
3589 "x \<notin> set xs \<Longrightarrow> List.insert x xs = x # xs"
3590 by (simp add: List.insert_def)
3592 lemma insert_Nil [simp]:
3593 "List.insert x [] = [x]"
3596 lemma set_insert [simp]:
3597 "set (List.insert x xs) = insert x (set xs)"
3598 by (auto simp add: List.insert_def)
3600 lemma distinct_insert [simp]:
3601 "distinct xs \<Longrightarrow> distinct (List.insert x xs)"
3602 by (simp add: List.insert_def)
3604 lemma insert_remdups:
3605 "List.insert x (remdups xs) = remdups (List.insert x xs)"
3606 by (simp add: List.insert_def)
3609 subsubsection {* @{const List.find} *}
3611 lemma find_None_iff: "List.find P xs = None \<longleftrightarrow> \<not> (\<exists>x. x \<in> set xs \<and> P x)"
3612 proof (induction xs)
3613 case Nil thus ?case by simp
3615 case (Cons x xs) thus ?case by (fastforce split: if_splits)
3618 lemma find_Some_iff:
3619 "List.find P xs = Some x \<longleftrightarrow>
3620 (\<exists>i<length xs. P (xs!i) \<and> x = xs!i \<and> (\<forall>j<i. \<not> P (xs!j)))"
3621 proof (induction xs)
3622 case Nil thus ?case by simp
3624 case (Cons x xs) thus ?case
3625 by(auto simp: nth_Cons' split: if_splits)
3626 (metis One_nat_def diff_Suc_1 less_Suc_eq_0_disj)
3629 lemma find_cong[fundef_cong]:
3630 assumes "xs = ys" and "\<And>x. x \<in> set ys \<Longrightarrow> P x = Q x"
3631 shows "List.find P xs = List.find Q ys"
3632 proof (cases "List.find P xs")
3633 case None thus ?thesis by (metis find_None_iff assms)
3636 hence "List.find Q ys = Some x" using assms
3637 by (auto simp add: find_Some_iff)
3638 thus ?thesis using Some by auto
3641 lemma find_dropWhile:
3642 "List.find P xs = (case dropWhile (Not \<circ> P) xs
3643 of [] \<Rightarrow> None
3644 | x # _ \<Rightarrow> Some x)"
3645 by (induct xs) simp_all
3648 subsubsection {* @{const remove1} *}
3650 lemma remove1_append:
3651 "remove1 x (xs @ ys) =
3652 (if x \<in> set xs then remove1 x xs @ ys else xs @ remove1 x ys)"
3655 lemma remove1_commute: "remove1 x (remove1 y zs) = remove1 y (remove1 x zs)"
3658 lemma in_set_remove1[simp]:
3659 "a \<noteq> b \<Longrightarrow> a : set(remove1 b xs) = (a : set xs)"
3664 lemma set_remove1_subset: "set(remove1 x xs) <= set xs"
3671 lemma set_remove1_eq [simp]: "distinct xs ==> set(remove1 x xs) = set xs - {x}"
3678 lemma length_remove1:
3679 "length(remove1 x xs) = (if x : set xs then length xs - 1 else length xs)"
3681 apply (auto dest!:length_pos_if_in_set)
3684 lemma remove1_filter_not[simp]:
3685 "\<not> P x \<Longrightarrow> remove1 x (filter P xs) = filter P xs"
3688 lemma filter_remove1:
3689 "filter Q (remove1 x xs) = remove1 x (filter Q xs)"
3692 lemma notin_set_remove1[simp]: "x ~: set xs ==> x ~: set(remove1 y xs)"
3693 apply(insert set_remove1_subset)
3697 lemma distinct_remove1[simp]: "distinct xs ==> distinct(remove1 x xs)"
3698 by (induct xs) simp_all
3700 lemma remove1_remdups:
3701 "distinct xs \<Longrightarrow> remove1 x (remdups xs) = remdups (remove1 x xs)"
3702 by (induct xs) simp_all
3705 assumes "x \<notin> set xs"
3706 shows "remove1 x xs = xs"
3707 using assms by (induct xs) simp_all
3710 subsubsection {* @{const removeAll} *}
3712 lemma removeAll_filter_not_eq:
3713 "removeAll x = filter (\<lambda>y. x \<noteq> y)"
3716 show "removeAll x xs = filter (\<lambda>y. x \<noteq> y) xs"
3720 lemma removeAll_append[simp]:
3721 "removeAll x (xs @ ys) = removeAll x xs @ removeAll x ys"
3724 lemma set_removeAll[simp]: "set(removeAll x xs) = set xs - {x}"
3727 lemma removeAll_id[simp]: "x \<notin> set xs \<Longrightarrow> removeAll x xs = xs"
3730 (* Needs count:: 'a \<Rightarrow> 'a list \<Rightarrow> nat
3731 lemma length_removeAll:
3732 "length(removeAll x xs) = length xs - count x xs"
3735 lemma removeAll_filter_not[simp]:
3736 "\<not> P x \<Longrightarrow> removeAll x (filter P xs) = filter P xs"
3739 lemma distinct_removeAll:
3740 "distinct xs \<Longrightarrow> distinct (removeAll x xs)"
3741 by (simp add: removeAll_filter_not_eq)
3743 lemma distinct_remove1_removeAll:
3744 "distinct xs ==> remove1 x xs = removeAll x xs"
3745 by (induct xs) simp_all
3747 lemma map_removeAll_inj_on: "inj_on f (insert x (set xs)) \<Longrightarrow>
3748 map f (removeAll x xs) = removeAll (f x) (map f xs)"
3749 by (induct xs) (simp_all add:inj_on_def)
3751 lemma map_removeAll_inj: "inj f \<Longrightarrow>
3752 map f (removeAll x xs) = removeAll (f x) (map f xs)"
3753 by(metis map_removeAll_inj_on subset_inj_on subset_UNIV)
3756 subsubsection {* @{const replicate} *}
3758 lemma length_replicate [simp]: "length (replicate n x) = n"
3761 lemma Ex_list_of_length: "\<exists>xs. length xs = n"
3762 by (rule exI[of _ "replicate n undefined"]) simp
3764 lemma map_replicate [simp]: "map f (replicate n x) = replicate n (f x)"
3767 lemma map_replicate_const:
3768 "map (\<lambda> x. k) lst = replicate (length lst) k"
3769 by (induct lst) auto
3771 lemma replicate_app_Cons_same:
3772 "(replicate n x) @ (x # xs) = x # replicate n x @ xs"
3775 lemma rev_replicate [simp]: "rev (replicate n x) = replicate n x"
3776 apply (induct n, simp)
3777 apply (simp add: replicate_app_Cons_same)
3780 lemma replicate_add: "replicate (n + m) x = replicate n x @ replicate m x"
3783 text{* Courtesy of Matthias Daum: *}
3784 lemma append_replicate_commute:
3785 "replicate n x @ replicate k x = replicate k x @ replicate n x"
3786 apply (simp add: replicate_add [THEN sym])
3787 apply (simp add: add_commute)
3790 text{* Courtesy of Andreas Lochbihler: *}
3791 lemma filter_replicate:
3792 "filter P (replicate n x) = (if P x then replicate n x else [])"
3795 lemma hd_replicate [simp]: "n \<noteq> 0 ==> hd (replicate n x) = x"
3798 lemma tl_replicate [simp]: "tl (replicate n x) = replicate (n - 1) x"
3801 lemma last_replicate [simp]: "n \<noteq> 0 ==> last (replicate n x) = x"
3802 by (atomize (full), induct n) auto
3804 lemma nth_replicate[simp]: "i < n ==> (replicate n x)!i = x"
3805 apply (induct n arbitrary: i, simp)
3806 apply (simp add: nth_Cons split: nat.split)
3809 text{* Courtesy of Matthias Daum (2 lemmas): *}
3810 lemma take_replicate[simp]: "take i (replicate k x) = replicate (min i k) x"
3811 apply (case_tac "k \<le> i")
3812 apply (simp add: min_def)
3813 apply (drule not_leE)
3814 apply (simp add: min_def)
3815 apply (subgoal_tac "replicate k x = replicate i x @ replicate (k - i) x")
3817 apply (simp add: replicate_add [symmetric])
3820 lemma drop_replicate[simp]: "drop i (replicate k x) = replicate (k-i) x"
3821 apply (induct k arbitrary: i)
3829 lemma set_replicate_Suc: "set (replicate (Suc n) x) = {x}"
3832 lemma set_replicate [simp]: "n \<noteq> 0 ==> set (replicate n x) = {x}"
3833 by (fast dest!: not0_implies_Suc intro!: set_replicate_Suc)
3835 lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})"
3838 lemma in_set_replicate[simp]: "(x : set (replicate n y)) = (x = y & n \<noteq> 0)"
3839 by (simp add: set_replicate_conv_if)
3841 lemma Ball_set_replicate[simp]:
3842 "(ALL x : set(replicate n a). P x) = (P a | n=0)"
3843 by(simp add: set_replicate_conv_if)
3845 lemma Bex_set_replicate[simp]:
3846 "(EX x : set(replicate n a). P x) = (P a & n\<noteq>0)"
3847 by(simp add: set_replicate_conv_if)
3849 lemma replicate_append_same:
3850 "replicate i x @ [x] = x # replicate i x"
3851 by (induct i) simp_all
3853 lemma map_replicate_trivial:
3854 "map (\<lambda>i. x) [0..<i] = replicate i x"
3855 by (induct i) (simp_all add: replicate_append_same)
3857 lemma concat_replicate_trivial[simp]:
3858 "concat (replicate i []) = []"
3859 by (induct i) (auto simp add: map_replicate_const)
3861 lemma replicate_empty[simp]: "(replicate n x = []) \<longleftrightarrow> n=0"
3864 lemma empty_replicate[simp]: "([] = replicate n x) \<longleftrightarrow> n=0"
3867 lemma replicate_eq_replicate[simp]:
3868 "(replicate m x = replicate n y) \<longleftrightarrow> (m=n & (m\<noteq>0 \<longrightarrow> x=y))"
3869 apply(induct m arbitrary: n)
3875 lemma replicate_length_filter:
3876 "replicate (length (filter (\<lambda>y. x = y) xs)) x = filter (\<lambda>y. x = y) xs"
3879 lemma comm_append_are_replicate:
3880 fixes xs ys :: "'a list"
3881 assumes "xs \<noteq> []" "ys \<noteq> []"
3882 assumes "xs @ ys = ys @ xs"
3883 shows "\<exists>m n zs. concat (replicate m zs) = xs \<and> concat (replicate n zs) = ys"
3885 proof (induct "length (xs @ ys)" arbitrary: xs ys rule: less_induct)
3888 def xs' \<equiv> "if (length xs \<le> length ys) then xs else ys"
3889 and ys' \<equiv> "if (length xs \<le> length ys) then ys else xs"
3891 prems': "length xs' \<le> length ys'"
3892 "xs' @ ys' = ys' @ xs'"
3893 and "xs' \<noteq> []"
3894 and len: "length (xs @ ys) = length (xs' @ ys')"
3895 using less by (auto intro: less.hyps)
3898 obtain ws where "ys' = xs' @ ws"
3899 by (auto simp: append_eq_append_conv2)
3901 have "\<exists>m n zs. concat (replicate m zs) = xs' \<and> concat (replicate n zs) = ys'"
3902 proof (cases "ws = []")
3904 then have "concat (replicate 1 xs') = xs'"
3905 and "concat (replicate 1 xs') = ys'"
3906 using `ys' = xs' @ ws` by auto
3907 then show ?thesis by blast
3910 from `ys' = xs' @ ws` and `xs' @ ys' = ys' @ xs'`
3911 have "xs' @ ws = ws @ xs'" by simp
3912 then have "\<exists>m n zs. concat (replicate m zs) = xs' \<and> concat (replicate n zs) = ws"
3913 using False and `xs' \<noteq> []` and `ys' = xs' @ ws` and len
3914 by (intro less.hyps) auto
3915 then obtain m n zs where *: "concat (replicate m zs) = xs'"
3916 and "concat (replicate n zs) = ws" by blast
3917 then have "concat (replicate (m + n) zs) = ys'"
3918 using `ys' = xs' @ ws`
3919 by (simp add: replicate_add)
3920 with * show ?thesis by blast
3923 using xs'_def ys'_def by metis
3926 lemma comm_append_is_replicate:
3927 fixes xs ys :: "'a list"
3928 assumes "xs \<noteq> []" "ys \<noteq> []"
3929 assumes "xs @ ys = ys @ xs"
3930 shows "\<exists>n zs. n > 1 \<and> concat (replicate n zs) = xs @ ys"
3933 obtain m n zs where "concat (replicate m zs) = xs"
3934 and "concat (replicate n zs) = ys"
3935 using assms by (metis comm_append_are_replicate)
3936 then have "m + n > 1" and "concat (replicate (m+n) zs) = xs @ ys"
3937 using `xs \<noteq> []` and `ys \<noteq> []`
3938 by (auto simp: replicate_add)
3939 then show ?thesis by blast
3942 lemma Cons_replicate_eq:
3943 "x # xs = replicate n y \<longleftrightarrow> x = y \<and> n > 0 \<and> xs = replicate (n - 1) x"
3946 lemma replicate_length_same:
3947 "(\<forall>y\<in>set xs. y = x) \<Longrightarrow> replicate (length xs) x = xs"
3948 by (induct xs) simp_all
3950 lemma foldr_replicate [simp]:
3951 "foldr f (replicate n x) = f x ^^ n"
3952 by (induct n) (simp_all)
3954 lemma fold_replicate [simp]:
3955 "fold f (replicate n x) = f x ^^ n"
3956 by (subst foldr_fold [symmetric]) simp_all
3959 subsubsection {* @{const enumerate} *}
3961 lemma enumerate_simps [simp, code]:
3962 "enumerate n [] = []"
3963 "enumerate n (x # xs) = (n, x) # enumerate (Suc n) xs"
3964 apply (auto simp add: enumerate_eq_zip not_le)
3965 apply (cases "n < n + length xs")
3966 apply (auto simp add: upt_conv_Cons)
3969 lemma length_enumerate [simp]:
3970 "length (enumerate n xs) = length xs"
3971 by (simp add: enumerate_eq_zip)
3973 lemma map_fst_enumerate [simp]:
3974 "map fst (enumerate n xs) = [n..<n + length xs]"
3975 by (simp add: enumerate_eq_zip)
3977 lemma map_snd_enumerate [simp]:
3978 "map snd (enumerate n xs) = xs"
3979 by (simp add: enumerate_eq_zip)
3981 lemma in_set_enumerate_eq:
3982 "p \<in> set (enumerate n xs) \<longleftrightarrow> n \<le> fst p \<and> fst p < length xs + n \<and> nth xs (fst p - n) = snd p"
3986 moreover assume "m < length xs + n"
3987 ultimately have "[n..<n + length xs] ! (m - n) = m \<and>
3988 xs ! (m - n) = xs ! (m - n) \<and> m - n < length xs" by auto
3989 then have "\<exists>q. [n..<n + length xs] ! q = m \<and>
3990 xs ! q = xs ! (m - n) \<and> q < length xs" ..
3991 } then show ?thesis by (cases p) (auto simp add: enumerate_eq_zip in_set_zip)
3994 lemma nth_enumerate_eq:
3995 assumes "m < length xs"
3996 shows "enumerate n xs ! m = (n + m, xs ! m)"
3997 using assms by (simp add: enumerate_eq_zip)
3999 lemma enumerate_replicate_eq:
4000 "enumerate n (replicate m a) = map (\<lambda>q. (q, a)) [n..<n + m]"
4001 by (rule pair_list_eqI)
4002 (simp_all add: enumerate_eq_zip comp_def map_replicate_const)
4004 lemma enumerate_Suc_eq:
4005 "enumerate (Suc n) xs = map (apfst Suc) (enumerate n xs)"
4006 by (rule pair_list_eqI)
4007 (simp_all add: not_le, simp del: map_map [simp del] add: map_Suc_upt map_map [symmetric])
4009 lemma distinct_enumerate [simp]:
4010 "distinct (enumerate n xs)"
4011 by (simp add: enumerate_eq_zip distinct_zipI1)
4014 subsubsection {* @{const rotate1} and @{const rotate} *}
4016 lemma rotate0[simp]: "rotate 0 = id"
4017 by(simp add:rotate_def)
4019 lemma rotate_Suc[simp]: "rotate (Suc n) xs = rotate1(rotate n xs)"
4020 by(simp add:rotate_def)
4023 "rotate (m+n) = rotate m o rotate n"
4024 by(simp add:rotate_def funpow_add)
4026 lemma rotate_rotate: "rotate m (rotate n xs) = rotate (m+n) xs"
4027 by(simp add:rotate_add)
4029 lemma rotate1_rotate_swap: "rotate1 (rotate n xs) = rotate n (rotate1 xs)"
4030 by(simp add:rotate_def funpow_swap1)
4032 lemma rotate1_length01[simp]: "length xs <= 1 \<Longrightarrow> rotate1 xs = xs"
4033 by(cases xs) simp_all
4035 lemma rotate_length01[simp]: "length xs <= 1 \<Longrightarrow> rotate n xs = xs"
4038 apply (simp add:rotate_def)
4041 lemma rotate1_hd_tl: "xs \<noteq> [] \<Longrightarrow> rotate1 xs = tl xs @ [hd xs]"
4042 by (cases xs) simp_all
4044 lemma rotate_drop_take:
4045 "rotate n xs = drop (n mod length xs) xs @ take (n mod length xs) xs"
4048 apply(simp add:rotate_def)
4049 apply(cases "xs = []")
4051 apply(case_tac "n mod length xs = 0")
4052 apply(simp add:mod_Suc)
4053 apply(simp add: rotate1_hd_tl drop_Suc take_Suc)
4054 apply(simp add:mod_Suc rotate1_hd_tl drop_Suc[symmetric] drop_tl[symmetric]
4055 take_hd_drop linorder_not_le)
4058 lemma rotate_conv_mod: "rotate n xs = rotate (n mod length xs) xs"
4059 by(simp add:rotate_drop_take)
4061 lemma rotate_id[simp]: "n mod length xs = 0 \<Longrightarrow> rotate n xs = xs"
4062 by(simp add:rotate_drop_take)
4064 lemma length_rotate1[simp]: "length(rotate1 xs) = length xs"
4065 by (cases xs) simp_all
4067 lemma length_rotate[simp]: "length(rotate n xs) = length xs"
4068 by (induct n arbitrary: xs) (simp_all add:rotate_def)
4070 lemma distinct1_rotate[simp]: "distinct(rotate1 xs) = distinct xs"
4073 lemma distinct_rotate[simp]: "distinct(rotate n xs) = distinct xs"
4074 by (induct n) (simp_all add:rotate_def)
4076 lemma rotate_map: "rotate n (map f xs) = map f (rotate n xs)"
4077 by(simp add:rotate_drop_take take_map drop_map)
4079 lemma set_rotate1[simp]: "set(rotate1 xs) = set xs"
4082 lemma set_rotate[simp]: "set(rotate n xs) = set xs"
4083 by (induct n) (simp_all add:rotate_def)
4085 lemma rotate1_is_Nil_conv[simp]: "(rotate1 xs = []) = (xs = [])"
4088 lemma rotate_is_Nil_conv[simp]: "(rotate n xs = []) = (xs = [])"
4089 by (induct n) (simp_all add:rotate_def)
4092 "rotate n (rev xs) = rev(rotate (length xs - (n mod length xs)) xs)"
4093 apply(simp add:rotate_drop_take rev_drop rev_take)
4094 apply(cases "length xs = 0")
4096 apply(cases "n mod length xs = 0")
4098 apply(simp add:rotate_drop_take rev_drop rev_take)
4101 lemma hd_rotate_conv_nth: "xs \<noteq> [] \<Longrightarrow> hd(rotate n xs) = xs!(n mod length xs)"
4102 apply(simp add:rotate_drop_take hd_append hd_drop_conv_nth hd_conv_nth)
4103 apply(subgoal_tac "length xs \<noteq> 0")
4105 using mod_less_divisor[of "length xs" n] by arith
4108 subsubsection {* @{const sublist} --- a generalization of @{const nth} to sets *}
4110 lemma sublist_empty [simp]: "sublist xs {} = []"
4111 by (auto simp add: sublist_def)
4113 lemma sublist_nil [simp]: "sublist [] A = []"
4114 by (auto simp add: sublist_def)
4116 lemma length_sublist:
4117 "length(sublist xs I) = card{i. i < length xs \<and> i : I}"
4118 by(simp add: sublist_def length_filter_conv_card cong:conj_cong)
4120 lemma sublist_shift_lemma_Suc:
4121 "map fst (filter (%p. P(Suc(snd p))) (zip xs is)) =
4122 map fst (filter (%p. P(snd p)) (zip xs (map Suc is)))"
4123 apply(induct xs arbitrary: "is")
4125 apply (case_tac "is")
4130 lemma sublist_shift_lemma:
4131 "map fst [p<-zip xs [i..<i + length xs] . snd p : A] =
4132 map fst [p<-zip xs [0..<length xs] . snd p + i : A]"
4133 by (induct xs rule: rev_induct) (simp_all add: add_commute)
4135 lemma sublist_append:
4136 "sublist (l @ l') A = sublist l A @ sublist l' {j. j + length l : A}"
4137 apply (unfold sublist_def)
4138 apply (induct l' rule: rev_induct, simp)
4139 apply (simp add: upt_add_eq_append[of 0] sublist_shift_lemma)
4140 apply (simp add: add_commute)
4144 "sublist (x # l) A = (if 0:A then [x] else []) @ sublist l {j. Suc j : A}"
4145 apply (induct l rule: rev_induct)
4146 apply (simp add: sublist_def)
4147 apply (simp del: append_Cons add: append_Cons[symmetric] sublist_append)
4150 lemma set_sublist: "set(sublist xs I) = {xs!i|i. i<size xs \<and> i \<in> I}"
4151 apply(induct xs arbitrary: I)
4152 apply(auto simp: sublist_Cons nth_Cons split:nat.split dest!: gr0_implies_Suc)
4155 lemma set_sublist_subset: "set(sublist xs I) \<subseteq> set xs"
4156 by(auto simp add:set_sublist)
4158 lemma notin_set_sublistI[simp]: "x \<notin> set xs \<Longrightarrow> x \<notin> set(sublist xs I)"
4159 by(auto simp add:set_sublist)
4161 lemma in_set_sublistD: "x \<in> set(sublist xs I) \<Longrightarrow> x \<in> set xs"
4162 by(auto simp add:set_sublist)
4164 lemma sublist_singleton [simp]: "sublist [x] A = (if 0 : A then [x] else [])"
4165 by (simp add: sublist_Cons)
4168 lemma distinct_sublistI[simp]: "distinct xs \<Longrightarrow> distinct(sublist xs I)"
4169 apply(induct xs arbitrary: I)
4171 apply(auto simp add:sublist_Cons)
4175 lemma sublist_upt_eq_take [simp]: "sublist l {..<n} = take n l"
4176 apply (induct l rule: rev_induct, simp)
4177 apply (simp split: nat_diff_split add: sublist_append)
4180 lemma filter_in_sublist:
4181 "distinct xs \<Longrightarrow> filter (%x. x \<in> set(sublist xs s)) xs = sublist xs s"
4182 proof (induct xs arbitrary: s)
4183 case Nil thus ?case by simp
4186 then have "!x. x: set xs \<longrightarrow> x \<noteq> a" by auto
4187 with Cons show ?case by(simp add: sublist_Cons cong:filter_cong)
4191 subsubsection {* @{const sublists} and @{const List.n_lists} *}
4193 lemma length_sublists:
4194 "length (sublists xs) = 2 ^ length xs"
4195 by (induct xs) (simp_all add: Let_def)
4197 lemma sublists_powset:
4198 "set ` set (sublists xs) = Pow (set xs)"
4200 have aux: "\<And>x A. set ` Cons x ` A = insert x ` set ` A"
4201 by (auto simp add: image_def)
4202 have "set (map set (sublists xs)) = Pow (set xs)"
4204 (simp_all add: aux Let_def Pow_insert Un_commute comp_def del: map_map)
4205 then show ?thesis by simp
4208 lemma distinct_set_sublists:
4209 assumes "distinct xs"
4210 shows "distinct (map set (sublists xs))"
4211 proof (rule card_distinct)
4212 have "finite (set xs)" by rule
4213 then have "card (Pow (set xs)) = 2 ^ card (set xs)" by (rule card_Pow)
4214 with assms distinct_card [of xs]
4215 have "card (Pow (set xs)) = 2 ^ length xs" by simp
4216 then show "card (set (map set (sublists xs))) = length (map set (sublists xs))"
4217 by (simp add: sublists_powset length_sublists)
4220 lemma n_lists_Nil [simp]: "List.n_lists n [] = (if n = 0 then [[]] else [])"
4221 by (induct n) simp_all
4223 lemma length_n_lists: "length (List.n_lists n xs) = length xs ^ n"
4224 by (induct n) (auto simp add: length_concat o_def listsum_triv)
4226 lemma length_n_lists_elem: "ys \<in> set (List.n_lists n xs) \<Longrightarrow> length ys = n"
4227 by (induct n arbitrary: ys) auto
4229 lemma set_n_lists: "set (List.n_lists n xs) = {ys. length ys = n \<and> set ys \<subseteq> set xs}"
4230 proof (rule set_eqI)
4232 show "ys \<in> set (List.n_lists n xs) \<longleftrightarrow> ys \<in> {ys. length ys = n \<and> set ys \<subseteq> set xs}"
4234 have "ys \<in> set (List.n_lists n xs) \<Longrightarrow> length ys = n"
4235 by (induct n arbitrary: ys) auto
4236 moreover have "\<And>x. ys \<in> set (List.n_lists n xs) \<Longrightarrow> x \<in> set ys \<Longrightarrow> x \<in> set xs"
4237 by (induct n arbitrary: ys) auto
4238 moreover have "set ys \<subseteq> set xs \<Longrightarrow> ys \<in> set (List.n_lists (length ys) xs)"
4240 ultimately show ?thesis by auto
4244 lemma distinct_n_lists:
4245 assumes "distinct xs"
4246 shows "distinct (List.n_lists n xs)"
4247 proof (rule card_distinct)
4248 from assms have card_length: "card (set xs) = length xs" by (rule distinct_card)
4249 have "card (set (List.n_lists n xs)) = card (set xs) ^ n"
4251 case 0 then show ?case by simp
4254 moreover have "card (\<Union>ys\<in>set (List.n_lists n xs). (\<lambda>y. y # ys) ` set xs)
4255 = (\<Sum>ys\<in>set (List.n_lists n xs). card ((\<lambda>y. y # ys) ` set xs))"
4256 by (rule card_UN_disjoint) auto
4257 moreover have "\<And>ys. card ((\<lambda>y. y # ys) ` set xs) = card (set xs)"
4258 by (rule card_image) (simp add: inj_on_def)
4259 ultimately show ?case by auto
4261 also have "\<dots> = length xs ^ n" by (simp add: card_length)
4262 finally show "card (set (List.n_lists n xs)) = length (List.n_lists n xs)"
4263 by (simp add: length_n_lists)
4267 subsubsection {* @{const splice} *}
4269 lemma splice_Nil2 [simp, code]: "splice xs [] = xs"
4270 by (cases xs) simp_all
4272 declare splice.simps(1,3)[code]
4273 declare splice.simps(2)[simp del]
4275 lemma length_splice[simp]: "length(splice xs ys) = length xs + length ys"
4276 by (induct xs ys rule: splice.induct) auto
4279 subsubsection {* Transpose *}
4281 function transpose where
4282 "transpose [] = []" |
4283 "transpose ([] # xss) = transpose xss" |
4284 "transpose ((x#xs) # xss) =
4285 (x # [h. (h#t) \<leftarrow> xss]) # transpose (xs # [t. (h#t) \<leftarrow> xss])"
4286 by pat_completeness auto
4288 lemma transpose_aux_filter_head:
4289 "concat (map (case_list [] (\<lambda>h t. [h])) xss) =
4290 map (\<lambda>xs. hd xs) [ys\<leftarrow>xss . ys \<noteq> []]"
4291 by (induct xss) (auto split: list.split)
4293 lemma transpose_aux_filter_tail:
4294 "concat (map (case_list [] (\<lambda>h t. [t])) xss) =
4295 map (\<lambda>xs. tl xs) [ys\<leftarrow>xss . ys \<noteq> []]"
4296 by (induct xss) (auto split: list.split)
4298 lemma transpose_aux_max:
4299 "max (Suc (length xs)) (foldr (\<lambda>xs. max (length xs)) xss 0) =
4300 Suc (max (length xs) (foldr (\<lambda>x. max (length x - Suc 0)) [ys\<leftarrow>xss . ys\<noteq>[]] 0))"
4301 (is "max _ ?foldB = Suc (max _ ?foldA)")
4302 proof (cases "[ys\<leftarrow>xss . ys\<noteq>[]] = []")
4304 hence "foldr (\<lambda>xs. max (length xs)) xss 0 = 0"
4307 then have "x = []" by (cases x) auto
4308 with Cons show ?case by auto
4310 thus ?thesis using True by simp
4314 have foldA: "?foldA = foldr (\<lambda>x. max (length x)) [ys\<leftarrow>xss . ys \<noteq> []] 0 - 1"
4315 by (induct xss) auto
4316 have foldB: "?foldB = foldr (\<lambda>x. max (length x)) [ys\<leftarrow>xss . ys \<noteq> []] 0"
4317 by (induct xss) auto
4322 obtain z zs where zs: "[ys\<leftarrow>xss . ys \<noteq> []] = z#zs" by (auto simp: neq_Nil_conv)
4323 hence "z \<in> set ([ys\<leftarrow>xss . ys \<noteq> []])" by auto
4324 hence "z \<noteq> []" by auto
4327 by (auto simp: max_def intro: less_le_trans)
4330 unfolding foldA foldB max_Suc_Suc[symmetric]
4334 termination transpose
4335 by (relation "measure (\<lambda>xs. foldr (\<lambda>xs. max (length xs)) xs 0 + length xs)")
4336 (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max less_Suc_eq_le)
4338 lemma transpose_empty: "(transpose xs = []) \<longleftrightarrow> (\<forall>x \<in> set xs. x = [])"
4339 by (induct rule: transpose.induct) simp_all
4341 lemma length_transpose:
4342 fixes xs :: "'a list list"
4343 shows "length (transpose xs) = foldr (\<lambda>xs. max (length xs)) xs 0"
4344 by (induct rule: transpose.induct)
4345 (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max
4346 max_Suc_Suc[symmetric] simp del: max_Suc_Suc)
4348 lemma nth_transpose:
4349 fixes xs :: "'a list list"
4350 assumes "i < length (transpose xs)"
4351 shows "transpose xs ! i = map (\<lambda>xs. xs ! i) [ys \<leftarrow> xs. i < length ys]"
4352 using assms proof (induct arbitrary: i rule: transpose.induct)
4354 def XS == "(x # xs) # xss"
4355 hence [simp]: "XS \<noteq> []" by auto
4359 thus ?thesis by (simp add: transpose_aux_filter_head hd_conv_nth)
4362 have *: "\<And>xss. xs # map tl xss = map tl ((x#xs)#xss)" by simp
4363 have **: "\<And>xss. (x#xs) # filter (\<lambda>ys. ys \<noteq> []) xss = filter (\<lambda>ys. ys \<noteq> []) ((x#xs)#xss)" by simp
4364 { fix x have "Suc j < length x \<longleftrightarrow> x \<noteq> [] \<and> j < length x - Suc 0"
4365 by (cases x) simp_all
4368 have j_less: "j < length (transpose (xs # concat (map (case_list [] (\<lambda>h t. [t])) xss)))"
4369 using "3.prems" by (simp add: transpose_aux_filter_tail length_transpose Suc)
4372 unfolding transpose.simps `i = Suc j` nth_Cons_Suc "3.hyps"[OF j_less]
4373 apply (auto simp: transpose_aux_filter_tail filter_map comp_def length_transpose * ** *** XS_def[symmetric])
4374 apply (rule list.exhaust)
4379 lemma transpose_map_map:
4380 "transpose (map (map f) xs) = map (map f) (transpose xs)"
4381 proof (rule nth_equalityI, safe)
4382 have [simp]: "length (transpose (map (map f) xs)) = length (transpose xs)"
4383 by (simp add: length_transpose foldr_map comp_def)
4384 show "length (transpose (map (map f) xs)) = length (map (map f) (transpose xs))" by simp
4386 fix i assume "i < length (transpose (map (map f) xs))"
4387 thus "transpose (map (map f) xs) ! i = map (map f) (transpose xs) ! i"
4388 by (simp add: nth_transpose filter_map comp_def)
4392 subsubsection {* (In)finiteness *}
4394 lemma finite_maxlen:
4395 "finite (M::'a list set) ==> EX n. ALL s:M. size s < n"
4396 proof (induct rule: finite.induct)
4397 case emptyI show ?case by simp
4400 then obtain n where "\<forall>s\<in>M. length s < n" by blast
4401 hence "ALL s:insert xs M. size s < max n (size xs) + 1" by auto
4405 lemma lists_length_Suc_eq:
4406 "{xs. set xs \<subseteq> A \<and> length xs = Suc n} =
4407 (\<lambda>(xs, n). n#xs) ` ({xs. set xs \<subseteq> A \<and> length xs = n} \<times> A)"
4408 by (auto simp: length_Suc_conv)
4412 shows finite_lists_length_eq: "finite {xs. set xs \<subseteq> A \<and> length xs = n}"
4413 and card_lists_length_eq: "card {xs. set xs \<subseteq> A \<and> length xs = n} = (card A)^n"
4416 (auto simp: card_image inj_split_Cons lists_length_Suc_eq cong: conj_cong)
4418 lemma finite_lists_length_le:
4419 assumes "finite A" shows "finite {xs. set xs \<subseteq> A \<and> length xs \<le> n}"
4422 have "?S = (\<Union>n\<in>{0..n}. {xs. set xs \<subseteq> A \<and> length xs = n})" by auto
4423 thus ?thesis by (auto intro!: finite_lists_length_eq[OF `finite A`] simp only:)
4426 lemma card_lists_length_le:
4427 assumes "finite A" shows "card {xs. set xs \<subseteq> A \<and> length xs \<le> n} = (\<Sum>i\<le>n. card A^i)"
4429 have "(\<Sum>i\<le>n. card A^i) = card (\<Union>i\<le>n. {xs. set xs \<subseteq> A \<and> length xs = i})"
4431 by (subst card_UN_disjoint)
4432 (auto simp add: card_lists_length_eq finite_lists_length_eq)
4433 also have "(\<Union>i\<le>n. {xs. set xs \<subseteq> A \<and> length xs = i}) = {xs. set xs \<subseteq> A \<and> length xs \<le> n}"
4435 finally show ?thesis by simp
4438 lemma card_lists_distinct_length_eq:
4439 assumes "k < card A"
4440 shows "card {xs. length xs = k \<and> distinct xs \<and> set xs \<subseteq> A} = \<Prod>{card A - k + 1 .. card A}"
4444 then have "{xs. length xs = 0 \<and> distinct xs \<and> set xs \<subseteq> A} = {[]}" by auto
4445 then show ?case by simp
4448 let "?k_list" = "\<lambda>k xs. length xs = k \<and> distinct xs \<and> set xs \<subseteq> A"
4449 have inj_Cons: "\<And>A. inj_on (\<lambda>(xs, n). n # xs) A" by (rule inj_onI) auto
4451 from Suc have "k < card A" by simp
4452 moreover have "finite A" using assms by (simp add: card_ge_0_finite)
4453 moreover have "finite {xs. ?k_list k xs}"
4454 using finite_lists_length_eq[OF `finite A`, of k]
4455 by - (rule finite_subset, auto)
4456 moreover have "\<And>i j. i \<noteq> j \<longrightarrow> {i} \<times> (A - set i) \<inter> {j} \<times> (A - set j) = {}"
4458 moreover have "\<And>i. i \<in>Collect (?k_list k) \<Longrightarrow> card (A - set i) = card A - k"
4459 by (simp add: card_Diff_subset distinct_card)
4460 moreover have "{xs. ?k_list (Suc k) xs} =
4461 (\<lambda>(xs, n). n#xs) ` \<Union>((\<lambda>xs. {xs} \<times> (A - set xs)) ` {xs. ?k_list k xs})"
4462 by (auto simp: length_Suc_conv)
4464 have "Suc (card A - Suc k) = card A - k" using Suc.prems by simp
4465 then have "(card A - k) * \<Prod>{Suc (card A - k)..card A} = \<Prod>{Suc (card A - Suc k)..card A}"
4466 by (subst setprod_insert[symmetric]) (simp add: atLeastAtMost_insertL)+
4467 ultimately show ?case
4468 by (simp add: card_image inj_Cons card_UN_disjoint Suc.hyps algebra_simps)
4471 lemma infinite_UNIV_listI: "~ finite(UNIV::'a list set)"
4473 apply(drule finite_maxlen)
4474 apply (metis UNIV_I length_replicate less_not_refl)
4478 subsection {* Sorting *}
4480 text{* Currently it is not shown that @{const sort} returns a
4481 permutation of its input because the nicest proof is via multisets,
4482 which are not yet available. Alternatively one could define a function
4483 that counts the number of occurrences of an element in a list and use
4484 that instead of multisets to state the correctness property. *}
4489 lemma set_insort_key:
4490 "set (insort_key f x xs) = insert x (set xs)"
4493 lemma length_insort [simp]:
4494 "length (insort_key f x xs) = Suc (length xs)"
4495 by (induct xs) simp_all
4497 lemma insort_key_left_comm:
4498 assumes "f x \<noteq> f y"
4499 shows "insort_key f y (insort_key f x xs) = insort_key f x (insort_key f y xs)"
4500 by (induct xs) (auto simp add: assms dest: antisym)
4502 lemma insort_left_comm:
4503 "insort x (insort y xs) = insort y (insort x xs)"
4504 by (cases "x = y") (auto intro: insort_key_left_comm)
4506 lemma comp_fun_commute_insort:
4507 "comp_fun_commute insort"
4509 qed (simp add: insort_left_comm fun_eq_iff)
4511 lemma sort_key_simps [simp]:
4512 "sort_key f [] = []"
4513 "sort_key f (x#xs) = insort_key f x (sort_key f xs)"
4514 by (simp_all add: sort_key_def)
4516 lemma (in linorder) sort_key_conv_fold:
4517 assumes "inj_on f (set xs)"
4518 shows "sort_key f xs = fold (insort_key f) xs []"
4520 have "fold (insort_key f) (rev xs) = fold (insort_key f) xs"
4521 proof (rule fold_rev, rule ext)
4524 assume "x \<in> set xs" "y \<in> set xs"
4525 with assms have *: "f y = f x \<Longrightarrow> y = x" by (auto dest: inj_onD)
4526 have **: "x = y \<longleftrightarrow> y = x" by auto
4527 show "(insort_key f y \<circ> insort_key f x) zs = (insort_key f x \<circ> insort_key f y) zs"
4528 by (induct zs) (auto intro: * simp add: **)
4530 then show ?thesis by (simp add: sort_key_def foldr_conv_fold)
4533 lemma (in linorder) sort_conv_fold:
4534 "sort xs = fold insort xs []"
4535 by (rule sort_key_conv_fold) simp
4537 lemma length_sort[simp]: "length (sort_key f xs) = length xs"
4538 by (induct xs, auto)
4540 lemma sorted_Cons: "sorted (x#xs) = (sorted xs & (ALL y:set xs. x <= y))"
4541 apply(induct xs arbitrary: x) apply simp
4542 by simp (blast intro: order_trans)
4545 "sorted xs \<Longrightarrow> sorted (tl xs)"
4546 by (cases xs) (simp_all add: sorted_Cons)
4548 lemma sorted_append:
4549 "sorted (xs@ys) = (sorted xs & sorted ys & (\<forall>x \<in> set xs. \<forall>y \<in> set ys. x\<le>y))"
4550 by (induct xs) (auto simp add:sorted_Cons)
4552 lemma sorted_nth_mono:
4553 "sorted xs \<Longrightarrow> i \<le> j \<Longrightarrow> j < length xs \<Longrightarrow> xs!i \<le> xs!j"
4554 by (induct xs arbitrary: i j) (auto simp:nth_Cons' sorted_Cons)
4556 lemma sorted_rev_nth_mono:
4557 "sorted (rev xs) \<Longrightarrow> i \<le> j \<Longrightarrow> j < length xs \<Longrightarrow> xs!j \<le> xs!i"
4558 using sorted_nth_mono[ of "rev xs" "length xs - j - 1" "length xs - i - 1"]
4559 rev_nth[of "length xs - i - 1" "xs"] rev_nth[of "length xs - j - 1" "xs"]
4562 lemma sorted_nth_monoI:
4563 "(\<And> i j. \<lbrakk> i \<le> j ; j < length xs \<rbrakk> \<Longrightarrow> xs ! i \<le> xs ! j) \<Longrightarrow> sorted xs"
4567 proof (rule Cons.hyps)
4568 fix i j assume "i \<le> j" and "j < length xs"
4569 with Cons.prems[of "Suc i" "Suc j"]
4570 show "xs ! i \<le> xs ! j" by auto
4574 fix y assume "y \<in> set xs"
4575 then obtain j where "j < length xs" and "xs ! j = y"
4576 unfolding in_set_conv_nth by blast
4577 with Cons.prems[of 0 "Suc j"]
4583 unfolding sorted_Cons by auto
4586 lemma sorted_equals_nth_mono:
4587 "sorted xs = (\<forall>j < length xs. \<forall>i \<le> j. xs ! i \<le> xs ! j)"
4588 by (auto intro: sorted_nth_monoI sorted_nth_mono)
4590 lemma set_insort: "set(insort_key f x xs) = insert x (set xs)"
4593 lemma set_sort[simp]: "set(sort_key f xs) = set xs"
4594 by (induct xs) (simp_all add:set_insort)
4596 lemma distinct_insort: "distinct (insort_key f x xs) = (x \<notin> set xs \<and> distinct xs)"
4597 by(induct xs)(auto simp:set_insort)
4599 lemma distinct_sort[simp]: "distinct (sort_key f xs) = distinct xs"
4600 by (induct xs) (simp_all add: distinct_insort)
4602 lemma sorted_insort_key: "sorted (map f (insort_key f x xs)) = sorted (map f xs)"
4603 by (induct xs) (auto simp:sorted_Cons set_insort)
4605 lemma sorted_insort: "sorted (insort x xs) = sorted xs"
4606 using sorted_insort_key [where f="\<lambda>x. x"] by simp
4608 theorem sorted_sort_key [simp]: "sorted (map f (sort_key f xs))"
4609 by (induct xs) (auto simp:sorted_insort_key)
4611 theorem sorted_sort [simp]: "sorted (sort xs)"
4612 using sorted_sort_key [where f="\<lambda>x. x"] by simp
4614 lemma sorted_butlast:
4615 assumes "xs \<noteq> []" and "sorted xs"
4616 shows "sorted (butlast xs)"
4618 from `xs \<noteq> []` obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto
4619 with `sorted xs` show ?thesis by (simp add: sorted_append)
4622 lemma insort_not_Nil [simp]:
4623 "insort_key f a xs \<noteq> []"
4624 by (induct xs) simp_all
4626 lemma insort_is_Cons: "\<forall>x\<in>set xs. f a \<le> f x \<Longrightarrow> insort_key f a xs = a # xs"
4629 lemma sorted_sort_id: "sorted xs \<Longrightarrow> sort xs = xs"
4630 by (induct xs) (auto simp add: sorted_Cons insort_is_Cons)
4632 lemma sorted_map_remove1:
4633 "sorted (map f xs) \<Longrightarrow> sorted (map f (remove1 x xs))"
4634 by (induct xs) (auto simp add: sorted_Cons)
4636 lemma sorted_remove1: "sorted xs \<Longrightarrow> sorted (remove1 a xs)"
4637 using sorted_map_remove1 [of "\<lambda>x. x"] by simp
4639 lemma insort_key_remove1:
4640 assumes "a \<in> set xs" and "sorted (map f xs)" and "hd (filter (\<lambda>x. f a = f x) xs) = a"
4641 shows "insort_key f a (remove1 a xs) = xs"
4642 using assms proof (induct xs)
4645 proof (cases "x = a")
4647 then have "f x \<noteq> f a" using Cons.prems by auto
4648 then have "f x < f a" using Cons.prems by (auto simp: sorted_Cons)
4649 with `f x \<noteq> f a` show ?thesis using Cons by (auto simp: sorted_Cons insort_is_Cons)
4650 qed (auto simp: sorted_Cons insort_is_Cons)
4653 lemma insort_remove1:
4654 assumes "a \<in> set xs" and "sorted xs"
4655 shows "insort a (remove1 a xs) = xs"
4656 proof (rule insort_key_remove1)
4657 from `a \<in> set xs` show "a \<in> set xs" .
4658 from `sorted xs` show "sorted (map (\<lambda>x. x) xs)" by simp
4659 from `a \<in> set xs` have "a \<in> set (filter (op = a) xs)" by auto
4660 then have "set (filter (op = a) xs) \<noteq> {}" by auto
4661 then have "filter (op = a) xs \<noteq> []" by (auto simp only: set_empty)
4662 then have "length (filter (op = a) xs) > 0" by simp
4663 then obtain n where n: "Suc n = length (filter (op = a) xs)"
4664 by (cases "length (filter (op = a) xs)") simp_all
4665 moreover have "replicate (Suc n) a = a # replicate n a"
4667 ultimately show "hd (filter (op = a) xs) = a" by (simp add: replicate_length_filter)
4670 lemma sorted_remdups[simp]:
4671 "sorted l \<Longrightarrow> sorted (remdups l)"
4672 by (induct l) (auto simp: sorted_Cons)
4674 lemma sorted_remdups_adj[simp]:
4675 "sorted xs \<Longrightarrow> sorted (remdups_adj xs)"
4676 by (induct xs rule: remdups_adj.induct, simp_all split: split_if_asm add: sorted_Cons)
4678 lemma sorted_distinct_set_unique:
4679 assumes "sorted xs" "distinct xs" "sorted ys" "distinct ys" "set xs = set ys"
4682 from assms have 1: "length xs = length ys" by (auto dest!: distinct_card)
4683 from assms show ?thesis
4684 proof(induct rule:list_induct2[OF 1])
4685 case 1 show ?case by simp
4687 case 2 thus ?case by (simp add:sorted_Cons)
4688 (metis Diff_insert_absorb antisym insertE insert_iff)
4692 lemma map_sorted_distinct_set_unique:
4693 assumes "inj_on f (set xs \<union> set ys)"
4694 assumes "sorted (map f xs)" "distinct (map f xs)"
4695 "sorted (map f ys)" "distinct (map f ys)"
4696 assumes "set xs = set ys"
4699 from assms have "map f xs = map f ys"
4700 by (simp add: sorted_distinct_set_unique)
4701 with `inj_on f (set xs \<union> set ys)` show "xs = ys"
4702 by (blast intro: map_inj_on)
4705 lemma finite_sorted_distinct_unique:
4706 shows "finite A \<Longrightarrow> EX! xs. set xs = A & sorted xs & distinct xs"
4707 apply(drule finite_distinct_list)
4709 apply(rule_tac a="sort xs" in ex1I)
4710 apply (auto simp: sorted_distinct_set_unique)
4715 shows sorted_take: "sorted (take n xs)"
4716 and sorted_drop: "sorted (drop n xs)"
4718 from assms have "sorted (take n xs @ drop n xs)" by simp
4719 then show "sorted (take n xs)" and "sorted (drop n xs)"
4720 unfolding sorted_append by simp_all
4723 lemma sorted_dropWhile: "sorted xs \<Longrightarrow> sorted (dropWhile P xs)"
4724 by (auto dest: sorted_drop simp add: dropWhile_eq_drop)
4726 lemma sorted_takeWhile: "sorted xs \<Longrightarrow> sorted (takeWhile P xs)"
4727 by (subst takeWhile_eq_take) (auto dest: sorted_take)
4729 lemma sorted_filter:
4730 "sorted (map f xs) \<Longrightarrow> sorted (map f (filter P xs))"
4731 by (induct xs) (simp_all add: sorted_Cons)
4733 lemma foldr_max_sorted:
4734 assumes "sorted (rev xs)"
4735 shows "foldr max xs y = (if xs = [] then y else max (xs ! 0) y)"
4739 then have "sorted (rev xs)" using sorted_append by auto
4740 with Cons show ?case
4741 by (cases xs) (auto simp add: sorted_append max_def)
4744 lemma filter_equals_takeWhile_sorted_rev:
4745 assumes sorted: "sorted (rev (map f xs))"
4746 shows "[x \<leftarrow> xs. t < f x] = takeWhile (\<lambda> x. t < f x) xs"
4747 (is "filter ?P xs = ?tW")
4748 proof (rule takeWhile_eq_filter[symmetric])
4749 let "?dW" = "dropWhile ?P xs"
4750 fix x assume "x \<in> set ?dW"
4751 then obtain i where i: "i < length ?dW" and nth_i: "x = ?dW ! i"
4752 unfolding in_set_conv_nth by auto
4753 hence "length ?tW + i < length (?tW @ ?dW)"
4754 unfolding length_append by simp
4755 hence i': "length (map f ?tW) + i < length (map f xs)" by simp
4756 have "(map f ?tW @ map f ?dW) ! (length (map f ?tW) + i) \<le>
4757 (map f ?tW @ map f ?dW) ! (length (map f ?tW) + 0)"
4758 using sorted_rev_nth_mono[OF sorted _ i', of "length ?tW"]
4759 unfolding map_append[symmetric] by simp
4760 hence "f x \<le> f (?dW ! 0)"
4761 unfolding nth_append_length_plus nth_i
4762 using i preorder_class.le_less_trans[OF le0 i] by simp
4763 also have "... \<le> t"
4764 using hd_dropWhile[of "?P" xs] le0[THEN preorder_class.le_less_trans, OF i]
4765 using hd_conv_nth[of "?dW"] by simp
4766 finally show "\<not> t < f x" by simp
4769 lemma insort_insert_key_triv:
4770 "f x \<in> f ` set xs \<Longrightarrow> insort_insert_key f x xs = xs"
4771 by (simp add: insort_insert_key_def)
4773 lemma insort_insert_triv:
4774 "x \<in> set xs \<Longrightarrow> insort_insert x xs = xs"
4775 using insort_insert_key_triv [of "\<lambda>x. x"] by simp
4777 lemma insort_insert_insort_key:
4778 "f x \<notin> f ` set xs \<Longrightarrow> insort_insert_key f x xs = insort_key f x xs"
4779 by (simp add: insort_insert_key_def)
4781 lemma insort_insert_insort:
4782 "x \<notin> set xs \<Longrightarrow> insort_insert x xs = insort x xs"
4783 using insort_insert_insort_key [of "\<lambda>x. x"] by simp
4785 lemma set_insort_insert:
4786 "set (insort_insert x xs) = insert x (set xs)"
4787 by (auto simp add: insort_insert_key_def set_insort)
4789 lemma distinct_insort_insert:
4790 assumes "distinct xs"
4791 shows "distinct (insort_insert_key f x xs)"
4792 using assms by (induct xs) (auto simp add: insort_insert_key_def set_insort)
4794 lemma sorted_insort_insert_key:
4795 assumes "sorted (map f xs)"
4796 shows "sorted (map f (insort_insert_key f x xs))"
4797 using assms by (simp add: insort_insert_key_def sorted_insort_key)
4799 lemma sorted_insort_insert:
4801 shows "sorted (insort_insert x xs)"
4802 using assms sorted_insort_insert_key [of "\<lambda>x. x"] by simp
4804 lemma filter_insort_triv:
4805 "\<not> P x \<Longrightarrow> filter P (insort_key f x xs) = filter P xs"
4806 by (induct xs) simp_all
4808 lemma filter_insort:
4809 "sorted (map f xs) \<Longrightarrow> P x \<Longrightarrow> filter P (insort_key f x xs) = insort_key f x (filter P xs)"
4810 using assms by (induct xs)
4811 (auto simp add: sorted_Cons, subst insort_is_Cons, auto)
4814 "filter P (sort_key f xs) = sort_key f (filter P xs)"
4815 by (induct xs) (simp_all add: filter_insort_triv filter_insort)
4817 lemma sorted_map_same:
4818 "sorted (map f [x\<leftarrow>xs. f x = g xs])"
4819 proof (induct xs arbitrary: g)
4820 case Nil then show ?case by simp
4823 then have "sorted (map f [y\<leftarrow>xs . f y = (\<lambda>xs. f x) xs])" .
4824 moreover from Cons have "sorted (map f [y\<leftarrow>xs . f y = (g \<circ> Cons x) xs])" .
4825 ultimately show ?case by (simp_all add: sorted_Cons)
4829 "sorted [x\<leftarrow>xs. x = g xs]"
4830 using sorted_map_same [of "\<lambda>x. x"] by simp
4832 lemma remove1_insort [simp]:
4833 "remove1 x (insort x xs) = xs"
4834 by (induct xs) simp_all
4838 lemma sorted_upt[simp]: "sorted[i..<j]"
4839 by (induct j) (simp_all add:sorted_append)
4841 lemma sorted_upto[simp]: "sorted[i..j]"
4842 apply(induct i j rule:upto.induct)
4843 apply(subst upto.simps)
4844 apply(simp add:sorted_Cons)
4847 lemma sorted_find_Min:
4849 assumes "\<exists>x \<in> set xs. P x"
4850 shows "List.find P xs = Some (Min {x\<in>set xs. P x})"
4851 using assms proof (induct xs rule: sorted.induct)
4852 case Nil then show ?case by simp
4854 case (Cons xs x) show ?case proof (cases "P x")
4855 case True with Cons show ?thesis by (auto intro: Min_eqI [symmetric])
4857 case False then have "{y. (y = x \<or> y \<in> set xs) \<and> P y} = {y \<in> set xs. P y}"
4859 with Cons False show ?thesis by simp_all
4864 subsubsection {* @{const transpose} on sorted lists *}
4866 lemma sorted_transpose[simp]:
4867 shows "sorted (rev (map length (transpose xs)))"
4868 by (auto simp: sorted_equals_nth_mono rev_nth nth_transpose
4869 length_filter_conv_card intro: card_mono)
4871 lemma transpose_max_length:
4872 "foldr (\<lambda>xs. max (length xs)) (transpose xs) 0 = length [x \<leftarrow> xs. x \<noteq> []]"
4874 proof (cases "transpose xs = []")
4876 have "?L = foldr max (map length (transpose xs)) 0"
4877 by (simp add: foldr_map comp_def)
4878 also have "... = length (transpose xs ! 0)"
4879 using False sorted_transpose by (simp add: foldr_max_sorted)
4880 finally show ?thesis
4881 using False by (simp add: nth_transpose)
4884 hence "[x \<leftarrow> xs. x \<noteq> []] = []"
4885 by (auto intro!: filter_False simp: transpose_empty)
4886 thus ?thesis by (simp add: transpose_empty True)
4889 lemma length_transpose_sorted:
4890 fixes xs :: "'a list list"
4891 assumes sorted: "sorted (rev (map length xs))"
4892 shows "length (transpose xs) = (if xs = [] then 0 else length (xs ! 0))"
4893 proof (cases "xs = []")
4896 using foldr_max_sorted[OF sorted] False
4897 unfolding length_transpose foldr_map comp_def
4901 lemma nth_nth_transpose_sorted[simp]:
4902 fixes xs :: "'a list list"
4903 assumes sorted: "sorted (rev (map length xs))"
4904 and i: "i < length (transpose xs)"
4905 and j: "j < length [ys \<leftarrow> xs. i < length ys]"
4906 shows "transpose xs ! i ! j = xs ! j ! i"
4907 using j filter_equals_takeWhile_sorted_rev[OF sorted, of i]
4908 nth_transpose[OF i] nth_map[OF j]
4909 by (simp add: takeWhile_nth)
4911 lemma transpose_column_length:
4912 fixes xs :: "'a list list"
4913 assumes sorted: "sorted (rev (map length xs))" and "i < length xs"
4914 shows "length (filter (\<lambda>ys. i < length ys) (transpose xs)) = length (xs ! i)"
4916 have "xs \<noteq> []" using `i < length xs` by auto
4917 note filter_equals_takeWhile_sorted_rev[OF sorted, simp]
4918 { fix j assume "j \<le> i"
4919 note sorted_rev_nth_mono[OF sorted, of j i, simplified, OF this `i < length xs`]
4920 } note sortedE = this[consumes 1]
4922 have "{j. j < length (transpose xs) \<and> i < length (transpose xs ! j)}
4923 = {..< length (xs ! i)}"
4926 assume "j < length (transpose xs)" and "i < length (transpose xs ! j)"
4927 with this(2) nth_transpose[OF this(1)]
4928 have "i < length (takeWhile (\<lambda>ys. j < length ys) xs)" by simp
4929 from nth_mem[OF this] takeWhile_nth[OF this]
4930 show "j < length (xs ! i)" by (auto dest: set_takeWhileD)
4932 fix j assume "j < length (xs ! i)"
4933 thus "j < length (transpose xs)"
4934 using foldr_max_sorted[OF sorted] `xs \<noteq> []` sortedE[OF le0]
4935 by (auto simp: length_transpose comp_def foldr_map)
4937 have "Suc i \<le> length (takeWhile (\<lambda>ys. j < length ys) xs)"
4938 using `i < length xs` `j < length (xs ! i)` less_Suc_eq_le
4939 by (auto intro!: length_takeWhile_less_P_nth dest!: sortedE)
4940 with nth_transpose[OF `j < length (transpose xs)`]
4941 show "i < length (transpose xs ! j)" by simp
4943 thus ?thesis by (simp add: length_filter_conv_card)
4946 lemma transpose_column:
4947 fixes xs :: "'a list list"
4948 assumes sorted: "sorted (rev (map length xs))" and "i < length xs"
4949 shows "map (\<lambda>ys. ys ! i) (filter (\<lambda>ys. i < length ys) (transpose xs))
4950 = xs ! i" (is "?R = _")
4951 proof (rule nth_equalityI, safe)
4952 show length: "length ?R = length (xs ! i)"
4953 using transpose_column_length[OF assms] by simp
4955 fix j assume j: "j < length ?R"
4956 note * = less_le_trans[OF this, unfolded length_map, OF length_filter_le]
4957 from j have j_less: "j < length (xs ! i)" using length by simp
4958 have i_less_tW: "Suc i \<le> length (takeWhile (\<lambda>ys. Suc j \<le> length ys) xs)"
4959 proof (rule length_takeWhile_less_P_nth)
4960 show "Suc i \<le> length xs" using `i < length xs` by simp
4961 fix k assume "k < Suc i"
4962 hence "k \<le> i" by auto
4963 with sorted_rev_nth_mono[OF sorted this] `i < length xs`
4964 have "length (xs ! i) \<le> length (xs ! k)" by simp
4965 thus "Suc j \<le> length (xs ! k)" using j_less by simp
4967 have i_less_filter: "i < length [ys\<leftarrow>xs . j < length ys]"
4968 unfolding filter_equals_takeWhile_sorted_rev[OF sorted, of j]
4969 using i_less_tW by (simp_all add: Suc_le_eq)
4970 from j show "?R ! j = xs ! i ! j"
4971 unfolding filter_equals_takeWhile_sorted_rev[OF sorted_transpose, of i]
4972 by (simp add: takeWhile_nth nth_nth_transpose_sorted[OF sorted * i_less_filter])
4975 lemma transpose_transpose:
4976 fixes xs :: "'a list list"
4977 assumes sorted: "sorted (rev (map length xs))"
4978 shows "transpose (transpose xs) = takeWhile (\<lambda>x. x \<noteq> []) xs" (is "?L = ?R")
4980 have len: "length ?L = length ?R"
4981 unfolding length_transpose transpose_max_length
4982 using filter_equals_takeWhile_sorted_rev[OF sorted, of 0]
4985 { fix i assume "i < length ?R"
4986 with less_le_trans[OF _ length_takeWhile_le[of _ xs]]
4987 have "i < length xs" by simp
4990 by (rule nth_equalityI)
4991 (simp_all add: len nth_transpose transpose_column[OF sorted] * takeWhile_nth)
4994 theorem transpose_rectangle:
4995 assumes "xs = [] \<Longrightarrow> n = 0"
4996 assumes rect: "\<And> i. i < length xs \<Longrightarrow> length (xs ! i) = n"
4997 shows "transpose xs = map (\<lambda> i. map (\<lambda> j. xs ! j ! i) [0..<length xs]) [0..<n]"
4998 (is "?trans = ?map")
4999 proof (rule nth_equalityI)
5000 have "sorted (rev (map length xs))"
5001 by (auto simp: rev_nth rect intro!: sorted_nth_monoI)
5002 from foldr_max_sorted[OF this] assms
5003 show len: "length ?trans = length ?map"
5004 by (simp_all add: length_transpose foldr_map comp_def)
5006 { fix i assume "i < n" hence "[ys\<leftarrow>xs . i < length ys] = xs"
5007 using rect by (auto simp: in_set_conv_nth intro!: filter_True) }
5008 ultimately show "\<forall>i < length ?trans. ?trans ! i = ?map ! i"
5009 by (auto simp: nth_transpose intro: nth_equalityI)
5013 subsubsection {* @{text sorted_list_of_set} *}
5015 text{* This function maps (finite) linearly ordered sets to sorted
5016 lists. Warning: in most cases it is not a good idea to convert from
5017 sets to lists but one should convert in the other direction (via
5020 subsubsection {* @{text sorted_list_of_set} *}
5022 text{* This function maps (finite) linearly ordered sets to sorted
5023 lists. Warning: in most cases it is not a good idea to convert from
5024 sets to lists but one should convert in the other direction (via
5030 definition sorted_list_of_set :: "'a set \<Rightarrow> 'a list" where
5031 "sorted_list_of_set = folding.F insort []"
5033 sublocale sorted_list_of_set!: folding insort Nil
5035 "folding.F insort [] = sorted_list_of_set"
5037 interpret comp_fun_commute insort by (fact comp_fun_commute_insort)
5038 show "folding insort" by default (fact comp_fun_commute)
5039 show "folding.F insort [] = sorted_list_of_set" by (simp only: sorted_list_of_set_def)
5042 lemma sorted_list_of_set_empty:
5043 "sorted_list_of_set {} = []"
5044 by (fact sorted_list_of_set.empty)
5046 lemma sorted_list_of_set_insert [simp]:
5047 "finite A \<Longrightarrow> sorted_list_of_set (insert x A) = insort x (sorted_list_of_set (A - {x}))"
5048 by (fact sorted_list_of_set.insert_remove)
5050 lemma sorted_list_of_set_eq_Nil_iff [simp]:
5051 "finite A \<Longrightarrow> sorted_list_of_set A = [] \<longleftrightarrow> A = {}"
5052 by (auto simp: sorted_list_of_set.remove)
5054 lemma sorted_list_of_set [simp]:
5055 "finite A \<Longrightarrow> set (sorted_list_of_set A) = A \<and> sorted (sorted_list_of_set A)
5056 \<and> distinct (sorted_list_of_set A)"
5057 by (induct A rule: finite_induct) (simp_all add: set_insort sorted_insort distinct_insort)
5059 lemma distinct_sorted_list_of_set:
5060 "distinct (sorted_list_of_set A)"
5061 using sorted_list_of_set by (cases "finite A") auto
5063 lemma sorted_list_of_set_sort_remdups [code]:
5064 "sorted_list_of_set (set xs) = sort (remdups xs)"
5066 interpret comp_fun_commute insort by (fact comp_fun_commute_insort)
5067 show ?thesis by (simp add: sorted_list_of_set.eq_fold sort_conv_fold fold_set_fold_remdups)
5070 lemma sorted_list_of_set_remove:
5072 shows "sorted_list_of_set (A - {x}) = remove1 x (sorted_list_of_set A)"
5073 proof (cases "x \<in> A")
5074 case False with assms have "x \<notin> set (sorted_list_of_set A)" by simp
5075 with False show ?thesis by (simp add: remove1_idem)
5077 case True then obtain B where A: "A = insert x B" by (rule Set.set_insert)
5078 with assms show ?thesis by simp
5083 lemma sorted_list_of_set_range [simp]:
5084 "sorted_list_of_set {m..<n} = [m..<n]"
5085 by (rule sorted_distinct_set_unique) simp_all
5088 subsubsection {* @{text lists}: the list-forming operator over sets *}
5091 lists :: "'a set => 'a list set"
5094 Nil [intro!, simp]: "[]: lists A"
5095 | Cons [intro!, simp]: "[| a: A; l: lists A|] ==> a#l : lists A"
5097 inductive_cases listsE [elim!]: "x#l : lists A"
5098 inductive_cases listspE [elim!]: "listsp A (x # l)"
5100 inductive_simps listsp_simps[code]:
5104 lemma listsp_mono [mono]: "A \<le> B ==> listsp A \<le> listsp B"
5105 by (rule predicate1I, erule listsp.induct, blast+)
5107 lemmas lists_mono = listsp_mono [to_set]
5110 assumes l: "listsp A l" shows "listsp B l ==> listsp (inf A B) l" using l
5113 lemmas lists_IntI = listsp_infI [to_set]
5115 lemma listsp_inf_eq [simp]: "listsp (inf A B) = inf (listsp A) (listsp B)"
5116 proof (rule mono_inf [where f=listsp, THEN order_antisym])
5117 show "mono listsp" by (simp add: mono_def listsp_mono)
5118 show "inf (listsp A) (listsp B) \<le> listsp (inf A B)" by (blast intro!: listsp_infI)
5121 lemmas listsp_conj_eq [simp] = listsp_inf_eq [simplified inf_fun_def inf_bool_def]
5123 lemmas lists_Int_eq [simp] = listsp_inf_eq [to_set]
5125 lemma Cons_in_lists_iff[simp]: "x#xs : lists A \<longleftrightarrow> x:A \<and> xs : lists A"
5128 lemma append_in_listsp_conv [iff]:
5129 "(listsp A (xs @ ys)) = (listsp A xs \<and> listsp A ys)"
5132 lemmas append_in_lists_conv [iff] = append_in_listsp_conv [to_set]
5134 lemma in_listsp_conv_set: "(listsp A xs) = (\<forall>x \<in> set xs. A x)"
5135 -- {* eliminate @{text listsp} in favour of @{text set} *}
5138 lemmas in_lists_conv_set [code_unfold] = in_listsp_conv_set [to_set]
5140 lemma in_listspD [dest!]: "listsp A xs ==> \<forall>x\<in>set xs. A x"
5141 by (rule in_listsp_conv_set [THEN iffD1])
5143 lemmas in_listsD [dest!] = in_listspD [to_set]
5145 lemma in_listspI [intro!]: "\<forall>x\<in>set xs. A x ==> listsp A xs"
5146 by (rule in_listsp_conv_set [THEN iffD2])
5148 lemmas in_listsI [intro!] = in_listspI [to_set]
5150 lemma lists_eq_set: "lists A = {xs. set xs <= A}"
5153 lemma lists_empty [simp]: "lists {} = {[]}"
5156 lemma lists_UNIV [simp]: "lists UNIV = UNIV"
5159 lemma lists_image: "lists (f`A) = map f ` lists A"
5161 { fix xs have "\<forall>x\<in>set xs. x \<in> f ` A \<Longrightarrow> xs \<in> map f ` lists A"
5162 by (induct xs) (auto simp del: map.simps simp add: map.simps[symmetric] intro!: imageI) }
5163 then show ?thesis by auto
5166 subsubsection {* Inductive definition for membership *}
5168 inductive ListMem :: "'a \<Rightarrow> 'a list \<Rightarrow> bool"
5170 elem: "ListMem x (x # xs)"
5171 | insert: "ListMem x xs \<Longrightarrow> ListMem x (y # xs)"
5173 lemma ListMem_iff: "(ListMem x xs) = (x \<in> set xs)"
5175 apply (induct set: ListMem)
5178 apply (auto intro: ListMem.intros)
5182 subsubsection {* Lists as Cartesian products *}
5184 text{*@{text"set_Cons A Xs"}: the set of lists with head drawn from
5185 @{term A} and tail drawn from @{term Xs}.*}
5187 definition set_Cons :: "'a set \<Rightarrow> 'a list set \<Rightarrow> 'a list set" where
5188 "set_Cons A XS = {z. \<exists>x xs. z = x # xs \<and> x \<in> A \<and> xs \<in> XS}"
5190 lemma set_Cons_sing_Nil [simp]: "set_Cons A {[]} = (%x. [x])`A"
5191 by (auto simp add: set_Cons_def)
5193 text{*Yields the set of lists, all of the same length as the argument and
5194 with elements drawn from the corresponding element of the argument.*}
5196 primrec listset :: "'a set list \<Rightarrow> 'a list set" where
5197 "listset [] = {[]}" |
5198 "listset (A # As) = set_Cons A (listset As)"
5201 subsection {* Relations on Lists *}
5203 subsubsection {* Length Lexicographic Ordering *}
5205 text{*These orderings preserve well-foundedness: shorter lists
5206 precede longer lists. These ordering are not used in dictionaries.*}
5208 primrec -- {*The lexicographic ordering for lists of the specified length*}
5209 lexn :: "('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> ('a list \<times> 'a list) set" where
5212 (map_pair (%(x, xs). x#xs) (%(x, xs). x#xs) ` (r <*lex*> lexn r n)) Int
5213 {(xs, ys). length xs = Suc n \<and> length ys = Suc n}"
5215 definition lex :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
5216 "lex r = (\<Union>n. lexn r n)" -- {*Holds only between lists of the same length*}
5218 definition lenlex :: "('a \<times> 'a) set => ('a list \<times> 'a list) set" where
5219 "lenlex r = inv_image (less_than <*lex*> lex r) (\<lambda>xs. (length xs, xs))"
5220 -- {*Compares lists by their length and then lexicographically*}
5222 lemma wf_lexn: "wf r ==> wf (lexn r n)"
5223 apply (induct n, simp, simp)
5224 apply(rule wf_subset)
5225 prefer 2 apply (rule Int_lower1)
5226 apply(rule wf_map_pair_image)
5227 prefer 2 apply (rule inj_onI, auto)
5231 "(xs, ys) : lexn r n ==> length xs = n \<and> length ys = n"
5232 by (induct n arbitrary: xs ys) auto
5234 lemma wf_lex [intro!]: "wf r ==> wf (lex r)"
5235 apply (unfold lex_def)
5237 apply (blast intro: wf_lexn, clarify)
5238 apply (rename_tac m n)
5239 apply (subgoal_tac "m \<noteq> n")
5240 prefer 2 apply blast
5241 apply (blast dest: lexn_length not_sym)
5246 {(xs,ys). length xs = n \<and> length ys = n \<and>
5247 (\<exists>xys x y xs' ys'. xs= xys @ x#xs' \<and> ys= xys @ y # ys' \<and> (x, y):r)}"
5248 apply (induct n, simp)
5249 apply (simp add: image_Collect lex_prod_def, safe, blast)
5250 apply (rule_tac x = "ab # xys" in exI, simp)
5251 apply (case_tac xys, simp_all, blast)
5256 {(xs,ys). length xs = length ys \<and>
5257 (\<exists>xys x y xs' ys'. xs = xys @ x # xs' \<and> ys = xys @ y # ys' \<and> (x, y):r)}"
5258 by (force simp add: lex_def lexn_conv)
5260 lemma wf_lenlex [intro!]: "wf r ==> wf (lenlex r)"
5261 by (unfold lenlex_def) blast
5264 "lenlex r = {(xs,ys). length xs < length ys |
5265 length xs = length ys \<and> (xs, ys) : lex r}"
5266 by (simp add: lenlex_def Id_on_def lex_prod_def inv_image_def)
5268 lemma Nil_notin_lex [iff]: "([], ys) \<notin> lex r"
5269 by (simp add: lex_conv)
5271 lemma Nil2_notin_lex [iff]: "(xs, []) \<notin> lex r"
5272 by (simp add:lex_conv)
5274 lemma Cons_in_lex [simp]:
5275 "((x # xs, y # ys) : lex r) =
5276 ((x, y) : r \<and> length xs = length ys | x = y \<and> (xs, ys) : lex r)"
5277 apply (simp add: lex_conv)
5279 prefer 2 apply (blast intro: Cons_eq_appendI, clarify)
5280 apply (case_tac xys, simp, simp)
5285 subsubsection {* Lexicographic Ordering *}
5287 text {* Classical lexicographic ordering on lists, ie. "a" < "ab" < "b".
5288 This ordering does \emph{not} preserve well-foundedness.
5289 Author: N. Voelker, March 2005. *}
5291 definition lexord :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
5292 "lexord r = {(x,y). \<exists> a v. y = x @ a # v \<or>
5293 (\<exists> u a b v w. (a,b) \<in> r \<and> x = u @ (a # v) \<and> y = u @ (b # w))}"
5295 lemma lexord_Nil_left[simp]: "([],y) \<in> lexord r = (\<exists> a x. y = a # x)"
5296 by (unfold lexord_def, induct_tac y, auto)
5298 lemma lexord_Nil_right[simp]: "(x,[]) \<notin> lexord r"
5299 by (unfold lexord_def, induct_tac x, auto)
5301 lemma lexord_cons_cons[simp]:
5302 "((a # x, b # y) \<in> lexord r) = ((a,b)\<in> r | (a = b & (x,y)\<in> lexord r))"
5303 apply (unfold lexord_def, safe, simp_all)
5304 apply (case_tac u, simp, simp)
5305 apply (case_tac u, simp, clarsimp, blast, blast, clarsimp)
5306 apply (erule_tac x="b # u" in allE)
5309 lemmas lexord_simps = lexord_Nil_left lexord_Nil_right lexord_cons_cons
5311 lemma lexord_append_rightI: "\<exists> b z. y = b # z \<Longrightarrow> (x, x @ y) \<in> lexord r"
5312 by (induct_tac x, auto)
5314 lemma lexord_append_left_rightI:
5315 "(a,b) \<in> r \<Longrightarrow> (u @ a # x, u @ b # y) \<in> lexord r"
5316 by (induct_tac u, auto)
5318 lemma lexord_append_leftI: " (u,v) \<in> lexord r \<Longrightarrow> (x @ u, x @ v) \<in> lexord r"
5321 lemma lexord_append_leftD:
5322 "\<lbrakk> (x @ u, x @ v) \<in> lexord r; (! a. (a,a) \<notin> r) \<rbrakk> \<Longrightarrow> (u,v) \<in> lexord r"
5323 by (erule rev_mp, induct_tac x, auto)
5325 lemma lexord_take_index_conv:
5326 "((x,y) : lexord r) =
5327 ((length x < length y \<and> take (length x) y = x) \<or>
5328 (\<exists>i. i < min(length x)(length y) & take i x = take i y & (x!i,y!i) \<in> r))"
5329 apply (unfold lexord_def Let_def, clarsimp)
5330 apply (rule_tac f = "(% a b. a \<or> b)" in arg_cong2)
5332 apply (rule_tac x="hd (drop (length x) y)" in exI)
5333 apply (rule_tac x="tl (drop (length x) y)" in exI)
5334 apply (erule subst, simp add: min_def)
5335 apply (rule_tac x ="length u" in exI, simp)
5336 apply (rule_tac x ="take i x" in exI)
5337 apply (rule_tac x ="x ! i" in exI)
5338 apply (rule_tac x ="y ! i" in exI, safe)
5339 apply (rule_tac x="drop (Suc i) x" in exI)
5340 apply (drule sym, simp add: drop_Suc_conv_tl)
5341 apply (rule_tac x="drop (Suc i) y" in exI)
5342 by (simp add: drop_Suc_conv_tl)
5344 -- {* lexord is extension of partial ordering List.lex *}
5345 lemma lexord_lex: "(x,y) \<in> lex r = ((x,y) \<in> lexord r \<and> length x = length y)"
5346 apply (rule_tac x = y in spec)
5347 apply (induct_tac x, clarsimp)
5348 by (clarify, case_tac x, simp, force)
5350 lemma lexord_irreflexive: "ALL x. (x,x) \<notin> r \<Longrightarrow> (xs,xs) \<notin> lexord r"
5353 text{* By Ren\'e Thiemann: *}
5354 lemma lexord_partial_trans:
5355 "(\<And>x y z. x \<in> set xs \<Longrightarrow> (x,y) \<in> r \<Longrightarrow> (y,z) \<in> r \<Longrightarrow> (x,z) \<in> r)
5356 \<Longrightarrow> (xs,ys) \<in> lexord r \<Longrightarrow> (ys,zs) \<in> lexord r \<Longrightarrow> (xs,zs) \<in> lexord r"
5357 proof (induct xs arbitrary: ys zs)
5359 from Nil(3) show ?case unfolding lexord_def by (cases zs, auto)
5361 case (Cons x xs yys zzs)
5362 from Cons(3) obtain y ys where yys: "yys = y # ys" unfolding lexord_def
5363 by (cases yys, auto)
5364 note Cons = Cons[unfolded yys]
5365 from Cons(3) have one: "(x,y) \<in> r \<or> x = y \<and> (xs,ys) \<in> lexord r" by auto
5366 from Cons(4) obtain z zs where zzs: "zzs = z # zs" unfolding lexord_def
5367 by (cases zzs, auto)
5368 note Cons = Cons[unfolded zzs]
5369 from Cons(4) have two: "(y,z) \<in> r \<or> y = z \<and> (ys,zs) \<in> lexord r" by auto
5371 assume "(xs,ys) \<in> lexord r" and "(ys,zs) \<in> lexord r"
5372 from Cons(1)[OF _ this] Cons(2)
5373 have "(xs,zs) \<in> lexord r" by auto
5376 assume "(x,y) \<in> r" and "(y,z) \<in> r"
5377 from Cons(2)[OF _ this] have "(x,z) \<in> r" by auto
5379 from one two ind1 ind2
5380 have "(x,z) \<in> r \<or> x = z \<and> (xs,zs) \<in> lexord r" by blast
5381 thus ?case unfolding zzs by auto
5385 "\<lbrakk> (x, y) \<in> lexord r; (y, z) \<in> lexord r; trans r \<rbrakk> \<Longrightarrow> (x, z) \<in> lexord r"
5386 by(auto simp: trans_def intro:lexord_partial_trans)
5388 lemma lexord_transI: "trans r \<Longrightarrow> trans (lexord r)"
5389 by (rule transI, drule lexord_trans, blast)
5391 lemma lexord_linear: "(! a b. (a,b)\<in> r | a = b | (b,a) \<in> r) \<Longrightarrow> (x,y) : lexord r | x = y | (y,x) : lexord r"
5392 apply (rule_tac x = y in spec)
5393 apply (induct_tac x, rule allI)
5394 apply (case_tac x, simp, simp)
5395 apply (rule allI, case_tac x, simp, simp)
5399 Predicate version of lexicographic order integrated with Isabelle's order type classes.
5400 Author: Andreas Lochbihler
5405 inductive lexordp :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
5407 Nil: "lexordp [] (y # ys)"
5408 | Cons: "x < y \<Longrightarrow> lexordp (x # xs) (y # ys)"
5410 "\<lbrakk> \<not> x < y; \<not> y < x; lexordp xs ys \<rbrakk> \<Longrightarrow> lexordp (x # xs) (y # ys)"
5412 lemma lexordp_simps [simp]:
5413 "lexordp [] ys = (ys \<noteq> [])"
5414 "lexordp xs [] = False"
5415 "lexordp (x # xs) (y # ys) \<longleftrightarrow> x < y \<or> \<not> y < x \<and> lexordp xs ys"
5416 by(subst lexordp.simps, fastforce simp add: neq_Nil_conv)+
5418 inductive lexordp_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
5419 Nil: "lexordp_eq [] ys"
5420 | Cons: "x < y \<Longrightarrow> lexordp_eq (x # xs) (y # ys)"
5421 | Cons_eq: "\<lbrakk> \<not> x < y; \<not> y < x; lexordp_eq xs ys \<rbrakk> \<Longrightarrow> lexordp_eq (x # xs) (y # ys)"
5423 lemma lexordp_eq_simps [simp]:
5424 "lexordp_eq [] ys = True"
5425 "lexordp_eq xs [] \<longleftrightarrow> xs = []"
5426 "lexordp_eq (x # xs) [] = False"
5427 "lexordp_eq (x # xs) (y # ys) \<longleftrightarrow> x < y \<or> \<not> y < x \<and> lexordp_eq xs ys"
5428 by(subst lexordp_eq.simps, fastforce)+
5430 lemma lexordp_append_rightI: "ys \<noteq> Nil \<Longrightarrow> lexordp xs (xs @ ys)"
5431 by(induct xs)(auto simp add: neq_Nil_conv)
5433 lemma lexordp_append_left_rightI: "x < y \<Longrightarrow> lexordp (us @ x # xs) (us @ y # ys)"
5436 lemma lexordp_eq_refl: "lexordp_eq xs xs"
5437 by(induct xs) simp_all
5439 lemma lexordp_append_leftI: "lexordp us vs \<Longrightarrow> lexordp (xs @ us) (xs @ vs)"
5442 lemma lexordp_append_leftD: "\<lbrakk> lexordp (xs @ us) (xs @ vs); \<forall>a. \<not> a < a \<rbrakk> \<Longrightarrow> lexordp us vs"
5445 lemma lexordp_irreflexive:
5446 assumes irrefl: "\<forall>x. \<not> x < x"
5447 shows "\<not> lexordp xs xs"
5449 assume "lexordp xs xs"
5450 thus False by(induct xs ys\<equiv>xs)(simp_all add: irrefl)
5453 lemma lexordp_into_lexordp_eq:
5454 assumes "lexordp xs ys"
5455 shows "lexordp_eq xs ys"
5456 using assms by induct simp_all
5460 declare ord.lexordp_simps [simp, code]
5461 declare ord.lexordp_eq_simps [code, simp]
5463 lemma lexord_code [code, code_unfold]: "lexordp = ord.lexordp less"
5464 unfolding lexordp_def ord.lexordp_def ..
5468 lemma lexordp_antisym:
5469 assumes "lexordp xs ys" "lexordp ys xs"
5471 using assms by induct auto
5473 lemma lexordp_irreflexive': "\<not> lexordp xs xs"
5474 by(rule lexordp_irreflexive) simp
5478 context linorder begin
5480 lemma lexordp_cases [consumes 1, case_names Nil Cons Cons_eq, cases pred: lexordp]:
5481 assumes "lexordp xs ys"
5482 obtains (Nil) y ys' where "xs = []" "ys = y # ys'"
5483 | (Cons) x xs' y ys' where "xs = x # xs'" "ys = y # ys'" "x < y"
5484 | (Cons_eq) x xs' ys' where "xs = x # xs'" "ys = x # ys'" "lexordp xs' ys'"
5485 using assms by cases (fastforce simp add: not_less_iff_gr_or_eq)+
5487 lemma lexordp_induct [consumes 1, case_names Nil Cons Cons_eq, induct pred: lexordp]:
5488 assumes major: "lexordp xs ys"
5489 and Nil: "\<And>y ys. P [] (y # ys)"
5490 and Cons: "\<And>x xs y ys. x < y \<Longrightarrow> P (x # xs) (y # ys)"
5491 and Cons_eq: "\<And>x xs ys. \<lbrakk> lexordp xs ys; P xs ys \<rbrakk> \<Longrightarrow> P (x # xs) (x # ys)"
5493 using major by induct (simp_all add: Nil Cons not_less_iff_gr_or_eq Cons_eq)
5496 "lexordp xs ys \<longleftrightarrow> (\<exists>x vs. ys = xs @ x # vs) \<or> (\<exists>us a b vs ws. a < b \<and> xs = us @ a # vs \<and> ys = us @ b # ws)"
5499 assume ?lhs thus ?rhs
5501 case Cons_eq thus ?case by simp (metis append.simps(2))
5502 qed(fastforce intro: disjI2 del: disjCI intro: exI[where x="[]"])+
5504 assume ?rhs thus ?lhs
5505 by(auto intro: lexordp_append_leftI[where us="[]", simplified] lexordp_append_leftI)
5508 lemma lexordp_conv_lexord:
5509 "lexordp xs ys \<longleftrightarrow> (xs, ys) \<in> lexord {(x, y). x < y}"
5510 by(simp add: lexordp_iff lexord_def)
5512 lemma lexordp_eq_antisym:
5513 assumes "lexordp_eq xs ys" "lexordp_eq ys xs"
5515 using assms by induct simp_all
5517 lemma lexordp_eq_trans:
5518 assumes "lexordp_eq xs ys" and "lexordp_eq ys zs"
5519 shows "lexordp_eq xs zs"
5521 apply(induct arbitrary: zs)
5522 apply(case_tac [2-3] zs)
5526 lemma lexordp_trans:
5527 assumes "lexordp xs ys" "lexordp ys zs"
5528 shows "lexordp xs zs"
5530 apply(induct arbitrary: zs)
5531 apply(case_tac [2-3] zs)
5535 lemma lexordp_linear: "lexordp xs ys \<or> xs = ys \<or> lexordp ys xs"
5536 proof(induct xs arbitrary: ys)
5537 case Nil thus ?case by(cases ys) simp_all
5539 case Cons thus ?case by(cases ys) auto
5542 lemma lexordp_conv_lexordp_eq: "lexordp xs ys \<longleftrightarrow> lexordp_eq xs ys \<and> \<not> lexordp_eq ys xs"
5543 (is "?lhs \<longleftrightarrow> ?rhs")
5546 moreover hence "\<not> lexordp_eq ys xs" by induct simp_all
5547 ultimately show ?rhs by(simp add: lexordp_into_lexordp_eq)
5550 hence "lexordp_eq xs ys" "\<not> lexordp_eq ys xs" by simp_all
5551 thus ?lhs by induct simp_all
5554 lemma lexordp_eq_conv_lexord: "lexordp_eq xs ys \<longleftrightarrow> xs = ys \<or> lexordp xs ys"
5555 by(auto simp add: lexordp_conv_lexordp_eq lexordp_eq_refl dest: lexordp_eq_antisym)
5557 lemma lexordp_eq_linear: "lexordp_eq xs ys \<or> lexordp_eq ys xs"
5558 apply(induct xs arbitrary: ys)
5559 apply(case_tac [!] ys)
5563 lemma lexordp_linorder: "class.linorder lexordp_eq lexordp"
5564 by unfold_locales(auto simp add: lexordp_conv_lexordp_eq lexordp_eq_refl lexordp_eq_antisym intro: lexordp_eq_trans del: disjCI intro: lexordp_eq_linear)
5568 subsubsection {* Lexicographic combination of measure functions *}
5570 text {* These are useful for termination proofs *}
5572 definition "measures fs = inv_image (lex less_than) (%a. map (%f. f a) fs)"
5574 lemma wf_measures[simp]: "wf (measures fs)"
5575 unfolding measures_def
5578 lemma in_measures[simp]:
5579 "(x, y) \<in> measures [] = False"
5580 "(x, y) \<in> measures (f # fs)
5581 = (f x < f y \<or> (f x = f y \<and> (x, y) \<in> measures fs))"
5582 unfolding measures_def
5585 lemma measures_less: "f x < f y ==> (x, y) \<in> measures (f#fs)"
5588 lemma measures_lesseq: "f x <= f y ==> (x, y) \<in> measures fs ==> (x, y) \<in> measures (f#fs)"
5592 subsubsection {* Lifting Relations to Lists: one element *}
5594 definition listrel1 :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
5595 "listrel1 r = {(xs,ys).
5596 \<exists>us z z' vs. xs = us @ z # vs \<and> (z,z') \<in> r \<and> ys = us @ z' # vs}"
5599 "\<lbrakk> (x, y) \<in> r; xs = us @ x # vs; ys = us @ y # vs \<rbrakk> \<Longrightarrow>
5600 (xs, ys) \<in> listrel1 r"
5601 unfolding listrel1_def by auto
5604 "\<lbrakk> (xs, ys) \<in> listrel1 r;
5605 !!x y us vs. \<lbrakk> (x, y) \<in> r; xs = us @ x # vs; ys = us @ y # vs \<rbrakk> \<Longrightarrow> P
5606 \<rbrakk> \<Longrightarrow> P"
5607 unfolding listrel1_def by auto
5609 lemma not_Nil_listrel1 [iff]: "([], xs) \<notin> listrel1 r"
5610 unfolding listrel1_def by blast
5612 lemma not_listrel1_Nil [iff]: "(xs, []) \<notin> listrel1 r"
5613 unfolding listrel1_def by blast
5615 lemma Cons_listrel1_Cons [iff]:
5616 "(x # xs, y # ys) \<in> listrel1 r \<longleftrightarrow>
5617 (x,y) \<in> r \<and> xs = ys \<or> x = y \<and> (xs, ys) \<in> listrel1 r"
5618 by (simp add: listrel1_def Cons_eq_append_conv) (blast)
5620 lemma listrel1I1: "(x,y) \<in> r \<Longrightarrow> (x # xs, y # xs) \<in> listrel1 r"
5621 by (metis Cons_listrel1_Cons)
5623 lemma listrel1I2: "(xs, ys) \<in> listrel1 r \<Longrightarrow> (x # xs, x # ys) \<in> listrel1 r"
5624 by (metis Cons_listrel1_Cons)
5626 lemma append_listrel1I:
5627 "(xs, ys) \<in> listrel1 r \<and> us = vs \<or> xs = ys \<and> (us, vs) \<in> listrel1 r
5628 \<Longrightarrow> (xs @ us, ys @ vs) \<in> listrel1 r"
5629 unfolding listrel1_def
5630 by auto (blast intro: append_eq_appendI)+
5632 lemma Cons_listrel1E1[elim!]:
5633 assumes "(x # xs, ys) \<in> listrel1 r"
5634 and "\<And>y. ys = y # xs \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> R"
5635 and "\<And>zs. ys = x # zs \<Longrightarrow> (xs, zs) \<in> listrel1 r \<Longrightarrow> R"
5637 using assms by (cases ys) blast+
5639 lemma Cons_listrel1E2[elim!]:
5640 assumes "(xs, y # ys) \<in> listrel1 r"
5641 and "\<And>x. xs = x # ys \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> R"
5642 and "\<And>zs. xs = y # zs \<Longrightarrow> (zs, ys) \<in> listrel1 r \<Longrightarrow> R"
5644 using assms by (cases xs) blast+
5646 lemma snoc_listrel1_snoc_iff:
5647 "(xs @ [x], ys @ [y]) \<in> listrel1 r
5648 \<longleftrightarrow> (xs, ys) \<in> listrel1 r \<and> x = y \<or> xs = ys \<and> (x,y) \<in> r" (is "?L \<longleftrightarrow> ?R")
5651 by (fastforce simp: listrel1_def snoc_eq_iff_butlast butlast_append)
5653 assume ?R then show ?L unfolding listrel1_def by force
5656 lemma listrel1_eq_len: "(xs,ys) \<in> listrel1 r \<Longrightarrow> length xs = length ys"
5657 unfolding listrel1_def by auto
5659 lemma listrel1_mono:
5660 "r \<subseteq> s \<Longrightarrow> listrel1 r \<subseteq> listrel1 s"
5661 unfolding listrel1_def by blast
5664 lemma listrel1_converse: "listrel1 (r^-1) = (listrel1 r)^-1"
5665 unfolding listrel1_def by blast
5667 lemma in_listrel1_converse:
5668 "(x,y) : listrel1 (r^-1) \<longleftrightarrow> (x,y) : (listrel1 r)^-1"
5669 unfolding listrel1_def by blast
5671 lemma listrel1_iff_update:
5672 "(xs,ys) \<in> (listrel1 r)
5673 \<longleftrightarrow> (\<exists>y n. (xs ! n, y) \<in> r \<and> n < length xs \<and> ys = xs[n:=y])" (is "?L \<longleftrightarrow> ?R")
5676 then obtain x y u v where "xs = u @ x # v" "ys = u @ y # v" "(x,y) \<in> r"
5677 unfolding listrel1_def by auto
5678 then have "ys = xs[length u := y]" and "length u < length xs"
5679 and "(xs ! length u, y) \<in> r" by auto
5680 then show "?R" by auto
5683 then obtain x y n where "(xs!n, y) \<in> r" "n < size xs" "ys = xs[n:=y]" "x = xs!n"
5685 then obtain u v where "xs = u @ x # v" and "ys = u @ y # v" and "(x, y) \<in> r"
5686 by (auto intro: upd_conv_take_nth_drop id_take_nth_drop)
5687 then show "?L" by (auto simp: listrel1_def)
5691 text{* Accessible part and wellfoundedness: *}
5693 lemma Cons_acc_listrel1I [intro!]:
5694 "x \<in> Wellfounded.acc r \<Longrightarrow> xs \<in> Wellfounded.acc (listrel1 r) \<Longrightarrow> (x # xs) \<in> Wellfounded.acc (listrel1 r)"
5695 apply (induct arbitrary: xs set: Wellfounded.acc)
5696 apply (erule thin_rl)
5697 apply (erule acc_induct)
5702 lemma lists_accD: "xs \<in> lists (Wellfounded.acc r) \<Longrightarrow> xs \<in> Wellfounded.acc (listrel1 r)"
5703 apply (induct set: lists)
5707 apply (fast dest: acc_downward)
5710 lemma lists_accI: "xs \<in> Wellfounded.acc (listrel1 r) \<Longrightarrow> xs \<in> lists (Wellfounded.acc r)"
5711 apply (induct set: Wellfounded.acc)
5714 apply (fastforce dest!: in_set_conv_decomp[THEN iffD1] simp: listrel1_def)
5717 lemma wf_listrel1_iff[simp]: "wf(listrel1 r) = wf r"
5718 by(metis wf_acc_iff in_lists_conv_set lists_accI lists_accD Cons_in_lists_iff)
5721 subsubsection {* Lifting Relations to Lists: all elements *}
5724 listrel :: "('a \<times> 'b) set \<Rightarrow> ('a list \<times> 'b list) set"
5725 for r :: "('a \<times> 'b) set"
5727 Nil: "([],[]) \<in> listrel r"
5728 | Cons: "[| (x,y) \<in> r; (xs,ys) \<in> listrel r |] ==> (x#xs, y#ys) \<in> listrel r"
5730 inductive_cases listrel_Nil1 [elim!]: "([],xs) \<in> listrel r"
5731 inductive_cases listrel_Nil2 [elim!]: "(xs,[]) \<in> listrel r"
5732 inductive_cases listrel_Cons1 [elim!]: "(y#ys,xs) \<in> listrel r"
5733 inductive_cases listrel_Cons2 [elim!]: "(xs,y#ys) \<in> listrel r"
5736 lemma listrel_eq_len: "(xs, ys) \<in> listrel r \<Longrightarrow> length xs = length ys"
5737 by(induct rule: listrel.induct) auto
5739 lemma listrel_iff_zip [code_unfold]: "(xs,ys) : listrel r \<longleftrightarrow>
5740 length xs = length ys & (\<forall>(x,y) \<in> set(zip xs ys). (x,y) \<in> r)" (is "?L \<longleftrightarrow> ?R")
5742 assume ?L thus ?R by induct (auto intro: listrel_eq_len)
5746 by (induct rule: list_induct2) (auto intro: listrel.intros)
5749 lemma listrel_iff_nth: "(xs,ys) : listrel r \<longleftrightarrow>
5750 length xs = length ys & (\<forall>n < length xs. (xs!n, ys!n) \<in> r)" (is "?L \<longleftrightarrow> ?R")
5751 by (auto simp add: all_set_conv_all_nth listrel_iff_zip)
5754 lemma listrel_mono: "r \<subseteq> s \<Longrightarrow> listrel r \<subseteq> listrel s"
5756 apply (erule listrel.induct)
5757 apply (blast intro: listrel.intros)+
5760 lemma listrel_subset: "r \<subseteq> A \<times> A \<Longrightarrow> listrel r \<subseteq> lists A \<times> lists A"
5762 apply (erule listrel.induct, auto)
5765 lemma listrel_refl_on: "refl_on A r \<Longrightarrow> refl_on (lists A) (listrel r)"
5766 apply (simp add: refl_on_def listrel_subset Ball_def)
5768 apply (induct_tac x)
5769 apply (auto intro: listrel.intros)
5772 lemma listrel_sym: "sym r \<Longrightarrow> sym (listrel r)"
5773 apply (auto simp add: sym_def)
5774 apply (erule listrel.induct)
5775 apply (blast intro: listrel.intros)+
5778 lemma listrel_trans: "trans r \<Longrightarrow> trans (listrel r)"
5779 apply (simp add: trans_def)
5782 apply (erule listrel.induct)
5783 apply (blast intro: listrel.intros)+
5786 theorem equiv_listrel: "equiv A r \<Longrightarrow> equiv (lists A) (listrel r)"
5787 by (simp add: equiv_def listrel_refl_on listrel_sym listrel_trans)
5789 lemma listrel_rtrancl_refl[iff]: "(xs,xs) : listrel(r^*)"
5790 using listrel_refl_on[of UNIV, OF refl_rtrancl]
5791 by(auto simp: refl_on_def)
5793 lemma listrel_rtrancl_trans:
5794 "\<lbrakk> (xs,ys) : listrel(r^*); (ys,zs) : listrel(r^*) \<rbrakk>
5795 \<Longrightarrow> (xs,zs) : listrel(r^*)"
5796 by (metis listrel_trans trans_def trans_rtrancl)
5799 lemma listrel_Nil [simp]: "listrel r `` {[]} = {[]}"
5800 by (blast intro: listrel.intros)
5803 "listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})"
5804 by (auto simp add: set_Cons_def intro: listrel.intros)
5806 text {* Relating @{term listrel1}, @{term listrel} and closures: *}
5808 lemma listrel1_rtrancl_subset_rtrancl_listrel1:
5809 "listrel1 (r^*) \<subseteq> (listrel1 r)^*"
5810 proof (rule subrelI)
5811 fix xs ys assume 1: "(xs,ys) \<in> listrel1 (r^*)"
5813 have "(x,y) : r^* \<Longrightarrow> (us @ x # vs, us @ y # vs) : (listrel1 r)^*"
5814 proof(induct rule: rtrancl.induct)
5815 case rtrancl_refl show ?case by simp
5817 case rtrancl_into_rtrancl thus ?case
5818 by (metis listrel1I rtrancl.rtrancl_into_rtrancl)
5820 thus "(xs,ys) \<in> (listrel1 r)^*" using 1 by(blast elim: listrel1E)
5823 lemma rtrancl_listrel1_eq_len: "(x,y) \<in> (listrel1 r)^* \<Longrightarrow> length x = length y"
5824 by (induct rule: rtrancl.induct) (auto intro: listrel1_eq_len)
5826 lemma rtrancl_listrel1_ConsI1:
5827 "(xs,ys) : (listrel1 r)^* \<Longrightarrow> (x#xs,x#ys) : (listrel1 r)^*"
5828 apply(induct rule: rtrancl.induct)
5830 by (metis listrel1I2 rtrancl.rtrancl_into_rtrancl)
5832 lemma rtrancl_listrel1_ConsI2:
5833 "(x,y) \<in> r^* \<Longrightarrow> (xs, ys) \<in> (listrel1 r)^*
5834 \<Longrightarrow> (x # xs, y # ys) \<in> (listrel1 r)^*"
5835 by (blast intro: rtrancl_trans rtrancl_listrel1_ConsI1
5836 subsetD[OF listrel1_rtrancl_subset_rtrancl_listrel1 listrel1I1])
5838 lemma listrel1_subset_listrel:
5839 "r \<subseteq> r' \<Longrightarrow> refl r' \<Longrightarrow> listrel1 r \<subseteq> listrel(r')"
5840 by(auto elim!: listrel1E simp add: listrel_iff_zip set_zip refl_on_def)
5842 lemma listrel_reflcl_if_listrel1:
5843 "(xs,ys) : listrel1 r \<Longrightarrow> (xs,ys) : listrel(r^*)"
5844 by(erule listrel1E)(auto simp add: listrel_iff_zip set_zip)
5846 lemma listrel_rtrancl_eq_rtrancl_listrel1: "listrel (r^*) = (listrel1 r)^*"
5848 { fix x y assume "(x,y) \<in> listrel (r^*)"
5849 then have "(x,y) \<in> (listrel1 r)^*"
5850 by induct (auto intro: rtrancl_listrel1_ConsI2) }
5851 then show "listrel (r^*) \<subseteq> (listrel1 r)^*"
5854 show "listrel (r^*) \<supseteq> (listrel1 r)^*"
5856 fix xs ys assume "(xs,ys) \<in> (listrel1 r)^*"
5857 then show "(xs,ys) \<in> listrel (r^*)"
5859 case base show ?case by(auto simp add: listrel_iff_zip set_zip)
5862 thus ?case by (metis listrel_reflcl_if_listrel1 listrel_rtrancl_trans)
5867 lemma rtrancl_listrel1_if_listrel:
5868 "(xs,ys) : listrel r \<Longrightarrow> (xs,ys) : (listrel1 r)^*"
5869 by(metis listrel_rtrancl_eq_rtrancl_listrel1 subsetD[OF listrel_mono] r_into_rtrancl subsetI)
5871 lemma listrel_subset_rtrancl_listrel1: "listrel r \<subseteq> (listrel1 r)^*"
5872 by(fast intro:rtrancl_listrel1_if_listrel)
5875 subsection {* Size function *}
5877 lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (list_size f)"
5878 by (rule is_measure_trivial)
5880 lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (option_size f)"
5881 by (rule is_measure_trivial)
5883 lemma list_size_estimation[termination_simp]:
5884 "x \<in> set xs \<Longrightarrow> y < f x \<Longrightarrow> y < list_size f xs"
5887 lemma list_size_estimation'[termination_simp]:
5888 "x \<in> set xs \<Longrightarrow> y \<le> f x \<Longrightarrow> y \<le> list_size f xs"
5891 lemma list_size_map[simp]: "list_size f (map g xs) = list_size (f o g) xs"
5894 lemma list_size_append[simp]: "list_size f (xs @ ys) = list_size f xs + list_size f ys"
5895 by (induct xs, auto)
5897 lemma list_size_pointwise[termination_simp]:
5898 "(\<And>x. x \<in> set xs \<Longrightarrow> f x \<le> g x) \<Longrightarrow> list_size f xs \<le> list_size g xs"
5899 by (induct xs) force+
5902 subsection {* Monad operation *}
5904 definition bind :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
5905 "bind xs f = concat (map f xs)"
5907 hide_const (open) bind
5909 lemma bind_simps [simp]:
5910 "List.bind [] f = []"
5911 "List.bind (x # xs) f = f x @ List.bind xs f"
5912 by (simp_all add: bind_def)
5915 subsection {* Transfer *}
5917 definition embed_list :: "nat list \<Rightarrow> int list" where
5918 "embed_list l = map int l"
5920 definition nat_list :: "int list \<Rightarrow> bool" where
5921 "nat_list l = nat_set (set l)"
5923 definition return_list :: "int list \<Rightarrow> nat list" where
5924 "return_list l = map nat l"
5926 lemma transfer_nat_int_list_return_embed: "nat_list l \<longrightarrow>
5927 embed_list (return_list l) = l"
5928 unfolding embed_list_def return_list_def nat_list_def nat_set_def
5933 lemma transfer_nat_int_list_functions:
5934 "l @ m = return_list (embed_list l @ embed_list m)"
5935 "[] = return_list []"
5936 unfolding return_list_def embed_list_def
5938 apply (induct l, auto)
5939 apply (induct m, auto)
5943 lemma transfer_nat_int_fold1: "fold f l x =
5944 fold (%x. f (nat x)) (embed_list l) x";
5948 subsection {* Code generation *}
5950 text{* Optional tail recursive version of @{const map}. Can avoid
5951 stack overflow in some target languages. *}
5953 fun map_tailrec_rev :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'b list" where
5954 "map_tailrec_rev f [] bs = bs" |
5955 "map_tailrec_rev f (a#as) bs = map_tailrec_rev f as (f a # bs)"
5957 lemma map_tailrec_rev:
5958 "map_tailrec_rev f as bs = rev(map f as) @ bs"
5959 by(induction as arbitrary: bs) simp_all
5961 definition map_tailrec :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
5962 "map_tailrec f as = rev (map_tailrec_rev f as [])"
5964 text{* Code equation: *}
5965 lemma map_eq_map_tailrec: "map = map_tailrec"
5966 by(simp add: fun_eq_iff map_tailrec_def map_tailrec_rev)
5969 subsubsection {* Counterparts for set-related operations *}
5971 definition member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
5972 [code_abbrev]: "member xs x \<longleftrightarrow> x \<in> set xs"
5975 Use @{text member} only for generating executable code. Otherwise use
5976 @{prop "x \<in> set xs"} instead --- it is much easier to reason about.
5979 lemma member_rec [code]:
5980 "member (x # xs) y \<longleftrightarrow> x = y \<or> member xs y"
5981 "member [] y \<longleftrightarrow> False"
5982 by (auto simp add: member_def)
5984 lemma in_set_member (* FIXME delete candidate *):
5985 "x \<in> set xs \<longleftrightarrow> member xs x"
5986 by (simp add: member_def)
5988 definition list_all :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
5989 list_all_iff [code_abbrev]: "list_all P xs \<longleftrightarrow> Ball (set xs) P"
5991 definition list_ex :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
5992 list_ex_iff [code_abbrev]: "list_ex P xs \<longleftrightarrow> Bex (set xs) P"
5994 definition list_ex1 :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
5995 list_ex1_iff [code_abbrev]: "list_ex1 P xs \<longleftrightarrow> (\<exists>! x. x \<in> set xs \<and> P x)"
5998 Usually you should prefer @{text "\<forall>x\<in>set xs"}, @{text "\<exists>x\<in>set xs"}
5999 and @{text "\<exists>!x. x\<in>set xs \<and> _"} over @{const list_all}, @{const list_ex}
6000 and @{const list_ex1} in specifications.
6003 lemma list_all_simps [simp, code]:
6004 "list_all P (x # xs) \<longleftrightarrow> P x \<and> list_all P xs"
6005 "list_all P [] \<longleftrightarrow> True"
6006 by (simp_all add: list_all_iff)
6008 lemma list_ex_simps [simp, code]:
6009 "list_ex P (x # xs) \<longleftrightarrow> P x \<or> list_ex P xs"
6010 "list_ex P [] \<longleftrightarrow> False"
6011 by (simp_all add: list_ex_iff)
6013 lemma list_ex1_simps [simp, code]:
6014 "list_ex1 P [] = False"
6015 "list_ex1 P (x # xs) = (if P x then list_all (\<lambda>y. \<not> P y \<or> x = y) xs else list_ex1 P xs)"
6016 by (auto simp add: list_ex1_iff list_all_iff)
6018 lemma Ball_set_list_all: (* FIXME delete candidate *)
6019 "Ball (set xs) P \<longleftrightarrow> list_all P xs"
6020 by (simp add: list_all_iff)
6022 lemma Bex_set_list_ex: (* FIXME delete candidate *)
6023 "Bex (set xs) P \<longleftrightarrow> list_ex P xs"
6024 by (simp add: list_ex_iff)
6026 lemma list_all_append [simp]:
6027 "list_all P (xs @ ys) \<longleftrightarrow> list_all P xs \<and> list_all P ys"
6028 by (auto simp add: list_all_iff)
6030 lemma list_ex_append [simp]:
6031 "list_ex P (xs @ ys) \<longleftrightarrow> list_ex P xs \<or> list_ex P ys"
6032 by (auto simp add: list_ex_iff)
6034 lemma list_all_rev [simp]:
6035 "list_all P (rev xs) \<longleftrightarrow> list_all P xs"
6036 by (simp add: list_all_iff)
6038 lemma list_ex_rev [simp]:
6039 "list_ex P (rev xs) \<longleftrightarrow> list_ex P xs"
6040 by (simp add: list_ex_iff)
6042 lemma list_all_length:
6043 "list_all P xs \<longleftrightarrow> (\<forall>n < length xs. P (xs ! n))"
6044 by (auto simp add: list_all_iff set_conv_nth)
6046 lemma list_ex_length:
6047 "list_ex P xs \<longleftrightarrow> (\<exists>n < length xs. P (xs ! n))"
6048 by (auto simp add: list_ex_iff set_conv_nth)
6050 lemma list_all_cong [fundef_cong]:
6051 "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> list_all f xs = list_all g ys"
6052 by (simp add: list_all_iff)
6054 lemma list_ex_cong [fundef_cong]:
6055 "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> list_ex f xs = list_ex g ys"
6056 by (simp add: list_ex_iff)
6058 definition can_select :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool" where
6059 [code_abbrev]: "can_select P A = (\<exists>!x\<in>A. P x)"
6061 lemma can_select_set_list_ex1 [code]:
6062 "can_select P (set A) = list_ex1 P A"
6063 by (simp add: list_ex1_iff can_select_def)
6066 text {* Executable checks for relations on sets *}
6068 definition listrel1p :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
6069 "listrel1p r xs ys = ((xs, ys) \<in> listrel1 {(x, y). r x y})"
6071 lemma [code_unfold]:
6072 "(xs, ys) \<in> listrel1 r = listrel1p (\<lambda>x y. (x, y) \<in> r) xs ys"
6073 unfolding listrel1p_def by auto
6076 "listrel1p r [] xs = False"
6077 "listrel1p r xs [] = False"
6078 "listrel1p r (x # xs) (y # ys) \<longleftrightarrow>
6079 r x y \<and> xs = ys \<or> x = y \<and> listrel1p r xs ys"
6080 by (simp add: listrel1p_def)+
6083 lexordp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
6084 "lexordp r xs ys = ((xs, ys) \<in> lexord {(x, y). r x y})"
6086 lemma [code_unfold]:
6087 "(xs, ys) \<in> lexord r = lexordp (\<lambda>x y. (x, y) \<in> r) xs ys"
6088 unfolding lexordp_def by auto
6091 "lexordp r xs [] = False"
6092 "lexordp r [] (y#ys) = True"
6093 "lexordp r (x # xs) (y # ys) = (r x y | (x = y & lexordp r xs ys))"
6094 unfolding lexordp_def by auto
6096 text {* Bounded quantification and summation over nats. *}
6098 lemma atMost_upto [code_unfold]:
6099 "{..n} = set [0..<Suc n]"
6102 lemma atLeast_upt [code_unfold]:
6103 "{..<n} = set [0..<n]"
6106 lemma greaterThanLessThan_upt [code_unfold]:
6107 "{n<..<m} = set [Suc n..<m]"
6110 lemmas atLeastLessThan_upt [code_unfold] = set_upt [symmetric]
6112 lemma greaterThanAtMost_upt [code_unfold]:
6113 "{n<..m} = set [Suc n..<Suc m]"
6116 lemma atLeastAtMost_upt [code_unfold]:
6117 "{n..m} = set [n..<Suc m]"
6120 lemma all_nat_less_eq [code_unfold]:
6121 "(\<forall>m<n\<Colon>nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..<n}. P m)"
6124 lemma ex_nat_less_eq [code_unfold]:
6125 "(\<exists>m<n\<Colon>nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..<n}. P m)"
6128 lemma all_nat_less [code_unfold]:
6129 "(\<forall>m\<le>n\<Colon>nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..n}. P m)"
6132 lemma ex_nat_less [code_unfold]:
6133 "(\<exists>m\<le>n\<Colon>nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..n}. P m)"
6136 lemma setsum_set_upt_conv_listsum_nat [code_unfold]:
6137 "setsum f (set [m..<n]) = listsum (map f [m..<n])"
6138 by (simp add: interv_listsum_conv_setsum_set_nat)
6140 text{* Bounded @{text LEAST} operator: *}
6142 definition "Bleast S P = (LEAST x. x \<in> S \<and> P x)"
6144 definition "abort_Bleast S P = (LEAST x. x \<in> S \<and> P x)"
6146 declare [[code abort: abort_Bleast]]
6148 lemma Bleast_code [code]:
6149 "Bleast (set xs) P = (case filter P (sort xs) of
6150 x#xs \<Rightarrow> x |
6151 [] \<Rightarrow> abort_Bleast (set xs) P)"
6152 proof (cases "filter P (sort xs)")
6153 case Nil thus ?thesis by (simp add: Bleast_def abort_Bleast_def)
6156 have "(LEAST x. x \<in> set xs \<and> P x) = x"
6157 proof (rule Least_equality)
6158 show "x \<in> set xs \<and> P x"
6159 by (metis Cons Cons_eq_filter_iff in_set_conv_decomp set_sort)
6161 fix y assume "y : set xs \<and> P y"
6162 hence "y : set (filter P xs)" by auto
6164 by (metis Cons eq_iff filter_sort set_ConsD set_sort sorted_Cons sorted_sort)
6166 thus ?thesis using Cons by (simp add: Bleast_def)
6169 declare Bleast_def[symmetric, code_unfold]
6171 text {* Summation over ints. *}
6173 lemma greaterThanLessThan_upto [code_unfold]:
6174 "{i<..<j::int} = set [i+1..j - 1]"
6177 lemma atLeastLessThan_upto [code_unfold]:
6178 "{i..<j::int} = set [i..j - 1]"
6181 lemma greaterThanAtMost_upto [code_unfold]:
6182 "{i<..j::int} = set [i+1..j]"
6185 lemmas atLeastAtMost_upto [code_unfold] = set_upto [symmetric]
6187 lemma setsum_set_upto_conv_listsum_int [code_unfold]:
6188 "setsum f (set [i..j::int]) = listsum (map f [i..j])"
6189 by (simp add: interv_listsum_conv_setsum_set_int)
6192 subsubsection {* Optimizing by rewriting *}
6194 definition null :: "'a list \<Rightarrow> bool" where
6195 [code_abbrev]: "null xs \<longleftrightarrow> xs = []"
6198 Efficient emptyness check is implemented by @{const null}.
6201 lemma null_rec [code]:
6202 "null (x # xs) \<longleftrightarrow> False"
6203 "null [] \<longleftrightarrow> True"
6204 by (simp_all add: null_def)
6206 lemma eq_Nil_null: (* FIXME delete candidate *)
6207 "xs = [] \<longleftrightarrow> null xs"
6208 by (simp add: null_def)
6210 lemma equal_Nil_null [code_unfold]:
6211 "HOL.equal xs [] \<longleftrightarrow> null xs"
6212 "HOL.equal [] = null"
6213 by (auto simp add: equal null_def)
6215 definition maps :: "('a \<Rightarrow> 'b list) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
6216 [code_abbrev]: "maps f xs = concat (map f xs)"
6218 definition map_filter :: "('a \<Rightarrow> 'b option) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
6219 [code_post]: "map_filter f xs = map (the \<circ> f) (filter (\<lambda>x. f x \<noteq> None) xs)"
6222 Operations @{const maps} and @{const map_filter} avoid
6223 intermediate lists on execution -- do not use for proving.
6226 lemma maps_simps [code]:
6227 "maps f (x # xs) = f x @ maps f xs"
6229 by (simp_all add: maps_def)
6231 lemma map_filter_simps [code]:
6232 "map_filter f (x # xs) = (case f x of None \<Rightarrow> map_filter f xs | Some y \<Rightarrow> y # map_filter f xs)"
6233 "map_filter f [] = []"
6234 by (simp_all add: map_filter_def split: option.split)
6236 lemma concat_map_maps: (* FIXME delete candidate *)
6237 "concat (map f xs) = maps f xs"
6238 by (simp add: maps_def)
6240 lemma map_filter_map_filter [code_unfold]:
6241 "map f (filter P xs) = map_filter (\<lambda>x. if P x then Some (f x) else None) xs"
6242 by (simp add: map_filter_def)
6244 text {* Optimized code for @{text"\<forall>i\<in>{a..b::int}"} and @{text"\<forall>n:{a..<b::nat}"}
6245 and similiarly for @{text"\<exists>"}. *}
6247 definition all_interval_nat :: "(nat \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
6248 "all_interval_nat P i j \<longleftrightarrow> (\<forall>n \<in> {i..<j}. P n)"
6251 "all_interval_nat P i j \<longleftrightarrow> i \<ge> j \<or> P i \<and> all_interval_nat P (Suc i) j"
6253 have *: "\<And>n. P i \<Longrightarrow> \<forall>n\<in>{Suc i..<j}. P n \<Longrightarrow> i \<le> n \<Longrightarrow> n < j \<Longrightarrow> P n"
6256 assume "P i" "\<forall>n\<in>{Suc i..<j}. P n" "i \<le> n" "n < j"
6257 then show "P n" by (cases "n = i") simp_all
6259 show ?thesis by (auto simp add: all_interval_nat_def intro: *)
6262 lemma list_all_iff_all_interval_nat [code_unfold]:
6263 "list_all P [i..<j] \<longleftrightarrow> all_interval_nat P i j"
6264 by (simp add: list_all_iff all_interval_nat_def)
6266 lemma list_ex_iff_not_all_inverval_nat [code_unfold]:
6267 "list_ex P [i..<j] \<longleftrightarrow> \<not> (all_interval_nat (Not \<circ> P) i j)"
6268 by (simp add: list_ex_iff all_interval_nat_def)
6270 definition all_interval_int :: "(int \<Rightarrow> bool) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool" where
6271 "all_interval_int P i j \<longleftrightarrow> (\<forall>k \<in> {i..j}. P k)"
6274 "all_interval_int P i j \<longleftrightarrow> i > j \<or> P i \<and> all_interval_int P (i + 1) j"
6276 have *: "\<And>k. P i \<Longrightarrow> \<forall>k\<in>{i+1..j}. P k \<Longrightarrow> i \<le> k \<Longrightarrow> k \<le> j \<Longrightarrow> P k"
6279 assume "P i" "\<forall>k\<in>{i+1..j}. P k" "i \<le> k" "k \<le> j"
6280 then show "P k" by (cases "k = i") simp_all
6282 show ?thesis by (auto simp add: all_interval_int_def intro: *)
6285 lemma list_all_iff_all_interval_int [code_unfold]:
6286 "list_all P [i..j] \<longleftrightarrow> all_interval_int P i j"
6287 by (simp add: list_all_iff all_interval_int_def)
6289 lemma list_ex_iff_not_all_inverval_int [code_unfold]:
6290 "list_ex P [i..j] \<longleftrightarrow> \<not> (all_interval_int (Not \<circ> P) i j)"
6291 by (simp add: list_ex_iff all_interval_int_def)
6293 text {* optimized code (tail-recursive) for @{term length} *}
6295 definition gen_length :: "nat \<Rightarrow> 'a list \<Rightarrow> nat"
6296 where "gen_length n xs = n + length xs"
6298 lemma gen_length_code [code]:
6299 "gen_length n [] = n"
6300 "gen_length n (x # xs) = gen_length (Suc n) xs"
6301 by(simp_all add: gen_length_def)
6303 declare list.size(3-4)[code del]
6305 lemma length_code [code]: "length = gen_length 0"
6306 by(simp add: gen_length_def fun_eq_iff)
6308 hide_const (open) member null maps map_filter all_interval_nat all_interval_int gen_length
6311 subsubsection {* Pretty lists *}
6314 (* Code generation for list literals. *)
6316 signature LIST_CODE =
6318 val implode_list: Code_Thingol.iterm -> Code_Thingol.iterm list option
6319 val default_list: int * string
6320 -> (Code_Printer.fixity -> Code_Thingol.iterm -> Pretty.T)
6321 -> Code_Printer.fixity -> Code_Thingol.iterm -> Code_Thingol.iterm -> Pretty.T
6322 val add_literal_list: string -> theory -> theory
6325 structure List_Code : LIST_CODE =
6328 open Basic_Code_Thingol;
6330 fun implode_list t =
6332 fun dest_cons (IConst { sym = Code_Symbol.Constant @{const_name Cons}, ... } `$ t1 `$ t2) = SOME (t1, t2)
6333 | dest_cons _ = NONE;
6334 val (ts, t') = Code_Thingol.unfoldr dest_cons t;
6336 of IConst { sym = Code_Symbol.Constant @{const_name Nil}, ... } => SOME ts
6340 fun default_list (target_fxy, target_cons) pr fxy t1 t2 =
6341 Code_Printer.brackify_infix (target_fxy, Code_Printer.R) fxy (
6342 pr (Code_Printer.INFX (target_fxy, Code_Printer.X)) t1,
6343 Code_Printer.str target_cons,
6344 pr (Code_Printer.INFX (target_fxy, Code_Printer.R)) t2
6347 fun add_literal_list target =
6349 fun pretty literals pr _ vars fxy [(t1, _), (t2, _)] =
6350 case Option.map (cons t1) (implode_list t2)
6352 Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts)
6354 default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;