fundamental treatment of undefined vs. universally partial replaces code_abort
5 header {* The datatype of finite lists *}
8 imports Presburger Code_Numeral Quotient Lifting_Set Lifting_Option Lifting_Product
13 | Cons 'a "'a list" (infixr "#" 65)
16 -- {* list Enumeration *}
17 "_list" :: "args => 'a list" ("[(_)]")
24 subsection {* Basic list processing functions *}
26 primrec hd :: "'a list \<Rightarrow> 'a" where
29 primrec tl :: "'a list \<Rightarrow> 'a list" where
33 primrec last :: "'a list \<Rightarrow> 'a" where
34 "last (x # xs) = (if xs = [] then x else last xs)"
36 primrec butlast :: "'a list \<Rightarrow> 'a list" where
38 "butlast (x # xs) = (if xs = [] then [] else x # butlast xs)"
40 primrec set :: "'a list \<Rightarrow> 'a set" where
42 "set (x # xs) = insert x (set xs)"
44 definition coset :: "'a list \<Rightarrow> 'a set" where
45 [simp]: "coset xs = - set xs"
47 primrec map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
49 "map f (x # xs) = f x # map f xs"
51 primrec append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "@" 65) where
52 append_Nil: "[] @ ys = ys" |
53 append_Cons: "(x#xs) @ ys = x # xs @ ys"
55 primrec rev :: "'a list \<Rightarrow> 'a list" where
57 "rev (x # xs) = rev xs @ [x]"
59 primrec filter:: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
61 "filter P (x # xs) = (if P x then x # filter P xs else filter P xs)"
64 -- {* Special syntax for filter *}
65 "_filter" :: "[pttrn, 'a list, bool] => 'a list" ("(1[_<-_./ _])")
68 "[x<-xs . P]"== "CONST filter (%x. P) xs"
71 "_filter" :: "[pttrn, 'a list, bool] => 'a list"("(1[_\<leftarrow>_ ./ _])")
73 "_filter" :: "[pttrn, 'a list, bool] => 'a list"("(1[_\<leftarrow>_ ./ _])")
75 primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
76 fold_Nil: "fold f [] = id" |
77 fold_Cons: "fold f (x # xs) = fold f xs \<circ> f x"
79 primrec foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
80 foldr_Nil: "foldr f [] = id" |
81 foldr_Cons: "foldr f (x # xs) = f x \<circ> foldr f xs"
83 primrec foldl :: "('b \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b" where
84 foldl_Nil: "foldl f a [] = a" |
85 foldl_Cons: "foldl f a (x # xs) = foldl f (f a x) xs"
87 primrec concat:: "'a list list \<Rightarrow> 'a list" where
89 "concat (x # xs) = x @ concat xs"
91 definition (in monoid_add) listsum :: "'a list \<Rightarrow> 'a" where
92 "listsum xs = foldr plus xs 0"
94 primrec drop:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
95 drop_Nil: "drop n [] = []" |
96 drop_Cons: "drop n (x # xs) = (case n of 0 \<Rightarrow> x # xs | Suc m \<Rightarrow> drop m xs)"
97 -- {*Warning: simpset does not contain this definition, but separate
98 theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
100 primrec take:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
101 take_Nil:"take n [] = []" |
102 take_Cons: "take n (x # xs) = (case n of 0 \<Rightarrow> [] | Suc m \<Rightarrow> x # take m xs)"
103 -- {*Warning: simpset does not contain this definition, but separate
104 theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
106 primrec nth :: "'a list => nat => 'a" (infixl "!" 100) where
107 nth_Cons: "(x # xs) ! n = (case n of 0 \<Rightarrow> x | Suc k \<Rightarrow> xs ! k)"
108 -- {*Warning: simpset does not contain this definition, but separate
109 theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
111 primrec list_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
112 "list_update [] i v = []" |
113 "list_update (x # xs) i v =
114 (case i of 0 \<Rightarrow> v # xs | Suc j \<Rightarrow> x # list_update xs j v)"
116 nonterminal lupdbinds and lupdbind
119 "_lupdbind":: "['a, 'a] => lupdbind" ("(2_ :=/ _)")
120 "" :: "lupdbind => lupdbinds" ("_")
121 "_lupdbinds" :: "[lupdbind, lupdbinds] => lupdbinds" ("_,/ _")
122 "_LUpdate" :: "['a, lupdbinds] => 'a" ("_/[(_)]" [900,0] 900)
125 "_LUpdate xs (_lupdbinds b bs)" == "_LUpdate (_LUpdate xs b) bs"
126 "xs[i:=x]" == "CONST list_update xs i x"
128 primrec takeWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
129 "takeWhile P [] = []" |
130 "takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])"
132 primrec dropWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
133 "dropWhile P [] = []" |
134 "dropWhile P (x # xs) = (if P x then dropWhile P xs else x # xs)"
136 primrec zip :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
138 zip_Cons: "zip xs (y # ys) =
139 (case xs of [] => [] | z # zs => (z, y) # zip zs ys)"
140 -- {*Warning: simpset does not contain this definition, but separate
141 theorems for @{text "xs = []"} and @{text "xs = z # zs"} *}
143 primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
144 "product [] _ = []" |
145 "product (x#xs) ys = map (Pair x) ys @ product xs ys"
147 hide_const (open) product
149 primrec product_lists :: "'a list list \<Rightarrow> 'a list list" where
150 "product_lists [] = [[]]" |
151 "product_lists (xs # xss) = concat (map (\<lambda>x. map (Cons x) (product_lists xss)) xs)"
153 primrec upt :: "nat \<Rightarrow> nat \<Rightarrow> nat list" ("(1[_..</_'])") where
154 upt_0: "[i..<0] = []" |
155 upt_Suc: "[i..<(Suc j)] = (if i <= j then [i..<j] @ [j] else [])"
157 definition insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
158 "insert x xs = (if x \<in> set xs then xs else x # xs)"
160 hide_const (open) insert
161 hide_fact (open) insert_def
163 primrec find :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a option" where
165 "find P (x#xs) = (if P x then Some x else find P xs)"
167 hide_const (open) find
169 primrec those :: "'a option list \<Rightarrow> 'a list option"
171 "those [] = Some []" |
172 "those (x # xs) = (case x of
173 None \<Rightarrow> None
174 | Some y \<Rightarrow> Option.map (Cons y) (those xs))"
176 primrec remove1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
177 "remove1 x [] = []" |
178 "remove1 x (y # xs) = (if x = y then xs else y # remove1 x xs)"
180 primrec removeAll :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
181 "removeAll x [] = []" |
182 "removeAll x (y # xs) = (if x = y then removeAll x xs else y # removeAll x xs)"
184 primrec distinct :: "'a list \<Rightarrow> bool" where
185 "distinct [] \<longleftrightarrow> True" |
186 "distinct (x # xs) \<longleftrightarrow> x \<notin> set xs \<and> distinct xs"
188 primrec remdups :: "'a list \<Rightarrow> 'a list" where
190 "remdups (x # xs) = (if x \<in> set xs then remdups xs else x # remdups xs)"
192 fun remdups_adj :: "'a list \<Rightarrow> 'a list" where
193 "remdups_adj [] = []" |
194 "remdups_adj [x] = [x]" |
195 "remdups_adj (x # y # xs) = (if x = y then remdups_adj (x # xs) else x # remdups_adj (y # xs))"
197 primrec replicate :: "nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
198 replicate_0: "replicate 0 x = []" |
199 replicate_Suc: "replicate (Suc n) x = x # replicate n x"
202 Function @{text size} is overloaded for all datatypes. Users may
203 refer to the list version as @{text length}. *}
205 abbreviation length :: "'a list \<Rightarrow> nat" where
206 "length \<equiv> size"
208 definition enumerate :: "nat \<Rightarrow> 'a list \<Rightarrow> (nat \<times> 'a) list" where
209 enumerate_eq_zip: "enumerate n xs = zip [n..<n + length xs] xs"
211 primrec rotate1 :: "'a list \<Rightarrow> 'a list" where
213 "rotate1 (x # xs) = xs @ [x]"
215 definition rotate :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
216 "rotate n = rotate1 ^^ n"
218 definition list_all2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> bool" where
220 (length xs = length ys \<and> (\<forall>(x, y) \<in> set (zip xs ys). P x y))"
222 definition sublist :: "'a list => nat set => 'a list" where
223 "sublist xs A = map fst (filter (\<lambda>p. snd p \<in> A) (zip xs [0..<size xs]))"
225 primrec sublists :: "'a list \<Rightarrow> 'a list list" where
226 "sublists [] = [[]]" |
227 "sublists (x#xs) = (let xss = sublists xs in map (Cons x) xss @ xss)"
229 primrec n_lists :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list list" where
230 "n_lists 0 xs = [[]]" |
231 "n_lists (Suc n) xs = concat (map (\<lambda>ys. map (\<lambda>y. y # ys) xs) (n_lists n xs))"
233 hide_const (open) n_lists
235 fun splice :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
236 "splice [] ys = ys" |
237 "splice xs [] = xs" |
238 "splice (x#xs) (y#ys) = x # y # splice xs ys"
244 @{lemma "[a,b]@[c,d] = [a,b,c,d]" by simp}\\
245 @{lemma "length [a,b,c] = 3" by simp}\\
246 @{lemma "set [a,b,c] = {a,b,c}" by simp}\\
247 @{lemma "map f [a,b,c] = [f a, f b, f c]" by simp}\\
248 @{lemma "rev [a,b,c] = [c,b,a]" by simp}\\
249 @{lemma "hd [a,b,c,d] = a" by simp}\\
250 @{lemma "tl [a,b,c,d] = [b,c,d]" by simp}\\
251 @{lemma "last [a,b,c,d] = d" by simp}\\
252 @{lemma "butlast [a,b,c,d] = [a,b,c]" by simp}\\
253 @{lemma[source] "filter (\<lambda>n::nat. n<2) [0,2,1] = [0,1]" by simp}\\
254 @{lemma "concat [[a,b],[c,d,e],[],[f]] = [a,b,c,d,e,f]" by simp}\\
255 @{lemma "fold f [a,b,c] x = f c (f b (f a x))" by simp}\\
256 @{lemma "foldr f [a,b,c] x = f a (f b (f c x))" by simp}\\
257 @{lemma "foldl f x [a,b,c] = f (f (f x a) b) c" by simp}\\
258 @{lemma "zip [a,b,c] [x,y,z] = [(a,x),(b,y),(c,z)]" by simp}\\
259 @{lemma "zip [a,b] [x,y,z] = [(a,x),(b,y)]" by simp}\\
260 @{lemma "enumerate 3 [a,b,c] = [(3,a),(4,b),(5,c)]" by normalization}\\
261 @{lemma "List.product [a,b] [c,d] = [(a, c), (a, d), (b, c), (b, d)]" by simp}\\
262 @{lemma "product_lists [[a,b], [c], [d,e]] = [[a,c,d], [a,c,e], [b,c,d], [b,c,e]]" by simp}\\
263 @{lemma "splice [a,b,c] [x,y,z] = [a,x,b,y,c,z]" by simp}\\
264 @{lemma "splice [a,b,c,d] [x,y] = [a,x,b,y,c,d]" by simp}\\
265 @{lemma "take 2 [a,b,c,d] = [a,b]" by simp}\\
266 @{lemma "take 6 [a,b,c,d] = [a,b,c,d]" by simp}\\
267 @{lemma "drop 2 [a,b,c,d] = [c,d]" by simp}\\
268 @{lemma "drop 6 [a,b,c,d] = []" by simp}\\
269 @{lemma "takeWhile (%n::nat. n<3) [1,2,3,0] = [1,2]" by simp}\\
270 @{lemma "dropWhile (%n::nat. n<3) [1,2,3,0] = [3,0]" by simp}\\
271 @{lemma "distinct [2,0,1::nat]" by simp}\\
272 @{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\
273 @{lemma "remdups_adj [2,2,3,1,1::nat,2,1] = [2,3,1,2,1]" by simp}\\
274 @{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\
275 @{lemma "List.insert 3 [0::nat,1,2] = [3,0,1,2]" by (simp add: List.insert_def)}\\
276 @{lemma "List.find (%i::int. i>0) [0,0] = None" by simp}\\
277 @{lemma "List.find (%i::int. i>0) [0,1,0,2] = Some 1" by simp}\\
278 @{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\
279 @{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\
280 @{lemma "nth [a,b,c,d] 2 = c" by simp}\\
281 @{lemma "[a,b,c,d][2 := x] = [a,b,x,d]" by simp}\\
282 @{lemma "sublist [a,b,c,d,e] {0,2,3} = [a,c,d]" by (simp add:sublist_def)}\\
283 @{lemma "sublists [a,b] = [[a, b], [a], [b], []]" by simp}\\
284 @{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)}\\
285 @{lemma "rotate1 [a,b,c,d] = [b,c,d,a]" by simp}\\
286 @{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate_def eval_nat_numeral)}\\
287 @{lemma "replicate 4 a = [a,a,a,a]" by (simp add:eval_nat_numeral)}\\
288 @{lemma "[2..<5] = [2,3,4]" by (simp add:eval_nat_numeral)}\\
289 @{lemma "listsum [1,2,3::nat] = 6" by (simp add: listsum_def)}
291 \caption{Characteristic examples}
292 \label{fig:Characteristic}
294 Figure~\ref{fig:Characteristic} shows characteristic examples
295 that should give an intuitive understanding of the above functions.
298 text{* The following simple sort functions are intended for proofs,
299 not for efficient implementations. *}
304 inductive sorted :: "'a list \<Rightarrow> bool" where
305 Nil [iff]: "sorted []"
306 | Cons: "\<forall>y\<in>set xs. x \<le> y \<Longrightarrow> sorted xs \<Longrightarrow> sorted (x # xs)"
308 lemma sorted_single [iff]:
310 by (rule sorted.Cons) auto
313 "x \<le> y \<Longrightarrow> sorted (y # zs) \<Longrightarrow> sorted (x # y # zs)"
314 by (rule sorted.Cons) (cases "y # zs" rule: sorted.cases, auto)
316 lemma sorted_many_eq [simp, code]:
317 "sorted (x # y # zs) \<longleftrightarrow> x \<le> y \<and> sorted (y # zs)"
318 by (auto intro: sorted_many elim: sorted.cases)
321 "sorted [] \<longleftrightarrow> True"
322 "sorted [x] \<longleftrightarrow> True"
325 primrec insort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where
326 "insort_key f x [] = [x]" |
327 "insort_key f x (y#ys) =
328 (if f x \<le> f y then (x#y#ys) else y#(insort_key f x ys))"
330 definition sort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b list \<Rightarrow> 'b list" where
331 "sort_key f xs = foldr (insort_key f) xs []"
333 definition insort_insert_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where
334 "insort_insert_key f x xs =
335 (if f x \<in> f ` set xs then xs else insort_key f x xs)"
337 abbreviation "sort \<equiv> sort_key (\<lambda>x. x)"
338 abbreviation "insort \<equiv> insort_key (\<lambda>x. x)"
339 abbreviation "insort_insert \<equiv> insort_insert_key (\<lambda>x. x)"
344 subsubsection {* List comprehension *}
346 text{* Input syntax for Haskell-like list comprehension notation.
347 Typical example: @{text"[(x,y). x \<leftarrow> xs, y \<leftarrow> ys, x \<noteq> y]"},
348 the list of all pairs of distinct elements from @{text xs} and @{text ys}.
349 The syntax is as in Haskell, except that @{text"|"} becomes a dot
350 (like in Isabelle's set comprehension): @{text"[e. x \<leftarrow> xs, \<dots>]"} rather than
351 \verb![e| x <- xs, ...]!.
353 The qualifiers after the dot are
355 \item[generators] @{text"p \<leftarrow> xs"},
356 where @{text p} is a pattern and @{text xs} an expression of list type, or
357 \item[guards] @{text"b"}, where @{text b} is a boolean expression.
358 %\item[local bindings] @ {text"let x = e"}.
361 Just like in Haskell, list comprehension is just a shorthand. To avoid
362 misunderstandings, the translation into desugared form is not reversed
363 upon output. Note that the translation of @{text"[e. x \<leftarrow> xs]"} is
364 optmized to @{term"map (%x. e) xs"}.
366 It is easy to write short list comprehensions which stand for complex
367 expressions. During proofs, they may become unreadable (and
368 mangled). In such cases it can be advisable to introduce separate
369 definitions for the list comprehensions in question. *}
371 nonterminal lc_qual and lc_quals
374 "_listcompr" :: "'a \<Rightarrow> lc_qual \<Rightarrow> lc_quals \<Rightarrow> 'a list" ("[_ . __")
375 "_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ <- _")
376 "_lc_test" :: "bool \<Rightarrow> lc_qual" ("_")
377 (*"_lc_let" :: "letbinds => lc_qual" ("let _")*)
378 "_lc_end" :: "lc_quals" ("]")
379 "_lc_quals" :: "lc_qual \<Rightarrow> lc_quals \<Rightarrow> lc_quals" (", __")
380 "_lc_abs" :: "'a => 'b list => 'b list"
382 (* These are easier than ML code but cannot express the optimized
383 translation of [e. p<-xs]
385 "[e. p<-xs]" => "concat(map (_lc_abs p [e]) xs)"
386 "_listcompr e (_lc_gen p xs) (_lc_quals Q Qs)"
387 => "concat (map (_lc_abs p (_listcompr e Q Qs)) xs)"
388 "[e. P]" => "if P then [e] else []"
389 "_listcompr e (_lc_test P) (_lc_quals Q Qs)"
390 => "if P then (_listcompr e Q Qs) else []"
391 "_listcompr e (_lc_let b) (_lc_quals Q Qs)"
392 => "_Let b (_listcompr e Q Qs)"
396 "_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
398 "_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual" ("_ \<leftarrow> _")
402 val NilC = Syntax.const @{const_syntax Nil};
403 val ConsC = Syntax.const @{const_syntax Cons};
404 val mapC = Syntax.const @{const_syntax map};
405 val concatC = Syntax.const @{const_syntax concat};
406 val IfC = Syntax.const @{const_syntax If};
408 fun single x = ConsC $ x $ NilC;
410 fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)
412 (* FIXME proper name context!? *)
414 Free (singleton (Name.variant_list (fold Term.add_free_names [p, e] [])) "x", dummyT);
415 val e = if opti then single e else e;
416 val case1 = Syntax.const @{syntax_const "_case1"} $ p $ e;
418 Syntax.const @{syntax_const "_case1"} $
419 Syntax.const @{const_syntax dummy_pattern} $ NilC;
420 val cs = Syntax.const @{syntax_const "_case2"} $ case1 $ case2;
421 in Syntax_Trans.abs_tr [x, Case_Translation.case_tr false ctxt [x, cs]] end;
423 fun abs_tr ctxt p e opti =
424 (case Term_Position.strip_positions p of
427 val thy = Proof_Context.theory_of ctxt;
428 val s' = Proof_Context.intern_const ctxt s;
430 if Sign.declared_const thy s'
431 then (pat_tr ctxt p e opti, false)
432 else (Syntax_Trans.abs_tr [p, e], true)
434 | _ => (pat_tr ctxt p e opti, false));
436 fun lc_tr ctxt [e, Const (@{syntax_const "_lc_test"}, _) $ b, qs] =
440 Const (@{syntax_const "_lc_end"}, _) => single e
441 | Const (@{syntax_const "_lc_quals"}, _) $ q $ qs => lc_tr ctxt [e, q, qs]);
442 in IfC $ b $ res $ NilC end
444 [e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es,
445 Const(@{syntax_const "_lc_end"}, _)] =
446 (case abs_tr ctxt p e true of
447 (f, true) => mapC $ f $ es
448 | (f, false) => concatC $ (mapC $ f $ es))
450 [e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es,
451 Const (@{syntax_const "_lc_quals"}, _) $ q $ qs] =
452 let val e' = lc_tr ctxt [e, q, qs];
453 in concatC $ (mapC $ (fst (abs_tr ctxt p e' false)) $ es) end;
455 in [(@{syntax_const "_listcompr"}, lc_tr)] end
460 val read = Syntax.read_term @{context};
461 fun check s1 s2 = read s1 aconv read s2 orelse error ("Check failed: " ^ quote s1);
463 check "[(x,y,z). b]" "if b then [(x, y, z)] else []";
464 check "[(x,y,z). x\<leftarrow>xs]" "map (\<lambda>x. (x, y, z)) xs";
465 check "[e x y. x\<leftarrow>xs, y\<leftarrow>ys]" "concat (map (\<lambda>x. map (\<lambda>y. e x y) ys) xs)";
466 check "[(x,y,z). x<a, x>b]" "if x < a then if b < x then [(x, y, z)] else [] else []";
467 check "[(x,y,z). x\<leftarrow>xs, x>b]" "concat (map (\<lambda>x. if b < x then [(x, y, z)] else []) xs)";
468 check "[(x,y,z). x<a, x\<leftarrow>xs]" "if x < a then map (\<lambda>x. (x, y, z)) xs else []";
469 check "[(x,y). Cons True x \<leftarrow> xs]"
470 "concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | True # x \<Rightarrow> [(x, y)] | False # x \<Rightarrow> []) xs)";
471 check "[(x,y,z). Cons x [] \<leftarrow> xs]"
472 "concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | [x] \<Rightarrow> [(x, y, z)] | x # aa # lista \<Rightarrow> []) xs)";
473 check "[(x,y,z). x<a, x>b, x=d]"
474 "if x < a then if b < x then if x = d then [(x, y, z)] else [] else [] else []";
475 check "[(x,y,z). x<a, x>b, y\<leftarrow>ys]"
476 "if x < a then if b < x then map (\<lambda>y. (x, y, z)) ys else [] else []";
477 check "[(x,y,z). x<a, x\<leftarrow>xs,y>b]"
478 "if x < a then concat (map (\<lambda>x. if b < y then [(x, y, z)] else []) xs) else []";
479 check "[(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys]"
480 "if x < a then concat (map (\<lambda>x. map (\<lambda>y. (x, y, z)) ys) xs) else []";
481 check "[(x,y,z). x\<leftarrow>xs, x>b, y<a]"
482 "concat (map (\<lambda>x. if b < x then if y < a then [(x, y, z)] else [] else []) xs)";
483 check "[(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys]"
484 "concat (map (\<lambda>x. if b < x then map (\<lambda>y. (x, y, z)) ys else []) xs)";
485 check "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x]"
486 "concat (map (\<lambda>x. concat (map (\<lambda>y. if x < y then [(x, y, z)] else []) ys)) xs)";
487 check "[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs]"
488 "concat (map (\<lambda>x. concat (map (\<lambda>y. map (\<lambda>z. (x, y, z)) zs) ys)) xs)"
493 term "[(x,y). x\<leftarrow>xs, let xx = x+x, y\<leftarrow>ys, y \<noteq> xx]"
498 (* Simproc for rewriting list comprehensions applied to List.set to set
501 signature LIST_TO_SET_COMPREHENSION =
503 val simproc : Proof.context -> cterm -> thm option
506 structure List_to_Set_Comprehension : LIST_TO_SET_COMPREHENSION =
511 fun all_exists_conv cv ctxt ct =
512 (case Thm.term_of ct of
513 Const (@{const_name HOL.Ex}, _) $ Abs _ =>
514 Conv.arg_conv (Conv.abs_conv (all_exists_conv cv o #2) ctxt) ct
517 fun all_but_last_exists_conv cv ctxt ct =
518 (case Thm.term_of ct of
519 Const (@{const_name HOL.Ex}, _) $ Abs (_, _, Const (@{const_name HOL.Ex}, _) $ _) =>
520 Conv.arg_conv (Conv.abs_conv (all_but_last_exists_conv cv o #2) ctxt) ct
523 fun Collect_conv cv ctxt ct =
524 (case Thm.term_of ct of
525 Const (@{const_name Set.Collect}, _) $ Abs _ => Conv.arg_conv (Conv.abs_conv cv ctxt) ct
526 | _ => raise CTERM ("Collect_conv", [ct]))
528 fun rewr_conv' th = Conv.rewr_conv (mk_meta_eq th)
530 fun conjunct_assoc_conv ct =
532 (rewr_conv' @{thm conj_assoc} then_conv HOLogic.conj_conv Conv.all_conv conjunct_assoc_conv) ct
534 fun right_hand_set_comprehension_conv conv ctxt =
535 HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv
536 (Collect_conv (all_exists_conv conv o #2) ctxt))
539 (* term abstraction of list comprehension patterns *)
541 datatype termlets = If | Case of (typ * int)
543 fun simproc ctxt redex =
545 val set_Nil_I = @{thm trans} OF [@{thm set.simps(1)}, @{thm empty_def}]
546 val set_singleton = @{lemma "set [a] = {x. x = a}" by simp}
547 val inst_Collect_mem_eq = @{lemma "set A = {x. x : set A}" by simp}
548 val del_refl_eq = @{lemma "(t = t & P) == P" by simp}
549 fun mk_set T = Const (@{const_name List.set}, HOLogic.listT T --> HOLogic.mk_setT T)
550 fun dest_set (Const (@{const_name List.set}, _) $ xs) = xs
551 fun dest_singleton_list (Const (@{const_name List.Cons}, _)
552 $ t $ (Const (@{const_name List.Nil}, _))) = t
553 | dest_singleton_list t = raise TERM ("dest_singleton_list", [t])
554 (* We check that one case returns a singleton list and all other cases
555 return [], and return the index of the one singleton list case *)
556 fun possible_index_of_singleton_case cases =
558 fun check (i, case_t) s =
559 (case strip_abs_body case_t of
560 (Const (@{const_name List.Nil}, _)) => s
561 | _ => (case s of SOME NONE => SOME (SOME i) | _ => NONE))
563 fold_index check cases (SOME NONE) |> the_default NONE
565 (* returns (case_expr type index chosen_case constr_name) option *)
566 fun dest_case case_term =
568 val (case_const, args) = strip_comb case_term
570 (case try dest_Const case_const of
572 (case Ctr_Sugar.ctr_sugar_of_case ctxt c of
574 (case possible_index_of_singleton_case (fst (split_last args)) of
577 val constr_names = map (fst o dest_Const) ctrs
578 val (Ts, _) = strip_type T
579 val T' = List.last Ts
580 in SOME (List.last args, T', i, nth args i, nth constr_names i) end
585 (* returns condition continuing term option *)
586 fun dest_if (Const (@{const_name If}, _) $ cond $ then_t $ Const (@{const_name Nil}, _)) =
589 fun tac _ [] = rtac set_singleton 1 ORELSE rtac inst_Collect_mem_eq 1
590 | tac ctxt (If :: cont) =
591 Splitter.split_tac [@{thm split_if}] 1
592 THEN rtac @{thm conjI} 1
593 THEN rtac @{thm impI} 1
594 THEN Subgoal.FOCUS (fn {prems, context, ...} =>
595 CONVERSION (right_hand_set_comprehension_conv (K
596 (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_TrueI})) Conv.all_conv
598 rewr_conv' @{lemma "(True & P) = P" by simp})) context) 1) ctxt 1
600 THEN rtac @{thm impI} 1
601 THEN Subgoal.FOCUS (fn {prems, context, ...} =>
602 CONVERSION (right_hand_set_comprehension_conv (K
603 (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_FalseI})) Conv.all_conv
604 then_conv rewr_conv' @{lemma "(False & P) = False" by simp})) context) 1) ctxt 1
605 THEN rtac set_Nil_I 1
606 | tac ctxt (Case (T, i) :: cont) =
608 val SOME {injects, distincts, case_thms, split, ...} =
609 Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type T))
611 (* do case distinction *)
612 Splitter.split_tac [split] 1
613 THEN EVERY (map_index (fn (i', _) =>
614 (if i' < length case_thms - 1 then rtac @{thm conjI} 1 else all_tac)
615 THEN REPEAT_DETERM (rtac @{thm allI} 1)
616 THEN rtac @{thm impI} 1
618 (* continue recursively *)
619 Subgoal.FOCUS (fn {prems, context, ...} =>
620 CONVERSION (Thm.eta_conversion then_conv right_hand_set_comprehension_conv (K
622 (HOLogic.eq_conv Conv.all_conv (rewr_conv' (List.last prems)) then_conv
623 (Conv.try_conv (Conv.rewrs_conv (map mk_meta_eq injects))))
625 then_conv (Conv.try_conv (Conv.rewr_conv del_refl_eq))
626 then_conv conjunct_assoc_conv)) context
627 then_conv (HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt) =>
629 (all_but_last_exists_conv
631 @{lemma "(EX x. x = t & P x) = P t" by simp})) ctxt)) context)))) 1) ctxt 1
634 Subgoal.FOCUS (fn {prems, context, ...} =>
636 (right_hand_set_comprehension_conv (K
638 ((HOLogic.eq_conv Conv.all_conv
639 (rewr_conv' (List.last prems))) then_conv
640 (Conv.rewrs_conv (map (fn th => th RS @{thm Eq_FalseI}) distincts)))
641 Conv.all_conv then_conv
642 (rewr_conv' @{lemma "(False & P) = False" by simp}))) context then_conv
643 HOLogic.Trueprop_conv
644 (HOLogic.eq_conv Conv.all_conv
645 (Collect_conv (fn (_, ctxt) =>
649 @{lemma "(EX x. P) = P" by simp})) ctxt)) context))) 1) ctxt 1
650 THEN rtac set_Nil_I 1)) case_thms)
652 fun make_inner_eqs bound_vs Tis eqs t =
654 SOME (x, T, i, cont, constr_name) =>
656 val (vs, body) = strip_abs (Envir.eta_long (map snd bound_vs) cont)
657 val x' = incr_boundvars (length vs) x
658 val eqs' = map (incr_boundvars (length vs)) eqs
661 (Const (constr_name, map snd vs ---> T), map Bound (((length vs) - 1) downto 0))
662 val constr_eq = Const (@{const_name HOL.eq}, T --> T --> @{typ bool}) $ constr_t $ x'
664 make_inner_eqs (rev vs @ bound_vs) (Case (T, i) :: Tis) (constr_eq :: eqs') body
668 SOME (condition, cont) => make_inner_eqs bound_vs (If :: Tis) (condition :: eqs) cont
670 if eqs = [] then NONE (* no rewriting, nothing to be done *)
673 val Type (@{type_name List.list}, [rT]) = fastype_of1 (map snd bound_vs, t)
675 (case try dest_singleton_list t of
677 Const (@{const_name HOL.eq}, rT --> rT --> @{typ bool}) $
678 Bound (length bound_vs) $ t'
680 Const (@{const_name Set.member}, rT --> HOLogic.mk_setT rT --> @{typ bool}) $
681 Bound (length bound_vs) $ (mk_set rT $ t))
682 val reverse_bounds = curry subst_bounds
683 ((map Bound ((length bound_vs - 1) downto 0)) @ [Bound (length bound_vs)])
684 val eqs' = map reverse_bounds eqs
685 val pat_eq' = reverse_bounds pat_eq
687 fold (fn (_, T) => fn t => HOLogic.exists_const T $ absdummy T t)
688 (rev bound_vs) (fold (curry HOLogic.mk_conj) eqs' pat_eq')
689 val lhs = term_of redex
690 val rhs = HOLogic.mk_Collect ("x", rT, inner_t)
691 val rewrite_rule_t = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
694 ((Goal.prove ctxt [] [] rewrite_rule_t
695 (fn {context, ...} => tac context (rev Tis))) RS @{thm eq_reflection})
698 make_inner_eqs [] [] [] (dest_set (term_of redex))
704 simproc_setup list_to_set_comprehension ("set xs") = {* K List_to_Set_Comprehension.simproc *}
706 code_datatype set coset
708 hide_const (open) coset
711 subsubsection {* @{const Nil} and @{const Cons} *}
713 lemma not_Cons_self [simp]:
717 lemma not_Cons_self2 [simp]:
719 by (rule not_Cons_self [symmetric])
721 lemma neq_Nil_conv: "(xs \<noteq> []) = (\<exists>y ys. xs = y # ys)"
724 lemma tl_Nil: "tl xs = [] \<longleftrightarrow> xs = [] \<or> (EX x. xs = [x])"
727 lemma Nil_tl: "[] = tl xs \<longleftrightarrow> xs = [] \<or> (EX x. xs = [x])"
731 "(\<And>xs. \<forall>ys. length ys < length xs \<longrightarrow> P ys \<Longrightarrow> P xs) \<Longrightarrow> P xs"
732 by (fact measure_induct)
734 lemma list_nonempty_induct [consumes 1, case_names single cons]:
735 assumes "xs \<noteq> []"
736 assumes single: "\<And>x. P [x]"
737 assumes cons: "\<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (x # xs)"
739 using `xs \<noteq> []` proof (induct xs)
740 case Nil then show ?case by simp
746 with single show ?thesis by simp
751 from Cons show "xs \<noteq> []" by simp
752 with Cons.hyps show "P xs" .
757 lemma inj_split_Cons: "inj_on (\<lambda>(xs, n). n#xs) X"
758 by (auto intro!: inj_onI)
761 subsubsection {* @{const length} *}
764 Needs to come before @{text "@"} because of theorem @{text
765 append_eq_append_conv}.
768 lemma length_append [simp]: "length (xs @ ys) = length xs + length ys"
771 lemma length_map [simp]: "length (map f xs) = length xs"
774 lemma length_rev [simp]: "length (rev xs) = length xs"
777 lemma length_tl [simp]: "length (tl xs) = length xs - 1"
780 lemma length_0_conv [iff]: "(length xs = 0) = (xs = [])"
783 lemma length_greater_0_conv [iff]: "(0 < length xs) = (xs \<noteq> [])"
786 lemma length_pos_if_in_set: "x : set xs \<Longrightarrow> length xs > 0"
789 lemma length_Suc_conv:
790 "(length xs = Suc n) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
793 lemma Suc_length_conv:
794 "(Suc n = length xs) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
795 apply (induct xs, simp, simp)
799 lemma impossible_Cons: "length xs <= length ys ==> xs = x # ys = False"
802 lemma list_induct2 [consumes 1, case_names Nil Cons]:
803 "length xs = length ys \<Longrightarrow> P [] [] \<Longrightarrow>
804 (\<And>x xs y ys. length xs = length ys \<Longrightarrow> P xs ys \<Longrightarrow> P (x#xs) (y#ys))
805 \<Longrightarrow> P xs ys"
806 proof (induct xs arbitrary: ys)
807 case Nil then show ?case by simp
809 case (Cons x xs ys) then show ?case by (cases ys) simp_all
812 lemma list_induct3 [consumes 2, case_names Nil Cons]:
813 "length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> P [] [] [] \<Longrightarrow>
814 (\<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))
815 \<Longrightarrow> P xs ys zs"
816 proof (induct xs arbitrary: ys zs)
817 case Nil then show ?case by simp
819 case (Cons x xs ys zs) then show ?case by (cases ys, simp_all)
823 lemma list_induct4 [consumes 3, case_names Nil Cons]:
824 "length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow>
825 P [] [] [] [] \<Longrightarrow> (\<And>x xs y ys z zs w ws. length xs = length ys \<Longrightarrow>
826 length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow> P xs ys zs ws \<Longrightarrow>
827 P (x#xs) (y#ys) (z#zs) (w#ws)) \<Longrightarrow> P xs ys zs ws"
828 proof (induct xs arbitrary: ys zs ws)
829 case Nil then show ?case by simp
831 case (Cons x xs ys zs ws) then show ?case by ((cases ys, simp_all), (cases zs,simp_all)) (cases ws, simp_all)
836 \<And>x xs. P (x#xs) [];
837 \<And>y ys. P [] (y#ys);
838 \<And>x xs y ys. P xs ys \<Longrightarrow> P (x#xs) (y#ys) \<rbrakk>
839 \<Longrightarrow> P xs ys"
840 by (induct xs arbitrary: ys) (case_tac x, auto)+
842 lemma neq_if_length_neq: "length xs \<noteq> length ys \<Longrightarrow> (xs = ys) == False"
843 by (rule Eq_FalseI) auto
845 simproc_setup list_neq ("(xs::'a list) = ys") = {*
847 Reduces xs=ys to False if xs and ys cannot be of the same length.
848 This is the case if the atomic sublists of one are a submultiset
849 of those of the other list and there are fewer Cons's in one than the other.
854 fun len (Const(@{const_name Nil},_)) acc = acc
855 | len (Const(@{const_name Cons},_) $ _ $ xs) (ts,n) = len xs (ts,n+1)
856 | len (Const(@{const_name append},_) $ xs $ ys) acc = len xs (len ys acc)
857 | len (Const(@{const_name rev},_) $ xs) acc = len xs acc
858 | len (Const(@{const_name map},_) $ _ $ xs) acc = len xs acc
859 | len t (ts,n) = (t::ts,n);
861 val ss = simpset_of @{context};
863 fun list_neq ctxt ct =
865 val (Const(_,eqT) $ lhs $ rhs) = Thm.term_of ct;
866 val (ls,m) = len lhs ([],0) and (rs,n) = len rhs ([],0);
869 val Type(_,listT::_) = eqT;
870 val size = HOLogic.size_const listT;
871 val eq_len = HOLogic.mk_eq (size $ lhs, size $ rhs);
872 val neq_len = HOLogic.mk_Trueprop (HOLogic.Not $ eq_len);
873 val thm = Goal.prove ctxt [] [] neq_len
874 (K (simp_tac (put_simpset ss ctxt) 1));
875 in SOME (thm RS @{thm neq_if_length_neq}) end
877 if m < n andalso submultiset (op aconv) (ls,rs) orelse
878 n < m andalso submultiset (op aconv) (rs,ls)
879 then prove_neq() else NONE
885 subsubsection {* @{text "@"} -- append *}
887 lemma append_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)"
890 lemma append_Nil2 [simp]: "xs @ [] = xs"
893 lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \<and> ys = [])"
896 lemma Nil_is_append_conv [iff]: "([] = xs @ ys) = (xs = [] \<and> ys = [])"
899 lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])"
902 lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])"
905 lemma append_eq_append_conv [simp]:
906 "length xs = length ys \<or> length us = length vs
907 ==> (xs@us = ys@vs) = (xs=ys \<and> us=vs)"
908 apply (induct xs arbitrary: ys)
909 apply (case_tac ys, simp, force)
910 apply (case_tac ys, force, simp)
913 lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) =
914 (EX us. xs = zs @ us & us @ ys = ts | xs @ us = zs & ys = us@ ts)"
915 apply (induct xs arbitrary: ys zs ts)
922 lemma same_append_eq [iff, induct_simp]: "(xs @ ys = xs @ zs) = (ys = zs)"
925 lemma append1_eq_conv [iff]: "(xs @ [x] = ys @ [y]) = (xs = ys \<and> x = y)"
928 lemma append_same_eq [iff, induct_simp]: "(ys @ xs = zs @ xs) = (ys = zs)"
931 lemma append_self_conv2 [iff]: "(xs @ ys = ys) = (xs = [])"
932 using append_same_eq [of _ _ "[]"] by auto
934 lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])"
935 using append_same_eq [of "[]"] by auto
937 lemma hd_Cons_tl [simp]: "xs \<noteq> [] ==> hd xs # tl xs = xs"
940 lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)"
943 lemma hd_append2 [simp]: "xs \<noteq> [] ==> hd (xs @ ys) = hd xs"
944 by (simp add: hd_append split: list.split)
946 lemma tl_append: "tl (xs @ ys) = (case xs of [] => tl ys | z#zs => zs @ ys)"
947 by (simp split: list.split)
949 lemma tl_append2 [simp]: "xs \<noteq> [] ==> tl (xs @ ys) = tl xs @ ys"
950 by (simp add: tl_append split: list.split)
953 lemma Cons_eq_append_conv: "x#xs = ys@zs =
954 (ys = [] & x#xs = zs | (EX ys'. x#ys' = ys & xs = ys'@zs))"
957 lemma append_eq_Cons_conv: "(ys@zs = x#xs) =
958 (ys = [] & zs = x#xs | (EX ys'. ys = x#ys' & ys'@zs = xs))"
962 text {* Trivial rules for solving @{text "@"}-equations automatically. *}
964 lemma eq_Nil_appendI: "xs = ys ==> xs = [] @ ys"
967 lemma Cons_eq_appendI:
968 "[| x # xs1 = ys; xs = xs1 @ zs |] ==> x # xs = ys @ zs"
971 lemma append_eq_appendI:
972 "[| xs @ xs1 = zs; ys = xs1 @ us |] ==> xs @ ys = zs @ us"
977 Simplification procedure for all list equalities.
978 Currently only tries to rearrange @{text "@"} to see if
979 - both lists end in a singleton list,
980 - or both lists end in the same list.
983 simproc_setup list_eq ("(xs::'a list) = ys") = {*
985 fun last (cons as Const (@{const_name Cons}, _) $ _ $ xs) =
986 (case xs of Const (@{const_name Nil}, _) => cons | _ => last xs)
987 | last (Const(@{const_name append},_) $ _ $ ys) = last ys
990 fun list1 (Const(@{const_name Cons},_) $ _ $ Const(@{const_name Nil},_)) = true
993 fun butlast ((cons as Const(@{const_name Cons},_) $ x) $ xs) =
994 (case xs of Const (@{const_name Nil}, _) => xs | _ => cons $ butlast xs)
995 | butlast ((app as Const (@{const_name append}, _) $ xs) $ ys) = app $ butlast ys
996 | butlast xs = Const(@{const_name Nil}, fastype_of xs);
999 simpset_of (put_simpset HOL_basic_ss @{context}
1000 addsimps [@{thm append_assoc}, @{thm append_Nil}, @{thm append_Cons}]);
1002 fun list_eq ctxt (F as (eq as Const(_,eqT)) $ lhs $ rhs) =
1004 val lastl = last lhs and lastr = last rhs;
1007 val lhs1 = butlast lhs and rhs1 = butlast rhs;
1008 val Type(_,listT::_) = eqT
1009 val appT = [listT,listT] ---> listT
1010 val app = Const(@{const_name append},appT)
1011 val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr)
1012 val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2));
1013 val thm = Goal.prove ctxt [] [] eq
1014 (K (simp_tac (put_simpset rearr_ss ctxt) 1));
1015 in SOME ((conv RS (thm RS trans)) RS eq_reflection) end;
1017 if list1 lastl andalso list1 lastr then rearr @{thm append1_eq_conv}
1018 else if lastl aconv lastr then rearr @{thm append_same_eq}
1021 in fn _ => fn ctxt => fn ct => list_eq ctxt (term_of ct) end;
1025 subsubsection {* @{const map} *}
1028 "xs \<noteq> [] \<Longrightarrow> hd (map f xs) = f (hd xs)"
1029 by (cases xs) simp_all
1032 "map f (tl xs) = tl (map f xs)"
1033 by (cases xs) simp_all
1035 lemma map_ext: "(!!x. x : set xs --> f x = g x) ==> map f xs = map g xs"
1036 by (induct xs) simp_all
1038 lemma map_ident [simp]: "map (\<lambda>x. x) = (\<lambda>xs. xs)"
1039 by (rule ext, induct_tac xs) auto
1041 lemma map_append [simp]: "map f (xs @ ys) = map f xs @ map f ys"
1044 lemma map_map [simp]: "map f (map g xs) = map (f \<circ> g) xs"
1047 lemma map_comp_map[simp]: "((map f) o (map g)) = map(f o g)"
1052 lemma rev_map: "rev (map f xs) = map f (rev xs)"
1055 lemma map_eq_conv[simp]: "(map f xs = map g xs) = (!x : set xs. f x = g x)"
1058 lemma map_cong [fundef_cong]:
1059 "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> map f xs = map g ys"
1062 lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])"
1065 lemma Nil_is_map_conv [iff]: "([] = map f xs) = (xs = [])"
1068 lemma map_eq_Cons_conv:
1069 "(map f xs = y#ys) = (\<exists>z zs. xs = z#zs \<and> f z = y \<and> map f zs = ys)"
1072 lemma Cons_eq_map_conv:
1073 "(x#xs = map f ys) = (\<exists>z zs. ys = z#zs \<and> x = f z \<and> xs = map f zs)"
1076 lemmas map_eq_Cons_D = map_eq_Cons_conv [THEN iffD1]
1077 lemmas Cons_eq_map_D = Cons_eq_map_conv [THEN iffD1]
1078 declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!]
1081 "(EX xs. ys = map f xs) = (ALL y : set ys. EX x. y = f x)"
1082 by(induct ys, auto simp add: Cons_eq_map_conv)
1084 lemma map_eq_imp_length_eq:
1085 assumes "map f xs = map g ys"
1086 shows "length xs = length ys"
1088 proof (induct ys arbitrary: xs)
1089 case Nil then show ?case by simp
1091 case (Cons y ys) then obtain z zs where xs: "xs = z # zs" by auto
1092 from Cons xs have "map f zs = map g ys" by simp
1093 with Cons have "length zs = length ys" by blast
1094 with xs show ?case by simp
1098 "[| map f xs = map f ys; inj_on f (set xs Un set ys) |]
1100 apply(frule map_eq_imp_length_eq)
1101 apply(rotate_tac -1)
1102 apply(induct rule:list_induct2)
1105 apply (blast intro:sym)
1108 lemma inj_on_map_eq_map:
1109 "inj_on f (set xs Un set ys) \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
1110 by(blast dest:map_inj_on)
1112 lemma map_injective:
1113 "map f xs = map f ys ==> inj f ==> xs = ys"
1114 by (induct ys arbitrary: xs) (auto dest!:injD)
1116 lemma inj_map_eq_map[simp]: "inj f \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
1117 by(blast dest:map_injective)
1119 lemma inj_mapI: "inj f ==> inj (map f)"
1120 by (iprover dest: map_injective injD intro: inj_onI)
1122 lemma inj_mapD: "inj (map f) ==> inj f"
1123 apply (unfold inj_on_def, clarify)
1124 apply (erule_tac x = "[x]" in ballE)
1125 apply (erule_tac x = "[y]" in ballE, simp, blast)
1129 lemma inj_map[iff]: "inj (map f) = inj f"
1130 by (blast dest: inj_mapD intro: inj_mapI)
1132 lemma inj_on_mapI: "inj_on f (\<Union>(set ` A)) \<Longrightarrow> inj_on (map f) A"
1134 apply(erule map_inj_on)
1135 apply(blast intro:inj_onI dest:inj_onD)
1138 lemma map_idI: "(\<And>x. x \<in> set xs \<Longrightarrow> f x = x) \<Longrightarrow> map f xs = xs"
1139 by (induct xs, auto)
1141 lemma map_fun_upd [simp]: "y \<notin> set xs \<Longrightarrow> map (f(y:=v)) xs = map f xs"
1144 lemma map_fst_zip[simp]:
1145 "length xs = length ys \<Longrightarrow> map fst (zip xs ys) = xs"
1146 by (induct rule:list_induct2, simp_all)
1148 lemma map_snd_zip[simp]:
1149 "length xs = length ys \<Longrightarrow> map snd (zip xs ys) = ys"
1150 by (induct rule:list_induct2, simp_all)
1152 enriched_type map: map
1153 by (simp_all add: id_def)
1155 declare map.id [simp]
1158 subsubsection {* @{const rev} *}
1160 lemma rev_append [simp]: "rev (xs @ ys) = rev ys @ rev xs"
1163 lemma rev_rev_ident [simp]: "rev (rev xs) = xs"
1166 lemma rev_swap: "(rev xs = ys) = (xs = rev ys)"
1169 lemma rev_is_Nil_conv [iff]: "(rev xs = []) = (xs = [])"
1172 lemma Nil_is_rev_conv [iff]: "([] = rev xs) = (xs = [])"
1175 lemma rev_singleton_conv [simp]: "(rev xs = [x]) = (xs = [x])"
1178 lemma singleton_rev_conv [simp]: "([x] = rev xs) = (xs = [x])"
1181 lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)"
1182 apply (induct xs arbitrary: ys, force)
1183 apply (case_tac ys, simp, force)
1186 lemma inj_on_rev[iff]: "inj_on rev A"
1187 by(simp add:inj_on_def)
1189 lemma rev_induct [case_names Nil snoc]:
1190 "[| P []; !!x xs. P xs ==> P (xs @ [x]) |] ==> P xs"
1191 apply(simplesubst rev_rev_ident[symmetric])
1192 apply(rule_tac list = "rev xs" in list.induct, simp_all)
1195 lemma rev_exhaust [case_names Nil snoc]:
1196 "(xs = [] ==> P) ==>(!!ys y. xs = ys @ [y] ==> P) ==> P"
1197 by (induct xs rule: rev_induct) auto
1199 lemmas rev_cases = rev_exhaust
1201 lemma rev_eq_Cons_iff[iff]: "(rev xs = y#ys) = (xs = rev ys @ [y])"
1202 by(rule rev_cases[of xs]) auto
1205 subsubsection {* @{const set} *}
1207 declare set.simps [code_post] --"pretty output"
1209 lemma finite_set [iff]: "finite (set xs)"
1212 lemma set_append [simp]: "set (xs @ ys) = (set xs \<union> set ys)"
1215 lemma hd_in_set[simp]: "xs \<noteq> [] \<Longrightarrow> hd xs : set xs"
1218 lemma set_subset_Cons: "set xs \<subseteq> set (x # xs)"
1221 lemma set_ConsD: "y \<in> set (x # xs) \<Longrightarrow> y=x \<or> y \<in> set xs"
1224 lemma set_empty [iff]: "(set xs = {}) = (xs = [])"
1227 lemma set_empty2[iff]: "({} = set xs) = (xs = [])"
1230 lemma set_rev [simp]: "set (rev xs) = set xs"
1233 lemma set_map [simp]: "set (map f xs) = f`(set xs)"
1236 lemma set_filter [simp]: "set (filter P xs) = {x. x : set xs \<and> P x}"
1239 lemma set_upt [simp]: "set[i..<j] = {i..<j}"
1243 lemma split_list: "x : set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs"
1245 case Nil thus ?case by simp
1247 case Cons thus ?case by (auto intro: Cons_eq_appendI)
1250 lemma in_set_conv_decomp: "x \<in> set xs \<longleftrightarrow> (\<exists>ys zs. xs = ys @ x # zs)"
1251 by (auto elim: split_list)
1253 lemma split_list_first: "x : set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys"
1255 case Nil thus ?case by simp
1260 assume "x = a" thus ?case using Cons by fastforce
1262 assume "x \<noteq> a" thus ?case using Cons by(fastforce intro!: Cons_eq_appendI)
1266 lemma in_set_conv_decomp_first:
1267 "(x : set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys)"
1268 by (auto dest!: split_list_first)
1270 lemma split_list_last: "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs"
1271 proof (induct xs rule: rev_induct)
1272 case Nil thus ?case by simp
1277 assume "x = a" thus ?case using snoc by (metis List.set.simps(1) emptyE)
1279 assume "x \<noteq> a" thus ?case using snoc by fastforce
1283 lemma in_set_conv_decomp_last:
1284 "(x : set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs)"
1285 by (auto dest!: split_list_last)
1287 lemma split_list_prop: "\<exists>x \<in> set xs. P x \<Longrightarrow> \<exists>ys x zs. xs = ys @ x # zs & P x"
1289 case Nil thus ?case by simp
1291 case Cons thus ?case
1292 by(simp add:Bex_def)(metis append_Cons append.simps(1))
1295 lemma split_list_propE:
1296 assumes "\<exists>x \<in> set xs. P x"
1297 obtains ys x zs where "xs = ys @ x # zs" and "P x"
1298 using split_list_prop [OF assms] by blast
1300 lemma split_list_first_prop:
1301 "\<exists>x \<in> set xs. P x \<Longrightarrow>
1302 \<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y)"
1304 case Nil thus ?case by simp
1310 thus ?thesis by simp (metis Un_upper1 contra_subsetD in_set_conv_decomp_first self_append_conv2 set_append)
1313 hence "\<exists>x\<in>set xs. P x" using Cons(2) by simp
1314 thus ?thesis using `\<not> P x` Cons(1) by (metis append_Cons set_ConsD)
1318 lemma split_list_first_propE:
1319 assumes "\<exists>x \<in> set xs. P x"
1320 obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>y \<in> set ys. \<not> P y"
1321 using split_list_first_prop [OF assms] by blast
1323 lemma split_list_first_prop_iff:
1324 "(\<exists>x \<in> set xs. P x) \<longleftrightarrow>
1325 (\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y))"
1326 by (rule, erule split_list_first_prop) auto
1328 lemma split_list_last_prop:
1329 "\<exists>x \<in> set xs. P x \<Longrightarrow>
1330 \<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z)"
1331 proof(induct xs rule:rev_induct)
1332 case Nil thus ?case by simp
1337 assume "P x" thus ?thesis by (metis emptyE set_empty)
1340 hence "\<exists>x\<in>set xs. P x" using snoc(2) by simp
1341 thus ?thesis using `\<not> P x` snoc(1) by fastforce
1345 lemma split_list_last_propE:
1346 assumes "\<exists>x \<in> set xs. P x"
1347 obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>z \<in> set zs. \<not> P z"
1348 using split_list_last_prop [OF assms] by blast
1350 lemma split_list_last_prop_iff:
1351 "(\<exists>x \<in> set xs. P x) \<longleftrightarrow>
1352 (\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z))"
1353 by (metis split_list_last_prop [where P=P] in_set_conv_decomp)
1355 lemma finite_list: "finite A ==> EX xs. set xs = A"
1356 by (erule finite_induct)
1357 (auto simp add: set.simps(2) [symmetric] simp del: set.simps(2))
1359 lemma card_length: "card (set xs) \<le> length xs"
1360 by (induct xs) (auto simp add: card_insert_if)
1362 lemma set_minus_filter_out:
1363 "set xs - {y} = set (filter (\<lambda>x. \<not> (x = y)) xs)"
1367 subsubsection {* @{const filter} *}
1369 lemma filter_append [simp]: "filter P (xs @ ys) = filter P xs @ filter P ys"
1372 lemma rev_filter: "rev (filter P xs) = filter P (rev xs)"
1373 by (induct xs) simp_all
1375 lemma filter_filter [simp]: "filter P (filter Q xs) = filter (\<lambda>x. Q x \<and> P x) xs"
1378 lemma length_filter_le [simp]: "length (filter P xs) \<le> length xs"
1379 by (induct xs) (auto simp add: le_SucI)
1381 lemma sum_length_filter_compl:
1382 "length(filter P xs) + length(filter (%x. ~P x) xs) = length xs"
1383 by(induct xs) simp_all
1385 lemma filter_True [simp]: "\<forall>x \<in> set xs. P x ==> filter P xs = xs"
1388 lemma filter_False [simp]: "\<forall>x \<in> set xs. \<not> P x ==> filter P xs = []"
1391 lemma filter_empty_conv: "(filter P xs = []) = (\<forall>x\<in>set xs. \<not> P x)"
1392 by (induct xs) simp_all
1394 lemma filter_id_conv: "(filter P xs = xs) = (\<forall>x\<in>set xs. P x)"
1397 apply(cut_tac P=P and xs=xs in length_filter_le)
1402 "filter P (map f xs) = map f (filter (P o f) xs)"
1403 by (induct xs) simp_all
1405 lemma length_filter_map[simp]:
1406 "length (filter P (map f xs)) = length(filter (P o f) xs)"
1407 by (simp add:filter_map)
1409 lemma filter_is_subset [simp]: "set (filter P xs) \<le> set xs"
1412 lemma length_filter_less:
1413 "\<lbrakk> x : set xs; ~ P x \<rbrakk> \<Longrightarrow> length(filter P xs) < length xs"
1415 case Nil thus ?case by simp
1417 case (Cons x xs) thus ?case
1418 apply (auto split:split_if_asm)
1419 using length_filter_le[of P xs] apply arith
1423 lemma length_filter_conv_card:
1424 "length(filter p xs) = card{i. i < length xs & p(xs!i)}"
1426 case Nil thus ?case by simp
1429 let ?S = "{i. i < length xs & p(xs!i)}"
1430 have fin: "finite ?S" by(fast intro: bounded_nat_set_is_finite)
1431 show ?case (is "?l = card ?S'")
1434 hence eq: "?S' = insert 0 (Suc ` ?S)"
1435 by(auto simp: image_def split:nat.split dest:gr0_implies_Suc)
1436 have "length (filter p (x # xs)) = Suc(card ?S)"
1437 using Cons `p x` by simp
1438 also have "\<dots> = Suc(card(Suc ` ?S))" using fin
1439 by (simp add: card_image)
1440 also have "\<dots> = card ?S'" using eq fin
1441 by (simp add:card_insert_if) (simp add:image_def)
1442 finally show ?thesis .
1445 hence eq: "?S' = Suc ` ?S"
1446 by(auto simp add: image_def split:nat.split elim:lessE)
1447 have "length (filter p (x # xs)) = card ?S"
1448 using Cons `\<not> p x` by simp
1449 also have "\<dots> = 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)
1453 finally show ?thesis .
1457 lemma Cons_eq_filterD:
1458 "x#xs = filter P ys \<Longrightarrow>
1459 \<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
1460 (is "_ \<Longrightarrow> \<exists>us vs. ?P ys us vs")
1462 case Nil thus ?case by simp
1465 show ?case (is "\<exists>x. ?Q x")
1471 with Py Cons.prems have "?Q []" by simp
1472 then show ?thesis ..
1474 assume "x \<noteq> y"
1475 with Py Cons.prems show ?thesis by simp
1479 with Cons obtain us vs where "?P (y#ys) (y#us) vs" by fastforce
1480 then have "?Q (y#us)" by simp
1481 then show ?thesis ..
1485 lemma filter_eq_ConsD:
1486 "filter P ys = x#xs \<Longrightarrow>
1487 \<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
1488 by(rule Cons_eq_filterD) simp
1490 lemma filter_eq_Cons_iff:
1491 "(filter P ys = x#xs) =
1492 (\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
1493 by(auto dest:filter_eq_ConsD)
1495 lemma Cons_eq_filter_iff:
1496 "(x#xs = filter P ys) =
1497 (\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
1498 by(auto dest:Cons_eq_filterD)
1500 lemma filter_cong[fundef_cong]:
1501 "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> P x = Q x) \<Longrightarrow> filter P xs = filter Q ys"
1503 apply(erule thin_rl)
1504 by (induct ys) simp_all
1507 subsubsection {* List partitioning *}
1509 primrec partition :: "('a \<Rightarrow> bool) \<Rightarrow>'a list \<Rightarrow> 'a list \<times> 'a list" where
1510 "partition P [] = ([], [])" |
1511 "partition P (x # xs) =
1512 (let (yes, no) = partition P xs
1513 in if P x then (x # yes, no) else (yes, x # no))"
1515 lemma partition_filter1:
1516 "fst (partition P xs) = filter P xs"
1517 by (induct xs) (auto simp add: Let_def split_def)
1519 lemma partition_filter2:
1520 "snd (partition P xs) = filter (Not o P) xs"
1521 by (induct xs) (auto simp add: Let_def split_def)
1524 assumes "partition P xs = (yes, no)"
1525 shows "(\<forall>p \<in> set yes. P p) \<and> (\<forall>p \<in> set no. \<not> P p)"
1527 from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)"
1529 then show ?thesis by (simp_all add: partition_filter1 partition_filter2)
1532 lemma partition_set:
1533 assumes "partition P xs = (yes, no)"
1534 shows "set yes \<union> set no = set xs"
1536 from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)"
1538 then show ?thesis by (auto simp add: partition_filter1 partition_filter2)
1541 lemma partition_filter_conv[simp]:
1542 "partition f xs = (filter f xs,filter (Not o f) xs)"
1543 unfolding partition_filter2[symmetric]
1544 unfolding partition_filter1[symmetric] by simp
1546 declare partition.simps[simp del]
1549 subsubsection {* @{const concat} *}
1551 lemma concat_append [simp]: "concat (xs @ ys) = concat xs @ concat ys"
1554 lemma concat_eq_Nil_conv [simp]: "(concat xss = []) = (\<forall>xs \<in> set xss. xs = [])"
1555 by (induct xss) auto
1557 lemma Nil_eq_concat_conv [simp]: "([] = concat xss) = (\<forall>xs \<in> set xss. xs = [])"
1558 by (induct xss) auto
1560 lemma set_concat [simp]: "set (concat xs) = (UN x:set xs. set x)"
1563 lemma concat_map_singleton[simp]: "concat(map (%x. [f x]) xs) = map f xs"
1566 lemma map_concat: "map f (concat xs) = concat (map (map f) xs)"
1569 lemma filter_concat: "filter p (concat xs) = concat (map (filter p) xs)"
1572 lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))"
1575 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)"
1576 proof (induct xs arbitrary: ys)
1578 thus ?case by (cases ys) auto
1581 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"
1582 by (simp add: concat_eq_concat_iff)
1585 subsubsection {* @{const nth} *}
1587 lemma nth_Cons_0 [simp, code]: "(x # xs)!0 = x"
1590 lemma nth_Cons_Suc [simp, code]: "(x # xs)!(Suc n) = xs!n"
1593 declare nth.simps [simp del]
1595 lemma nth_Cons_pos[simp]: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)"
1596 by(auto simp: Nat.gr0_conv_Suc)
1599 "(xs @ ys)!n = (if n < length xs then xs!n else ys!(n - length xs))"
1600 apply (induct xs arbitrary: n, simp)
1601 apply (case_tac n, auto)
1604 lemma nth_append_length [simp]: "(xs @ x # ys) ! length xs = x"
1607 lemma nth_append_length_plus[simp]: "(xs @ ys) ! (length xs + n) = ys ! n"
1610 lemma nth_map [simp]: "n < length xs ==> (map f xs)!n = f(xs!n)"
1611 apply (induct xs arbitrary: n, simp)
1612 apply (case_tac n, auto)
1616 assumes "n < length (tl x)" shows "tl x ! n = x ! Suc n"
1617 using assms by (induct x) auto
1619 lemma hd_conv_nth: "xs \<noteq> [] \<Longrightarrow> hd xs = xs!0"
1620 by(cases xs) simp_all
1623 lemma list_eq_iff_nth_eq:
1624 "(xs = ys) = (length xs = length ys \<and> (ALL i<length xs. xs!i = ys!i))"
1625 apply(induct xs arbitrary: ys)
1629 apply(simp add:nth_Cons split:nat.split)apply blast
1632 lemma set_conv_nth: "set xs = {xs!i | i. i < length xs}"
1633 apply (induct xs, simp, simp)
1635 apply (metis nat_case_0 nth.simps zero_less_Suc)
1636 apply (metis less_Suc_eq_0_disj nth_Cons_Suc)
1637 apply (case_tac i, simp)
1638 apply (metis diff_Suc_Suc nat_case_Suc nth.simps zero_less_diff)
1641 lemma in_set_conv_nth: "(x \<in> set xs) = (\<exists>i < length xs. xs!i = x)"
1642 by(auto simp:set_conv_nth)
1644 lemma nth_equal_first_eq:
1645 assumes "x \<notin> set xs"
1646 assumes "n \<le> length xs"
1647 shows "(x # xs) ! n = x \<longleftrightarrow> n = 0" (is "?lhs \<longleftrightarrow> ?rhs")
1652 assume "n \<noteq> 0"
1653 then have "n > 0" by simp
1654 with `?lhs` have "xs ! (n - 1) = x" by simp
1655 moreover from `n > 0` `n \<le> length xs` have "n - 1 < length xs" by simp
1656 ultimately have "\<exists>i<length xs. xs ! i = x" by auto
1657 with `x \<notin> set xs` in_set_conv_nth [of x xs] show False by simp
1660 assume ?rhs then show ?lhs by simp
1663 lemma nth_non_equal_first_eq:
1664 assumes "x \<noteq> y"
1665 shows "(x # xs) ! n = y \<longleftrightarrow> xs ! (n - 1) = y \<and> n > 0" (is "?lhs \<longleftrightarrow> ?rhs")
1667 assume "?lhs" with assms have "n > 0" by (cases n) simp_all
1668 with `?lhs` show ?rhs by simp
1670 assume "?rhs" then show "?lhs" by simp
1673 lemma list_ball_nth: "[| n < length xs; !x : set xs. P x|] ==> P(xs!n)"
1674 by (auto simp add: set_conv_nth)
1676 lemma nth_mem [simp]: "n < length xs ==> xs!n : set xs"
1677 by (auto simp add: set_conv_nth)
1679 lemma all_nth_imp_all_set:
1680 "[| !i < length xs. P(xs!i); x : set xs|] ==> P x"
1681 by (auto simp add: set_conv_nth)
1683 lemma all_set_conv_all_nth:
1684 "(\<forall>x \<in> set xs. P x) = (\<forall>i. i < length xs --> P (xs ! i))"
1685 by (auto simp add: set_conv_nth)
1688 "n < size xs \<Longrightarrow> rev xs ! n = xs ! (length xs - Suc n)"
1689 proof (induct xs arbitrary: n)
1690 case Nil thus ?case by simp
1693 hence n: "n < Suc (length xs)" by simp
1695 { assume "n < length xs"
1696 with n obtain n' where n': "length xs - n = Suc n'"
1697 by (cases "length xs - n", auto)
1699 from n' have "length xs - Suc n = n'" by simp
1701 have "xs ! (length xs - Suc n) = (x # xs) ! (length xs - n)" by simp
1704 show ?case by (clarsimp simp add: Cons nth_append)
1707 lemma Skolem_list_nth:
1708 "(ALL i<k. EX x. P i x) = (EX xs. size xs = k & (ALL i<k. P i (xs!i)))"
1709 (is "_ = (EX xs. ?P k xs)")
1711 case 0 show ?case by simp
1714 show ?case (is "?L = ?R" is "_ = (EX xs. ?P' xs)")
1716 assume "?R" thus "?L" using Suc by auto
1719 with Suc obtain x xs where "?P k xs & P k x" by (metis less_Suc_eq)
1720 hence "?P'(xs@[x])" by(simp add:nth_append less_Suc_eq)
1726 subsubsection {* @{const list_update} *}
1728 lemma length_list_update [simp]: "length(xs[i:=x]) = length xs"
1729 by (induct xs arbitrary: i) (auto split: nat.split)
1731 lemma nth_list_update:
1732 "i < length xs==> (xs[i:=x])!j = (if i = j then x else xs!j)"
1733 by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
1735 lemma nth_list_update_eq [simp]: "i < length xs ==> (xs[i:=x])!i = x"
1736 by (simp add: nth_list_update)
1738 lemma nth_list_update_neq [simp]: "i \<noteq> j ==> xs[i:=x]!j = xs!j"
1739 by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
1741 lemma list_update_id[simp]: "xs[i := xs!i] = xs"
1742 by (induct xs arbitrary: i) (simp_all split:nat.splits)
1744 lemma list_update_beyond[simp]: "length xs \<le> i \<Longrightarrow> xs[i:=x] = xs"
1745 apply (induct xs arbitrary: i)
1751 lemma list_update_nonempty[simp]: "xs[k:=x] = [] \<longleftrightarrow> xs=[]"
1752 by(metis length_0_conv length_list_update)
1754 lemma list_update_same_conv:
1755 "i < length xs ==> (xs[i := x] = xs) = (xs!i = x)"
1756 by (induct xs arbitrary: i) (auto split: nat.split)
1758 lemma list_update_append1:
1759 "i < size xs \<Longrightarrow> (xs @ ys)[i:=x] = xs[i:=x] @ ys"
1760 apply (induct xs arbitrary: i, simp)
1761 apply(simp split:nat.split)
1764 lemma list_update_append:
1765 "(xs @ ys) [n:= x] =
1766 (if n < length xs then xs[n:= x] @ ys else xs @ (ys [n-length xs:= x]))"
1767 by (induct xs arbitrary: n) (auto split:nat.splits)
1769 lemma list_update_length [simp]:
1770 "(xs @ x # ys)[length xs := y] = (xs @ y # ys)"
1771 by (induct xs, auto)
1773 lemma map_update: "map f (xs[k:= y]) = (map f xs)[k := f y]"
1774 by(induct xs arbitrary: k)(auto split:nat.splits)
1777 "k < length xs \<Longrightarrow> rev (xs[k:= y]) = (rev xs)[length xs - k - 1 := y]"
1778 by (induct xs arbitrary: k) (auto simp: list_update_append split:nat.splits)
1781 "(zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])"
1782 by (induct ys arbitrary: i xy xs) (auto, case_tac xs, auto split: nat.split)
1784 lemma set_update_subset_insert: "set(xs[i:=x]) <= insert x (set xs)"
1785 by (induct xs arbitrary: i) (auto split: nat.split)
1787 lemma set_update_subsetI: "[| set xs <= A; x:A |] ==> set(xs[i := x]) <= A"
1788 by (blast dest!: set_update_subset_insert [THEN subsetD])
1790 lemma set_update_memI: "n < length xs \<Longrightarrow> x \<in> set (xs[n := x])"
1791 by (induct xs arbitrary: n) (auto split:nat.splits)
1793 lemma list_update_overwrite[simp]:
1794 "xs [i := x, i := y] = xs [i := y]"
1795 apply (induct xs arbitrary: i) apply simp
1796 apply (case_tac i, simp_all)
1799 lemma list_update_swap:
1800 "i \<noteq> i' \<Longrightarrow> xs [i := x, i' := x'] = xs [i' := x', i := x]"
1801 apply (induct xs arbitrary: i i')
1803 apply (case_tac i, case_tac i')
1809 lemma list_update_code [code]:
1811 "(x # xs)[0 := y] = y # xs"
1812 "(x # xs)[Suc i := y] = x # xs[i := y]"
1816 subsubsection {* @{const last} and @{const butlast} *}
1818 lemma last_snoc [simp]: "last (xs @ [x]) = x"
1821 lemma butlast_snoc [simp]: "butlast (xs @ [x]) = xs"
1824 lemma last_ConsL: "xs = [] \<Longrightarrow> last(x#xs) = x"
1827 lemma last_ConsR: "xs \<noteq> [] \<Longrightarrow> last(x#xs) = last xs"
1830 lemma last_append: "last(xs @ ys) = (if ys = [] then last xs else last ys)"
1831 by (induct xs) (auto)
1833 lemma last_appendL[simp]: "ys = [] \<Longrightarrow> last(xs @ ys) = last xs"
1834 by(simp add:last_append)
1836 lemma last_appendR[simp]: "ys \<noteq> [] \<Longrightarrow> last(xs @ ys) = last ys"
1837 by(simp add:last_append)
1839 lemma last_tl: "xs = [] \<or> tl xs \<noteq> [] \<Longrightarrow>last (tl xs) = last xs"
1840 by (induct xs) simp_all
1842 lemma butlast_tl: "butlast (tl xs) = tl (butlast xs)"
1843 by (induct xs) simp_all
1845 lemma hd_rev: "xs \<noteq> [] \<Longrightarrow> hd(rev xs) = last xs"
1846 by(rule rev_exhaust[of xs]) simp_all
1848 lemma last_rev: "xs \<noteq> [] \<Longrightarrow> last(rev xs) = hd xs"
1849 by(cases xs) simp_all
1851 lemma last_in_set[simp]: "as \<noteq> [] \<Longrightarrow> last as \<in> set as"
1854 lemma length_butlast [simp]: "length (butlast xs) = length xs - 1"
1855 by (induct xs rule: rev_induct) auto
1857 lemma butlast_append:
1858 "butlast (xs @ ys) = (if ys = [] then butlast xs else xs @ butlast ys)"
1859 by (induct xs arbitrary: ys) auto
1861 lemma append_butlast_last_id [simp]:
1862 "xs \<noteq> [] ==> butlast xs @ [last xs] = xs"
1865 lemma in_set_butlastD: "x : set (butlast xs) ==> x : set xs"
1866 by (induct xs) (auto split: split_if_asm)
1868 lemma in_set_butlast_appendI:
1869 "x : set (butlast xs) | x : set (butlast ys) ==> x : set (butlast (xs @ ys))"
1870 by (auto dest: in_set_butlastD simp add: butlast_append)
1872 lemma last_drop[simp]: "n < length xs \<Longrightarrow> last (drop n xs) = last xs"
1873 apply (induct xs arbitrary: n)
1875 apply (auto split:nat.split)
1879 assumes "n < length (butlast xs)" shows "butlast xs ! n = xs ! n"
1882 moreover from assms have "butlast xs ! n = (butlast xs @ [last xs]) ! n"
1883 by (simp add: nth_append)
1884 ultimately show ?thesis using append_butlast_last_id by simp
1887 lemma last_conv_nth: "xs\<noteq>[] \<Longrightarrow> last xs = xs!(length xs - 1)"
1888 by(induct xs)(auto simp:neq_Nil_conv)
1890 lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs"
1891 by (induct xs, simp, case_tac xs, simp_all)
1893 lemma last_list_update:
1894 "xs \<noteq> [] \<Longrightarrow> last(xs[k:=x]) = (if k = size xs - 1 then x else last xs)"
1895 by (auto simp: last_conv_nth)
1897 lemma butlast_list_update:
1898 "butlast(xs[k:=x]) =
1899 (if k = size xs - 1 then butlast xs else (butlast xs)[k:=x])"
1900 apply(cases xs rule:rev_cases)
1902 apply(simp add:list_update_append split:nat.splits)
1906 "xs \<noteq> [] \<Longrightarrow> last (map f xs) = f (last xs)"
1907 by (cases xs rule: rev_cases) simp_all
1910 "map f (butlast xs) = butlast (map f xs)"
1911 by (induct xs) simp_all
1913 lemma snoc_eq_iff_butlast:
1914 "xs @ [x] = ys \<longleftrightarrow> (ys \<noteq> [] & butlast ys = xs & last ys = x)"
1915 by (metis append_butlast_last_id append_is_Nil_conv butlast_snoc last_snoc not_Cons_self)
1918 subsubsection {* @{const take} and @{const drop} *}
1920 lemma take_0 [simp]: "take 0 xs = []"
1923 lemma drop_0 [simp]: "drop 0 xs = xs"
1926 lemma take_Suc_Cons [simp]: "take (Suc n) (x # xs) = x # take n xs"
1929 lemma drop_Suc_Cons [simp]: "drop (Suc n) (x # xs) = drop n xs"
1932 declare take_Cons [simp del] and drop_Cons [simp del]
1934 lemma take_1_Cons [simp]: "take 1 (x # xs) = [x]"
1935 unfolding One_nat_def by simp
1937 lemma drop_1_Cons [simp]: "drop 1 (x # xs) = xs"
1938 unfolding One_nat_def by simp
1940 lemma take_Suc: "xs ~= [] ==> take (Suc n) xs = hd xs # take n (tl xs)"
1941 by(clarsimp simp add:neq_Nil_conv)
1943 lemma drop_Suc: "drop (Suc n) xs = drop n (tl xs)"
1944 by(cases xs, simp_all)
1946 lemma take_tl: "take n (tl xs) = tl (take (Suc n) xs)"
1947 by (induct xs arbitrary: n) simp_all
1949 lemma drop_tl: "drop n (tl xs) = tl(drop n xs)"
1950 by(induct xs arbitrary: n, simp_all add:drop_Cons drop_Suc split:nat.split)
1952 lemma tl_take: "tl (take n xs) = take (n - 1) (tl xs)"
1953 by (cases n, simp, cases xs, auto)
1955 lemma tl_drop: "tl (drop n xs) = drop n (tl xs)"
1956 by (simp only: drop_tl)
1958 lemma nth_via_drop: "drop n xs = y#ys \<Longrightarrow> xs!n = y"
1959 apply (induct xs arbitrary: n, simp)
1960 apply(simp add:drop_Cons nth_Cons split:nat.splits)
1963 lemma take_Suc_conv_app_nth:
1964 "i < length xs \<Longrightarrow> take (Suc i) xs = take i xs @ [xs!i]"
1965 apply (induct xs arbitrary: i, simp)
1966 apply (case_tac i, auto)
1969 lemma drop_Suc_conv_tl:
1970 "i < length xs \<Longrightarrow> (xs!i) # (drop (Suc i) xs) = drop i xs"
1971 apply (induct xs arbitrary: i, simp)
1972 apply (case_tac i, auto)
1975 lemma length_take [simp]: "length (take n xs) = min (length xs) n"
1976 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
1978 lemma length_drop [simp]: "length (drop n xs) = (length xs - n)"
1979 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
1981 lemma take_all [simp]: "length xs <= n ==> take n xs = xs"
1982 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
1984 lemma drop_all [simp]: "length xs <= n ==> drop n xs = []"
1985 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
1987 lemma take_append [simp]:
1988 "take n (xs @ ys) = (take n xs @ take (n - length xs) ys)"
1989 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
1991 lemma drop_append [simp]:
1992 "drop n (xs @ ys) = drop n xs @ drop (n - length xs) ys"
1993 by (induct n arbitrary: xs) (auto, case_tac xs, auto)
1995 lemma take_take [simp]: "take n (take m xs) = take (min n m) xs"
1996 apply (induct m arbitrary: xs n, auto)
1997 apply (case_tac xs, auto)
1998 apply (case_tac n, auto)
2001 lemma drop_drop [simp]: "drop n (drop m xs) = drop (n + m) xs"
2002 apply (induct m arbitrary: xs, auto)
2003 apply (case_tac xs, auto)
2006 lemma take_drop: "take n (drop m xs) = drop m (take (n + m) xs)"
2007 apply (induct m arbitrary: xs n, auto)
2008 apply (case_tac xs, auto)
2011 lemma drop_take: "drop n (take m xs) = take (m-n) (drop n xs)"
2012 apply(induct xs arbitrary: m n)
2014 apply(simp add: take_Cons drop_Cons split:nat.split)
2017 lemma append_take_drop_id [simp]: "take n xs @ drop n xs = xs"
2018 apply (induct n arbitrary: xs, auto)
2019 apply (case_tac xs, auto)
2022 lemma take_eq_Nil[simp]: "(take n xs = []) = (n = 0 \<or> xs = [])"
2023 apply(induct xs arbitrary: n)
2025 apply(simp add:take_Cons split:nat.split)
2028 lemma drop_eq_Nil[simp]: "(drop n xs = []) = (length xs <= n)"
2029 apply(induct xs arbitrary: n)
2031 apply(simp add:drop_Cons split:nat.split)
2034 lemma take_map: "take n (map f xs) = map f (take n xs)"
2035 apply (induct n arbitrary: xs, auto)
2036 apply (case_tac xs, auto)
2039 lemma drop_map: "drop n (map f xs) = map f (drop n xs)"
2040 apply (induct n arbitrary: xs, auto)
2041 apply (case_tac xs, auto)
2044 lemma rev_take: "rev (take i xs) = drop (length xs - i) (rev xs)"
2045 apply (induct xs arbitrary: i, auto)
2046 apply (case_tac i, auto)
2049 lemma rev_drop: "rev (drop i xs) = take (length xs - i) (rev xs)"
2050 apply (induct xs arbitrary: i, auto)
2051 apply (case_tac i, auto)
2054 lemma nth_take [simp]: "i < n ==> (take n xs)!i = xs!i"
2055 apply (induct xs arbitrary: i n, auto)
2056 apply (case_tac n, blast)
2057 apply (case_tac i, auto)
2060 lemma nth_drop [simp]:
2061 "n + i <= length xs ==> (drop n xs)!i = xs!(n + i)"
2062 apply (induct n arbitrary: xs i, auto)
2063 apply (case_tac xs, auto)
2067 "n <= length xs ==> butlast (take n xs) = take (n - 1) xs"
2068 by (simp add: butlast_conv_take min.absorb1 min.absorb2)
2070 lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)"
2071 by (simp add: butlast_conv_take drop_take add_ac)
2073 lemma take_butlast: "n < length xs ==> take n (butlast xs) = take n xs"
2074 by (simp add: butlast_conv_take min.absorb1)
2076 lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)"
2077 by (simp add: butlast_conv_take drop_take add_ac)
2079 lemma hd_drop_conv_nth: "n < length xs \<Longrightarrow> hd(drop n xs) = xs!n"
2080 by(simp add: hd_conv_nth)
2082 lemma set_take_subset_set_take:
2083 "m <= n \<Longrightarrow> set(take m xs) <= set(take n xs)"
2084 apply (induct xs arbitrary: m n)
2087 apply (auto simp: take_Cons)
2090 lemma set_take_subset: "set(take n xs) \<subseteq> set xs"
2091 by(induct xs arbitrary: n)(auto simp:take_Cons split:nat.split)
2093 lemma set_drop_subset: "set(drop n xs) \<subseteq> set xs"
2094 by(induct xs arbitrary: n)(auto simp:drop_Cons split:nat.split)
2096 lemma set_drop_subset_set_drop:
2097 "m >= n \<Longrightarrow> set(drop m xs) <= set(drop n xs)"
2098 apply(induct xs arbitrary: m n)
2099 apply(auto simp:drop_Cons split:nat.split)
2100 apply (metis set_drop_subset subset_iff)
2103 lemma in_set_takeD: "x : set(take n xs) \<Longrightarrow> x : set xs"
2104 using set_take_subset by fast
2106 lemma in_set_dropD: "x : set(drop n xs) \<Longrightarrow> x : set xs"
2107 using set_drop_subset by fast
2109 lemma append_eq_conv_conj:
2110 "(xs @ ys = zs) = (xs = take (length xs) zs \<and> ys = drop (length xs) zs)"
2111 apply (induct xs arbitrary: zs, simp, clarsimp)
2112 apply (case_tac zs, auto)
2116 "take (i+j) xs = take i xs @ take j (drop i xs)"
2117 apply (induct xs arbitrary: i, auto)
2118 apply (case_tac i, simp_all)
2121 lemma append_eq_append_conv_if:
2122 "(xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>1 @ ys\<^sub>2) =
2123 (if size xs\<^sub>1 \<le> size ys\<^sub>1
2124 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
2125 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)"
2126 apply(induct xs\<^sub>1 arbitrary: ys\<^sub>1)
2128 apply(case_tac ys\<^sub>1)
2133 "n < length xs \<Longrightarrow> take n xs @ [hd (drop n xs)] = take (Suc n) xs"
2134 apply(induct xs arbitrary: n)
2136 apply(simp add:drop_Cons split:nat.split)
2139 lemma id_take_nth_drop:
2140 "i < length xs \<Longrightarrow> xs = take i xs @ xs!i # drop (Suc i) xs"
2142 assume si: "i < length xs"
2143 hence "xs = take (Suc i) xs @ drop (Suc i) xs" by auto
2145 from si have "take (Suc i) xs = take i xs @ [xs!i]"
2146 apply (rule_tac take_Suc_conv_app_nth) by arith
2147 ultimately show ?thesis by auto
2150 lemma upd_conv_take_nth_drop:
2151 "i < length xs \<Longrightarrow> xs[i:=a] = take i xs @ a # drop (Suc i) xs"
2153 assume i: "i < length xs"
2154 have "xs[i:=a] = (take i xs @ xs!i # drop (Suc i) xs)[i:=a]"
2155 by(rule arg_cong[OF id_take_nth_drop[OF i]])
2156 also have "\<dots> = take i xs @ a # drop (Suc i) xs"
2157 using i by (simp add: list_update_append)
2158 finally show ?thesis .
2162 "i < length xs \<Longrightarrow> xs ! i # drop (Suc i) xs = drop i xs"
2163 apply (induct i arbitrary: xs)
2164 apply (simp add: neq_Nil_conv)
2172 subsubsection {* @{const takeWhile} and @{const dropWhile} *}
2174 lemma length_takeWhile_le: "length (takeWhile P xs) \<le> length xs"
2177 lemma takeWhile_dropWhile_id [simp]: "takeWhile P xs @ dropWhile P xs = xs"
2180 lemma takeWhile_append1 [simp]:
2181 "[| x:set xs; ~P(x)|] ==> takeWhile P (xs @ ys) = takeWhile P xs"
2184 lemma takeWhile_append2 [simp]:
2185 "(!!x. x : set xs ==> P x) ==> takeWhile P (xs @ ys) = xs @ takeWhile P ys"
2188 lemma takeWhile_tail: "\<not> P x ==> takeWhile P (xs @ (x#l)) = takeWhile P xs"
2191 lemma takeWhile_nth: "j < length (takeWhile P xs) \<Longrightarrow> takeWhile P xs ! j = xs ! j"
2192 apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto
2194 lemma dropWhile_nth: "j < length (dropWhile P xs) \<Longrightarrow> dropWhile P xs ! j = xs ! (j + length (takeWhile P xs))"
2195 apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto
2197 lemma length_dropWhile_le: "length (dropWhile P xs) \<le> length xs"
2200 lemma dropWhile_append1 [simp]:
2201 "[| x : set xs; ~P(x)|] ==> dropWhile P (xs @ ys) = (dropWhile P xs)@ys"
2204 lemma dropWhile_append2 [simp]:
2205 "(!!x. x:set xs ==> P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys"
2208 lemma dropWhile_append3:
2209 "\<not> P y \<Longrightarrow>dropWhile P (xs @ y # ys) = dropWhile P xs @ y # ys"
2212 lemma dropWhile_last:
2213 "x \<in> set xs \<Longrightarrow> \<not> P x \<Longrightarrow> last (dropWhile P xs) = last xs"
2214 by (auto simp add: dropWhile_append3 in_set_conv_decomp)
2216 lemma set_dropWhileD: "x \<in> set (dropWhile P xs) \<Longrightarrow> x \<in> set xs"
2217 by (induct xs) (auto split: split_if_asm)
2219 lemma set_takeWhileD: "x : set (takeWhile P xs) ==> x : set xs \<and> P x"
2220 by (induct xs) (auto split: split_if_asm)
2222 lemma takeWhile_eq_all_conv[simp]:
2223 "(takeWhile P xs = xs) = (\<forall>x \<in> set xs. P x)"
2226 lemma dropWhile_eq_Nil_conv[simp]:
2227 "(dropWhile P xs = []) = (\<forall>x \<in> set xs. P x)"
2230 lemma dropWhile_eq_Cons_conv:
2231 "(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys & \<not> P y)"
2234 lemma distinct_takeWhile[simp]: "distinct xs ==> distinct (takeWhile P xs)"
2235 by (induct xs) (auto dest: set_takeWhileD)
2237 lemma distinct_dropWhile[simp]: "distinct xs ==> distinct (dropWhile P xs)"
2240 lemma takeWhile_map: "takeWhile P (map f xs) = map f (takeWhile (P \<circ> f) xs)"
2243 lemma dropWhile_map: "dropWhile P (map f xs) = map f (dropWhile (P \<circ> f) xs)"
2246 lemma takeWhile_eq_take: "takeWhile P xs = take (length (takeWhile P xs)) xs"
2249 lemma dropWhile_eq_drop: "dropWhile P xs = drop (length (takeWhile P xs)) xs"
2253 "dropWhile P xs \<noteq> [] \<Longrightarrow> \<not> P (hd (dropWhile P xs))"
2254 using assms by (induct xs) auto
2256 lemma takeWhile_eq_filter:
2257 assumes "\<And> x. x \<in> set (dropWhile P xs) \<Longrightarrow> \<not> P x"
2258 shows "takeWhile P xs = filter P xs"
2260 have A: "filter P xs = filter P (takeWhile P xs @ dropWhile P xs)"
2262 have B: "filter P (dropWhile P xs) = []"
2263 unfolding filter_empty_conv using assms by blast
2264 have "filter P xs = takeWhile P xs"
2265 unfolding A filter_append B
2266 by (auto simp add: filter_id_conv dest: set_takeWhileD)
2270 lemma takeWhile_eq_take_P_nth:
2271 "\<lbrakk> \<And> i. \<lbrakk> i < n ; i < length xs \<rbrakk> \<Longrightarrow> P (xs ! i) ; n < length xs \<Longrightarrow> \<not> P (xs ! n) \<rbrakk> \<Longrightarrow>
2272 takeWhile P xs = take n xs"
2273 proof (induct xs arbitrary: n)
2277 case (Suc n') note this[simp]
2278 have "P x" using Cons.prems(1)[of 0] by simp
2279 moreover have "takeWhile P xs = take n' xs"
2280 proof (rule Cons.hyps)
2281 case goal1 thus "P (xs ! i)" using Cons.prems(1)[of "Suc i"] by simp
2282 next case goal2 thus ?case using Cons by auto
2284 ultimately show ?thesis by simp
2288 lemma nth_length_takeWhile:
2289 "length (takeWhile P xs) < length xs \<Longrightarrow> \<not> P (xs ! length (takeWhile P xs))"
2292 lemma length_takeWhile_less_P_nth:
2293 assumes all: "\<And> i. i < j \<Longrightarrow> P (xs ! i)" and "j \<le> length xs"
2294 shows "j \<le> length (takeWhile P xs)"
2295 proof (rule classical)
2296 assume "\<not> ?thesis"
2297 hence "length (takeWhile P xs) < length xs" using assms by simp
2298 thus ?thesis using all `\<not> ?thesis` nth_length_takeWhile[of P xs] by auto
2301 text{* The following two lemmmas could be generalized to an arbitrary
2304 lemma takeWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow>
2305 takeWhile (\<lambda>y. y \<noteq> x) (rev xs) = rev (tl (dropWhile (\<lambda>y. y \<noteq> x) xs))"
2306 by(induct xs) (auto simp: takeWhile_tail[where l="[]"])
2308 lemma dropWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow>
2309 dropWhile (\<lambda>y. y \<noteq> x) (rev xs) = x # rev (takeWhile (\<lambda>y. y \<noteq> x) xs)"
2313 apply(subst dropWhile_append2)
2317 lemma takeWhile_not_last:
2318 "distinct xs \<Longrightarrow> takeWhile (\<lambda>y. y \<noteq> last xs) xs = butlast xs"
2325 lemma takeWhile_cong [fundef_cong]:
2326 "[| l = k; !!x. x : set l ==> P x = Q x |]
2327 ==> takeWhile P l = takeWhile Q k"
2328 by (induct k arbitrary: l) (simp_all)
2330 lemma dropWhile_cong [fundef_cong]:
2331 "[| l = k; !!x. x : set l ==> P x = Q x |]
2332 ==> dropWhile P l = dropWhile Q k"
2333 by (induct k arbitrary: l, simp_all)
2335 lemma takeWhile_idem [simp]:
2336 "takeWhile P (takeWhile P xs) = takeWhile P xs"
2339 lemma dropWhile_idem [simp]:
2340 "dropWhile P (dropWhile P xs) = dropWhile P xs"
2344 subsubsection {* @{const zip} *}
2346 lemma zip_Nil [simp]: "zip [] ys = []"
2349 lemma zip_Cons_Cons [simp]: "zip (x # xs) (y # ys) = (x, y) # zip xs ys"
2352 declare zip_Cons [simp del]
2357 "zip (x # xs) (y # ys) = (x, y) # zip xs ys"
2358 by (fact zip_Nil zip.simps(1) zip_Cons_Cons)+
2361 "zip (x#xs) ys = (case ys of [] \<Rightarrow> [] | y#ys \<Rightarrow> (x,y)#zip xs ys)"
2362 by(auto split:list.split)
2364 lemma length_zip [simp]:
2365 "length (zip xs ys) = min (length xs) (length ys)"
2366 by (induct xs ys rule:list_induct2') auto
2368 lemma zip_obtain_same_length:
2369 assumes "\<And>zs ws n. length zs = length ws \<Longrightarrow> n = min (length xs) (length ys)
2370 \<Longrightarrow> zs = take n xs \<Longrightarrow> ws = take n ys \<Longrightarrow> P (zip zs ws)"
2371 shows "P (zip xs ys)"
2373 let ?n = "min (length xs) (length ys)"
2374 have "P (zip (take ?n xs) (take ?n ys))"
2375 by (rule assms) simp_all
2376 moreover have "zip xs ys = zip (take ?n xs) (take ?n ys)"
2377 proof (induct xs arbitrary: ys)
2378 case Nil then show ?case by simp
2380 case (Cons x xs) then show ?case by (cases ys) simp_all
2382 ultimately show ?thesis by simp
2387 zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)"
2388 by (induct xs zs rule:list_induct2') auto
2392 zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs"
2393 by (induct xs ys rule:list_induct2') auto
2395 lemma zip_append [simp]:
2396 "[| length xs = length us |] ==>
2397 zip (xs@ys) (us@vs) = zip xs us @ zip ys vs"
2398 by (simp add: zip_append1)
2401 "length xs = length ys ==> zip (rev xs) (rev ys) = rev (zip xs ys)"
2402 by (induct rule:list_induct2, simp_all)
2405 "zip (map f xs) (map g ys) = map (\<lambda> (x, y). (f x, g y)) (zip xs ys)"
2406 proof (induct xs arbitrary: ys)
2407 case (Cons x xs) note Cons_x_xs = Cons.hyps
2411 show ?thesis unfolding Cons using Cons_x_xs by simp
2416 "zip (map f xs) ys = map (\<lambda>(x, y). (f x, y)) (zip xs ys)"
2417 using zip_map_map[of f xs "\<lambda>x. x" ys] by simp
2420 "zip xs (map f ys) = map (\<lambda>(x, y). (x, f y)) (zip xs ys)"
2421 using zip_map_map[of "\<lambda>x. x" xs f ys] by simp
2424 "map f (zip (map g xs) ys) = map (%(x,y). f(g x, y)) (zip xs ys)"
2425 unfolding zip_map1 by auto
2428 "map f (zip xs (map g ys)) = map (%(x,y). f(x, g y)) (zip xs ys)"
2429 unfolding zip_map2 by auto
2431 text{* Courtesy of Andreas Lochbihler: *}
2432 lemma zip_same_conv_map: "zip xs xs = map (\<lambda>x. (x, x)) xs"
2435 lemma nth_zip [simp]:
2436 "[| i < length xs; i < length ys|] ==> (zip xs ys)!i = (xs!i, ys!i)"
2437 apply (induct ys arbitrary: i xs, simp)
2439 apply (simp_all add: nth.simps split: nat.split)
2443 "set (zip xs ys) = {(xs!i, ys!i) | i. i < min (length xs) (length ys)}"
2444 by(simp add: set_conv_nth cong: rev_conj_cong)
2446 lemma zip_same: "((a,b) \<in> set (zip xs xs)) = (a \<in> set xs \<and> a = b)"
2450 "zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]"
2451 by(rule sym, simp add: update_zip)
2453 lemma zip_replicate [simp]:
2454 "zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)"
2455 apply (induct i arbitrary: j, auto)
2456 apply (case_tac j, auto)
2460 "take n (zip xs ys) = zip (take n xs) (take n ys)"
2461 apply (induct n arbitrary: xs ys)
2463 apply (case_tac xs, simp)
2464 apply (case_tac ys, simp_all)
2468 "drop n (zip xs ys) = zip (drop n xs) (drop n ys)"
2469 apply (induct n arbitrary: xs ys)
2471 apply (case_tac xs, simp)
2472 apply (case_tac ys, simp_all)
2475 lemma zip_takeWhile_fst: "zip (takeWhile P xs) ys = takeWhile (P \<circ> fst) (zip xs ys)"
2476 proof (induct xs arbitrary: ys)
2477 case (Cons x xs) thus ?case by (cases ys) auto
2480 lemma zip_takeWhile_snd: "zip xs (takeWhile P ys) = takeWhile (P \<circ> snd) (zip xs ys)"
2481 proof (induct xs arbitrary: ys)
2482 case (Cons x xs) thus ?case by (cases ys) auto
2485 lemma set_zip_leftD:
2486 "(x,y)\<in> set (zip xs ys) \<Longrightarrow> x \<in> set xs"
2487 by (induct xs ys rule:list_induct2') auto
2489 lemma set_zip_rightD:
2490 "(x,y)\<in> set (zip xs ys) \<Longrightarrow> y \<in> set ys"
2491 by (induct xs ys rule:list_induct2') auto
2494 "(x,y) : set(zip xs ys) \<Longrightarrow> (\<lbrakk> x : set xs; y : set ys \<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R"
2495 by(blast dest: set_zip_leftD set_zip_rightD)
2497 lemma zip_map_fst_snd:
2498 "zip (map fst zs) (map snd zs) = zs"
2499 by (induct zs) simp_all
2502 "length xs = length ys \<Longrightarrow> zip xs ys = zs \<longleftrightarrow> map fst zs = xs \<and> map snd zs = ys"
2503 by (auto simp add: zip_map_fst_snd)
2506 "p \<in> set (zip xs ys) \<longleftrightarrow> (\<exists>n. xs ! n = fst p \<and> ys ! n = snd p
2507 \<and> n < length xs \<and> n < length ys)"
2508 by (cases p) (auto simp add: set_zip)
2510 lemma pair_list_eqI:
2511 assumes "map fst xs = map fst ys" and "map snd xs = map snd ys"
2514 from assms(1) have "length xs = length ys" by (rule map_eq_imp_length_eq)
2515 from this assms show ?thesis
2516 by (induct xs ys rule: list_induct2) (simp_all add: prod_eqI)
2520 subsubsection {* @{const list_all2} *}
2522 lemma list_all2_lengthD [intro?]:
2523 "list_all2 P xs ys ==> length xs = length ys"
2524 by (simp add: list_all2_def)
2526 lemma list_all2_Nil [iff, code]: "list_all2 P [] ys = (ys = [])"
2527 by (simp add: list_all2_def)
2529 lemma list_all2_Nil2 [iff, code]: "list_all2 P xs [] = (xs = [])"
2530 by (simp add: list_all2_def)
2532 lemma list_all2_Cons [iff, code]:
2533 "list_all2 P (x # xs) (y # ys) = (P x y \<and> list_all2 P xs ys)"
2534 by (auto simp add: list_all2_def)
2536 lemma list_all2_Cons1:
2537 "list_all2 P (x # xs) ys = (\<exists>z zs. ys = z # zs \<and> P x z \<and> list_all2 P xs zs)"
2540 lemma list_all2_Cons2:
2541 "list_all2 P xs (y # ys) = (\<exists>z zs. xs = z # zs \<and> P z y \<and> list_all2 P zs ys)"
2544 lemma list_all2_induct
2545 [consumes 1, case_names Nil Cons, induct set: list_all2]:
2546 assumes P: "list_all2 P xs ys"
2547 assumes Nil: "R [] []"
2548 assumes Cons: "\<And>x xs y ys.
2549 \<lbrakk>P x y; list_all2 P xs ys; R xs ys\<rbrakk> \<Longrightarrow> R (x # xs) (y # ys)"
2552 by (induct xs arbitrary: ys) (auto simp add: list_all2_Cons1 Nil Cons)
2554 lemma list_all2_rev [iff]:
2555 "list_all2 P (rev xs) (rev ys) = list_all2 P xs ys"
2556 by (simp add: list_all2_def zip_rev cong: conj_cong)
2558 lemma list_all2_rev1:
2559 "list_all2 P (rev xs) ys = list_all2 P xs (rev ys)"
2560 by (subst list_all2_rev [symmetric]) simp
2562 lemma list_all2_append1:
2563 "list_all2 P (xs @ ys) zs =
2564 (EX us vs. zs = us @ vs \<and> length us = length xs \<and> length vs = length ys \<and>
2565 list_all2 P xs us \<and> list_all2 P ys vs)"
2566 apply (simp add: list_all2_def zip_append1)
2568 apply (rule_tac x = "take (length xs) zs" in exI)
2569 apply (rule_tac x = "drop (length xs) zs" in exI)
2570 apply (force split: nat_diff_split simp add: min_def, clarify)
2571 apply (simp add: ball_Un)
2574 lemma list_all2_append2:
2575 "list_all2 P xs (ys @ zs) =
2576 (EX us vs. xs = us @ vs \<and> length us = length ys \<and> length vs = length zs \<and>
2577 list_all2 P us ys \<and> list_all2 P vs zs)"
2578 apply (simp add: list_all2_def zip_append2)
2580 apply (rule_tac x = "take (length ys) xs" in exI)
2581 apply (rule_tac x = "drop (length ys) xs" in exI)
2582 apply (force split: nat_diff_split simp add: min_def, clarify)
2583 apply (simp add: ball_Un)
2586 lemma list_all2_append:
2587 "length xs = length ys \<Longrightarrow>
2588 list_all2 P (xs@us) (ys@vs) = (list_all2 P xs ys \<and> list_all2 P us vs)"
2589 by (induct rule:list_induct2, simp_all)
2591 lemma list_all2_appendI [intro?, trans]:
2592 "\<lbrakk> list_all2 P a b; list_all2 P c d \<rbrakk> \<Longrightarrow> list_all2 P (a@c) (b@d)"
2593 by (simp add: list_all2_append list_all2_lengthD)
2595 lemma list_all2_conv_all_nth:
2596 "list_all2 P xs ys =
2597 (length xs = length ys \<and> (\<forall>i < length xs. P (xs!i) (ys!i)))"
2598 by (force simp add: list_all2_def set_zip)
2600 lemma list_all2_trans:
2601 assumes tr: "!!a b c. P1 a b ==> P2 b c ==> P3 a c"
2602 shows "!!bs cs. list_all2 P1 as bs ==> list_all2 P2 bs cs ==> list_all2 P3 as cs"
2603 (is "!!bs cs. PROP ?Q as bs cs")
2605 fix x xs bs assume I1: "!!bs cs. PROP ?Q xs bs cs"
2606 show "!!cs. PROP ?Q (x # xs) bs cs"
2608 fix y ys cs assume I2: "!!cs. PROP ?Q (x # xs) ys cs"
2609 show "PROP ?Q (x # xs) (y # ys) cs"
2610 by (induct cs) (auto intro: tr I1 I2)
2614 lemma list_all2_all_nthI [intro?]:
2615 "length a = length b \<Longrightarrow> (\<And>n. n < length a \<Longrightarrow> P (a!n) (b!n)) \<Longrightarrow> list_all2 P a b"
2616 by (simp add: list_all2_conv_all_nth)
2619 "\<forall>x \<in> set (zip a b). split P x \<Longrightarrow> length a = length b \<Longrightarrow> list_all2 P a b"
2620 by (simp add: list_all2_def)
2622 lemma list_all2_nthD:
2623 "\<lbrakk> list_all2 P xs ys; p < size xs \<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)"
2624 by (simp add: list_all2_conv_all_nth)
2626 lemma list_all2_nthD2:
2627 "\<lbrakk>list_all2 P xs ys; p < size ys\<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)"
2628 by (frule list_all2_lengthD) (auto intro: list_all2_nthD)
2630 lemma list_all2_map1:
2631 "list_all2 P (map f as) bs = list_all2 (\<lambda>x y. P (f x) y) as bs"
2632 by (simp add: list_all2_conv_all_nth)
2634 lemma list_all2_map2:
2635 "list_all2 P as (map f bs) = list_all2 (\<lambda>x y. P x (f y)) as bs"
2636 by (auto simp add: list_all2_conv_all_nth)
2638 lemma list_all2_refl [intro?]:
2639 "(\<And>x. P x x) \<Longrightarrow> list_all2 P xs xs"
2640 by (simp add: list_all2_conv_all_nth)
2642 lemma list_all2_update_cong:
2643 "\<lbrakk> list_all2 P xs ys; P x y \<rbrakk> \<Longrightarrow> list_all2 P (xs[i:=x]) (ys[i:=y])"
2644 by (cases "i < length ys") (auto simp add: list_all2_conv_all_nth nth_list_update)
2646 lemma list_all2_takeI [simp,intro?]:
2647 "list_all2 P xs ys \<Longrightarrow> list_all2 P (take n xs) (take n ys)"
2648 apply (induct xs arbitrary: n ys)
2650 apply (clarsimp simp add: list_all2_Cons1)
2655 lemma list_all2_dropI [simp,intro?]:
2656 "list_all2 P as bs \<Longrightarrow> list_all2 P (drop n as) (drop n bs)"
2657 apply (induct as arbitrary: n bs, simp)
2658 apply (clarsimp simp add: list_all2_Cons1)
2659 apply (case_tac n, simp, simp)
2662 lemma list_all2_mono [intro?]:
2663 "list_all2 P xs ys \<Longrightarrow> (\<And>xs ys. P xs ys \<Longrightarrow> Q xs ys) \<Longrightarrow> list_all2 Q xs ys"
2664 apply (induct xs arbitrary: ys, simp)
2665 apply (case_tac ys, auto)
2669 "xs = ys \<longleftrightarrow> list_all2 (op =) xs ys"
2670 by (induct xs ys rule: list_induct2') auto
2672 lemma list_eq_iff_zip_eq:
2673 "xs = ys \<longleftrightarrow> length xs = length ys \<and> (\<forall>(x,y) \<in> set (zip xs ys). x = y)"
2674 by(auto simp add: set_zip list_all2_eq list_all2_conv_all_nth cong: conj_cong)
2677 subsubsection {* @{const List.product} and @{const product_lists} *}
2679 lemma product_list_set:
2680 "set (List.product xs ys) = set xs \<times> set ys"
2683 lemma length_product [simp]:
2684 "length (List.product xs ys) = length xs * length ys"
2685 by (induct xs) simp_all
2688 assumes "n < length xs * length ys"
2689 shows "List.product xs ys ! n = (xs ! (n div length ys), ys ! (n mod length ys))"
2690 using assms proof (induct xs arbitrary: n)
2691 case Nil then show ?case by simp
2694 then have "length ys > 0" by auto
2695 with Cons show ?case
2696 by (auto simp add: nth_append not_less le_mod_geq le_div_geq)
2699 lemma in_set_product_lists_length:
2700 "xs \<in> set (product_lists xss) \<Longrightarrow> length xs = length xss"
2701 by (induct xss arbitrary: xs) auto
2703 lemma product_lists_set:
2704 "set (product_lists xss) = {xs. list_all2 (\<lambda>x ys. x \<in> set ys) xs xss}" (is "?L = Collect ?R")
2705 proof (intro equalityI subsetI, unfold mem_Collect_eq)
2706 fix xs assume "xs \<in> ?L"
2707 then have "length xs = length xss" by (rule in_set_product_lists_length)
2708 from this `xs \<in> ?L` show "?R xs" by (induct xs xss rule: list_induct2) auto
2710 fix xs assume "?R xs"
2711 then show "xs \<in> ?L" by induct auto
2715 subsubsection {* @{const fold} with natural argument order *}
2717 lemma fold_simps [code]: -- {* eta-expanded variant for generated code -- enables tail-recursion optimisation in Scala *}
2719 "fold f (x # xs) s = fold f xs (f x s)"
2722 lemma fold_remove1_split:
2723 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"
2724 and x: "x \<in> set xs"
2725 shows "fold f xs = fold f (remove1 x xs) \<circ> f x"
2726 using assms by (induct xs) (auto simp add: comp_assoc)
2728 lemma fold_cong [fundef_cong]:
2729 "a = b \<Longrightarrow> xs = ys \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> f x = g x)
2730 \<Longrightarrow> fold f xs a = fold g ys b"
2731 by (induct ys arbitrary: a b xs) simp_all
2734 assumes "\<And>x. x \<in> set xs \<Longrightarrow> f x = id"
2735 shows "fold f xs = id"
2736 using assms by (induct xs) simp_all
2739 assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<circ> g x = f x \<circ> h"
2740 shows "h \<circ> fold g xs = fold f xs \<circ> h"
2741 using assms by (induct xs) (simp_all add: fun_eq_iff)
2743 lemma fold_commute_apply:
2744 assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<circ> g x = f x \<circ> h"
2745 shows "h (fold g xs s) = fold f xs (h s)"
2747 from assms have "h \<circ> fold g xs = fold f xs \<circ> h" by (rule fold_commute)
2748 then show ?thesis by (simp add: fun_eq_iff)
2751 lemma fold_invariant:
2752 assumes "\<And>x. x \<in> set xs \<Longrightarrow> Q x" and "P s"
2753 and "\<And>x s. Q x \<Longrightarrow> P s \<Longrightarrow> P (f x s)"
2754 shows "P (fold f xs s)"
2755 using assms by (induct xs arbitrary: s) simp_all
2757 lemma fold_append [simp]:
2758 "fold f (xs @ ys) = fold f ys \<circ> fold f xs"
2759 by (induct xs) simp_all
2761 lemma fold_map [code_unfold]:
2762 "fold g (map f xs) = fold (g o f) xs"
2763 by (induct xs) simp_all
2766 assumes "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y"
2767 shows "fold f (rev xs) = fold f xs"
2768 using assms by (induct xs) (simp_all add: fold_commute_apply fun_eq_iff)
2770 lemma fold_Cons_rev:
2771 "fold Cons xs = append (rev xs)"
2772 by (induct xs) simp_all
2774 lemma rev_conv_fold [code]:
2775 "rev xs = fold Cons xs []"
2776 by (simp add: fold_Cons_rev)
2778 lemma fold_append_concat_rev:
2779 "fold append xss = append (concat (rev xss))"
2780 by (induct xss) simp_all
2782 text {* @{const Finite_Set.fold} and @{const fold} *}
2784 lemma (in comp_fun_commute) fold_set_fold_remdups:
2785 "Finite_Set.fold f y (set xs) = fold f (remdups xs) y"
2786 by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm insert_absorb)
2788 lemma (in comp_fun_idem) fold_set_fold:
2789 "Finite_Set.fold f y (set xs) = fold f xs y"
2790 by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm)
2792 lemma union_set_fold [code]:
2793 "set xs \<union> A = fold Set.insert xs A"
2795 interpret comp_fun_idem Set.insert
2796 by (fact comp_fun_idem_insert)
2797 show ?thesis by (simp add: union_fold_insert fold_set_fold)
2800 lemma union_coset_filter [code]:
2801 "List.coset xs \<union> A = List.coset (List.filter (\<lambda>x. x \<notin> A) xs)"
2804 lemma minus_set_fold [code]:
2805 "A - set xs = fold Set.remove xs A"
2807 interpret comp_fun_idem Set.remove
2808 by (fact comp_fun_idem_remove)
2810 by (simp add: minus_fold_remove [of _ A] fold_set_fold)
2813 lemma minus_coset_filter [code]:
2814 "A - List.coset xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
2817 lemma inter_set_filter [code]:
2818 "A \<inter> set xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
2821 lemma inter_coset_fold [code]:
2822 "A \<inter> List.coset xs = fold Set.remove xs A"
2823 by (simp add: Diff_eq [symmetric] minus_set_fold)
2825 lemma (in semilattice_set) set_eq_fold [code]:
2826 "F (set (x # xs)) = fold f xs x"
2828 interpret comp_fun_idem f
2829 by default (simp_all add: fun_eq_iff left_commute)
2830 show ?thesis by (simp add: eq_fold fold_set_fold)
2833 lemma (in complete_lattice) Inf_set_fold:
2834 "Inf (set xs) = fold inf xs top"
2836 interpret comp_fun_idem "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
2837 by (fact comp_fun_idem_inf)
2838 show ?thesis by (simp add: Inf_fold_inf fold_set_fold inf_commute)
2841 declare Inf_set_fold [where 'a = "'a set", code]
2843 lemma (in complete_lattice) Sup_set_fold:
2844 "Sup (set xs) = fold sup xs bot"
2846 interpret comp_fun_idem "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
2847 by (fact comp_fun_idem_sup)
2848 show ?thesis by (simp add: Sup_fold_sup fold_set_fold sup_commute)
2851 declare Sup_set_fold [where 'a = "'a set", code]
2853 lemma (in complete_lattice) INF_set_fold:
2854 "INFI (set xs) f = fold (inf \<circ> f) xs top"
2855 unfolding INF_def set_map [symmetric] Inf_set_fold fold_map ..
2857 declare INF_set_fold [code]
2859 lemma (in complete_lattice) SUP_set_fold:
2860 "SUPR (set xs) f = fold (sup \<circ> f) xs bot"
2861 unfolding SUP_def set_map [symmetric] Sup_set_fold fold_map ..
2863 declare SUP_set_fold [code]
2866 subsubsection {* Fold variants: @{const foldr} and @{const foldl} *}
2868 text {* Correspondence *}
2870 lemma foldr_conv_fold [code_abbrev]:
2871 "foldr f xs = fold f (rev xs)"
2872 by (induct xs) simp_all
2874 lemma foldl_conv_fold:
2875 "foldl f s xs = fold (\<lambda>x s. f s x) xs s"
2876 by (induct xs arbitrary: s) simp_all
2878 lemma foldr_conv_foldl: -- {* The ``Third Duality Theorem'' in Bird \& Wadler: *}
2879 "foldr f xs a = foldl (\<lambda>x y. f y x) a (rev xs)"
2880 by (simp add: foldr_conv_fold foldl_conv_fold)
2882 lemma foldl_conv_foldr:
2883 "foldl f a xs = foldr (\<lambda>x y. f y x) (rev xs) a"
2884 by (simp add: foldr_conv_fold foldl_conv_fold)
2887 assumes "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y"
2888 shows "foldr f xs = fold f xs"
2889 using assms unfolding foldr_conv_fold by (rule fold_rev)
2891 lemma foldr_cong [fundef_cong]:
2892 "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"
2893 by (auto simp add: foldr_conv_fold intro!: fold_cong)
2895 lemma foldl_cong [fundef_cong]:
2896 "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"
2897 by (auto simp add: foldl_conv_fold intro!: fold_cong)
2899 lemma foldr_append [simp]:
2900 "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)"
2901 by (simp add: foldr_conv_fold)
2903 lemma foldl_append [simp]:
2904 "foldl f a (xs @ ys) = foldl f (foldl f a xs) ys"
2905 by (simp add: foldl_conv_fold)
2907 lemma foldr_map [code_unfold]:
2908 "foldr g (map f xs) a = foldr (g o f) xs a"
2909 by (simp add: foldr_conv_fold fold_map rev_map)
2911 lemma foldl_map [code_unfold]:
2912 "foldl g a (map f xs) = foldl (\<lambda>a x. g a (f x)) a xs"
2913 by (simp add: foldl_conv_fold fold_map comp_def)
2915 lemma concat_conv_foldr [code]:
2916 "concat xss = foldr append xss []"
2917 by (simp add: fold_append_concat_rev foldr_conv_fold)
2920 subsubsection {* @{const upt} *}
2922 lemma upt_rec[code]: "[i..<j] = (if i<j then i#[Suc i..<j] else [])"
2923 -- {* simp does not terminate! *}
2926 lemmas upt_rec_numeral[simp] = upt_rec[of "numeral m" "numeral n"] for m n
2928 lemma upt_conv_Nil [simp]: "j <= i ==> [i..<j] = []"
2929 by (subst upt_rec) simp
2931 lemma upt_eq_Nil_conv[simp]: "([i..<j] = []) = (j = 0 \<or> j <= i)"
2932 by(induct j)simp_all
2934 lemma upt_eq_Cons_conv:
2935 "([i..<j] = x#xs) = (i < j & i = x & [i+1..<j] = xs)"
2936 apply(induct j arbitrary: x xs)
2938 apply(clarsimp simp add: append_eq_Cons_conv)
2942 lemma upt_Suc_append: "i <= j ==> [i..<(Suc j)] = [i..<j]@[j]"
2943 -- {* Only needed if @{text upt_Suc} is deleted from the simpset. *}
2946 lemma upt_conv_Cons: "i < j ==> [i..<j] = i # [Suc i..<j]"
2947 by (simp add: upt_rec)
2949 lemma upt_add_eq_append: "i<=j ==> [i..<j+k] = [i..<j]@[j..<j+k]"
2950 -- {* LOOPS as a simprule, since @{text "j <= j"}. *}
2953 lemma length_upt [simp]: "length [i..<j] = j - i"
2954 by (induct j) (auto simp add: Suc_diff_le)
2956 lemma nth_upt [simp]: "i + k < j ==> [i..<j] ! k = i + k"
2958 apply (auto simp add: less_Suc_eq nth_append split: nat_diff_split)
2962 lemma hd_upt[simp]: "i < j \<Longrightarrow> hd[i..<j] = i"
2963 by(simp add:upt_conv_Cons)
2965 lemma last_upt[simp]: "i < j \<Longrightarrow> last[i..<j] = j - 1"
2968 by(simp add:upt_Suc_append)
2970 lemma take_upt [simp]: "i+m <= n ==> take m [i..<n] = [i..<i+m]"
2971 apply (induct m arbitrary: i, simp)
2972 apply (subst upt_rec)
2974 apply (subst upt_rec)
2975 apply (simp del: upt.simps)
2978 lemma drop_upt[simp]: "drop m [i..<j] = [i+m..<j]"
2983 lemma map_Suc_upt: "map Suc [m..<n] = [Suc m..<Suc n]"
2986 lemma map_add_upt: "map (\<lambda>i. i + n) [0..<m] = [n..<m + n]"
2987 by (induct m) simp_all
2989 lemma nth_map_upt: "i < n-m ==> (map f [m..<n]) ! i = f(m+i)"
2990 apply (induct n m arbitrary: i rule: diff_induct)
2991 prefer 3 apply (subst map_Suc_upt[symmetric])
2992 apply (auto simp add: less_diff_conv)
2996 "map (\<lambda>n. n - Suc 0) [Suc m..<Suc n] = [m..<n]"
2997 by (induct n) simp_all
2999 lemma nth_take_lemma:
3000 "k <= length xs ==> k <= length ys ==>
3001 (!!i. i < k --> xs!i = ys!i) ==> take k xs = take k ys"
3002 apply (atomize, induct k arbitrary: xs ys)
3003 apply (simp_all add: less_Suc_eq_0_disj all_conj_distrib, clarify)
3004 txt {* Both lists must be non-empty *}
3005 apply (case_tac xs, simp)
3006 apply (case_tac ys, clarify)
3007 apply (simp (no_asm_use))
3009 txt {* prenexing's needed, not miniscoping *}
3010 apply (simp (no_asm_use) add: all_simps [symmetric] del: all_simps)
3014 lemma nth_equalityI:
3015 "[| length xs = length ys; ALL i < length xs. xs!i = ys!i |] ==> xs = ys"
3016 by (frule nth_take_lemma [OF le_refl eq_imp_le]) simp_all
3019 "map (\<lambda>i. xs ! i) [0..<length xs] = xs"
3020 by (rule nth_equalityI, auto)
3022 (* needs nth_equalityI *)
3023 lemma list_all2_antisym:
3024 "\<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>
3025 \<Longrightarrow> xs = ys"
3026 apply (simp add: list_all2_conv_all_nth)
3027 apply (rule nth_equalityI, blast, simp)
3030 lemma take_equalityI: "(\<forall>i. take i xs = take i ys) ==> xs = ys"
3031 -- {* The famous take-lemma. *}
3032 apply (drule_tac x = "max (length xs) (length ys)" in spec)
3033 apply (simp add: le_max_iff_disj)
3038 "take n (x # xs) = (if n = 0 then [] else x # take (n - 1) xs)"
3039 by (cases n) simp_all
3042 "drop n (x # xs) = (if n = 0 then x # xs else drop (n - 1) xs)"
3043 by (cases n) simp_all
3045 lemma nth_Cons': "(x # xs)!n = (if n = 0 then x else xs!(n - 1))"
3046 by (cases n) simp_all
3048 lemma take_Cons_numeral [simp]:
3049 "take (numeral v) (x # xs) = x # take (numeral v - 1) xs"
3050 by (simp add: take_Cons')
3052 lemma drop_Cons_numeral [simp]:
3053 "drop (numeral v) (x # xs) = drop (numeral v - 1) xs"
3054 by (simp add: drop_Cons')
3056 lemma nth_Cons_numeral [simp]:
3057 "(x # xs) ! numeral v = xs ! (numeral v - 1)"
3058 by (simp add: nth_Cons')
3061 subsubsection {* @{text upto}: interval-list on @{typ int} *}
3063 function upto :: "int \<Rightarrow> int \<Rightarrow> int list" ("(1[_../_])") where
3064 "upto i j = (if i \<le> j then i # [i+1..j] else [])"
3067 by(relation "measure(%(i::int,j). nat(j - i + 1))") auto
3069 declare upto.simps[simp del]
3071 lemmas upto_rec_numeral [simp] =
3072 upto.simps[of "numeral m" "numeral n"]
3073 upto.simps[of "numeral m" "- numeral n"]
3074 upto.simps[of "- numeral m" "numeral n"]
3075 upto.simps[of "- numeral m" "- numeral n"] for m n
3077 lemma upto_empty[simp]: "j < i \<Longrightarrow> [i..j] = []"
3078 by(simp add: upto.simps)
3080 lemma upto_rec1: "i \<le> j \<Longrightarrow> [i..j] = i#[i+1..j]"
3081 by(simp add: upto.simps)
3083 lemma upto_rec2: "i \<le> j \<Longrightarrow> [i..j] = [i..j - 1]@[j]"
3084 proof(induct "nat(j-i)" arbitrary: i j)
3085 case 0 thus ?case by(simp add: upto.simps)
3088 hence "n = nat (j - (i + 1))" "i < j" by linarith+
3089 from this(2) Suc.hyps(1)[OF this(1)] Suc(2,3) upto_rec1 show ?case by simp
3092 lemma set_upto[simp]: "set[i..j] = {i..j}"
3093 proof(induct i j rule:upto.induct)
3095 from this show ?case
3096 unfolding upto.simps[of i j] simp_from_to[of i j] by auto
3099 text{* Tail recursive version for code generation: *}
3101 definition upto_aux :: "int \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" where
3102 "upto_aux i j js = [i..j] @ js"
3104 lemma upto_aux_rec [code]:
3105 "upto_aux i j js = (if j<i then js else upto_aux i (j - 1) (j#js))"
3106 by (simp add: upto_aux_def upto_rec2)
3108 lemma upto_code[code]: "[i..j] = upto_aux i j []"
3109 by(simp add: upto_aux_def)
3112 subsubsection {* @{const distinct} and @{const remdups} and @{const remdups_adj} *}
3115 "distinct xs \<Longrightarrow> distinct (tl xs)"
3116 by (cases xs) simp_all
3118 lemma distinct_append [simp]:
3119 "distinct (xs @ ys) = (distinct xs \<and> distinct ys \<and> set xs \<inter> set ys = {})"
3122 lemma distinct_rev[simp]: "distinct(rev xs) = distinct xs"
3125 lemma set_remdups [simp]: "set (remdups xs) = set xs"
3126 by (induct xs) (auto simp add: insert_absorb)
3128 lemma distinct_remdups [iff]: "distinct (remdups xs)"
3131 lemma distinct_remdups_id: "distinct xs ==> remdups xs = xs"
3132 by (induct xs, auto)
3134 lemma remdups_id_iff_distinct [simp]: "remdups xs = xs \<longleftrightarrow> distinct xs"
3135 by (metis distinct_remdups distinct_remdups_id)
3137 lemma finite_distinct_list: "finite A \<Longrightarrow> EX xs. set xs = A & distinct xs"
3138 by (metis distinct_remdups finite_list set_remdups)
3140 lemma remdups_eq_nil_iff [simp]: "(remdups x = []) = (x = [])"
3143 lemma remdups_eq_nil_right_iff [simp]: "([] = remdups x) = (x = [])"
3146 lemma length_remdups_leq[iff]: "length(remdups xs) <= length xs"
3149 lemma length_remdups_eq[iff]:
3150 "(length (remdups xs) = length xs) = (remdups xs = xs)"
3153 apply(subgoal_tac "length (remdups xs) <= length xs")
3155 apply(rule length_remdups_leq)
3158 lemma remdups_filter: "remdups(filter P xs) = filter P (remdups xs)"
3164 "distinct(map f xs) = (distinct xs & inj_on f (set xs))"
3167 lemma distinct_filter [simp]: "distinct xs ==> distinct (filter P xs)"
3170 lemma distinct_upt[simp]: "distinct[i..<j]"
3173 lemma distinct_upto[simp]: "distinct[i..j]"
3174 apply(induct i j rule:upto.induct)
3175 apply(subst upto.simps)
3179 lemma distinct_take[simp]: "distinct xs \<Longrightarrow> distinct (take i xs)"
3180 apply(induct xs arbitrary: i)
3184 apply(blast dest:in_set_takeD)
3187 lemma distinct_drop[simp]: "distinct xs \<Longrightarrow> distinct (drop i xs)"
3188 apply(induct xs arbitrary: i)
3194 lemma distinct_list_update:
3195 assumes d: "distinct xs" and a: "a \<notin> set xs - {xs!i}"
3196 shows "distinct (xs[i:=a])"
3197 proof (cases "i < length xs")
3199 with a have "a \<notin> set (take i xs @ xs ! i # drop (Suc i) xs) - {xs!i}"
3200 apply (drule_tac id_take_nth_drop) by simp
3201 with d True show ?thesis
3202 apply (simp add: upd_conv_take_nth_drop)
3203 apply (drule subst [OF id_take_nth_drop]) apply assumption
3204 apply simp apply (cases "a = xs!i") apply simp by blast
3206 case False with d show ?thesis by auto
3209 lemma distinct_concat:
3210 assumes "distinct xs"
3211 and "\<And> ys. ys \<in> set xs \<Longrightarrow> distinct ys"
3212 and "\<And> ys zs. \<lbrakk> ys \<in> set xs ; zs \<in> set xs ; ys \<noteq> zs \<rbrakk> \<Longrightarrow> set ys \<inter> set zs = {}"
3213 shows "distinct (concat xs)"
3214 using assms by (induct xs) auto
3216 text {* It is best to avoid this indexed version of distinct, but
3217 sometimes it is useful. *}
3219 lemma distinct_conv_nth:
3220 "distinct xs = (\<forall>i < size xs. \<forall>j < size xs. i \<noteq> j --> xs!i \<noteq> xs!j)"
3221 apply (induct xs, simp, simp)
3222 apply (rule iffI, clarsimp)
3224 apply (case_tac j, simp)
3225 apply (simp add: set_conv_nth)
3227 apply (clarsimp simp add: set_conv_nth, simp)
3230 apply (metis Zero_neq_Suc gr0_conv_Suc in_set_conv_nth lessI less_trans_Suc nth_Cons' nth_Cons_Suc)
3232 apply (clarsimp simp add: set_conv_nth)
3233 apply (erule_tac x = 0 in allE, simp)
3234 apply (erule_tac x = "Suc i" in allE, simp, clarsimp)
3236 apply (metis Suc_Suc_eq lessI less_trans_Suc nth_Cons_Suc)
3238 apply (erule_tac x = "Suc i" in allE, simp)
3239 apply (erule_tac x = "Suc j" in allE, simp)
3242 lemma nth_eq_iff_index_eq:
3243 "\<lbrakk> distinct xs; i < length xs; j < length xs \<rbrakk> \<Longrightarrow> (xs!i = xs!j) = (i = j)"
3244 by(auto simp: distinct_conv_nth)
3246 lemma distinct_card: "distinct xs ==> card (set xs) = size xs"
3249 lemma card_distinct: "card (set xs) = size xs ==> distinct xs"
3251 case Nil thus ?case by simp
3255 proof (cases "x \<in> set xs")
3256 case False with Cons show ?thesis by simp
3258 case True with Cons.prems
3259 have "card (set xs) = Suc (length xs)"
3260 by (simp add: card_insert_if split: split_if_asm)
3261 moreover have "card (set xs) \<le> length xs" by (rule card_length)
3262 ultimately have False by simp
3267 lemma distinct_length_filter: "distinct xs \<Longrightarrow> length (filter P xs) = card ({x. P x} Int set xs)"
3268 by (induct xs) (auto)
3270 lemma not_distinct_decomp: "~ distinct ws ==> EX xs ys zs y. ws = xs@[y]@ys@[y]@zs"
3271 apply (induct n == "length ws" arbitrary:ws) apply simp
3272 apply(case_tac ws) apply simp
3273 apply (simp split:split_if_asm)
3274 apply (metis Cons_eq_appendI eq_Nil_appendI split_list)
3277 lemma not_distinct_conv_prefix:
3278 defines "dec as xs y ys \<equiv> y \<in> set xs \<and> distinct xs \<and> as = xs @ y # ys"
3279 shows "\<not>distinct as \<longleftrightarrow> (\<exists>xs y ys. dec as xs y ys)" (is "?L = ?R")
3281 assume "?L" then show "?R"
3282 proof (induct "length as" arbitrary: as rule: less_induct)
3284 obtain xs ys zs y where decomp: "as = (xs @ y # ys) @ y # zs"
3285 using not_distinct_decomp[OF less.prems] by auto
3287 proof (cases "distinct (xs @ y # ys)")
3289 with decomp have "dec as (xs @ y # ys) y zs" by (simp add: dec_def)
3290 then show ?thesis by blast
3293 with less decomp obtain xs' y' ys' where "dec (xs @ y # ys) xs' y' ys'"
3294 by atomize_elim auto
3295 with decomp have "dec as xs' y' (ys' @ y # zs)" by (simp add: dec_def)
3296 then show ?thesis by blast
3299 qed (auto simp: dec_def)
3301 lemma distinct_product:
3302 assumes "distinct xs" and "distinct ys"
3303 shows "distinct (List.product xs ys)"
3304 using assms by (induct xs)
3305 (auto intro: inj_onI simp add: product_list_set distinct_map)
3307 lemma distinct_product_lists:
3308 assumes "\<forall>xs \<in> set xss. distinct xs"
3309 shows "distinct (product_lists xss)"
3310 using assms proof (induction xss)
3311 case (Cons xs xss) note * = this
3313 proof (cases "product_lists xss")
3314 case Nil then show ?thesis by (induct xs) simp_all
3316 case (Cons ps pss) with * show ?thesis
3317 by (auto intro!: inj_onI distinct_concat simp add: distinct_map)
3321 lemma length_remdups_concat:
3322 "length (remdups (concat xss)) = card (\<Union>xs\<in>set xss. set xs)"
3323 by (simp add: distinct_card [symmetric])
3325 lemma length_remdups_card_conv: "length(remdups xs) = card(set xs)"
3327 have xs: "concat[xs] = xs" by simp
3328 from length_remdups_concat[of "[xs]"] show ?thesis unfolding xs by simp
3331 lemma remdups_remdups:
3332 "remdups (remdups xs) = remdups xs"
3333 by (induct xs) simp_all
3335 lemma distinct_butlast:
3336 assumes "distinct xs"
3337 shows "distinct (butlast xs)"
3338 proof (cases "xs = []")
3340 from `xs \<noteq> []` obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto
3341 with `distinct xs` show ?thesis by simp
3344 lemma remdups_map_remdups:
3345 "remdups (map f (remdups xs)) = remdups (map f xs)"
3346 by (induct xs) simp_all
3348 lemma distinct_zipI1:
3349 assumes "distinct xs"
3350 shows "distinct (zip xs ys)"
3351 proof (rule zip_obtain_same_length)
3352 fix xs' :: "'a list" and ys' :: "'b list" and n
3353 assume "length xs' = length ys'"
3354 assume "xs' = take n xs"
3355 with assms have "distinct xs'" by simp
3356 with `length xs' = length ys'` show "distinct (zip xs' ys')"
3357 by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE)
3360 lemma distinct_zipI2:
3361 assumes "distinct ys"
3362 shows "distinct (zip xs ys)"
3363 proof (rule zip_obtain_same_length)
3364 fix xs' :: "'b list" and ys' :: "'a list" and n
3365 assume "length xs' = length ys'"
3366 assume "ys' = take n ys"
3367 with assms have "distinct ys'" by simp
3368 with `length xs' = length ys'` show "distinct (zip xs' ys')"
3369 by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE)
3372 lemma set_take_disj_set_drop_if_distinct:
3373 "distinct vs \<Longrightarrow> i \<le> j \<Longrightarrow> set (take i vs) \<inter> set (drop j vs) = {}"
3374 by (auto simp: in_set_conv_nth distinct_conv_nth)
3376 (* The next two lemmas help Sledgehammer. *)
3378 lemma distinct_singleton: "distinct [x]" by simp
3380 lemma distinct_length_2_or_more:
3381 "distinct (a # b # xs) \<longleftrightarrow> (a \<noteq> b \<and> distinct (a # xs) \<and> distinct (b # xs))"
3382 by (metis distinct.simps(2) hd.simps hd_in_set list.simps(2) set_ConsD set_rev_mp set_subset_Cons)
3384 lemma remdups_adj_Cons: "remdups_adj (x # xs) =
3385 (case remdups_adj xs of [] \<Rightarrow> [x] | y # xs \<Rightarrow> if x = y then y # xs else x # y # xs)"
3386 by (induct xs arbitrary: x) (auto split: list.splits)
3388 lemma remdups_adj_append_two:
3389 "remdups_adj (xs @ [x,y]) = remdups_adj (xs @ [x]) @ (if x = y then [] else [y])"
3390 by (induct xs rule: remdups_adj.induct, simp_all)
3392 lemma remdups_adj_rev[simp]: "remdups_adj (rev xs) = rev (remdups_adj xs)"
3393 by (induct xs rule: remdups_adj.induct, simp_all add: remdups_adj_append_two)
3395 lemma remdups_adj_length[simp]: "length (remdups_adj xs) \<le> length xs"
3396 by (induct xs rule: remdups_adj.induct, auto)
3398 lemma remdups_adj_length_ge1[simp]: "xs \<noteq> [] \<Longrightarrow> length (remdups_adj xs) \<ge> Suc 0"
3399 by (induct xs rule: remdups_adj.induct, simp_all)
3401 lemma remdups_adj_Nil_iff[simp]: "remdups_adj xs = [] \<longleftrightarrow> xs = []"
3402 by (induct xs rule: remdups_adj.induct, simp_all)
3404 lemma remdups_adj_set[simp]: "set (remdups_adj xs) = set xs"
3405 by (induct xs rule: remdups_adj.induct, simp_all)
3407 lemma remdups_adj_Cons_alt[simp]: "x # tl (remdups_adj (x # xs)) = remdups_adj (x # xs)"
3408 by (induct xs rule: remdups_adj.induct, auto)
3410 lemma remdups_adj_distinct: "distinct xs \<Longrightarrow> remdups_adj xs = xs"
3411 by (induct xs rule: remdups_adj.induct, simp_all)
3413 lemma remdups_adj_append:
3414 "remdups_adj (xs\<^sub>1 @ x # xs\<^sub>2) = remdups_adj (xs\<^sub>1 @ [x]) @ tl (remdups_adj (x # xs\<^sub>2))"
3415 by (induct xs\<^sub>1 rule: remdups_adj.induct, simp_all)
3417 lemma remdups_adj_singleton:
3418 "remdups_adj xs = [x] \<Longrightarrow> xs = replicate (length xs) x"
3419 by (induct xs rule: remdups_adj.induct, auto split: split_if_asm)
3421 lemma remdups_adj_map_injective:
3423 shows "remdups_adj (map f xs) = map f (remdups_adj xs)"
3424 by (induct xs rule: remdups_adj.induct,
3425 auto simp add: injD[OF assms])
3428 subsubsection {* List summation: @{const listsum} and @{text"\<Sum>"}*}
3430 lemma (in monoid_add) listsum_simps [simp]:
3432 "listsum (x # xs) = x + listsum xs"
3433 by (simp_all add: listsum_def)
3435 lemma (in monoid_add) listsum_append [simp]:
3436 "listsum (xs @ ys) = listsum xs + listsum ys"
3437 by (induct xs) (simp_all add: add.assoc)
3439 lemma (in comm_monoid_add) listsum_rev [simp]:
3440 "listsum (rev xs) = listsum xs"
3441 by (simp add: listsum_def foldr_fold fold_rev fun_eq_iff add_ac)
3443 lemma (in monoid_add) fold_plus_listsum_rev:
3444 "fold plus xs = plus (listsum (rev xs))"
3447 have "fold plus xs x = fold plus xs (x + 0)" by simp
3448 also have "\<dots> = fold plus (x # xs) 0" by simp
3449 also have "\<dots> = foldr plus (rev xs @ [x]) 0" by (simp add: foldr_conv_fold)
3450 also have "\<dots> = listsum (rev xs @ [x])" by (simp add: listsum_def)
3451 also have "\<dots> = listsum (rev xs) + listsum [x]" by simp
3452 finally show "fold plus xs x = listsum (rev xs) + x" by simp
3455 text{* Some syntactic sugar for summing a function over a list: *}
3458 "_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3SUM _<-_. _)" [0, 51, 10] 10)
3460 "_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3\<Sum>_\<leftarrow>_. _)" [0, 51, 10] 10)
3461 syntax (HTML output)
3462 "_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3\<Sum>_\<leftarrow>_. _)" [0, 51, 10] 10)
3464 translations -- {* Beware of argument permutation! *}
3465 "SUM x<-xs. b" == "CONST listsum (CONST map (%x. b) xs)"
3466 "\<Sum>x\<leftarrow>xs. b" == "CONST listsum (CONST map (%x. b) xs)"
3468 lemma (in comm_monoid_add) listsum_map_remove1:
3469 "x \<in> set xs \<Longrightarrow> listsum (map f xs) = f x + listsum (map f (remove1 x xs))"
3470 by (induct xs) (auto simp add: ac_simps)
3472 lemma (in monoid_add) list_size_conv_listsum:
3473 "list_size f xs = listsum (map f xs) + size xs"
3476 lemma (in monoid_add) length_concat:
3477 "length (concat xss) = listsum (map length xss)"
3478 by (induct xss) simp_all
3480 lemma (in monoid_add) length_product_lists:
3481 "length (product_lists xss) = foldr op * (map length xss) 1"
3483 case (Cons xs xss) then show ?case by (induct xs) (auto simp: length_concat o_def)
3486 lemma (in monoid_add) listsum_map_filter:
3487 assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<not> P x \<Longrightarrow> f x = 0"
3488 shows "listsum (map f (filter P xs)) = listsum (map f xs)"
3489 using assms by (induct xs) auto
3491 lemma (in comm_monoid_add) distinct_listsum_conv_Setsum:
3492 "distinct xs \<Longrightarrow> listsum xs = Setsum (set xs)"
3493 by (induct xs) simp_all
3495 lemma listsum_eq_0_nat_iff_nat [simp]:
3496 "listsum ns = (0::nat) \<longleftrightarrow> (\<forall>n \<in> set ns. n = 0)"
3497 by (induct ns) simp_all
3499 lemma member_le_listsum_nat:
3500 "(n :: nat) \<in> set ns \<Longrightarrow> n \<le> listsum ns"
3503 lemma elem_le_listsum_nat:
3504 "k < size ns \<Longrightarrow> ns ! k \<le> listsum (ns::nat list)"
3505 by (rule member_le_listsum_nat) simp
3507 lemma listsum_update_nat:
3508 "k < size ns \<Longrightarrow> listsum (ns[k := (n::nat)]) = listsum ns + n - ns ! k"
3509 apply(induct ns arbitrary:k)
3510 apply (auto split:nat.split)
3511 apply(drule elem_le_listsum_nat)
3515 lemma (in monoid_add) listsum_triv:
3516 "(\<Sum>x\<leftarrow>xs. r) = of_nat (length xs) * r"
3517 by (induct xs) (simp_all add: distrib_right)
3519 lemma (in monoid_add) listsum_0 [simp]:
3520 "(\<Sum>x\<leftarrow>xs. 0) = 0"
3521 by (induct xs) (simp_all add: distrib_right)
3523 text{* For non-Abelian groups @{text xs} needs to be reversed on one side: *}
3524 lemma (in ab_group_add) uminus_listsum_map:
3525 "- listsum (map f xs) = listsum (map (uminus \<circ> f) xs)"
3526 by (induct xs) simp_all
3528 lemma (in comm_monoid_add) listsum_addf:
3529 "(\<Sum>x\<leftarrow>xs. f x + g x) = listsum (map f xs) + listsum (map g xs)"
3530 by (induct xs) (simp_all add: algebra_simps)
3532 lemma (in ab_group_add) listsum_subtractf:
3533 "(\<Sum>x\<leftarrow>xs. f x - g x) = listsum (map f xs) - listsum (map g xs)"
3534 by (induct xs) (simp_all add: algebra_simps)
3536 lemma (in semiring_0) listsum_const_mult:
3537 "(\<Sum>x\<leftarrow>xs. c * f x) = c * (\<Sum>x\<leftarrow>xs. f x)"
3538 by (induct xs) (simp_all add: algebra_simps)
3540 lemma (in semiring_0) listsum_mult_const:
3541 "(\<Sum>x\<leftarrow>xs. f x * c) = (\<Sum>x\<leftarrow>xs. f x) * c"
3542 by (induct xs) (simp_all add: algebra_simps)
3544 lemma (in ordered_ab_group_add_abs) listsum_abs:
3545 "\<bar>listsum xs\<bar> \<le> listsum (map abs xs)"
3546 by (induct xs) (simp_all add: order_trans [OF abs_triangle_ineq])
3549 fixes f g :: "'a \<Rightarrow> 'b::{monoid_add, ordered_ab_semigroup_add}"
3550 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)"
3551 by (induct xs) (simp, simp add: add_mono)
3553 lemma (in monoid_add) listsum_distinct_conv_setsum_set:
3554 "distinct xs \<Longrightarrow> listsum (map f xs) = setsum f (set xs)"
3555 by (induct xs) simp_all
3557 lemma (in monoid_add) interv_listsum_conv_setsum_set_nat:
3558 "listsum (map f [m..<n]) = setsum f (set [m..<n])"
3559 by (simp add: listsum_distinct_conv_setsum_set)
3561 lemma (in monoid_add) interv_listsum_conv_setsum_set_int:
3562 "listsum (map f [k..l]) = setsum f (set [k..l])"
3563 by (simp add: listsum_distinct_conv_setsum_set)
3565 text {* General equivalence between @{const listsum} and @{const setsum} *}
3566 lemma (in monoid_add) listsum_setsum_nth:
3567 "listsum xs = (\<Sum> i = 0 ..< length xs. xs ! i)"
3568 using interv_listsum_conv_setsum_set_nat [of "op ! xs" 0 "length xs"] by (simp add: map_nth)
3571 subsubsection {* @{const insert} *}
3573 lemma in_set_insert [simp]:
3574 "x \<in> set xs \<Longrightarrow> List.insert x xs = xs"
3575 by (simp add: List.insert_def)
3577 lemma not_in_set_insert [simp]:
3578 "x \<notin> set xs \<Longrightarrow> List.insert x xs = x # xs"
3579 by (simp add: List.insert_def)
3581 lemma insert_Nil [simp]:
3582 "List.insert x [] = [x]"
3585 lemma set_insert [simp]:
3586 "set (List.insert x xs) = insert x (set xs)"
3587 by (auto simp add: List.insert_def)
3589 lemma distinct_insert [simp]:
3590 "distinct xs \<Longrightarrow> distinct (List.insert x xs)"
3591 by (simp add: List.insert_def)
3593 lemma insert_remdups:
3594 "List.insert x (remdups xs) = remdups (List.insert x xs)"
3595 by (simp add: List.insert_def)
3598 subsubsection {* @{const List.find} *}
3600 lemma find_None_iff: "List.find P xs = None \<longleftrightarrow> \<not> (\<exists>x. x \<in> set xs \<and> P x)"
3601 proof (induction xs)
3602 case Nil thus ?case by simp
3604 case (Cons x xs) thus ?case by (fastforce split: if_splits)
3607 lemma find_Some_iff:
3608 "List.find P xs = Some x \<longleftrightarrow>
3609 (\<exists>i<length xs. P (xs!i) \<and> x = xs!i \<and> (\<forall>j<i. \<not> P (xs!j)))"
3610 proof (induction xs)
3611 case Nil thus ?case by simp
3613 case (Cons x xs) thus ?case
3614 by(auto simp: nth_Cons' split: if_splits)
3615 (metis One_nat_def diff_Suc_1 less_Suc_eq_0_disj)
3618 lemma find_cong[fundef_cong]:
3619 assumes "xs = ys" and "\<And>x. x \<in> set ys \<Longrightarrow> P x = Q x"
3620 shows "List.find P xs = List.find Q ys"
3621 proof (cases "List.find P xs")
3622 case None thus ?thesis by (metis find_None_iff assms)
3625 hence "List.find Q ys = Some x" using assms
3626 by (auto simp add: find_Some_iff)
3627 thus ?thesis using Some by auto
3630 lemma find_dropWhile:
3631 "List.find P xs = (case dropWhile (Not \<circ> P) xs
3632 of [] \<Rightarrow> None
3633 | x # _ \<Rightarrow> Some x)"
3634 by (induct xs) simp_all
3637 subsubsection {* @{const remove1} *}
3639 lemma remove1_append:
3640 "remove1 x (xs @ ys) =
3641 (if x \<in> set xs then remove1 x xs @ ys else xs @ remove1 x ys)"
3644 lemma remove1_commute: "remove1 x (remove1 y zs) = remove1 y (remove1 x zs)"
3647 lemma in_set_remove1[simp]:
3648 "a \<noteq> b \<Longrightarrow> a : set(remove1 b xs) = (a : set xs)"
3653 lemma set_remove1_subset: "set(remove1 x xs) <= set xs"
3660 lemma set_remove1_eq [simp]: "distinct xs ==> set(remove1 x xs) = set xs - {x}"
3667 lemma length_remove1:
3668 "length(remove1 x xs) = (if x : set xs then length xs - 1 else length xs)"
3670 apply (auto dest!:length_pos_if_in_set)
3673 lemma remove1_filter_not[simp]:
3674 "\<not> P x \<Longrightarrow> remove1 x (filter P xs) = filter P xs"
3677 lemma filter_remove1:
3678 "filter Q (remove1 x xs) = remove1 x (filter Q xs)"
3681 lemma notin_set_remove1[simp]: "x ~: set xs ==> x ~: set(remove1 y xs)"
3682 apply(insert set_remove1_subset)
3686 lemma distinct_remove1[simp]: "distinct xs ==> distinct(remove1 x xs)"
3687 by (induct xs) simp_all
3689 lemma remove1_remdups:
3690 "distinct xs \<Longrightarrow> remove1 x (remdups xs) = remdups (remove1 x xs)"
3691 by (induct xs) simp_all
3694 assumes "x \<notin> set xs"
3695 shows "remove1 x xs = xs"
3696 using assms by (induct xs) simp_all
3699 subsubsection {* @{const removeAll} *}
3701 lemma removeAll_filter_not_eq:
3702 "removeAll x = filter (\<lambda>y. x \<noteq> y)"
3705 show "removeAll x xs = filter (\<lambda>y. x \<noteq> y) xs"
3709 lemma removeAll_append[simp]:
3710 "removeAll x (xs @ ys) = removeAll x xs @ removeAll x ys"
3713 lemma set_removeAll[simp]: "set(removeAll x xs) = set xs - {x}"
3716 lemma removeAll_id[simp]: "x \<notin> set xs \<Longrightarrow> removeAll x xs = xs"
3719 (* Needs count:: 'a \<Rightarrow> 'a list \<Rightarrow> nat
3720 lemma length_removeAll:
3721 "length(removeAll x xs) = length xs - count x xs"
3724 lemma removeAll_filter_not[simp]:
3725 "\<not> P x \<Longrightarrow> removeAll x (filter P xs) = filter P xs"
3728 lemma distinct_removeAll:
3729 "distinct xs \<Longrightarrow> distinct (removeAll x xs)"
3730 by (simp add: removeAll_filter_not_eq)
3732 lemma distinct_remove1_removeAll:
3733 "distinct xs ==> remove1 x xs = removeAll x xs"
3734 by (induct xs) simp_all
3736 lemma map_removeAll_inj_on: "inj_on f (insert x (set xs)) \<Longrightarrow>
3737 map f (removeAll x xs) = removeAll (f x) (map f xs)"
3738 by (induct xs) (simp_all add:inj_on_def)
3740 lemma map_removeAll_inj: "inj f \<Longrightarrow>
3741 map f (removeAll x xs) = removeAll (f x) (map f xs)"
3742 by(metis map_removeAll_inj_on subset_inj_on subset_UNIV)
3745 subsubsection {* @{const replicate} *}
3747 lemma length_replicate [simp]: "length (replicate n x) = n"
3750 lemma Ex_list_of_length: "\<exists>xs. length xs = n"
3751 by (rule exI[of _ "replicate n undefined"]) simp
3753 lemma map_replicate [simp]: "map f (replicate n x) = replicate n (f x)"
3756 lemma map_replicate_const:
3757 "map (\<lambda> x. k) lst = replicate (length lst) k"
3758 by (induct lst) auto
3760 lemma replicate_app_Cons_same:
3761 "(replicate n x) @ (x # xs) = x # replicate n x @ xs"
3764 lemma rev_replicate [simp]: "rev (replicate n x) = replicate n x"
3765 apply (induct n, simp)
3766 apply (simp add: replicate_app_Cons_same)
3769 lemma replicate_add: "replicate (n + m) x = replicate n x @ replicate m x"
3772 text{* Courtesy of Matthias Daum: *}
3773 lemma append_replicate_commute:
3774 "replicate n x @ replicate k x = replicate k x @ replicate n x"
3775 apply (simp add: replicate_add [THEN sym])
3776 apply (simp add: add_commute)
3779 text{* Courtesy of Andreas Lochbihler: *}
3780 lemma filter_replicate:
3781 "filter P (replicate n x) = (if P x then replicate n x else [])"
3784 lemma hd_replicate [simp]: "n \<noteq> 0 ==> hd (replicate n x) = x"
3787 lemma tl_replicate [simp]: "tl (replicate n x) = replicate (n - 1) x"
3790 lemma last_replicate [simp]: "n \<noteq> 0 ==> last (replicate n x) = x"
3791 by (atomize (full), induct n) auto
3793 lemma nth_replicate[simp]: "i < n ==> (replicate n x)!i = x"
3794 apply (induct n arbitrary: i, simp)
3795 apply (simp add: nth_Cons split: nat.split)
3798 text{* Courtesy of Matthias Daum (2 lemmas): *}
3799 lemma take_replicate[simp]: "take i (replicate k x) = replicate (min i k) x"
3800 apply (case_tac "k \<le> i")
3801 apply (simp add: min_def)
3802 apply (drule not_leE)
3803 apply (simp add: min_def)
3804 apply (subgoal_tac "replicate k x = replicate i x @ replicate (k - i) x")
3806 apply (simp add: replicate_add [symmetric])
3809 lemma drop_replicate[simp]: "drop i (replicate k x) = replicate (k-i) x"
3810 apply (induct k arbitrary: i)
3818 lemma set_replicate_Suc: "set (replicate (Suc n) x) = {x}"
3821 lemma set_replicate [simp]: "n \<noteq> 0 ==> set (replicate n x) = {x}"
3822 by (fast dest!: not0_implies_Suc intro!: set_replicate_Suc)
3824 lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})"
3827 lemma in_set_replicate[simp]: "(x : set (replicate n y)) = (x = y & n \<noteq> 0)"
3828 by (simp add: set_replicate_conv_if)
3830 lemma Ball_set_replicate[simp]:
3831 "(ALL x : set(replicate n a). P x) = (P a | n=0)"
3832 by(simp add: set_replicate_conv_if)
3834 lemma Bex_set_replicate[simp]:
3835 "(EX x : set(replicate n a). P x) = (P a & n\<noteq>0)"
3836 by(simp add: set_replicate_conv_if)
3838 lemma replicate_append_same:
3839 "replicate i x @ [x] = x # replicate i x"
3840 by (induct i) simp_all
3842 lemma map_replicate_trivial:
3843 "map (\<lambda>i. x) [0..<i] = replicate i x"
3844 by (induct i) (simp_all add: replicate_append_same)
3846 lemma concat_replicate_trivial[simp]:
3847 "concat (replicate i []) = []"
3848 by (induct i) (auto simp add: map_replicate_const)
3850 lemma replicate_empty[simp]: "(replicate n x = []) \<longleftrightarrow> n=0"
3853 lemma empty_replicate[simp]: "([] = replicate n x) \<longleftrightarrow> n=0"
3856 lemma replicate_eq_replicate[simp]:
3857 "(replicate m x = replicate n y) \<longleftrightarrow> (m=n & (m\<noteq>0 \<longrightarrow> x=y))"
3858 apply(induct m arbitrary: n)
3864 lemma replicate_length_filter:
3865 "replicate (length (filter (\<lambda>y. x = y) xs)) x = filter (\<lambda>y. x = y) xs"
3868 lemma comm_append_are_replicate:
3869 fixes xs ys :: "'a list"
3870 assumes "xs \<noteq> []" "ys \<noteq> []"
3871 assumes "xs @ ys = ys @ xs"
3872 shows "\<exists>m n zs. concat (replicate m zs) = xs \<and> concat (replicate n zs) = ys"
3874 proof (induct "length (xs @ ys)" arbitrary: xs ys rule: less_induct)
3877 def xs' \<equiv> "if (length xs \<le> length ys) then xs else ys"
3878 and ys' \<equiv> "if (length xs \<le> length ys) then ys else xs"
3880 prems': "length xs' \<le> length ys'"
3881 "xs' @ ys' = ys' @ xs'"
3882 and "xs' \<noteq> []"
3883 and len: "length (xs @ ys) = length (xs' @ ys')"
3884 using less by (auto intro: less.hyps)
3887 obtain ws where "ys' = xs' @ ws"
3888 by (auto simp: append_eq_append_conv2)
3890 have "\<exists>m n zs. concat (replicate m zs) = xs' \<and> concat (replicate n zs) = ys'"
3891 proof (cases "ws = []")
3893 then have "concat (replicate 1 xs') = xs'"
3894 and "concat (replicate 1 xs') = ys'"
3895 using `ys' = xs' @ ws` by auto
3896 then show ?thesis by blast
3899 from `ys' = xs' @ ws` and `xs' @ ys' = ys' @ xs'`
3900 have "xs' @ ws = ws @ xs'" by simp
3901 then have "\<exists>m n zs. concat (replicate m zs) = xs' \<and> concat (replicate n zs) = ws"
3902 using False and `xs' \<noteq> []` and `ys' = xs' @ ws` and len
3903 by (intro less.hyps) auto
3904 then obtain m n zs where *: "concat (replicate m zs) = xs'"
3905 and "concat (replicate n zs) = ws" by blast
3906 then have "concat (replicate (m + n) zs) = ys'"
3907 using `ys' = xs' @ ws`
3908 by (simp add: replicate_add)
3909 with * show ?thesis by blast
3912 using xs'_def ys'_def by metis
3915 lemma comm_append_is_replicate:
3916 fixes xs ys :: "'a list"
3917 assumes "xs \<noteq> []" "ys \<noteq> []"
3918 assumes "xs @ ys = ys @ xs"
3919 shows "\<exists>n zs. n > 1 \<and> concat (replicate n zs) = xs @ ys"
3922 obtain m n zs where "concat (replicate m zs) = xs"
3923 and "concat (replicate n zs) = ys"
3924 using assms by (metis comm_append_are_replicate)
3925 then have "m + n > 1" and "concat (replicate (m+n) zs) = xs @ ys"
3926 using `xs \<noteq> []` and `ys \<noteq> []`
3927 by (auto simp: replicate_add)
3928 then show ?thesis by blast
3931 lemma Cons_replicate_eq:
3932 "x # xs = replicate n y \<longleftrightarrow> x = y \<and> n > 0 \<and> xs = replicate (n - 1) x"
3935 lemma replicate_length_same:
3936 "(\<forall>y\<in>set xs. y = x) \<Longrightarrow> replicate (length xs) x = xs"
3937 by (induct xs) simp_all
3939 lemma foldr_replicate [simp]:
3940 "foldr f (replicate n x) = f x ^^ n"
3941 by (induct n) (simp_all)
3943 lemma fold_replicate [simp]:
3944 "fold f (replicate n x) = f x ^^ n"
3945 by (subst foldr_fold [symmetric]) simp_all
3948 subsubsection {* @{const enumerate} *}
3950 lemma enumerate_simps [simp, code]:
3951 "enumerate n [] = []"
3952 "enumerate n (x # xs) = (n, x) # enumerate (Suc n) xs"
3953 apply (auto simp add: enumerate_eq_zip not_le)
3954 apply (cases "n < n + length xs")
3955 apply (auto simp add: upt_conv_Cons)
3958 lemma length_enumerate [simp]:
3959 "length (enumerate n xs) = length xs"
3960 by (simp add: enumerate_eq_zip)
3962 lemma map_fst_enumerate [simp]:
3963 "map fst (enumerate n xs) = [n..<n + length xs]"
3964 by (simp add: enumerate_eq_zip)
3966 lemma map_snd_enumerate [simp]:
3967 "map snd (enumerate n xs) = xs"
3968 by (simp add: enumerate_eq_zip)
3970 lemma in_set_enumerate_eq:
3971 "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"
3975 moreover assume "m < length xs + n"
3976 ultimately have "[n..<n + length xs] ! (m - n) = m \<and>
3977 xs ! (m - n) = xs ! (m - n) \<and> m - n < length xs" by auto
3978 then have "\<exists>q. [n..<n + length xs] ! q = m \<and>
3979 xs ! q = xs ! (m - n) \<and> q < length xs" ..
3980 } then show ?thesis by (cases p) (auto simp add: enumerate_eq_zip in_set_zip)
3983 lemma nth_enumerate_eq:
3984 assumes "m < length xs"
3985 shows "enumerate n xs ! m = (n + m, xs ! m)"
3986 using assms by (simp add: enumerate_eq_zip)
3988 lemma enumerate_replicate_eq:
3989 "enumerate n (replicate m a) = map (\<lambda>q. (q, a)) [n..<n + m]"
3990 by (rule pair_list_eqI)
3991 (simp_all add: enumerate_eq_zip comp_def map_replicate_const)
3993 lemma enumerate_Suc_eq:
3994 "enumerate (Suc n) xs = map (apfst Suc) (enumerate n xs)"
3995 by (rule pair_list_eqI)
3996 (simp_all add: not_le, simp del: map_map [simp del] add: map_Suc_upt map_map [symmetric])
3998 lemma distinct_enumerate [simp]:
3999 "distinct (enumerate n xs)"
4000 by (simp add: enumerate_eq_zip distinct_zipI1)
4003 subsubsection {* @{const rotate1} and @{const rotate} *}
4005 lemma rotate0[simp]: "rotate 0 = id"
4006 by(simp add:rotate_def)
4008 lemma rotate_Suc[simp]: "rotate (Suc n) xs = rotate1(rotate n xs)"
4009 by(simp add:rotate_def)
4012 "rotate (m+n) = rotate m o rotate n"
4013 by(simp add:rotate_def funpow_add)
4015 lemma rotate_rotate: "rotate m (rotate n xs) = rotate (m+n) xs"
4016 by(simp add:rotate_add)
4018 lemma rotate1_rotate_swap: "rotate1 (rotate n xs) = rotate n (rotate1 xs)"
4019 by(simp add:rotate_def funpow_swap1)
4021 lemma rotate1_length01[simp]: "length xs <= 1 \<Longrightarrow> rotate1 xs = xs"
4022 by(cases xs) simp_all
4024 lemma rotate_length01[simp]: "length xs <= 1 \<Longrightarrow> rotate n xs = xs"
4027 apply (simp add:rotate_def)
4030 lemma rotate1_hd_tl: "xs \<noteq> [] \<Longrightarrow> rotate1 xs = tl xs @ [hd xs]"
4031 by (cases xs) simp_all
4033 lemma rotate_drop_take:
4034 "rotate n xs = drop (n mod length xs) xs @ take (n mod length xs) xs"
4037 apply(simp add:rotate_def)
4038 apply(cases "xs = []")
4040 apply(case_tac "n mod length xs = 0")
4041 apply(simp add:mod_Suc)
4042 apply(simp add: rotate1_hd_tl drop_Suc take_Suc)
4043 apply(simp add:mod_Suc rotate1_hd_tl drop_Suc[symmetric] drop_tl[symmetric]
4044 take_hd_drop linorder_not_le)
4047 lemma rotate_conv_mod: "rotate n xs = rotate (n mod length xs) xs"
4048 by(simp add:rotate_drop_take)
4050 lemma rotate_id[simp]: "n mod length xs = 0 \<Longrightarrow> rotate n xs = xs"
4051 by(simp add:rotate_drop_take)
4053 lemma length_rotate1[simp]: "length(rotate1 xs) = length xs"
4054 by (cases xs) simp_all
4056 lemma length_rotate[simp]: "length(rotate n xs) = length xs"
4057 by (induct n arbitrary: xs) (simp_all add:rotate_def)
4059 lemma distinct1_rotate[simp]: "distinct(rotate1 xs) = distinct xs"
4062 lemma distinct_rotate[simp]: "distinct(rotate n xs) = distinct xs"
4063 by (induct n) (simp_all add:rotate_def)
4065 lemma rotate_map: "rotate n (map f xs) = map f (rotate n xs)"
4066 by(simp add:rotate_drop_take take_map drop_map)
4068 lemma set_rotate1[simp]: "set(rotate1 xs) = set xs"
4071 lemma set_rotate[simp]: "set(rotate n xs) = set xs"
4072 by (induct n) (simp_all add:rotate_def)
4074 lemma rotate1_is_Nil_conv[simp]: "(rotate1 xs = []) = (xs = [])"
4077 lemma rotate_is_Nil_conv[simp]: "(rotate n xs = []) = (xs = [])"
4078 by (induct n) (simp_all add:rotate_def)
4081 "rotate n (rev xs) = rev(rotate (length xs - (n mod length xs)) xs)"
4082 apply(simp add:rotate_drop_take rev_drop rev_take)
4083 apply(cases "length xs = 0")
4085 apply(cases "n mod length xs = 0")
4087 apply(simp add:rotate_drop_take rev_drop rev_take)
4090 lemma hd_rotate_conv_nth: "xs \<noteq> [] \<Longrightarrow> hd(rotate n xs) = xs!(n mod length xs)"
4091 apply(simp add:rotate_drop_take hd_append hd_drop_conv_nth hd_conv_nth)
4092 apply(subgoal_tac "length xs \<noteq> 0")
4094 using mod_less_divisor[of "length xs" n] by arith
4097 subsubsection {* @{const sublist} --- a generalization of @{const nth} to sets *}
4099 lemma sublist_empty [simp]: "sublist xs {} = []"
4100 by (auto simp add: sublist_def)
4102 lemma sublist_nil [simp]: "sublist [] A = []"
4103 by (auto simp add: sublist_def)
4105 lemma length_sublist:
4106 "length(sublist xs I) = card{i. i < length xs \<and> i : I}"
4107 by(simp add: sublist_def length_filter_conv_card cong:conj_cong)
4109 lemma sublist_shift_lemma_Suc:
4110 "map fst (filter (%p. P(Suc(snd p))) (zip xs is)) =
4111 map fst (filter (%p. P(snd p)) (zip xs (map Suc is)))"
4112 apply(induct xs arbitrary: "is")
4114 apply (case_tac "is")
4119 lemma sublist_shift_lemma:
4120 "map fst [p<-zip xs [i..<i + length xs] . snd p : A] =
4121 map fst [p<-zip xs [0..<length xs] . snd p + i : A]"
4122 by (induct xs rule: rev_induct) (simp_all add: add_commute)
4124 lemma sublist_append:
4125 "sublist (l @ l') A = sublist l A @ sublist l' {j. j + length l : A}"
4126 apply (unfold sublist_def)
4127 apply (induct l' rule: rev_induct, simp)
4128 apply (simp add: upt_add_eq_append[of 0] sublist_shift_lemma)
4129 apply (simp add: add_commute)
4133 "sublist (x # l) A = (if 0:A then [x] else []) @ sublist l {j. Suc j : A}"
4134 apply (induct l rule: rev_induct)
4135 apply (simp add: sublist_def)
4136 apply (simp del: append_Cons add: append_Cons[symmetric] sublist_append)
4139 lemma set_sublist: "set(sublist xs I) = {xs!i|i. i<size xs \<and> i \<in> I}"
4140 apply(induct xs arbitrary: I)
4141 apply(auto simp: sublist_Cons nth_Cons split:nat.split dest!: gr0_implies_Suc)
4144 lemma set_sublist_subset: "set(sublist xs I) \<subseteq> set xs"
4145 by(auto simp add:set_sublist)
4147 lemma notin_set_sublistI[simp]: "x \<notin> set xs \<Longrightarrow> x \<notin> set(sublist xs I)"
4148 by(auto simp add:set_sublist)
4150 lemma in_set_sublistD: "x \<in> set(sublist xs I) \<Longrightarrow> x \<in> set xs"
4151 by(auto simp add:set_sublist)
4153 lemma sublist_singleton [simp]: "sublist [x] A = (if 0 : A then [x] else [])"
4154 by (simp add: sublist_Cons)
4157 lemma distinct_sublistI[simp]: "distinct xs \<Longrightarrow> distinct(sublist xs I)"
4158 apply(induct xs arbitrary: I)
4160 apply(auto simp add:sublist_Cons)
4164 lemma sublist_upt_eq_take [simp]: "sublist l {..<n} = take n l"
4165 apply (induct l rule: rev_induct, simp)
4166 apply (simp split: nat_diff_split add: sublist_append)
4169 lemma filter_in_sublist:
4170 "distinct xs \<Longrightarrow> filter (%x. x \<in> set(sublist xs s)) xs = sublist xs s"
4171 proof (induct xs arbitrary: s)
4172 case Nil thus ?case by simp
4175 then have "!x. x: set xs \<longrightarrow> x \<noteq> a" by auto
4176 with Cons show ?case by(simp add: sublist_Cons cong:filter_cong)
4180 subsubsection {* @{const sublists} and @{const List.n_lists} *}
4182 lemma length_sublists:
4183 "length (sublists xs) = 2 ^ length xs"
4184 by (induct xs) (simp_all add: Let_def)
4186 lemma sublists_powset:
4187 "set ` set (sublists xs) = Pow (set xs)"
4189 have aux: "\<And>x A. set ` Cons x ` A = insert x ` set ` A"
4190 by (auto simp add: image_def)
4191 have "set (map set (sublists xs)) = Pow (set xs)"
4193 (simp_all add: aux Let_def Pow_insert Un_commute comp_def del: map_map)
4194 then show ?thesis by simp
4197 lemma distinct_set_sublists:
4198 assumes "distinct xs"
4199 shows "distinct (map set (sublists xs))"
4200 proof (rule card_distinct)
4201 have "finite (set xs)" by rule
4202 then have "card (Pow (set xs)) = 2 ^ card (set xs)" by (rule card_Pow)
4203 with assms distinct_card [of xs]
4204 have "card (Pow (set xs)) = 2 ^ length xs" by simp
4205 then show "card (set (map set (sublists xs))) = length (map set (sublists xs))"
4206 by (simp add: sublists_powset length_sublists)
4209 lemma n_lists_Nil [simp]: "List.n_lists n [] = (if n = 0 then [[]] else [])"
4210 by (induct n) simp_all
4212 lemma length_n_lists: "length (List.n_lists n xs) = length xs ^ n"
4213 by (induct n) (auto simp add: length_concat o_def listsum_triv)
4215 lemma length_n_lists_elem: "ys \<in> set (List.n_lists n xs) \<Longrightarrow> length ys = n"
4216 by (induct n arbitrary: ys) auto
4218 lemma set_n_lists: "set (List.n_lists n xs) = {ys. length ys = n \<and> set ys \<subseteq> set xs}"
4219 proof (rule set_eqI)
4221 show "ys \<in> set (List.n_lists n xs) \<longleftrightarrow> ys \<in> {ys. length ys = n \<and> set ys \<subseteq> set xs}"
4223 have "ys \<in> set (List.n_lists n xs) \<Longrightarrow> length ys = n"
4224 by (induct n arbitrary: ys) auto
4225 moreover have "\<And>x. ys \<in> set (List.n_lists n xs) \<Longrightarrow> x \<in> set ys \<Longrightarrow> x \<in> set xs"
4226 by (induct n arbitrary: ys) auto
4227 moreover have "set ys \<subseteq> set xs \<Longrightarrow> ys \<in> set (List.n_lists (length ys) xs)"
4229 ultimately show ?thesis by auto
4233 lemma distinct_n_lists:
4234 assumes "distinct xs"
4235 shows "distinct (List.n_lists n xs)"
4236 proof (rule card_distinct)
4237 from assms have card_length: "card (set xs) = length xs" by (rule distinct_card)
4238 have "card (set (List.n_lists n xs)) = card (set xs) ^ n"
4240 case 0 then show ?case by simp
4243 moreover have "card (\<Union>ys\<in>set (List.n_lists n xs). (\<lambda>y. y # ys) ` set xs)
4244 = (\<Sum>ys\<in>set (List.n_lists n xs). card ((\<lambda>y. y # ys) ` set xs))"
4245 by (rule card_UN_disjoint) auto
4246 moreover have "\<And>ys. card ((\<lambda>y. y # ys) ` set xs) = card (set xs)"
4247 by (rule card_image) (simp add: inj_on_def)
4248 ultimately show ?case by auto
4250 also have "\<dots> = length xs ^ n" by (simp add: card_length)
4251 finally show "card (set (List.n_lists n xs)) = length (List.n_lists n xs)"
4252 by (simp add: length_n_lists)
4256 subsubsection {* @{const splice} *}
4258 lemma splice_Nil2 [simp, code]: "splice xs [] = xs"
4259 by (cases xs) simp_all
4261 declare splice.simps(1,3)[code]
4262 declare splice.simps(2)[simp del]
4264 lemma length_splice[simp]: "length(splice xs ys) = length xs + length ys"
4265 by (induct xs ys rule: splice.induct) auto
4268 subsubsection {* Transpose *}
4270 function transpose where
4271 "transpose [] = []" |
4272 "transpose ([] # xss) = transpose xss" |
4273 "transpose ((x#xs) # xss) =
4274 (x # [h. (h#t) \<leftarrow> xss]) # transpose (xs # [t. (h#t) \<leftarrow> xss])"
4275 by pat_completeness auto
4277 lemma transpose_aux_filter_head:
4278 "concat (map (list_case [] (\<lambda>h t. [h])) xss) =
4279 map (\<lambda>xs. hd xs) [ys\<leftarrow>xss . ys \<noteq> []]"
4280 by (induct xss) (auto split: list.split)
4282 lemma transpose_aux_filter_tail:
4283 "concat (map (list_case [] (\<lambda>h t. [t])) xss) =
4284 map (\<lambda>xs. tl xs) [ys\<leftarrow>xss . ys \<noteq> []]"
4285 by (induct xss) (auto split: list.split)
4287 lemma transpose_aux_max:
4288 "max (Suc (length xs)) (foldr (\<lambda>xs. max (length xs)) xss 0) =
4289 Suc (max (length xs) (foldr (\<lambda>x. max (length x - Suc 0)) [ys\<leftarrow>xss . ys\<noteq>[]] 0))"
4290 (is "max _ ?foldB = Suc (max _ ?foldA)")
4291 proof (cases "[ys\<leftarrow>xss . ys\<noteq>[]] = []")
4293 hence "foldr (\<lambda>xs. max (length xs)) xss 0 = 0"
4296 then have "x = []" by (cases x) auto
4297 with Cons show ?case by auto
4299 thus ?thesis using True by simp
4303 have foldA: "?foldA = foldr (\<lambda>x. max (length x)) [ys\<leftarrow>xss . ys \<noteq> []] 0 - 1"
4304 by (induct xss) auto
4305 have foldB: "?foldB = foldr (\<lambda>x. max (length x)) [ys\<leftarrow>xss . ys \<noteq> []] 0"
4306 by (induct xss) auto
4311 obtain z zs where zs: "[ys\<leftarrow>xss . ys \<noteq> []] = z#zs" by (auto simp: neq_Nil_conv)
4312 hence "z \<in> set ([ys\<leftarrow>xss . ys \<noteq> []])" by auto
4313 hence "z \<noteq> []" by auto
4316 by (auto simp: max_def intro: less_le_trans)
4319 unfolding foldA foldB max_Suc_Suc[symmetric]
4323 termination transpose
4324 by (relation "measure (\<lambda>xs. foldr (\<lambda>xs. max (length xs)) xs 0 + length xs)")
4325 (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max less_Suc_eq_le)
4327 lemma transpose_empty: "(transpose xs = []) \<longleftrightarrow> (\<forall>x \<in> set xs. x = [])"
4328 by (induct rule: transpose.induct) simp_all
4330 lemma length_transpose:
4331 fixes xs :: "'a list list"
4332 shows "length (transpose xs) = foldr (\<lambda>xs. max (length xs)) xs 0"
4333 by (induct rule: transpose.induct)
4334 (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max
4335 max_Suc_Suc[symmetric] simp del: max_Suc_Suc)
4337 lemma nth_transpose:
4338 fixes xs :: "'a list list"
4339 assumes "i < length (transpose xs)"
4340 shows "transpose xs ! i = map (\<lambda>xs. xs ! i) [ys \<leftarrow> xs. i < length ys]"
4341 using assms proof (induct arbitrary: i rule: transpose.induct)
4343 def XS == "(x # xs) # xss"
4344 hence [simp]: "XS \<noteq> []" by auto
4348 thus ?thesis by (simp add: transpose_aux_filter_head hd_conv_nth)
4351 have *: "\<And>xss. xs # map tl xss = map tl ((x#xs)#xss)" by simp
4352 have **: "\<And>xss. (x#xs) # filter (\<lambda>ys. ys \<noteq> []) xss = filter (\<lambda>ys. ys \<noteq> []) ((x#xs)#xss)" by simp
4353 { fix x have "Suc j < length x \<longleftrightarrow> x \<noteq> [] \<and> j < length x - Suc 0"
4354 by (cases x) simp_all
4357 have j_less: "j < length (transpose (xs # concat (map (list_case [] (\<lambda>h t. [t])) xss)))"
4358 using "3.prems" by (simp add: transpose_aux_filter_tail length_transpose Suc)
4361 unfolding transpose.simps `i = Suc j` nth_Cons_Suc "3.hyps"[OF j_less]
4362 apply (auto simp: transpose_aux_filter_tail filter_map comp_def length_transpose * ** *** XS_def[symmetric])
4363 apply (rule_tac y=x in list.exhaust)
4368 lemma transpose_map_map:
4369 "transpose (map (map f) xs) = map (map f) (transpose xs)"
4370 proof (rule nth_equalityI, safe)
4371 have [simp]: "length (transpose (map (map f) xs)) = length (transpose xs)"
4372 by (simp add: length_transpose foldr_map comp_def)
4373 show "length (transpose (map (map f) xs)) = length (map (map f) (transpose xs))" by simp
4375 fix i assume "i < length (transpose (map (map f) xs))"
4376 thus "transpose (map (map f) xs) ! i = map (map f) (transpose xs) ! i"
4377 by (simp add: nth_transpose filter_map comp_def)
4381 subsubsection {* (In)finiteness *}
4383 lemma finite_maxlen:
4384 "finite (M::'a list set) ==> EX n. ALL s:M. size s < n"
4385 proof (induct rule: finite.induct)
4386 case emptyI show ?case by simp
4389 then obtain n where "\<forall>s\<in>M. length s < n" by blast
4390 hence "ALL s:insert xs M. size s < max n (size xs) + 1" by auto
4394 lemma lists_length_Suc_eq:
4395 "{xs. set xs \<subseteq> A \<and> length xs = Suc n} =
4396 (\<lambda>(xs, n). n#xs) ` ({xs. set xs \<subseteq> A \<and> length xs = n} \<times> A)"
4397 by (auto simp: length_Suc_conv)
4401 shows finite_lists_length_eq: "finite {xs. set xs \<subseteq> A \<and> length xs = n}"
4402 and card_lists_length_eq: "card {xs. set xs \<subseteq> A \<and> length xs = n} = (card A)^n"
4405 (auto simp: card_image inj_split_Cons lists_length_Suc_eq cong: conj_cong)
4407 lemma finite_lists_length_le:
4408 assumes "finite A" shows "finite {xs. set xs \<subseteq> A \<and> length xs \<le> n}"
4411 have "?S = (\<Union>n\<in>{0..n}. {xs. set xs \<subseteq> A \<and> length xs = n})" by auto
4412 thus ?thesis by (auto intro!: finite_lists_length_eq[OF `finite A`] simp only:)
4415 lemma card_lists_length_le:
4416 assumes "finite A" shows "card {xs. set xs \<subseteq> A \<and> length xs \<le> n} = (\<Sum>i\<le>n. card A^i)"
4418 have "(\<Sum>i\<le>n. card A^i) = card (\<Union>i\<le>n. {xs. set xs \<subseteq> A \<and> length xs = i})"
4420 by (subst card_UN_disjoint)
4421 (auto simp add: card_lists_length_eq finite_lists_length_eq)
4422 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}"
4424 finally show ?thesis by simp
4427 lemma card_lists_distinct_length_eq:
4428 assumes "k < card A"
4429 shows "card {xs. length xs = k \<and> distinct xs \<and> set xs \<subseteq> A} = \<Prod>{card A - k + 1 .. card A}"
4433 then have "{xs. length xs = 0 \<and> distinct xs \<and> set xs \<subseteq> A} = {[]}" by auto
4434 then show ?case by simp
4437 let "?k_list" = "\<lambda>k xs. length xs = k \<and> distinct xs \<and> set xs \<subseteq> A"
4438 have inj_Cons: "\<And>A. inj_on (\<lambda>(xs, n). n # xs) A" by (rule inj_onI) auto
4440 from Suc have "k < card A" by simp
4441 moreover have "finite A" using assms by (simp add: card_ge_0_finite)
4442 moreover have "finite {xs. ?k_list k xs}"
4443 using finite_lists_length_eq[OF `finite A`, of k]
4444 by - (rule finite_subset, auto)
4445 moreover have "\<And>i j. i \<noteq> j \<longrightarrow> {i} \<times> (A - set i) \<inter> {j} \<times> (A - set j) = {}"
4447 moreover have "\<And>i. i \<in>Collect (?k_list k) \<Longrightarrow> card (A - set i) = card A - k"
4448 by (simp add: card_Diff_subset distinct_card)
4449 moreover have "{xs. ?k_list (Suc k) xs} =
4450 (\<lambda>(xs, n). n#xs) ` \<Union>((\<lambda>xs. {xs} \<times> (A - set xs)) ` {xs. ?k_list k xs})"
4451 by (auto simp: length_Suc_conv)
4453 have "Suc (card A - Suc k) = card A - k" using Suc.prems by simp
4454 then have "(card A - k) * \<Prod>{Suc (card A - k)..card A} = \<Prod>{Suc (card A - Suc k)..card A}"
4455 by (subst setprod_insert[symmetric]) (simp add: atLeastAtMost_insertL)+
4456 ultimately show ?case
4457 by (simp add: card_image inj_Cons card_UN_disjoint Suc.hyps algebra_simps)
4460 lemma infinite_UNIV_listI: "~ finite(UNIV::'a list set)"
4462 apply(drule finite_maxlen)
4463 apply (metis UNIV_I length_replicate less_not_refl)
4467 subsection {* Sorting *}
4469 text{* Currently it is not shown that @{const sort} returns a
4470 permutation of its input because the nicest proof is via multisets,
4471 which are not yet available. Alternatively one could define a function
4472 that counts the number of occurrences of an element in a list and use
4473 that instead of multisets to state the correctness property. *}
4478 lemma set_insort_key:
4479 "set (insort_key f x xs) = insert x (set xs)"
4482 lemma length_insort [simp]:
4483 "length (insort_key f x xs) = Suc (length xs)"
4484 by (induct xs) simp_all
4486 lemma insort_key_left_comm:
4487 assumes "f x \<noteq> f y"
4488 shows "insort_key f y (insort_key f x xs) = insort_key f x (insort_key f y xs)"
4489 by (induct xs) (auto simp add: assms dest: antisym)
4491 lemma insort_left_comm:
4492 "insort x (insort y xs) = insort y (insort x xs)"
4493 by (cases "x = y") (auto intro: insort_key_left_comm)
4495 lemma comp_fun_commute_insort:
4496 "comp_fun_commute insort"
4498 qed (simp add: insort_left_comm fun_eq_iff)
4500 lemma sort_key_simps [simp]:
4501 "sort_key f [] = []"
4502 "sort_key f (x#xs) = insort_key f x (sort_key f xs)"
4503 by (simp_all add: sort_key_def)
4505 lemma (in linorder) sort_key_conv_fold:
4506 assumes "inj_on f (set xs)"
4507 shows "sort_key f xs = fold (insort_key f) xs []"
4509 have "fold (insort_key f) (rev xs) = fold (insort_key f) xs"
4510 proof (rule fold_rev, rule ext)
4513 assume "x \<in> set xs" "y \<in> set xs"
4514 with assms have *: "f y = f x \<Longrightarrow> y = x" by (auto dest: inj_onD)
4515 have **: "x = y \<longleftrightarrow> y = x" by auto
4516 show "(insort_key f y \<circ> insort_key f x) zs = (insort_key f x \<circ> insort_key f y) zs"
4517 by (induct zs) (auto intro: * simp add: **)
4519 then show ?thesis by (simp add: sort_key_def foldr_conv_fold)
4522 lemma (in linorder) sort_conv_fold:
4523 "sort xs = fold insort xs []"
4524 by (rule sort_key_conv_fold) simp
4526 lemma length_sort[simp]: "length (sort_key f xs) = length xs"
4527 by (induct xs, auto)
4529 lemma sorted_Cons: "sorted (x#xs) = (sorted xs & (ALL y:set xs. x <= y))"
4530 apply(induct xs arbitrary: x) apply simp
4531 by simp (blast intro: order_trans)
4534 "sorted xs \<Longrightarrow> sorted (tl xs)"
4535 by (cases xs) (simp_all add: sorted_Cons)
4537 lemma sorted_append:
4538 "sorted (xs@ys) = (sorted xs & sorted ys & (\<forall>x \<in> set xs. \<forall>y \<in> set ys. x\<le>y))"
4539 by (induct xs) (auto simp add:sorted_Cons)
4541 lemma sorted_nth_mono:
4542 "sorted xs \<Longrightarrow> i \<le> j \<Longrightarrow> j < length xs \<Longrightarrow> xs!i \<le> xs!j"
4543 by (induct xs arbitrary: i j) (auto simp:nth_Cons' sorted_Cons)
4545 lemma sorted_rev_nth_mono:
4546 "sorted (rev xs) \<Longrightarrow> i \<le> j \<Longrightarrow> j < length xs \<Longrightarrow> xs!j \<le> xs!i"
4547 using sorted_nth_mono[ of "rev xs" "length xs - j - 1" "length xs - i - 1"]
4548 rev_nth[of "length xs - i - 1" "xs"] rev_nth[of "length xs - j - 1" "xs"]
4551 lemma sorted_nth_monoI:
4552 "(\<And> i j. \<lbrakk> i \<le> j ; j < length xs \<rbrakk> \<Longrightarrow> xs ! i \<le> xs ! j) \<Longrightarrow> sorted xs"
4556 proof (rule Cons.hyps)
4557 fix i j assume "i \<le> j" and "j < length xs"
4558 with Cons.prems[of "Suc i" "Suc j"]
4559 show "xs ! i \<le> xs ! j" by auto
4563 fix y assume "y \<in> set xs"
4564 then obtain j where "j < length xs" and "xs ! j = y"
4565 unfolding in_set_conv_nth by blast
4566 with Cons.prems[of 0 "Suc j"]
4572 unfolding sorted_Cons by auto
4575 lemma sorted_equals_nth_mono:
4576 "sorted xs = (\<forall>j < length xs. \<forall>i \<le> j. xs ! i \<le> xs ! j)"
4577 by (auto intro: sorted_nth_monoI sorted_nth_mono)
4579 lemma set_insort: "set(insort_key f x xs) = insert x (set xs)"
4582 lemma set_sort[simp]: "set(sort_key f xs) = set xs"
4583 by (induct xs) (simp_all add:set_insort)
4585 lemma distinct_insort: "distinct (insort_key f x xs) = (x \<notin> set xs \<and> distinct xs)"
4586 by(induct xs)(auto simp:set_insort)
4588 lemma distinct_sort[simp]: "distinct (sort_key f xs) = distinct xs"
4589 by (induct xs) (simp_all add: distinct_insort)
4591 lemma sorted_insort_key: "sorted (map f (insort_key f x xs)) = sorted (map f xs)"
4592 by (induct xs) (auto simp:sorted_Cons set_insort)
4594 lemma sorted_insort: "sorted (insort x xs) = sorted xs"
4595 using sorted_insort_key [where f="\<lambda>x. x"] by simp
4597 theorem sorted_sort_key [simp]: "sorted (map f (sort_key f xs))"
4598 by (induct xs) (auto simp:sorted_insort_key)
4600 theorem sorted_sort [simp]: "sorted (sort xs)"
4601 using sorted_sort_key [where f="\<lambda>x. x"] by simp
4603 lemma sorted_butlast:
4604 assumes "xs \<noteq> []" and "sorted xs"
4605 shows "sorted (butlast xs)"
4607 from `xs \<noteq> []` obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto
4608 with `sorted xs` show ?thesis by (simp add: sorted_append)
4611 lemma insort_not_Nil [simp]:
4612 "insort_key f a xs \<noteq> []"
4613 by (induct xs) simp_all
4615 lemma insort_is_Cons: "\<forall>x\<in>set xs. f a \<le> f x \<Longrightarrow> insort_key f a xs = a # xs"
4618 lemma sorted_sort_id: "sorted xs \<Longrightarrow> sort xs = xs"
4619 by (induct xs) (auto simp add: sorted_Cons insort_is_Cons)
4621 lemma sorted_map_remove1:
4622 "sorted (map f xs) \<Longrightarrow> sorted (map f (remove1 x xs))"
4623 by (induct xs) (auto simp add: sorted_Cons)
4625 lemma sorted_remove1: "sorted xs \<Longrightarrow> sorted (remove1 a xs)"
4626 using sorted_map_remove1 [of "\<lambda>x. x"] by simp
4628 lemma insort_key_remove1:
4629 assumes "a \<in> set xs" and "sorted (map f xs)" and "hd (filter (\<lambda>x. f a = f x) xs) = a"
4630 shows "insort_key f a (remove1 a xs) = xs"
4631 using assms proof (induct xs)
4634 proof (cases "x = a")
4636 then have "f x \<noteq> f a" using Cons.prems by auto
4637 then have "f x < f a" using Cons.prems by (auto simp: sorted_Cons)
4638 with `f x \<noteq> f a` show ?thesis using Cons by (auto simp: sorted_Cons insort_is_Cons)
4639 qed (auto simp: sorted_Cons insort_is_Cons)
4642 lemma insort_remove1:
4643 assumes "a \<in> set xs" and "sorted xs"
4644 shows "insort a (remove1 a xs) = xs"
4645 proof (rule insort_key_remove1)
4646 from `a \<in> set xs` show "a \<in> set xs" .
4647 from `sorted xs` show "sorted (map (\<lambda>x. x) xs)" by simp
4648 from `a \<in> set xs` have "a \<in> set (filter (op = a) xs)" by auto
4649 then have "set (filter (op = a) xs) \<noteq> {}" by auto
4650 then have "filter (op = a) xs \<noteq> []" by (auto simp only: set_empty)
4651 then have "length (filter (op = a) xs) > 0" by simp
4652 then obtain n where n: "Suc n = length (filter (op = a) xs)"
4653 by (cases "length (filter (op = a) xs)") simp_all
4654 moreover have "replicate (Suc n) a = a # replicate n a"
4656 ultimately show "hd (filter (op = a) xs) = a" by (simp add: replicate_length_filter)
4659 lemma sorted_remdups[simp]:
4660 "sorted l \<Longrightarrow> sorted (remdups l)"
4661 by (induct l) (auto simp: sorted_Cons)
4663 lemma sorted_remdups_adj[simp]:
4664 "sorted xs \<Longrightarrow> sorted (remdups_adj xs)"
4665 by (induct xs rule: remdups_adj.induct, simp_all split: split_if_asm add: sorted_Cons)
4667 lemma sorted_distinct_set_unique:
4668 assumes "sorted xs" "distinct xs" "sorted ys" "distinct ys" "set xs = set ys"
4671 from assms have 1: "length xs = length ys" by (auto dest!: distinct_card)
4672 from assms show ?thesis
4673 proof(induct rule:list_induct2[OF 1])
4674 case 1 show ?case by simp
4676 case 2 thus ?case by (simp add:sorted_Cons)
4677 (metis Diff_insert_absorb antisym insertE insert_iff)
4681 lemma map_sorted_distinct_set_unique:
4682 assumes "inj_on f (set xs \<union> set ys)"
4683 assumes "sorted (map f xs)" "distinct (map f xs)"
4684 "sorted (map f ys)" "distinct (map f ys)"
4685 assumes "set xs = set ys"
4688 from assms have "map f xs = map f ys"
4689 by (simp add: sorted_distinct_set_unique)
4690 with `inj_on f (set xs \<union> set ys)` show "xs = ys"
4691 by (blast intro: map_inj_on)
4694 lemma finite_sorted_distinct_unique:
4695 shows "finite A \<Longrightarrow> EX! xs. set xs = A & sorted xs & distinct xs"
4696 apply(drule finite_distinct_list)
4698 apply(rule_tac a="sort xs" in ex1I)
4699 apply (auto simp: sorted_distinct_set_unique)
4704 shows sorted_take: "sorted (take n xs)"
4705 and sorted_drop: "sorted (drop n xs)"
4707 from assms have "sorted (take n xs @ drop n xs)" by simp
4708 then show "sorted (take n xs)" and "sorted (drop n xs)"
4709 unfolding sorted_append by simp_all
4712 lemma sorted_dropWhile: "sorted xs \<Longrightarrow> sorted (dropWhile P xs)"
4713 by (auto dest: sorted_drop simp add: dropWhile_eq_drop)
4715 lemma sorted_takeWhile: "sorted xs \<Longrightarrow> sorted (takeWhile P xs)"
4716 by (subst takeWhile_eq_take) (auto dest: sorted_take)
4718 lemma sorted_filter:
4719 "sorted (map f xs) \<Longrightarrow> sorted (map f (filter P xs))"
4720 by (induct xs) (simp_all add: sorted_Cons)
4722 lemma foldr_max_sorted:
4723 assumes "sorted (rev xs)"
4724 shows "foldr max xs y = (if xs = [] then y else max (xs ! 0) y)"
4728 then have "sorted (rev xs)" using sorted_append by auto
4729 with Cons show ?case
4730 by (cases xs) (auto simp add: sorted_append max_def)
4733 lemma filter_equals_takeWhile_sorted_rev:
4734 assumes sorted: "sorted (rev (map f xs))"
4735 shows "[x \<leftarrow> xs. t < f x] = takeWhile (\<lambda> x. t < f x) xs"
4736 (is "filter ?P xs = ?tW")
4737 proof (rule takeWhile_eq_filter[symmetric])
4738 let "?dW" = "dropWhile ?P xs"
4739 fix x assume "x \<in> set ?dW"
4740 then obtain i where i: "i < length ?dW" and nth_i: "x = ?dW ! i"
4741 unfolding in_set_conv_nth by auto
4742 hence "length ?tW + i < length (?tW @ ?dW)"
4743 unfolding length_append by simp
4744 hence i': "length (map f ?tW) + i < length (map f xs)" by simp
4745 have "(map f ?tW @ map f ?dW) ! (length (map f ?tW) + i) \<le>
4746 (map f ?tW @ map f ?dW) ! (length (map f ?tW) + 0)"
4747 using sorted_rev_nth_mono[OF sorted _ i', of "length ?tW"]
4748 unfolding map_append[symmetric] by simp
4749 hence "f x \<le> f (?dW ! 0)"
4750 unfolding nth_append_length_plus nth_i
4751 using i preorder_class.le_less_trans[OF le0 i] by simp
4752 also have "... \<le> t"
4753 using hd_dropWhile[of "?P" xs] le0[THEN preorder_class.le_less_trans, OF i]
4754 using hd_conv_nth[of "?dW"] by simp
4755 finally show "\<not> t < f x" by simp
4758 lemma insort_insert_key_triv:
4759 "f x \<in> f ` set xs \<Longrightarrow> insort_insert_key f x xs = xs"
4760 by (simp add: insort_insert_key_def)
4762 lemma insort_insert_triv:
4763 "x \<in> set xs \<Longrightarrow> insort_insert x xs = xs"
4764 using insort_insert_key_triv [of "\<lambda>x. x"] by simp
4766 lemma insort_insert_insort_key:
4767 "f x \<notin> f ` set xs \<Longrightarrow> insort_insert_key f x xs = insort_key f x xs"
4768 by (simp add: insort_insert_key_def)
4770 lemma insort_insert_insort:
4771 "x \<notin> set xs \<Longrightarrow> insort_insert x xs = insort x xs"
4772 using insort_insert_insort_key [of "\<lambda>x. x"] by simp
4774 lemma set_insort_insert:
4775 "set (insort_insert x xs) = insert x (set xs)"
4776 by (auto simp add: insort_insert_key_def set_insort)
4778 lemma distinct_insort_insert:
4779 assumes "distinct xs"
4780 shows "distinct (insort_insert_key f x xs)"
4781 using assms by (induct xs) (auto simp add: insort_insert_key_def set_insort)
4783 lemma sorted_insort_insert_key:
4784 assumes "sorted (map f xs)"
4785 shows "sorted (map f (insort_insert_key f x xs))"
4786 using assms by (simp add: insort_insert_key_def sorted_insort_key)
4788 lemma sorted_insort_insert:
4790 shows "sorted (insort_insert x xs)"
4791 using assms sorted_insort_insert_key [of "\<lambda>x. x"] by simp
4793 lemma filter_insort_triv:
4794 "\<not> P x \<Longrightarrow> filter P (insort_key f x xs) = filter P xs"
4795 by (induct xs) simp_all
4797 lemma filter_insort:
4798 "sorted (map f xs) \<Longrightarrow> P x \<Longrightarrow> filter P (insort_key f x xs) = insort_key f x (filter P xs)"
4799 using assms by (induct xs)
4800 (auto simp add: sorted_Cons, subst insort_is_Cons, auto)
4803 "filter P (sort_key f xs) = sort_key f (filter P xs)"
4804 by (induct xs) (simp_all add: filter_insort_triv filter_insort)
4806 lemma sorted_map_same:
4807 "sorted (map f [x\<leftarrow>xs. f x = g xs])"
4808 proof (induct xs arbitrary: g)
4809 case Nil then show ?case by simp
4812 then have "sorted (map f [y\<leftarrow>xs . f y = (\<lambda>xs. f x) xs])" .
4813 moreover from Cons have "sorted (map f [y\<leftarrow>xs . f y = (g \<circ> Cons x) xs])" .
4814 ultimately show ?case by (simp_all add: sorted_Cons)
4818 "sorted [x\<leftarrow>xs. x = g xs]"
4819 using sorted_map_same [of "\<lambda>x. x"] by simp
4821 lemma remove1_insort [simp]:
4822 "remove1 x (insort x xs) = xs"
4823 by (induct xs) simp_all
4827 lemma sorted_upt[simp]: "sorted[i..<j]"
4828 by (induct j) (simp_all add:sorted_append)
4830 lemma sorted_upto[simp]: "sorted[i..j]"
4831 apply(induct i j rule:upto.induct)
4832 apply(subst upto.simps)
4833 apply(simp add:sorted_Cons)
4836 lemma sorted_find_Min:
4838 assumes "\<exists>x \<in> set xs. P x"
4839 shows "List.find P xs = Some (Min {x\<in>set xs. P x})"
4840 using assms proof (induct xs rule: sorted.induct)
4841 case Nil then show ?case by simp
4843 case (Cons xs x) show ?case proof (cases "P x")
4844 case True with Cons show ?thesis by (auto intro: Min_eqI [symmetric])
4846 case False then have "{y. (y = x \<or> y \<in> set xs) \<and> P y} = {y \<in> set xs. P y}"
4848 with Cons False show ?thesis by simp_all
4853 subsubsection {* @{const transpose} on sorted lists *}
4855 lemma sorted_transpose[simp]:
4856 shows "sorted (rev (map length (transpose xs)))"
4857 by (auto simp: sorted_equals_nth_mono rev_nth nth_transpose
4858 length_filter_conv_card intro: card_mono)
4860 lemma transpose_max_length:
4861 "foldr (\<lambda>xs. max (length xs)) (transpose xs) 0 = length [x \<leftarrow> xs. x \<noteq> []]"
4863 proof (cases "transpose xs = []")
4865 have "?L = foldr max (map length (transpose xs)) 0"
4866 by (simp add: foldr_map comp_def)
4867 also have "... = length (transpose xs ! 0)"
4868 using False sorted_transpose by (simp add: foldr_max_sorted)
4869 finally show ?thesis
4870 using False by (simp add: nth_transpose)
4873 hence "[x \<leftarrow> xs. x \<noteq> []] = []"
4874 by (auto intro!: filter_False simp: transpose_empty)
4875 thus ?thesis by (simp add: transpose_empty True)
4878 lemma length_transpose_sorted:
4879 fixes xs :: "'a list list"
4880 assumes sorted: "sorted (rev (map length xs))"
4881 shows "length (transpose xs) = (if xs = [] then 0 else length (xs ! 0))"
4882 proof (cases "xs = []")
4885 using foldr_max_sorted[OF sorted] False
4886 unfolding length_transpose foldr_map comp_def
4890 lemma nth_nth_transpose_sorted[simp]:
4891 fixes xs :: "'a list list"
4892 assumes sorted: "sorted (rev (map length xs))"
4893 and i: "i < length (transpose xs)"
4894 and j: "j < length [ys \<leftarrow> xs. i < length ys]"
4895 shows "transpose xs ! i ! j = xs ! j ! i"
4896 using j filter_equals_takeWhile_sorted_rev[OF sorted, of i]
4897 nth_transpose[OF i] nth_map[OF j]
4898 by (simp add: takeWhile_nth)
4900 lemma transpose_column_length:
4901 fixes xs :: "'a list list"
4902 assumes sorted: "sorted (rev (map length xs))" and "i < length xs"
4903 shows "length (filter (\<lambda>ys. i < length ys) (transpose xs)) = length (xs ! i)"
4905 have "xs \<noteq> []" using `i < length xs` by auto
4906 note filter_equals_takeWhile_sorted_rev[OF sorted, simp]
4907 { fix j assume "j \<le> i"
4908 note sorted_rev_nth_mono[OF sorted, of j i, simplified, OF this `i < length xs`]
4909 } note sortedE = this[consumes 1]
4911 have "{j. j < length (transpose xs) \<and> i < length (transpose xs ! j)}
4912 = {..< length (xs ! i)}"
4915 assume "j < length (transpose xs)" and "i < length (transpose xs ! j)"
4916 with this(2) nth_transpose[OF this(1)]
4917 have "i < length (takeWhile (\<lambda>ys. j < length ys) xs)" by simp
4918 from nth_mem[OF this] takeWhile_nth[OF this]
4919 show "j < length (xs ! i)" by (auto dest: set_takeWhileD)
4921 fix j assume "j < length (xs ! i)"
4922 thus "j < length (transpose xs)"
4923 using foldr_max_sorted[OF sorted] `xs \<noteq> []` sortedE[OF le0]
4924 by (auto simp: length_transpose comp_def foldr_map)
4926 have "Suc i \<le> length (takeWhile (\<lambda>ys. j < length ys) xs)"
4927 using `i < length xs` `j < length (xs ! i)` less_Suc_eq_le
4928 by (auto intro!: length_takeWhile_less_P_nth dest!: sortedE)
4929 with nth_transpose[OF `j < length (transpose xs)`]
4930 show "i < length (transpose xs ! j)" by simp
4932 thus ?thesis by (simp add: length_filter_conv_card)
4935 lemma transpose_column:
4936 fixes xs :: "'a list list"
4937 assumes sorted: "sorted (rev (map length xs))" and "i < length xs"
4938 shows "map (\<lambda>ys. ys ! i) (filter (\<lambda>ys. i < length ys) (transpose xs))
4939 = xs ! i" (is "?R = _")
4940 proof (rule nth_equalityI, safe)
4941 show length: "length ?R = length (xs ! i)"
4942 using transpose_column_length[OF assms] by simp
4944 fix j assume j: "j < length ?R"
4945 note * = less_le_trans[OF this, unfolded length_map, OF length_filter_le]
4946 from j have j_less: "j < length (xs ! i)" using length by simp
4947 have i_less_tW: "Suc i \<le> length (takeWhile (\<lambda>ys. Suc j \<le> length ys) xs)"
4948 proof (rule length_takeWhile_less_P_nth)
4949 show "Suc i \<le> length xs" using `i < length xs` by simp
4950 fix k assume "k < Suc i"
4951 hence "k \<le> i" by auto
4952 with sorted_rev_nth_mono[OF sorted this] `i < length xs`
4953 have "length (xs ! i) \<le> length (xs ! k)" by simp
4954 thus "Suc j \<le> length (xs ! k)" using j_less by simp
4956 have i_less_filter: "i < length [ys\<leftarrow>xs . j < length ys]"
4957 unfolding filter_equals_takeWhile_sorted_rev[OF sorted, of j]
4958 using i_less_tW by (simp_all add: Suc_le_eq)
4959 from j show "?R ! j = xs ! i ! j"
4960 unfolding filter_equals_takeWhile_sorted_rev[OF sorted_transpose, of i]
4961 by (simp add: takeWhile_nth nth_nth_transpose_sorted[OF sorted * i_less_filter])
4964 lemma transpose_transpose:
4965 fixes xs :: "'a list list"
4966 assumes sorted: "sorted (rev (map length xs))"
4967 shows "transpose (transpose xs) = takeWhile (\<lambda>x. x \<noteq> []) xs" (is "?L = ?R")
4969 have len: "length ?L = length ?R"
4970 unfolding length_transpose transpose_max_length
4971 using filter_equals_takeWhile_sorted_rev[OF sorted, of 0]
4974 { fix i assume "i < length ?R"
4975 with less_le_trans[OF _ length_takeWhile_le[of _ xs]]
4976 have "i < length xs" by simp
4979 by (rule nth_equalityI)
4980 (simp_all add: len nth_transpose transpose_column[OF sorted] * takeWhile_nth)
4983 theorem transpose_rectangle:
4984 assumes "xs = [] \<Longrightarrow> n = 0"
4985 assumes rect: "\<And> i. i < length xs \<Longrightarrow> length (xs ! i) = n"
4986 shows "transpose xs = map (\<lambda> i. map (\<lambda> j. xs ! j ! i) [0..<length xs]) [0..<n]"
4987 (is "?trans = ?map")
4988 proof (rule nth_equalityI)
4989 have "sorted (rev (map length xs))"
4990 by (auto simp: rev_nth rect intro!: sorted_nth_monoI)
4991 from foldr_max_sorted[OF this] assms
4992 show len: "length ?trans = length ?map"
4993 by (simp_all add: length_transpose foldr_map comp_def)
4995 { fix i assume "i < n" hence "[ys\<leftarrow>xs . i < length ys] = xs"
4996 using rect by (auto simp: in_set_conv_nth intro!: filter_True) }
4997 ultimately show "\<forall>i < length ?trans. ?trans ! i = ?map ! i"
4998 by (auto simp: nth_transpose intro: nth_equalityI)
5002 subsubsection {* @{text sorted_list_of_set} *}
5004 text{* This function maps (finite) linearly ordered sets to sorted
5005 lists. Warning: in most cases it is not a good idea to convert from
5006 sets to lists but one should convert in the other direction (via
5009 subsubsection {* @{text sorted_list_of_set} *}
5011 text{* This function maps (finite) linearly ordered sets to sorted
5012 lists. Warning: in most cases it is not a good idea to convert from
5013 sets to lists but one should convert in the other direction (via
5019 definition sorted_list_of_set :: "'a set \<Rightarrow> 'a list" where
5020 "sorted_list_of_set = folding.F insort []"
5022 sublocale sorted_list_of_set!: folding insort Nil
5024 "folding.F insort [] = sorted_list_of_set"
5026 interpret comp_fun_commute insort by (fact comp_fun_commute_insort)
5027 show "folding insort" by default (fact comp_fun_commute)
5028 show "folding.F insort [] = sorted_list_of_set" by (simp only: sorted_list_of_set_def)
5031 lemma sorted_list_of_set_empty:
5032 "sorted_list_of_set {} = []"
5033 by (fact sorted_list_of_set.empty)
5035 lemma sorted_list_of_set_insert [simp]:
5036 "finite A \<Longrightarrow> sorted_list_of_set (insert x A) = insort x (sorted_list_of_set (A - {x}))"
5037 by (fact sorted_list_of_set.insert_remove)
5039 lemma sorted_list_of_set_eq_Nil_iff [simp]:
5040 "finite A \<Longrightarrow> sorted_list_of_set A = [] \<longleftrightarrow> A = {}"
5041 by (auto simp: sorted_list_of_set.remove)
5043 lemma sorted_list_of_set [simp]:
5044 "finite A \<Longrightarrow> set (sorted_list_of_set A) = A \<and> sorted (sorted_list_of_set A)
5045 \<and> distinct (sorted_list_of_set A)"
5046 by (induct A rule: finite_induct) (simp_all add: set_insort sorted_insort distinct_insort)
5048 lemma distinct_sorted_list_of_set:
5049 "distinct (sorted_list_of_set A)"
5050 using sorted_list_of_set by (cases "finite A") auto
5052 lemma sorted_list_of_set_sort_remdups [code]:
5053 "sorted_list_of_set (set xs) = sort (remdups xs)"
5055 interpret comp_fun_commute insort by (fact comp_fun_commute_insort)
5056 show ?thesis by (simp add: sorted_list_of_set.eq_fold sort_conv_fold fold_set_fold_remdups)
5059 lemma sorted_list_of_set_remove:
5061 shows "sorted_list_of_set (A - {x}) = remove1 x (sorted_list_of_set A)"
5062 proof (cases "x \<in> A")
5063 case False with assms have "x \<notin> set (sorted_list_of_set A)" by simp
5064 with False show ?thesis by (simp add: remove1_idem)
5066 case True then obtain B where A: "A = insert x B" by (rule Set.set_insert)
5067 with assms show ?thesis by simp
5072 lemma sorted_list_of_set_range [simp]:
5073 "sorted_list_of_set {m..<n} = [m..<n]"
5074 by (rule sorted_distinct_set_unique) simp_all
5077 subsubsection {* @{text lists}: the list-forming operator over sets *}
5080 lists :: "'a set => 'a list set"
5083 Nil [intro!, simp]: "[]: lists A"
5084 | Cons [intro!, simp]: "[| a: A; l: lists A|] ==> a#l : lists A"
5086 inductive_cases listsE [elim!]: "x#l : lists A"
5087 inductive_cases listspE [elim!]: "listsp A (x # l)"
5089 inductive_simps listsp_simps[code]:
5093 lemma listsp_mono [mono]: "A \<le> B ==> listsp A \<le> listsp B"
5094 by (rule predicate1I, erule listsp.induct, blast+)
5096 lemmas lists_mono = listsp_mono [to_set]
5099 assumes l: "listsp A l" shows "listsp B l ==> listsp (inf A B) l" using l
5102 lemmas lists_IntI = listsp_infI [to_set]
5104 lemma listsp_inf_eq [simp]: "listsp (inf A B) = inf (listsp A) (listsp B)"
5105 proof (rule mono_inf [where f=listsp, THEN order_antisym])
5106 show "mono listsp" by (simp add: mono_def listsp_mono)
5107 show "inf (listsp A) (listsp B) \<le> listsp (inf A B)" by (blast intro!: listsp_infI)
5110 lemmas listsp_conj_eq [simp] = listsp_inf_eq [simplified inf_fun_def inf_bool_def]
5112 lemmas lists_Int_eq [simp] = listsp_inf_eq [to_set]
5114 lemma Cons_in_lists_iff[simp]: "x#xs : lists A \<longleftrightarrow> x:A \<and> xs : lists A"
5117 lemma append_in_listsp_conv [iff]:
5118 "(listsp A (xs @ ys)) = (listsp A xs \<and> listsp A ys)"
5121 lemmas append_in_lists_conv [iff] = append_in_listsp_conv [to_set]
5123 lemma in_listsp_conv_set: "(listsp A xs) = (\<forall>x \<in> set xs. A x)"
5124 -- {* eliminate @{text listsp} in favour of @{text set} *}
5127 lemmas in_lists_conv_set [code_unfold] = in_listsp_conv_set [to_set]
5129 lemma in_listspD [dest!]: "listsp A xs ==> \<forall>x\<in>set xs. A x"
5130 by (rule in_listsp_conv_set [THEN iffD1])
5132 lemmas in_listsD [dest!] = in_listspD [to_set]
5134 lemma in_listspI [intro!]: "\<forall>x\<in>set xs. A x ==> listsp A xs"
5135 by (rule in_listsp_conv_set [THEN iffD2])
5137 lemmas in_listsI [intro!] = in_listspI [to_set]
5139 lemma lists_eq_set: "lists A = {xs. set xs <= A}"
5142 lemma lists_empty [simp]: "lists {} = {[]}"
5145 lemma lists_UNIV [simp]: "lists UNIV = UNIV"
5148 lemma lists_image: "lists (f`A) = map f ` lists A"
5150 { fix xs have "\<forall>x\<in>set xs. x \<in> f ` A \<Longrightarrow> xs \<in> map f ` lists A"
5151 by (induct xs) (auto simp del: map.simps simp add: map.simps[symmetric] intro!: imageI) }
5152 then show ?thesis by auto
5155 subsubsection {* Inductive definition for membership *}
5157 inductive ListMem :: "'a \<Rightarrow> 'a list \<Rightarrow> bool"
5159 elem: "ListMem x (x # xs)"
5160 | insert: "ListMem x xs \<Longrightarrow> ListMem x (y # xs)"
5162 lemma ListMem_iff: "(ListMem x xs) = (x \<in> set xs)"
5164 apply (induct set: ListMem)
5167 apply (auto intro: ListMem.intros)
5171 subsubsection {* Lists as Cartesian products *}
5173 text{*@{text"set_Cons A Xs"}: the set of lists with head drawn from
5174 @{term A} and tail drawn from @{term Xs}.*}
5176 definition set_Cons :: "'a set \<Rightarrow> 'a list set \<Rightarrow> 'a list set" where
5177 "set_Cons A XS = {z. \<exists>x xs. z = x # xs \<and> x \<in> A \<and> xs \<in> XS}"
5179 lemma set_Cons_sing_Nil [simp]: "set_Cons A {[]} = (%x. [x])`A"
5180 by (auto simp add: set_Cons_def)
5182 text{*Yields the set of lists, all of the same length as the argument and
5183 with elements drawn from the corresponding element of the argument.*}
5185 primrec listset :: "'a set list \<Rightarrow> 'a list set" where
5186 "listset [] = {[]}" |
5187 "listset (A # As) = set_Cons A (listset As)"
5190 subsection {* Relations on Lists *}
5192 subsubsection {* Length Lexicographic Ordering *}
5194 text{*These orderings preserve well-foundedness: shorter lists
5195 precede longer lists. These ordering are not used in dictionaries.*}
5197 primrec -- {*The lexicographic ordering for lists of the specified length*}
5198 lexn :: "('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> ('a list \<times> 'a list) set" where
5201 (map_pair (%(x, xs). x#xs) (%(x, xs). x#xs) ` (r <*lex*> lexn r n)) Int
5202 {(xs, ys). length xs = Suc n \<and> length ys = Suc n}"
5204 definition lex :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
5205 "lex r = (\<Union>n. lexn r n)" -- {*Holds only between lists of the same length*}
5207 definition lenlex :: "('a \<times> 'a) set => ('a list \<times> 'a list) set" where
5208 "lenlex r = inv_image (less_than <*lex*> lex r) (\<lambda>xs. (length xs, xs))"
5209 -- {*Compares lists by their length and then lexicographically*}
5211 lemma wf_lexn: "wf r ==> wf (lexn r n)"
5212 apply (induct n, simp, simp)
5213 apply(rule wf_subset)
5214 prefer 2 apply (rule Int_lower1)
5215 apply(rule wf_map_pair_image)
5216 prefer 2 apply (rule inj_onI, auto)
5220 "(xs, ys) : lexn r n ==> length xs = n \<and> length ys = n"
5221 by (induct n arbitrary: xs ys) auto
5223 lemma wf_lex [intro!]: "wf r ==> wf (lex r)"
5224 apply (unfold lex_def)
5226 apply (blast intro: wf_lexn, clarify)
5227 apply (rename_tac m n)
5228 apply (subgoal_tac "m \<noteq> n")
5229 prefer 2 apply blast
5230 apply (blast dest: lexn_length not_sym)
5235 {(xs,ys). length xs = n \<and> length ys = n \<and>
5236 (\<exists>xys x y xs' ys'. xs= xys @ x#xs' \<and> ys= xys @ y # ys' \<and> (x, y):r)}"
5237 apply (induct n, simp)
5238 apply (simp add: image_Collect lex_prod_def, safe, blast)
5239 apply (rule_tac x = "ab # xys" in exI, simp)
5240 apply (case_tac xys, simp_all, blast)
5245 {(xs,ys). length xs = length ys \<and>
5246 (\<exists>xys x y xs' ys'. xs = xys @ x # xs' \<and> ys = xys @ y # ys' \<and> (x, y):r)}"
5247 by (force simp add: lex_def lexn_conv)
5249 lemma wf_lenlex [intro!]: "wf r ==> wf (lenlex r)"
5250 by (unfold lenlex_def) blast
5253 "lenlex r = {(xs,ys). length xs < length ys |
5254 length xs = length ys \<and> (xs, ys) : lex r}"
5255 by (simp add: lenlex_def Id_on_def lex_prod_def inv_image_def)
5257 lemma Nil_notin_lex [iff]: "([], ys) \<notin> lex r"
5258 by (simp add: lex_conv)
5260 lemma Nil2_notin_lex [iff]: "(xs, []) \<notin> lex r"
5261 by (simp add:lex_conv)
5263 lemma Cons_in_lex [simp]:
5264 "((x # xs, y # ys) : lex r) =
5265 ((x, y) : r \<and> length xs = length ys | x = y \<and> (xs, ys) : lex r)"
5266 apply (simp add: lex_conv)
5268 prefer 2 apply (blast intro: Cons_eq_appendI, clarify)
5269 apply (case_tac xys, simp, simp)
5274 subsubsection {* Lexicographic Ordering *}
5276 text {* Classical lexicographic ordering on lists, ie. "a" < "ab" < "b".
5277 This ordering does \emph{not} preserve well-foundedness.
5278 Author: N. Voelker, March 2005. *}
5280 definition lexord :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
5281 "lexord r = {(x,y). \<exists> a v. y = x @ a # v \<or>
5282 (\<exists> u a b v w. (a,b) \<in> r \<and> x = u @ (a # v) \<and> y = u @ (b # w))}"
5284 lemma lexord_Nil_left[simp]: "([],y) \<in> lexord r = (\<exists> a x. y = a # x)"
5285 by (unfold lexord_def, induct_tac y, auto)
5287 lemma lexord_Nil_right[simp]: "(x,[]) \<notin> lexord r"
5288 by (unfold lexord_def, induct_tac x, auto)
5290 lemma lexord_cons_cons[simp]:
5291 "((a # x, b # y) \<in> lexord r) = ((a,b)\<in> r | (a = b & (x,y)\<in> lexord r))"
5292 apply (unfold lexord_def, safe, simp_all)
5293 apply (case_tac u, simp, simp)
5294 apply (case_tac u, simp, clarsimp, blast, blast, clarsimp)
5295 apply (erule_tac x="b # u" in allE)
5298 lemmas lexord_simps = lexord_Nil_left lexord_Nil_right lexord_cons_cons
5300 lemma lexord_append_rightI: "\<exists> b z. y = b # z \<Longrightarrow> (x, x @ y) \<in> lexord r"
5301 by (induct_tac x, auto)
5303 lemma lexord_append_left_rightI:
5304 "(a,b) \<in> r \<Longrightarrow> (u @ a # x, u @ b # y) \<in> lexord r"
5305 by (induct_tac u, auto)
5307 lemma lexord_append_leftI: " (u,v) \<in> lexord r \<Longrightarrow> (x @ u, x @ v) \<in> lexord r"
5310 lemma lexord_append_leftD:
5311 "\<lbrakk> (x @ u, x @ v) \<in> lexord r; (! a. (a,a) \<notin> r) \<rbrakk> \<Longrightarrow> (u,v) \<in> lexord r"
5312 by (erule rev_mp, induct_tac x, auto)
5314 lemma lexord_take_index_conv:
5315 "((x,y) : lexord r) =
5316 ((length x < length y \<and> take (length x) y = x) \<or>
5317 (\<exists>i. i < min(length x)(length y) & take i x = take i y & (x!i,y!i) \<in> r))"
5318 apply (unfold lexord_def Let_def, clarsimp)
5319 apply (rule_tac f = "(% a b. a \<or> b)" in arg_cong2)
5321 apply (rule_tac x="hd (drop (length x) y)" in exI)
5322 apply (rule_tac x="tl (drop (length x) y)" in exI)
5323 apply (erule subst, simp add: min_def)
5324 apply (rule_tac x ="length u" in exI, simp)
5325 apply (rule_tac x ="take i x" in exI)
5326 apply (rule_tac x ="x ! i" in exI)
5327 apply (rule_tac x ="y ! i" in exI, safe)
5328 apply (rule_tac x="drop (Suc i) x" in exI)
5329 apply (drule sym, simp add: drop_Suc_conv_tl)
5330 apply (rule_tac x="drop (Suc i) y" in exI)
5331 by (simp add: drop_Suc_conv_tl)
5333 -- {* lexord is extension of partial ordering List.lex *}
5334 lemma lexord_lex: "(x,y) \<in> lex r = ((x,y) \<in> lexord r \<and> length x = length y)"
5335 apply (rule_tac x = y in spec)
5336 apply (induct_tac x, clarsimp)
5337 by (clarify, case_tac x, simp, force)
5339 lemma lexord_irreflexive: "ALL x. (x,x) \<notin> r \<Longrightarrow> (xs,xs) \<notin> lexord r"
5342 text{* By Ren\'e Thiemann: *}
5343 lemma lexord_partial_trans:
5344 "(\<And>x y z. x \<in> set xs \<Longrightarrow> (x,y) \<in> r \<Longrightarrow> (y,z) \<in> r \<Longrightarrow> (x,z) \<in> r)
5345 \<Longrightarrow> (xs,ys) \<in> lexord r \<Longrightarrow> (ys,zs) \<in> lexord r \<Longrightarrow> (xs,zs) \<in> lexord r"
5346 proof (induct xs arbitrary: ys zs)
5348 from Nil(3) show ?case unfolding lexord_def by (cases zs, auto)
5350 case (Cons x xs yys zzs)
5351 from Cons(3) obtain y ys where yys: "yys = y # ys" unfolding lexord_def
5352 by (cases yys, auto)
5353 note Cons = Cons[unfolded yys]
5354 from Cons(3) have one: "(x,y) \<in> r \<or> x = y \<and> (xs,ys) \<in> lexord r" by auto
5355 from Cons(4) obtain z zs where zzs: "zzs = z # zs" unfolding lexord_def
5356 by (cases zzs, auto)
5357 note Cons = Cons[unfolded zzs]
5358 from Cons(4) have two: "(y,z) \<in> r \<or> y = z \<and> (ys,zs) \<in> lexord r" by auto
5360 assume "(xs,ys) \<in> lexord r" and "(ys,zs) \<in> lexord r"
5361 from Cons(1)[OF _ this] Cons(2)
5362 have "(xs,zs) \<in> lexord r" by auto
5365 assume "(x,y) \<in> r" and "(y,z) \<in> r"
5366 from Cons(2)[OF _ this] have "(x,z) \<in> r" by auto
5368 from one two ind1 ind2
5369 have "(x,z) \<in> r \<or> x = z \<and> (xs,zs) \<in> lexord r" by blast
5370 thus ?case unfolding zzs by auto
5374 "\<lbrakk> (x, y) \<in> lexord r; (y, z) \<in> lexord r; trans r \<rbrakk> \<Longrightarrow> (x, z) \<in> lexord r"
5375 by(auto simp: trans_def intro:lexord_partial_trans)
5377 lemma lexord_transI: "trans r \<Longrightarrow> trans (lexord r)"
5378 by (rule transI, drule lexord_trans, blast)
5380 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"
5381 apply (rule_tac x = y in spec)
5382 apply (induct_tac x, rule allI)
5383 apply (case_tac x, simp, simp)
5384 apply (rule allI, case_tac x, simp, simp)
5388 Predicate version of lexicographic order integrated with Isabelle's order type classes.
5389 Author: Andreas Lochbihler
5394 inductive lexordp :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
5396 Nil: "lexordp [] (y # ys)"
5397 | Cons: "x < y \<Longrightarrow> lexordp (x # xs) (y # ys)"
5399 "\<lbrakk> \<not> x < y; \<not> y < x; lexordp xs ys \<rbrakk> \<Longrightarrow> lexordp (x # xs) (y # ys)"
5401 lemma lexordp_simps [simp]:
5402 "lexordp [] ys = (ys \<noteq> [])"
5403 "lexordp xs [] = False"
5404 "lexordp (x # xs) (y # ys) \<longleftrightarrow> x < y \<or> \<not> y < x \<and> lexordp xs ys"
5405 by(subst lexordp.simps, fastforce simp add: neq_Nil_conv)+
5407 inductive lexordp_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
5408 Nil: "lexordp_eq [] ys"
5409 | Cons: "x < y \<Longrightarrow> lexordp_eq (x # xs) (y # ys)"
5410 | Cons_eq: "\<lbrakk> \<not> x < y; \<not> y < x; lexordp_eq xs ys \<rbrakk> \<Longrightarrow> lexordp_eq (x # xs) (y # ys)"
5412 lemma lexordp_eq_simps [simp]:
5413 "lexordp_eq [] ys = True"
5414 "lexordp_eq xs [] \<longleftrightarrow> xs = []"
5415 "lexordp_eq (x # xs) [] = False"
5416 "lexordp_eq (x # xs) (y # ys) \<longleftrightarrow> x < y \<or> \<not> y < x \<and> lexordp_eq xs ys"
5417 by(subst lexordp_eq.simps, fastforce)+
5419 lemma lexordp_append_rightI: "ys \<noteq> Nil \<Longrightarrow> lexordp xs (xs @ ys)"
5420 by(induct xs)(auto simp add: neq_Nil_conv)
5422 lemma lexordp_append_left_rightI: "x < y \<Longrightarrow> lexordp (us @ x # xs) (us @ y # ys)"
5425 lemma lexordp_eq_refl: "lexordp_eq xs xs"
5426 by(induct xs) simp_all
5428 lemma lexordp_append_leftI: "lexordp us vs \<Longrightarrow> lexordp (xs @ us) (xs @ vs)"
5431 lemma lexordp_append_leftD: "\<lbrakk> lexordp (xs @ us) (xs @ vs); \<forall>a. \<not> a < a \<rbrakk> \<Longrightarrow> lexordp us vs"
5434 lemma lexordp_irreflexive:
5435 assumes irrefl: "\<forall>x. \<not> x < x"
5436 shows "\<not> lexordp xs xs"
5438 assume "lexordp xs xs"
5439 thus False by(induct xs ys\<equiv>xs)(simp_all add: irrefl)
5442 lemma lexordp_into_lexordp_eq:
5443 assumes "lexordp xs ys"
5444 shows "lexordp_eq xs ys"
5445 using assms by induct simp_all
5449 declare ord.lexordp_simps [simp, code]
5450 declare ord.lexordp_eq_simps [code, simp]
5452 lemma lexord_code [code, code_unfold]: "lexordp = ord.lexordp less"
5453 unfolding lexordp_def ord.lexordp_def ..
5457 lemma lexordp_antisym:
5458 assumes "lexordp xs ys" "lexordp ys xs"
5460 using assms by induct auto
5462 lemma lexordp_irreflexive': "\<not> lexordp xs xs"
5463 by(rule lexordp_irreflexive) simp
5467 context linorder begin
5469 lemma lexordp_cases [consumes 1, case_names Nil Cons Cons_eq, cases pred: lexordp]:
5470 assumes "lexordp xs ys"
5471 obtains (Nil) y ys' where "xs = []" "ys = y # ys'"
5472 | (Cons) x xs' y ys' where "xs = x # xs'" "ys = y # ys'" "x < y"
5473 | (Cons_eq) x xs' ys' where "xs = x # xs'" "ys = x # ys'" "lexordp xs' ys'"
5474 using assms by cases (fastforce simp add: not_less_iff_gr_or_eq)+
5476 lemma lexordp_induct [consumes 1, case_names Nil Cons Cons_eq, induct pred: lexordp]:
5477 assumes major: "lexordp xs ys"
5478 and Nil: "\<And>y ys. P [] (y # ys)"
5479 and Cons: "\<And>x xs y ys. x < y \<Longrightarrow> P (x # xs) (y # ys)"
5480 and Cons_eq: "\<And>x xs ys. \<lbrakk> lexordp xs ys; P xs ys \<rbrakk> \<Longrightarrow> P (x # xs) (x # ys)"
5482 using major by induct (simp_all add: Nil Cons not_less_iff_gr_or_eq Cons_eq)
5485 "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)"
5488 assume ?lhs thus ?rhs
5490 case Cons_eq thus ?case by simp (metis append.simps(2))
5491 qed(fastforce intro: disjI2 del: disjCI intro: exI[where x="[]"])+
5493 assume ?rhs thus ?lhs
5494 by(auto intro: lexordp_append_leftI[where us="[]", simplified] lexordp_append_leftI)
5497 lemma lexordp_conv_lexord:
5498 "lexordp xs ys \<longleftrightarrow> (xs, ys) \<in> lexord {(x, y). x < y}"
5499 by(simp add: lexordp_iff lexord_def)
5501 lemma lexordp_eq_antisym:
5502 assumes "lexordp_eq xs ys" "lexordp_eq ys xs"
5504 using assms by induct simp_all
5506 lemma lexordp_eq_trans:
5507 assumes "lexordp_eq xs ys" and "lexordp_eq ys zs"
5508 shows "lexordp_eq xs zs"
5510 apply(induct arbitrary: zs)
5511 apply(case_tac [2-3] zs)
5515 lemma lexordp_trans:
5516 assumes "lexordp xs ys" "lexordp ys zs"
5517 shows "lexordp xs zs"
5519 apply(induct arbitrary: zs)
5520 apply(case_tac [2-3] zs)
5524 lemma lexordp_linear: "lexordp xs ys \<or> xs = ys \<or> lexordp ys xs"
5525 proof(induct xs arbitrary: ys)
5526 case Nil thus ?case by(cases ys) simp_all
5528 case Cons thus ?case by(cases ys) auto
5531 lemma lexordp_conv_lexordp_eq: "lexordp xs ys \<longleftrightarrow> lexordp_eq xs ys \<and> \<not> lexordp_eq ys xs"
5532 (is "?lhs \<longleftrightarrow> ?rhs")
5535 moreover hence "\<not> lexordp_eq ys xs" by induct simp_all
5536 ultimately show ?rhs by(simp add: lexordp_into_lexordp_eq)
5539 hence "lexordp_eq xs ys" "\<not> lexordp_eq ys xs" by simp_all
5540 thus ?lhs by induct simp_all
5543 lemma lexordp_eq_conv_lexord: "lexordp_eq xs ys \<longleftrightarrow> xs = ys \<or> lexordp xs ys"
5544 by(auto simp add: lexordp_conv_lexordp_eq lexordp_eq_refl dest: lexordp_eq_antisym)
5546 lemma lexordp_eq_linear: "lexordp_eq xs ys \<or> lexordp_eq ys xs"
5547 apply(induct xs arbitrary: ys)
5548 apply(case_tac [!] ys)
5552 lemma lexordp_linorder: "class.linorder lexordp_eq lexordp"
5553 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)
5557 subsubsection {* Lexicographic combination of measure functions *}
5559 text {* These are useful for termination proofs *}
5561 definition "measures fs = inv_image (lex less_than) (%a. map (%f. f a) fs)"
5563 lemma wf_measures[simp]: "wf (measures fs)"
5564 unfolding measures_def
5567 lemma in_measures[simp]:
5568 "(x, y) \<in> measures [] = False"
5569 "(x, y) \<in> measures (f # fs)
5570 = (f x < f y \<or> (f x = f y \<and> (x, y) \<in> measures fs))"
5571 unfolding measures_def
5574 lemma measures_less: "f x < f y ==> (x, y) \<in> measures (f#fs)"
5577 lemma measures_lesseq: "f x <= f y ==> (x, y) \<in> measures fs ==> (x, y) \<in> measures (f#fs)"
5581 subsubsection {* Lifting Relations to Lists: one element *}
5583 definition listrel1 :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
5584 "listrel1 r = {(xs,ys).
5585 \<exists>us z z' vs. xs = us @ z # vs \<and> (z,z') \<in> r \<and> ys = us @ z' # vs}"
5588 "\<lbrakk> (x, y) \<in> r; xs = us @ x # vs; ys = us @ y # vs \<rbrakk> \<Longrightarrow>
5589 (xs, ys) \<in> listrel1 r"
5590 unfolding listrel1_def by auto
5593 "\<lbrakk> (xs, ys) \<in> listrel1 r;
5594 !!x y us vs. \<lbrakk> (x, y) \<in> r; xs = us @ x # vs; ys = us @ y # vs \<rbrakk> \<Longrightarrow> P
5595 \<rbrakk> \<Longrightarrow> P"
5596 unfolding listrel1_def by auto
5598 lemma not_Nil_listrel1 [iff]: "([], xs) \<notin> listrel1 r"
5599 unfolding listrel1_def by blast
5601 lemma not_listrel1_Nil [iff]: "(xs, []) \<notin> listrel1 r"
5602 unfolding listrel1_def by blast
5604 lemma Cons_listrel1_Cons [iff]:
5605 "(x # xs, y # ys) \<in> listrel1 r \<longleftrightarrow>
5606 (x,y) \<in> r \<and> xs = ys \<or> x = y \<and> (xs, ys) \<in> listrel1 r"
5607 by (simp add: listrel1_def Cons_eq_append_conv) (blast)
5609 lemma listrel1I1: "(x,y) \<in> r \<Longrightarrow> (x # xs, y # xs) \<in> listrel1 r"
5610 by (metis Cons_listrel1_Cons)
5612 lemma listrel1I2: "(xs, ys) \<in> listrel1 r \<Longrightarrow> (x # xs, x # ys) \<in> listrel1 r"
5613 by (metis Cons_listrel1_Cons)
5615 lemma append_listrel1I:
5616 "(xs, ys) \<in> listrel1 r \<and> us = vs \<or> xs = ys \<and> (us, vs) \<in> listrel1 r
5617 \<Longrightarrow> (xs @ us, ys @ vs) \<in> listrel1 r"
5618 unfolding listrel1_def
5619 by auto (blast intro: append_eq_appendI)+
5621 lemma Cons_listrel1E1[elim!]:
5622 assumes "(x # xs, ys) \<in> listrel1 r"
5623 and "\<And>y. ys = y # xs \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> R"
5624 and "\<And>zs. ys = x # zs \<Longrightarrow> (xs, zs) \<in> listrel1 r \<Longrightarrow> R"
5626 using assms by (cases ys) blast+
5628 lemma Cons_listrel1E2[elim!]:
5629 assumes "(xs, y # ys) \<in> listrel1 r"
5630 and "\<And>x. xs = x # ys \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> R"
5631 and "\<And>zs. xs = y # zs \<Longrightarrow> (zs, ys) \<in> listrel1 r \<Longrightarrow> R"
5633 using assms by (cases xs) blast+
5635 lemma snoc_listrel1_snoc_iff:
5636 "(xs @ [x], ys @ [y]) \<in> listrel1 r
5637 \<longleftrightarrow> (xs, ys) \<in> listrel1 r \<and> x = y \<or> xs = ys \<and> (x,y) \<in> r" (is "?L \<longleftrightarrow> ?R")
5640 by (fastforce simp: listrel1_def snoc_eq_iff_butlast butlast_append)
5642 assume ?R then show ?L unfolding listrel1_def by force
5645 lemma listrel1_eq_len: "(xs,ys) \<in> listrel1 r \<Longrightarrow> length xs = length ys"
5646 unfolding listrel1_def by auto
5648 lemma listrel1_mono:
5649 "r \<subseteq> s \<Longrightarrow> listrel1 r \<subseteq> listrel1 s"
5650 unfolding listrel1_def by blast
5653 lemma listrel1_converse: "listrel1 (r^-1) = (listrel1 r)^-1"
5654 unfolding listrel1_def by blast
5656 lemma in_listrel1_converse:
5657 "(x,y) : listrel1 (r^-1) \<longleftrightarrow> (x,y) : (listrel1 r)^-1"
5658 unfolding listrel1_def by blast
5660 lemma listrel1_iff_update:
5661 "(xs,ys) \<in> (listrel1 r)
5662 \<longleftrightarrow> (\<exists>y n. (xs ! n, y) \<in> r \<and> n < length xs \<and> ys = xs[n:=y])" (is "?L \<longleftrightarrow> ?R")
5665 then obtain x y u v where "xs = u @ x # v" "ys = u @ y # v" "(x,y) \<in> r"
5666 unfolding listrel1_def by auto
5667 then have "ys = xs[length u := y]" and "length u < length xs"
5668 and "(xs ! length u, y) \<in> r" by auto
5669 then show "?R" by auto
5672 then obtain x y n where "(xs!n, y) \<in> r" "n < size xs" "ys = xs[n:=y]" "x = xs!n"
5674 then obtain u v where "xs = u @ x # v" and "ys = u @ y # v" and "(x, y) \<in> r"
5675 by (auto intro: upd_conv_take_nth_drop id_take_nth_drop)
5676 then show "?L" by (auto simp: listrel1_def)
5680 text{* Accessible part and wellfoundedness: *}
5682 lemma Cons_acc_listrel1I [intro!]:
5683 "x \<in> Wellfounded.acc r \<Longrightarrow> xs \<in> Wellfounded.acc (listrel1 r) \<Longrightarrow> (x # xs) \<in> Wellfounded.acc (listrel1 r)"
5684 apply (induct arbitrary: xs set: Wellfounded.acc)
5685 apply (erule thin_rl)
5686 apply (erule acc_induct)
5691 lemma lists_accD: "xs \<in> lists (Wellfounded.acc r) \<Longrightarrow> xs \<in> Wellfounded.acc (listrel1 r)"
5692 apply (induct set: lists)
5696 apply (fast dest: acc_downward)
5699 lemma lists_accI: "xs \<in> Wellfounded.acc (listrel1 r) \<Longrightarrow> xs \<in> lists (Wellfounded.acc r)"
5700 apply (induct set: Wellfounded.acc)
5703 apply (fastforce dest!: in_set_conv_decomp[THEN iffD1] simp: listrel1_def)
5706 lemma wf_listrel1_iff[simp]: "wf(listrel1 r) = wf r"
5707 by(metis wf_acc_iff in_lists_conv_set lists_accI lists_accD Cons_in_lists_iff)
5710 subsubsection {* Lifting Relations to Lists: all elements *}
5713 listrel :: "('a \<times> 'b) set \<Rightarrow> ('a list \<times> 'b list) set"
5714 for r :: "('a \<times> 'b) set"
5716 Nil: "([],[]) \<in> listrel r"
5717 | Cons: "[| (x,y) \<in> r; (xs,ys) \<in> listrel r |] ==> (x#xs, y#ys) \<in> listrel r"
5719 inductive_cases listrel_Nil1 [elim!]: "([],xs) \<in> listrel r"
5720 inductive_cases listrel_Nil2 [elim!]: "(xs,[]) \<in> listrel r"
5721 inductive_cases listrel_Cons1 [elim!]: "(y#ys,xs) \<in> listrel r"
5722 inductive_cases listrel_Cons2 [elim!]: "(xs,y#ys) \<in> listrel r"
5725 lemma listrel_eq_len: "(xs, ys) \<in> listrel r \<Longrightarrow> length xs = length ys"
5726 by(induct rule: listrel.induct) auto
5728 lemma listrel_iff_zip [code_unfold]: "(xs,ys) : listrel r \<longleftrightarrow>
5729 length xs = length ys & (\<forall>(x,y) \<in> set(zip xs ys). (x,y) \<in> r)" (is "?L \<longleftrightarrow> ?R")
5731 assume ?L thus ?R by induct (auto intro: listrel_eq_len)
5735 by (induct rule: list_induct2) (auto intro: listrel.intros)
5738 lemma listrel_iff_nth: "(xs,ys) : listrel r \<longleftrightarrow>
5739 length xs = length ys & (\<forall>n < length xs. (xs!n, ys!n) \<in> r)" (is "?L \<longleftrightarrow> ?R")
5740 by (auto simp add: all_set_conv_all_nth listrel_iff_zip)
5743 lemma listrel_mono: "r \<subseteq> s \<Longrightarrow> listrel r \<subseteq> listrel s"
5745 apply (erule listrel.induct)
5746 apply (blast intro: listrel.intros)+
5749 lemma listrel_subset: "r \<subseteq> A \<times> A \<Longrightarrow> listrel r \<subseteq> lists A \<times> lists A"
5751 apply (erule listrel.induct, auto)
5754 lemma listrel_refl_on: "refl_on A r \<Longrightarrow> refl_on (lists A) (listrel r)"
5755 apply (simp add: refl_on_def listrel_subset Ball_def)
5757 apply (induct_tac x)
5758 apply (auto intro: listrel.intros)
5761 lemma listrel_sym: "sym r \<Longrightarrow> sym (listrel r)"
5762 apply (auto simp add: sym_def)
5763 apply (erule listrel.induct)
5764 apply (blast intro: listrel.intros)+
5767 lemma listrel_trans: "trans r \<Longrightarrow> trans (listrel r)"
5768 apply (simp add: trans_def)
5771 apply (erule listrel.induct)
5772 apply (blast intro: listrel.intros)+
5775 theorem equiv_listrel: "equiv A r \<Longrightarrow> equiv (lists A) (listrel r)"
5776 by (simp add: equiv_def listrel_refl_on listrel_sym listrel_trans)
5778 lemma listrel_rtrancl_refl[iff]: "(xs,xs) : listrel(r^*)"
5779 using listrel_refl_on[of UNIV, OF refl_rtrancl]
5780 by(auto simp: refl_on_def)
5782 lemma listrel_rtrancl_trans:
5783 "\<lbrakk> (xs,ys) : listrel(r^*); (ys,zs) : listrel(r^*) \<rbrakk>
5784 \<Longrightarrow> (xs,zs) : listrel(r^*)"
5785 by (metis listrel_trans trans_def trans_rtrancl)
5788 lemma listrel_Nil [simp]: "listrel r `` {[]} = {[]}"
5789 by (blast intro: listrel.intros)
5792 "listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})"
5793 by (auto simp add: set_Cons_def intro: listrel.intros)
5795 text {* Relating @{term listrel1}, @{term listrel} and closures: *}
5797 lemma listrel1_rtrancl_subset_rtrancl_listrel1:
5798 "listrel1 (r^*) \<subseteq> (listrel1 r)^*"
5799 proof (rule subrelI)
5800 fix xs ys assume 1: "(xs,ys) \<in> listrel1 (r^*)"
5802 have "(x,y) : r^* \<Longrightarrow> (us @ x # vs, us @ y # vs) : (listrel1 r)^*"
5803 proof(induct rule: rtrancl.induct)
5804 case rtrancl_refl show ?case by simp
5806 case rtrancl_into_rtrancl thus ?case
5807 by (metis listrel1I rtrancl.rtrancl_into_rtrancl)
5809 thus "(xs,ys) \<in> (listrel1 r)^*" using 1 by(blast elim: listrel1E)
5812 lemma rtrancl_listrel1_eq_len: "(x,y) \<in> (listrel1 r)^* \<Longrightarrow> length x = length y"
5813 by (induct rule: rtrancl.induct) (auto intro: listrel1_eq_len)
5815 lemma rtrancl_listrel1_ConsI1:
5816 "(xs,ys) : (listrel1 r)^* \<Longrightarrow> (x#xs,x#ys) : (listrel1 r)^*"
5817 apply(induct rule: rtrancl.induct)
5819 by (metis listrel1I2 rtrancl.rtrancl_into_rtrancl)
5821 lemma rtrancl_listrel1_ConsI2:
5822 "(x,y) \<in> r^* \<Longrightarrow> (xs, ys) \<in> (listrel1 r)^*
5823 \<Longrightarrow> (x # xs, y # ys) \<in> (listrel1 r)^*"
5824 by (blast intro: rtrancl_trans rtrancl_listrel1_ConsI1
5825 subsetD[OF listrel1_rtrancl_subset_rtrancl_listrel1 listrel1I1])
5827 lemma listrel1_subset_listrel:
5828 "r \<subseteq> r' \<Longrightarrow> refl r' \<Longrightarrow> listrel1 r \<subseteq> listrel(r')"
5829 by(auto elim!: listrel1E simp add: listrel_iff_zip set_zip refl_on_def)
5831 lemma listrel_reflcl_if_listrel1:
5832 "(xs,ys) : listrel1 r \<Longrightarrow> (xs,ys) : listrel(r^*)"
5833 by(erule listrel1E)(auto simp add: listrel_iff_zip set_zip)
5835 lemma listrel_rtrancl_eq_rtrancl_listrel1: "listrel (r^*) = (listrel1 r)^*"
5837 { fix x y assume "(x,y) \<in> listrel (r^*)"
5838 then have "(x,y) \<in> (listrel1 r)^*"
5839 by induct (auto intro: rtrancl_listrel1_ConsI2) }
5840 then show "listrel (r^*) \<subseteq> (listrel1 r)^*"
5843 show "listrel (r^*) \<supseteq> (listrel1 r)^*"
5845 fix xs ys assume "(xs,ys) \<in> (listrel1 r)^*"
5846 then show "(xs,ys) \<in> listrel (r^*)"
5848 case base show ?case by(auto simp add: listrel_iff_zip set_zip)
5851 thus ?case by (metis listrel_reflcl_if_listrel1 listrel_rtrancl_trans)
5856 lemma rtrancl_listrel1_if_listrel:
5857 "(xs,ys) : listrel r \<Longrightarrow> (xs,ys) : (listrel1 r)^*"
5858 by(metis listrel_rtrancl_eq_rtrancl_listrel1 subsetD[OF listrel_mono] r_into_rtrancl subsetI)
5860 lemma listrel_subset_rtrancl_listrel1: "listrel r \<subseteq> (listrel1 r)^*"
5861 by(fast intro:rtrancl_listrel1_if_listrel)
5864 subsection {* Size function *}
5866 lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (list_size f)"
5867 by (rule is_measure_trivial)
5869 lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (option_size f)"
5870 by (rule is_measure_trivial)
5872 lemma list_size_estimation[termination_simp]:
5873 "x \<in> set xs \<Longrightarrow> y < f x \<Longrightarrow> y < list_size f xs"
5876 lemma list_size_estimation'[termination_simp]:
5877 "x \<in> set xs \<Longrightarrow> y \<le> f x \<Longrightarrow> y \<le> list_size f xs"
5880 lemma list_size_map[simp]: "list_size f (map g xs) = list_size (f o g) xs"
5883 lemma list_size_append[simp]: "list_size f (xs @ ys) = list_size f xs + list_size f ys"
5884 by (induct xs, auto)
5886 lemma list_size_pointwise[termination_simp]:
5887 "(\<And>x. x \<in> set xs \<Longrightarrow> f x \<le> g x) \<Longrightarrow> list_size f xs \<le> list_size g xs"
5888 by (induct xs) force+
5891 subsection {* Monad operation *}
5893 definition bind :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
5894 "bind xs f = concat (map f xs)"
5896 hide_const (open) bind
5898 lemma bind_simps [simp]:
5899 "List.bind [] f = []"
5900 "List.bind (x # xs) f = f x @ List.bind xs f"
5901 by (simp_all add: bind_def)
5904 subsection {* Transfer *}
5906 definition embed_list :: "nat list \<Rightarrow> int list" where
5907 "embed_list l = map int l"
5909 definition nat_list :: "int list \<Rightarrow> bool" where
5910 "nat_list l = nat_set (set l)"
5912 definition return_list :: "int list \<Rightarrow> nat list" where
5913 "return_list l = map nat l"
5915 lemma transfer_nat_int_list_return_embed: "nat_list l \<longrightarrow>
5916 embed_list (return_list l) = l"
5917 unfolding embed_list_def return_list_def nat_list_def nat_set_def
5922 lemma transfer_nat_int_list_functions:
5923 "l @ m = return_list (embed_list l @ embed_list m)"
5924 "[] = return_list []"
5925 unfolding return_list_def embed_list_def
5927 apply (induct l, auto)
5928 apply (induct m, auto)
5932 lemma transfer_nat_int_fold1: "fold f l x =
5933 fold (%x. f (nat x)) (embed_list l) x";
5937 subsection {* Code generation *}
5940 text{* Optional tail recursive version of @{const map}. Can avoid
5941 stack overflow in some target languages. *}
5943 fun map_tailrec_rev :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'b list" where
5944 "map_tailrec_rev f [] bs = bs" |
5945 "map_tailrec_rev f (a#as) bs = map_tailrec_rev f as (f a # bs)"
5947 lemma map_tailrec_rev:
5948 "map_tailrec_rev f as bs = rev(map f as) @ bs"
5949 by(induction as arbitrary: bs) simp_all
5951 definition map_tailrec :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
5952 "map_tailrec f as = rev (map_tailrec_rev f as [])"
5954 text{* Code equation: *}
5955 lemma map_eq_map_tailrec: "map = map_tailrec"
5956 by(simp add: fun_eq_iff map_tailrec_def map_tailrec_rev)
5959 subsubsection {* Counterparts for set-related operations *}
5961 definition member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
5962 [code_abbrev]: "member xs x \<longleftrightarrow> x \<in> set xs"
5965 Use @{text member} only for generating executable code. Otherwise use
5966 @{prop "x \<in> set xs"} instead --- it is much easier to reason about.
5969 lemma member_rec [code]:
5970 "member (x # xs) y \<longleftrightarrow> x = y \<or> member xs y"
5971 "member [] y \<longleftrightarrow> False"
5972 by (auto simp add: member_def)
5974 lemma in_set_member (* FIXME delete candidate *):
5975 "x \<in> set xs \<longleftrightarrow> member xs x"
5976 by (simp add: member_def)
5978 definition list_all :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
5979 list_all_iff [code_abbrev]: "list_all P xs \<longleftrightarrow> Ball (set xs) P"
5981 definition list_ex :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
5982 list_ex_iff [code_abbrev]: "list_ex P xs \<longleftrightarrow> Bex (set xs) P"
5984 definition list_ex1 :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
5985 list_ex1_iff [code_abbrev]: "list_ex1 P xs \<longleftrightarrow> (\<exists>! x. x \<in> set xs \<and> P x)"
5988 Usually you should prefer @{text "\<forall>x\<in>set xs"}, @{text "\<exists>x\<in>set xs"}
5989 and @{text "\<exists>!x. x\<in>set xs \<and> _"} over @{const list_all}, @{const list_ex}
5990 and @{const list_ex1} in specifications.
5993 lemma list_all_simps [simp, code]:
5994 "list_all P (x # xs) \<longleftrightarrow> P x \<and> list_all P xs"
5995 "list_all P [] \<longleftrightarrow> True"
5996 by (simp_all add: list_all_iff)
5998 lemma list_ex_simps [simp, code]:
5999 "list_ex P (x # xs) \<longleftrightarrow> P x \<or> list_ex P xs"
6000 "list_ex P [] \<longleftrightarrow> False"
6001 by (simp_all add: list_ex_iff)
6003 lemma list_ex1_simps [simp, code]:
6004 "list_ex1 P [] = False"
6005 "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)"
6006 by (auto simp add: list_ex1_iff list_all_iff)
6008 lemma Ball_set_list_all: (* FIXME delete candidate *)
6009 "Ball (set xs) P \<longleftrightarrow> list_all P xs"
6010 by (simp add: list_all_iff)
6012 lemma Bex_set_list_ex: (* FIXME delete candidate *)
6013 "Bex (set xs) P \<longleftrightarrow> list_ex P xs"
6014 by (simp add: list_ex_iff)
6016 lemma list_all_append [simp]:
6017 "list_all P (xs @ ys) \<longleftrightarrow> list_all P xs \<and> list_all P ys"
6018 by (auto simp add: list_all_iff)
6020 lemma list_ex_append [simp]:
6021 "list_ex P (xs @ ys) \<longleftrightarrow> list_ex P xs \<or> list_ex P ys"
6022 by (auto simp add: list_ex_iff)
6024 lemma list_all_rev [simp]:
6025 "list_all P (rev xs) \<longleftrightarrow> list_all P xs"
6026 by (simp add: list_all_iff)
6028 lemma list_ex_rev [simp]:
6029 "list_ex P (rev xs) \<longleftrightarrow> list_ex P xs"
6030 by (simp add: list_ex_iff)
6032 lemma list_all_length:
6033 "list_all P xs \<longleftrightarrow> (\<forall>n < length xs. P (xs ! n))"
6034 by (auto simp add: list_all_iff set_conv_nth)
6036 lemma list_ex_length:
6037 "list_ex P xs \<longleftrightarrow> (\<exists>n < length xs. P (xs ! n))"
6038 by (auto simp add: list_ex_iff set_conv_nth)
6040 lemma list_all_cong [fundef_cong]:
6041 "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> list_all f xs = list_all g ys"
6042 by (simp add: list_all_iff)
6044 lemma list_ex_cong [fundef_cong]:
6045 "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> list_ex f xs = list_ex g ys"
6046 by (simp add: list_ex_iff)
6048 definition can_select :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool" where
6049 [code_abbrev]: "can_select P A = (\<exists>!x\<in>A. P x)"
6051 lemma can_select_set_list_ex1 [code]:
6052 "can_select P (set A) = list_ex1 P A"
6053 by (simp add: list_ex1_iff can_select_def)
6056 text {* Executable checks for relations on sets *}
6058 definition listrel1p :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
6059 "listrel1p r xs ys = ((xs, ys) \<in> listrel1 {(x, y). r x y})"
6061 lemma [code_unfold]:
6062 "(xs, ys) \<in> listrel1 r = listrel1p (\<lambda>x y. (x, y) \<in> r) xs ys"
6063 unfolding listrel1p_def by auto
6066 "listrel1p r [] xs = False"
6067 "listrel1p r xs [] = False"
6068 "listrel1p r (x # xs) (y # ys) \<longleftrightarrow>
6069 r x y \<and> xs = ys \<or> x = y \<and> listrel1p r xs ys"
6070 by (simp add: listrel1p_def)+
6073 lexordp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
6074 "lexordp r xs ys = ((xs, ys) \<in> lexord {(x, y). r x y})"
6076 lemma [code_unfold]:
6077 "(xs, ys) \<in> lexord r = lexordp (\<lambda>x y. (x, y) \<in> r) xs ys"
6078 unfolding lexordp_def by auto
6081 "lexordp r xs [] = False"
6082 "lexordp r [] (y#ys) = True"
6083 "lexordp r (x # xs) (y # ys) = (r x y | (x = y & lexordp r xs ys))"
6084 unfolding lexordp_def by auto
6086 text {* Bounded quantification and summation over nats. *}
6088 lemma atMost_upto [code_unfold]:
6089 "{..n} = set [0..<Suc n]"
6092 lemma atLeast_upt [code_unfold]:
6093 "{..<n} = set [0..<n]"
6096 lemma greaterThanLessThan_upt [code_unfold]:
6097 "{n<..<m} = set [Suc n..<m]"
6100 lemmas atLeastLessThan_upt [code_unfold] = set_upt [symmetric]
6102 lemma greaterThanAtMost_upt [code_unfold]:
6103 "{n<..m} = set [Suc n..<Suc m]"
6106 lemma atLeastAtMost_upt [code_unfold]:
6107 "{n..m} = set [n..<Suc m]"
6110 lemma all_nat_less_eq [code_unfold]:
6111 "(\<forall>m<n\<Colon>nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..<n}. P m)"
6114 lemma ex_nat_less_eq [code_unfold]:
6115 "(\<exists>m<n\<Colon>nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..<n}. P m)"
6118 lemma all_nat_less [code_unfold]:
6119 "(\<forall>m\<le>n\<Colon>nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..n}. P m)"
6122 lemma ex_nat_less [code_unfold]:
6123 "(\<exists>m\<le>n\<Colon>nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..n}. P m)"
6126 lemma setsum_set_upt_conv_listsum_nat [code_unfold]:
6127 "setsum f (set [m..<n]) = listsum (map f [m..<n])"
6128 by (simp add: interv_listsum_conv_setsum_set_nat)
6130 text{* Bounded @{text LEAST} operator: *}
6132 definition "Bleast S P = (LEAST x. x \<in> S \<and> P x)"
6134 definition "abort_Bleast S P = (LEAST x. x \<in> S \<and> P x)"
6136 declare [[code abort: abort_Bleast]]
6138 lemma Bleast_code [code]:
6139 "Bleast (set xs) P = (case filter P (sort xs) of
6140 x#xs \<Rightarrow> x |
6141 [] \<Rightarrow> abort_Bleast (set xs) P)"
6142 proof (cases "filter P (sort xs)")
6143 case Nil thus ?thesis by (simp add: Bleast_def abort_Bleast_def)
6146 have "(LEAST x. x \<in> set xs \<and> P x) = x"
6147 proof (rule Least_equality)
6148 show "x \<in> set xs \<and> P x"
6149 by (metis Cons Cons_eq_filter_iff in_set_conv_decomp set_sort)
6151 fix y assume "y : set xs \<and> P y"
6152 hence "y : set (filter P xs)" by auto
6154 by (metis Cons eq_iff filter_sort set_ConsD set_sort sorted_Cons sorted_sort)
6156 thus ?thesis using Cons by (simp add: Bleast_def)
6159 declare Bleast_def[symmetric, code_unfold]
6161 text {* Summation over ints. *}
6163 lemma greaterThanLessThan_upto [code_unfold]:
6164 "{i<..<j::int} = set [i+1..j - 1]"
6167 lemma atLeastLessThan_upto [code_unfold]:
6168 "{i..<j::int} = set [i..j - 1]"
6171 lemma greaterThanAtMost_upto [code_unfold]:
6172 "{i<..j::int} = set [i+1..j]"
6175 lemmas atLeastAtMost_upto [code_unfold] = set_upto [symmetric]
6177 lemma setsum_set_upto_conv_listsum_int [code_unfold]:
6178 "setsum f (set [i..j::int]) = listsum (map f [i..j])"
6179 by (simp add: interv_listsum_conv_setsum_set_int)
6182 subsubsection {* Optimizing by rewriting *}
6184 definition null :: "'a list \<Rightarrow> bool" where
6185 [code_abbrev]: "null xs \<longleftrightarrow> xs = []"
6188 Efficient emptyness check is implemented by @{const null}.
6191 lemma null_rec [code]:
6192 "null (x # xs) \<longleftrightarrow> False"
6193 "null [] \<longleftrightarrow> True"
6194 by (simp_all add: null_def)
6196 lemma eq_Nil_null: (* FIXME delete candidate *)
6197 "xs = [] \<longleftrightarrow> null xs"
6198 by (simp add: null_def)
6200 lemma equal_Nil_null [code_unfold]:
6201 "HOL.equal xs [] \<longleftrightarrow> null xs"
6202 "HOL.equal [] = null"
6203 by (auto simp add: equal null_def)
6205 definition maps :: "('a \<Rightarrow> 'b list) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
6206 [code_abbrev]: "maps f xs = concat (map f xs)"
6208 definition map_filter :: "('a \<Rightarrow> 'b option) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
6209 [code_post]: "map_filter f xs = map (the \<circ> f) (filter (\<lambda>x. f x \<noteq> None) xs)"
6212 Operations @{const maps} and @{const map_filter} avoid
6213 intermediate lists on execution -- do not use for proving.
6216 lemma maps_simps [code]:
6217 "maps f (x # xs) = f x @ maps f xs"
6219 by (simp_all add: maps_def)
6221 lemma map_filter_simps [code]:
6222 "map_filter f (x # xs) = (case f x of None \<Rightarrow> map_filter f xs | Some y \<Rightarrow> y # map_filter f xs)"
6223 "map_filter f [] = []"
6224 by (simp_all add: map_filter_def split: option.split)
6226 lemma concat_map_maps: (* FIXME delete candidate *)
6227 "concat (map f xs) = maps f xs"
6228 by (simp add: maps_def)
6230 lemma map_filter_map_filter [code_unfold]:
6231 "map f (filter P xs) = map_filter (\<lambda>x. if P x then Some (f x) else None) xs"
6232 by (simp add: map_filter_def)
6234 text {* Optimized code for @{text"\<forall>i\<in>{a..b::int}"} and @{text"\<forall>n:{a..<b::nat}"}
6235 and similiarly for @{text"\<exists>"}. *}
6237 definition all_interval_nat :: "(nat \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
6238 "all_interval_nat P i j \<longleftrightarrow> (\<forall>n \<in> {i..<j}. P n)"
6241 "all_interval_nat P i j \<longleftrightarrow> i \<ge> j \<or> P i \<and> all_interval_nat P (Suc i) j"
6243 have *: "\<And>n. P i \<Longrightarrow> \<forall>n\<in>{Suc i..<j}. P n \<Longrightarrow> i \<le> n \<Longrightarrow> n < j \<Longrightarrow> P n"
6246 assume "P i" "\<forall>n\<in>{Suc i..<j}. P n" "i \<le> n" "n < j"
6247 then show "P n" by (cases "n = i") simp_all
6249 show ?thesis by (auto simp add: all_interval_nat_def intro: *)
6252 lemma list_all_iff_all_interval_nat [code_unfold]:
6253 "list_all P [i..<j] \<longleftrightarrow> all_interval_nat P i j"
6254 by (simp add: list_all_iff all_interval_nat_def)
6256 lemma list_ex_iff_not_all_inverval_nat [code_unfold]:
6257 "list_ex P [i..<j] \<longleftrightarrow> \<not> (all_interval_nat (Not \<circ> P) i j)"
6258 by (simp add: list_ex_iff all_interval_nat_def)
6260 definition all_interval_int :: "(int \<Rightarrow> bool) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool" where
6261 "all_interval_int P i j \<longleftrightarrow> (\<forall>k \<in> {i..j}. P k)"
6264 "all_interval_int P i j \<longleftrightarrow> i > j \<or> P i \<and> all_interval_int P (i + 1) j"
6266 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"
6269 assume "P i" "\<forall>k\<in>{i+1..j}. P k" "i \<le> k" "k \<le> j"
6270 then show "P k" by (cases "k = i") simp_all
6272 show ?thesis by (auto simp add: all_interval_int_def intro: *)
6275 lemma list_all_iff_all_interval_int [code_unfold]:
6276 "list_all P [i..j] \<longleftrightarrow> all_interval_int P i j"
6277 by (simp add: list_all_iff all_interval_int_def)
6279 lemma list_ex_iff_not_all_inverval_int [code_unfold]:
6280 "list_ex P [i..j] \<longleftrightarrow> \<not> (all_interval_int (Not \<circ> P) i j)"
6281 by (simp add: list_ex_iff all_interval_int_def)
6283 text {* optimized code (tail-recursive) for @{term length} *}
6285 definition gen_length :: "nat \<Rightarrow> 'a list \<Rightarrow> nat"
6286 where "gen_length n xs = n + length xs"
6288 lemma gen_length_code [code]:
6289 "gen_length n [] = n"
6290 "gen_length n (x # xs) = gen_length (Suc n) xs"
6291 by(simp_all add: gen_length_def)
6293 declare list.size(3-4)[code del]
6295 lemma length_code [code]: "length = gen_length 0"
6296 by(simp add: gen_length_def fun_eq_iff)
6298 hide_const (open) member null maps map_filter all_interval_nat all_interval_int gen_length
6301 subsubsection {* Pretty lists *}
6304 (* Code generation for list literals. *)
6306 signature LIST_CODE =
6308 val implode_list: string -> string -> Code_Thingol.iterm -> Code_Thingol.iterm list option
6309 val default_list: int * string
6310 -> (Code_Printer.fixity -> Code_Thingol.iterm -> Pretty.T)
6311 -> Code_Printer.fixity -> Code_Thingol.iterm -> Code_Thingol.iterm -> Pretty.T
6312 val add_literal_list: string -> theory -> theory
6315 structure List_Code : LIST_CODE =
6318 open Basic_Code_Thingol;
6320 fun implode_list nil' cons' t =
6322 fun dest_cons (IConst { name = c, ... } `$ t1 `$ t2) =
6326 | dest_cons _ = NONE;
6327 val (ts, t') = Code_Thingol.unfoldr dest_cons t;
6329 of IConst { name = c, ... } => if c = nil' then SOME ts else NONE
6333 fun default_list (target_fxy, target_cons) pr fxy t1 t2 =
6334 Code_Printer.brackify_infix (target_fxy, Code_Printer.R) fxy (
6335 pr (Code_Printer.INFX (target_fxy, Code_Printer.X)) t1,
6336 Code_Printer.str target_cons,
6337 pr (Code_Printer.INFX (target_fxy, Code_Printer.R)) t2
6340 fun add_literal_list target =
6342 fun pretty literals [nil', cons'] pr thm vars fxy [(t1, _), (t2, _)] =
6343 case Option.map (cons t1) (implode_list nil' cons' t2)
6345 Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts)
6347 default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
6349 Code_Target.set_printings (Code_Symbol.Constant (@{const_name Cons},
6350 [(target, SOME (Code_Printer.complex_const_syntax (2, ([@{const_name Nil}, @{const_name Cons}], pretty))))]))
<