1(*  Title:      HOL/List.thy
2    Author:     Tobias Nipkow
3*)
4
5section \<open>The datatype of finite lists\<close>
6
7theory List
8imports Sledgehammer Code_Numeral Lifting_Set
9begin
10
11datatype (set: 'a) list =
12    Nil  ("[]")
13  | Cons (hd: 'a) (tl: "'a list")  (infixr "#" 65)
14for
15  map: map
16  rel: list_all2
17  pred: list_all
18where
19  "tl [] = []"
20
21datatype_compat list
22
23lemma [case_names Nil Cons, cases type: list]:
24  \<comment> \<open>for backward compatibility -- names of variables differ\<close>
25  "(y = [] \<Longrightarrow> P) \<Longrightarrow> (\<And>a list. y = a # list \<Longrightarrow> P) \<Longrightarrow> P"
26by (rule list.exhaust)
27
28lemma [case_names Nil Cons, induct type: list]:
29  \<comment> \<open>for backward compatibility -- names of variables differ\<close>
30  "P [] \<Longrightarrow> (\<And>a list. P list \<Longrightarrow> P (a # list)) \<Longrightarrow> P list"
31by (rule list.induct)
32
33text \<open>Compatibility:\<close>
34
35setup \<open>Sign.mandatory_path "list"\<close>
36
37lemmas inducts = list.induct
38lemmas recs = list.rec
39lemmas cases = list.case
40
41setup \<open>Sign.parent_path\<close>
42
43lemmas set_simps = list.set (* legacy *)
44
45syntax
46  \<comment> \<open>list Enumeration\<close>
47  "_list" :: "args => 'a list"    ("[(_)]")
48
49translations
50  "[x, xs]" == "x#[xs]"
51  "[x]" == "x#[]"
52
53
54subsection \<open>Basic list processing functions\<close>
55
56primrec (nonexhaustive) last :: "'a list \<Rightarrow> 'a" where
57"last (x # xs) = (if xs = [] then x else last xs)"
58
59primrec butlast :: "'a list \<Rightarrow> 'a list" where
60"butlast [] = []" |
61"butlast (x # xs) = (if xs = [] then [] else x # butlast xs)"
62
63lemma set_rec: "set xs = rec_list {} (\<lambda>x _. insert x) xs"
64  by (induct xs) auto
65
66definition coset :: "'a list \<Rightarrow> 'a set" where
67[simp]: "coset xs = - set xs"
68
69primrec append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "@" 65) where
70append_Nil: "[] @ ys = ys" |
71append_Cons: "(x#xs) @ ys = x # xs @ ys"
72
73primrec rev :: "'a list \<Rightarrow> 'a list" where
74"rev [] = []" |
75"rev (x # xs) = rev xs @ [x]"
76
77primrec filter:: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
78"filter P [] = []" |
79"filter P (x # xs) = (if P x then x # filter P xs else filter P xs)"
80
81text \<open>Special input syntax for filter:\<close>
82syntax (ASCII)
83  "_filter" :: "[pttrn, 'a list, bool] => 'a list"  ("(1[_<-_./ _])")
84syntax
85  "_filter" :: "[pttrn, 'a list, bool] => 'a list"  ("(1[_\<leftarrow>_ ./ _])")
86translations
87  "[x<-xs . P]" \<rightharpoonup> "CONST filter (\<lambda>x. P) xs"
88
89primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
90fold_Nil:  "fold f [] = id" |
91fold_Cons: "fold f (x # xs) = fold f xs \<circ> f x"
92
93primrec foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
94foldr_Nil:  "foldr f [] = id" |
95foldr_Cons: "foldr f (x # xs) = f x \<circ> foldr f xs"
96
97primrec foldl :: "('b \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b" where
98foldl_Nil:  "foldl f a [] = a" |
99foldl_Cons: "foldl f a (x # xs) = foldl f (f a x) xs"
100
101primrec concat:: "'a list list \<Rightarrow> 'a list" where
102"concat [] = []" |
103"concat (x # xs) = x @ concat xs"
104
105primrec drop:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
106drop_Nil: "drop n [] = []" |
107drop_Cons: "drop n (x # xs) = (case n of 0 \<Rightarrow> x # xs | Suc m \<Rightarrow> drop m xs)"
108  \<comment> \<open>Warning: simpset does not contain this definition, but separate
109       theorems for \<open>n = 0\<close> and \<open>n = Suc k\<close>\<close>
110
111primrec take:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
112take_Nil:"take n [] = []" |
113take_Cons: "take n (x # xs) = (case n of 0 \<Rightarrow> [] | Suc m \<Rightarrow> x # take m xs)"
114  \<comment> \<open>Warning: simpset does not contain this definition, but separate
115       theorems for \<open>n = 0\<close> and \<open>n = Suc k\<close>\<close>
116
117primrec (nonexhaustive) nth :: "'a list => nat => 'a" (infixl "!" 100) where
118nth_Cons: "(x # xs) ! n = (case n of 0 \<Rightarrow> x | Suc k \<Rightarrow> xs ! k)"
119  \<comment> \<open>Warning: simpset does not contain this definition, but separate
120       theorems for \<open>n = 0\<close> and \<open>n = Suc k\<close>\<close>
121
122primrec list_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
123"list_update [] i v = []" |
124"list_update (x # xs) i v =
125  (case i of 0 \<Rightarrow> v # xs | Suc j \<Rightarrow> x # list_update xs j v)"
126
127nonterminal lupdbinds and lupdbind
128
129syntax
130  "_lupdbind":: "['a, 'a] => lupdbind"    ("(2_ :=/ _)")
131  "" :: "lupdbind => lupdbinds"    ("_")
132  "_lupdbinds" :: "[lupdbind, lupdbinds] => lupdbinds"    ("_,/ _")
133  "_LUpdate" :: "['a, lupdbinds] => 'a"    ("_/[(_)]" [900,0] 900)
134
135translations
136  "_LUpdate xs (_lupdbinds b bs)" == "_LUpdate (_LUpdate xs b) bs"
137  "xs[i:=x]" == "CONST list_update xs i x"
138
139primrec takeWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
140"takeWhile P [] = []" |
141"takeWhile P (x # xs) = (if P x then x # takeWhile P xs else [])"
142
143primrec dropWhile :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
144"dropWhile P [] = []" |
145"dropWhile P (x # xs) = (if P x then dropWhile P xs else x # xs)"
146
147primrec zip :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
148"zip xs [] = []" |
149zip_Cons: "zip xs (y # ys) =
150  (case xs of [] \<Rightarrow> [] | z # zs \<Rightarrow> (z, y) # zip zs ys)"
151  \<comment> \<open>Warning: simpset does not contain this definition, but separate
152       theorems for \<open>xs = []\<close> and \<open>xs = z # zs\<close>\<close>
153
154abbreviation map2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list" where
155"map2 f xs ys \<equiv> map (\<lambda>(x,y). f x y) (zip xs ys)"
156
157primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
158"product [] _ = []" |
159"product (x#xs) ys = map (Pair x) ys @ product xs ys"
160
161hide_const (open) product
162
163primrec product_lists :: "'a list list \<Rightarrow> 'a list list" where
164"product_lists [] = [[]]" |
165"product_lists (xs # xss) = concat (map (\<lambda>x. map (Cons x) (product_lists xss)) xs)"
166
167primrec upt :: "nat \<Rightarrow> nat \<Rightarrow> nat list" ("(1[_..</_'])") where
168upt_0: "[i..<0] = []" |
169upt_Suc: "[i..<(Suc j)] = (if i <= j then [i..<j] @ [j] else [])"
170
171definition insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
172"insert x xs = (if x \<in> set xs then xs else x # xs)"
173
174definition union :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
175"union = fold insert"
176
177hide_const (open) insert union
178hide_fact (open) insert_def union_def
179
180primrec find :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a option" where
181"find _ [] = None" |
182"find P (x#xs) = (if P x then Some x else find P xs)"
183
184text \<open>In the context of multisets, \<open>count_list\<close> is equivalent to
185  @{term "count \<circ> mset"} and it it advisable to use the latter.\<close>
186primrec count_list :: "'a list \<Rightarrow> 'a \<Rightarrow> nat" where
187"count_list [] y = 0" |
188"count_list (x#xs) y = (if x=y then count_list xs y + 1 else count_list xs y)"
189
190definition
191   "extract" :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> ('a list * 'a * 'a list) option"
192where "extract P xs =
193  (case dropWhile (Not \<circ> P) xs of
194     [] \<Rightarrow> None |
195     y#ys \<Rightarrow> Some(takeWhile (Not \<circ> P) xs, y, ys))"
196
197hide_const (open) "extract"
198
199primrec those :: "'a option list \<Rightarrow> 'a list option"
200where
201"those [] = Some []" |
202"those (x # xs) = (case x of
203  None \<Rightarrow> None
204| Some y \<Rightarrow> map_option (Cons y) (those xs))"
205
206primrec remove1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
207"remove1 x [] = []" |
208"remove1 x (y # xs) = (if x = y then xs else y # remove1 x xs)"
209
210primrec removeAll :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
211"removeAll x [] = []" |
212"removeAll x (y # xs) = (if x = y then removeAll x xs else y # removeAll x xs)"
213
214primrec distinct :: "'a list \<Rightarrow> bool" where
215"distinct [] \<longleftrightarrow> True" |
216"distinct (x # xs) \<longleftrightarrow> x \<notin> set xs \<and> distinct xs"
217
218primrec remdups :: "'a list \<Rightarrow> 'a list" where
219"remdups [] = []" |
220"remdups (x # xs) = (if x \<in> set xs then remdups xs else x # remdups xs)"
221
222fun remdups_adj :: "'a list \<Rightarrow> 'a list" where
223"remdups_adj [] = []" |
224"remdups_adj [x] = [x]" |
225"remdups_adj (x # y # xs) = (if x = y then remdups_adj (x # xs) else x # remdups_adj (y # xs))"
226
227primrec replicate :: "nat \<Rightarrow> 'a \<Rightarrow> 'a list" where
228replicate_0: "replicate 0 x = []" |
229replicate_Suc: "replicate (Suc n) x = x # replicate n x"
230
231text \<open>
232  Function \<open>size\<close> is overloaded for all datatypes. Users may
233  refer to the list version as \<open>length\<close>.\<close>
234
235abbreviation length :: "'a list \<Rightarrow> nat" where
236"length \<equiv> size"
237
238definition enumerate :: "nat \<Rightarrow> 'a list \<Rightarrow> (nat \<times> 'a) list" where
239enumerate_eq_zip: "enumerate n xs = zip [n..<n + length xs] xs"
240
241primrec rotate1 :: "'a list \<Rightarrow> 'a list" where
242"rotate1 [] = []" |
243"rotate1 (x # xs) = xs @ [x]"
244
245definition rotate :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
246"rotate n = rotate1 ^^ n"
247
248definition nths :: "'a list => nat set => 'a list" where
249"nths xs A = map fst (filter (\<lambda>p. snd p \<in> A) (zip xs [0..<size xs]))"
250
251primrec subseqs :: "'a list \<Rightarrow> 'a list list" where
252"subseqs [] = [[]]" |
253"subseqs (x#xs) = (let xss = subseqs xs in map (Cons x) xss @ xss)"
254
255primrec n_lists :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list list" where
256"n_lists 0 xs = [[]]" |
257"n_lists (Suc n) xs = concat (map (\<lambda>ys. map (\<lambda>y. y # ys) xs) (n_lists n xs))"
258
259hide_const (open) n_lists
260
261fun splice :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
262"splice [] ys = ys" |
263"splice xs [] = xs" |
264"splice (x#xs) (y#ys) = x # y # splice xs ys"
265
266function shuffle where
267  "shuffle [] ys = {ys}"
268| "shuffle xs [] = {xs}"
269| "shuffle (x # xs) (y # ys) = (#) x ` shuffle xs (y # ys) \<union> (#) y ` shuffle (x # xs) ys"
270  by pat_completeness simp_all
271termination by lexicographic_order
272
273text\<open>Use only if you cannot use @{const Min} instead:\<close>
274fun min_list :: "'a::ord list \<Rightarrow> 'a" where
275"min_list (x # xs) = (case xs of [] \<Rightarrow> x | _ \<Rightarrow> min x (min_list xs))"
276
277text\<open>Returns first minimum:\<close>
278fun arg_min_list :: "('a \<Rightarrow> ('b::linorder)) \<Rightarrow> 'a list \<Rightarrow> 'a" where
279"arg_min_list f [x] = x" |
280"arg_min_list f (x#y#zs) = (let m = arg_min_list f (y#zs) in if f x \<le> f m then x else m)"
281
282text\<open>
283\begin{figure}[htbp]
284\fbox{
285\begin{tabular}{l}
286@{lemma "[a,b]@[c,d] = [a,b,c,d]" by simp}\\
287@{lemma "length [a,b,c] = 3" by simp}\\
288@{lemma "set [a,b,c] = {a,b,c}" by simp}\\
289@{lemma "map f [a,b,c] = [f a, f b, f c]" by simp}\\
290@{lemma "rev [a,b,c] = [c,b,a]" by simp}\\
291@{lemma "hd [a,b,c,d] = a" by simp}\\
292@{lemma "tl [a,b,c,d] = [b,c,d]" by simp}\\
293@{lemma "last [a,b,c,d] = d" by simp}\\
294@{lemma "butlast [a,b,c,d] = [a,b,c]" by simp}\\
295@{lemma[source] "filter (\<lambda>n::nat. n<2) [0,2,1] = [0,1]" by simp}\\
296@{lemma "concat [[a,b],[c,d,e],[],[f]] = [a,b,c,d,e,f]" by simp}\\
297@{lemma "fold f [a,b,c] x = f c (f b (f a x))" by simp}\\
298@{lemma "foldr f [a,b,c] x = f a (f b (f c x))" by simp}\\
299@{lemma "foldl f x [a,b,c] = f (f (f x a) b) c" by simp}\\
300@{lemma "zip [a,b,c] [x,y,z] = [(a,x),(b,y),(c,z)]" by simp}\\
301@{lemma "zip [a,b] [x,y,z] = [(a,x),(b,y)]" by simp}\\
302@{lemma "enumerate 3 [a,b,c] = [(3,a),(4,b),(5,c)]" by normalization}\\
303@{lemma "List.product [a,b] [c,d] = [(a, c), (a, d), (b, c), (b, d)]" by simp}\\
304@{lemma "product_lists [[a,b], [c], [d,e]] = [[a,c,d], [a,c,e], [b,c,d], [b,c,e]]" by simp}\\
305@{lemma "splice [a,b,c] [x,y,z] = [a,x,b,y,c,z]" by simp}\\
306@{lemma "splice [a,b,c,d] [x,y] = [a,x,b,y,c,d]" by simp}\\
307@{lemma "shuffle [a,b] [c,d] =  {[a,b,c,d],[a,c,b,d],[a,c,d,b],[c,a,b,d],[c,a,d,b],[c,d,a,b]}"
308    by (simp add: insert_commute)}\\
309@{lemma "take 2 [a,b,c,d] = [a,b]" by simp}\\
310@{lemma "take 6 [a,b,c,d] = [a,b,c,d]" by simp}\\
311@{lemma "drop 2 [a,b,c,d] = [c,d]" by simp}\\
312@{lemma "drop 6 [a,b,c,d] = []" by simp}\\
313@{lemma "takeWhile (%n::nat. n<3) [1,2,3,0] = [1,2]" by simp}\\
314@{lemma "dropWhile (%n::nat. n<3) [1,2,3,0] = [3,0]" by simp}\\
315@{lemma "distinct [2,0,1::nat]" by simp}\\
316@{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\
317@{lemma "remdups_adj [2,2,3,1,1::nat,2,1] = [2,3,1,2,1]" by simp}\\
318@{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\
319@{lemma "List.insert 3 [0::nat,1,2] = [3,0,1,2]" by (simp add: List.insert_def)}\\
320@{lemma "List.union [2,3,4] [0::int,1,2] = [4,3,0,1,2]" by (simp add: List.insert_def List.union_def)}\\
321@{lemma "List.find (%i::int. i>0) [0,0] = None" by simp}\\
322@{lemma "List.find (%i::int. i>0) [0,1,0,2] = Some 1" by simp}\\
323@{lemma "count_list [0,1,0,2::int] 0 = 2" by (simp)}\\
324@{lemma "List.extract (%i::int. i>0) [0,0] = None" by(simp add: extract_def)}\\
325@{lemma "List.extract (%i::int. i>0) [0,1,0,2] = Some([0], 1, [0,2])" by(simp add: extract_def)}\\
326@{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\
327@{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\
328@{lemma "nth [a,b,c,d] 2 = c" by simp}\\
329@{lemma "[a,b,c,d][2 := x] = [a,b,x,d]" by simp}\\
330@{lemma "nths [a,b,c,d,e] {0,2,3} = [a,c,d]" by (simp add:nths_def)}\\
331@{lemma "subseqs [a,b] = [[a, b], [a], [b], []]" by simp}\\
332@{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)}\\
333@{lemma "rotate1 [a,b,c,d] = [b,c,d,a]" by simp}\\
334@{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate_def eval_nat_numeral)}\\
335@{lemma "replicate 4 a = [a,a,a,a]" by (simp add:eval_nat_numeral)}\\
336@{lemma "[2..<5] = [2,3,4]" by (simp add:eval_nat_numeral)}\\
337@{lemma "min_list [3,1,-2::int] = -2" by (simp)}\\
338@{lemma "arg_min_list (\<lambda>i. i*i) [3,-1,1,-2::int] = -1" by (simp)}
339\end{tabular}}
340\caption{Characteristic examples}
341\label{fig:Characteristic}
342\end{figure}
343Figure~\ref{fig:Characteristic} shows characteristic examples
344that should give an intuitive understanding of the above functions.
345\<close>
346
347text\<open>The following simple sort functions are intended for proofs,
348not for efficient implementations.\<close>
349
350text \<open>A sorted predicate w.r.t. a relation:\<close>
351
352fun sorted_wrt :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
353"sorted_wrt P [] = True" |
354"sorted_wrt P (x # ys) = ((\<forall>y \<in> set ys. P x y) \<and> sorted_wrt P ys)"
355
356(* FIXME: define sorted in terms of sorted_wrt *)
357
358text \<open>A class-based sorted predicate:\<close>
359
360context linorder
361begin
362
363fun sorted :: "'a list \<Rightarrow> bool" where
364"sorted [] = True" |
365"sorted (x # ys) = ((\<forall>y \<in> set ys. x \<le> y) \<and> sorted ys)"
366
367lemma sorted_sorted_wrt: "sorted = sorted_wrt (\<le>)"
368proof (rule ext)
369  fix xs show "sorted xs = sorted_wrt (\<le>) xs"
370    by(induction xs rule: sorted.induct) auto
371qed
372
373primrec insort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where
374"insort_key f x [] = [x]" |
375"insort_key f x (y#ys) =
376  (if f x \<le> f y then (x#y#ys) else y#(insort_key f x ys))"
377
378definition sort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b list \<Rightarrow> 'b list" where
379"sort_key f xs = foldr (insort_key f) xs []"
380
381definition insort_insert_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where
382"insort_insert_key f x xs =
383  (if f x \<in> f ` set xs then xs else insort_key f x xs)"
384
385abbreviation "sort \<equiv> sort_key (\<lambda>x. x)"
386abbreviation "insort \<equiv> insort_key (\<lambda>x. x)"
387abbreviation "insort_insert \<equiv> insort_insert_key (\<lambda>x. x)"
388
389definition stable_sort_key :: "(('b \<Rightarrow> 'a) \<Rightarrow> 'b list \<Rightarrow> 'b list) \<Rightarrow> bool" where
390"stable_sort_key sk =
391   (\<forall>f xs k. filter (\<lambda>y. f y = k) (sk f xs) = filter (\<lambda>y. f y = k) xs)"
392
393end
394
395
396subsubsection \<open>List comprehension\<close>
397
398text\<open>Input syntax for Haskell-like list comprehension notation.
399Typical example: \<open>[(x,y). x \<leftarrow> xs, y \<leftarrow> ys, x \<noteq> y]\<close>,
400the list of all pairs of distinct elements from \<open>xs\<close> and \<open>ys\<close>.
401The syntax is as in Haskell, except that \<open>|\<close> becomes a dot
402(like in Isabelle's set comprehension): \<open>[e. x \<leftarrow> xs, \<dots>]\<close> rather than
403\verb![e| x <- xs, ...]!.
404
405The qualifiers after the dot are
406\begin{description}
407\item[generators] \<open>p \<leftarrow> xs\<close>,
408 where \<open>p\<close> is a pattern and \<open>xs\<close> an expression of list type, or
409\item[guards] \<open>b\<close>, where \<open>b\<close> is a boolean expression.
410%\item[local bindings] @ {text"let x = e"}.
411\end{description}
412
413Just like in Haskell, list comprehension is just a shorthand. To avoid
414misunderstandings, the translation into desugared form is not reversed
415upon output. Note that the translation of \<open>[e. x \<leftarrow> xs]\<close> is
416optmized to @{term"map (%x. e) xs"}.
417
418It is easy to write short list comprehensions which stand for complex
419expressions. During proofs, they may become unreadable (and
420mangled). In such cases it can be advisable to introduce separate
421definitions for the list comprehensions in question.\<close>
422
423nonterminal lc_qual and lc_quals
424
425syntax
426  "_listcompr" :: "'a \<Rightarrow> lc_qual \<Rightarrow> lc_quals \<Rightarrow> 'a list"  ("[_ . __")
427  "_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual"  ("_ \<leftarrow> _")
428  "_lc_test" :: "bool \<Rightarrow> lc_qual" ("_")
429  (*"_lc_let" :: "letbinds => lc_qual"  ("let _")*)
430  "_lc_end" :: "lc_quals" ("]")
431  "_lc_quals" :: "lc_qual \<Rightarrow> lc_quals \<Rightarrow> lc_quals"  (", __")
432
433syntax (ASCII)
434  "_lc_gen" :: "'a \<Rightarrow> 'a list \<Rightarrow> lc_qual"  ("_ <- _")
435
436parse_translation \<open>
437let
438  val NilC = Syntax.const @{const_syntax Nil};
439  val ConsC = Syntax.const @{const_syntax Cons};
440  val mapC = Syntax.const @{const_syntax map};
441  val concatC = Syntax.const @{const_syntax concat};
442  val IfC = Syntax.const @{const_syntax If};
443  val dummyC = Syntax.const @{const_syntax Pure.dummy_pattern}
444
445  fun single x = ConsC $ x $ NilC;
446
447  fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)
448    let
449      (* FIXME proper name context!? *)
450      val x =
451        Free (singleton (Name.variant_list (fold Term.add_free_names [p, e] [])) "x", dummyT);
452      val e = if opti then single e else e;
453      val case1 = Syntax.const @{syntax_const "_case1"} $ p $ e;
454      val case2 =
455        Syntax.const @{syntax_const "_case1"} $ dummyC $ NilC;
456      val cs = Syntax.const @{syntax_const "_case2"} $ case1 $ case2;
457    in Syntax_Trans.abs_tr [x, Case_Translation.case_tr false ctxt [x, cs]] end;
458
459  fun pair_pat_tr (x as Free _) e = Syntax_Trans.abs_tr [x, e]
460    | pair_pat_tr (_ $ p1 $ p2) e =
461        Syntax.const @{const_syntax case_prod} $ pair_pat_tr p1 (pair_pat_tr p2 e)
462    | pair_pat_tr dummy e = Syntax_Trans.abs_tr [Syntax.const "_idtdummy", e]
463
464  fun pair_pat ctxt (Const (@{const_syntax "Pair"},_) $ s $ t) =
465        pair_pat ctxt s andalso pair_pat ctxt t
466    | pair_pat ctxt (Free (s,_)) =
467        let
468          val thy = Proof_Context.theory_of ctxt;
469          val s' = Proof_Context.intern_const ctxt s;
470        in not (Sign.declared_const thy s') end
471    | pair_pat _ t = (t = dummyC);
472
473  fun abs_tr ctxt p e opti =
474    let val p = Term_Position.strip_positions p
475    in if pair_pat ctxt p
476       then (pair_pat_tr p e, true)
477       else (pat_tr ctxt p e opti, false)
478    end
479
480  fun lc_tr ctxt [e, Const (@{syntax_const "_lc_test"}, _) $ b, qs] =
481    let
482      val res =
483        (case qs of
484           Const (@{syntax_const "_lc_end"}, _) => single e
485         | Const (@{syntax_const "_lc_quals"}, _) $ q $ qs => lc_tr ctxt [e, q, qs]);
486    in IfC $ b $ res $ NilC end
487  | lc_tr ctxt
488      [e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es,
489          Const(@{syntax_const "_lc_end"}, _)] =
490      (case abs_tr ctxt p e true of
491         (f, true) => mapC $ f $ es
492       | (f, false) => concatC $ (mapC $ f $ es))
493  | lc_tr ctxt
494      [e, Const (@{syntax_const "_lc_gen"}, _) $ p $ es,
495          Const (@{syntax_const "_lc_quals"}, _) $ q $ qs] =
496      let val e' = lc_tr ctxt [e, q, qs];
497      in concatC $ (mapC $ (fst (abs_tr ctxt p e' false)) $ es) end;
498
499in [(@{syntax_const "_listcompr"}, lc_tr)] end
500\<close>
501
502ML_val \<open>
503  let
504    val read = Syntax.read_term @{context} o Syntax.implode_input;
505    fun check s1 s2 =
506      read s1 aconv read s2 orelse
507        error ("Check failed: " ^
508          quote (Input.source_content s1) ^ Position.here_list [Input.pos_of s1, Input.pos_of s2]);
509  in
510    check \<open>[(x,y,z). b]\<close> \<open>if b then [(x, y, z)] else []\<close>;
511    check \<open>[(x,y,z). (x,_,y)\<leftarrow>xs]\<close> \<open>map (\<lambda>(x,_,y). (x, y, z)) xs\<close>;
512    check \<open>[e x y. (x,_)\<leftarrow>xs, y\<leftarrow>ys]\<close> \<open>concat (map (\<lambda>(x,_). map (\<lambda>y. e x y) ys) xs)\<close>;
513    check \<open>[(x,y,z). x<a, x>b]\<close> \<open>if x < a then if b < x then [(x, y, z)] else [] else []\<close>;
514    check \<open>[(x,y,z). x\<leftarrow>xs, x>b]\<close> \<open>concat (map (\<lambda>x. if b < x then [(x, y, z)] else []) xs)\<close>;
515    check \<open>[(x,y,z). x<a, x\<leftarrow>xs]\<close> \<open>if x < a then map (\<lambda>x. (x, y, z)) xs else []\<close>;
516    check \<open>[(x,y). Cons True x \<leftarrow> xs]\<close>
517      \<open>concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | True # x \<Rightarrow> [(x, y)] | False # x \<Rightarrow> []) xs)\<close>;
518    check \<open>[(x,y,z). Cons x [] \<leftarrow> xs]\<close>
519      \<open>concat (map (\<lambda>xa. case xa of [] \<Rightarrow> [] | [x] \<Rightarrow> [(x, y, z)] | x # aa # lista \<Rightarrow> []) xs)\<close>;
520    check \<open>[(x,y,z). x<a, x>b, x=d]\<close>
521      \<open>if x < a then if b < x then if x = d then [(x, y, z)] else [] else [] else []\<close>;
522    check \<open>[(x,y,z). x<a, x>b, y\<leftarrow>ys]\<close>
523      \<open>if x < a then if b < x then map (\<lambda>y. (x, y, z)) ys else [] else []\<close>;
524    check \<open>[(x,y,z). x<a, (_,x)\<leftarrow>xs,y>b]\<close>
525      \<open>if x < a then concat (map (\<lambda>(_,x). if b < y then [(x, y, z)] else []) xs) else []\<close>;
526    check \<open>[(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys]\<close>
527      \<open>if x < a then concat (map (\<lambda>x. map (\<lambda>y. (x, y, z)) ys) xs) else []\<close>;
528    check \<open>[(x,y,z). x\<leftarrow>xs, x>b, y<a]\<close>
529      \<open>concat (map (\<lambda>x. if b < x then if y < a then [(x, y, z)] else [] else []) xs)\<close>;
530    check \<open>[(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys]\<close>
531      \<open>concat (map (\<lambda>x. if b < x then map (\<lambda>y. (x, y, z)) ys else []) xs)\<close>;
532    check \<open>[(x,y,z). x\<leftarrow>xs, (y,_)\<leftarrow>ys,y>x]\<close>
533      \<open>concat (map (\<lambda>x. concat (map (\<lambda>(y,_). if x < y then [(x, y, z)] else []) ys)) xs)\<close>;
534    check \<open>[(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs]\<close>
535      \<open>concat (map (\<lambda>x. concat (map (\<lambda>y. map (\<lambda>z. (x, y, z)) zs) ys)) xs)\<close>
536  end;
537\<close>
538
539ML \<open>
540(* Simproc for rewriting list comprehensions applied to List.set to set
541   comprehension. *)
542
543signature LIST_TO_SET_COMPREHENSION =
544sig
545  val simproc : Proof.context -> cterm -> thm option
546end
547
548structure List_to_Set_Comprehension : LIST_TO_SET_COMPREHENSION =
549struct
550
551(* conversion *)
552
553fun all_exists_conv cv ctxt ct =
554  (case Thm.term_of ct of
555    Const (@{const_name Ex}, _) $ Abs _ =>
556      Conv.arg_conv (Conv.abs_conv (all_exists_conv cv o #2) ctxt) ct
557  | _ => cv ctxt ct)
558
559fun all_but_last_exists_conv cv ctxt ct =
560  (case Thm.term_of ct of
561    Const (@{const_name Ex}, _) $ Abs (_, _, Const (@{const_name Ex}, _) $ _) =>
562      Conv.arg_conv (Conv.abs_conv (all_but_last_exists_conv cv o #2) ctxt) ct
563  | _ => cv ctxt ct)
564
565fun Collect_conv cv ctxt ct =
566  (case Thm.term_of ct of
567    Const (@{const_name Collect}, _) $ Abs _ => Conv.arg_conv (Conv.abs_conv cv ctxt) ct
568  | _ => raise CTERM ("Collect_conv", [ct]))
569
570fun rewr_conv' th = Conv.rewr_conv (mk_meta_eq th)
571
572fun conjunct_assoc_conv ct =
573  Conv.try_conv
574    (rewr_conv' @{thm conj_assoc} then_conv HOLogic.conj_conv Conv.all_conv conjunct_assoc_conv) ct
575
576fun right_hand_set_comprehension_conv conv ctxt =
577  HOLogic.Trueprop_conv (HOLogic.eq_conv Conv.all_conv
578    (Collect_conv (all_exists_conv conv o #2) ctxt))
579
580
581(* term abstraction of list comprehension patterns *)
582
583datatype termlets = If | Case of typ * int
584
585local
586
587val set_Nil_I = @{lemma "set [] = {x. False}" by (simp add: empty_def [symmetric])}
588val set_singleton = @{lemma "set [a] = {x. x = a}" by simp}
589val inst_Collect_mem_eq = @{lemma "set A = {x. x \<in> set A}" by simp}
590val del_refl_eq = @{lemma "(t = t \<and> P) \<equiv> P" by simp}
591
592fun mk_set T = Const (@{const_name set}, HOLogic.listT T --> HOLogic.mk_setT T)
593fun dest_set (Const (@{const_name set}, _) $ xs) = xs
594
595fun dest_singleton_list (Const (@{const_name Cons}, _) $ t $ (Const (@{const_name Nil}, _))) = t
596  | dest_singleton_list t = raise TERM ("dest_singleton_list", [t])
597
598(*We check that one case returns a singleton list and all other cases
599  return [], and return the index of the one singleton list case.*)
600fun possible_index_of_singleton_case cases =
601  let
602    fun check (i, case_t) s =
603      (case strip_abs_body case_t of
604        (Const (@{const_name Nil}, _)) => s
605      | _ => (case s of SOME NONE => SOME (SOME i) | _ => NONE))
606  in
607    fold_index check cases (SOME NONE) |> the_default NONE
608  end
609
610(*returns condition continuing term option*)
611fun dest_if (Const (@{const_name If}, _) $ cond $ then_t $ Const (@{const_name Nil}, _)) =
612      SOME (cond, then_t)
613  | dest_if _ = NONE
614
615(*returns (case_expr type index chosen_case constr_name) option*)
616fun dest_case ctxt case_term =
617  let
618    val (case_const, args) = strip_comb case_term
619  in
620    (case try dest_Const case_const of
621      SOME (c, T) =>
622        (case Ctr_Sugar.ctr_sugar_of_case ctxt c of
623          SOME {ctrs, ...} =>
624            (case possible_index_of_singleton_case (fst (split_last args)) of
625              SOME i =>
626                let
627                  val constr_names = map (fst o dest_Const) ctrs
628                  val (Ts, _) = strip_type T
629                  val T' = List.last Ts
630                in SOME (List.last args, T', i, nth args i, nth constr_names i) end
631            | NONE => NONE)
632        | NONE => NONE)
633    | NONE => NONE)
634  end
635
636fun tac ctxt [] =
637      resolve_tac ctxt [set_singleton] 1 ORELSE
638      resolve_tac ctxt [inst_Collect_mem_eq] 1
639  | tac ctxt (If :: cont) =
640      Splitter.split_tac ctxt @{thms if_split} 1
641      THEN resolve_tac ctxt @{thms conjI} 1
642      THEN resolve_tac ctxt @{thms impI} 1
643      THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
644        CONVERSION (right_hand_set_comprehension_conv (K
645          (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_TrueI})) Conv.all_conv
646           then_conv
647           rewr_conv' @{lemma "(True \<and> P) = P" by simp})) ctxt') 1) ctxt 1
648      THEN tac ctxt cont
649      THEN resolve_tac ctxt @{thms impI} 1
650      THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
651          CONVERSION (right_hand_set_comprehension_conv (K
652            (HOLogic.conj_conv (Conv.rewr_conv (List.last prems RS @{thm Eq_FalseI})) Conv.all_conv
653             then_conv rewr_conv' @{lemma "(False \<and> P) = False" by simp})) ctxt') 1) ctxt 1
654      THEN resolve_tac ctxt [set_Nil_I] 1
655  | tac ctxt (Case (T, i) :: cont) =
656      let
657        val SOME {injects, distincts, case_thms, split, ...} =
658          Ctr_Sugar.ctr_sugar_of ctxt (fst (dest_Type T))
659      in
660        (* do case distinction *)
661        Splitter.split_tac ctxt [split] 1
662        THEN EVERY (map_index (fn (i', _) =>
663          (if i' < length case_thms - 1 then resolve_tac ctxt @{thms conjI} 1 else all_tac)
664          THEN REPEAT_DETERM (resolve_tac ctxt @{thms allI} 1)
665          THEN resolve_tac ctxt @{thms impI} 1
666          THEN (if i' = i then
667            (* continue recursively *)
668            Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
669              CONVERSION (Thm.eta_conversion then_conv right_hand_set_comprehension_conv (K
670                  ((HOLogic.conj_conv
671                    (HOLogic.eq_conv Conv.all_conv (rewr_conv' (List.last prems)) then_conv
672                      (Conv.try_conv (Conv.rewrs_conv (map mk_meta_eq injects))))
673                    Conv.all_conv)
674                    then_conv (Conv.try_conv (Conv.rewr_conv del_refl_eq))
675                    then_conv conjunct_assoc_conv)) ctxt'
676                then_conv
677                  (HOLogic.Trueprop_conv
678                    (HOLogic.eq_conv Conv.all_conv (Collect_conv (fn (_, ctxt'') =>
679                      Conv.repeat_conv
680                        (all_but_last_exists_conv
681                          (K (rewr_conv'
682                            @{lemma "(\<exists>x. x = t \<and> P x) = P t" by simp})) ctxt'')) ctxt')))) 1) ctxt 1
683            THEN tac ctxt cont
684          else
685            Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
686              CONVERSION
687                (right_hand_set_comprehension_conv (K
688                  (HOLogic.conj_conv
689                    ((HOLogic.eq_conv Conv.all_conv
690                      (rewr_conv' (List.last prems))) then_conv
691                      (Conv.rewrs_conv (map (fn th => th RS @{thm Eq_FalseI}) distincts)))
692                    Conv.all_conv then_conv
693                    (rewr_conv' @{lemma "(False \<and> P) = False" by simp}))) ctxt' then_conv
694                  HOLogic.Trueprop_conv
695                    (HOLogic.eq_conv Conv.all_conv
696                      (Collect_conv (fn (_, ctxt'') =>
697                        Conv.repeat_conv
698                          (Conv.bottom_conv
699                            (K (rewr_conv' @{lemma "(\<exists>x. P) = P" by simp})) ctxt'')) ctxt'))) 1) ctxt 1
700            THEN resolve_tac ctxt [set_Nil_I] 1)) case_thms)
701      end
702
703in
704
705fun simproc ctxt redex =
706  let
707    fun make_inner_eqs bound_vs Tis eqs t =
708      (case dest_case ctxt t of
709        SOME (x, T, i, cont, constr_name) =>
710          let
711            val (vs, body) = strip_abs (Envir.eta_long (map snd bound_vs) cont)
712            val x' = incr_boundvars (length vs) x
713            val eqs' = map (incr_boundvars (length vs)) eqs
714            val constr_t =
715              list_comb
716                (Const (constr_name, map snd vs ---> T), map Bound (((length vs) - 1) downto 0))
717            val constr_eq = Const (@{const_name HOL.eq}, T --> T --> @{typ bool}) $ constr_t $ x'
718          in
719            make_inner_eqs (rev vs @ bound_vs) (Case (T, i) :: Tis) (constr_eq :: eqs') body
720          end
721      | NONE =>
722          (case dest_if t of
723            SOME (condition, cont) => make_inner_eqs bound_vs (If :: Tis) (condition :: eqs) cont
724          | NONE =>
725            if null eqs then NONE (*no rewriting, nothing to be done*)
726            else
727              let
728                val Type (@{type_name list}, [rT]) = fastype_of1 (map snd bound_vs, t)
729                val pat_eq =
730                  (case try dest_singleton_list t of
731                    SOME t' =>
732                      Const (@{const_name HOL.eq}, rT --> rT --> @{typ bool}) $
733                        Bound (length bound_vs) $ t'
734                  | NONE =>
735                      Const (@{const_name Set.member}, rT --> HOLogic.mk_setT rT --> @{typ bool}) $
736                        Bound (length bound_vs) $ (mk_set rT $ t))
737                val reverse_bounds = curry subst_bounds
738                  ((map Bound ((length bound_vs - 1) downto 0)) @ [Bound (length bound_vs)])
739                val eqs' = map reverse_bounds eqs
740                val pat_eq' = reverse_bounds pat_eq
741                val inner_t =
742                  fold (fn (_, T) => fn t => HOLogic.exists_const T $ absdummy T t)
743                    (rev bound_vs) (fold (curry HOLogic.mk_conj) eqs' pat_eq')
744                val lhs = Thm.term_of redex
745                val rhs = HOLogic.mk_Collect ("x", rT, inner_t)
746                val rewrite_rule_t = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
747              in
748                SOME
749                  ((Goal.prove ctxt [] [] rewrite_rule_t
750                    (fn {context = ctxt', ...} => tac ctxt' (rev Tis))) RS @{thm eq_reflection})
751              end))
752  in
753    make_inner_eqs [] [] [] (dest_set (Thm.term_of redex))
754  end
755
756end
757
758end
759\<close>
760
761simproc_setup list_to_set_comprehension ("set xs") =
762  \<open>K List_to_Set_Comprehension.simproc\<close>
763
764code_datatype set coset
765hide_const (open) coset
766
767
768subsubsection \<open>@{const Nil} and @{const Cons}\<close>
769
770lemma not_Cons_self [simp]:
771  "xs \<noteq> x # xs"
772by (induct xs) auto
773
774lemma not_Cons_self2 [simp]: "x # xs \<noteq> xs"
775by (rule not_Cons_self [symmetric])
776
777lemma neq_Nil_conv: "(xs \<noteq> []) = (\<exists>y ys. xs = y # ys)"
778by (induct xs) auto
779
780lemma tl_Nil: "tl xs = [] \<longleftrightarrow> xs = [] \<or> (\<exists>x. xs = [x])"
781by (cases xs) auto
782
783lemma Nil_tl: "[] = tl xs \<longleftrightarrow> xs = [] \<or> (\<exists>x. xs = [x])"
784by (cases xs) auto
785
786lemma length_induct:
787  "(\<And>xs. \<forall>ys. length ys < length xs \<longrightarrow> P ys \<Longrightarrow> P xs) \<Longrightarrow> P xs"
788by (fact measure_induct)
789
790lemma induct_list012:
791  "\<lbrakk>P []; \<And>x. P [x]; \<And>x y zs. P (y # zs) \<Longrightarrow> P (x # y # zs)\<rbrakk> \<Longrightarrow> P xs"
792by induction_schema (pat_completeness, lexicographic_order)
793
794lemma list_nonempty_induct [consumes 1, case_names single cons]:
795  "\<lbrakk> xs \<noteq> []; \<And>x. P [x]; \<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (x # xs)\<rbrakk> \<Longrightarrow> P xs"
796by(induction xs rule: induct_list012) auto
797
798lemma inj_split_Cons: "inj_on (\<lambda>(xs, n). n#xs) X"
799  by (auto intro!: inj_onI)
800
801lemma inj_on_Cons1 [simp]: "inj_on ((#) x) A"
802by(simp add: inj_on_def)
803
804subsubsection \<open>@{const length}\<close>
805
806text \<open>
807  Needs to come before \<open>@\<close> because of theorem \<open>append_eq_append_conv\<close>.
808\<close>
809
810lemma length_append [simp]: "length (xs @ ys) = length xs + length ys"
811by (induct xs) auto
812
813lemma length_map [simp]: "length (map f xs) = length xs"
814by (induct xs) auto
815
816lemma length_rev [simp]: "length (rev xs) = length xs"
817by (induct xs) auto
818
819lemma length_tl [simp]: "length (tl xs) = length xs - 1"
820by (cases xs) auto
821
822lemma length_0_conv [iff]: "(length xs = 0) = (xs = [])"
823by (induct xs) auto
824
825lemma length_greater_0_conv [iff]: "(0 < length xs) = (xs \<noteq> [])"
826by (induct xs) auto
827
828lemma length_pos_if_in_set: "x \<in> set xs \<Longrightarrow> length xs > 0"
829by auto
830
831lemma length_Suc_conv:
832"(length xs = Suc n) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
833by (induct xs) auto
834
835lemma Suc_length_conv:
836  "(Suc n = length xs) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
837apply (induct xs, simp, simp)
838apply blast
839done
840
841lemma impossible_Cons: "length xs <= length ys ==> xs = x # ys = False"
842by (induct xs) auto
843
844lemma list_induct2 [consumes 1, case_names Nil Cons]:
845  "length xs = length ys \<Longrightarrow> P [] [] \<Longrightarrow>
846   (\<And>x xs y ys. length xs = length ys \<Longrightarrow> P xs ys \<Longrightarrow> P (x#xs) (y#ys))
847   \<Longrightarrow> P xs ys"
848proof (induct xs arbitrary: ys)
849  case Nil then show ?case by simp
850next
851  case (Cons x xs ys) then show ?case by (cases ys) simp_all
852qed
853
854lemma list_induct3 [consumes 2, case_names Nil Cons]:
855  "length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> P [] [] [] \<Longrightarrow>
856   (\<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))
857   \<Longrightarrow> P xs ys zs"
858proof (induct xs arbitrary: ys zs)
859  case Nil then show ?case by simp
860next
861  case (Cons x xs ys zs) then show ?case by (cases ys, simp_all)
862    (cases zs, simp_all)
863qed
864
865lemma list_induct4 [consumes 3, case_names Nil Cons]:
866  "length xs = length ys \<Longrightarrow> length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow>
867   P [] [] [] [] \<Longrightarrow> (\<And>x xs y ys z zs w ws. length xs = length ys \<Longrightarrow>
868   length ys = length zs \<Longrightarrow> length zs = length ws \<Longrightarrow> P xs ys zs ws \<Longrightarrow>
869   P (x#xs) (y#ys) (z#zs) (w#ws)) \<Longrightarrow> P xs ys zs ws"
870proof (induct xs arbitrary: ys zs ws)
871  case Nil then show ?case by simp
872next
873  case (Cons x xs ys zs ws) then show ?case by ((cases ys, simp_all), (cases zs,simp_all)) (cases ws, simp_all)
874qed
875
876lemma list_induct2':
877  "\<lbrakk> P [] [];
878  \<And>x xs. P (x#xs) [];
879  \<And>y ys. P [] (y#ys);
880   \<And>x xs y ys. P xs ys  \<Longrightarrow> P (x#xs) (y#ys) \<rbrakk>
881 \<Longrightarrow> P xs ys"
882by (induct xs arbitrary: ys) (case_tac x, auto)+
883
884lemma list_all2_iff:
885  "list_all2 P xs ys \<longleftrightarrow> length xs = length ys \<and> (\<forall>(x, y) \<in> set (zip xs ys). P x y)"
886by (induct xs ys rule: list_induct2') auto
887
888lemma neq_if_length_neq: "length xs \<noteq> length ys \<Longrightarrow> (xs = ys) == False"
889by (rule Eq_FalseI) auto
890
891simproc_setup list_neq ("(xs::'a list) = ys") = \<open>
892(*
893Reduces xs=ys to False if xs and ys cannot be of the same length.
894This is the case if the atomic sublists of one are a submultiset
895of those of the other list and there are fewer Cons's in one than the other.
896*)
897
898let
899
900fun len (Const(@{const_name Nil},_)) acc = acc
901  | len (Const(@{const_name Cons},_) $ _ $ xs) (ts,n) = len xs (ts,n+1)
902  | len (Const(@{const_name append},_) $ xs $ ys) acc = len xs (len ys acc)
903  | len (Const(@{const_name rev},_) $ xs) acc = len xs acc
904  | len (Const(@{const_name map},_) $ _ $ xs) acc = len xs acc
905  | len t (ts,n) = (t::ts,n);
906
907val ss = simpset_of @{context};
908
909fun list_neq ctxt ct =
910  let
911    val (Const(_,eqT) $ lhs $ rhs) = Thm.term_of ct;
912    val (ls,m) = len lhs ([],0) and (rs,n) = len rhs ([],0);
913    fun prove_neq() =
914      let
915        val Type(_,listT::_) = eqT;
916        val size = HOLogic.size_const listT;
917        val eq_len = HOLogic.mk_eq (size $ lhs, size $ rhs);
918        val neq_len = HOLogic.mk_Trueprop (HOLogic.Not $ eq_len);
919        val thm = Goal.prove ctxt [] [] neq_len
920          (K (simp_tac (put_simpset ss ctxt) 1));
921      in SOME (thm RS @{thm neq_if_length_neq}) end
922  in
923    if m < n andalso submultiset (aconv) (ls,rs) orelse
924       n < m andalso submultiset (aconv) (rs,ls)
925    then prove_neq() else NONE
926  end;
927in K list_neq end;
928\<close>
929
930
931subsubsection \<open>\<open>@\<close> -- append\<close>
932
933global_interpretation append: monoid append Nil
934proof
935  fix xs ys zs :: "'a list"
936  show "(xs @ ys) @ zs = xs @ (ys @ zs)"
937    by (induct xs) simp_all
938  show "xs @ [] = xs"
939    by (induct xs) simp_all
940qed simp
941
942lemma append_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)"
943  by (fact append.assoc)
944
945lemma append_Nil2: "xs @ [] = xs"
946  by (fact append.right_neutral)
947
948lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \<and> ys = [])"
949by (induct xs) auto
950
951lemma Nil_is_append_conv [iff]: "([] = xs @ ys) = (xs = [] \<and> ys = [])"
952by (induct xs) auto
953
954lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])"
955by (induct xs) auto
956
957lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])"
958by (induct xs) auto
959
960lemma append_eq_append_conv [simp]:
961  "length xs = length ys \<or> length us = length vs
962  ==> (xs@us = ys@vs) = (xs=ys \<and> us=vs)"
963apply (induct xs arbitrary: ys)
964 apply (case_tac ys, simp, force)
965apply (case_tac ys, force, simp)
966done
967
968lemma append_eq_append_conv2: "(xs @ ys = zs @ ts) =
969  (\<exists>us. xs = zs @ us \<and> us @ ys = ts \<or> xs @ us = zs \<and> ys = us @ ts)"
970apply (induct xs arbitrary: ys zs ts)
971 apply fastforce
972apply(case_tac zs)
973 apply simp
974apply fastforce
975done
976
977lemma same_append_eq [iff, induct_simp]: "(xs @ ys = xs @ zs) = (ys = zs)"
978by simp
979
980lemma append1_eq_conv [iff]: "(xs @ [x] = ys @ [y]) = (xs = ys \<and> x = y)"
981by simp
982
983lemma append_same_eq [iff, induct_simp]: "(ys @ xs = zs @ xs) = (ys = zs)"
984by simp
985
986lemma append_self_conv2 [iff]: "(xs @ ys = ys) = (xs = [])"
987using append_same_eq [of _ _ "[]"] by auto
988
989lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])"
990using append_same_eq [of "[]"] by auto
991
992lemma hd_Cons_tl: "xs \<noteq> [] ==> hd xs # tl xs = xs"
993  by (fact list.collapse)
994
995lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)"
996by (induct xs) auto
997
998lemma hd_append2 [simp]: "xs \<noteq> [] ==> hd (xs @ ys) = hd xs"
999by (simp add: hd_append split: list.split)
1000
1001lemma tl_append: "tl (xs @ ys) = (case xs of [] \<Rightarrow> tl ys | z#zs \<Rightarrow> zs @ ys)"
1002by (simp split: list.split)
1003
1004lemma tl_append2 [simp]: "xs \<noteq> [] ==> tl (xs @ ys) = tl xs @ ys"
1005by (simp add: tl_append split: list.split)
1006
1007
1008lemma Cons_eq_append_conv: "x#xs = ys@zs =
1009 (ys = [] \<and> x#xs = zs \<or> (\<exists>ys'. x#ys' = ys \<and> xs = ys'@zs))"
1010by(cases ys) auto
1011
1012lemma append_eq_Cons_conv: "(ys@zs = x#xs) =
1013 (ys = [] \<and> zs = x#xs \<or> (\<exists>ys'. ys = x#ys' \<and> ys'@zs = xs))"
1014by(cases ys) auto
1015
1016lemma longest_common_prefix:
1017  "\<exists>ps xs' ys'. xs = ps @ xs' \<and> ys = ps @ ys'
1018       \<and> (xs' = [] \<or> ys' = [] \<or> hd xs' \<noteq> hd ys')"
1019by (induct xs ys rule: list_induct2')
1020   (blast, blast, blast,
1021    metis (no_types, hide_lams) append_Cons append_Nil list.sel(1))
1022
1023text \<open>Trivial rules for solving \<open>@\<close>-equations automatically.\<close>
1024
1025lemma eq_Nil_appendI: "xs = ys ==> xs = [] @ ys"
1026by simp
1027
1028lemma Cons_eq_appendI:
1029"[| x # xs1 = ys; xs = xs1 @ zs |] ==> x # xs = ys @ zs"
1030by (drule sym) simp
1031
1032lemma append_eq_appendI:
1033"[| xs @ xs1 = zs; ys = xs1 @ us |] ==> xs @ ys = zs @ us"
1034by (drule sym) simp
1035
1036
1037text \<open>
1038Simplification procedure for all list equalities.
1039Currently only tries to rearrange \<open>@\<close> to see if
1040- both lists end in a singleton list,
1041- or both lists end in the same list.
1042\<close>
1043
1044simproc_setup list_eq ("(xs::'a list) = ys")  = \<open>
1045  let
1046    fun last (cons as Const (@{const_name Cons}, _) $ _ $ xs) =
1047          (case xs of Const (@{const_name Nil}, _) => cons | _ => last xs)
1048      | last (Const(@{const_name append},_) $ _ $ ys) = last ys
1049      | last t = t;
1050
1051    fun list1 (Const(@{const_name Cons},_) $ _ $ Const(@{const_name Nil},_)) = true
1052      | list1 _ = false;
1053
1054    fun butlast ((cons as Const(@{const_name Cons},_) $ x) $ xs) =
1055          (case xs of Const (@{const_name Nil}, _) => xs | _ => cons $ butlast xs)
1056      | butlast ((app as Const (@{const_name append}, _) $ xs) $ ys) = app $ butlast ys
1057      | butlast xs = Const(@{const_name Nil}, fastype_of xs);
1058
1059    val rearr_ss =
1060      simpset_of (put_simpset HOL_basic_ss @{context}
1061        addsimps [@{thm append_assoc}, @{thm append_Nil}, @{thm append_Cons}]);
1062
1063    fun list_eq ctxt (F as (eq as Const(_,eqT)) $ lhs $ rhs) =
1064      let
1065        val lastl = last lhs and lastr = last rhs;
1066        fun rearr conv =
1067          let
1068            val lhs1 = butlast lhs and rhs1 = butlast rhs;
1069            val Type(_,listT::_) = eqT
1070            val appT = [listT,listT] ---> listT
1071            val app = Const(@{const_name append},appT)
1072            val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr)
1073            val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2));
1074            val thm = Goal.prove ctxt [] [] eq
1075              (K (simp_tac (put_simpset rearr_ss ctxt) 1));
1076          in SOME ((conv RS (thm RS trans)) RS eq_reflection) end;
1077      in
1078        if list1 lastl andalso list1 lastr then rearr @{thm append1_eq_conv}
1079        else if lastl aconv lastr then rearr @{thm append_same_eq}
1080        else NONE
1081      end;
1082  in fn _ => fn ctxt => fn ct => list_eq ctxt (Thm.term_of ct) end;
1083\<close>
1084
1085
1086subsubsection \<open>@{const map}\<close>
1087
1088lemma hd_map: "xs \<noteq> [] \<Longrightarrow> hd (map f xs) = f (hd xs)"
1089by (cases xs) simp_all
1090
1091lemma map_tl: "map f (tl xs) = tl (map f xs)"
1092by (cases xs) simp_all
1093
1094lemma map_ext: "(\<And>x. x \<in> set xs \<longrightarrow> f x = g x) ==> map f xs = map g xs"
1095by (induct xs) simp_all
1096
1097lemma map_ident [simp]: "map (\<lambda>x. x) = (\<lambda>xs. xs)"
1098by (rule ext, induct_tac xs) auto
1099
1100lemma map_append [simp]: "map f (xs @ ys) = map f xs @ map f ys"
1101by (induct xs) auto
1102
1103lemma map_map [simp]: "map f (map g xs) = map (f \<circ> g) xs"
1104by (induct xs) auto
1105
1106lemma map_comp_map[simp]: "((map f) \<circ> (map g)) = map(f \<circ> g)"
1107by (rule ext) simp
1108
1109lemma rev_map: "rev (map f xs) = map f (rev xs)"
1110by (induct xs) auto
1111
1112lemma map_eq_conv[simp]: "(map f xs = map g xs) = (\<forall>x \<in> set xs. f x = g x)"
1113by (induct xs) auto
1114
1115lemma map_cong [fundef_cong]:
1116  "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> map f xs = map g ys"
1117by simp
1118
1119lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])"
1120by (cases xs) auto
1121
1122lemma Nil_is_map_conv [iff]: "([] = map f xs) = (xs = [])"
1123by (cases xs) auto
1124
1125lemma map_eq_Cons_conv:
1126  "(map f xs = y#ys) = (\<exists>z zs. xs = z#zs \<and> f z = y \<and> map f zs = ys)"
1127by (cases xs) auto
1128
1129lemma Cons_eq_map_conv:
1130  "(x#xs = map f ys) = (\<exists>z zs. ys = z#zs \<and> x = f z \<and> xs = map f zs)"
1131by (cases ys) auto
1132
1133lemmas map_eq_Cons_D = map_eq_Cons_conv [THEN iffD1]
1134lemmas Cons_eq_map_D = Cons_eq_map_conv [THEN iffD1]
1135declare map_eq_Cons_D [dest!]  Cons_eq_map_D [dest!]
1136
1137lemma ex_map_conv:
1138  "(\<exists>xs. ys = map f xs) = (\<forall>y \<in> set ys. \<exists>x. y = f x)"
1139by(induct ys, auto simp add: Cons_eq_map_conv)
1140
1141lemma map_eq_imp_length_eq:
1142  assumes "map f xs = map g ys"
1143  shows "length xs = length ys"
1144  using assms
1145proof (induct ys arbitrary: xs)
1146  case Nil then show ?case by simp
1147next
1148  case (Cons y ys) then obtain z zs where xs: "xs = z # zs" by auto
1149  from Cons xs have "map f zs = map g ys" by simp
1150  with Cons have "length zs = length ys" by blast
1151  with xs show ?case by simp
1152qed
1153
1154lemma map_inj_on:
1155 "[| map f xs = map f ys; inj_on f (set xs Un set ys) |]
1156  ==> xs = ys"
1157apply(frule map_eq_imp_length_eq)
1158apply(rotate_tac -1)
1159apply(induct rule:list_induct2)
1160 apply simp
1161apply(simp)
1162apply (blast intro:sym)
1163done
1164
1165lemma inj_on_map_eq_map:
1166  "inj_on f (set xs Un set ys) \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
1167by(blast dest:map_inj_on)
1168
1169lemma map_injective:
1170  "map f xs = map f ys ==> inj f ==> xs = ys"
1171by (induct ys arbitrary: xs) (auto dest!:injD)
1172
1173lemma inj_map_eq_map[simp]: "inj f \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
1174by(blast dest:map_injective)
1175
1176lemma inj_mapI: "inj f ==> inj (map f)"
1177by (iprover dest: map_injective injD intro: inj_onI)
1178
1179lemma inj_mapD: "inj (map f) ==> inj f"
1180  apply (unfold inj_def)
1181  apply clarify
1182  apply (erule_tac x = "[x]" in allE)
1183  apply (erule_tac x = "[y]" in allE)
1184  apply auto
1185  done
1186
1187lemma inj_map[iff]: "inj (map f) = inj f"
1188by (blast dest: inj_mapD intro: inj_mapI)
1189
1190lemma inj_on_mapI: "inj_on f (\<Union>(set ` A)) \<Longrightarrow> inj_on (map f) A"
1191apply(rule inj_onI)
1192apply(erule map_inj_on)
1193apply(blast intro:inj_onI dest:inj_onD)
1194done
1195
1196lemma map_idI: "(\<And>x. x \<in> set xs \<Longrightarrow> f x = x) \<Longrightarrow> map f xs = xs"
1197by (induct xs, auto)
1198
1199lemma map_fun_upd [simp]: "y \<notin> set xs \<Longrightarrow> map (f(y:=v)) xs = map f xs"
1200by (induct xs) auto
1201
1202lemma map_fst_zip[simp]:
1203  "length xs = length ys \<Longrightarrow> map fst (zip xs ys) = xs"
1204by (induct rule:list_induct2, simp_all)
1205
1206lemma map_snd_zip[simp]:
1207  "length xs = length ys \<Longrightarrow> map snd (zip xs ys) = ys"
1208by (induct rule:list_induct2, simp_all)
1209
1210lemma map_fst_zip_take:
1211  "map fst (zip xs ys) = take (min (length xs) (length ys)) xs"
1212by (induct xs ys rule: list_induct2') simp_all
1213
1214lemma map_snd_zip_take:
1215  "map snd (zip xs ys) = take (min (length xs) (length ys)) ys"
1216by (induct xs ys rule: list_induct2') simp_all
1217
1218lemma map2_map_map: "map2 h (map f xs) (map g xs) = map (\<lambda>x. h (f x) (g x)) xs"
1219by (induction xs) (auto)
1220
1221functor map: map
1222by (simp_all add: id_def)
1223
1224declare map.id [simp]
1225
1226
1227subsubsection \<open>@{const rev}\<close>
1228
1229lemma rev_append [simp]: "rev (xs @ ys) = rev ys @ rev xs"
1230by (induct xs) auto
1231
1232lemma rev_rev_ident [simp]: "rev (rev xs) = xs"
1233by (induct xs) auto
1234
1235lemma rev_swap: "(rev xs = ys) = (xs = rev ys)"
1236by auto
1237
1238lemma rev_is_Nil_conv [iff]: "(rev xs = []) = (xs = [])"
1239by (induct xs) auto
1240
1241lemma Nil_is_rev_conv [iff]: "([] = rev xs) = (xs = [])"
1242by (induct xs) auto
1243
1244lemma rev_singleton_conv [simp]: "(rev xs = [x]) = (xs = [x])"
1245by (cases xs) auto
1246
1247lemma singleton_rev_conv [simp]: "([x] = rev xs) = (xs = [x])"
1248by (cases xs) auto
1249
1250lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)"
1251apply (induct xs arbitrary: ys, force)
1252apply (case_tac ys, simp, force)
1253done
1254
1255lemma inj_on_rev[iff]: "inj_on rev A"
1256by(simp add:inj_on_def)
1257
1258lemma rev_induct [case_names Nil snoc]:
1259  "[| P []; !!x xs. P xs ==> P (xs @ [x]) |] ==> P xs"
1260apply(simplesubst rev_rev_ident[symmetric])
1261apply(rule_tac list = "rev xs" in list.induct, simp_all)
1262done
1263
1264lemma rev_exhaust [case_names Nil snoc]:
1265  "(xs = [] ==> P) ==>(!!ys y. xs = ys @ [y] ==> P) ==> P"
1266by (induct xs rule: rev_induct) auto
1267
1268lemmas rev_cases = rev_exhaust
1269
1270lemma rev_nonempty_induct [consumes 1, case_names single snoc]:
1271  assumes "xs \<noteq> []"
1272  and single: "\<And>x. P [x]"
1273  and snoc': "\<And>x xs. xs \<noteq> [] \<Longrightarrow> P xs \<Longrightarrow> P (xs@[x])"
1274  shows "P xs"
1275using \<open>xs \<noteq> []\<close> proof (induct xs rule: rev_induct)
1276  case (snoc x xs) then show ?case
1277  proof (cases xs)
1278    case Nil thus ?thesis by (simp add: single)
1279  next
1280    case Cons with snoc show ?thesis by (fastforce intro!: snoc')
1281  qed
1282qed simp
1283
1284lemma rev_eq_Cons_iff[iff]: "(rev xs = y#ys) = (xs = rev ys @ [y])"
1285by(rule rev_cases[of xs]) auto
1286
1287
1288subsubsection \<open>@{const set}\<close>
1289
1290declare list.set[code_post]  \<comment> \<open>pretty output\<close>
1291
1292lemma finite_set [iff]: "finite (set xs)"
1293by (induct xs) auto
1294
1295lemma set_append [simp]: "set (xs @ ys) = (set xs \<union> set ys)"
1296by (induct xs) auto
1297
1298lemma hd_in_set[simp]: "xs \<noteq> [] \<Longrightarrow> hd xs \<in> set xs"
1299by(cases xs) auto
1300
1301lemma set_subset_Cons: "set xs \<subseteq> set (x # xs)"
1302by auto
1303
1304lemma set_ConsD: "y \<in> set (x # xs) \<Longrightarrow> y=x \<or> y \<in> set xs"
1305by auto
1306
1307lemma set_empty [iff]: "(set xs = {}) = (xs = [])"
1308by (induct xs) auto
1309
1310lemma set_empty2[iff]: "({} = set xs) = (xs = [])"
1311by(induct xs) auto
1312
1313lemma set_rev [simp]: "set (rev xs) = set xs"
1314by (induct xs) auto
1315
1316lemma set_map [simp]: "set (map f xs) = f`(set xs)"
1317by (induct xs) auto
1318
1319lemma set_filter [simp]: "set (filter P xs) = {x. x \<in> set xs \<and> P x}"
1320by (induct xs) auto
1321
1322lemma set_upt [simp]: "set[i..<j] = {i..<j}"
1323by (induct j) auto
1324
1325
1326lemma split_list: "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs"
1327proof (induct xs)
1328  case Nil thus ?case by simp
1329next
1330  case Cons thus ?case by (auto intro: Cons_eq_appendI)
1331qed
1332
1333lemma in_set_conv_decomp: "x \<in> set xs \<longleftrightarrow> (\<exists>ys zs. xs = ys @ x # zs)"
1334  by (auto elim: split_list)
1335
1336lemma split_list_first: "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys"
1337proof (induct xs)
1338  case Nil thus ?case by simp
1339next
1340  case (Cons a xs)
1341  show ?case
1342  proof cases
1343    assume "x = a" thus ?case using Cons by fastforce
1344  next
1345    assume "x \<noteq> a" thus ?case using Cons by(fastforce intro!: Cons_eq_appendI)
1346  qed
1347qed
1348
1349lemma in_set_conv_decomp_first:
1350  "(x \<in> set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys)"
1351  by (auto dest!: split_list_first)
1352
1353lemma split_list_last: "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs"
1354proof (induct xs rule: rev_induct)
1355  case Nil thus ?case by simp
1356next
1357  case (snoc a xs)
1358  show ?case
1359  proof cases
1360    assume "x = a" thus ?case using snoc by (auto intro!: exI)
1361  next
1362    assume "x \<noteq> a" thus ?case using snoc by fastforce
1363  qed
1364qed
1365
1366lemma in_set_conv_decomp_last:
1367  "(x \<in> set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs)"
1368  by (auto dest!: split_list_last)
1369
1370lemma split_list_prop: "\<exists>x \<in> set xs. P x \<Longrightarrow> \<exists>ys x zs. xs = ys @ x # zs \<and> P x"
1371proof (induct xs)
1372  case Nil thus ?case by simp
1373next
1374  case Cons thus ?case
1375    by(simp add:Bex_def)(metis append_Cons append.simps(1))
1376qed
1377
1378lemma split_list_propE:
1379  assumes "\<exists>x \<in> set xs. P x"
1380  obtains ys x zs where "xs = ys @ x # zs" and "P x"
1381using split_list_prop [OF assms] by blast
1382
1383lemma split_list_first_prop:
1384  "\<exists>x \<in> set xs. P x \<Longrightarrow>
1385   \<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y)"
1386proof (induct xs)
1387  case Nil thus ?case by simp
1388next
1389  case (Cons x xs)
1390  show ?case
1391  proof cases
1392    assume "P x"
1393    hence "x # xs = [] @ x # xs \<and> P x \<and> (\<forall>y\<in>set []. \<not> P y)" by simp
1394    thus ?thesis by fast
1395  next
1396    assume "\<not> P x"
1397    hence "\<exists>x\<in>set xs. P x" using Cons(2) by simp
1398    thus ?thesis using \<open>\<not> P x\<close> Cons(1) by (metis append_Cons set_ConsD)
1399  qed
1400qed
1401
1402lemma split_list_first_propE:
1403  assumes "\<exists>x \<in> set xs. P x"
1404  obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>y \<in> set ys. \<not> P y"
1405using split_list_first_prop [OF assms] by blast
1406
1407lemma split_list_first_prop_iff:
1408  "(\<exists>x \<in> set xs. P x) \<longleftrightarrow>
1409   (\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>y \<in> set ys. \<not> P y))"
1410by (rule, erule split_list_first_prop) auto
1411
1412lemma split_list_last_prop:
1413  "\<exists>x \<in> set xs. P x \<Longrightarrow>
1414   \<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z)"
1415proof(induct xs rule:rev_induct)
1416  case Nil thus ?case by simp
1417next
1418  case (snoc x xs)
1419  show ?case
1420  proof cases
1421    assume "P x" thus ?thesis by (auto intro!: exI)
1422  next
1423    assume "\<not> P x"
1424    hence "\<exists>x\<in>set xs. P x" using snoc(2) by simp
1425    thus ?thesis using \<open>\<not> P x\<close> snoc(1) by fastforce
1426  qed
1427qed
1428
1429lemma split_list_last_propE:
1430  assumes "\<exists>x \<in> set xs. P x"
1431  obtains ys x zs where "xs = ys @ x # zs" and "P x" and "\<forall>z \<in> set zs. \<not> P z"
1432using split_list_last_prop [OF assms] by blast
1433
1434lemma split_list_last_prop_iff:
1435  "(\<exists>x \<in> set xs. P x) \<longleftrightarrow>
1436   (\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z))"
1437  by rule (erule split_list_last_prop, auto)
1438
1439
1440lemma finite_list: "finite A \<Longrightarrow> \<exists>xs. set xs = A"
1441  by (erule finite_induct) (auto simp add: list.set(2)[symmetric] simp del: list.set(2))
1442
1443lemma card_length: "card (set xs) \<le> length xs"
1444by (induct xs) (auto simp add: card_insert_if)
1445
1446lemma set_minus_filter_out:
1447  "set xs - {y} = set (filter (\<lambda>x. \<not> (x = y)) xs)"
1448  by (induct xs) auto
1449
1450lemma append_Cons_eq_iff:
1451  "\<lbrakk> x \<notin> set xs; x \<notin> set ys \<rbrakk> \<Longrightarrow>
1452   xs @ x # ys = xs' @ x # ys' \<longleftrightarrow> (xs = xs' \<and> ys = ys')"
1453by(auto simp: append_eq_Cons_conv Cons_eq_append_conv append_eq_append_conv2)
1454
1455
1456subsubsection \<open>@{const filter}\<close>
1457
1458lemma filter_append [simp]: "filter P (xs @ ys) = filter P xs @ filter P ys"
1459by (induct xs) auto
1460
1461lemma rev_filter: "rev (filter P xs) = filter P (rev xs)"
1462by (induct xs) simp_all
1463
1464lemma filter_filter [simp]: "filter P (filter Q xs) = filter (\<lambda>x. Q x \<and> P x) xs"
1465by (induct xs) auto
1466
1467lemma length_filter_le [simp]: "length (filter P xs) \<le> length xs"
1468by (induct xs) (auto simp add: le_SucI)
1469
1470lemma sum_length_filter_compl:
1471  "length(filter P xs) + length(filter (\<lambda>x. \<not>P x) xs) = length xs"
1472by(induct xs) simp_all
1473
1474lemma filter_True [simp]: "\<forall>x \<in> set xs. P x ==> filter P xs = xs"
1475by (induct xs) auto
1476
1477lemma filter_False [simp]: "\<forall>x \<in> set xs. \<not> P x ==> filter P xs = []"
1478by (induct xs) auto
1479
1480lemma filter_empty_conv: "(filter P xs = []) = (\<forall>x\<in>set xs. \<not> P x)"
1481by (induct xs) simp_all
1482
1483lemma filter_id_conv: "(filter P xs = xs) = (\<forall>x\<in>set xs. P x)"
1484apply (induct xs)
1485 apply auto
1486apply(cut_tac P=P and xs=xs in length_filter_le)
1487apply simp
1488done
1489
1490lemma filter_map: "filter P (map f xs) = map f (filter (P \<circ> f) xs)"
1491by (induct xs) simp_all
1492
1493lemma length_filter_map[simp]:
1494  "length (filter P (map f xs)) = length(filter (P \<circ> f) xs)"
1495by (simp add:filter_map)
1496
1497lemma filter_is_subset [simp]: "set (filter P xs) \<le> set xs"
1498by auto
1499
1500lemma length_filter_less:
1501  "\<lbrakk> x \<in> set xs; \<not> P x \<rbrakk> \<Longrightarrow> length(filter P xs) < length xs"
1502proof (induct xs)
1503  case Nil thus ?case by simp
1504next
1505  case (Cons x xs) thus ?case
1506    apply (auto split:if_split_asm)
1507    using length_filter_le[of P xs] apply arith
1508  done
1509qed
1510
1511lemma length_filter_conv_card:
1512  "length(filter p xs) = card{i. i < length xs \<and> p(xs!i)}"
1513proof (induct xs)
1514  case Nil thus ?case by simp
1515next
1516  case (Cons x xs)
1517  let ?S = "{i. i < length xs \<and> p(xs!i)}"
1518  have fin: "finite ?S" by(fast intro: bounded_nat_set_is_finite)
1519  show ?case (is "?l = card ?S'")
1520  proof (cases)
1521    assume "p x"
1522    hence eq: "?S' = insert 0 (Suc ` ?S)"
1523      by(auto simp: image_def split:nat.split dest:gr0_implies_Suc)
1524    have "length (filter p (x # xs)) = Suc(card ?S)"
1525      using Cons \<open>p x\<close> by simp
1526    also have "\<dots> = Suc(card(Suc ` ?S))" using fin
1527      by (simp add: card_image)
1528    also have "\<dots> = card ?S'" using eq fin
1529      by (simp add:card_insert_if) (simp add:image_def)
1530    finally show ?thesis .
1531  next
1532    assume "\<not> p x"
1533    hence eq: "?S' = Suc ` ?S"
1534      by(auto simp add: image_def split:nat.split elim:lessE)
1535    have "length (filter p (x # xs)) = card ?S"
1536      using Cons \<open>\<not> p x\<close> by simp
1537    also have "\<dots> = card(Suc ` ?S)" using fin
1538      by (simp add: card_image)
1539    also have "\<dots> = card ?S'" using eq fin
1540      by (simp add:card_insert_if)
1541    finally show ?thesis .
1542  qed
1543qed
1544
1545lemma Cons_eq_filterD:
1546  "x#xs = filter P ys \<Longrightarrow>
1547  \<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
1548  (is "_ \<Longrightarrow> \<exists>us vs. ?P ys us vs")
1549proof(induct ys)
1550  case Nil thus ?case by simp
1551next
1552  case (Cons y ys)
1553  show ?case (is "\<exists>x. ?Q x")
1554  proof cases
1555    assume Py: "P y"
1556    show ?thesis
1557    proof cases
1558      assume "x = y"
1559      with Py Cons.prems have "?Q []" by simp
1560      then show ?thesis ..
1561    next
1562      assume "x \<noteq> y"
1563      with Py Cons.prems show ?thesis by simp
1564    qed
1565  next
1566    assume "\<not> P y"
1567    with Cons obtain us vs where "?P (y#ys) (y#us) vs" by fastforce
1568    then have "?Q (y#us)" by simp
1569    then show ?thesis ..
1570  qed
1571qed
1572
1573lemma filter_eq_ConsD:
1574  "filter P ys = x#xs \<Longrightarrow>
1575  \<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
1576by(rule Cons_eq_filterD) simp
1577
1578lemma filter_eq_Cons_iff:
1579  "(filter P ys = x#xs) =
1580  (\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
1581by(auto dest:filter_eq_ConsD)
1582
1583lemma Cons_eq_filter_iff:
1584  "(x#xs = filter P ys) =
1585  (\<exists>us vs. ys = us @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs)"
1586by(auto dest:Cons_eq_filterD)
1587
1588lemma inj_on_filter_key_eq:
1589  assumes "inj_on f (insert y (set xs))"
1590  shows "filter (\<lambda>x. f y = f x) xs = filter (HOL.eq y) xs"
1591  using assms by (induct xs) auto
1592
1593lemma filter_cong[fundef_cong]:
1594  "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> P x = Q x) \<Longrightarrow> filter P xs = filter Q ys"
1595apply simp
1596apply(erule thin_rl)
1597by (induct ys) simp_all
1598
1599
1600subsubsection \<open>List partitioning\<close>
1601
1602primrec partition :: "('a \<Rightarrow> bool) \<Rightarrow>'a list \<Rightarrow> 'a list \<times> 'a list" where
1603"partition P [] = ([], [])" |
1604"partition P (x # xs) =
1605  (let (yes, no) = partition P xs
1606   in if P x then (x # yes, no) else (yes, x # no))"
1607
1608lemma partition_filter1: "fst (partition P xs) = filter P xs"
1609by (induct xs) (auto simp add: Let_def split_def)
1610
1611lemma partition_filter2: "snd (partition P xs) = filter (Not \<circ> P) xs"
1612by (induct xs) (auto simp add: Let_def split_def)
1613
1614lemma partition_P:
1615  assumes "partition P xs = (yes, no)"
1616  shows "(\<forall>p \<in> set yes.  P p) \<and> (\<forall>p  \<in> set no. \<not> P p)"
1617proof -
1618  from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)"
1619    by simp_all
1620  then show ?thesis by (simp_all add: partition_filter1 partition_filter2)
1621qed
1622
1623lemma partition_set:
1624  assumes "partition P xs = (yes, no)"
1625  shows "set yes \<union> set no = set xs"
1626proof -
1627  from assms have "yes = fst (partition P xs)" and "no = snd (partition P xs)"
1628    by simp_all
1629  then show ?thesis by (auto simp add: partition_filter1 partition_filter2)
1630qed
1631
1632lemma partition_filter_conv[simp]:
1633  "partition f xs = (filter f xs,filter (Not \<circ> f) xs)"
1634unfolding partition_filter2[symmetric]
1635unfolding partition_filter1[symmetric] by simp
1636
1637declare partition.simps[simp del]
1638
1639
1640subsubsection \<open>@{const concat}\<close>
1641
1642lemma concat_append [simp]: "concat (xs @ ys) = concat xs @ concat ys"
1643by (induct xs) auto
1644
1645lemma concat_eq_Nil_conv [simp]: "(concat xss = []) = (\<forall>xs \<in> set xss. xs = [])"
1646by (induct xss) auto
1647
1648lemma Nil_eq_concat_conv [simp]: "([] = concat xss) = (\<forall>xs \<in> set xss. xs = [])"
1649by (induct xss) auto
1650
1651lemma set_concat [simp]: "set (concat xs) = (UN x:set xs. set x)"
1652by (induct xs) auto
1653
1654lemma concat_map_singleton[simp]: "concat(map (%x. [f x]) xs) = map f xs"
1655by (induct xs) auto
1656
1657lemma map_concat: "map f (concat xs) = concat (map (map f) xs)"
1658by (induct xs) auto
1659
1660lemma filter_concat: "filter p (concat xs) = concat (map (filter p) xs)"
1661by (induct xs) auto
1662
1663lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))"
1664by (induct xs) auto
1665
1666lemma 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)"
1667proof (induct xs arbitrary: ys)
1668  case (Cons x xs ys)
1669  thus ?case by (cases ys) auto
1670qed (auto)
1671
1672lemma concat_injective: "concat xs = concat ys ==> length xs = length ys ==> \<forall>(x, y) \<in> set (zip xs ys). length x = length y ==> xs = ys"
1673by (simp add: concat_eq_concat_iff)
1674
1675
1676subsubsection \<open>@{const nth}\<close>
1677
1678lemma nth_Cons_0 [simp, code]: "(x # xs)!0 = x"
1679by auto
1680
1681lemma nth_Cons_Suc [simp, code]: "(x # xs)!(Suc n) = xs!n"
1682by auto
1683
1684declare nth.simps [simp del]
1685
1686lemma nth_Cons_pos[simp]: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)"
1687by(auto simp: Nat.gr0_conv_Suc)
1688
1689lemma nth_append:
1690  "(xs @ ys)!n = (if n < length xs then xs!n else ys!(n - length xs))"
1691apply (induct xs arbitrary: n, simp)
1692apply (case_tac n, auto)
1693done
1694
1695lemma nth_append_length [simp]: "(xs @ x # ys) ! length xs = x"
1696by (induct xs) auto
1697
1698lemma nth_append_length_plus[simp]: "(xs @ ys) ! (length xs + n) = ys ! n"
1699by (induct xs) auto
1700
1701lemma nth_map [simp]: "n < length xs ==> (map f xs)!n = f(xs!n)"
1702apply (induct xs arbitrary: n, simp)
1703apply (case_tac n, auto)
1704done
1705
1706lemma nth_tl: "n < length (tl xs) \<Longrightarrow> tl xs ! n = xs ! Suc n"
1707by (induction xs) auto
1708
1709lemma hd_conv_nth: "xs \<noteq> [] \<Longrightarrow> hd xs = xs!0"
1710by(cases xs) simp_all
1711
1712
1713lemma list_eq_iff_nth_eq:
1714  "(xs = ys) = (length xs = length ys \<and> (\<forall>i<length xs. xs!i = ys!i))"
1715apply(induct xs arbitrary: ys)
1716 apply force
1717apply(case_tac ys)
1718 apply simp
1719apply(simp add:nth_Cons split:nat.split)apply blast
1720done
1721
1722lemma set_conv_nth: "set xs = {xs!i | i. i < length xs}"
1723apply (induct xs, simp, simp)
1724apply safe
1725apply (metis nat.case(1) nth.simps zero_less_Suc)
1726apply (metis less_Suc_eq_0_disj nth_Cons_Suc)
1727apply (case_tac i, simp)
1728apply (metis diff_Suc_Suc nat.case(2) nth.simps zero_less_diff)
1729done
1730
1731lemma in_set_conv_nth: "(x \<in> set xs) = (\<exists>i < length xs. xs!i = x)"
1732by(auto simp:set_conv_nth)
1733
1734lemma nth_equal_first_eq:
1735  assumes "x \<notin> set xs"
1736  assumes "n \<le> length xs"
1737  shows "(x # xs) ! n = x \<longleftrightarrow> n = 0" (is "?lhs \<longleftrightarrow> ?rhs")
1738proof
1739  assume ?lhs
1740  show ?rhs
1741  proof (rule ccontr)
1742    assume "n \<noteq> 0"
1743    then have "n > 0" by simp
1744    with \<open>?lhs\<close> have "xs ! (n - 1) = x" by simp
1745    moreover from \<open>n > 0\<close> \<open>n \<le> length xs\<close> have "n - 1 < length xs" by simp
1746    ultimately have "\<exists>i<length xs. xs ! i = x" by auto
1747    with \<open>x \<notin> set xs\<close> in_set_conv_nth [of x xs] show False by simp
1748  qed
1749next
1750  assume ?rhs then show ?lhs by simp
1751qed
1752
1753lemma nth_non_equal_first_eq:
1754  assumes "x \<noteq> y"
1755  shows "(x # xs) ! n = y \<longleftrightarrow> xs ! (n - 1) = y \<and> n > 0" (is "?lhs \<longleftrightarrow> ?rhs")
1756proof
1757  assume "?lhs" with assms have "n > 0" by (cases n) simp_all
1758  with \<open>?lhs\<close> show ?rhs by simp
1759next
1760  assume "?rhs" then show "?lhs" by simp
1761qed
1762
1763lemma list_ball_nth: "\<lbrakk>n < length xs; \<forall>x \<in> set xs. P x\<rbrakk> \<Longrightarrow> P(xs!n)"
1764by (auto simp add: set_conv_nth)
1765
1766lemma nth_mem [simp]: "n < length xs \<Longrightarrow> xs!n \<in> set xs"
1767by (auto simp add: set_conv_nth)
1768
1769lemma all_nth_imp_all_set:
1770  "\<lbrakk>\<forall>i < length xs. P(xs!i); x \<in> set xs\<rbrakk> \<Longrightarrow> P x"
1771by (auto simp add: set_conv_nth)
1772
1773lemma all_set_conv_all_nth:
1774  "(\<forall>x \<in> set xs. P x) = (\<forall>i. i < length xs \<longrightarrow> P (xs ! i))"
1775by (auto simp add: set_conv_nth)
1776
1777lemma rev_nth:
1778  "n < size xs \<Longrightarrow> rev xs ! n = xs ! (length xs - Suc n)"
1779proof (induct xs arbitrary: n)
1780  case Nil thus ?case by simp
1781next
1782  case (Cons x xs)
1783  hence n: "n < Suc (length xs)" by simp
1784  moreover
1785  { assume "n < length xs"
1786    with n obtain n' where n': "length xs - n = Suc n'"
1787      by (cases "length xs - n", auto)
1788    moreover
1789    from n' have "length xs - Suc n = n'" by simp
1790    ultimately
1791    have "xs ! (length xs - Suc n) = (x # xs) ! (length xs - n)" by simp
1792  }
1793  ultimately
1794  show ?case by (clarsimp simp add: Cons nth_append)
1795qed
1796
1797lemma Skolem_list_nth:
1798  "(\<forall>i<k. \<exists>x. P i x) = (\<exists>xs. size xs = k \<and> (\<forall>i<k. P i (xs!i)))"
1799  (is "_ = (\<exists>xs. ?P k xs)")
1800proof(induct k)
1801  case 0 show ?case by simp
1802next
1803  case (Suc k)
1804  show ?case (is "?L = ?R" is "_ = (\<exists>xs. ?P' xs)")
1805  proof
1806    assume "?R" thus "?L" using Suc by auto
1807  next
1808    assume "?L"
1809    with Suc obtain x xs where "?P k xs \<and> P k x" by (metis less_Suc_eq)
1810    hence "?P'(xs@[x])" by(simp add:nth_append less_Suc_eq)
1811    thus "?R" ..
1812  qed
1813qed
1814
1815
1816subsubsection \<open>@{const list_update}\<close>
1817
1818lemma length_list_update [simp]: "length(xs[i:=x]) = length xs"
1819by (induct xs arbitrary: i) (auto split: nat.split)
1820
1821lemma nth_list_update:
1822"i < length xs==> (xs[i:=x])!j = (if i = j then x else xs!j)"
1823by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
1824
1825lemma nth_list_update_eq [simp]: "i < length xs ==> (xs[i:=x])!i = x"
1826by (simp add: nth_list_update)
1827
1828lemma nth_list_update_neq [simp]: "i \<noteq> j ==> xs[i:=x]!j = xs!j"
1829by (induct xs arbitrary: i j) (auto simp add: nth_Cons split: nat.split)
1830
1831lemma list_update_id[simp]: "xs[i := xs!i] = xs"
1832by (induct xs arbitrary: i) (simp_all split:nat.splits)
1833
1834lemma list_update_beyond[simp]: "length xs \<le> i \<Longrightarrow> xs[i:=x] = xs"
1835apply (induct xs arbitrary: i)
1836 apply simp
1837apply (case_tac i)
1838apply simp_all
1839done
1840
1841lemma list_update_nonempty[simp]: "xs[k:=x] = [] \<longleftrightarrow> xs=[]"
1842by (simp only: length_0_conv[symmetric] length_list_update)
1843
1844lemma list_update_same_conv:
1845  "i < length xs ==> (xs[i := x] = xs) = (xs!i = x)"
1846by (induct xs arbitrary: i) (auto split: nat.split)
1847
1848lemma list_update_append1:
1849  "i < size xs \<Longrightarrow> (xs @ ys)[i:=x] = xs[i:=x] @ ys"
1850by (induct xs arbitrary: i)(auto split:nat.split)
1851
1852lemma list_update_append:
1853  "(xs @ ys) [n:= x] =
1854  (if n < length xs then xs[n:= x] @ ys else xs @ (ys [n-length xs:= x]))"
1855by (induct xs arbitrary: n) (auto split:nat.splits)
1856
1857lemma list_update_length [simp]:
1858  "(xs @ x # ys)[length xs := y] = (xs @ y # ys)"
1859by (induct xs, auto)
1860
1861lemma map_update: "map f (xs[k:= y]) = (map f xs)[k := f y]"
1862by(induct xs arbitrary: k)(auto split:nat.splits)
1863
1864lemma rev_update:
1865  "k < length xs \<Longrightarrow> rev (xs[k:= y]) = (rev xs)[length xs - k - 1 := y]"
1866by (induct xs arbitrary: k) (auto simp: list_update_append split:nat.splits)
1867
1868lemma update_zip:
1869  "(zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])"
1870by (induct ys arbitrary: i xy xs) (auto, case_tac xs, auto split: nat.split)
1871
1872lemma set_update_subset_insert: "set(xs[i:=x]) <= insert x (set xs)"
1873by (induct xs arbitrary: i) (auto split: nat.split)
1874
1875lemma set_update_subsetI: "\<lbrakk>set xs \<subseteq> A; x \<in> A\<rbrakk> \<Longrightarrow> set(xs[i := x]) \<subseteq> A"
1876by (blast dest!: set_update_subset_insert [THEN subsetD])
1877
1878lemma set_update_memI: "n < length xs \<Longrightarrow> x \<in> set (xs[n := x])"
1879by (induct xs arbitrary: n) (auto split:nat.splits)
1880
1881lemma list_update_overwrite[simp]:
1882  "xs [i := x, i := y] = xs [i := y]"
1883apply (induct xs arbitrary: i) apply simp
1884apply (case_tac i, simp_all)
1885done
1886
1887lemma list_update_swap:
1888  "i \<noteq> i' \<Longrightarrow> xs [i := x, i' := x'] = xs [i' := x', i := x]"
1889apply (induct xs arbitrary: i i')
1890 apply simp
1891apply (case_tac i, case_tac i')
1892  apply auto
1893apply (case_tac i')
1894apply auto
1895done
1896
1897lemma list_update_code [code]:
1898  "[][i := y] = []"
1899  "(x # xs)[0 := y] = y # xs"
1900  "(x # xs)[Suc i := y] = x # xs[i := y]"
1901by simp_all
1902
1903
1904subsubsection \<open>@{const last} and @{const butlast}\<close>
1905
1906lemma last_snoc [simp]: "last (xs @ [x]) = x"
1907by (induct xs) auto
1908
1909lemma butlast_snoc [simp]: "butlast (xs @ [x]) = xs"
1910by (induct xs) auto
1911
1912lemma last_ConsL: "xs = [] \<Longrightarrow> last(x#xs) = x"
1913by simp
1914
1915lemma last_ConsR: "xs \<noteq> [] \<Longrightarrow> last(x#xs) = last xs"
1916by simp
1917
1918lemma last_append: "last(xs @ ys) = (if ys = [] then last xs else last ys)"
1919by (induct xs) (auto)
1920
1921lemma last_appendL[simp]: "ys = [] \<Longrightarrow> last(xs @ ys) = last xs"
1922by(simp add:last_append)
1923
1924lemma last_appendR[simp]: "ys \<noteq> [] \<Longrightarrow> last(xs @ ys) = last ys"
1925by(simp add:last_append)
1926
1927lemma last_tl: "xs = [] \<or> tl xs \<noteq> [] \<Longrightarrow>last (tl xs) = last xs"
1928by (induct xs) simp_all
1929
1930lemma butlast_tl: "butlast (tl xs) = tl (butlast xs)"
1931by (induct xs) simp_all
1932
1933lemma hd_rev: "xs \<noteq> [] \<Longrightarrow> hd(rev xs) = last xs"
1934by(rule rev_exhaust[of xs]) simp_all
1935
1936lemma last_rev: "xs \<noteq> [] \<Longrightarrow> last(rev xs) = hd xs"
1937by(cases xs) simp_all
1938
1939lemma last_in_set[simp]: "as \<noteq> [] \<Longrightarrow> last as \<in> set as"
1940by (induct as) auto
1941
1942lemma length_butlast [simp]: "length (butlast xs) = length xs - 1"
1943by (induct xs rule: rev_induct) auto
1944
1945lemma butlast_append:
1946  "butlast (xs @ ys) = (if ys = [] then butlast xs else xs @ butlast ys)"
1947by (induct xs arbitrary: ys) auto
1948
1949lemma append_butlast_last_id [simp]:
1950  "xs \<noteq> [] \<Longrightarrow> butlast xs @ [last xs] = xs"
1951by (induct xs) auto
1952
1953lemma in_set_butlastD: "x \<in> set (butlast xs) \<Longrightarrow> x \<in> set xs"
1954by (induct xs) (auto split: if_split_asm)
1955
1956lemma in_set_butlast_appendI:
1957  "x \<in> set (butlast xs) \<or> x \<in> set (butlast ys) \<Longrightarrow> x \<in> set (butlast (xs @ ys))"
1958by (auto dest: in_set_butlastD simp add: butlast_append)
1959
1960lemma last_drop[simp]: "n < length xs \<Longrightarrow> last (drop n xs) = last xs"
1961by (induct xs arbitrary: n)(auto split:nat.split)
1962
1963lemma nth_butlast:
1964  assumes "n < length (butlast xs)" shows "butlast xs ! n = xs ! n"
1965proof (cases xs)
1966  case (Cons y ys)
1967  moreover from assms have "butlast xs ! n = (butlast xs @ [last xs]) ! n"
1968    by (simp add: nth_append)
1969  ultimately show ?thesis using append_butlast_last_id by simp
1970qed simp
1971
1972lemma last_conv_nth: "xs\<noteq>[] \<Longrightarrow> last xs = xs!(length xs - 1)"
1973by(induct xs)(auto simp:neq_Nil_conv)
1974
1975lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs"
1976by (induction xs rule: induct_list012) simp_all
1977
1978lemma last_list_update:
1979  "xs \<noteq> [] \<Longrightarrow> last(xs[k:=x]) = (if k = size xs - 1 then x else last xs)"
1980by (auto simp: last_conv_nth)
1981
1982lemma butlast_list_update:
1983  "butlast(xs[k:=x]) =
1984  (if k = size xs - 1 then butlast xs else (butlast xs)[k:=x])"
1985by(cases xs rule:rev_cases)(auto simp: list_update_append split: nat.splits)
1986
1987lemma last_map: "xs \<noteq> [] \<Longrightarrow> last (map f xs) = f (last xs)"
1988by (cases xs rule: rev_cases) simp_all
1989
1990lemma map_butlast: "map f (butlast xs) = butlast (map f xs)"
1991by (induct xs) simp_all
1992
1993lemma snoc_eq_iff_butlast:
1994  "xs @ [x] = ys \<longleftrightarrow> (ys \<noteq> [] \<and> butlast ys = xs \<and> last ys = x)"
1995by fastforce
1996
1997corollary longest_common_suffix:
1998  "\<exists>ss xs' ys'. xs = xs' @ ss \<and> ys = ys' @ ss
1999       \<and> (xs' = [] \<or> ys' = [] \<or> last xs' \<noteq> last ys')"
2000using longest_common_prefix[of "rev xs" "rev ys"]
2001unfolding rev_swap rev_append by (metis last_rev rev_is_Nil_conv)
2002
2003
2004subsubsection \<open>@{const take} and @{const drop}\<close>
2005
2006lemma take_0: "take 0 xs = []"
2007by (induct xs) auto
2008
2009lemma drop_0: "drop 0 xs = xs"
2010by (induct xs) auto
2011
2012lemma take0[simp]: "take 0 = (\<lambda>xs. [])"
2013by(rule ext) (rule take_0)
2014
2015lemma drop0[simp]: "drop 0 = (\<lambda>x. x)"
2016by(rule ext) (rule drop_0)
2017
2018lemma take_Suc_Cons [simp]: "take (Suc n) (x # xs) = x # take n xs"
2019by simp
2020
2021lemma drop_Suc_Cons [simp]: "drop (Suc n) (x # xs) = drop n xs"
2022by simp
2023
2024declare take_Cons [simp del] and drop_Cons [simp del]
2025
2026lemma take_Suc: "xs \<noteq> [] \<Longrightarrow> take (Suc n) xs = hd xs # take n (tl xs)"
2027by(clarsimp simp add:neq_Nil_conv)
2028
2029lemma drop_Suc: "drop (Suc n) xs = drop n (tl xs)"
2030by(cases xs, simp_all)
2031
2032lemma hd_take[simp]: "j > 0 \<Longrightarrow> hd (take j xs) = hd xs"
2033by (metis gr0_conv_Suc list.sel(1) take.simps(1) take_Suc)
2034
2035lemma take_tl: "take n (tl xs) = tl (take (Suc n) xs)"
2036by (induct xs arbitrary: n) simp_all
2037
2038lemma drop_tl: "drop n (tl xs) = tl(drop n xs)"
2039by(induct xs arbitrary: n, simp_all add:drop_Cons drop_Suc split:nat.split)
2040
2041lemma tl_take: "tl (take n xs) = take (n - 1) (tl xs)"
2042by (cases n, simp, cases xs, auto)
2043
2044lemma tl_drop: "tl (drop n xs) = drop n (tl xs)"
2045by (simp only: drop_tl)
2046
2047lemma nth_via_drop: "drop n xs = y#ys \<Longrightarrow> xs!n = y"
2048by (induct xs arbitrary: n, simp)(auto simp: drop_Cons nth_Cons split: nat.splits)
2049
2050lemma take_Suc_conv_app_nth:
2051  "i < length xs \<Longrightarrow> take (Suc i) xs = take i xs @ [xs!i]"
2052apply (induct xs arbitrary: i, simp)
2053apply (case_tac i, auto)
2054done
2055
2056lemma Cons_nth_drop_Suc:
2057  "i < length xs \<Longrightarrow> (xs!i) # (drop (Suc i) xs) = drop i xs"
2058apply (induct xs arbitrary: i, simp)
2059apply (case_tac i, auto)
2060done
2061
2062lemma length_take [simp]: "length (take n xs) = min (length xs) n"
2063by (induct n arbitrary: xs) (auto, case_tac xs, auto)
2064
2065lemma length_drop [simp]: "length (drop n xs) = (length xs - n)"
2066by (induct n arbitrary: xs) (auto, case_tac xs, auto)
2067
2068lemma take_all [simp]: "length xs <= n ==> take n xs = xs"
2069by (induct n arbitrary: xs) (auto, case_tac xs, auto)
2070
2071lemma drop_all [simp]: "length xs <= n ==> drop n xs = []"
2072by (induct n arbitrary: xs) (auto, case_tac xs, auto)
2073
2074lemma take_append [simp]:
2075  "take n (xs @ ys) = (take n xs @ take (n - length xs) ys)"
2076by (induct n arbitrary: xs) (auto, case_tac xs, auto)
2077
2078lemma drop_append [simp]:
2079  "drop n (xs @ ys) = drop n xs @ drop (n - length xs) ys"
2080by (induct n arbitrary: xs) (auto, case_tac xs, auto)
2081
2082lemma take_take [simp]: "take n (take m xs) = take (min n m) xs"
2083apply (induct m arbitrary: xs n, auto)
2084 apply (case_tac xs, auto)
2085apply (case_tac n, auto)
2086done
2087
2088lemma drop_drop [simp]: "drop n (drop m xs) = drop (n + m) xs"
2089apply (induct m arbitrary: xs, auto)
2090 apply (case_tac xs, auto)
2091done
2092
2093lemma take_drop: "take n (drop m xs) = drop m (take (n + m) xs)"
2094apply (induct m arbitrary: xs n, auto)
2095 apply (case_tac xs, auto)
2096done
2097
2098lemma drop_take: "drop n (take m xs) = take (m-n) (drop n xs)"
2099by(induct xs arbitrary: m n)(auto simp: take_Cons drop_Cons split: nat.split)
2100
2101lemma append_take_drop_id [simp]: "take n xs @ drop n xs = xs"
2102apply (induct n arbitrary: xs, auto)
2103apply (case_tac xs, auto)
2104done
2105
2106lemma take_eq_Nil[simp]: "(take n xs = []) = (n = 0 \<or> xs = [])"
2107by(induct xs arbitrary: n)(auto simp: take_Cons split:nat.split)
2108
2109lemma drop_eq_Nil[simp]: "(drop n xs = []) = (length xs <= n)"
2110by (induct xs arbitrary: n) (auto simp: drop_Cons split:nat.split)
2111
2112lemma take_map: "take n (map f xs) = map f (take n xs)"
2113apply (induct n arbitrary: xs, auto)
2114 apply (case_tac xs, auto)
2115done
2116
2117lemma drop_map: "drop n (map f xs) = map f (drop n xs)"
2118apply (induct n arbitrary: xs, auto)
2119 apply (case_tac xs, auto)
2120done
2121
2122lemma rev_take: "rev (take i xs) = drop (length xs - i) (rev xs)"
2123apply (induct xs arbitrary: i, auto)
2124 apply (case_tac i, auto)
2125done
2126
2127lemma rev_drop: "rev (drop i xs) = take (length xs - i) (rev xs)"
2128apply (induct xs arbitrary: i, auto)
2129 apply (case_tac i, auto)
2130done
2131
2132lemma drop_rev: "drop n (rev xs) = rev (take (length xs - n) xs)"
2133  by (cases "length xs < n") (auto simp: rev_take)
2134
2135lemma take_rev: "take n (rev xs) = rev (drop (length xs - n) xs)"
2136  by (cases "length xs < n") (auto simp: rev_drop)
2137
2138lemma nth_take [simp]: "i < n ==> (take n xs)!i = xs!i"
2139apply (induct xs arbitrary: i n, auto)
2140 apply (case_tac n, blast)
2141apply (case_tac i, auto)
2142done
2143
2144lemma nth_drop [simp]:
2145  "n <= length xs ==> (drop n xs)!i = xs!(n + i)"
2146apply (induct n arbitrary: xs i, auto)
2147 apply (case_tac xs, auto)
2148done
2149
2150lemma butlast_take:
2151  "n <= length xs ==> butlast (take n xs) = take (n - 1) xs"
2152by (simp add: butlast_conv_take min.absorb1 min.absorb2)
2153
2154lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)"
2155by (simp add: butlast_conv_take drop_take ac_simps)
2156
2157lemma take_butlast: "n < length xs ==> take n (butlast xs) = take n xs"
2158by (simp add: butlast_conv_take min.absorb1)
2159
2160lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)"
2161by (simp add: butlast_conv_take drop_take ac_simps)
2162
2163lemma hd_drop_conv_nth: "n < length xs \<Longrightarrow> hd(drop n xs) = xs!n"
2164by(simp add: hd_conv_nth)
2165
2166lemma set_take_subset_set_take:
2167  "m <= n \<Longrightarrow> set(take m xs) <= set(take n xs)"
2168apply (induct xs arbitrary: m n)
2169 apply simp
2170apply (case_tac n)
2171apply (auto simp: take_Cons)
2172done
2173
2174lemma set_take_subset: "set(take n xs) \<subseteq> set xs"
2175by(induct xs arbitrary: n)(auto simp:take_Cons split:nat.split)
2176
2177lemma set_drop_subset: "set(drop n xs) \<subseteq> set xs"
2178by(induct xs arbitrary: n)(auto simp:drop_Cons split:nat.split)
2179
2180lemma set_drop_subset_set_drop:
2181  "m >= n \<Longrightarrow> set(drop m xs) <= set(drop n xs)"
2182apply(induct xs arbitrary: m n)
2183 apply(auto simp:drop_Cons split:nat.split)
2184by (metis set_drop_subset subset_iff)
2185
2186lemma in_set_takeD: "x \<in> set(take n xs) \<Longrightarrow> x \<in> set xs"
2187using set_take_subset by fast
2188
2189lemma in_set_dropD: "x \<in> set(drop n xs) \<Longrightarrow> x \<in> set xs"
2190using set_drop_subset by fast
2191
2192lemma append_eq_conv_conj:
2193  "(xs @ ys = zs) = (xs = take (length xs) zs \<and> ys = drop (length xs) zs)"
2194apply (induct xs arbitrary: zs, simp, clarsimp)
2195 apply (case_tac zs, auto)
2196done
2197
2198lemma take_add:  "take (i+j) xs = take i xs @ take j (drop i xs)"
2199apply (induct xs arbitrary: i, auto)
2200 apply (case_tac i, simp_all)
2201done
2202
2203lemma append_eq_append_conv_if:
2204  "(xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>1 @ ys\<^sub>2) =
2205  (if size xs\<^sub>1 \<le> size ys\<^sub>1
2206   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
2207   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)"
2208apply(induct xs\<^sub>1 arbitrary: ys\<^sub>1)
2209 apply simp
2210apply(case_tac ys\<^sub>1)
2211apply simp_all
2212done
2213
2214lemma take_hd_drop:
2215  "n < length xs \<Longrightarrow> take n xs @ [hd (drop n xs)] = take (Suc n) xs"
2216apply(induct xs arbitrary: n)
2217 apply simp
2218apply(simp add:drop_Cons split:nat.split)
2219done
2220
2221lemma id_take_nth_drop:
2222  "i < length xs \<Longrightarrow> xs = take i xs @ xs!i # drop (Suc i) xs"
2223proof -
2224  assume si: "i < length xs"
2225  hence "xs = take (Suc i) xs @ drop (Suc i) xs" by auto
2226  moreover
2227  from si have "take (Suc i) xs = take i xs @ [xs!i]"
2228    apply (rule_tac take_Suc_conv_app_nth) by arith
2229  ultimately show ?thesis by auto
2230qed
2231
2232lemma take_update_cancel[simp]: "n \<le> m \<Longrightarrow> take n (xs[m := y]) = take n xs"
2233by(simp add: list_eq_iff_nth_eq)
2234
2235lemma drop_update_cancel[simp]: "n < m \<Longrightarrow> drop m (xs[n := x]) = drop m xs"
2236by(simp add: list_eq_iff_nth_eq)
2237
2238lemma upd_conv_take_nth_drop:
2239  "i < length xs \<Longrightarrow> xs[i:=a] = take i xs @ a # drop (Suc i) xs"
2240proof -
2241  assume i: "i < length xs"
2242  have "xs[i:=a] = (take i xs @ xs!i # drop (Suc i) xs)[i:=a]"
2243    by(rule arg_cong[OF id_take_nth_drop[OF i]])
2244  also have "\<dots> = take i xs @ a # drop (Suc i) xs"
2245    using i by (simp add: list_update_append)
2246  finally show ?thesis .
2247qed
2248
2249lemma take_update_swap: "take m (xs[n := x]) = (take m xs)[n := x]"
2250apply(cases "n \<ge> length xs")
2251 apply simp
2252apply(simp add: upd_conv_take_nth_drop take_Cons drop_take min_def diff_Suc
2253  split: nat.split)
2254done
2255
2256lemma drop_update_swap: "m \<le> n \<Longrightarrow> drop m (xs[n := x]) = (drop m xs)[n-m := x]"
2257apply(cases "n \<ge> length xs")
2258 apply simp
2259apply(simp add: upd_conv_take_nth_drop drop_take)
2260done
2261
2262lemma nth_image: "l \<le> size xs \<Longrightarrow> nth xs ` {0..<l} = set(take l xs)"
2263by(auto simp: set_conv_nth image_def) (metis Suc_le_eq nth_take order_trans)
2264
2265
2266subsubsection \<open>@{const takeWhile} and @{const dropWhile}\<close>
2267
2268lemma length_takeWhile_le: "length (takeWhile P xs) \<le> length xs"
2269by (induct xs) auto
2270
2271lemma takeWhile_dropWhile_id [simp]: "takeWhile P xs @ dropWhile P xs = xs"
2272by (induct xs) auto
2273
2274lemma takeWhile_append1 [simp]:
2275  "\<lbrakk>x \<in> set xs; \<not>P(x)\<rbrakk> \<Longrightarrow> takeWhile P (xs @ ys) = takeWhile P xs"
2276by (induct xs) auto
2277
2278lemma takeWhile_append2 [simp]:
2279  "(\<And>x. x \<in> set xs \<Longrightarrow> P x) \<Longrightarrow> takeWhile P (xs @ ys) = xs @ takeWhile P ys"
2280by (induct xs) auto
2281
2282lemma takeWhile_tail: "\<not> P x \<Longrightarrow> takeWhile P (xs @ (x#l)) = takeWhile P xs"
2283by (induct xs) auto
2284
2285lemma takeWhile_nth: "j < length (takeWhile P xs) \<Longrightarrow> takeWhile P xs ! j = xs ! j"
2286apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto
2287
2288lemma dropWhile_nth: "j < length (dropWhile P xs) \<Longrightarrow>
2289  dropWhile P xs ! j = xs ! (j + length (takeWhile P xs))"
2290apply (subst (3) takeWhile_dropWhile_id[symmetric]) unfolding nth_append by auto
2291
2292lemma length_dropWhile_le: "length (dropWhile P xs) \<le> length xs"
2293by (induct xs) auto
2294
2295lemma dropWhile_append1 [simp]:
2296  "\<lbrakk>x \<in> set xs; \<not>P(x)\<rbrakk> \<Longrightarrow> dropWhile P (xs @ ys) = (dropWhile P xs)@ys"
2297by (induct xs) auto
2298
2299lemma dropWhile_append2 [simp]:
2300  "(\<And>x. x \<in> set xs \<Longrightarrow> P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys"
2301by (induct xs) auto
2302
2303lemma dropWhile_append3:
2304  "\<not> P y \<Longrightarrow>dropWhile P (xs @ y # ys) = dropWhile P xs @ y # ys"
2305by (induct xs) auto
2306
2307lemma dropWhile_last:
2308  "x \<in> set xs \<Longrightarrow> \<not> P x \<Longrightarrow> last (dropWhile P xs) = last xs"
2309by (auto simp add: dropWhile_append3 in_set_conv_decomp)
2310
2311lemma set_dropWhileD: "x \<in> set (dropWhile P xs) \<Longrightarrow> x \<in> set xs"
2312by (induct xs) (auto split: if_split_asm)
2313
2314lemma set_takeWhileD: "x \<in> set (takeWhile P xs) \<Longrightarrow> x \<in> set xs \<and> P x"
2315by (induct xs) (auto split: if_split_asm)
2316
2317lemma takeWhile_eq_all_conv[simp]:
2318  "(takeWhile P xs = xs) = (\<forall>x \<in> set xs. P x)"
2319by(induct xs, auto)
2320
2321lemma dropWhile_eq_Nil_conv[simp]:
2322  "(dropWhile P xs = []) = (\<forall>x \<in> set xs. P x)"
2323by(induct xs, auto)
2324
2325lemma dropWhile_eq_Cons_conv:
2326  "(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys \<and> \<not> P y)"
2327by(induct xs, auto)
2328
2329lemma distinct_takeWhile[simp]: "distinct xs ==> distinct (takeWhile P xs)"
2330by (induct xs) (auto dest: set_takeWhileD)
2331
2332lemma distinct_dropWhile[simp]: "distinct xs ==> distinct (dropWhile P xs)"
2333by (induct xs) auto
2334
2335lemma takeWhile_map: "takeWhile P (map f xs) = map f (takeWhile (P \<circ> f) xs)"
2336by (induct xs) auto
2337
2338lemma dropWhile_map: "dropWhile P (map f xs) = map f (dropWhile (P \<circ> f) xs)"
2339by (induct xs) auto
2340
2341lemma takeWhile_eq_take: "takeWhile P xs = take (length (takeWhile P xs)) xs"
2342by (induct xs) auto
2343
2344lemma dropWhile_eq_drop: "dropWhile P xs = drop (length (takeWhile P xs)) xs"
2345by (induct xs) auto
2346
2347lemma hd_dropWhile: "dropWhile P xs \<noteq> [] \<Longrightarrow> \<not> P (hd (dropWhile P xs))"
2348by (induct xs) auto
2349
2350lemma takeWhile_eq_filter:
2351  assumes "\<And> x. x \<in> set (dropWhile P xs) \<Longrightarrow> \<not> P x"
2352  shows "takeWhile P xs = filter P xs"
2353proof -
2354  have A: "filter P xs = filter P (takeWhile P xs @ dropWhile P xs)"
2355    by simp
2356  have B: "filter P (dropWhile P xs) = []"
2357    unfolding filter_empty_conv using assms by blast
2358  have "filter P xs = takeWhile P xs"
2359    unfolding A filter_append B
2360    by (auto simp add: filter_id_conv dest: set_takeWhileD)
2361  thus ?thesis ..
2362qed
2363
2364lemma takeWhile_eq_take_P_nth:
2365  "\<lbrakk> \<And> i. \<lbrakk> i < n ; i < length xs \<rbrakk> \<Longrightarrow> P (xs ! i) ; n < length xs \<Longrightarrow> \<not> P (xs ! n) \<rbrakk> \<Longrightarrow>
2366  takeWhile P xs = take n xs"
2367proof (induct xs arbitrary: n)
2368  case Nil
2369  thus ?case by simp
2370next
2371  case (Cons x xs)
2372  show ?case
2373  proof (cases n)
2374    case 0
2375    with Cons show ?thesis by simp
2376  next
2377    case [simp]: (Suc n')
2378    have "P x" using Cons.prems(1)[of 0] by simp
2379    moreover have "takeWhile P xs = take n' xs"
2380    proof (rule Cons.hyps)
2381      fix i
2382      assume "i < n'" "i < length xs"
2383      thus "P (xs ! i)" using Cons.prems(1)[of "Suc i"] by simp
2384    next
2385      assume "n' < length xs"
2386      thus "\<not> P (xs ! n')" using Cons by auto
2387    qed
2388    ultimately show ?thesis by simp
2389   qed
2390qed
2391
2392lemma nth_length_takeWhile:
2393  "length (takeWhile P xs) < length xs \<Longrightarrow> \<not> P (xs ! length (takeWhile P xs))"
2394by (induct xs) auto
2395
2396lemma length_takeWhile_less_P_nth:
2397  assumes all: "\<And> i. i < j \<Longrightarrow> P (xs ! i)" and "j \<le> length xs"
2398  shows "j \<le> length (takeWhile P xs)"
2399proof (rule classical)
2400  assume "\<not> ?thesis"
2401  hence "length (takeWhile P xs) < length xs" using assms by simp
2402  thus ?thesis using all \<open>\<not> ?thesis\<close> nth_length_takeWhile[of P xs] by auto
2403qed
2404
2405lemma takeWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow>
2406  takeWhile (\<lambda>y. y \<noteq> x) (rev xs) = rev (tl (dropWhile (\<lambda>y. y \<noteq> x) xs))"
2407by(induct xs) (auto simp: takeWhile_tail[where l="[]"])
2408
2409lemma dropWhile_neq_rev: "\<lbrakk>distinct xs; x \<in> set xs\<rbrakk> \<Longrightarrow>
2410  dropWhile (\<lambda>y. y \<noteq> x) (rev xs) = x # rev (takeWhile (\<lambda>y. y \<noteq> x) xs)"
2411apply(induct xs)
2412 apply simp
2413apply auto
2414apply(subst dropWhile_append2)
2415apply auto
2416done
2417
2418lemma takeWhile_not_last:
2419  "distinct xs \<Longrightarrow> takeWhile (\<lambda>y. y \<noteq> last xs) xs = butlast xs"
2420by(induction xs rule: induct_list012) auto
2421
2422lemma takeWhile_cong [fundef_cong]:
2423  "\<lbrakk>l = k; \<And>x. x \<in> set l \<Longrightarrow> P x = Q x\<rbrakk>
2424  \<Longrightarrow> takeWhile P l = takeWhile Q k"
2425by (induct k arbitrary: l) (simp_all)
2426
2427lemma dropWhile_cong [fundef_cong]:
2428  "\<lbrakk>l = k; \<And>x. x \<in> set l \<Longrightarrow> P x = Q x\<rbrakk>
2429  \<Longrightarrow> dropWhile P l = dropWhile Q k"
2430by (induct k arbitrary: l, simp_all)
2431
2432lemma takeWhile_idem [simp]:
2433  "takeWhile P (takeWhile P xs) = takeWhile P xs"
2434by (induct xs) auto
2435
2436lemma dropWhile_idem [simp]:
2437  "dropWhile P (dropWhile P xs) = dropWhile P xs"
2438by (induct xs) auto
2439
2440
2441subsubsection \<open>@{const zip}\<close>
2442
2443lemma zip_Nil [simp]: "zip [] ys = []"
2444by (induct ys) auto
2445
2446lemma zip_Cons_Cons [simp]: "zip (x # xs) (y # ys) = (x, y) # zip xs ys"
2447by simp
2448
2449declare zip_Cons [simp del]
2450
2451lemma [code]:
2452  "zip [] ys = []"
2453  "zip xs [] = []"
2454  "zip (x # xs) (y # ys) = (x, y) # zip xs ys"
2455by (fact zip_Nil zip.simps(1) zip_Cons_Cons)+
2456
2457lemma zip_Cons1:
2458  "zip (x#xs) ys = (case ys of [] \<Rightarrow> [] | y#ys \<Rightarrow> (x,y)#zip xs ys)"
2459by(auto split:list.split)
2460
2461lemma length_zip [simp]:
2462  "length (zip xs ys) = min (length xs) (length ys)"
2463by (induct xs ys rule:list_induct2') auto
2464
2465lemma zip_obtain_same_length:
2466  assumes "\<And>zs ws n. length zs = length ws \<Longrightarrow> n = min (length xs) (length ys)
2467    \<Longrightarrow> zs = take n xs \<Longrightarrow> ws = take n ys \<Longrightarrow> P (zip zs ws)"
2468  shows "P (zip xs ys)"
2469proof -
2470  let ?n = "min (length xs) (length ys)"
2471  have "P (zip (take ?n xs) (take ?n ys))"
2472    by (rule assms) simp_all
2473  moreover have "zip xs ys = zip (take ?n xs) (take ?n ys)"
2474  proof (induct xs arbitrary: ys)
2475    case Nil then show ?case by simp
2476  next
2477    case (Cons x xs) then show ?case by (cases ys) simp_all
2478  qed
2479  ultimately show ?thesis by simp
2480qed
2481
2482lemma zip_append1:
2483  "zip (xs @ ys) zs =
2484  zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)"
2485by (induct xs zs rule:list_induct2') auto
2486
2487lemma zip_append2:
2488  "zip xs (ys @ zs) =
2489  zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs"
2490by (induct xs ys rule:list_induct2') auto
2491
2492lemma zip_append [simp]:
2493  "[| length xs = length us |] ==>
2494  zip (xs@ys) (us@vs) = zip xs us @ zip ys vs"
2495by (simp add: zip_append1)
2496
2497lemma zip_rev:
2498  "length xs = length ys ==> zip (rev xs) (rev ys) = rev (zip xs ys)"
2499by (induct rule:list_induct2, simp_all)
2500
2501lemma zip_map_map:
2502  "zip (map f xs) (map g ys) = map (\<lambda> (x, y). (f x, g y)) (zip xs ys)"
2503proof (induct xs arbitrary: ys)
2504  case (Cons x xs) note Cons_x_xs = Cons.hyps
2505  show ?case
2506  proof (cases ys)
2507    case (Cons y ys')
2508    show ?thesis unfolding Cons using Cons_x_xs by simp
2509  qed simp
2510qed simp
2511
2512lemma zip_map1:
2513  "zip (map f xs) ys = map (\<lambda>(x, y). (f x, y)) (zip xs ys)"
2514using zip_map_map[of f xs "\<lambda>x. x" ys] by simp
2515
2516lemma zip_map2:
2517  "zip xs (map f ys) = map (\<lambda>(x, y). (x, f y)) (zip xs ys)"
2518using zip_map_map[of "\<lambda>x. x" xs f ys] by simp
2519
2520lemma map_zip_map:
2521  "map f (zip (map g xs) ys) = map (%(x,y). f(g x, y)) (zip xs ys)"
2522by (auto simp: zip_map1)
2523
2524lemma map_zip_map2:
2525  "map f (zip xs (map g ys)) = map (%(x,y). f(x, g y)) (zip xs ys)"
2526by (auto simp: zip_map2)
2527
2528text\<open>Courtesy of Andreas Lochbihler:\<close>
2529lemma zip_same_conv_map: "zip xs xs = map (\<lambda>x. (x, x)) xs"
2530by(induct xs) auto
2531
2532lemma nth_zip [simp]:
2533  "[| i < length xs; i < length ys|] ==> (zip xs ys)!i = (xs!i, ys!i)"
2534apply (induct ys arbitrary: i xs, simp)
2535apply (case_tac xs)
2536 apply (simp_all add: nth.simps split: nat.split)
2537done
2538
2539lemma set_zip:
2540  "set (zip xs ys) = {(xs!i, ys!i) | i. i < min (length xs) (length ys)}"
2541by(simp add: set_conv_nth cong: rev_conj_cong)
2542
2543lemma zip_same: "((a,b) \<in> set (zip xs xs)) = (a \<in> set xs \<and> a = b)"
2544by(induct xs) auto
2545
2546lemma zip_update:
2547  "zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]"
2548by(rule sym, simp add: update_zip)
2549
2550lemma zip_replicate [simp]:
2551  "zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)"
2552apply (induct i arbitrary: j, auto)
2553apply (case_tac j, auto)
2554done
2555
2556lemma zip_replicate1: "zip (replicate n x) ys = map (Pair x) (take n ys)"
2557by(induction ys arbitrary: n)(case_tac [2] n, simp_all)
2558
2559lemma take_zip:
2560  "take n (zip xs ys) = zip (take n xs) (take n ys)"
2561apply (induct n arbitrary: xs ys)
2562 apply simp
2563apply (case_tac xs, simp)
2564apply (case_tac ys, simp_all)
2565done
2566
2567lemma drop_zip:
2568  "drop n (zip xs ys) = zip (drop n xs) (drop n ys)"
2569apply (induct n arbitrary: xs ys)
2570 apply simp
2571apply (case_tac xs, simp)
2572apply (case_tac ys, simp_all)
2573done
2574
2575lemma zip_takeWhile_fst: "zip (takeWhile P xs) ys = takeWhile (P \<circ> fst) (zip xs ys)"
2576proof (induct xs arbitrary: ys)
2577  case (Cons x xs) thus ?case by (cases ys) auto
2578qed simp
2579
2580lemma zip_takeWhile_snd: "zip xs (takeWhile P ys) = takeWhile (P \<circ> snd) (zip xs ys)"
2581proof (induct xs arbitrary: ys)
2582  case (Cons x xs) thus ?case by (cases ys) auto
2583qed simp
2584
2585lemma set_zip_leftD: "(x,y)\<in> set (zip xs ys) \<Longrightarrow> x \<in> set xs"
2586by (induct xs ys rule:list_induct2') auto
2587
2588lemma set_zip_rightD: "(x,y)\<in> set (zip xs ys) \<Longrightarrow> y \<in> set ys"
2589by (induct xs ys rule:list_induct2') auto
2590
2591lemma in_set_zipE:
2592  "(x,y) \<in> set(zip xs ys) \<Longrightarrow> (\<lbrakk> x \<in> set xs; y \<in> set ys \<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R"
2593by(blast dest: set_zip_leftD set_zip_rightD)
2594
2595lemma zip_map_fst_snd: "zip (map fst zs) (map snd zs) = zs"
2596by (induct zs) simp_all
2597
2598lemma zip_eq_conv:
2599  "length xs = length ys \<Longrightarrow> zip xs ys = zs \<longleftrightarrow> map fst zs = xs \<and> map snd zs = ys"
2600by (auto simp add: zip_map_fst_snd)
2601
2602lemma in_set_zip:
2603  "p \<in> set (zip xs ys) \<longleftrightarrow> (\<exists>n. xs ! n = fst p \<and> ys ! n = snd p
2604  \<and> n < length xs \<and> n < length ys)"
2605by (cases p) (auto simp add: set_zip)
2606
2607lemma in_set_impl_in_set_zip1:
2608  assumes "length xs = length ys"
2609  assumes "x \<in> set xs"
2610  obtains y where "(x, y) \<in> set (zip xs ys)"
2611proof -
2612  from assms have "x \<in> set (map fst (zip xs ys))" by simp
2613  from this that show ?thesis by fastforce
2614qed
2615
2616lemma in_set_impl_in_set_zip2:
2617  assumes "length xs = length ys"
2618  assumes "y \<in> set ys"
2619  obtains x where "(x, y) \<in> set (zip xs ys)"
2620proof -
2621  from assms have "y \<in> set (map snd (zip xs ys))" by simp
2622  from this that show ?thesis by fastforce
2623qed
2624
2625lemma pair_list_eqI:
2626  assumes "map fst xs = map fst ys" and "map snd xs = map snd ys"
2627  shows "xs = ys"
2628proof -
2629  from assms(1) have "length xs = length ys" by (rule map_eq_imp_length_eq)
2630  from this assms show ?thesis
2631    by (induct xs ys rule: list_induct2) (simp_all add: prod_eqI)
2632qed
2633
2634
2635subsubsection \<open>@{const list_all2}\<close>
2636
2637lemma list_all2_lengthD [intro?]:
2638  "list_all2 P xs ys ==> length xs = length ys"
2639by (simp add: list_all2_iff)
2640
2641lemma list_all2_Nil [iff, code]: "list_all2 P [] ys = (ys = [])"
2642by (simp add: list_all2_iff)
2643
2644lemma list_all2_Nil2 [iff, code]: "list_all2 P xs [] = (xs = [])"
2645by (simp add: list_all2_iff)
2646
2647lemma list_all2_Cons [iff, code]:
2648  "list_all2 P (x # xs) (y # ys) = (P x y \<and> list_all2 P xs ys)"
2649by (auto simp add: list_all2_iff)
2650
2651lemma list_all2_Cons1:
2652  "list_all2 P (x # xs) ys = (\<exists>z zs. ys = z # zs \<and> P x z \<and> list_all2 P xs zs)"
2653by (cases ys) auto
2654
2655lemma list_all2_Cons2:
2656  "list_all2 P xs (y # ys) = (\<exists>z zs. xs = z # zs \<and> P z y \<and> list_all2 P zs ys)"
2657by (cases xs) auto
2658
2659lemma list_all2_induct
2660  [consumes 1, case_names Nil Cons, induct set: list_all2]:
2661  assumes P: "list_all2 P xs ys"
2662  assumes Nil: "R [] []"
2663  assumes Cons: "\<And>x xs y ys.
2664    \<lbrakk>P x y; list_all2 P xs ys; R xs ys\<rbrakk> \<Longrightarrow> R (x # xs) (y # ys)"
2665  shows "R xs ys"
2666using P
2667by (induct xs arbitrary: ys) (auto simp add: list_all2_Cons1 Nil Cons)
2668
2669lemma list_all2_rev [iff]:
2670  "list_all2 P (rev xs) (rev ys) = list_all2 P xs ys"
2671by (simp add: list_all2_iff zip_rev cong: conj_cong)
2672
2673lemma list_all2_rev1:
2674  "list_all2 P (rev xs) ys = list_all2 P xs (rev ys)"
2675by (subst list_all2_rev [symmetric]) simp
2676
2677lemma list_all2_append1:
2678  "list_all2 P (xs @ ys) zs =
2679  (\<exists>us vs. zs = us @ vs \<and> length us = length xs \<and> length vs = length ys \<and>
2680    list_all2 P xs us \<and> list_all2 P ys vs)"
2681apply (simp add: list_all2_iff zip_append1)
2682apply (rule iffI)
2683 apply (rule_tac x = "take (length xs) zs" in exI)
2684 apply (rule_tac x = "drop (length xs) zs" in exI)
2685 apply (force split: nat_diff_split simp add: min_def, clarify)
2686apply (simp add: ball_Un)
2687done
2688
2689lemma list_all2_append2:
2690  "list_all2 P xs (ys @ zs) =
2691  (\<exists>us vs. xs = us @ vs \<and> length us = length ys \<and> length vs = length zs \<and>
2692    list_all2 P us ys \<and> list_all2 P vs zs)"
2693apply (simp add: list_all2_iff zip_append2)
2694apply (rule iffI)
2695 apply (rule_tac x = "take (length ys) xs" in exI)
2696 apply (rule_tac x = "drop (length ys) xs" in exI)
2697 apply (force split: nat_diff_split simp add: min_def, clarify)
2698apply (simp add: ball_Un)
2699done
2700
2701lemma list_all2_append:
2702  "length xs = length ys \<Longrightarrow>
2703  list_all2 P (xs@us) (ys@vs) = (list_all2 P xs ys \<and> list_all2 P us vs)"
2704by (induct rule:list_induct2, simp_all)
2705
2706lemma list_all2_appendI [intro?, trans]:
2707  "\<lbrakk> list_all2 P a b; list_all2 P c d \<rbrakk> \<Longrightarrow> list_all2 P (a@c) (b@d)"
2708by (simp add: list_all2_append list_all2_lengthD)
2709
2710lemma list_all2_conv_all_nth:
2711  "list_all2 P xs ys =
2712  (length xs = length ys \<and> (\<forall>i < length xs. P (xs!i) (ys!i)))"
2713by (force simp add: list_all2_iff set_zip)
2714
2715lemma list_all2_trans:
2716  assumes tr: "!!a b c. P1 a b ==> P2 b c ==> P3 a c"
2717  shows "!!bs cs. list_all2 P1 as bs ==> list_all2 P2 bs cs ==> list_all2 P3 as cs"
2718        (is "!!bs cs. PROP ?Q as bs cs")
2719proof (induct as)
2720  fix x xs bs assume I1: "!!bs cs. PROP ?Q xs bs cs"
2721  show "!!cs. PROP ?Q (x # xs) bs cs"
2722  proof (induct bs)
2723    fix y ys cs assume I2: "!!cs. PROP ?Q (x # xs) ys cs"
2724    show "PROP ?Q (x # xs) (y # ys) cs"
2725      by (induct cs) (auto intro: tr I1 I2)
2726  qed simp
2727qed simp
2728
2729lemma list_all2_all_nthI [intro?]:
2730  "length a = length b \<Longrightarrow> (\<And>n. n < length a \<Longrightarrow> P (a!n) (b!n)) \<Longrightarrow> list_all2 P a b"
2731by (simp add: list_all2_conv_all_nth)
2732
2733lemma list_all2I:
2734  "\<forall>x \<in> set (zip a b). case_prod P x \<Longrightarrow> length a = length b \<Longrightarrow> list_all2 P a b"
2735by (simp add: list_all2_iff)
2736
2737lemma list_all2_nthD:
2738  "\<lbrakk> list_all2 P xs ys; p < size xs \<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)"
2739by (simp add: list_all2_conv_all_nth)
2740
2741lemma list_all2_nthD2:
2742  "\<lbrakk>list_all2 P xs ys; p < size ys\<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)"
2743by (frule list_all2_lengthD) (auto intro: list_all2_nthD)
2744
2745lemma list_all2_map1:
2746  "list_all2 P (map f as) bs = list_all2 (\<lambda>x y. P (f x) y) as bs"
2747by (simp add: list_all2_conv_all_nth)
2748
2749lemma list_all2_map2:
2750  "list_all2 P as (map f bs) = list_all2 (\<lambda>x y. P x (f y)) as bs"
2751by (auto simp add: list_all2_conv_all_nth)
2752
2753lemma list_all2_refl [intro?]:
2754  "(\<And>x. P x x) \<Longrightarrow> list_all2 P xs xs"
2755by (simp add: list_all2_conv_all_nth)
2756
2757lemma list_all2_update_cong:
2758  "\<lbrakk> list_all2 P xs ys; P x y \<rbrakk> \<Longrightarrow> list_all2 P (xs[i:=x]) (ys[i:=y])"
2759by (cases "i < length ys") (auto simp add: list_all2_conv_all_nth nth_list_update)
2760
2761lemma list_all2_takeI [simp,intro?]:
2762  "list_all2 P xs ys \<Longrightarrow> list_all2 P (take n xs) (take n ys)"
2763apply (induct xs arbitrary: n ys)
2764 apply simp
2765apply (clarsimp simp add: list_all2_Cons1)
2766apply (case_tac n)
2767apply auto
2768done
2769
2770lemma list_all2_dropI [simp,intro?]:
2771  "list_all2 P as bs \<Longrightarrow> list_all2 P (drop n as) (drop n bs)"
2772apply (induct as arbitrary: n bs, simp)
2773apply (clarsimp simp add: list_all2_Cons1)
2774apply (case_tac n, simp, simp)
2775done
2776
2777lemma list_all2_mono [intro?]:
2778  "list_all2 P xs ys \<Longrightarrow> (\<And>xs ys. P xs ys \<Longrightarrow> Q xs ys) \<Longrightarrow> list_all2 Q xs ys"
2779apply (induct xs arbitrary: ys, simp)
2780apply (case_tac ys, auto)
2781done
2782
2783lemma list_all2_eq:
2784  "xs = ys \<longleftrightarrow> list_all2 (=) xs ys"
2785by (induct xs ys rule: list_induct2') auto
2786
2787lemma list_eq_iff_zip_eq:
2788  "xs = ys \<longleftrightarrow> length xs = length ys \<and> (\<forall>(x,y) \<in> set (zip xs ys). x = y)"
2789by(auto simp add: set_zip list_all2_eq list_all2_conv_all_nth cong: conj_cong)
2790
2791lemma list_all2_same: "list_all2 P xs xs \<longleftrightarrow> (\<forall>x\<in>set xs. P x x)"
2792by(auto simp add: list_all2_conv_all_nth set_conv_nth)
2793
2794lemma zip_assoc:
2795  "zip xs (zip ys zs) = map (\<lambda>((x, y), z). (x, y, z)) (zip (zip xs ys) zs)"
2796by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all
2797
2798lemma zip_commute: "zip xs ys = map (\<lambda>(x, y). (y, x)) (zip ys xs)"
2799by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all
2800
2801lemma zip_left_commute:
2802  "zip xs (zip ys zs) = map (\<lambda>(y, (x, z)). (x, y, z)) (zip ys (zip xs zs))"
2803by(rule list_all2_all_nthI[where P="(=)", unfolded list.rel_eq]) simp_all
2804
2805lemma zip_replicate2: "zip xs (replicate n y) = map (\<lambda>x. (x, y)) (take n xs)"
2806by(subst zip_commute)(simp add: zip_replicate1)
2807
2808subsubsection \<open>@{const List.product} and @{const product_lists}\<close>
2809
2810lemma product_concat_map:
2811  "List.product xs ys = concat (map (\<lambda>x. map (\<lambda>y. (x,y)) ys) xs)"
2812by(induction xs) (simp)+
2813
2814lemma set_product[simp]: "set (List.product xs ys) = set xs \<times> set ys"
2815by (induct xs) auto
2816
2817lemma length_product [simp]:
2818  "length (List.product xs ys) = length xs * length ys"
2819by (induct xs) simp_all
2820
2821lemma product_nth:
2822  assumes "n < length xs * length ys"
2823  shows "List.product xs ys ! n = (xs ! (n div length ys), ys ! (n mod length ys))"
2824using assms proof (induct xs arbitrary: n)
2825  case Nil then show ?case by simp
2826next
2827  case (Cons x xs n)
2828  then have "length ys > 0" by auto
2829  with Cons show ?case
2830    by (auto simp add: nth_append not_less le_mod_geq le_div_geq)
2831qed
2832
2833lemma in_set_product_lists_length:
2834  "xs \<in> set (product_lists xss) \<Longrightarrow> length xs = length xss"
2835by (induct xss arbitrary: xs) auto
2836
2837lemma product_lists_set:
2838  "set (product_lists xss) = {xs. list_all2 (\<lambda>x ys. x \<in> set ys) xs xss}" (is "?L = Collect ?R")
2839proof (intro equalityI subsetI, unfold mem_Collect_eq)
2840  fix xs assume "xs \<in> ?L"
2841  then have "length xs = length xss" by (rule in_set_product_lists_length)
2842  from this \<open>xs \<in> ?L\<close> show "?R xs" by (induct xs xss rule: list_induct2) auto
2843next
2844  fix xs assume "?R xs"
2845  then show "xs \<in> ?L" by induct auto
2846qed
2847
2848
2849subsubsection \<open>@{const fold} with natural argument order\<close>
2850
2851lemma fold_simps [code]: \<comment> \<open>eta-expanded variant for generated code -- enables tail-recursion optimisation in Scala\<close>
2852  "fold f [] s = s"
2853  "fold f (x # xs) s = fold f xs (f x s)"
2854by simp_all
2855
2856lemma fold_remove1_split:
2857  "\<lbrakk> \<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f x \<circ> f y = f y \<circ> f x;
2858    x \<in> set xs \<rbrakk>
2859  \<Longrightarrow> fold f xs = fold f (remove1 x xs) \<circ> f x"
2860by (induct xs) (auto simp add: comp_assoc)
2861
2862lemma fold_cong [fundef_cong]:
2863  "a = b \<Longrightarrow> xs = ys \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> f x = g x)
2864    \<Longrightarrow> fold f xs a = fold g ys b"
2865by (induct ys arbitrary: a b xs) simp_all
2866
2867lemma fold_id: "(\<And>x. x \<in> set xs \<Longrightarrow> f x = id) \<Longrightarrow> fold f xs = id"
2868by (induct xs) simp_all
2869
2870lemma fold_commute:
2871  "(\<And>x. x \<in> set xs \<Longrightarrow> h \<circ> g x = f x \<circ> h) \<Longrightarrow> h \<circ> fold g xs = fold f xs \<circ> h"
2872by (induct xs) (simp_all add: fun_eq_iff)
2873
2874lemma fold_commute_apply:
2875  assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<circ> g x = f x \<circ> h"
2876  shows "h (fold g xs s) = fold f xs (h s)"
2877proof -
2878  from assms have "h \<circ> fold g xs = fold f xs \<circ> h" by (rule fold_commute)
2879  then show ?thesis by (simp add: fun_eq_iff)
2880qed
2881
2882lemma fold_invariant:
2883  "\<lbrakk> \<And>x. x \<in> set xs \<Longrightarrow> Q x;  P s;  \<And>x s. Q x \<Longrightarrow> P s \<Longrightarrow> P (f x s) \<rbrakk>
2884  \<Longrightarrow> P (fold f xs s)"
2885by (induct xs arbitrary: s) simp_all
2886
2887lemma fold_append [simp]: "fold f (xs @ ys) = fold f ys \<circ> fold f xs"
2888by (induct xs) simp_all
2889
2890lemma fold_map [code_unfold]: "fold g (map f xs) = fold (g \<circ> f) xs"
2891by (induct xs) simp_all
2892
2893lemma fold_filter:
2894  "fold f (filter P xs) = fold (\<lambda>x. if P x then f x else id) xs"
2895by (induct xs) simp_all
2896
2897lemma fold_rev:
2898  "(\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y)
2899  \<Longrightarrow> fold f (rev xs) = fold f xs"
2900by (induct xs) (simp_all add: fold_commute_apply fun_eq_iff)
2901
2902lemma fold_Cons_rev: "fold Cons xs = append (rev xs)"
2903by (induct xs) simp_all
2904
2905lemma rev_conv_fold [code]: "rev xs = fold Cons xs []"
2906by (simp add: fold_Cons_rev)
2907
2908lemma fold_append_concat_rev: "fold append xss = append (concat (rev xss))"
2909by (induct xss) simp_all
2910
2911text \<open>@{const Finite_Set.fold} and @{const fold}\<close>
2912
2913lemma (in comp_fun_commute) fold_set_fold_remdups:
2914  "Finite_Set.fold f y (set xs) = fold f (remdups xs) y"
2915by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm insert_absorb)
2916
2917lemma (in comp_fun_idem) fold_set_fold:
2918  "Finite_Set.fold f y (set xs) = fold f xs y"
2919by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_left_comm)
2920
2921lemma union_set_fold [code]: "set xs \<union> A = fold Set.insert xs A"
2922proof -
2923  interpret comp_fun_idem Set.insert
2924    by (fact comp_fun_idem_insert)
2925  show ?thesis by (simp add: union_fold_insert fold_set_fold)
2926qed
2927
2928lemma union_coset_filter [code]:
2929  "List.coset xs \<union> A = List.coset (List.filter (\<lambda>x. x \<notin> A) xs)"
2930by auto
2931
2932lemma minus_set_fold [code]: "A - set xs = fold Set.remove xs A"
2933proof -
2934  interpret comp_fun_idem Set.remove
2935    by (fact comp_fun_idem_remove)
2936  show ?thesis
2937    by (simp add: minus_fold_remove [of _ A] fold_set_fold)
2938qed
2939
2940lemma minus_coset_filter [code]:
2941  "A - List.coset xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
2942by auto
2943
2944lemma inter_set_filter [code]:
2945  "A \<inter> set xs = set (List.filter (\<lambda>x. x \<in> A) xs)"
2946by auto
2947
2948lemma inter_coset_fold [code]:
2949  "A \<inter> List.coset xs = fold Set.remove xs A"
2950by (simp add: Diff_eq [symmetric] minus_set_fold)
2951
2952lemma (in semilattice_set) set_eq_fold [code]:
2953  "F (set (x # xs)) = fold f xs x"
2954proof -
2955  interpret comp_fun_idem f
2956    by standard (simp_all add: fun_eq_iff left_commute)
2957  show ?thesis by (simp add: eq_fold fold_set_fold)
2958qed
2959
2960lemma (in complete_lattice) Inf_set_fold:
2961  "Inf (set xs) = fold inf xs top"
2962proof -
2963  interpret comp_fun_idem "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
2964    by (fact comp_fun_idem_inf)
2965  show ?thesis by (simp add: Inf_fold_inf fold_set_fold inf_commute)
2966qed
2967
2968declare Inf_set_fold [where 'a = "'a set", code]
2969
2970lemma (in complete_lattice) Sup_set_fold:
2971  "Sup (set xs) = fold sup xs bot"
2972proof -
2973  interpret comp_fun_idem "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a"
2974    by (fact comp_fun_idem_sup)
2975  show ?thesis by (simp add: Sup_fold_sup fold_set_fold sup_commute)
2976qed
2977
2978declare Sup_set_fold [where 'a = "'a set", code]
2979
2980lemma (in complete_lattice) INF_set_fold:
2981  "INFIMUM (set xs) f = fold (inf \<circ> f) xs top"
2982  using Inf_set_fold [of "map f xs "] by (simp add: fold_map)
2983
2984declare INF_set_fold [code]
2985
2986lemma (in complete_lattice) SUP_set_fold:
2987  "SUPREMUM (set xs) f = fold (sup \<circ> f) xs bot"
2988  using Sup_set_fold [of "map f xs "] by (simp add: fold_map)
2989
2990declare SUP_set_fold [code]
2991
2992
2993subsubsection \<open>Fold variants: @{const foldr} and @{const foldl}\<close>
2994
2995text \<open>Correspondence\<close>
2996
2997lemma foldr_conv_fold [code_abbrev]: "foldr f xs = fold f (rev xs)"
2998by (induct xs) simp_all
2999
3000lemma foldl_conv_fold: "foldl f s xs = fold (\<lambda>x s. f s x) xs s"
3001by (induct xs arbitrary: s) simp_all
3002
3003lemma foldr_conv_foldl: \<comment> \<open>The ``Third Duality Theorem'' in Bird \& Wadler:\<close>
3004  "foldr f xs a = foldl (\<lambda>x y. f y x) a (rev xs)"
3005by (simp add: foldr_conv_fold foldl_conv_fold)
3006
3007lemma foldl_conv_foldr:
3008  "foldl f a xs = foldr (\<lambda>x y. f y x) (rev xs) a"
3009by (simp add: foldr_conv_fold foldl_conv_fold)
3010
3011lemma foldr_fold:
3012  "(\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y)
3013  \<Longrightarrow> foldr f xs = fold f xs"
3014unfolding foldr_conv_fold by (rule fold_rev)
3015
3016lemma foldr_cong [fundef_cong]:
3017  "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"
3018by (auto simp add: foldr_conv_fold intro!: fold_cong)
3019
3020lemma foldl_cong [fundef_cong]:
3021  "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"
3022by (auto simp add: foldl_conv_fold intro!: fold_cong)
3023
3024lemma foldr_append [simp]: "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)"
3025by (simp add: foldr_conv_fold)
3026
3027lemma foldl_append [simp]: "foldl f a (xs @ ys) = foldl f (foldl f a xs) ys"
3028by (simp add: foldl_conv_fold)
3029
3030lemma foldr_map [code_unfold]: "foldr g (map f xs) a = foldr (g \<circ> f) xs a"
3031by (simp add: foldr_conv_fold fold_map rev_map)
3032
3033lemma foldr_filter:
3034  "foldr f (filter P xs) = foldr (\<lambda>x. if P x then f x else id) xs"
3035by (simp add: foldr_conv_fold rev_filter fold_filter)
3036
3037lemma foldl_map [code_unfold]:
3038  "foldl g a (map f xs) = foldl (\<lambda>a x. g a (f x)) a xs"
3039by (simp add: foldl_conv_fold fold_map comp_def)
3040
3041lemma concat_conv_foldr [code]:
3042  "concat xss = foldr append xss []"
3043by (simp add: fold_append_concat_rev foldr_conv_fold)
3044
3045
3046subsubsection \<open>@{const upt}\<close>
3047
3048lemma upt_rec[code]: "[i..<j] = (if i<j then i#[Suc i..<j] else [])"
3049\<comment> \<open>simp does not terminate!\<close>
3050by (induct j) auto
3051
3052lemmas upt_rec_numeral[simp] = upt_rec[of "numeral m" "numeral n"] for m n
3053
3054lemma upt_conv_Nil [simp]: "j <= i ==> [i..<j] = []"
3055by (subst upt_rec) simp
3056
3057lemma upt_eq_Nil_conv[simp]: "([i..<j] = []) = (j = 0 \<or> j <= i)"
3058by(induct j)simp_all
3059
3060lemma upt_eq_Cons_conv:
3061 "([i..<j] = x#xs) = (i < j \<and> i = x \<and> [i+1..<j] = xs)"
3062apply(induct j arbitrary: x xs)
3063 apply simp
3064apply(clarsimp simp add: append_eq_Cons_conv)
3065apply arith
3066done
3067
3068lemma upt_Suc_append: "i <= j ==> [i..<(Suc j)] = [i..<j]@[j]"
3069\<comment> \<open>Only needed if \<open>upt_Suc\<close> is deleted from the simpset.\<close>
3070by simp
3071
3072lemma upt_conv_Cons: "i < j ==> [i..<j] = i # [Suc i..<j]"
3073by (simp add: upt_rec)
3074
3075lemma upt_conv_Cons_Cons: \<comment> \<open>no precondition\<close>
3076  "m # n # ns = [m..<q] \<longleftrightarrow> n # ns = [Suc m..<q]"
3077proof (cases "m < q")
3078  case False then show ?thesis by simp
3079next
3080  case True then show ?thesis by (simp add: upt_conv_Cons)
3081qed
3082
3083lemma upt_add_eq_append: "i<=j ==> [i..<j+k] = [i..<j]@[j..<j+k]"
3084\<comment> \<open>LOOPS as a simprule, since \<open>j <= j\<close>.\<close>
3085by (induct k) auto
3086
3087lemma length_upt [simp]: "length [i..<j] = j - i"
3088by (induct j) (auto simp add: Suc_diff_le)
3089
3090lemma nth_upt [simp]: "i + k < j ==> [i..<j] ! k = i + k"
3091by (induct j) (auto simp add: less_Suc_eq nth_append split: nat_diff_split)
3092
3093lemma hd_upt[simp]: "i < j \<Longrightarrow> hd[i..<j] = i"
3094by(simp add:upt_conv_Cons)
3095
3096lemma tl_upt: "tl [m..<n] = [Suc m..<n]"
3097by (simp add: upt_rec)
3098
3099lemma last_upt[simp]: "i < j \<Longrightarrow> last[i..<j] = j - 1"
3100by(cases j)(auto simp: upt_Suc_append)
3101
3102lemma take_upt [simp]: "i+m <= n ==> take m [i..<n] = [i..<i+m]"
3103apply (induct m arbitrary: i, simp)
3104apply (subst upt_rec)
3105apply (rule sym)
3106apply (subst upt_rec)
3107apply (simp del: upt.simps)
3108done
3109
3110lemma drop_upt[simp]: "drop m [i..<j] = [i+m..<j]"
3111by(induct j) auto
3112
3113lemma map_Suc_upt: "map Suc [m..<n] = [Suc m..<Suc n]"
3114by (induct n) auto
3115
3116lemma map_add_upt: "map (\<lambda>i. i + n) [0..<m] = [n..<m + n]"
3117by (induct m) simp_all
3118
3119lemma nth_map_upt: "i < n-m ==> (map f [m..<n]) ! i = f(m+i)"
3120apply (induct n m  arbitrary: i rule: diff_induct)
3121  prefer 3 apply (subst map_Suc_upt[symmetric])
3122  apply (auto simp add: less_diff_conv)
3123done
3124
3125lemma map_decr_upt: "map (\<lambda>n. n - Suc 0) [Suc m..<Suc n] = [m..<n]"
3126by (induct n) simp_all
3127
3128lemma map_upt_Suc: "map f [0 ..< Suc n] = f 0 # map (\<lambda>i. f (Suc i)) [0 ..< n]"
3129by (induct n arbitrary: f) auto
3130
3131
3132lemma nth_take_lemma:
3133  "k \<le> length xs \<Longrightarrow> k \<le> length ys \<Longrightarrow>
3134     (\<And>i. i < k \<longrightarrow> xs!i = ys!i) \<Longrightarrow> take k xs = take k ys"
3135apply (atomize, induct k arbitrary: xs ys)
3136apply (simp_all add: less_Suc_eq_0_disj all_conj_distrib, clarify)
3137txt \<open>Both lists must be non-empty\<close>
3138apply (case_tac xs, simp)
3139apply (case_tac ys, clarify)
3140 apply (simp (no_asm_use))
3141apply clarify
3142txt \<open>prenexing's needed, not miniscoping\<close>
3143apply (simp (no_asm_use) add: all_simps [symmetric] del: all_simps)
3144apply blast
3145done
3146
3147lemma nth_equalityI:
3148  "[| length xs = length ys; \<forall>i < length xs. xs!i = ys!i |] ==> xs = ys"
3149by (frule nth_take_lemma [OF le_refl eq_imp_le]) simp_all
3150
3151lemma map_nth:
3152  "map (\<lambda>i. xs ! i) [0..<length xs] = xs"
3153by (rule nth_equalityI, auto)
3154
3155lemma list_all2_antisym:
3156  "\<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>
3157  \<Longrightarrow> xs = ys"
3158apply (simp add: list_all2_conv_all_nth)
3159apply (rule nth_equalityI, blast, simp)
3160done
3161
3162lemma take_equalityI: "(\<forall>i. take i xs = take i ys) ==> xs = ys"
3163\<comment> \<open>The famous take-lemma.\<close>
3164apply (drule_tac x = "max (length xs) (length ys)" in spec)
3165apply (simp add: le_max_iff_disj)
3166done
3167
3168
3169lemma take_Cons':
3170  "take n (x # xs) = (if n = 0 then [] else x # take (n - 1) xs)"
3171by (cases n) simp_all
3172
3173lemma drop_Cons':
3174  "drop n (x # xs) = (if n = 0 then x # xs else drop (n - 1) xs)"
3175by (cases n) simp_all
3176
3177lemma nth_Cons': "(x # xs)!n = (if n = 0 then x else xs!(n - 1))"
3178by (cases n) simp_all
3179
3180lemma take_Cons_numeral [simp]:
3181  "take (numeral v) (x # xs) = x # take (numeral v - 1) xs"
3182by (simp add: take_Cons')
3183
3184lemma drop_Cons_numeral [simp]:
3185  "drop (numeral v) (x # xs) = drop (numeral v - 1) xs"
3186by (simp add: drop_Cons')
3187
3188lemma nth_Cons_numeral [simp]:
3189  "(x # xs) ! numeral v = xs ! (numeral v - 1)"
3190by (simp add: nth_Cons')
3191
3192
3193subsubsection \<open>\<open>upto\<close>: interval-list on @{typ int}\<close>
3194
3195function upto :: "int \<Rightarrow> int \<Rightarrow> int list" ("(1[_../_])") where
3196  "upto i j = (if i \<le> j then i # [i+1..j] else [])"
3197by auto
3198termination
3199by(relation "measure(%(i::int,j). nat(j - i + 1))") auto
3200
3201declare upto.simps[simp del]
3202
3203lemmas upto_rec_numeral [simp] =
3204  upto.simps[of "numeral m" "numeral n"]
3205  upto.simps[of "numeral m" "- numeral n"]
3206  upto.simps[of "- numeral m" "numeral n"]
3207  upto.simps[of "- numeral m" "- numeral n"] for m n
3208
3209lemma upto_empty[simp]: "j < i \<Longrightarrow> [i..j] = []"
3210by(simp add: upto.simps)
3211
3212lemma upto_single[simp]: "[i..i] = [i]"
3213by(simp add: upto.simps)
3214
3215lemma upto_Nil[simp]: "[i..j] = [] \<longleftrightarrow> j < i"
3216by (simp add: upto.simps)
3217
3218lemma upto_Nil2[simp]: "[] = [i..j] \<longleftrightarrow> j < i"
3219by (simp add: upto.simps)
3220
3221lemma upto_rec1: "i \<le> j \<Longrightarrow> [i..j] = i#[i+1..j]"
3222by(simp add: upto.simps)
3223
3224lemma upto_rec2: "i \<le> j \<Longrightarrow> [i..j] = [i..j - 1]@[j]"
3225proof(induct "nat(j-i)" arbitrary: i j)
3226  case 0 thus ?case by(simp add: upto.simps)
3227next
3228  case (Suc n)
3229  hence "n = nat (j - (i + 1))" "i < j" by linarith+
3230  from this(2) Suc.hyps(1)[OF this(1)] Suc(2,3) upto_rec1 show ?case by simp
3231qed
3232
3233lemma length_upto[simp]: "length [i..j] = nat(j - i + 1)"
3234by(induction i j rule: upto.induct) (auto simp: upto.simps)
3235
3236lemma set_upto[simp]: "set[i..j] = {i..j}"
3237proof(induct i j rule:upto.induct)
3238  case (1 i j)
3239  from this show ?case
3240    unfolding upto.simps[of i j] by auto
3241qed
3242
3243lemma nth_upto: "i + int k \<le> j \<Longrightarrow> [i..j] ! k = i + int k"
3244apply(induction i j arbitrary: k rule: upto.induct)
3245apply(subst upto_rec1)
3246apply(auto simp add: nth_Cons')
3247done
3248
3249lemma upto_split1: 
3250  "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> [i..k] = [i..j-1] @ [j..k]"
3251proof (induction j rule: int_ge_induct)
3252  case base thus ?case by (simp add: upto_rec1)
3253next
3254  case step thus ?case using upto_rec1 upto_rec2 by simp
3255qed
3256
3257lemma upto_split2: 
3258  "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> [i..k] = [i..j] @ [j+1..k]"
3259using upto_rec1 upto_rec2 upto_split1 by auto
3260
3261lemma upto_split3: "\<lbrakk> i \<le> j; j \<le> k \<rbrakk> \<Longrightarrow> [i..k] = [i..j-1] @ j # [j+1..k]"
3262using upto_rec1 upto_split1 by auto
3263
3264text\<open>Tail recursive version for code generation:\<close>
3265
3266definition upto_aux :: "int \<Rightarrow> int \<Rightarrow> int list \<Rightarrow> int list" where
3267  "upto_aux i j js = [i..j] @ js"
3268
3269lemma upto_aux_rec [code]:
3270  "upto_aux i j js = (if j<i then js else upto_aux i (j - 1) (j#js))"
3271by (simp add: upto_aux_def upto_rec2)
3272
3273lemma upto_code[code]: "[i..j] = upto_aux i j []"
3274by(simp add: upto_aux_def)
3275
3276
3277subsubsection \<open>@{const distinct} and @{const remdups} and @{const remdups_adj}\<close>
3278
3279lemma distinct_tl: "distinct xs \<Longrightarrow> distinct (tl xs)"
3280by (cases xs) simp_all
3281
3282lemma distinct_append [simp]:
3283  "distinct (xs @ ys) = (distinct xs \<and> distinct ys \<and> set xs \<inter> set ys = {})"
3284by (induct xs) auto
3285
3286lemma distinct_rev[simp]: "distinct(rev xs) = distinct xs"
3287by(induct xs) auto
3288
3289lemma set_remdups [simp]: "set (remdups xs) = set xs"
3290by (induct xs) (auto simp add: insert_absorb)
3291
3292lemma distinct_remdups [iff]: "distinct (remdups xs)"
3293by (induct xs) auto
3294
3295lemma distinct_remdups_id: "distinct xs ==> remdups xs = xs"
3296by (induct xs, auto)
3297
3298lemma remdups_id_iff_distinct [simp]: "remdups xs = xs \<longleftrightarrow> distinct xs"
3299by (metis distinct_remdups distinct_remdups_id)
3300
3301lemma finite_distinct_list: "finite A \<Longrightarrow> \<exists>xs. set xs = A \<and> distinct xs"
3302by (metis distinct_remdups finite_list set_remdups)
3303
3304lemma remdups_eq_nil_iff [simp]: "(remdups x = []) = (x = [])"
3305by (induct x, auto)
3306
3307lemma remdups_eq_nil_right_iff [simp]: "([] = remdups x) = (x = [])"
3308by (induct x, auto)
3309
3310lemma length_remdups_leq[iff]: "length(remdups xs) <= length xs"
3311by (induct xs) auto
3312
3313lemma length_remdups_eq[iff]:
3314  "(length (remdups xs) = length xs) = (remdups xs = xs)"
3315apply(induct xs)
3316 apply auto
3317apply(subgoal_tac "length (remdups xs) <= length xs")
3318 apply arith
3319apply(rule length_remdups_leq)
3320done
3321
3322lemma remdups_filter: "remdups(filter P xs) = filter P (remdups xs)"
3323by (induct xs) auto
3324
3325lemma distinct_map:
3326  "distinct(map f xs) = (distinct xs \<and> inj_on f (set xs))"
3327by (induct xs) auto
3328
3329lemma distinct_map_filter:
3330  "distinct (map f xs) \<Longrightarrow> distinct (map f (filter P xs))"
3331by (induct xs) auto
3332
3333lemma distinct_filter [simp]: "distinct xs ==> distinct (filter P xs)"
3334by (induct xs) auto
3335
3336lemma distinct_upt[simp]: "distinct[i..<j]"
3337by (induct j) auto
3338
3339lemma distinct_upto[simp]: "distinct[i..j]"
3340apply(induct i j rule:upto.induct)
3341apply(subst upto.simps)
3342apply(simp)
3343done
3344
3345lemma distinct_take[simp]: "distinct xs \<Longrightarrow> distinct (take i xs)"
3346apply(induct xs arbitrary: i)
3347 apply simp
3348apply (case_tac i)
3349 apply simp_all
3350apply(blast dest:in_set_takeD)
3351done
3352
3353lemma distinct_drop[simp]: "distinct xs \<Longrightarrow> distinct (drop i xs)"
3354apply(induct xs arbitrary: i)
3355 apply simp
3356apply (case_tac i)
3357 apply simp_all
3358done
3359
3360lemma distinct_list_update:
3361assumes d: "distinct xs" and a: "a \<notin> set xs - {xs!i}"
3362shows "distinct (xs[i:=a])"
3363proof (cases "i < length xs")
3364  case True
3365  with a have "a \<notin> set (take i xs @ xs ! i # drop (Suc i) xs) - {xs!i}"
3366    apply (drule_tac id_take_nth_drop) by simp
3367  with d True show ?thesis
3368    apply (simp add: upd_conv_take_nth_drop)
3369    apply (drule subst [OF id_take_nth_drop]) apply assumption
3370    apply simp apply (cases "a = xs!i") apply simp by blast
3371next
3372  case False with d show ?thesis by auto
3373qed
3374
3375lemma distinct_concat:
3376  "\<lbrakk> distinct xs;
3377     \<And> ys. ys \<in> set xs \<Longrightarrow> distinct ys;
3378     \<And> ys zs. \<lbrakk> ys \<in> set xs ; zs \<in> set xs ; ys \<noteq> zs \<rbrakk> \<Longrightarrow> set ys \<inter> set zs = {}
3379   \<rbrakk> \<Longrightarrow> distinct (concat xs)"
3380by (induct xs) auto
3381
3382text \<open>It is best to avoid this indexed version of distinct, but
3383sometimes it is useful.\<close>
3384
3385lemma distinct_conv_nth:
3386"distinct xs = (\<forall>i < size xs. \<forall>j < size xs. i \<noteq> j \<longrightarrow> xs!i \<noteq> xs!j)"
3387apply (induct xs, simp, simp)
3388apply (rule iffI, clarsimp)
3389 apply (case_tac i)
3390apply (case_tac j, simp)
3391apply (simp add: set_conv_nth)
3392 apply (case_tac j)
3393apply (clarsimp simp add: set_conv_nth, simp)
3394apply (rule conjI)
3395 apply (clarsimp simp add: set_conv_nth)
3396 apply (erule_tac x = 0 in allE, simp)
3397 apply (erule_tac x = "Suc i" in allE, simp, clarsimp)
3398apply (erule_tac x = "Suc i" in allE, simp)
3399apply (erule_tac x = "Suc j" in allE, simp)
3400done
3401
3402lemma nth_eq_iff_index_eq:
3403  "\<lbrakk> distinct xs; i < length xs; j < length xs \<rbrakk> \<Longrightarrow> (xs!i = xs!j) = (i = j)"
3404by(auto simp: distinct_conv_nth)
3405
3406lemma distinct_Ex1:
3407  "distinct xs \<Longrightarrow> x \<in> set xs \<Longrightarrow> (\<exists>!i. i < length xs \<and> xs ! i = x)"
3408  by (auto simp: in_set_conv_nth nth_eq_iff_index_eq)
3409
3410lemma inj_on_nth: "distinct xs \<Longrightarrow> \<forall>i \<in> I. i < length xs \<Longrightarrow> inj_on (nth xs) I"
3411by (rule inj_onI) (simp add: nth_eq_iff_index_eq)
3412
3413lemma bij_betw_nth:
3414  assumes "distinct xs" "A = {..<length xs}" "B = set xs"
3415  shows   "bij_betw ((!) xs) A B"
3416  using assms unfolding bij_betw_def
3417  by (auto intro!: inj_on_nth simp: set_conv_nth)
3418
3419lemma set_update_distinct: "\<lbrakk> distinct xs;  n < length xs \<rbrakk> \<Longrightarrow>
3420  set(xs[n := x]) = insert x (set xs - {xs!n})"
3421by(auto simp: set_eq_iff in_set_conv_nth nth_list_update nth_eq_iff_index_eq)
3422
3423lemma distinct_swap[simp]: "\<lbrakk> i < size xs; j < size xs \<rbrakk> \<Longrightarrow>
3424  distinct(xs[i := xs!j, j := xs!i]) = distinct xs"
3425apply (simp add: distinct_conv_nth nth_list_update)
3426apply safe
3427apply metis+
3428done
3429
3430lemma set_swap[simp]:
3431  "\<lbrakk> i < size xs; j < size xs \<rbrakk> \<Longrightarrow> set(xs[i := xs!j, j := xs!i]) = set xs"
3432by(simp add: set_conv_nth nth_list_update) metis
3433
3434lemma distinct_card: "distinct xs ==> card (set xs) = size xs"
3435by (induct xs) auto
3436
3437lemma card_distinct: "card (set xs) = size xs ==> distinct xs"
3438proof (induct xs)
3439  case Nil thus ?case by simp
3440next
3441  case (Cons x xs)
3442  show ?case
3443  proof (cases "x \<in> set xs")
3444    case False with Cons show ?thesis by simp
3445  next
3446    case True with Cons.prems
3447    have "card (set xs) = Suc (length xs)"
3448      by (simp add: card_insert_if split: if_split_asm)
3449    moreover have "card (set xs) \<le> length xs" by (rule card_length)
3450    ultimately have False by simp
3451    thus ?thesis ..
3452  qed
3453qed
3454
3455lemma distinct_length_filter: "distinct xs \<Longrightarrow> length (filter P xs) = card ({x. P x} Int set xs)"
3456by (induct xs) (auto)
3457
3458lemma not_distinct_decomp: "\<not> distinct ws \<Longrightarrow> \<exists>xs ys zs y. ws = xs@[y]@ys@[y]@zs"
3459apply (induct n == "length ws" arbitrary:ws) apply simp
3460apply(case_tac ws) apply simp
3461apply (simp split:if_split_asm)
3462apply (metis Cons_eq_appendI eq_Nil_appendI split_list)
3463done
3464
3465lemma not_distinct_conv_prefix:
3466  defines "dec as xs y ys \<equiv> y \<in> set xs \<and> distinct xs \<and> as = xs @ y # ys"
3467  shows "\<not>distinct as \<longleftrightarrow> (\<exists>xs y ys. dec as xs y ys)" (is "?L = ?R")
3468proof
3469  assume "?L" then show "?R"
3470  proof (induct "length as" arbitrary: as rule: less_induct)
3471    case less
3472    obtain xs ys zs y where decomp: "as = (xs @ y # ys) @ y # zs"
3473      using not_distinct_decomp[OF less.prems] by auto
3474    show ?case
3475    proof (cases "distinct (xs @ y # ys)")
3476      case True
3477      with decomp have "dec as (xs @ y # ys) y zs" by (simp add: dec_def)
3478      then show ?thesis by blast
3479    next
3480      case False
3481      with less decomp obtain xs' y' ys' where "dec (xs @ y # ys) xs' y' ys'"
3482        by atomize_elim auto
3483      with decomp have "dec as xs' y' (ys' @ y # zs)" by (simp add: dec_def)
3484      then show ?thesis by blast
3485    qed
3486  qed
3487qed (auto simp: dec_def)
3488
3489lemma distinct_product:
3490  "distinct xs \<Longrightarrow> distinct ys \<Longrightarrow> distinct (List.product xs ys)"
3491by (induct xs) (auto intro: inj_onI simp add: distinct_map)
3492
3493lemma distinct_product_lists:
3494  assumes "\<forall>xs \<in> set xss. distinct xs"
3495  shows "distinct (product_lists xss)"
3496using assms proof (induction xss)
3497  case (Cons xs xss) note * = this
3498  then show ?case
3499  proof (cases "product_lists xss")
3500    case Nil then show ?thesis by (induct xs) simp_all
3501  next
3502    case (Cons ps pss) with * show ?thesis
3503      by (auto intro!: inj_onI distinct_concat simp add: distinct_map)
3504  qed
3505qed simp
3506
3507lemma length_remdups_concat:
3508  "length (remdups (concat xss)) = card (\<Union>xs\<in>set xss. set xs)"
3509by (simp add: distinct_card [symmetric])
3510
3511lemma length_remdups_card_conv: "length(remdups xs) = card(set xs)"
3512proof -
3513  have xs: "concat[xs] = xs" by simp
3514  from length_remdups_concat[of "[xs]"] show ?thesis unfolding xs by simp
3515qed
3516
3517lemma remdups_remdups: "remdups (remdups xs) = remdups xs"
3518by (induct xs) simp_all
3519
3520lemma distinct_butlast:
3521  assumes "distinct xs"
3522  shows "distinct (butlast xs)"
3523proof (cases "xs = []")
3524  case False
3525    from \<open>xs \<noteq> []\<close> obtain ys y where "xs = ys @ [y]" by (cases xs rule: rev_cases) auto
3526    with \<open>distinct xs\<close> show ?thesis by simp
3527qed (auto)
3528
3529lemma remdups_map_remdups:
3530  "remdups (map f (remdups xs)) = remdups (map f xs)"
3531by (induct xs) simp_all
3532
3533lemma distinct_zipI1:
3534  assumes "distinct xs"
3535  shows "distinct (zip xs ys)"
3536proof (rule zip_obtain_same_length)
3537  fix xs' :: "'a list" and ys' :: "'b list" and n
3538  assume "length xs' = length ys'"
3539  assume "xs' = take n xs"
3540  with assms have "distinct xs'" by simp
3541  with \<open>length xs' = length ys'\<close> show "distinct (zip xs' ys')"
3542    by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE)
3543qed
3544
3545lemma distinct_zipI2:
3546  assumes "distinct ys"
3547  shows "distinct (zip xs ys)"
3548proof (rule zip_obtain_same_length)
3549  fix xs' :: "'b list" and ys' :: "'a list" and n
3550  assume "length xs' = length ys'"
3551  assume "ys' = take n ys"
3552  with assms have "distinct ys'" by simp
3553  with \<open>length xs' = length ys'\<close> show "distinct (zip xs' ys')"
3554    by (induct xs' ys' rule: list_induct2) (auto elim: in_set_zipE)
3555qed
3556
3557lemma set_take_disj_set_drop_if_distinct:
3558  "distinct vs \<Longrightarrow> i \<le> j \<Longrightarrow> set (take i vs) \<inter> set (drop j vs) = {}"
3559by (auto simp: in_set_conv_nth distinct_conv_nth)
3560
3561(* The next two lemmas help Sledgehammer. *)
3562
3563lemma distinct_singleton: "distinct [x]" by simp
3564
3565lemma distinct_length_2_or_more:
3566  "distinct (a # b # xs) \<longleftrightarrow> (a \<noteq> b \<and> distinct (a # xs) \<and> distinct (b # xs))"
3567by force
3568
3569lemma remdups_adj_altdef: "(remdups_adj xs = ys) \<longleftrightarrow>
3570  (\<exists>f::nat => nat. mono f \<and> f ` {0 ..< size xs} = {0 ..< size ys}
3571    \<and> (\<forall>i < size xs. xs!i = ys!(f i))
3572    \<and> (\<forall>i. i + 1 < size xs \<longrightarrow> (xs!i = xs!(i+1) \<longleftrightarrow> f i = f(i+1))))" (is "?L \<longleftrightarrow> (\<exists>f. ?p f xs ys)")
3573proof
3574  assume ?L
3575  then show "\<exists>f. ?p f xs ys"
3576  proof (induct xs arbitrary: ys rule: remdups_adj.induct)
3577    case (1 ys)
3578    thus ?case by (intro exI[of _ id]) (auto simp: mono_def)
3579  next
3580    case (2 x ys)
3581    thus ?case by (intro exI[of _ id]) (auto simp: mono_def)
3582  next
3583    case (3 x1 x2 xs ys)
3584    let ?xs = "x1 # x2 # xs"
3585    let ?cond = "x1 = x2"
3586    define zs where "zs = remdups_adj (x2 # xs)"
3587    from 3(1-2)[of zs]
3588    obtain f where p: "?p f (x2 # xs) zs" unfolding zs_def by (cases ?cond) auto
3589    then have f0: "f 0 = 0"
3590      by (intro mono_image_least[where f=f]) blast+
3591    from p have mono: "mono f" and f_xs_zs: "f ` {0..<length (x2 # xs)} = {0..<length zs}" by auto
3592    have ys: "ys = (if x1 = x2 then zs else x1 # zs)"
3593      unfolding 3(3)[symmetric] zs_def by auto
3594    have zs0: "zs ! 0 = x2" unfolding zs_def by (induct xs) auto
3595    have zsne: "zs \<noteq> []" unfolding zs_def by (induct xs) auto
3596    let ?Succ = "if ?cond then id else Suc"
3597    let ?x1 = "if ?cond then id else Cons x1"
3598    let ?f = "\<lambda> i. if i = 0 then 0 else ?Succ (f (i - 1))"
3599    have ys: "ys = ?x1 zs" unfolding ys by (cases ?cond, auto)
3600    have mono: "mono ?f" using \<open>mono f\<close> unfolding mono_def by auto
3601    show ?case unfolding ys
3602    proof (intro exI[of _ ?f] conjI allI impI)
3603      show "mono ?f" by fact
3604    next
3605      fix i assume i: "i < length ?xs"
3606      with p show "?xs ! i = ?x1 zs ! (?f i)" using zs0 by auto
3607    next
3608      fix i assume i: "i + 1 < length ?xs"
3609      with p show "(?xs ! i = ?xs ! (i + 1)) = (?f i = ?f (i + 1))"
3610        by (cases i) (auto simp: f0)
3611    next
3612      have id: "{0 ..< length (?x1 zs)} = insert 0 (?Succ ` {0 ..< length zs})"
3613        using zsne by (cases ?cond, auto)
3614      { fix i  assume "i < Suc (length xs)"
3615        hence "Suc i \<in> {0..<Suc (Suc (length xs))} \<inter> Collect ((<) 0)" by auto
3616        from imageI[OF this, of "\<lambda>i. ?Succ (f (i - Suc 0))"]
3617        have "?Succ (f i) \<in> (\<lambda>i. ?Succ (f (i - Suc 0))) ` ({0..<Suc (Suc (length xs))} \<inter> Collect ((<) 0))" by auto
3618      }
3619      then show "?f ` {0 ..< length ?xs} = {0 ..< length (?x1  zs)}"
3620        unfolding id f_xs_zs[symmetric] by auto
3621    qed
3622  qed
3623next
3624  assume "\<exists> f. ?p f xs ys"
3625  then show ?L
3626  proof (induct xs arbitrary: ys rule: remdups_adj.induct)
3627    case 1 then show ?case by auto
3628  next
3629    case (2 x) then obtain f where f_img: "f ` {0 ..< size [x]} = {0 ..< size ys}"
3630        and f_nth: "\<And>i. i < size [x] \<Longrightarrow> [x]!i = ys!(f i)"
3631      by blast
3632
3633    have "length ys = card (f ` {0 ..< size [x]})"
3634      using f_img by auto
3635    then have *: "length ys = 1" by auto
3636    then have "f 0 = 0" using f_img by auto
3637    with * show ?case using f_nth by (cases ys) auto
3638  next
3639    case (3 x1 x2 xs)
3640    from "3.prems" obtain f where f_mono: "mono f"
3641      and f_img: "f ` {0..<length (x1 # x2 # xs)} = {0..<length ys}"
3642      and f_nth:
3643        "\<And>i. i < length (x1 # x2 # xs) \<Longrightarrow> (x1 # x2 # xs) ! i = ys ! f i"
3644        "\<And>i. i + 1 < length (x1 # x2 #xs) \<Longrightarrow>
3645          ((x1 # x2 # xs) ! i = (x1 # x2 # xs) ! (i + 1)) = (f i = f (i + 1))"
3646      by blast
3647
3648    show ?case
3649    proof cases
3650      assume "x1 = x2"
3651
3652      let ?f' = "f \<circ> Suc"
3653
3654      have "remdups_adj (x1 # xs) = ys"
3655      proof (intro "3.hyps" exI conjI impI allI)
3656        show "mono ?f'"
3657          using f_mono by (simp add: mono_iff_le_Suc)
3658      next
3659        have "?f' ` {0 ..< length (x1 # xs)} = f ` {Suc 0 ..< length (x1 # x2 # xs)}"
3660          by safe (fastforce, rename_tac x, case_tac x, auto)
3661        also have "\<dots> = f ` {0 ..< length (x1 # x2 # xs)}"
3662        proof -
3663          have "f 0 = f (Suc 0)" using \<open>x1 = x2\<close> f_nth[of 0] by simp
3664          then show ?thesis by safe (fastforce, rename_tac x, case_tac x, auto)
3665        qed
3666        also have "\<dots> = {0 ..< length ys}" by fact
3667        finally show  "?f' ` {0 ..< length (x1 # xs)} = {0 ..< length ys}" .
3668      qed (insert f_nth[of "Suc i" for i], auto simp: \<open>x1 = x2\<close>)
3669      then show ?thesis using \<open>x1 = x2\<close> by simp
3670    next
3671      assume "x1 \<noteq> x2"
3672
3673      have "2 \<le> length ys"
3674      proof -
3675        have "2 = card {f 0, f 1}" using \<open>x1 \<noteq> x2\<close> f_nth[of 0] by auto
3676        also have "\<dots> \<le> card (f ` {0..< length (x1 # x2 # xs)})"
3677          by (rule card_mono) auto
3678        finally show ?thesis using f_img by simp
3679      qed
3680
3681      have "f 0 = 0" using f_mono f_img by (rule mono_image_least) simp
3682
3683      have "f (Suc 0) = Suc 0"
3684      proof (rule ccontr)
3685        assume "f (Suc 0) \<noteq> Suc 0"
3686        then have "Suc 0 < f (Suc 0)" using f_nth[of 0] \<open>x1 \<noteq> x2\<close> \<open>f 0 = 0\<close> by auto
3687        then have "\<And>i. Suc 0 < f (Suc i)"
3688          using f_mono
3689          by (meson Suc_le_mono le0 less_le_trans monoD)
3690        then have "\<And>i. Suc 0 \<noteq> f i" using \<open>f 0 = 0\<close>
3691          by (case_tac i) fastforce+
3692        then have "Suc 0 \<notin> f ` {0 ..< length (x1 # x2 # xs)}" by auto
3693        then show False using f_img \<open>2 \<le> length ys\<close> by auto
3694      qed
3695
3696      obtain ys' where "ys = x1 # x2 # ys'"
3697        using \<open>2 \<le> length ys\<close> f_nth[of 0] f_nth[of 1]
3698        apply (cases ys)
3699        apply (rename_tac [2] ys', case_tac [2] ys')
3700        by (auto simp: \<open>f 0 = 0\<close> \<open>f (Suc 0) = Suc 0\<close>)
3701
3702      define f' where "f' x = f (Suc x) - 1" for x
3703
3704      { fix i
3705        have "Suc 0 \<le> f (Suc 0)" using f_nth[of 0] \<open>x1 \<noteq> x2\<close> \<open>f 0 = 0\<close>  by auto
3706        also have "\<dots> \<le> f (Suc i)" using f_mono by (rule monoD) arith
3707        finally have "Suc 0 \<le> f (Suc i)" .
3708      } note Suc0_le_f_Suc = this
3709
3710      { fix i have "f (Suc i) = Suc (f' i)"
3711          using Suc0_le_f_Suc[of i] by (auto simp: f'_def)
3712      } note f_Suc = this
3713
3714      have "remdups_adj (x2 # xs) = (x2 # ys')"
3715      proof (intro "3.hyps" exI conjI impI allI)
3716        show "mono f'"
3717          using Suc0_le_f_Suc f_mono by (auto simp: f'_def mono_iff_le_Suc le_diff_iff)
3718      next
3719        have "f' ` {0 ..< length (x2 # xs)} = (\<lambda>x. f x - 1) ` {0 ..< length (x1 # x2 #xs)}"
3720          apply safe
3721          apply (rename_tac [!] n,  case_tac [!] n)
3722          apply (auto simp: f'_def \<open>f 0 = 0\<close> \<open>f (Suc 0) = Suc 0\<close> intro: rev_image_eqI)
3723          done
3724        also have "\<dots> = (\<lambda>x. x - 1) ` f ` {0 ..< length (x1 # x2 #xs)}"
3725          by (auto simp: image_comp)
3726        also have "\<dots> = (\<lambda>x. x - 1) ` {0 ..< length ys}"
3727          by (simp only: f_img)
3728        also have "\<dots> = {0 ..< length (x2 # ys')}"
3729          using \<open>ys = _\<close> by (fastforce intro: rev_image_eqI)
3730        finally show  "f' ` {0 ..< length (x2 # xs)} = {0 ..< length (x2 # ys')}" .
3731      qed (insert f_nth[of "Suc i" for i] \<open>x1 \<noteq> x2\<close>, auto simp add: f_Suc \<open>ys = _\<close>)
3732      then show ?case using \<open>ys = _\<close> \<open>x1 \<noteq> x2\<close> by simp
3733    qed
3734  qed
3735qed
3736
3737lemma hd_remdups_adj[simp]: "hd (remdups_adj xs) = hd xs"
3738by (induction xs rule: remdups_adj.induct) simp_all
3739
3740lemma remdups_adj_Cons: "remdups_adj (x # xs) =
3741  (case remdups_adj xs of [] \<Rightarrow> [x] | y # xs \<Rightarrow> if x = y then y # xs else x # y # xs)"
3742by (induct xs arbitrary: x) (auto split: list.splits)
3743
3744lemma remdups_adj_append_two:
3745  "remdups_adj (xs @ [x,y]) = remdups_adj (xs @ [x]) @ (if x = y then [] else [y])"
3746by (induct xs rule: remdups_adj.induct, simp_all)
3747
3748lemma remdups_adj_adjacent:
3749  "Suc i < length (remdups_adj xs) \<Longrightarrow> remdups_adj xs ! i \<noteq> remdups_adj xs ! Suc i"
3750proof (induction xs arbitrary: i rule: remdups_adj.induct)
3751  case (3 x y xs i)
3752  thus ?case by (cases i, cases "x = y") (simp, auto simp: hd_conv_nth[symmetric])
3753qed simp_all
3754
3755lemma remdups_adj_rev[simp]: "remdups_adj (rev xs) = rev (remdups_adj xs)"
3756by (induct xs rule: remdups_adj.induct, simp_all add: remdups_adj_append_two)
3757
3758lemma remdups_adj_length[simp]: "length (remdups_adj xs) \<le> length xs"
3759by (induct xs rule: remdups_adj.induct, auto)
3760
3761lemma remdups_adj_length_ge1[simp]: "xs \<noteq> [] \<Longrightarrow> length (remdups_adj xs) \<ge> Suc 0"
3762by (induct xs rule: remdups_adj.induct, simp_all)
3763
3764lemma remdups_adj_Nil_iff[simp]: "remdups_adj xs = [] \<longleftrightarrow> xs = []"
3765by (induct xs rule: remdups_adj.induct, simp_all)
3766
3767lemma remdups_adj_set[simp]: "set (remdups_adj xs) = set xs"
3768by (induct xs rule: remdups_adj.induct, simp_all)
3769
3770lemma remdups_adj_Cons_alt[simp]: "x # tl (remdups_adj (x # xs)) = remdups_adj (x # xs)"
3771by (induct xs rule: remdups_adj.induct, auto)
3772
3773lemma remdups_adj_distinct: "distinct xs \<Longrightarrow> remdups_adj xs = xs"
3774by (induct xs rule: remdups_adj.induct, simp_all)
3775
3776lemma remdups_adj_append:
3777  "remdups_adj (xs\<^sub>1 @ x # xs\<^sub>2) = remdups_adj (xs\<^sub>1 @ [x]) @ tl (remdups_adj (x # xs\<^sub>2))"
3778by (induct xs\<^sub>1 rule: remdups_adj.induct, simp_all)
3779
3780lemma remdups_adj_singleton:
3781  "remdups_adj xs = [x] \<Longrightarrow> xs = replicate (length xs) x"
3782by (induct xs rule: remdups_adj.induct, auto split: if_split_asm)
3783
3784lemma remdups_adj_map_injective:
3785  assumes "inj f"
3786  shows "remdups_adj (map f xs) = map f (remdups_adj xs)"
3787by (induct xs rule: remdups_adj.induct) (auto simp add: injD[OF assms])
3788
3789lemma remdups_adj_replicate:
3790  "remdups_adj (replicate n x) = (if n = 0 then [] else [x])"
3791  by (induction n) (auto simp: remdups_adj_Cons)
3792
3793lemma remdups_upt [simp]: "remdups [m..<n] = [m..<n]"
3794proof (cases "m \<le> n")
3795  case False then show ?thesis by simp
3796next
3797  case True then obtain q where "n = m + q"
3798    by (auto simp add: le_iff_add)
3799  moreover have "remdups [m..<m + q] = [m..<m + q]"
3800    by (induct q) simp_all
3801  ultimately show ?thesis by simp
3802qed
3803
3804
3805subsubsection \<open>@{const insert}\<close>
3806
3807lemma in_set_insert [simp]:
3808  "x \<in> set xs \<Longrightarrow> List.insert x xs = xs"
3809by (simp add: List.insert_def)
3810
3811lemma not_in_set_insert [simp]:
3812  "x \<notin> set xs \<Longrightarrow> List.insert x xs = x # xs"
3813by (simp add: List.insert_def)
3814
3815lemma insert_Nil [simp]: "List.insert x [] = [x]"
3816by simp
3817
3818lemma set_insert [simp]: "set (List.insert x xs) = insert x (set xs)"
3819by (auto simp add: List.insert_def)
3820
3821lemma distinct_insert [simp]: "distinct (List.insert x xs) = distinct xs"
3822by (simp add: List.insert_def)
3823
3824lemma insert_remdups:
3825  "List.insert x (remdups xs) = remdups (List.insert x xs)"
3826by (simp add: List.insert_def)
3827
3828
3829subsubsection \<open>@{const List.union}\<close>
3830
3831text\<open>This is all one should need to know about union:\<close>
3832lemma set_union[simp]: "set (List.union xs ys) = set xs \<union> set ys"
3833unfolding List.union_def
3834by(induct xs arbitrary: ys) simp_all
3835
3836lemma distinct_union[simp]: "distinct(List.union xs ys) = distinct ys"
3837unfolding List.union_def
3838by(induct xs arbitrary: ys) simp_all
3839
3840
3841subsubsection \<open>@{const List.find}\<close>
3842
3843lemma find_None_iff: "List.find P xs = None \<longleftrightarrow> \<not> (\<exists>x. x \<in> set xs \<and> P x)"
3844proof (induction xs)
3845  case Nil thus ?case by simp
3846next
3847  case (Cons x xs) thus ?case by (fastforce split: if_splits)
3848qed
3849
3850lemma find_Some_iff:
3851  "List.find P xs = Some x \<longleftrightarrow>
3852  (\<exists>i<length xs. P (xs!i) \<and> x = xs!i \<and> (\<forall>j<i. \<not> P (xs!j)))"
3853proof (induction xs)
3854  case Nil thus ?case by simp
3855next
3856  case (Cons x xs) thus ?case
3857    apply(auto simp: nth_Cons' split: if_splits)
3858    using diff_Suc_1[unfolded One_nat_def] less_Suc_eq_0_disj by fastforce
3859qed
3860
3861lemma find_cong[fundef_cong]:
3862  assumes "xs = ys" and "\<And>x. x \<in> set ys \<Longrightarrow> P x = Q x"
3863  shows "List.find P xs = List.find Q ys"
3864proof (cases "List.find P xs")
3865  case None thus ?thesis by (metis find_None_iff assms)
3866next
3867  case (Some x)
3868  hence "List.find Q ys = Some x" using assms
3869    by (auto simp add: find_Some_iff)
3870  thus ?thesis using Some by auto
3871qed
3872
3873lemma find_dropWhile:
3874  "List.find P xs = (case dropWhile (Not \<circ> P) xs
3875   of [] \<Rightarrow> None
3876    | x # _ \<Rightarrow> Some x)"
3877by (induct xs) simp_all
3878
3879
3880subsubsection \<open>@{const count_list}\<close>
3881
3882lemma count_notin[simp]: "x \<notin> set xs \<Longrightarrow> count_list xs x = 0"
3883by (induction xs) auto
3884
3885lemma count_le_length: "count_list xs x \<le> length xs"
3886by (induction xs) auto
3887
3888lemma sum_count_set:
3889  "set xs \<subseteq> X \<Longrightarrow> finite X \<Longrightarrow> sum (count_list xs) X = length xs"
3890apply(induction xs arbitrary: X)
3891 apply simp
3892apply (simp add: sum.If_cases)
3893by (metis Diff_eq sum.remove)
3894
3895
3896subsubsection \<open>@{const List.extract}\<close>
3897
3898lemma extract_None_iff: "List.extract P xs = None \<longleftrightarrow> \<not> (\<exists> x\<in>set xs. P x)"
3899by(auto simp: extract_def dropWhile_eq_Cons_conv split: list.splits)
3900  (metis in_set_conv_decomp)
3901
3902lemma extract_SomeE:
3903 "List.extract P xs = Some (ys, y, zs) \<Longrightarrow>
3904  xs = ys @ y # zs \<and> P y \<and> \<not> (\<exists> y \<in> set ys. P y)"
3905by(auto simp: extract_def dropWhile_eq_Cons_conv split: list.splits)
3906
3907lemma extract_Some_iff:
3908  "List.extract P xs = Some (ys, y, zs) \<longleftrightarrow>
3909   xs = ys @ y # zs \<and> P y \<and> \<not> (\<exists> y \<in> set ys. P y)"
3910by(auto simp: extract_def dropWhile_eq_Cons_conv dest: set_takeWhileD split: list.splits)
3911
3912lemma extract_Nil_code[code]: "List.extract P [] = None"
3913by(simp add: extract_def)
3914
3915lemma extract_Cons_code[code]:
3916  "List.extract P (x # xs) = (if P x then Some ([], x, xs) else
3917   (case List.extract P xs of
3918      None \<Rightarrow> None |
3919      Some (ys, y, zs) \<Rightarrow> Some (x#ys, y, zs)))"
3920by(auto simp add: extract_def comp_def split: list.splits)
3921  (metis dropWhile_eq_Nil_conv list.distinct(1))
3922
3923
3924subsubsection \<open>@{const remove1}\<close>
3925
3926lemma remove1_append:
3927  "remove1 x (xs @ ys) =
3928  (if x \<in> set xs then remove1 x xs @ ys else xs @ remove1 x ys)"
3929by (induct xs) auto
3930
3931lemma remove1_commute: "remove1 x (remove1 y zs) = remove1 y (remove1 x zs)"
3932by (induct zs) auto
3933
3934lemma in_set_remove1[simp]:
3935  "a \<noteq> b \<Longrightarrow> a \<in> set(remove1 b xs) = (a \<in> set xs)"
3936apply (induct xs)
3937 apply auto
3938done
3939
3940lemma set_remove1_subset: "set(remove1 x xs) \<subseteq> set xs"
3941apply(induct xs)
3942 apply simp
3943apply simp
3944apply blast
3945done
3946
3947lemma set_remove1_eq [simp]: "distinct xs \<Longrightarrow> set(remove1 x xs) = set xs - {x}"
3948apply(induct xs)
3949 apply simp
3950apply simp
3951apply blast
3952done
3953
3954lemma length_remove1:
3955  "length(remove1 x xs) = (if x \<in> set xs then length xs - 1 else length xs)"
3956by (induct xs) (auto dest!:length_pos_if_in_set)
3957
3958lemma remove1_filter_not[simp]:
3959  "\<not> P x \<Longrightarrow> remove1 x (filter P xs) = filter P xs"
3960by(induct xs) auto
3961
3962lemma filter_remove1:
3963  "filter Q (remove1 x xs) = remove1 x (filter Q xs)"
3964by (induct xs) auto
3965
3966lemma notin_set_remove1[simp]: "x \<notin> set xs \<Longrightarrow> x \<notin> set(remove1 y xs)"
3967by(insert set_remove1_subset) fast
3968
3969lemma distinct_remove1[simp]: "distinct xs \<Longrightarrow> distinct(remove1 x xs)"
3970by (induct xs) simp_all
3971
3972lemma remove1_remdups:
3973  "distinct xs \<Longrightarrow> remove1 x (remdups xs) = remdups (remove1 x xs)"
3974by (induct xs) simp_all
3975
3976lemma remove1_idem: "x \<notin> set xs \<Longrightarrow> remove1 x xs = xs"
3977by (induct xs) simp_all
3978
3979
3980subsubsection \<open>@{const removeAll}\<close>
3981
3982lemma removeAll_filter_not_eq:
3983  "removeAll x = filter (\<lambda>y. x \<noteq> y)"
3984proof
3985  fix xs
3986  show "removeAll x xs = filter (\<lambda>y. x \<noteq> y) xs"
3987    by (induct xs) auto
3988qed
3989
3990lemma removeAll_append[simp]:
3991  "removeAll x (xs @ ys) = removeAll x xs @ removeAll x ys"
3992by (induct xs) auto
3993
3994lemma set_removeAll[simp]: "set(removeAll x xs) = set xs - {x}"
3995by (induct xs) auto
3996
3997lemma removeAll_id[simp]: "x \<notin> set xs \<Longrightarrow> removeAll x xs = xs"
3998by (induct xs) auto
3999
4000(* Needs count:: 'a \<Rightarrow> 'a list \<Rightarrow> nat
4001lemma length_removeAll:
4002  "length(removeAll x xs) = length xs - count x xs"
4003*)
4004
4005lemma removeAll_filter_not[simp]:
4006  "\<not> P x \<Longrightarrow> removeAll x (filter P xs) = filter P xs"
4007by(induct xs) auto
4008
4009lemma distinct_removeAll:
4010  "distinct xs \<Longrightarrow> distinct (removeAll x xs)"
4011by (simp add: removeAll_filter_not_eq)
4012
4013lemma distinct_remove1_removeAll:
4014  "distinct xs \<Longrightarrow> remove1 x xs = removeAll x xs"
4015by (induct xs) simp_all
4016
4017lemma map_removeAll_inj_on: "inj_on f (insert x (set xs)) \<Longrightarrow>
4018  map f (removeAll x xs) = removeAll (f x) (map f xs)"
4019by (induct xs) (simp_all add:inj_on_def)
4020
4021lemma map_removeAll_inj: "inj f \<Longrightarrow>
4022  map f (removeAll x xs) = removeAll (f x) (map f xs)"
4023by (rule map_removeAll_inj_on, erule subset_inj_on, rule subset_UNIV)
4024
4025lemma length_removeAll_less_eq [simp]:
4026  "length (removeAll x xs) \<le> length xs"
4027  by (simp add: removeAll_filter_not_eq)
4028
4029lemma length_removeAll_less [termination_simp]:
4030  "x \<in> set xs \<Longrightarrow> length (removeAll x xs) < length xs"
4031  by (auto dest: length_filter_less simp add: removeAll_filter_not_eq)
4032
4033
4034subsubsection \<open>@{const replicate}\<close>
4035
4036lemma length_replicate [simp]: "length (replicate n x) = n"
4037by (induct n) auto
4038
4039lemma replicate_eqI:
4040  assumes "length xs = n" and "\<And>y. y \<in> set xs \<Longrightarrow> y = x"
4041  shows "xs = replicate n x"
4042  using assms
4043proof (induct xs arbitrary: n)
4044  case Nil then show ?case by simp
4045next
4046  case (Cons x xs) then show ?case by (cases n) simp_all
4047qed
4048
4049lemma Ex_list_of_length: "\<exists>xs. length xs = n"
4050by (rule exI[of _ "replicate n undefined"]) simp
4051
4052lemma map_replicate [simp]: "map f (replicate n x) = replicate n (f x)"
4053by (induct n) auto
4054
4055lemma map_replicate_const:
4056  "map (\<lambda> x. k) lst = replicate (length lst) k"
4057by (induct lst) auto
4058
4059lemma replicate_app_Cons_same:
4060"(replicate n x) @ (x # xs) = x # replicate n x @ xs"
4061by (induct n) auto
4062
4063lemma rev_replicate [simp]: "rev (replicate n x) = replicate n x"
4064by (induct n) (auto simp: replicate_app_Cons_same)
4065
4066lemma replicate_add: "replicate (n + m) x = replicate n x @ replicate m x"
4067by (induct n) auto
4068
4069text\<open>Courtesy of Matthias Daum:\<close>
4070lemma append_replicate_commute:
4071  "replicate n x @ replicate k x = replicate k x @ replicate n x"
4072apply (simp add: replicate_add [symmetric])
4073apply (simp add: add.commute)
4074done
4075
4076text\<open>Courtesy of Andreas Lochbihler:\<close>
4077lemma filter_replicate:
4078  "filter P (replicate n x) = (if P x then replicate n x else [])"
4079by(induct n) auto
4080
4081lemma hd_replicate [simp]: "n \<noteq> 0 ==> hd (replicate n x) = x"
4082by (induct n) auto
4083
4084lemma tl_replicate [simp]: "tl (replicate n x) = replicate (n - 1) x"
4085by (induct n) auto
4086
4087lemma last_replicate [simp]: "n \<noteq> 0 ==> last (replicate n x) = x"
4088by (atomize (full), induct n) auto
4089
4090lemma nth_replicate[simp]: "i < n ==> (replicate n x)!i = x"
4091by (induct n arbitrary: i)(auto simp: nth_Cons split: nat.split)
4092
4093text\<open>Courtesy of Matthias Daum (2 lemmas):\<close>
4094lemma take_replicate[simp]: "take i (replicate k x) = replicate (min i k) x"
4095apply (case_tac "k \<le> i")
4096 apply  (simp add: min_def)
4097apply (drule not_le_imp_less)
4098apply (simp add: min_def)
4099apply (subgoal_tac "replicate k x = replicate i x @ replicate (k - i) x")
4100 apply  simp
4101apply (simp add: replicate_add [symmetric])
4102done
4103
4104lemma drop_replicate[simp]: "drop i (replicate k x) = replicate (k-i) x"
4105apply (induct k arbitrary: i)
4106 apply simp
4107apply clarsimp
4108apply (case_tac i)
4109 apply simp
4110apply clarsimp
4111done
4112
4113lemma set_replicate_Suc: "set (replicate (Suc n) x) = {x}"
4114by (induct n) auto
4115
4116lemma set_replicate [simp]: "n \<noteq> 0 ==> set (replicate n x) = {x}"
4117by (fast dest!: not0_implies_Suc intro!: set_replicate_Suc)
4118
4119lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})"
4120by auto
4121
4122lemma in_set_replicate[simp]: "(x \<in> set (replicate n y)) = (x = y \<and> n \<noteq> 0)"
4123by (simp add: set_replicate_conv_if)
4124
4125lemma Ball_set_replicate[simp]:
4126  "(\<forall>x \<in> set(replicate n a). P x) = (P a \<or> n=0)"
4127by(simp add: set_replicate_conv_if)
4128
4129lemma Bex_set_replicate[simp]:
4130  "(\<exists>x \<in> set(replicate n a). P x) = (P a \<and> n\<noteq>0)"
4131by(simp add: set_replicate_conv_if)
4132
4133lemma replicate_append_same:
4134  "replicate i x @ [x] = x # replicate i x"
4135  by (induct i) simp_all
4136
4137lemma map_replicate_trivial:
4138  "map (\<lambda>i. x) [0..<i] = replicate i x"
4139by (induct i) (simp_all add: replicate_append_same)
4140
4141lemma concat_replicate_trivial[simp]:
4142  "concat (replicate i []) = []"
4143  by (induct i) (auto simp add: map_replicate_const)
4144
4145lemma replicate_empty[simp]: "(replicate n x = []) \<longleftrightarrow> n=0"
4146by (induct n) auto
4147
4148lemma empty_replicate[simp]: "([] = replicate n x) \<longleftrightarrow> n=0"
4149by (induct n) auto
4150
4151lemma replicate_eq_replicate[simp]:
4152  "(replicate m x = replicate n y) \<longleftrightarrow> (m=n \<and> (m\<noteq>0 \<longrightarrow> x=y))"
4153apply(induct m arbitrary: n)
4154 apply simp
4155apply(induct_tac n)
4156apply auto
4157done
4158
4159lemma replicate_length_filter:
4160  "replicate (length (filter (\<lambda>y. x = y) xs)) x = filter (\<lambda>y. x = y) xs"
4161  by (induct xs) auto
4162
4163lemma comm_append_are_replicate:
4164  "\<lbrakk> xs \<noteq> []; ys \<noteq> []; xs @ ys = ys @ xs \<rbrakk>
4165  \<Longrightarrow> \<exists>m n zs. concat (replicate m zs) = xs \<and> concat (replicate n zs) = ys"
4166proof (induction "length (xs @ ys)" arbitrary: xs ys rule: less_induct)
4167  case less
4168
4169  define xs' ys' where "xs' = (if (length xs \<le> length ys) then xs else ys)"
4170    and "ys' = (if (length xs \<le> length ys) then ys else xs)"
4171  then have
4172    prems': "length xs' \<le> length ys'"
4173            "xs' @ ys' = ys' @ xs'"
4174      and "xs' \<noteq> []"
4175      and len: "length (xs @ ys) = length (xs' @ ys')"
4176    using less by (auto intro: less.hyps)
4177
4178  from prems'
4179  obtain ws where "ys' = xs' @ ws"
4180    by (auto simp: append_eq_append_conv2)
4181
4182  have "\<exists>m n zs. concat (replicate m zs) = xs' \<and> concat (replicate n zs) = ys'"
4183  proof (cases "ws = []")
4184    case True
4185    then have "concat (replicate 1 xs') = xs'"
4186      and "concat (replicate 1 xs') = ys'"
4187      using \<open>ys' = xs' @ ws\<close> by auto
4188    then show ?thesis by blast
4189  next
4190    case False
4191    from \<open>ys' = xs' @ ws\<close> and \<open>xs' @ ys' = ys' @ xs'\<close>
4192    have "xs' @ ws = ws @ xs'" by simp
4193    then have "\<exists>m n zs. concat (replicate m zs) = xs' \<and> concat (replicate n zs) = ws"
4194      using False and \<open>xs' \<noteq> []\<close> and \<open>ys' = xs' @ ws\<close> and len
4195      by (intro less.hyps) auto
4196    then obtain m n zs where *: "concat (replicate m zs) = xs'"
4197      and "concat (replicate n zs) = ws" by blast
4198    then have "concat (replicate (m + n) zs) = ys'"
4199      using \<open>ys' = xs' @ ws\<close>
4200      by (simp add: replicate_add)
4201    with * show ?thesis by blast
4202  qed
4203  then show ?case
4204    using xs'_def ys'_def by meson
4205qed
4206
4207lemma comm_append_is_replicate:
4208  fixes xs ys :: "'a list"
4209  assumes "xs \<noteq> []" "ys \<noteq> []"
4210  assumes "xs @ ys = ys @ xs"
4211  shows "\<exists>n zs. n > 1 \<and> concat (replicate n zs) = xs @ ys"
4212proof -
4213  obtain m n zs where "concat (replicate m zs) = xs"
4214    and "concat (replicate n zs) = ys"
4215    using comm_append_are_replicate[of xs ys, OF assms] by blast
4216  then have "m + n > 1" and "concat (replicate (m+n) zs) = xs @ ys"
4217    using \<open>xs \<noteq> []\<close> and \<open>ys \<noteq> []\<close>
4218    by (auto simp: replicate_add)
4219  then show ?thesis by blast
4220qed
4221
4222lemma Cons_replicate_eq:
4223  "x # xs = replicate n y \<longleftrightarrow> x = y \<and> n > 0 \<and> xs = replicate (n - 1) x"
4224by (induct n) auto
4225
4226lemma replicate_length_same:
4227  "(\<forall>y\<in>set xs. y = x) \<Longrightarrow> replicate (length xs) x = xs"
4228by (induct xs) simp_all
4229
4230lemma foldr_replicate [simp]:
4231  "foldr f (replicate n x) = f x ^^ n"
4232by (induct n) (simp_all)
4233
4234lemma fold_replicate [simp]:
4235  "fold f (replicate n x) = f x ^^ n"
4236by (subst foldr_fold [symmetric]) simp_all
4237
4238
4239subsubsection \<open>@{const enumerate}\<close>
4240
4241lemma enumerate_simps [simp, code]:
4242  "enumerate n [] = []"
4243  "enumerate n (x # xs) = (n, x) # enumerate (Suc n) xs"
4244apply (auto simp add: enumerate_eq_zip not_le)
4245apply (cases "n < n + length xs")
4246 apply (auto simp add: upt_conv_Cons)
4247done
4248
4249lemma length_enumerate [simp]:
4250  "length (enumerate n xs) = length xs"
4251by (simp add: enumerate_eq_zip)
4252
4253lemma map_fst_enumerate [simp]:
4254  "map fst (enumerate n xs) = [n..<n + length xs]"
4255by (simp add: enumerate_eq_zip)
4256
4257lemma map_snd_enumerate [simp]:
4258  "map snd (enumerate n xs) = xs"
4259by (simp add: enumerate_eq_zip)
4260
4261lemma in_set_enumerate_eq:
4262  "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"
4263proof -
4264  { fix m
4265    assume "n \<le> m"
4266    moreover assume "m < length xs + n"
4267    ultimately have "[n..<n + length xs] ! (m - n) = m \<and>
4268      xs ! (m - n) = xs ! (m - n) \<and> m - n < length xs" by auto
4269    then have "\<exists>q. [n..<n + length xs] ! q = m \<and>
4270        xs ! q = xs ! (m - n) \<and> q < length xs" ..
4271  } then show ?thesis by (cases p) (auto simp add: enumerate_eq_zip in_set_zip)
4272qed
4273
4274lemma nth_enumerate_eq: "m < length xs \<Longrightarrow> enumerate n xs ! m = (n + m, xs ! m)"
4275by (simp add: enumerate_eq_zip)
4276
4277lemma enumerate_replicate_eq:
4278  "enumerate n (replicate m a) = map (\<lambda>q. (q, a)) [n..<n + m]"
4279by (rule pair_list_eqI)
4280   (simp_all add: enumerate_eq_zip comp_def map_replicate_const)
4281
4282lemma enumerate_Suc_eq:
4283  "enumerate (Suc n) xs = map (apfst Suc) (enumerate n xs)"
4284by (rule pair_list_eqI)
4285   (simp_all add: not_le, simp del: map_map add: map_Suc_upt map_map [symmetric])
4286
4287lemma distinct_enumerate [simp]:
4288  "distinct (enumerate n xs)"
4289by (simp add: enumerate_eq_zip distinct_zipI1)
4290
4291lemma enumerate_append_eq:
4292  "enumerate n (xs @ ys) = enumerate n xs @ enumerate (n + length xs) ys"
4293unfolding enumerate_eq_zip
4294apply auto
4295 apply (subst zip_append [symmetric]) apply simp
4296 apply (subst upt_add_eq_append [symmetric])
4297 apply (simp_all add: ac_simps)
4298done
4299
4300lemma enumerate_map_upt:
4301  "enumerate n (map f [n..<m]) = map (\<lambda>k. (k, f k)) [n..<m]"
4302by (cases "n \<le> m") (simp_all add: zip_map2 zip_same_conv_map enumerate_eq_zip)
4303
4304
4305subsubsection \<open>@{const rotate1} and @{const rotate}\<close>
4306
4307lemma rotate0[simp]: "rotate 0 = id"
4308by(simp add:rotate_def)
4309
4310lemma rotate_Suc[simp]: "rotate (Suc n) xs = rotate1(rotate n xs)"
4311by(simp add:rotate_def)
4312
4313lemma rotate_add:
4314  "rotate (m+n) = rotate m \<circ> rotate n"
4315by(simp add:rotate_def funpow_add)
4316
4317lemma rotate_rotate: "rotate m (rotate n xs) = rotate (m+n) xs"
4318by(simp add:rotate_add)
4319
4320lemma rotate1_rotate_swap: "rotate1 (rotate n xs) = rotate n (rotate1 xs)"
4321by(simp add:rotate_def funpow_swap1)
4322
4323lemma rotate1_length01[simp]: "length xs <= 1 \<Longrightarrow> rotate1 xs = xs"
4324by(cases xs) simp_all
4325
4326lemma rotate_length01[simp]: "length xs <= 1 \<Longrightarrow> rotate n xs = xs"
4327apply(induct n)
4328 apply simp
4329apply (simp add:rotate_def)
4330done
4331
4332lemma rotate1_hd_tl: "xs \<noteq> [] \<Longrightarrow> rotate1 xs = tl xs @ [hd xs]"
4333by (cases xs) simp_all
4334
4335lemma rotate_drop_take:
4336  "rotate n xs = drop (n mod length xs) xs @ take (n mod length xs) xs"
4337apply(induct n)
4338 apply simp
4339apply(simp add:rotate_def)
4340apply(cases "xs = []")
4341 apply (simp)
4342apply(case_tac "n mod length xs = 0")
4343 apply(simp add:mod_Suc)
4344 apply(simp add: rotate1_hd_tl drop_Suc take_Suc)
4345apply(simp add:mod_Suc rotate1_hd_tl drop_Suc[symmetric] drop_tl[symmetric]
4346                take_hd_drop linorder_not_le)
4347done
4348
4349lemma rotate_conv_mod: "rotate n xs = rotate (n mod length xs) xs"
4350by(simp add:rotate_drop_take)
4351
4352lemma rotate_id[simp]: "n mod length xs = 0 \<Longrightarrow> rotate n xs = xs"
4353by(simp add:rotate_drop_take)
4354
4355lemma length_rotate1[simp]: "length(rotate1 xs) = length xs"
4356by (cases xs) simp_all
4357
4358lemma length_rotate[simp]: "length(rotate n xs) = length xs"
4359by (induct n arbitrary: xs) (simp_all add:rotate_def)
4360
4361lemma distinct1_rotate[simp]: "distinct(rotate1 xs) = distinct xs"
4362by (cases xs) auto
4363
4364lemma distinct_rotate[simp]: "distinct(rotate n xs) = distinct xs"
4365by (induct n) (simp_all add:rotate_def)
4366
4367lemma rotate_map: "rotate n (map f xs) = map f (rotate n xs)"
4368by(simp add:rotate_drop_take take_map drop_map)
4369
4370lemma set_rotate1[simp]: "set(rotate1 xs) = set xs"
4371by (cases xs) auto
4372
4373lemma set_rotate[simp]: "set(rotate n xs) = set xs"
4374by (induct n) (simp_all add:rotate_def)
4375
4376lemma rotate1_is_Nil_conv[simp]: "(rotate1 xs = []) = (xs = [])"
4377by (cases xs) auto
4378
4379lemma rotate_is_Nil_conv[simp]: "(rotate n xs = []) = (xs = [])"
4380by (induct n) (simp_all add:rotate_def)
4381
4382lemma rotate_rev:
4383  "rotate n (rev xs) = rev(rotate (length xs - (n mod length xs)) xs)"
4384proof (cases "length xs = 0 \<or> n mod length xs = 0")
4385  case False
4386  then show ?thesis
4387    by(simp add:rotate_drop_take rev_drop rev_take)
4388qed force
4389
4390lemma hd_rotate_conv_nth: "xs \<noteq> [] \<Longrightarrow> hd(rotate n xs) = xs!(n mod length xs)"
4391apply(simp add:rotate_drop_take hd_append hd_drop_conv_nth hd_conv_nth)
4392apply(subgoal_tac "length xs \<noteq> 0")
4393 prefer 2 apply simp
4394using mod_less_divisor[of "length xs" n] by arith
4395
4396lemma rotate_append: "rotate (length l) (l @ q) = q @ l"
4397  by (induct l arbitrary: q) (auto simp add: rotate1_rotate_swap)
4398
4399
4400subsubsection \<open>@{const nths} --- a generalization of @{const nth} to sets\<close>
4401
4402lemma nths_empty [simp]: "nths xs {} = []"
4403by (auto simp add: nths_def)
4404
4405lemma nths_nil [simp]: "nths [] A = []"
4406by (auto simp add: nths_def)
4407
4408lemma length_nths:
4409  "length (nths xs I) = card{i. i < length xs \<and> i \<in> I}"
4410by(simp add: nths_def length_filter_conv_card cong:conj_cong)
4411
4412lemma nths_shift_lemma_Suc:
4413  "map fst (filter (%p. P(Suc(snd p))) (zip xs is)) =
4414   map fst (filter (%p. P(snd p)) (zip xs (map Suc is)))"
4415apply(induct xs arbitrary: "is")
4416 apply simp
4417apply (case_tac "is")
4418 apply simp
4419apply simp
4420done
4421
4422lemma nths_shift_lemma:
4423  "map fst (filter (\<lambda>p. snd p \<in> A) (zip xs [i..<i + length xs])) =
4424   map fst (filter (\<lambda>p. snd p + i \<in> A) (zip xs [0..<length xs]))"
4425by (induct xs rule: rev_induct) (simp_all add: add.commute)
4426
4427lemma nths_append:
4428  "nths (l @ l') A = nths l A @ nths l' {j. j + length l \<in> A}"
4429apply (unfold nths_def)
4430apply (induct l' rule: rev_induct, simp)
4431apply (simp add: upt_add_eq_append[of 0] nths_shift_lemma)
4432apply (simp add: add.commute)
4433done
4434
4435lemma nths_Cons:
4436  "nths (x # l) A = (if 0 \<in> A then [x] else []) @ nths l {j. Suc j \<in> A}"
4437apply (induct l rule: rev_induct)
4438 apply (simp add: nths_def)
4439apply (simp del: append_Cons add: append_Cons[symmetric] nths_append)
4440done
4441
4442lemma nths_map: "nths (map f xs) I = map f (nths xs I)"
4443by(induction xs arbitrary: I) (simp_all add: nths_Cons)
4444
4445lemma set_nths: "set(nths xs I) = {xs!i|i. i<size xs \<and> i \<in> I}"
4446apply(induct xs arbitrary: I)
4447apply(auto simp: nths_Cons nth_Cons split:nat.split dest!: gr0_implies_Suc)
4448done
4449
4450lemma set_nths_subset: "set(nths xs I) \<subseteq> set xs"
4451by(auto simp add:set_nths)
4452
4453lemma notin_set_nthsI[simp]: "x \<notin> set xs \<Longrightarrow> x \<notin> set(nths xs I)"
4454by(auto simp add:set_nths)
4455
4456lemma in_set_nthsD: "x \<in> set(nths xs I) \<Longrightarrow> x \<in> set xs"
4457by(auto simp add:set_nths)
4458
4459lemma nths_singleton [simp]: "nths [x] A = (if 0 \<in> A then [x] else [])"
4460by (simp add: nths_Cons)
4461
4462lemma distinct_nthsI[simp]: "distinct xs \<Longrightarrow> distinct (nths xs I)"
4463by (induct xs arbitrary: I) (auto simp: nths_Cons)
4464
4465lemma nths_upt_eq_take [simp]: "nths l {..<n} = take n l"
4466by (induct l rule: rev_induct)
4467   (simp_all split: nat_diff_split add: nths_append)
4468
4469lemma filter_eq_nths: "filter P xs = nths xs {i. i<length xs \<and> P(xs!i)}"
4470by(induction xs) (auto simp: nths_Cons)
4471
4472lemma filter_in_nths:
4473  "distinct xs \<Longrightarrow> filter (%x. x \<in> set (nths xs s)) xs = nths xs s"
4474proof (induct xs arbitrary: s)
4475  case Nil thus ?case by simp
4476next
4477  case (Cons a xs)
4478  then have "\<forall>x. x \<in> set xs \<longrightarrow> x \<noteq> a" by auto
4479  with Cons show ?case by(simp add: nths_Cons cong:filter_cong)
4480qed
4481
4482
4483subsubsection \<open>@{const subseqs} and @{const List.n_lists}\<close>
4484
4485lemma length_subseqs: "length (subseqs xs) = 2 ^ length xs"
4486  by (induct xs) (simp_all add: Let_def)
4487
4488lemma subseqs_powset: "set ` set (subseqs xs) = Pow (set xs)"
4489proof -
4490  have aux: "\<And>x A. set ` Cons x ` A = insert x ` set ` A"
4491    by (auto simp add: image_def)
4492  have "set (map set (subseqs xs)) = Pow (set xs)"
4493    by (induct xs) (simp_all add: aux Let_def Pow_insert Un_commute comp_def del: map_map)
4494  then show ?thesis by simp
4495qed
4496
4497lemma distinct_set_subseqs:
4498  assumes "distinct xs"
4499  shows "distinct (map set (subseqs xs))"
4500proof (rule card_distinct)
4501  have "finite (set xs)" ..
4502  then have "card (Pow (set xs)) = 2 ^ card (set xs)"
4503    by (rule card_Pow)
4504  with assms distinct_card [of xs] have "card (Pow (set xs)) = 2 ^ length xs"
4505    by simp
4506  then show "card (set (map set (subseqs xs))) = length (map set (subseqs xs))"
4507    by (simp add: subseqs_powset length_subseqs)
4508qed
4509
4510lemma n_lists_Nil [simp]: "List.n_lists n [] = (if n = 0 then [[]] else [])"
4511  by (induct n) simp_all
4512
4513lemma length_n_lists_elem: "ys \<in> set (List.n_lists n xs) \<Longrightarrow> length ys = n"
4514  by (induct n arbitrary: ys) auto
4515
4516lemma set_n_lists: "set (List.n_lists n xs) = {ys. length ys = n \<and> set ys \<subseteq> set xs}"
4517proof (rule set_eqI)
4518  fix ys :: "'a list"
4519  show "ys \<in> set (List.n_lists n xs) \<longleftrightarrow> ys \<in> {ys. length ys = n \<and> set ys \<subseteq> set xs}"
4520  proof -
4521    have "ys \<in> set (List.n_lists n xs) \<Longrightarrow> length ys = n"
4522      by (induct n arbitrary: ys) auto
4523    moreover have "\<And>x. ys \<in> set (List.n_lists n xs) \<Longrightarrow> x \<in> set ys \<Longrightarrow> x \<in> set xs"
4524      by (induct n arbitrary: ys) auto
4525    moreover have "set ys \<subseteq> set xs \<Longrightarrow> ys \<in> set (List.n_lists (length ys) xs)"
4526      by (induct ys) auto
4527    ultimately show ?thesis by auto
4528  qed
4529qed
4530
4531lemma subseqs_refl: "xs \<in> set (subseqs xs)"
4532  by (induct xs) (simp_all add: Let_def)
4533
4534lemma subset_subseqs: "X \<subseteq> set xs \<Longrightarrow> X \<in> set ` set (subseqs xs)"
4535  unfolding subseqs_powset by simp
4536
4537lemma Cons_in_subseqsD: "y # ys \<in> set (subseqs xs) \<Longrightarrow> ys \<in> set (subseqs xs)"
4538  by (induct xs) (auto simp: Let_def)
4539
4540lemma subseqs_distinctD: "\<lbrakk> ys \<in> set (subseqs xs); distinct xs \<rbrakk> \<Longrightarrow> distinct ys"
4541proof (induct xs arbitrary: ys)
4542  case (Cons x xs ys)
4543  then show ?case
4544    by (auto simp: Let_def) (metis Pow_iff contra_subsetD image_eqI subseqs_powset)
4545qed simp
4546
4547
4548subsubsection \<open>@{const splice}\<close>
4549
4550lemma splice_Nil2 [simp, code]: "splice xs [] = xs"
4551by (cases xs) simp_all
4552
4553declare splice.simps(1,3)[code]
4554declare splice.simps(2)[simp del]
4555
4556lemma length_splice[simp]: "length(splice xs ys) = length xs + length ys"
4557  by (induct xs ys rule: splice.induct) auto
4558
4559
4560subsubsection \<open>@{const shuffle}\<close>
4561
4562lemma Nil_in_shuffle[simp]: "[] \<in> shuffle xs ys \<longleftrightarrow> xs = [] \<and> ys = []"
4563  by (induct xs ys rule: shuffle.induct) auto
4564
4565lemma shuffleE:
4566  "zs \<in> shuffle xs ys \<Longrightarrow>
4567    (zs = xs \<Longrightarrow> ys = [] \<Longrightarrow> P) \<Longrightarrow>
4568    (zs = ys \<Longrightarrow> xs = [] \<Longrightarrow> P) \<Longrightarrow>
4569    (\<And>x xs' z zs'. xs = x # xs' \<Longrightarrow> zs = z # zs' \<Longrightarrow> x = z \<Longrightarrow> zs' \<in> shuffle xs' ys \<Longrightarrow> P) \<Longrightarrow>
4570    (\<And>y ys' z zs'. ys = y # ys' \<Longrightarrow> zs = z # zs' \<Longrightarrow> y = z \<Longrightarrow> zs' \<in> shuffle xs ys' \<Longrightarrow> P) \<Longrightarrow> P"
4571  by (induct xs ys rule: shuffle.induct) auto
4572
4573lemma Cons_in_shuffle_iff:
4574  "z # zs \<in> shuffle xs ys \<longleftrightarrow>
4575    (xs \<noteq> [] \<and> hd xs = z \<and> zs \<in> shuffle (tl xs) ys \<or>
4576     ys \<noteq> [] \<and> hd ys = z \<and> zs \<in> shuffle xs (tl ys))"
4577  by (induct xs ys rule: shuffle.induct) auto
4578
4579lemma splice_in_shuffle [simp, intro]: "splice xs ys \<in> shuffle xs ys"
4580  by (induction xs ys rule: splice.induct) (simp_all add: Cons_in_shuffle_iff)
4581
4582lemma Nil_in_shuffleI: "xs = [] \<Longrightarrow> ys = [] \<Longrightarrow> [] \<in> shuffle xs ys"
4583  by simp
4584
4585lemma Cons_in_shuffle_leftI: "zs \<in> shuffle xs ys \<Longrightarrow> z # zs \<in> shuffle (z # xs) ys"
4586  by (cases ys) auto
4587
4588lemma Cons_in_shuffle_rightI: "zs \<in> shuffle xs ys \<Longrightarrow> z # zs \<in> shuffle xs (z # ys)"
4589  by (cases xs) auto
4590
4591lemma finite_shuffle [simp, intro]: "finite (shuffle xs ys)"
4592  by (induction xs ys rule: shuffle.induct) simp_all
4593
4594lemma length_shuffle: "zs \<in> shuffle xs ys \<Longrightarrow> length zs = length xs + length ys"
4595  by (induction xs ys arbitrary: zs rule: shuffle.induct) auto
4596
4597lemma set_shuffle: "zs \<in> shuffle xs ys \<Longrightarrow> set zs = set xs \<union> set ys"
4598  by (induction xs ys arbitrary: zs rule: shuffle.induct) auto
4599
4600lemma distinct_disjoint_shuffle:
4601  assumes "distinct xs" "distinct ys" "set xs \<inter> set ys = {}" "zs \<in> shuffle xs ys"
4602  shows   "distinct zs"
4603using assms
4604proof (induction xs ys arbitrary: zs rule: shuffle.induct)
4605  case (3 x xs y ys)
4606  show ?case
4607  proof (cases zs)
4608    case (Cons z zs')
4609    with "3.prems" and "3.IH"[of zs'] show ?thesis by (force dest: set_shuffle)
4610  qed simp_all
4611qed simp_all
4612
4613lemma shuffle_commutes: "shuffle xs ys = shuffle ys xs"
4614  by (induction xs ys rule: shuffle.induct) (simp_all add: Un_commute)
4615
4616lemma Cons_shuffle_subset1: "(#) x ` shuffle xs ys \<subseteq> shuffle (x # xs) ys"
4617  by (cases ys) auto
4618
4619lemma Cons_shuffle_subset2: "(#) y ` shuffle xs ys \<subseteq> shuffle xs (y # ys)"
4620  by (cases xs) auto
4621
4622lemma filter_shuffle:
4623  "filter P ` shuffle xs ys = shuffle (filter P xs) (filter P ys)"
4624proof -
4625  have *: "filter P ` (#) x ` A = (if P x then (#) x ` filter P ` A else filter P ` A)" for x A
4626    by (auto simp: image_image)
4627  show ?thesis
4628  by (induction xs ys rule: shuffle.induct)
4629     (simp_all split: if_splits add: image_Un * Un_absorb1 Un_absorb2
4630           Cons_shuffle_subset1 Cons_shuffle_subset2)
4631qed
4632
4633lemma filter_shuffle_disjoint1:
4634  assumes "set xs \<inter> set ys = {}" "zs \<in> shuffle xs ys"
4635  shows   "filter (\<lambda>x. x \<in> set xs) zs = xs" (is "filter ?P _ = _")
4636    and   "filter (\<lambda>x. x \<notin> set xs) zs = ys" (is "filter ?Q _ = _")
4637  using assms
4638proof -
4639  from assms have "filter ?P zs \<in> filter ?P ` shuffle xs ys" by blast
4640  also have "filter ?P ` shuffle xs ys = shuffle (filter ?P xs) (filter ?P ys)"
4641    by (rule filter_shuffle)
4642  also have "filter ?P xs = xs" by (rule filter_True) simp_all
4643  also have "filter ?P ys = []" by (rule filter_False) (insert assms(1), auto)
4644  also have "shuffle xs [] = {xs}" by simp
4645  finally show "filter ?P zs = xs" by simp
4646next
4647  from assms have "filter ?Q zs \<in> filter ?Q ` shuffle xs ys" by blast
4648  also have "filter ?Q ` shuffle xs ys = shuffle (filter ?Q xs) (filter ?Q ys)"
4649    by (rule filter_shuffle)
4650  also have "filter ?Q ys = ys" by (rule filter_True) (insert assms(1), auto)
4651  also have "filter ?Q xs = []" by (rule filter_False) (insert assms(1), auto)
4652  also have "shuffle [] ys = {ys}" by simp
4653  finally show "filter ?Q zs = ys" by simp
4654qed
4655
4656lemma filter_shuffle_disjoint2:
4657  assumes "set xs \<inter> set ys = {}" "zs \<in> shuffle xs ys"
4658  shows   "filter (\<lambda>x. x \<in> set ys) zs = ys" "filter (\<lambda>x. x \<notin> set ys) zs = xs"
4659  using filter_shuffle_disjoint1[of ys xs zs] assms
4660  by (simp_all add: shuffle_commutes Int_commute)
4661
4662lemma partition_in_shuffle:
4663  "xs \<in> shuffle (filter P xs) (filter (\<lambda>x. \<not>P x) xs)"
4664proof (induction xs)
4665  case (Cons x xs)
4666  show ?case
4667  proof (cases "P x")
4668    case True
4669    hence "x # xs \<in> (#) x ` shuffle (filter P xs) (filter (\<lambda>x. \<not>P x) xs)"
4670      by (intro imageI Cons.IH)
4671    also have "\<dots> \<subseteq> shuffle (filter P (x # xs)) (filter (\<lambda>x. \<not>P x) (x # xs))"
4672      by (simp add: True Cons_shuffle_subset1)
4673    finally show ?thesis .
4674  next
4675    case False
4676    hence "x # xs \<in> (#) x ` shuffle (filter P xs) (filter (\<lambda>x. \<not>P x) xs)"
4677      by (intro imageI Cons.IH)
4678    also have "\<dots> \<subseteq> shuffle (filter P (x # xs)) (filter (\<lambda>x. \<not>P x) (x # xs))"
4679      by (simp add: False Cons_shuffle_subset2)
4680    finally show ?thesis .
4681  qed
4682qed auto
4683
4684lemma inv_image_partition:
4685  assumes "\<And>x. x \<in> set xs \<Longrightarrow> P x" "\<And>y. y \<in> set ys \<Longrightarrow> \<not>P y"
4686  shows   "partition P -` {(xs, ys)} = shuffle xs ys"
4687proof (intro equalityI subsetI)
4688  fix zs assume zs: "zs \<in> shuffle xs ys"
4689  hence [simp]: "set zs = set xs \<union> set ys" by (rule set_shuffle)
4690  from assms have "filter P zs = filter (\<lambda>x. x \<in> set xs) zs"
4691                  "filter (\<lambda>x. \<not>P x) zs = filter (\<lambda>x. x \<in> set ys) zs"
4692    by (intro filter_cong refl; force)+
4693  moreover from assms have "set xs \<inter> set ys = {}" by auto
4694  ultimately show "zs \<in> partition P -` {(xs, ys)}" using zs
4695    by (simp add: o_def filter_shuffle_disjoint1 filter_shuffle_disjoint2)
4696next
4697  fix zs assume "zs \<in> partition P -` {(xs, ys)}"
4698  thus "zs \<in> shuffle xs ys" using partition_in_shuffle[of zs] by (auto simp: o_def)
4699qed
4700
4701
4702subsubsection \<open>Transpose\<close>
4703
4704function transpose where
4705"transpose []             = []" |
4706"transpose ([]     # xss) = transpose xss" |
4707"transpose ((x#xs) # xss) =
4708  (x # [h. (h#t) \<leftarrow> xss]) # transpose (xs # [t. (h#t) \<leftarrow> xss])"
4709by pat_completeness auto
4710
4711lemma transpose_aux_filter_head:
4712  "concat (map (case_list [] (\<lambda>h t. [h])) xss) =
4713  map (\<lambda>xs. hd xs) (filter (\<lambda>ys. ys \<noteq> []) xss)"
4714  by (induct xss) (auto split: list.split)
4715
4716lemma transpose_aux_filter_tail:
4717  "concat (map (case_list [] (\<lambda>h t. [t])) xss) =
4718  map (\<lambda>xs. tl xs) (filter (\<lambda>ys. ys \<noteq> []) xss)"
4719  by (induct xss) (auto split: list.split)
4720
4721lemma transpose_aux_max:
4722  "max (Suc (length xs)) (foldr (\<lambda>xs. max (length xs)) xss 0) =
4723  Suc (max (length xs) (foldr (\<lambda>x. max (length x - Suc 0)) (filter (\<lambda>ys. ys \<noteq> []) xss) 0))"
4724  (is "max _ ?foldB = Suc (max _ ?foldA)")
4725proof (cases "(filter (\<lambda>ys. ys \<noteq> []) xss) = []")
4726  case True
4727  hence "foldr (\<lambda>xs. max (length xs)) xss 0 = 0"
4728  proof (induct xss)
4729    case (Cons x xs)
4730    then have "x = []" by (cases x) auto
4731    with Cons show ?case by auto
4732  qed simp
4733  thus ?thesis using True by simp
4734next
4735  case False
4736
4737  have foldA: "?foldA = foldr (\<lambda>x. max (length x)) (filter (\<lambda>ys. ys \<noteq> []) xss) 0 - 1"
4738    by (induct xss) auto
4739  have foldB: "?foldB = foldr (\<lambda>x. max (length x)) (filter (\<lambda>ys. ys \<noteq> []) xss) 0"
4740    by (induct xss) auto
4741
4742  have "0 < ?foldB"
4743  proof -
4744    from False
4745    obtain z zs where zs: "(filter (\<lambda>ys. ys \<noteq> []) xss) = z#zs" by (auto simp: neq_Nil_conv)
4746    hence "z \<in> set (filter (\<lambda>ys. ys \<noteq> []) xss)" by auto
4747    hence "z \<noteq> []" by auto
4748    thus ?thesis
4749      unfolding foldB zs
4750      by (auto simp: max_def intro: less_le_trans)
4751  qed
4752  thus ?thesis
4753    unfolding foldA foldB max_Suc_Suc[symmetric]
4754    by simp
4755qed
4756
4757termination transpose
4758  by (relation "measure (\<lambda>xs. foldr (\<lambda>xs. max (length xs)) xs 0 + length xs)")
4759     (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max less_Suc_eq_le)
4760
4761lemma transpose_empty: "(transpose xs = []) \<longleftrightarrow> (\<forall>x \<in> set xs. x = [])"
4762  by (induct rule: transpose.induct) simp_all
4763
4764lemma length_transpose:
4765  fixes xs :: "'a list list"
4766  shows "length (transpose xs) = foldr (\<lambda>xs. max (length xs)) xs 0"
4767  by (induct rule: transpose.induct)
4768    (auto simp: transpose_aux_filter_tail foldr_map comp_def transpose_aux_max
4769                max_Suc_Suc[symmetric] simp del: max_Suc_Suc)
4770
4771lemma nth_transpose:
4772  fixes xs :: "'a list list"
4773  assumes "i < length (transpose xs)"
4774  shows "transpose xs ! i = map (\<lambda>xs. xs ! i) (filter (\<lambda>ys. i < length ys) xs)"
4775using assms proof (induct arbitrary: i rule: transpose.induct)
4776  case (3 x xs xss)
4777  define XS where "XS = (x # xs) # xss"
4778  hence [simp]: "XS \<noteq> []" by auto
4779  thus ?case
4780  proof (cases i)
4781    case 0
4782    thus ?thesis by (simp add: transpose_aux_filter_head hd_conv_nth)
4783  next
4784    case (Suc j)
4785    have *: "\<And>xss. xs # map tl xss = map tl ((x#xs)#xss)" by simp
4786    have **: "\<And>xss. (x#xs) # filter (\<lambda>ys. ys \<noteq> []) xss = filter (\<lambda>ys. ys \<noteq> []) ((x#xs)#xss)" by simp
4787    { fix x have "Suc j < length x \<longleftrightarrow> x \<noteq> [] \<and> j < length x - Suc 0"
4788      by (cases x) simp_all
4789    } note *** = this
4790
4791    have j_less: "j < length (transpose (xs # concat (map (case_list [] (\<lambda>h t. [t])) xss)))"
4792      using "3.prems" by (simp add: transpose_aux_filter_tail length_transpose Suc)
4793
4794    show ?thesis
4795      unfolding transpose.simps \<open>i = Suc j\<close> nth_Cons_Suc "3.hyps"[OF j_less]
4796      apply (auto simp: transpose_aux_filter_tail filter_map comp_def length_transpose * ** *** XS_def[symmetric])
4797      apply (rule list.exhaust)
4798      by auto
4799  qed
4800qed simp_all
4801
4802lemma transpose_map_map:
4803  "transpose (map (map f) xs) = map (map f) (transpose xs)"
4804proof (rule nth_equalityI, safe)
4805  have [simp]: "length (transpose (map (map f) xs)) = length (transpose xs)"
4806    by (simp add: length_transpose foldr_map comp_def)
4807  show "length (transpose (map (map f) xs)) = length (map (map f) (transpose xs))" by simp
4808
4809  fix i assume "i < length (transpose (map (map f) xs))"
4810  thus "transpose (map (map f) xs) ! i = map (map f) (transpose xs) ! i"
4811    by (simp add: nth_transpose filter_map comp_def)
4812qed
4813
4814subsubsection \<open>@{const min} and @{const arg_min}\<close>
4815
4816lemma min_list_Min: "xs \<noteq> [] \<Longrightarrow> min_list xs = Min (set xs)"
4817by (induction xs rule: induct_list012)(auto)
4818
4819lemma f_arg_min_list_f: "xs \<noteq> [] \<Longrightarrow> f (arg_min_list f xs) = Min (f ` (set xs))"
4820by(induction f xs rule: arg_min_list.induct) (auto simp: min_def intro!: antisym)
4821
4822lemma arg_min_list_in: "xs \<noteq> [] \<Longrightarrow> arg_min_list f xs \<in> set xs"
4823by(induction xs rule: induct_list012) (auto simp: Let_def)
4824
4825
4826subsubsection \<open>(In)finiteness\<close>
4827
4828lemma finite_maxlen:
4829  "finite (M::'a list set) \<Longrightarrow> \<exists>n. \<forall>s\<in>M. size s < n"
4830proof (induct rule: finite.induct)
4831  case emptyI show ?case by simp
4832next
4833  case (insertI M xs)
4834  then obtain n where "\<forall>s\<in>M. length s < n" by blast
4835  hence "\<forall>s\<in>insert xs M. size s < max n (size xs) + 1" by auto
4836  thus ?case ..
4837qed
4838
4839lemma lists_length_Suc_eq:
4840  "{xs. set xs \<subseteq> A \<and> length xs = Suc n} =
4841    (\<lambda>(xs, n). n#xs) ` ({xs. set xs \<subseteq> A \<and> length xs = n} \<times> A)"
4842  by (auto simp: length_Suc_conv)
4843
4844lemma
4845  assumes "finite A"
4846  shows finite_lists_length_eq: "finite {xs. set xs \<subseteq> A \<and> length xs = n}"
4847  and card_lists_length_eq: "card {xs. set xs \<subseteq> A \<and> length xs = n} = (card A)^n"
4848  using \<open>finite A\<close>
4849  by (induct n)
4850     (auto simp: card_image inj_split_Cons lists_length_Suc_eq cong: conj_cong)
4851
4852lemma finite_lists_length_le:
4853  assumes "finite A" shows "finite {xs. set xs \<subseteq> A \<and> length xs \<le> n}"
4854 (is "finite ?S")
4855proof-
4856  have "?S = (\<Union>n\<in>{0..n}. {xs. set xs \<subseteq> A \<and> length xs = n})" by auto
4857  thus ?thesis by (auto intro!: finite_lists_length_eq[OF \<open>finite A\<close>] simp only:)
4858qed
4859
4860lemma card_lists_length_le:
4861  assumes "finite A" shows "card {xs. set xs \<subseteq> A \<and> length xs \<le> n} = (\<Sum>i\<le>n. card A^i)"
4862proof -
4863  have "(\<Sum>i\<le>n. card A^i) = card (\<Union>i\<le>n. {xs. set xs \<subseteq> A \<and> length xs = i})"
4864    using \<open>finite A\<close>
4865    by (subst card_UN_disjoint)
4866       (auto simp add: card_lists_length_eq finite_lists_length_eq)
4867  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}"
4868    by auto
4869  finally show ?thesis by simp
4870qed
4871
4872lemma finite_lists_distinct_length_eq [intro]:
4873  assumes "finite A"
4874  shows "finite {xs. length xs = n \<and> distinct xs \<and> set xs \<subseteq> A}" (is "finite ?S")
4875proof -
4876  have "finite {xs. set xs \<subseteq> A \<and> length xs = n}"
4877    using \<open>finite A\<close> by (rule finite_lists_length_eq)
4878  moreover have "?S \<subseteq> {xs. set xs \<subseteq> A \<and> length xs = n}" by auto
4879  ultimately show ?thesis using finite_subset by auto
4880qed
4881
4882lemma card_lists_distinct_length_eq:
4883  assumes "finite A" "k \<le> card A"
4884  shows "card {xs. length xs = k \<and> distinct xs \<and> set xs \<subseteq> A} = \<Prod>{card A - k + 1 .. card A}"
4885using assms
4886proof (induct k)
4887  case 0
4888  then have "{xs. length xs = 0 \<and> distinct xs \<and> set xs \<subseteq> A} = {[]}" by auto
4889  then show ?case by simp
4890next
4891  case (Suc k)
4892  let "?k_list" = "\<lambda>k xs. length xs = k \<and> distinct xs \<and> set xs \<subseteq> A"
4893  have inj_Cons: "\<And>A. inj_on (\<lambda>(xs, n). n # xs) A"  by (rule inj_onI) auto
4894
4895  from Suc have "k \<le> card A" by simp
4896  moreover note \<open>finite A\<close>
4897  moreover have "finite {xs. ?k_list k xs}"
4898    by (rule finite_subset) (use finite_lists_length_eq[OF \<open>finite A\<close>, of k] in auto)
4899  moreover have "\<And>i j. i \<noteq> j \<longrightarrow> {i} \<times> (A - set i) \<inter> {j} \<times> (A - set j) = {}"
4900    by auto
4901  moreover have "\<And>i. i \<in> {xs. ?k_list k xs} \<Longrightarrow> card (A - set i) = card A - k"
4902    by (simp add: card_Diff_subset distinct_card)
4903  moreover have "{xs. ?k_list (Suc k) xs} =
4904      (\<lambda>(xs, n). n#xs) ` \<Union>((\<lambda>xs. {xs} \<times> (A - set xs)) ` {xs. ?k_list k xs})"
4905    by (auto simp: length_Suc_conv)
4906  moreover have "Suc (card A - Suc k) = card A - k" using Suc.prems by simp
4907  then have "(card A - k) * \<Prod>{Suc (card A - k)..card A} = \<Prod>{Suc (card A - Suc k)..card A}"
4908    by (subst prod.insert[symmetric]) (simp add: atLeastAtMost_insertL)+
4909  ultimately show ?case
4910    by (simp add: card_image inj_Cons card_UN_disjoint Suc.hyps algebra_simps)
4911qed
4912
4913lemma card_lists_distinct_length_eq':
4914  assumes "k < card A"
4915  shows "card {xs. length xs = k \<and> distinct xs \<and> set xs \<subseteq> A} = \<Prod>{card A - k + 1 .. card A}"
4916proof -
4917  from \<open>k < card A\<close> have "finite A" and "k \<le> card A" using card_infinite by force+
4918  from this show ?thesis by (rule card_lists_distinct_length_eq)
4919qed
4920
4921lemma infinite_UNIV_listI: "\<not> finite(UNIV::'a list set)"
4922apply (rule notI)
4923apply (drule finite_maxlen)
4924apply clarsimp
4925apply (erule_tac x = "replicate n undefined" in allE)
4926by simp
4927
4928
4929subsection \<open>Sorting\<close>
4930
4931subsubsection \<open>@{const sorted_wrt}\<close>
4932
4933text \<open>Sometimes the second equation in the definition of @{const sorted_wrt} is too aggressive
4934because it relates each list element to \emph{all} its successors. Then this equation
4935should be removed and \<open>sorted_wrt2_simps\<close> should be added instead.\<close>
4936
4937lemma sorted_wrt1: "sorted_wrt P [x] = True"
4938by(simp)
4939
4940lemma sorted_wrt2: "transp P \<Longrightarrow> sorted_wrt P (x # y # zs) = (P x y \<and> sorted_wrt P (y # zs))"
4941apply(induction zs arbitrary: x y)
4942apply(auto dest: transpD)
4943apply (meson transpD)
4944done
4945
4946lemmas sorted_wrt2_simps = sorted_wrt1 sorted_wrt2
4947
4948lemma sorted_wrt_true [simp]:
4949  "sorted_wrt (\<lambda>_ _. True) xs"
4950by (induction xs) simp_all
4951
4952lemma sorted_wrt_append:
4953  "sorted_wrt P (xs @ ys) \<longleftrightarrow>
4954  sorted_wrt P xs \<and> sorted_wrt P ys \<and> (\<forall>x\<in>set xs. \<forall>y\<in>set ys. P x y)"
4955by (induction xs) auto
4956
4957lemma sorted_wrt_map:
4958  "sorted_wrt R (map f xs) = sorted_wrt (\<lambda>x y. R (f x) (f y)) xs"
4959by (induction xs) simp_all
4960
4961lemma sorted_wrt_rev:
4962  "sorted_wrt P (rev xs) = sorted_wrt (\<lambda>x y. P y x) xs"
4963by (induction xs) (auto simp add: sorted_wrt_append)
4964
4965lemma sorted_wrt_mono_rel:
4966  "(\<And>x y. \<lbrakk> x \<in> set xs; y \<in> set xs; P x y \<rbrakk> \<Longrightarrow> Q x y) \<Longrightarrow> sorted_wrt P xs \<Longrightarrow> sorted_wrt Q xs"
4967by(induction xs)(auto)
4968
4969lemma sorted_wrt01: "length xs \<le> 1 \<Longrightarrow> sorted_wrt P xs"
4970by(auto simp: le_Suc_eq length_Suc_conv)
4971
4972lemma sorted_wrt_iff_nth_less:
4973  "sorted_wrt P xs = (\<forall>i j. i < j \<longrightarrow> j < length xs \<longrightarrow> P (xs ! i) (xs ! j))"
4974apply(induction xs)
4975apply(auto simp add: in_set_conv_nth Ball_def nth_Cons split: nat.split)
4976done
4977
4978lemma sorted_wrt_nth_less:
4979  "\<lbrakk> sorted_wrt P xs; i < j; j < length xs \<rbrakk> \<Longrightarrow> P (xs ! i) (xs ! j)"
4980by(auto simp: sorted_wrt_iff_nth_less)
4981
4982lemma sorted_wrt_upt[simp]: "sorted_wrt (<) [m..<n]"
4983by(induction n) (auto simp: sorted_wrt_append)
4984
4985lemma sorted_wrt_upto[simp]: "sorted_wrt (<) [i..j]"
4986apply(induction i j rule: upto.induct)
4987apply(subst upto.simps)
4988apply(simp)
4989done
4990
4991text \<open>Each element is greater or equal to its index:\<close>
4992
4993lemma sorted_wrt_less_idx:
4994  "sorted_wrt (<) ns \<Longrightarrow> i < length ns \<Longrightarrow> i \<le> ns!i"
4995proof (induction ns arbitrary: i rule: rev_induct)
4996  case Nil thus ?case by simp
4997next
4998  case snoc
4999  thus ?case
5000    by (auto simp: nth_append sorted_wrt_append)
5001       (metis less_antisym not_less nth_mem)
5002qed
5003
5004
5005subsubsection \<open>@{const sorted}\<close>
5006
5007context linorder
5008begin
5009
5010text \<open>Sometimes the second equation in the definition of @{const sorted} is too aggressive
5011because it relates each list element to \emph{all} its successors. Then this equation
5012should be removed and \<open>sorted2_simps\<close> should be added instead.
5013Executable code is one such use case.\<close>
5014
5015lemma sorted1: "sorted [x] = True"
5016by simp
5017
5018lemma sorted2: "sorted (x # y # zs) = (x \<le> y \<and> sorted (y # zs))"
5019by(induction zs) auto
5020
5021lemmas sorted2_simps = sorted1 sorted2
5022
5023lemmas [code] = sorted.simps(1) sorted2_simps
5024
5025lemma sorted_append:
5026  "sorted (xs@ys) = (sorted xs \<and> sorted ys \<and> (\<forall>x \<in> set xs. \<forall>y \<in> set ys. x\<le>y))"
5027by (simp add: sorted_sorted_wrt sorted_wrt_append)
5028
5029lemma sorted_map:
5030  "sorted (map f xs) = sorted_wrt (\<lambda>x y. f x \<le> f y) xs"
5031by (simp add: sorted_sorted_wrt sorted_wrt_map)
5032
5033lemma sorted01: "length xs \<le> 1 \<Longrightarrow> sorted xs"
5034by (simp add: sorted_sorted_wrt sorted_wrt01)
5035
5036lemma sorted_tl:
5037  "sorted xs \<Longrightarrow> sorted (tl xs)"
5038by (cases xs) (simp_all)
5039
5040lemma sorted_iff_nth_mono_less:
5041  "sorted xs = (\<forall>i j. i < j \<longrightarrow> j < length xs \<longrightarrow> xs ! i \<le> xs ! j)"
5042by (simp add: sorted_sorted_wrt sorted_wrt_iff_nth_less)
5043
5044lemma sorted_iff_nth_mono:
5045  "sorted xs = (\<forall>i j. i \<le> j \<longrightarrow> j < length xs \<longrightarrow> xs ! i \<le> xs ! j)"
5046by (auto simp: sorted_iff_nth_mono_less nat_less_le)
5047
5048lemma sorted_nth_mono:
5049  "sorted xs \<Longrightarrow> i \<le> j \<Longrightarrow> j < length xs \<Longrightarrow> xs!i \<le> xs!j"
5050by (auto simp: sorted_iff_nth_mono)
5051
5052lemma sorted_rev_nth_mono:
5053  "sorted (rev xs) \<Longrightarrow> i \<le> j \<Longrightarrow> j < length xs \<Longrightarrow> xs!j \<le> xs!i"
5054using sorted_nth_mono[ of "rev xs" "length xs - j - 1" "length xs - i - 1"]
5055      rev_nth[of "length xs - i - 1" "xs"] rev_nth[of "length xs - j - 1" "xs"]
5056by auto
5057
5058lemma sorted_map_remove1:
5059  "sorted (map f xs) \<Longrightarrow> sorted (map f (remove1 x xs))"
5060by (induct xs) (auto)
5061
5062lemma sorted_remove1: "sorted xs \<Longrightarrow> sorted (remove1 a xs)"
5063using sorted_map_remove1 [of "\<lambda>x. x"] by simp
5064
5065lemma sorted_butlast:
5066  assumes "xs \<noteq> []" and "sorted xs"
5067  shows "sorted (butlast xs)"
5068proof -
5069  from \<open>xs \<noteq> []\<close> obtain ys y where "xs = ys @ [y]"
5070    by (cases xs rule: rev_cases) auto
5071  with \<open>sorted xs\<close> show ?thesis by (simp add: sorted_append)
5072qed
5073
5074lemma sorted_replicate [simp]: "sorted(replicate n x)"
5075by(induction n) (auto)
5076
5077lemma sorted_remdups[simp]:
5078  "sorted xs \<Longrightarrow> sorted (remdups xs)"
5079by (induct xs) (auto)
5080
5081lemma sorted_remdups_adj[simp]:
5082  "sorted xs \<Longrightarrow> sorted (remdups_adj xs)"
5083by (induct xs rule: remdups_adj.induct, simp_all split: if_split_asm)
5084
5085lemma sorted_nths: "sorted xs \<Longrightarrow> sorted (nths xs I)"
5086by(induction xs arbitrary: I)(auto simp: nths_Cons)
5087
5088lemma sorted_distinct_set_unique:
5089assumes "sorted xs" "distinct xs" "sorted ys" "distinct ys" "set xs = set ys"
5090shows "xs = ys"
5091proof -
5092  from assms have 1: "length xs = length ys" by (auto dest!: distinct_card)
5093  from assms show ?thesis
5094  proof(induct rule:list_induct2[OF 1])
5095    case 1 show ?case by simp
5096  next
5097    case 2 thus ?case by simp (metis Diff_insert_absorb antisym insertE insert_iff)
5098  qed
5099qed
5100
5101lemma map_sorted_distinct_set_unique:
5102  assumes "inj_on f (set xs \<union> set ys)"
5103  assumes "sorted (map f xs)" "distinct (map f xs)"
5104    "sorted (map f ys)" "distinct (map f ys)"
5105  assumes "set xs = set ys"
5106  shows "xs = ys"
5107proof -
5108  from assms have "map f xs = map f ys"
5109    by (simp add: sorted_distinct_set_unique)
5110  with \<open>inj_on f (set xs \<union> set ys)\<close> show "xs = ys"
5111    by (blast intro: map_inj_on)
5112qed
5113
5114lemma
5115  assumes "sorted xs"
5116  shows sorted_take: "sorted (take n xs)"
5117  and sorted_drop: "sorted (drop n xs)"
5118proof -
5119  from assms have "sorted (take n xs @ drop n xs)" by simp
5120  then show "sorted (take n xs)" and "sorted (drop n xs)"
5121    unfolding sorted_append by simp_all
5122qed
5123
5124lemma sorted_dropWhile: "sorted xs \<Longrightarrow> sorted (dropWhile P xs)"
5125  by (auto dest: sorted_drop simp add: dropWhile_eq_drop)
5126
5127lemma sorted_takeWhile: "sorted xs \<Longrightarrow> sorted (takeWhile P xs)"
5128  by (subst takeWhile_eq_take) (auto dest: sorted_take)
5129
5130lemma sorted_filter:
5131  "sorted (map f xs) \<Longrightarrow> sorted (map f (filter P xs))"
5132  by (induct xs) simp_all
5133
5134lemma foldr_max_sorted:
5135  assumes "sorted (rev xs)"
5136  shows "foldr max xs y = (if xs = [] then y else max (xs ! 0) y)"
5137  using assms
5138proof (induct xs)
5139  case (Cons x xs)
5140  then have "sorted (rev xs)" using sorted_append by auto
5141  with Cons show ?case
5142    by (cases xs) (auto simp add: sorted_append max_def)
5143qed simp
5144
5145lemma filter_equals_takeWhile_sorted_rev:
5146  assumes sorted: "sorted (rev (map f xs))"
5147  shows "filter (\<lambda>x. t < f x) xs = takeWhile (\<lambda> x. t < f x) xs"
5148    (is "filter ?P xs = ?tW")
5149proof (rule takeWhile_eq_filter[symmetric])
5150  let "?dW" = "dropWhile ?P xs"
5151  fix x assume "x \<in> set ?dW"
5152  then obtain i where i: "i < length ?dW" and nth_i: "x = ?dW ! i"
5153    unfolding in_set_conv_nth by auto
5154  hence "length ?tW + i < length (?tW @ ?dW)"
5155    unfolding length_append by simp
5156  hence i': "length (map f ?tW) + i < length (map f xs)" by simp
5157  have "(map f ?tW @ map f ?dW) ! (length (map f ?tW) + i) \<le>
5158        (map f ?tW @ map f ?dW) ! (length (map f ?tW) + 0)"
5159    using sorted_rev_nth_mono[OF sorted _ i', of "length ?tW"]
5160    unfolding map_append[symmetric] by simp
5161  hence "f x \<le> f (?dW ! 0)"
5162    unfolding nth_append_length_plus nth_i
5163    using i preorder_class.le_less_trans[OF le0 i] by simp
5164  also have "... \<le> t"
5165    using hd_dropWhile[of "?P" xs] le0[THEN preorder_class.le_less_trans, OF i]
5166    using hd_conv_nth[of "?dW"] by simp
5167  finally show "\<not> t < f x" by simp
5168qed
5169
5170lemma sorted_map_same:
5171  "sorted (map f (filter (\<lambda>x. f x = g xs) xs))"
5172proof (induct xs arbitrary: g)
5173  case Nil then show ?case by simp
5174next
5175  case (Cons x xs)
5176  then have "sorted (map f (filter (\<lambda>y. f y = (\<lambda>xs. f x) xs) xs))" .
5177  moreover from Cons have "sorted (map f (filter (\<lambda>y. f y = (g \<circ> Cons x) xs) xs))" .
5178  ultimately show ?case by simp_all
5179qed
5180
5181lemma sorted_same:
5182  "sorted (filter (\<lambda>x. x = g xs) xs)"
5183using sorted_map_same [of "\<lambda>x. x"] by simp
5184
5185end
5186
5187lemma sorted_upt[simp]: "sorted [m..<n]"
5188by(simp add: sorted_sorted_wrt sorted_wrt_mono_rel[OF _ sorted_wrt_upt])
5189
5190lemma sorted_upto[simp]: "sorted [m..n]"
5191by(simp add: sorted_sorted_wrt sorted_wrt_mono_rel[OF _ sorted_wrt_upto])
5192
5193
5194subsubsection \<open>Sorting functions\<close>
5195
5196text\<open>Currently it is not shown that @{const sort} returns a
5197permutation of its input because the nicest proof is via multisets,
5198which are not part of Main. Alternatively one could define a function
5199that counts the number of occurrences of an element in a list and use
5200that instead of multisets to state the correctness property.\<close>
5201
5202context linorder
5203begin
5204
5205lemma set_insort_key:
5206  "set (insort_key f x xs) = insert x (set xs)"
5207by (induct xs) auto
5208
5209lemma length_insort [simp]:
5210  "length (insort_key f x xs) = Suc (length xs)"
5211by (induct xs) simp_all
5212
5213lemma insort_key_left_comm:
5214  assumes "f x \<noteq> f y"
5215  shows "insort_key f y (insort_key f x xs) = insort_key f x (insort_key f y xs)"
5216by (induct xs) (auto simp add: assms dest: antisym)
5217
5218lemma insort_left_comm:
5219  "insort x (insort y xs) = insort y (insort x xs)"
5220by (cases "x = y") (auto intro: insort_key_left_comm)
5221
5222lemma comp_fun_commute_insort: "comp_fun_commute insort"
5223proof
5224qed (simp add: insort_left_comm fun_eq_iff)
5225
5226lemma sort_key_simps [simp]:
5227  "sort_key f [] = []"
5228  "sort_key f (x#xs) = insort_key f x (sort_key f xs)"
5229by (simp_all add: sort_key_def)
5230
5231lemma sort_key_conv_fold:
5232  assumes "inj_on f (set xs)"
5233  shows "sort_key f xs = fold (insort_key f) xs []"
5234proof -
5235  have "fold (insort_key f) (rev xs) = fold (insort_key f) xs"
5236  proof (rule fold_rev, rule ext)
5237    fix zs
5238    fix x y
5239    assume "x \<in> set xs" "y \<in> set xs"
5240    with assms have *: "f y = f x \<Longrightarrow> y = x" by (auto dest: inj_onD)
5241    have **: "x = y \<longleftrightarrow> y = x" by auto
5242    show "(insort_key f y \<circ> insort_key f x) zs = (insort_key f x \<circ> insort_key f y) zs"
5243      by (induct zs) (auto intro: * simp add: **)
5244  qed
5245  then show ?thesis by (simp add: sort_key_def foldr_conv_fold)
5246qed
5247
5248lemma sort_conv_fold:
5249  "sort xs = fold insort xs []"
5250by (rule sort_key_conv_fold) simp
5251
5252lemma length_sort[simp]: "length (sort_key f xs) = length xs"
5253by (induct xs, auto)
5254
5255lemma set_sort[simp]: "set(sort_key f xs) = set xs"
5256by (induct xs) (simp_all add: set_insort_key)
5257
5258lemma distinct_insort: "distinct (insort_key f x xs) = (x \<notin> set xs \<and> distinct xs)"
5259by(induct xs)(auto simp: set_insort_key)
5260
5261lemma distinct_sort[simp]: "distinct (sort_key f xs) = distinct xs"
5262by (induct xs) (simp_all add: distinct_insort)
5263
5264lemma sorted_insort_key: "sorted (map f (insort_key f x xs)) = sorted (map f xs)"
5265by (induct xs) (auto simp: set_insort_key)
5266
5267lemma sorted_insort: "sorted (insort x xs) = sorted xs"
5268using sorted_insort_key [where f="\<lambda>x. x"] by simp
5269
5270theorem sorted_sort_key [simp]: "sorted (map f (sort_key f xs))"
5271by (induct xs) (auto simp:sorted_insort_key)
5272
5273theorem sorted_sort [simp]: "sorted (sort xs)"
5274using sorted_sort_key [where f="\<lambda>x. x"] by simp
5275
5276lemma insort_not_Nil [simp]:
5277  "insort_key f a xs \<noteq> []"
5278by (induction xs) simp_all
5279
5280lemma insort_is_Cons: "\<forall>x\<in>set xs. f a \<le> f x \<Longrightarrow> insort_key f a xs = a # xs"
5281by (cases xs) auto
5282
5283lemma sorted_sort_id: "sorted xs \<Longrightarrow> sort xs = xs"
5284by (induct xs) (auto simp add: insort_is_Cons)
5285
5286lemma insort_key_remove1:
5287  assumes "a \<in> set xs" and "sorted (map f xs)" and "hd (filter (\<lambda>x. f a = f x) xs) = a"
5288  shows "insort_key f a (remove1 a xs) = xs"
5289using assms proof (induct xs)
5290  case (Cons x xs)
5291  then show ?case
5292  proof (cases "x = a")
5293    case False
5294    then have "f x \<noteq> f a" using Cons.prems by auto
5295    then have "f x < f a" using Cons.prems by auto
5296    with \<open>f x \<noteq> f a\<close> show ?thesis using Cons by (auto simp: insort_is_Cons)
5297  qed (auto simp: insort_is_Cons)
5298qed simp
5299
5300lemma insort_remove1:
5301  assumes "a \<in> set xs" and "sorted xs"
5302  shows "insort a (remove1 a xs) = xs"
5303proof (rule insort_key_remove1)
5304  define n where "n = length (filter ((=) a) xs) - 1"
5305  from \<open>a \<in> set xs\<close> show "a \<in> set xs" .
5306  from \<open>sorted xs\<close> show "sorted (map (\<lambda>x. x) xs)" by simp
5307  from \<open>a \<in> set xs\<close> have "a \<in> set (filter ((=) a) xs)" by auto
5308  then have "set (filter ((=) a) xs) \<noteq> {}" by auto
5309  then have "filter ((=) a) xs \<noteq> []" by (auto simp only: set_empty)
5310  then have "length (filter ((=) a) xs) > 0" by simp
5311  then have n: "Suc n = length (filter ((=) a) xs)" by (simp add: n_def)
5312  moreover have "replicate (Suc n) a = a # replicate n a"
5313    by simp
5314  ultimately show "hd (filter ((=) a) xs) = a" by (simp add: replicate_length_filter)
5315qed
5316
5317lemma finite_sorted_distinct_unique:
5318shows "finite A \<Longrightarrow> \<exists>!xs. set xs = A \<and> sorted xs \<and> distinct xs"
5319apply(drule finite_distinct_list)
5320apply clarify
5321apply(rule_tac a="sort xs" in ex1I)
5322apply (auto simp: sorted_distinct_set_unique)
5323done
5324
5325lemma insort_insert_key_triv:
5326  "f x \<in> f ` set xs \<Longrightarrow> insort_insert_key f x xs = xs"
5327  by (simp add: insort_insert_key_def)
5328
5329lemma insort_insert_triv:
5330  "x \<in> set xs \<Longrightarrow> insort_insert x xs = xs"
5331  using insort_insert_key_triv [of "\<lambda>x. x"] by simp
5332
5333lemma insort_insert_insort_key:
5334  "f x \<notin> f ` set xs \<Longrightarrow> insort_insert_key f x xs = insort_key f x xs"
5335  by (simp add: insort_insert_key_def)
5336
5337lemma insort_insert_insort:
5338  "x \<notin> set xs \<Longrightarrow> insort_insert x xs = insort x xs"
5339  using insort_insert_insort_key [of "\<lambda>x. x"] by simp
5340
5341lemma set_insort_insert:
5342  "set (insort_insert x xs) = insert x (set xs)"
5343  by (auto simp add: insort_insert_key_def set_insort_key)
5344
5345lemma distinct_insort_insert:
5346  assumes "distinct xs"
5347  shows "distinct (insort_insert_key f x xs)"
5348using assms by (induct xs) (auto simp add: insort_insert_key_def set_insort_key)
5349
5350lemma sorted_insort_insert_key:
5351  assumes "sorted (map f xs)"
5352  shows "sorted (map f (insort_insert_key f x xs))"
5353  using assms by (simp add: insort_insert_key_def sorted_insort_key)
5354
5355lemma sorted_insort_insert:
5356  assumes "sorted xs"
5357  shows "sorted (insort_insert x xs)"
5358  using assms sorted_insort_insert_key [of "\<lambda>x. x"] by simp
5359
5360lemma filter_insort_triv:
5361  "\<not> P x \<Longrightarrow> filter P (insort_key f x xs) = filter P xs"
5362  by (induct xs) simp_all
5363
5364lemma filter_insort:
5365  "sorted (map f xs) \<Longrightarrow> P x \<Longrightarrow> filter P (insort_key f x xs) = insort_key f x (filter P xs)"
5366  by (induct xs) (auto, subst insort_is_Cons, auto)
5367
5368lemma filter_sort:
5369  "filter P (sort_key f xs) = sort_key f (filter P xs)"
5370  by (induct xs) (simp_all add: filter_insort_triv filter_insort)
5371
5372lemma remove1_insort [simp]:
5373  "remove1 x (insort x xs) = xs"
5374  by (induct xs) simp_all
5375
5376end
5377
5378lemma sort_upt [simp]: "sort [m..<n] = [m..<n]"
5379by (rule sorted_sort_id) simp
5380
5381lemma sort_upto [simp]: "sort [i..j] = [i..j]"
5382by (rule sorted_sort_id) simp
5383
5384lemma sorted_find_Min:
5385  "sorted xs \<Longrightarrow> \<exists>x \<in> set xs. P x \<Longrightarrow> List.find P xs = Some (Min {x\<in>set xs. P x})"
5386proof (induct xs)
5387  case Nil then show ?case by simp
5388next
5389  case (Cons x xs) show ?case proof (cases "P x")
5390    case True
5391    with Cons show ?thesis by (auto intro: Min_eqI [symmetric])
5392  next
5393    case False then have "{y. (y = x \<or> y \<in> set xs) \<and> P y} = {y \<in> set xs. P y}"
5394      by auto
5395    with Cons False show ?thesis by (simp_all)
5396  qed
5397qed
5398
5399lemma sorted_enumerate [simp]: "sorted (map fst (enumerate n xs))"
5400by (simp add: enumerate_eq_zip)
5401
5402text \<open>Stability of @{const sort_key}:\<close>
5403
5404lemma sort_key_stable: "filter (\<lambda>y. f y = k) (sort_key f xs) = filter (\<lambda>y. f y = k) xs"
5405proof (induction xs)
5406  case Nil thus ?case by simp
5407next
5408  case (Cons a xs)
5409  thus ?case
5410  proof (cases "f a = k")
5411    case False thus ?thesis
5412      using Cons.IH by (metis (mono_tags) filter.simps(2) filter_sort)
5413  next
5414    case True
5415    hence ler: "filter (\<lambda>y. f y = k) (a # xs) = a # filter (\<lambda>y. f y = f a) xs" by simp
5416    have "\<forall>y \<in> set (sort_key f (filter (\<lambda>y. f y = f a) xs)). f y = f a" by simp
5417    hence "insort_key f a (sort_key f (filter (\<lambda>y. f y = f a) xs))
5418            = a # (sort_key f (filter (\<lambda>y. f y = f a) xs))"
5419      by (simp add: insort_is_Cons)
5420    hence lel: "filter (\<lambda>y. f y = k) (sort_key f (a # xs)) = a # filter (\<lambda>y. f y = f a) (sort_key f xs)"
5421      by (metis True filter_sort ler sort_key_simps(2))
5422    from lel ler show ?thesis using Cons.IH True by (auto)
5423  qed
5424qed
5425
5426corollary stable_sort_key_sort_key: "stable_sort_key sort_key"
5427by(simp add: stable_sort_key_def sort_key_stable)
5428
5429lemma sort_key_const: "sort_key (\<lambda>x. c) xs = xs"
5430by (metis (mono_tags) filter_True sort_key_stable)
5431
5432
5433subsubsection \<open>@{const transpose} on sorted lists\<close>
5434
5435lemma sorted_transpose[simp]: "sorted (rev (map length (transpose xs)))"
5436by (auto simp: sorted_iff_nth_mono rev_nth nth_transpose
5437    length_filter_conv_card intro: card_mono)
5438
5439lemma transpose_max_length:
5440  "foldr (\<lambda>xs. max (length xs)) (transpose xs) 0 = length (filter (\<lambda>x. x \<noteq> []) xs)"
5441  (is "?L = ?R")
5442proof (cases "transpose xs = []")
5443  case False
5444  have "?L = foldr max (map length (transpose xs)) 0"
5445    by (simp add: foldr_map comp_def)
5446  also have "... = length (transpose xs ! 0)"
5447    using False sorted_transpose by (simp add: foldr_max_sorted)
5448  finally show ?thesis
5449    using False by (simp add: nth_transpose)
5450next
5451  case True
5452  hence "filter (\<lambda>x. x \<noteq> []) xs = []"
5453    by (auto intro!: filter_False simp: transpose_empty)
5454  thus ?thesis by (simp add: transpose_empty True)
5455qed
5456
5457lemma length_transpose_sorted:
5458  fixes xs :: "'a list list"
5459  assumes sorted: "sorted (rev (map length xs))"
5460  shows "length (transpose xs) = (if xs = [] then 0 else length (xs ! 0))"
5461proof (cases "xs = []")
5462  case False
5463  thus ?thesis
5464    using foldr_max_sorted[OF sorted] False
5465    unfolding length_transpose foldr_map comp_def
5466    by simp
5467qed simp
5468
5469lemma nth_nth_transpose_sorted[simp]:
5470  fixes xs :: "'a list list"
5471  assumes sorted: "sorted (rev (map length xs))"
5472  and i: "i < length (transpose xs)"
5473  and j: "j < length (filter (\<lambda>ys. i < length ys) xs)"
5474  shows "transpose xs ! i ! j = xs ! j  ! i"
5475using j filter_equals_takeWhile_sorted_rev[OF sorted, of i]
5476    nth_transpose[OF i] nth_map[OF j]
5477by (simp add: takeWhile_nth)
5478
5479lemma transpose_column_length:
5480  fixes xs :: "'a list list"
5481  assumes sorted: "sorted (rev (map length xs))" and "i < length xs"
5482  shows "length (filter (\<lambda>ys. i < length ys) (transpose xs)) = length (xs ! i)"
5483proof -
5484  have "xs \<noteq> []" using \<open>i < length xs\<close> by auto
5485  note filter_equals_takeWhile_sorted_rev[OF sorted, simp]
5486  { fix j assume "j \<le> i"
5487    note sorted_rev_nth_mono[OF sorted, of j i, simplified, OF this \<open>i < length xs\<close>]
5488  } note sortedE = this[consumes 1]
5489
5490  have "{j. j < length (transpose xs) \<and> i < length (transpose xs ! j)}
5491    = {..< length (xs ! i)}"
5492  proof safe
5493    fix j
5494    assume "j < length (transpose xs)" and "i < length (transpose xs ! j)"
5495    with this(2) nth_transpose[OF this(1)]
5496    have "i < length (takeWhile (\<lambda>ys. j < length ys) xs)" by simp
5497    from nth_mem[OF this] takeWhile_nth[OF this]
5498    show "j < length (xs ! i)" by (auto dest: set_takeWhileD)
5499  next
5500    fix j assume "j < length (xs ! i)"
5501    thus "j < length (transpose xs)"
5502      using foldr_max_sorted[OF sorted] \<open>xs \<noteq> []\<close> sortedE[OF le0]
5503      by (auto simp: length_transpose comp_def foldr_map)
5504
5505    have "Suc i \<le> length (takeWhile (\<lambda>ys. j < length ys) xs)"
5506      using \<open>i < length xs\<close> \<open>j < length (xs ! i)\<close> less_Suc_eq_le
5507      by (auto intro!: length_takeWhile_less_P_nth dest!: sortedE)
5508    with nth_transpose[OF \<open>j < length (transpose xs)\<close>]
5509    show "i < length (transpose xs ! j)" by simp
5510  qed
5511  thus ?thesis by (simp add: length_filter_conv_card)
5512qed
5513
5514lemma transpose_column:
5515  fixes xs :: "'a list list"
5516  assumes sorted: "sorted (rev (map length xs))" and "i < length xs"
5517  shows "map (\<lambda>ys. ys ! i) (filter (\<lambda>ys. i < length ys) (transpose xs))
5518    = xs ! i" (is "?R = _")
5519proof (rule nth_equalityI, safe)
5520  show length: "length ?R = length (xs ! i)"
5521    using transpose_column_length[OF assms] by simp
5522
5523  fix j assume j: "j < length ?R"
5524  note * = less_le_trans[OF this, unfolded length_map, OF length_filter_le]
5525  from j have j_less: "j < length (xs ! i)" using length by simp
5526  have i_less_tW: "Suc i \<le> length (takeWhile (\<lambda>ys. Suc j \<le> length ys) xs)"
5527  proof (rule length_takeWhile_less_P_nth)
5528    show "Suc i \<le> length xs" using \<open>i < length xs\<close> by simp
5529    fix k assume "k < Suc i"
5530    hence "k \<le> i" by auto
5531    with sorted_rev_nth_mono[OF sorted this] \<open>i < length xs\<close>
5532    have "length (xs ! i) \<le> length (xs ! k)" by simp
5533    thus "Suc j \<le> length (xs ! k)" using j_less by simp
5534  qed
5535  have i_less_filter: "i < length (filter (\<lambda>ys. j < length ys) xs) "
5536    unfolding filter_equals_takeWhile_sorted_rev[OF sorted, of j]
5537    using i_less_tW by (simp_all add: Suc_le_eq)
5538  from j show "?R ! j = xs ! i ! j"
5539    unfolding filter_equals_takeWhile_sorted_rev[OF sorted_transpose, of i]
5540    by (simp add: takeWhile_nth nth_nth_transpose_sorted[OF sorted * i_less_filter])
5541qed
5542
5543lemma transpose_transpose:
5544  fixes xs :: "'a list list"
5545  assumes sorted: "sorted (rev (map length xs))"
5546  shows "transpose (transpose xs) = takeWhile (\<lambda>x. x \<noteq> []) xs" (is "?L = ?R")
5547proof -
5548  have len: "length ?L = length ?R"
5549    unfolding length_transpose transpose_max_length
5550    using filter_equals_takeWhile_sorted_rev[OF sorted, of 0]
5551    by simp
5552
5553  { fix i assume "i < length ?R"
5554    with less_le_trans[OF _ length_takeWhile_le[of _ xs]]
5555    have "i < length xs" by simp
5556  } note * = this
5557  show ?thesis
5558    by (rule nth_equalityI)
5559       (simp_all add: len nth_transpose transpose_column[OF sorted] * takeWhile_nth)
5560qed
5561
5562theorem transpose_rectangle:
5563  assumes "xs = [] \<Longrightarrow> n = 0"
5564  assumes rect: "\<And> i. i < length xs \<Longrightarrow> length (xs ! i) = n"
5565  shows "transpose xs = map (\<lambda> i. map (\<lambda> j. xs ! j ! i) [0..<length xs]) [0..<n]"
5566    (is "?trans = ?map")
5567proof (rule nth_equalityI)
5568  have "sorted (rev (map length xs))"
5569    by (auto simp: rev_nth rect sorted_iff_nth_mono)
5570  from foldr_max_sorted[OF this] assms
5571  show len: "length ?trans = length ?map"
5572    by (simp_all add: length_transpose foldr_map comp_def)
5573  moreover
5574  { fix i assume "i < n" hence "filter (\<lambda>ys. i < length ys) xs = xs"
5575      using rect by (auto simp: in_set_conv_nth intro!: filter_True) }
5576  ultimately show "\<forall>i < length ?trans. ?trans ! i = ?map ! i"
5577    by (auto simp: nth_transpose intro: nth_equalityI)
5578qed
5579
5580
5581subsubsection \<open>\<open>sorted_list_of_set\<close>\<close>
5582
5583text\<open>This function maps (finite) linearly ordered sets to sorted
5584lists. Warning: in most cases it is not a good idea to convert from
5585sets to lists but one should convert in the other direction (via
5586@{const set}).\<close>
5587
5588context linorder
5589begin
5590
5591definition sorted_list_of_set :: "'a set \<Rightarrow> 'a list" where
5592  "sorted_list_of_set = folding.F insort []"
5593
5594sublocale sorted_list_of_set: folding insort Nil
5595rewrites
5596  "folding.F insort [] = sorted_list_of_set"
5597proof -
5598  interpret comp_fun_commute insort by (fact comp_fun_commute_insort)
5599  show "folding insort" by standard (fact comp_fun_commute)
5600  show "folding.F insort [] = sorted_list_of_set" by (simp only: sorted_list_of_set_def)
5601qed
5602
5603lemma sorted_list_of_set_empty:
5604  "sorted_list_of_set {} = []"
5605by (fact sorted_list_of_set.empty)
5606
5607lemma sorted_list_of_set_insert [simp]:
5608  "finite A \<Longrightarrow> sorted_list_of_set (insert x A) = insort x (sorted_list_of_set (A - {x}))"
5609by (fact sorted_list_of_set.insert_remove)
5610
5611lemma sorted_list_of_set_eq_Nil_iff [simp]:
5612  "finite A \<Longrightarrow> sorted_list_of_set A = [] \<longleftrightarrow> A = {}"
5613by (auto simp: sorted_list_of_set.remove)
5614
5615lemma set_sorted_list_of_set [simp]:
5616  "finite A \<Longrightarrow> set (sorted_list_of_set A) = A"
5617by(induct A rule: finite_induct) (simp_all add: set_insort_key)
5618
5619lemma sorted_sorted_list_of_set [simp]: "sorted (sorted_list_of_set A)"
5620proof (cases "finite A")
5621  case True thus ?thesis by(induction A) (simp_all add: sorted_insort)
5622next
5623  case False thus ?thesis by simp
5624qed
5625
5626lemma distinct_sorted_list_of_set [simp]: "distinct (sorted_list_of_set A)"
5627proof (cases "finite A")
5628  case True thus ?thesis by(induction A) (simp_all add: distinct_insort)
5629next
5630  case False thus ?thesis by simp
5631qed
5632
5633lemmas sorted_list_of_set = set_sorted_list_of_set sorted_sorted_list_of_set distinct_sorted_list_of_set
5634
5635lemma sorted_list_of_set_sort_remdups [code]:
5636  "sorted_list_of_set (set xs) = sort (remdups xs)"
5637proof -
5638  interpret comp_fun_commute insort by (fact comp_fun_commute_insort)
5639  show ?thesis by (simp add: sorted_list_of_set.eq_fold sort_conv_fold fold_set_fold_remdups)
5640qed
5641
5642lemma sorted_list_of_set_remove:
5643  assumes "finite A"
5644  shows "sorted_list_of_set (A - {x}) = remove1 x (sorted_list_of_set A)"
5645proof (cases "x \<in> A")
5646  case False with assms have "x \<notin> set (sorted_list_of_set A)" by simp
5647  with False show ?thesis by (simp add: remove1_idem)
5648next
5649  case True then obtain B where A: "A = insert x B" by (rule Set.set_insert)
5650  with assms show ?thesis by simp
5651qed
5652
5653end
5654
5655lemma sorted_list_of_set_range [simp]:
5656  "sorted_list_of_set {m..<n} = [m..<n]"
5657by (rule sorted_distinct_set_unique) simp_all
5658
5659
5660subsubsection \<open>\<open>lists\<close>: the list-forming operator over sets\<close>
5661
5662inductive_set
5663  lists :: "'a set => 'a list set"
5664  for A :: "'a set"
5665where
5666    Nil [intro!, simp]: "[] \<in> lists A"
5667  | Cons [intro!, simp]: "\<lbrakk>a \<in> A; l \<in> lists A\<rbrakk> \<Longrightarrow> a#l \<in> lists A"
5668
5669inductive_cases listsE [elim!]: "x#l \<in> lists A"
5670inductive_cases listspE [elim!]: "listsp A (x # l)"
5671
5672inductive_simps listsp_simps[code]:
5673  "listsp A []"
5674  "listsp A (x # xs)"
5675
5676lemma listsp_mono [mono]: "A \<le> B ==> listsp A \<le> listsp B"
5677by (rule predicate1I, erule listsp.induct, blast+)
5678
5679lemmas lists_mono = listsp_mono [to_set]
5680
5681lemma listsp_infI:
5682  assumes l: "listsp A l" shows "listsp B l ==> listsp (inf A B) l" using l
5683by induct blast+
5684
5685lemmas lists_IntI = listsp_infI [to_set]
5686
5687lemma listsp_inf_eq [simp]: "listsp (inf A B) = inf (listsp A) (listsp B)"
5688proof (rule mono_inf [where f=listsp, THEN order_antisym])
5689  show "mono listsp" by (simp add: mono_def listsp_mono)
5690  show "inf (listsp A) (listsp B) \<le> listsp (inf A B)" by (blast intro!: listsp_infI)
5691qed
5692
5693lemmas listsp_conj_eq [simp] = listsp_inf_eq [simplified inf_fun_def inf_bool_def]
5694
5695lemmas lists_Int_eq [simp] = listsp_inf_eq [to_set]
5696
5697lemma Cons_in_lists_iff[simp]: "x#xs \<in> lists A \<longleftrightarrow> x \<in> A \<and> xs \<in> lists A"
5698by auto
5699
5700lemma append_in_listsp_conv [iff]:
5701     "(listsp A (xs @ ys)) = (listsp A xs \<and> listsp A ys)"
5702by (induct xs) auto
5703
5704lemmas append_in_lists_conv [iff] = append_in_listsp_conv [to_set]
5705
5706lemma in_listsp_conv_set: "(listsp A xs) = (\<forall>x \<in> set xs. A x)"
5707\<comment> \<open>eliminate \<open>listsp\<close> in favour of \<open>set\<close>\<close>
5708by (induct xs) auto
5709
5710lemmas in_lists_conv_set [code_unfold] = in_listsp_conv_set [to_set]
5711
5712lemma in_listspD [dest!]: "listsp A xs ==> \<forall>x\<in>set xs. A x"
5713by (rule in_listsp_conv_set [THEN iffD1])
5714
5715lemmas in_listsD [dest!] = in_listspD [to_set]
5716
5717lemma in_listspI [intro!]: "\<forall>x\<in>set xs. A x ==> listsp A xs"
5718by (rule in_listsp_conv_set [THEN iffD2])
5719
5720lemmas in_listsI [intro!] = in_listspI [to_set]
5721
5722lemma lists_eq_set: "lists A = {xs. set xs <= A}"
5723by auto
5724
5725lemma lists_empty [simp]: "lists {} = {[]}"
5726by auto
5727
5728lemma lists_UNIV [simp]: "lists UNIV = UNIV"
5729by auto
5730
5731lemma lists_image: "lists (f`A) = map f ` lists A"
5732proof -
5733  { fix xs have "\<forall>x\<in>set xs. x \<in> f ` A \<Longrightarrow> xs \<in> map f ` lists A"
5734      by (induct xs) (auto simp del: list.map simp add: list.map[symmetric] intro!: imageI) }
5735  then show ?thesis by auto
5736qed
5737
5738subsubsection \<open>Inductive definition for membership\<close>
5739
5740inductive ListMem :: "'a \<Rightarrow> 'a list \<Rightarrow> bool"
5741where
5742    elem:  "ListMem x (x # xs)"
5743  | insert:  "ListMem x xs \<Longrightarrow> ListMem x (y # xs)"
5744
5745lemma ListMem_iff: "(ListMem x xs) = (x \<in> set xs)"
5746apply (rule iffI)
5747 apply (induct set: ListMem)
5748  apply auto
5749apply (induct xs)
5750 apply (auto intro: ListMem.intros)
5751done
5752
5753
5754subsubsection \<open>Lists as Cartesian products\<close>
5755
5756text\<open>\<open>set_Cons A Xs\<close>: the set of lists with head drawn from
5757@{term A} and tail drawn from @{term Xs}.\<close>
5758
5759definition set_Cons :: "'a set \<Rightarrow> 'a list set \<Rightarrow> 'a list set" where
5760"set_Cons A XS = {z. \<exists>x xs. z = x # xs \<and> x \<in> A \<and> xs \<in> XS}"
5761
5762lemma set_Cons_sing_Nil [simp]: "set_Cons A {[]} = (%x. [x])`A"
5763by (auto simp add: set_Cons_def)
5764
5765text\<open>Yields the set of lists, all of the same length as the argument and
5766with elements drawn from the corresponding element of the argument.\<close>
5767
5768primrec listset :: "'a set list \<Rightarrow> 'a list set" where
5769"listset [] = {[]}" |
5770"listset (A # As) = set_Cons A (listset As)"
5771
5772
5773subsection \<open>Relations on Lists\<close>
5774
5775subsubsection \<open>Length Lexicographic Ordering\<close>
5776
5777text\<open>These orderings preserve well-foundedness: shorter lists
5778  precede longer lists. These ordering are not used in dictionaries.\<close>
5779
5780primrec \<comment> \<open>The lexicographic ordering for lists of the specified length\<close>
5781  lexn :: "('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> ('a list \<times> 'a list) set" where
5782"lexn r 0 = {}" |
5783"lexn r (Suc n) =
5784  (map_prod (%(x, xs). x#xs) (%(x, xs). x#xs) ` (r <*lex*> lexn r n)) Int
5785  {(xs, ys). length xs = Suc n \<and> length ys = Suc n}"
5786
5787definition lex :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
5788"lex r = (\<Union>n. lexn r n)" \<comment> \<open>Holds only between lists of the same length\<close>
5789
5790definition lenlex :: "('a \<times> 'a) set => ('a list \<times> 'a list) set" where
5791"lenlex r = inv_image (less_than <*lex*> lex r) (\<lambda>xs. (length xs, xs))"
5792        \<comment> \<open>Compares lists by their length and then lexicographically\<close>
5793
5794lemma wf_lexn: "wf r ==> wf (lexn r n)"
5795apply (induct n, simp, simp)
5796apply(rule wf_subset)
5797 prefer 2 apply (rule Int_lower1)
5798apply(rule wf_map_prod_image)
5799 prefer 2 apply (rule inj_onI, auto)
5800done
5801
5802lemma lexn_length:
5803  "(xs, ys) \<in> lexn r n \<Longrightarrow> length xs = n \<and> length ys = n"
5804by (induct n arbitrary: xs ys) auto
5805
5806lemma wf_lex [intro!]: "wf r ==> wf (lex r)"
5807  apply (unfold lex_def)
5808  apply (rule wf_UN)
5809   apply (simp add: wf_lexn)
5810  apply (metis DomainE Int_emptyI RangeE lexn_length)
5811  done
5812
5813lemma lexn_conv:
5814  "lexn r n =
5815    {(xs,ys). length xs = n \<and> length ys = n \<and>
5816    (\<exists>xys x y xs' ys'. xs= xys @ x#xs' \<and> ys= xys @ y # ys' \<and> (x, y) \<in> r)}"
5817apply (induct n, simp)
5818apply (simp add: image_Collect lex_prod_def, safe, blast)
5819 apply (rule_tac x = "ab # xys" in exI, simp)
5820apply (case_tac xys, simp_all, blast)
5821done
5822
5823text\<open>By Mathias Fleury:\<close>
5824lemma lexn_transI:
5825  assumes "trans r" shows "trans (lexn r n)"
5826unfolding trans_def
5827proof (intro allI impI)
5828  fix as bs cs
5829  assume asbs: "(as, bs) \<in> lexn r n" and bscs: "(bs, cs) \<in> lexn r n"
5830  obtain abs a b as' bs' where
5831    n: "length as = n" and "length bs = n" and
5832    as: "as = abs @ a # as'" and
5833    bs: "bs = abs @ b # bs'" and
5834    abr: "(a, b) \<in> r"
5835    using asbs unfolding lexn_conv by blast
5836  obtain bcs b' c' cs' bs' where
5837    n': "length cs = n" and "length bs = n" and
5838    bs': "bs = bcs @ b' # bs'" and
5839    cs: "cs = bcs @ c' # cs'" and
5840    b'c'r: "(b', c') \<in> r"
5841    using bscs unfolding lexn_conv by blast
5842  consider (le) "length bcs < length abs"
5843    | (eq) "length bcs = length abs"
5844    | (ge) "length bcs > length abs" by linarith
5845  thus "(as, cs) \<in> lexn r n"
5846  proof cases
5847    let ?k = "length bcs"
5848    case le
5849    hence "as ! ?k = bs ! ?k" unfolding as bs by (simp add: nth_append)
5850    hence "(as ! ?k, cs ! ?k) \<in> r" using b'c'r unfolding bs' cs by auto
5851    moreover
5852    have "length bcs < length as" using le unfolding as by simp
5853    from id_take_nth_drop[OF this]
5854    have "as = take ?k as @ as ! ?k # drop (Suc ?k) as" .
5855    moreover
5856    have "length bcs < length cs" unfolding cs by simp
5857    from id_take_nth_drop[OF this]
5858    have "cs = take ?k cs @ cs ! ?k # drop (Suc ?k) cs" .
5859    moreover have "take ?k as = take ?k cs"
5860      using le arg_cong[OF bs, of "take (length bcs)"]
5861      unfolding cs as bs' by auto
5862    ultimately show ?thesis using n n' unfolding lexn_conv by auto
5863  next
5864    let ?k = "length abs"
5865    case ge
5866    hence "bs ! ?k = cs ! ?k" unfolding bs' cs by (simp add: nth_append)
5867    hence "(as ! ?k, cs ! ?k) \<in> r" using abr unfolding as bs by auto
5868    moreover
5869    have "length abs < length as" using ge unfolding as by simp
5870    from id_take_nth_drop[OF this]
5871    have "as = take ?k as @ as ! ?k # drop (Suc ?k) as" .
5872    moreover have "length abs < length cs" using n n' unfolding as by simp
5873    from id_take_nth_drop[OF this]
5874    have "cs = take ?k cs @ cs ! ?k # drop (Suc ?k) cs" .
5875    moreover have "take ?k as = take ?k cs"
5876      using ge arg_cong[OF bs', of "take (length abs)"]
5877      unfolding cs as bs by auto
5878    ultimately show ?thesis using n n' unfolding lexn_conv by auto
5879  next
5880    let ?k = "length abs"
5881    case eq
5882    hence *: "abs = bcs" "b = b'" using bs bs' by auto
5883    hence "(a, c') \<in> r"
5884      using abr b'c'r assms unfolding trans_def by blast
5885    with * show ?thesis using n n' unfolding lexn_conv as bs cs by auto
5886  qed
5887qed
5888
5889lemma lex_conv:
5890  "lex r =
5891    {(xs,ys). length xs = length ys \<and>
5892    (\<exists>xys x y xs' ys'. xs = xys @ x # xs' \<and> ys = xys @ y # ys' \<and> (x, y) \<in> r)}"
5893by (force simp add: lex_def lexn_conv)
5894
5895lemma wf_lenlex [intro!]: "wf r ==> wf (lenlex r)"
5896by (unfold lenlex_def) blast
5897
5898lemma lenlex_conv:
5899    "lenlex r = {(xs,ys). length xs < length ys \<or>
5900                 length xs = length ys \<and> (xs, ys) \<in> lex r}"
5901by (simp add: lenlex_def Id_on_def lex_prod_def inv_image_def)
5902
5903lemma Nil_notin_lex [iff]: "([], ys) \<notin> lex r"
5904by (simp add: lex_conv)
5905
5906lemma Nil2_notin_lex [iff]: "(xs, []) \<notin> lex r"
5907by (simp add:lex_conv)
5908
5909lemma Cons_in_lex [simp]:
5910    "((x # xs, y # ys) \<in> lex r) =
5911      ((x, y) \<in> r \<and> length xs = length ys \<or> x = y \<and> (xs, ys) \<in> lex r)"
5912apply (simp add: lex_conv)
5913apply (rule iffI)
5914 prefer 2 apply (blast intro: Cons_eq_appendI, clarify)
5915apply (case_tac xys, simp, simp)
5916apply blast
5917  done
5918
5919lemma lex_append_rightI:
5920  "(xs, ys) \<in> lex r \<Longrightarrow> length vs = length us \<Longrightarrow> (xs @ us, ys @ vs) \<in> lex r"
5921by (fastforce simp: lex_def lexn_conv)
5922
5923lemma lex_append_leftI:
5924  "(ys, zs) \<in> lex r \<Longrightarrow> (xs @ ys, xs @ zs) \<in> lex r"
5925by (induct xs) auto
5926
5927lemma lex_append_leftD:
5928  "\<forall>x. (x,x) \<notin> r \<Longrightarrow> (xs @ ys, xs @ zs) \<in> lex r \<Longrightarrow> (ys, zs) \<in> lex r"
5929by (induct xs) auto
5930
5931lemma lex_append_left_iff:
5932  "\<forall>x. (x,x) \<notin> r \<Longrightarrow> (xs @ ys, xs @ zs) \<in> lex r \<longleftrightarrow> (ys, zs) \<in> lex r"
5933by(metis lex_append_leftD lex_append_leftI)
5934
5935lemma lex_take_index:
5936  assumes "(xs, ys) \<in> lex r"
5937  obtains i where "i < length xs" and "i < length ys" and "take i xs =
5938take i ys"
5939    and "(xs ! i, ys ! i) \<in> r"
5940proof -
5941  obtain n us x xs' y ys' where "(xs, ys) \<in> lexn r n" and "length xs = n" and "length ys = n"
5942    and "xs = us @ x # xs'" and "ys = us @ y # ys'" and "(x, y) \<in> r"
5943    using assms by (fastforce simp: lex_def lexn_conv)
5944  then show ?thesis by (intro that [of "length us"]) auto
5945qed
5946
5947
5948subsubsection \<open>Lexicographic Ordering\<close>
5949
5950text \<open>Classical lexicographic ordering on lists, ie. "a" < "ab" < "b".
5951    This ordering does \emph{not} preserve well-foundedness.
5952     Author: N. Voelker, March 2005.\<close>
5953
5954definition lexord :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
5955"lexord r = {(x,y). \<exists> a v. y = x @ a # v \<or>
5956            (\<exists> u a b v w. (a,b) \<in> r \<and> x = u @ (a # v) \<and> y = u @ (b # w))}"
5957
5958lemma lexord_Nil_left[simp]:  "([],y) \<in> lexord r = (\<exists> a x. y = a # x)"
5959by (unfold lexord_def, induct_tac y, auto)
5960
5961lemma lexord_Nil_right[simp]: "(x,[]) \<notin> lexord r"
5962by (unfold lexord_def, induct_tac x, auto)
5963
5964lemma lexord_cons_cons[simp]:
5965     "((a # x, b # y) \<in> lexord r) = ((a,b)\<in> r \<or> (a = b \<and> (x,y)\<in> lexord r))"
5966  apply (unfold lexord_def, safe, simp_all)
5967  apply (case_tac u, simp, simp)
5968  apply (case_tac u, simp, clarsimp, blast, blast, clarsimp)
5969  apply (erule_tac x="b # u" in allE)
5970  by force
5971
5972lemmas lexord_simps = lexord_Nil_left lexord_Nil_right lexord_cons_cons
5973
5974lemma lexord_append_rightI: "\<exists> b z. y = b # z \<Longrightarrow> (x, x @ y) \<in> lexord r"
5975by (induct_tac x, auto)
5976
5977lemma lexord_append_left_rightI:
5978     "(a,b) \<in> r \<Longrightarrow> (u @ a # x, u @ b # y) \<in> lexord r"
5979by (induct_tac u, auto)
5980
5981lemma lexord_append_leftI: " (u,v) \<in> lexord r \<Longrightarrow> (x @ u, x @ v) \<in> lexord r"
5982by (induct x, auto)
5983
5984lemma lexord_append_leftD:
5985     "\<lbrakk> (x @ u, x @ v) \<in> lexord r; (\<forall>a. (a,a) \<notin> r) \<rbrakk> \<Longrightarrow> (u,v) \<in> lexord r"
5986by (erule rev_mp, induct_tac x, auto)
5987
5988lemma lexord_take_index_conv:
5989   "((x,y) \<in> lexord r) =
5990    ((length x < length y \<and> take (length x) y = x) \<or>
5991     (\<exists>i. i < min(length x)(length y) \<and> take i x = take i y \<and> (x!i,y!i) \<in> r))"
5992  apply (unfold lexord_def Let_def, clarsimp)
5993  apply (rule_tac f = "(% a b. a \<or> b)" in arg_cong2)
5994  apply auto
5995  apply (rule_tac x="hd (drop (length x) y)" in exI)
5996  apply (rule_tac x="tl (drop (length x) y)" in exI)
5997  apply (erule subst, simp add: min_def)
5998  apply (rule_tac x ="length u" in exI, simp)
5999  apply (rule_tac x ="take i x" in exI)
6000  apply (rule_tac x ="x ! i" in exI)
6001  apply (rule_tac x ="y ! i" in exI, safe)
6002  apply (rule_tac x="drop (Suc i) x" in exI)
6003  apply (drule sym, simp add: Cons_nth_drop_Suc)
6004  apply (rule_tac x="drop (Suc i) y" in exI)
6005  by (simp add: Cons_nth_drop_Suc)
6006
6007\<comment> \<open>lexord is extension of partial ordering List.lex\<close>
6008lemma lexord_lex: "(x,y) \<in> lex r = ((x,y) \<in> lexord r \<and> length x = length y)"
6009  apply (rule_tac x = y in spec)
6010  apply (induct_tac x, clarsimp)
6011  by (clarify, case_tac x, simp, force)
6012
6013lemma lexord_irreflexive: "\<forall>x. (x,x) \<notin> r \<Longrightarrow> (xs,xs) \<notin> lexord r"
6014by (induct xs) auto
6015
6016text\<open>By Ren\'e Thiemann:\<close>
6017lemma lexord_partial_trans:
6018  "(\<And>x y z. x \<in> set xs \<Longrightarrow> (x,y) \<in> r \<Longrightarrow> (y,z) \<in> r \<Longrightarrow> (x,z) \<in> r)
6019   \<Longrightarrow>  (xs,ys) \<in> lexord r  \<Longrightarrow>  (ys,zs) \<in> lexord r \<Longrightarrow>  (xs,zs) \<in> lexord r"
6020proof (induct xs arbitrary: ys zs)
6021  case Nil
6022  from Nil(3) show ?case unfolding lexord_def by (cases zs, auto)
6023next
6024  case (Cons x xs yys zzs)
6025  from Cons(3) obtain y ys where yys: "yys = y # ys" unfolding lexord_def
6026    by (cases yys, auto)
6027  note Cons = Cons[unfolded yys]
6028  from Cons(3) have one: "(x,y) \<in> r \<or> x = y \<and> (xs,ys) \<in> lexord r" by auto
6029  from Cons(4) obtain z zs where zzs: "zzs = z # zs" unfolding lexord_def
6030    by (cases zzs, auto)
6031  note Cons = Cons[unfolded zzs]
6032  from Cons(4) have two: "(y,z) \<in> r \<or> y = z \<and> (ys,zs) \<in> lexord r" by auto
6033  {
6034    assume "(xs,ys) \<in> lexord r" and "(ys,zs) \<in> lexord r"
6035    from Cons(1)[OF _ this] Cons(2)
6036    have "(xs,zs) \<in> lexord r" by auto
6037  } note ind1 = this
6038  {
6039    assume "(x,y) \<in> r" and "(y,z) \<in> r"
6040    from Cons(2)[OF _ this] have "(x,z) \<in> r" by auto
6041  } note ind2 = this
6042  from one two ind1 ind2
6043  have "(x,z) \<in> r \<or> x = z \<and> (xs,zs) \<in> lexord r" by blast
6044  thus ?case unfolding zzs by auto
6045qed
6046
6047lemma lexord_trans:
6048    "\<lbrakk> (x, y) \<in> lexord r; (y, z) \<in> lexord r; trans r \<rbrakk> \<Longrightarrow> (x, z) \<in> lexord r"
6049by(auto simp: trans_def intro:lexord_partial_trans)
6050
6051lemma lexord_transI:  "trans r \<Longrightarrow> trans (lexord r)"
6052by (rule transI, drule lexord_trans, blast)
6053
6054lemma lexord_linear: "(\<forall>a b. (a,b)\<in> r \<or> a = b \<or> (b,a) \<in> r) \<Longrightarrow> (x,y) \<in> lexord r \<or> x = y \<or> (y,x) \<in> lexord r"
6055  apply (rule_tac x = y in spec)
6056  apply (induct_tac x, rule allI)
6057  apply (case_tac x, simp, simp)
6058  apply (rule allI, case_tac x, simp, simp)
6059  by blast
6060
6061lemma lexord_irrefl:
6062  "irrefl R \<Longrightarrow> irrefl (lexord R)"
6063by (simp add: irrefl_def lexord_irreflexive)
6064
6065lemma lexord_asym:
6066  assumes "asym R"
6067  shows "asym (lexord R)"
6068proof
6069  from assms obtain "irrefl R" by (blast elim: asym.cases)
6070  then show "irrefl (lexord R)" by (rule lexord_irrefl)
6071next
6072  fix xs ys
6073  assume "(xs, ys) \<in> lexord R"
6074  then show "(ys, xs) \<notin> lexord R"
6075  proof (induct xs arbitrary: ys)
6076    case Nil
6077    then show ?case by simp
6078  next
6079    case (Cons x xs)
6080    then obtain z zs where ys: "ys = z # zs" by (cases ys) auto
6081    with assms Cons show ?case by (auto elim: asym.cases)
6082  qed
6083qed
6084
6085lemma lexord_asymmetric:
6086  assumes "asym R"
6087  assumes hyp: "(a, b) \<in> lexord R"
6088  shows "(b, a) \<notin> lexord R"
6089proof -
6090  from \<open>asym R\<close> have "asym (lexord R)" by (rule lexord_asym)
6091  then show ?thesis by (rule asym.cases) (auto simp add: hyp)
6092qed
6093
6094
6095text \<open>
6096  Predicate version of lexicographic order integrated with Isabelle's order type classes.
6097  Author: Andreas Lochbihler
6098\<close>
6099
6100context ord
6101begin
6102
6103context
6104  notes [[inductive_internals]]
6105begin
6106
6107inductive lexordp :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
6108where
6109  Nil: "lexordp [] (y # ys)"
6110| Cons: "x < y \<Longrightarrow> lexordp (x # xs) (y # ys)"
6111| Cons_eq:
6112  "\<lbrakk> \<not> x < y; \<not> y < x; lexordp xs ys \<rbrakk> \<Longrightarrow> lexordp (x # xs) (y # ys)"
6113
6114end
6115
6116lemma lexordp_simps [simp]:
6117  "lexordp [] ys = (ys \<noteq> [])"
6118  "lexordp xs [] = False"
6119  "lexordp (x # xs) (y # ys) \<longleftrightarrow> x < y \<or> \<not> y < x \<and> lexordp xs ys"
6120by(subst lexordp.simps, fastforce simp add: neq_Nil_conv)+
6121
6122inductive lexordp_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
6123  Nil: "lexordp_eq [] ys"
6124| Cons: "x < y \<Longrightarrow> lexordp_eq (x # xs) (y # ys)"
6125| Cons_eq: "\<lbrakk> \<not> x < y; \<not> y < x; lexordp_eq xs ys \<rbrakk> \<Longrightarrow> lexordp_eq (x # xs) (y # ys)"
6126
6127lemma lexordp_eq_simps [simp]:
6128  "lexordp_eq [] ys = True"
6129  "lexordp_eq xs [] \<longleftrightarrow> xs = []"
6130  "lexordp_eq (x # xs) [] = False"
6131  "lexordp_eq (x # xs) (y # ys) \<longleftrightarrow> x < y \<or> \<not> y < x \<and> lexordp_eq xs ys"
6132by(subst lexordp_eq.simps, fastforce)+
6133
6134lemma lexordp_append_rightI: "ys \<noteq> Nil \<Longrightarrow> lexordp xs (xs @ ys)"
6135by(induct xs)(auto simp add: neq_Nil_conv)
6136
6137lemma lexordp_append_left_rightI: "x < y \<Longrightarrow> lexordp (us @ x # xs) (us @ y # ys)"
6138by(induct us) auto
6139
6140lemma lexordp_eq_refl: "lexordp_eq xs xs"
6141by(induct xs) simp_all
6142
6143lemma lexordp_append_leftI: "lexordp us vs \<Longrightarrow> lexordp (xs @ us) (xs @ vs)"
6144by(induct xs) auto
6145
6146lemma lexordp_append_leftD: "\<lbrakk> lexordp (xs @ us) (xs @ vs); \<forall>a. \<not> a < a \<rbrakk> \<Longrightarrow> lexordp us vs"
6147by(induct xs) auto
6148
6149lemma lexordp_irreflexive:
6150  assumes irrefl: "\<forall>x. \<not> x < x"
6151  shows "\<not> lexordp xs xs"
6152proof
6153  assume "lexordp xs xs"
6154  thus False by(induct xs ys\<equiv>xs)(simp_all add: irrefl)
6155qed
6156
6157lemma lexordp_into_lexordp_eq:
6158  assumes "lexordp xs ys"
6159  shows "lexordp_eq xs ys"
6160using assms by induct simp_all
6161
6162end
6163
6164declare ord.lexordp_simps [simp, code]
6165declare ord.lexordp_eq_simps [code, simp]
6166
6167lemma lexord_code [code, code_unfold]: "lexordp = ord.lexordp less"
6168unfolding lexordp_def ord.lexordp_def ..
6169
6170context order
6171begin
6172
6173lemma lexordp_antisym:
6174  assumes "lexordp xs ys" "lexordp ys xs"
6175  shows False
6176using assms by induct auto
6177
6178lemma lexordp_irreflexive': "\<not> lexordp xs xs"
6179by(rule lexordp_irreflexive) simp
6180
6181end
6182
6183context linorder begin
6184
6185lemma lexordp_cases [consumes 1, case_names Nil Cons Cons_eq, cases pred: lexordp]:
6186  assumes "lexordp xs ys"
6187  obtains (Nil) y ys' where "xs = []" "ys = y # ys'"
6188  | (Cons) x xs' y ys' where "xs = x # xs'" "ys = y # ys'" "x < y"
6189  | (Cons_eq) x xs' ys' where "xs = x # xs'" "ys = x # ys'" "lexordp xs' ys'"
6190using assms by cases (fastforce simp add: not_less_iff_gr_or_eq)+
6191
6192lemma lexordp_induct [consumes 1, case_names Nil Cons Cons_eq, induct pred: lexordp]:
6193  assumes major: "lexordp xs ys"
6194  and Nil: "\<And>y ys. P [] (y # ys)"
6195  and Cons: "\<And>x xs y ys. x < y \<Longrightarrow> P (x # xs) (y # ys)"
6196  and Cons_eq: "\<And>x xs ys. \<lbrakk> lexordp xs ys; P xs ys \<rbrakk> \<Longrightarrow> P (x # xs) (x # ys)"
6197  shows "P xs ys"
6198using major by induct (simp_all add: Nil Cons not_less_iff_gr_or_eq Cons_eq)
6199
6200lemma lexordp_iff:
6201  "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)"
6202  (is "?lhs = ?rhs")
6203proof
6204  assume ?lhs thus ?rhs
6205  proof induct
6206    case Cons_eq thus ?case by simp (metis append.simps(2))
6207  qed(fastforce intro: disjI2 del: disjCI intro: exI[where x="[]"])+
6208next
6209  assume ?rhs thus ?lhs
6210    by(auto intro: lexordp_append_leftI[where us="[]", simplified] lexordp_append_leftI)
6211qed
6212
6213lemma lexordp_conv_lexord:
6214  "lexordp xs ys \<longleftrightarrow> (xs, ys) \<in> lexord {(x, y). x < y}"
6215by(simp add: lexordp_iff lexord_def)
6216
6217lemma lexordp_eq_antisym:
6218  assumes "lexordp_eq xs ys" "lexordp_eq ys xs"
6219  shows "xs = ys"
6220using assms by induct simp_all
6221
6222lemma lexordp_eq_trans:
6223  assumes "lexordp_eq xs ys" and "lexordp_eq ys zs"
6224  shows "lexordp_eq xs zs"
6225using assms
6226apply(induct arbitrary: zs)
6227apply(case_tac [2-3] zs)
6228apply auto
6229done
6230
6231lemma lexordp_trans:
6232  assumes "lexordp xs ys" "lexordp ys zs"
6233  shows "lexordp xs zs"
6234using assms
6235apply(induct arbitrary: zs)
6236apply(case_tac [2-3] zs)
6237apply auto
6238done
6239
6240lemma lexordp_linear: "lexordp xs ys \<or> xs = ys \<or> lexordp ys xs"
6241proof(induct xs arbitrary: ys)
6242  case Nil thus ?case by(cases ys) simp_all
6243next
6244  case Cons thus ?case by(cases ys) auto
6245qed
6246
6247lemma lexordp_conv_lexordp_eq: "lexordp xs ys \<longleftrightarrow> lexordp_eq xs ys \<and> \<not> lexordp_eq ys xs"
6248  (is "?lhs \<longleftrightarrow> ?rhs")
6249proof
6250  assume ?lhs
6251  hence "\<not> lexordp_eq ys xs" by induct simp_all
6252  with \<open>?lhs\<close> show ?rhs by (simp add: lexordp_into_lexordp_eq)
6253next
6254  assume ?rhs
6255  hence "lexordp_eq xs ys" "\<not> lexordp_eq ys xs" by simp_all
6256  thus ?lhs by induct simp_all
6257qed
6258
6259lemma lexordp_eq_conv_lexord: "lexordp_eq xs ys \<longleftrightarrow> xs = ys \<or> lexordp xs ys"
6260by(auto simp add: lexordp_conv_lexordp_eq lexordp_eq_refl dest: lexordp_eq_antisym)
6261
6262lemma lexordp_eq_linear: "lexordp_eq xs ys \<or> lexordp_eq ys xs"
6263apply(induct xs arbitrary: ys)
6264apply(case_tac [!] ys)
6265apply auto
6266done
6267
6268lemma lexordp_linorder: "class.linorder lexordp_eq lexordp"
6269by 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)
6270
6271end
6272
6273lemma sorted_insort_is_snoc: "sorted xs \<Longrightarrow> \<forall>x \<in> set xs. a \<ge> x \<Longrightarrow> insort a xs = xs @ [a]"
6274 by (induct xs) (auto dest!: insort_is_Cons)
6275
6276
6277subsubsection \<open>Lexicographic combination of measure functions\<close>
6278
6279text \<open>These are useful for termination proofs\<close>
6280
6281definition "measures fs = inv_image (lex less_than) (%a. map (%f. f a) fs)"
6282
6283lemma wf_measures[simp]: "wf (measures fs)"
6284unfolding measures_def
6285by blast
6286
6287lemma in_measures[simp]:
6288  "(x, y) \<in> measures [] = False"
6289  "(x, y) \<in> measures (f # fs)
6290         = (f x < f y \<or> (f x = f y \<and> (x, y) \<in> measures fs))"
6291unfolding measures_def
6292by auto
6293
6294lemma measures_less: "f x < f y ==> (x, y) \<in> measures (f#fs)"
6295by simp
6296
6297lemma measures_lesseq: "f x <= f y ==> (x, y) \<in> measures fs ==> (x, y) \<in> measures (f#fs)"
6298by auto
6299
6300
6301subsubsection \<open>Lifting Relations to Lists: one element\<close>
6302
6303definition listrel1 :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
6304"listrel1 r = {(xs,ys).
6305   \<exists>us z z' vs. xs = us @ z # vs \<and> (z,z') \<in> r \<and> ys = us @ z' # vs}"
6306
6307lemma listrel1I:
6308  "\<lbrakk> (x, y) \<in> r;  xs = us @ x # vs;  ys = us @ y # vs \<rbrakk> \<Longrightarrow>
6309  (xs, ys) \<in> listrel1 r"
6310unfolding listrel1_def by auto
6311
6312lemma listrel1E:
6313  "\<lbrakk> (xs, ys) \<in> listrel1 r;
6314     !!x y us vs. \<lbrakk> (x, y) \<in> r;  xs = us @ x # vs;  ys = us @ y # vs \<rbrakk> \<Longrightarrow> P
6315   \<rbrakk> \<Longrightarrow> P"
6316unfolding listrel1_def by auto
6317
6318lemma not_Nil_listrel1 [iff]: "([], xs) \<notin> listrel1 r"
6319unfolding listrel1_def by blast
6320
6321lemma not_listrel1_Nil [iff]: "(xs, []) \<notin> listrel1 r"
6322unfolding listrel1_def by blast
6323
6324lemma Cons_listrel1_Cons [iff]:
6325  "(x # xs, y # ys) \<in> listrel1 r \<longleftrightarrow>
6326   (x,y) \<in> r \<and> xs = ys \<or> x = y \<and> (xs, ys) \<in> listrel1 r"
6327by (simp add: listrel1_def Cons_eq_append_conv) (blast)
6328
6329lemma listrel1I1: "(x,y) \<in> r \<Longrightarrow> (x # xs, y # xs) \<in> listrel1 r"
6330by fast
6331
6332lemma listrel1I2: "(xs, ys) \<in> listrel1 r \<Longrightarrow> (x # xs, x # ys) \<in> listrel1 r"
6333by fast
6334
6335lemma append_listrel1I:
6336  "(xs, ys) \<in> listrel1 r \<and> us = vs \<or> xs = ys \<and> (us, vs) \<in> listrel1 r
6337    \<Longrightarrow> (xs @ us, ys @ vs) \<in> listrel1 r"
6338unfolding listrel1_def
6339by auto (blast intro: append_eq_appendI)+
6340
6341lemma Cons_listrel1E1[elim!]:
6342  assumes "(x # xs, ys) \<in> listrel1 r"
6343    and "\<And>y. ys = y # xs \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> R"
6344    and "\<And>zs. ys = x # zs \<Longrightarrow> (xs, zs) \<in> listrel1 r \<Longrightarrow> R"
6345  shows R
6346using assms by (cases ys) blast+
6347
6348lemma Cons_listrel1E2[elim!]:
6349  assumes "(xs, y # ys) \<in> listrel1 r"
6350    and "\<And>x. xs = x # ys \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> R"
6351    and "\<And>zs. xs = y # zs \<Longrightarrow> (zs, ys) \<in> listrel1 r \<Longrightarrow> R"
6352  shows R
6353using assms by (cases xs) blast+
6354
6355lemma snoc_listrel1_snoc_iff:
6356  "(xs @ [x], ys @ [y]) \<in> listrel1 r
6357    \<longleftrightarrow> (xs, ys) \<in> listrel1 r \<and> x = y \<or> xs = ys \<and> (x,y) \<in> r" (is "?L \<longleftrightarrow> ?R")
6358proof
6359  assume ?L thus ?R
6360    by (fastforce simp: listrel1_def snoc_eq_iff_butlast butlast_append)
6361next
6362  assume ?R then show ?L unfolding listrel1_def by force
6363qed
6364
6365lemma listrel1_eq_len: "(xs,ys) \<in> listrel1 r \<Longrightarrow> length xs = length ys"
6366unfolding listrel1_def by auto
6367
6368lemma listrel1_mono:
6369  "r \<subseteq> s \<Longrightarrow> listrel1 r \<subseteq> listrel1 s"
6370unfolding listrel1_def by blast
6371
6372
6373lemma listrel1_converse: "listrel1 (r\<inverse>) = (listrel1 r)\<inverse>"
6374unfolding listrel1_def by blast
6375
6376lemma in_listrel1_converse:
6377  "(x,y) \<in> listrel1 (r\<inverse>) \<longleftrightarrow> (x,y) \<in> (listrel1 r)\<inverse>"
6378unfolding listrel1_def by blast
6379
6380lemma listrel1_iff_update:
6381  "(xs,ys) \<in> (listrel1 r)
6382   \<longleftrightarrow> (\<exists>y n. (xs ! n, y) \<in> r \<and> n < length xs \<and> ys = xs[n:=y])" (is "?L \<longleftrightarrow> ?R")
6383proof
6384  assume "?L"
6385  then obtain x y u v where "xs = u @ x # v"  "ys = u @ y # v"  "(x,y) \<in> r"
6386    unfolding listrel1_def by auto
6387  then have "ys = xs[length u := y]" and "length u < length xs"
6388    and "(xs ! length u, y) \<in> r" by auto
6389  then show "?R" by auto
6390next
6391  assume "?R"
6392  then obtain x y n where "(xs!n, y) \<in> r" "n < size xs" "ys = xs[n:=y]" "x = xs!n"
6393    by auto
6394  then obtain u v where "xs = u @ x # v" and "ys = u @ y # v" and "(x, y) \<in> r"
6395    by (auto intro: upd_conv_take_nth_drop id_take_nth_drop)
6396  then show "?L" by (auto simp: listrel1_def)
6397qed
6398
6399
6400text\<open>Accessible part and wellfoundedness:\<close>
6401
6402lemma Cons_acc_listrel1I [intro!]:
6403  "x \<in> Wellfounded.acc r \<Longrightarrow> xs \<in> Wellfounded.acc (listrel1 r) \<Longrightarrow> (x # xs) \<in> Wellfounded.acc (listrel1 r)"
6404apply (induct arbitrary: xs set: Wellfounded.acc)
6405apply (erule thin_rl)
6406apply (erule acc_induct)
6407apply (rule accI)
6408apply (blast)
6409done
6410
6411lemma lists_accD: "xs \<in> lists (Wellfounded.acc r) \<Longrightarrow> xs \<in> Wellfounded.acc (listrel1 r)"
6412apply (induct set: lists)
6413 apply (rule accI)
6414 apply simp
6415apply (rule accI)
6416apply (fast dest: acc_downward)
6417done
6418
6419lemma lists_accI: "xs \<in> Wellfounded.acc (listrel1 r) \<Longrightarrow> xs \<in> lists (Wellfounded.acc r)"
6420apply (induct set: Wellfounded.acc)
6421apply clarify
6422apply (rule accI)
6423apply (fastforce dest!: in_set_conv_decomp[THEN iffD1] simp: listrel1_def)
6424done
6425
6426lemma wf_listrel1_iff[simp]: "wf(listrel1 r) = wf r"
6427by (auto simp: wf_acc_iff
6428      intro: lists_accD lists_accI[THEN Cons_in_lists_iff[THEN iffD1, THEN conjunct1]])
6429
6430subsubsection \<open>Lifting Relations to Lists: all elements\<close>
6431
6432inductive_set
6433  listrel :: "('a \<times> 'b) set \<Rightarrow> ('a list \<times> 'b list) set"
6434  for r :: "('a \<times> 'b) set"
6435where
6436    Nil:  "([],[]) \<in> listrel r"
6437  | Cons: "[| (x,y) \<in> r; (xs,ys) \<in> listrel r |] ==> (x#xs, y#ys) \<in> listrel r"
6438
6439inductive_cases listrel_Nil1 [elim!]: "([],xs) \<in> listrel r"
6440inductive_cases listrel_Nil2 [elim!]: "(xs,[]) \<in> listrel r"
6441inductive_cases listrel_Cons1 [elim!]: "(y#ys,xs) \<in> listrel r"
6442inductive_cases listrel_Cons2 [elim!]: "(xs,y#ys) \<in> listrel r"
6443
6444
6445lemma listrel_eq_len:  "(xs, ys) \<in> listrel r \<Longrightarrow> length xs = length ys"
6446by(induct rule: listrel.induct) auto
6447
6448lemma listrel_iff_zip [code_unfold]: "(xs,ys) \<in> listrel r \<longleftrightarrow>
6449  length xs = length ys \<and> (\<forall>(x,y) \<in> set(zip xs ys). (x,y) \<in> r)" (is "?L \<longleftrightarrow> ?R")
6450proof
6451  assume ?L thus ?R by induct (auto intro: listrel_eq_len)
6452next
6453  assume ?R thus ?L
6454    apply (clarify)
6455    by (induct rule: list_induct2) (auto intro: listrel.intros)
6456qed
6457
6458lemma listrel_iff_nth: "(xs,ys) \<in> listrel r \<longleftrightarrow>
6459  length xs = length ys \<and> (\<forall>n < length xs. (xs!n, ys!n) \<in> r)" (is "?L \<longleftrightarrow> ?R")
6460by (auto simp add: all_set_conv_all_nth listrel_iff_zip)
6461
6462
6463lemma listrel_mono: "r \<subseteq> s \<Longrightarrow> listrel r \<subseteq> listrel s"
6464apply clarify
6465apply (erule listrel.induct)
6466apply (blast intro: listrel.intros)+
6467done
6468
6469lemma listrel_subset: "r \<subseteq> A \<times> A \<Longrightarrow> listrel r \<subseteq> lists A \<times> lists A"
6470apply clarify
6471apply (erule listrel.induct, auto)
6472done
6473
6474lemma listrel_refl_on: "refl_on A r \<Longrightarrow> refl_on (lists A) (listrel r)"
6475apply (simp add: refl_on_def listrel_subset Ball_def)
6476apply (rule allI)
6477apply (induct_tac x)
6478apply (auto intro: listrel.intros)
6479done
6480
6481lemma listrel_sym: "sym r \<Longrightarrow> sym (listrel r)"
6482apply (auto simp add: sym_def)
6483apply (erule listrel.induct)
6484apply (blast intro: listrel.intros)+
6485done
6486
6487lemma listrel_trans: "trans r \<Longrightarrow> trans (listrel r)"
6488apply (simp add: trans_def)
6489apply (intro allI)
6490apply (rule impI)
6491apply (erule listrel.induct)
6492apply (blast intro: listrel.intros)+
6493done
6494
6495theorem equiv_listrel: "equiv A r \<Longrightarrow> equiv (lists A) (listrel r)"
6496by (simp add: equiv_def listrel_refl_on listrel_sym listrel_trans)
6497
6498lemma listrel_rtrancl_refl[iff]: "(xs,xs) \<in> listrel(r\<^sup>*)"
6499using listrel_refl_on[of UNIV, OF refl_rtrancl]
6500by(auto simp: refl_on_def)
6501
6502lemma listrel_rtrancl_trans:
6503  "\<lbrakk>(xs,ys) \<in> listrel(r\<^sup>*);  (ys,zs) \<in> listrel(r\<^sup>*)\<rbrakk>
6504  \<Longrightarrow> (xs,zs) \<in> listrel(r\<^sup>*)"
6505by (metis listrel_trans trans_def trans_rtrancl)
6506
6507
6508lemma listrel_Nil [simp]: "listrel r `` {[]} = {[]}"
6509by (blast intro: listrel.intros)
6510
6511lemma listrel_Cons:
6512     "listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})"
6513by (auto simp add: set_Cons_def intro: listrel.intros)
6514
6515text \<open>Relating @{term listrel1}, @{term listrel} and closures:\<close>
6516
6517lemma listrel1_rtrancl_subset_rtrancl_listrel1:
6518  "listrel1 (r\<^sup>*) \<subseteq> (listrel1 r)\<^sup>*"
6519proof (rule subrelI)
6520  fix xs ys assume 1: "(xs,ys) \<in> listrel1 (r\<^sup>*)"
6521  { fix x y us vs
6522    have "(x,y) \<in> r\<^sup>* \<Longrightarrow> (us @ x # vs, us @ y # vs) \<in> (listrel1 r)\<^sup>*"
6523    proof(induct rule: rtrancl.induct)
6524      case rtrancl_refl show ?case by simp
6525    next
6526      case rtrancl_into_rtrancl thus ?case
6527        by (metis listrel1I rtrancl.rtrancl_into_rtrancl)
6528    qed }
6529  thus "(xs,ys) \<in> (listrel1 r)\<^sup>*" using 1 by(blast elim: listrel1E)
6530qed
6531
6532lemma rtrancl_listrel1_eq_len: "(x,y) \<in> (listrel1 r)\<^sup>* \<Longrightarrow> length x = length y"
6533by (induct rule: rtrancl.induct) (auto intro: listrel1_eq_len)
6534
6535lemma rtrancl_listrel1_ConsI1:
6536  "(xs,ys) \<in> (listrel1 r)\<^sup>* \<Longrightarrow> (x#xs,x#ys) \<in> (listrel1 r)\<^sup>*"
6537apply(induct rule: rtrancl.induct)
6538 apply simp
6539by (metis listrel1I2 rtrancl.rtrancl_into_rtrancl)
6540
6541lemma rtrancl_listrel1_ConsI2:
6542  "(x,y) \<in> r\<^sup>* \<Longrightarrow> (xs, ys) \<in> (listrel1 r)\<^sup>*
6543  \<Longrightarrow> (x # xs, y # ys) \<in> (listrel1 r)\<^sup>*"
6544  by (blast intro: rtrancl_trans rtrancl_listrel1_ConsI1
6545    subsetD[OF listrel1_rtrancl_subset_rtrancl_listrel1 listrel1I1])
6546
6547lemma listrel1_subset_listrel:
6548  "r \<subseteq> r' \<Longrightarrow> refl r' \<Longrightarrow> listrel1 r \<subseteq> listrel(r')"
6549by(auto elim!: listrel1E simp add: listrel_iff_zip set_zip refl_on_def)
6550
6551lemma listrel_reflcl_if_listrel1:
6552  "(xs,ys) \<in> listrel1 r \<Longrightarrow> (xs,ys) \<in> listrel(r\<^sup>*)"
6553by(erule listrel1E)(auto simp add: listrel_iff_zip set_zip)
6554
6555lemma listrel_rtrancl_eq_rtrancl_listrel1: "listrel (r\<^sup>*) = (listrel1 r)\<^sup>*"
6556proof
6557  { fix x y assume "(x,y) \<in> listrel (r\<^sup>*)"
6558    then have "(x,y) \<in> (listrel1 r)\<^sup>*"
6559    by induct (auto intro: rtrancl_listrel1_ConsI2) }
6560  then show "listrel (r\<^sup>*) \<subseteq> (listrel1 r)\<^sup>*"
6561    by (rule subrelI)
6562next
6563  show "listrel (r\<^sup>*) \<supseteq> (listrel1 r)\<^sup>*"
6564  proof(rule subrelI)
6565    fix xs ys assume "(xs,ys) \<in> (listrel1 r)\<^sup>*"
6566    then show "(xs,ys) \<in> listrel (r\<^sup>*)"
6567    proof induct
6568      case base show ?case by(auto simp add: listrel_iff_zip set_zip)
6569    next
6570      case (step ys zs)
6571      thus ?case by (metis listrel_reflcl_if_listrel1 listrel_rtrancl_trans)
6572    qed
6573  qed
6574qed
6575
6576lemma rtrancl_listrel1_if_listrel:
6577  "(xs,ys) \<in> listrel r \<Longrightarrow> (xs,ys) \<in> (listrel1 r)\<^sup>*"
6578by(metis listrel_rtrancl_eq_rtrancl_listrel1 subsetD[OF listrel_mono] r_into_rtrancl subsetI)
6579
6580lemma listrel_subset_rtrancl_listrel1: "listrel r \<subseteq> (listrel1 r)\<^sup>*"
6581by(fast intro:rtrancl_listrel1_if_listrel)
6582
6583
6584subsection \<open>Size function\<close>
6585
6586lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (size_list f)"
6587by (rule is_measure_trivial)
6588
6589lemma [measure_function]: "is_measure f \<Longrightarrow> is_measure (size_option f)"
6590by (rule is_measure_trivial)
6591
6592lemma size_list_estimation[termination_simp]:
6593  "x \<in> set xs \<Longrightarrow> y < f x \<Longrightarrow> y < size_list f xs"
6594by (induct xs) auto
6595
6596lemma size_list_estimation'[termination_simp]:
6597  "x \<in> set xs \<Longrightarrow> y \<le> f x \<Longrightarrow> y \<le> size_list f xs"
6598by (induct xs) auto
6599
6600lemma size_list_map[simp]: "size_list f (map g xs) = size_list (f \<circ> g) xs"
6601by (induct xs) auto
6602
6603lemma size_list_append[simp]: "size_list f (xs @ ys) = size_list f xs + size_list f ys"
6604by (induct xs, auto)
6605
6606lemma size_list_pointwise[termination_simp]:
6607  "(\<And>x. x \<in> set xs \<Longrightarrow> f x \<le> g x) \<Longrightarrow> size_list f xs \<le> size_list g xs"
6608by (induct xs) force+
6609
6610
6611subsection \<open>Monad operation\<close>
6612
6613definition bind :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
6614"bind xs f = concat (map f xs)"
6615
6616hide_const (open) bind
6617
6618lemma bind_simps [simp]:
6619  "List.bind [] f = []"
6620  "List.bind (x # xs) f = f x @ List.bind xs f"
6621  by (simp_all add: bind_def)
6622
6623lemma list_bind_cong [fundef_cong]:
6624  assumes "xs = ys" "(\<And>x. x \<in> set xs \<Longrightarrow> f x = g x)"
6625  shows   "List.bind xs f = List.bind ys g"
6626proof -
6627  from assms(2) have "List.bind xs f = List.bind xs g"
6628    by (induction xs) simp_all
6629  with assms(1) show ?thesis by simp
6630qed
6631
6632lemma set_list_bind: "set (List.bind xs f) = (\<Union>x\<in>set xs. set (f x))"
6633  by (induction xs) simp_all
6634
6635
6636subsection \<open>Code generation\<close>
6637
6638text\<open>Optional tail recursive version of @{const map}. Can avoid
6639stack overflow in some target languages.\<close>
6640
6641fun map_tailrec_rev ::  "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'b list" where
6642"map_tailrec_rev f [] bs = bs" |
6643"map_tailrec_rev f (a#as) bs = map_tailrec_rev f as (f a # bs)"
6644
6645lemma map_tailrec_rev:
6646  "map_tailrec_rev f as bs = rev(map f as) @ bs"
6647by(induction as arbitrary: bs) simp_all
6648
6649definition map_tailrec :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
6650"map_tailrec f as = rev (map_tailrec_rev f as [])"
6651
6652text\<open>Code equation:\<close>
6653lemma map_eq_map_tailrec: "map = map_tailrec"
6654by(simp add: fun_eq_iff map_tailrec_def map_tailrec_rev)
6655
6656
6657subsubsection \<open>Counterparts for set-related operations\<close>
6658
6659definition member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
6660[code_abbrev]: "member xs x \<longleftrightarrow> x \<in> set xs"
6661
6662text \<open>
6663  Use \<open>member\<close> only for generating executable code.  Otherwise use
6664  @{prop "x \<in> set xs"} instead --- it is much easier to reason about.
6665\<close>
6666
6667lemma member_rec [code]:
6668  "member (x # xs) y \<longleftrightarrow> x = y \<or> member xs y"
6669  "member [] y \<longleftrightarrow> False"
6670  by (auto simp add: member_def)
6671
6672lemma in_set_member (* FIXME delete candidate *):
6673  "x \<in> set xs \<longleftrightarrow> member xs x"
6674  by (simp add: member_def)
6675
6676lemmas list_all_iff [code_abbrev] = fun_cong[OF list.pred_set]
6677
6678definition list_ex :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
6679list_ex_iff [code_abbrev]: "list_ex P xs \<longleftrightarrow> Bex (set xs) P"
6680
6681definition list_ex1 :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
6682list_ex1_iff [code_abbrev]: "list_ex1 P xs \<longleftrightarrow> (\<exists>! x. x \<in> set xs \<and> P x)"
6683
6684text \<open>
6685  Usually you should prefer \<open>\<forall>x\<in>set xs\<close>, \<open>\<exists>x\<in>set xs\<close>
6686  and \<open>\<exists>!x. x\<in>set xs \<and> _\<close> over @{const list_all}, @{const list_ex}
6687  and @{const list_ex1} in specifications.
6688\<close>
6689
6690lemma list_all_simps [code]:
6691  "list_all P (x # xs) \<longleftrightarrow> P x \<and> list_all P xs"
6692  "list_all P [] \<longleftrightarrow> True"
6693  by (simp_all add: list_all_iff)
6694
6695lemma list_ex_simps [simp, code]:
6696  "list_ex P (x # xs) \<longleftrightarrow> P x \<or> list_ex P xs"
6697  "list_ex P [] \<longleftrightarrow> False"
6698  by (simp_all add: list_ex_iff)
6699
6700lemma list_ex1_simps [simp, code]:
6701  "list_ex1 P [] = False"
6702  "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)"
6703  by (auto simp add: list_ex1_iff list_all_iff)
6704
6705lemma Ball_set_list_all: (* FIXME delete candidate *)
6706  "Ball (set xs) P \<longleftrightarrow> list_all P xs"
6707  by (simp add: list_all_iff)
6708
6709lemma Bex_set_list_ex: (* FIXME delete candidate *)
6710  "Bex (set xs) P \<longleftrightarrow> list_ex P xs"
6711  by (simp add: list_ex_iff)
6712
6713lemma list_all_append [simp]:
6714  "list_all P (xs @ ys) \<longleftrightarrow> list_all P xs \<and> list_all P ys"
6715  by (auto simp add: list_all_iff)
6716
6717lemma list_ex_append [simp]:
6718  "list_ex P (xs @ ys) \<longleftrightarrow> list_ex P xs \<or> list_ex P ys"
6719  by (auto simp add: list_ex_iff)
6720
6721lemma list_all_rev [simp]:
6722  "list_all P (rev xs) \<longleftrightarrow> list_all P xs"
6723  by (simp add: list_all_iff)
6724
6725lemma list_ex_rev [simp]:
6726  "list_ex P (rev xs) \<longleftrightarrow> list_ex P xs"
6727  by (simp add: list_ex_iff)
6728
6729lemma list_all_length:
6730  "list_all P xs \<longleftrightarrow> (\<forall>n < length xs. P (xs ! n))"
6731  by (auto simp add: list_all_iff set_conv_nth)
6732
6733lemma list_ex_length:
6734  "list_ex P xs \<longleftrightarrow> (\<exists>n < length xs. P (xs ! n))"
6735  by (auto simp add: list_ex_iff set_conv_nth)
6736
6737lemmas list_all_cong [fundef_cong] = list.pred_cong
6738
6739lemma list_ex_cong [fundef_cong]:
6740  "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> list_ex f xs = list_ex g ys"
6741by (simp add: list_ex_iff)
6742
6743definition can_select :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool" where
6744[code_abbrev]: "can_select P A = (\<exists>!x\<in>A. P x)"
6745
6746lemma can_select_set_list_ex1 [code]:
6747  "can_select P (set A) = list_ex1 P A"
6748  by (simp add: list_ex1_iff can_select_def)
6749
6750
6751text \<open>Executable checks for relations on sets\<close>
6752
6753definition listrel1p :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
6754"listrel1p r xs ys = ((xs, ys) \<in> listrel1 {(x, y). r x y})"
6755
6756lemma [code_unfold]:
6757  "(xs, ys) \<in> listrel1 r = listrel1p (\<lambda>x y. (x, y) \<in> r) xs ys"
6758unfolding listrel1p_def by auto
6759
6760lemma [code]:
6761  "listrel1p r [] xs = False"
6762  "listrel1p r xs [] =  False"
6763  "listrel1p r (x # xs) (y # ys) \<longleftrightarrow>
6764     r x y \<and> xs = ys \<or> x = y \<and> listrel1p r xs ys"
6765by (simp add: listrel1p_def)+
6766
6767definition
6768  lexordp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
6769  "lexordp r xs ys = ((xs, ys) \<in> lexord {(x, y). r x y})"
6770
6771lemma [code_unfold]:
6772  "(xs, ys) \<in> lexord r = lexordp (\<lambda>x y. (x, y) \<in> r) xs ys"
6773unfolding lexordp_def by auto
6774
6775lemma [code]:
6776  "lexordp r xs [] = False"
6777  "lexordp r [] (y#ys) = True"
6778  "lexordp r (x # xs) (y # ys) = (r x y \<or> (x = y \<and> lexordp r xs ys))"
6779unfolding lexordp_def by auto
6780
6781text \<open>Bounded quantification and summation over nats.\<close>
6782
6783lemma atMost_upto [code_unfold]:
6784  "{..n} = set [0..<Suc n]"
6785  by auto
6786
6787lemma atLeast_upt [code_unfold]:
6788  "{..<n} = set [0..<n]"
6789  by auto
6790
6791lemma greaterThanLessThan_upt [code_unfold]:
6792  "{n<..<m} = set [Suc n..<m]"
6793  by auto
6794
6795lemmas atLeastLessThan_upt [code_unfold] = set_upt [symmetric]
6796
6797lemma greaterThanAtMost_upt [code_unfold]:
6798  "{n<..m} = set [Suc n..<Suc m]"
6799  by auto
6800
6801lemma atLeastAtMost_upt [code_unfold]:
6802  "{n..m} = set [n..<Suc m]"
6803  by auto
6804
6805lemma all_nat_less_eq [code_unfold]:
6806  "(\<forall>m<n::nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..<n}. P m)"
6807  by auto
6808
6809lemma ex_nat_less_eq [code_unfold]:
6810  "(\<exists>m<n::nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..<n}. P m)"
6811  by auto
6812
6813lemma all_nat_less [code_unfold]:
6814  "(\<forall>m\<le>n::nat. P m) \<longleftrightarrow> (\<forall>m \<in> {0..n}. P m)"
6815  by auto
6816
6817lemma ex_nat_less [code_unfold]:
6818  "(\<exists>m\<le>n::nat. P m) \<longleftrightarrow> (\<exists>m \<in> {0..n}. P m)"
6819  by auto
6820
6821text\<open>Bounded \<open>LEAST\<close> operator:\<close>
6822
6823definition "Bleast S P = (LEAST x. x \<in> S \<and> P x)"
6824
6825definition "abort_Bleast S P = (LEAST x. x \<in> S \<and> P x)"
6826
6827declare [[code abort: abort_Bleast]]
6828
6829lemma Bleast_code [code]:
6830 "Bleast (set xs) P = (case filter P (sort xs) of
6831    x#xs \<Rightarrow> x |
6832    [] \<Rightarrow> abort_Bleast (set xs) P)"
6833proof (cases "filter P (sort xs)")
6834  case Nil thus ?thesis by (simp add: Bleast_def abort_Bleast_def)
6835next
6836  case (Cons x ys)
6837  have "(LEAST x. x \<in> set xs \<and> P x) = x"
6838  proof (rule Least_equality)
6839    show "x \<in> set xs \<and> P x"
6840      by (metis Cons Cons_eq_filter_iff in_set_conv_decomp set_sort)
6841    next
6842      fix y assume "y \<in> set xs \<and> P y"
6843      hence "y \<in> set (filter P xs)" by auto
6844      thus "x \<le> y"
6845        by (metis Cons eq_iff filter_sort set_ConsD set_sort sorted.simps(2) sorted_sort)
6846  qed
6847  thus ?thesis using Cons by (simp add: Bleast_def)
6848qed
6849
6850declare Bleast_def[symmetric, code_unfold]
6851
6852text \<open>Summation over ints.\<close>
6853
6854lemma greaterThanLessThan_upto [code_unfold]:
6855  "{i<..<j::int} = set [i+1..j - 1]"
6856by auto
6857
6858lemma atLeastLessThan_upto [code_unfold]:
6859  "{i..<j::int} = set [i..j - 1]"
6860by auto
6861
6862lemma greaterThanAtMost_upto [code_unfold]:
6863  "{i<..j::int} = set [i+1..j]"
6864by auto
6865
6866lemmas atLeastAtMost_upto [code_unfold] = set_upto [symmetric]
6867
6868
6869subsubsection \<open>Optimizing by rewriting\<close>
6870
6871definition null :: "'a list \<Rightarrow> bool" where
6872  [code_abbrev]: "null xs \<longleftrightarrow> xs = []"
6873
6874text \<open>
6875  Efficient emptyness check is implemented by @{const null}.
6876\<close>
6877
6878lemma null_rec [code]:
6879  "null (x # xs) \<longleftrightarrow> False"
6880  "null [] \<longleftrightarrow> True"
6881  by (simp_all add: null_def)
6882
6883lemma eq_Nil_null: (* FIXME delete candidate *)
6884  "xs = [] \<longleftrightarrow> null xs"
6885  by (simp add: null_def)
6886
6887lemma equal_Nil_null [code_unfold]:
6888  "HOL.equal xs [] \<longleftrightarrow> null xs"
6889  "HOL.equal [] = null"
6890  by (auto simp add: equal null_def)
6891
6892definition maps :: "('a \<Rightarrow> 'b list) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
6893  [code_abbrev]: "maps f xs = concat (map f xs)"
6894
6895definition map_filter :: "('a \<Rightarrow> 'b option) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
6896  [code_post]: "map_filter f xs = map (the \<circ> f) (filter (\<lambda>x. f x \<noteq> None) xs)"
6897
6898text \<open>
6899  Operations @{const maps} and @{const map_filter} avoid
6900  intermediate lists on execution -- do not use for proving.
6901\<close>
6902
6903lemma maps_simps [code]:
6904  "maps f (x # xs) = f x @ maps f xs"
6905  "maps f [] = []"
6906  by (simp_all add: maps_def)
6907
6908lemma map_filter_simps [code]:
6909  "map_filter f (x # xs) = (case f x of None \<Rightarrow> map_filter f xs | Some y \<Rightarrow> y # map_filter f xs)"
6910  "map_filter f [] = []"
6911  by (simp_all add: map_filter_def split: option.split)
6912
6913lemma concat_map_maps: (* FIXME delete candidate *)
6914  "concat (map f xs) = maps f xs"
6915  by (simp add: maps_def)
6916
6917lemma map_filter_map_filter [code_unfold]:
6918  "map f (filter P xs) = map_filter (\<lambda>x. if P x then Some (f x) else None) xs"
6919  by (simp add: map_filter_def)
6920
6921text \<open>Optimized code for \<open>\<forall>i\<in>{a..b::int}\<close> and \<open>\<forall>n:{a..<b::nat}\<close>
6922and similiarly for \<open>\<exists>\<close>.\<close>
6923
6924definition all_interval_nat :: "(nat \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
6925  "all_interval_nat P i j \<longleftrightarrow> (\<forall>n \<in> {i..<j}. P n)"
6926
6927lemma [code]:
6928  "all_interval_nat P i j \<longleftrightarrow> i \<ge> j \<or> P i \<and> all_interval_nat P (Suc i) j"
6929proof -
6930  have *: "\<And>n. P i \<Longrightarrow> \<forall>n\<in>{Suc i..<j}. P n \<Longrightarrow> i \<le> n \<Longrightarrow> n < j \<Longrightarrow> P n"
6931  proof -
6932    fix n
6933    assume "P i" "\<forall>n\<in>{Suc i..<j}. P n" "i \<le> n" "n < j"
6934    then show "P n" by (cases "n = i") simp_all
6935  qed
6936  show ?thesis by (auto simp add: all_interval_nat_def intro: *)
6937qed
6938
6939lemma list_all_iff_all_interval_nat [code_unfold]:
6940  "list_all P [i..<j] \<longleftrightarrow> all_interval_nat P i j"
6941  by (simp add: list_all_iff all_interval_nat_def)
6942
6943lemma list_ex_iff_not_all_inverval_nat [code_unfold]:
6944  "list_ex P [i..<j] \<longleftrightarrow> \<not> (all_interval_nat (Not \<circ> P) i j)"
6945  by (simp add: list_ex_iff all_interval_nat_def)
6946
6947definition all_interval_int :: "(int \<Rightarrow> bool) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool" where
6948  "all_interval_int P i j \<longleftrightarrow> (\<forall>k \<in> {i..j}. P k)"
6949
6950lemma [code]:
6951  "all_interval_int P i j \<longleftrightarrow> i > j \<or> P i \<and> all_interval_int P (i + 1) j"
6952proof -
6953  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"
6954  proof -
6955    fix k
6956    assume "P i" "\<forall>k\<in>{i+1..j}. P k" "i \<le> k" "k \<le> j"
6957    then show "P k" by (cases "k = i") simp_all
6958  qed
6959  show ?thesis by (auto simp add: all_interval_int_def intro: *)
6960qed
6961
6962lemma list_all_iff_all_interval_int [code_unfold]:
6963  "list_all P [i..j] \<longleftrightarrow> all_interval_int P i j"
6964  by (simp add: list_all_iff all_interval_int_def)
6965
6966lemma list_ex_iff_not_all_inverval_int [code_unfold]:
6967  "list_ex P [i..j] \<longleftrightarrow> \<not> (all_interval_int (Not \<circ> P) i j)"
6968  by (simp add: list_ex_iff all_interval_int_def)
6969
6970text \<open>optimized code (tail-recursive) for @{term length}\<close>
6971
6972definition gen_length :: "nat \<Rightarrow> 'a list \<Rightarrow> nat"
6973where "gen_length n xs = n + length xs"
6974
6975lemma gen_length_code [code]:
6976  "gen_length n [] = n"
6977  "gen_length n (x # xs) = gen_length (Suc n) xs"
6978by(simp_all add: gen_length_def)
6979
6980declare list.size(3-4)[code del]
6981
6982lemma length_code [code]: "length = gen_length 0"
6983by(simp add: gen_length_def fun_eq_iff)
6984
6985hide_const (open) member null maps map_filter all_interval_nat all_interval_int gen_length
6986
6987
6988subsubsection \<open>Pretty lists\<close>
6989
6990ML \<open>
6991(* Code generation for list literals. *)
6992
6993signature LIST_CODE =
6994sig
6995  val add_literal_list: string -> theory -> theory
6996end;
6997
6998structure List_Code : LIST_CODE =
6999struct
7000
7001open Basic_Code_Thingol;
7002
7003fun implode_list t =
7004  let
7005    fun dest_cons (IConst { sym = Code_Symbol.Constant @{const_name Cons}, ... } `$ t1 `$ t2) = SOME (t1, t2)
7006      | dest_cons _ = NONE;
7007    val (ts, t') = Code_Thingol.unfoldr dest_cons t;
7008  in case t'
7009   of IConst { sym = Code_Symbol.Constant @{const_name Nil}, ... } => SOME ts
7010    | _ => NONE
7011  end;
7012
7013fun print_list (target_fxy, target_cons) pr fxy t1 t2 =
7014  Code_Printer.brackify_infix (target_fxy, Code_Printer.R) fxy (
7015    pr (Code_Printer.INFX (target_fxy, Code_Printer.X)) t1,
7016    Code_Printer.str target_cons,
7017    pr (Code_Printer.INFX (target_fxy, Code_Printer.R)) t2
7018  );
7019
7020fun add_literal_list target =
7021  let
7022    fun pretty literals pr _ vars fxy [(t1, _), (t2, _)] =
7023      case Option.map (cons t1) (implode_list t2)
7024       of SOME ts =>
7025            Code_Printer.literal_list literals (map (pr vars Code_Printer.NOBR) ts)
7026        | NONE =>
7027            print_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
7028  in
7029    Code_Target.set_printings (Code_Symbol.Constant (@{const_name Cons},
7030      [(target, SOME (Code_Printer.complex_const_syntax (2, pretty)))]))
7031  end
7032
7033end;
7034\<close>
7035
7036code_printing
7037  type_constructor list \<rightharpoonup>
7038    (SML) "_ list"
7039    and (OCaml) "_ list"
7040    and (Haskell) "![(_)]"
7041    and (Scala) "List[(_)]"
7042| constant Nil \<rightharpoonup>
7043    (SML) "[]"
7044    and (OCaml) "[]"
7045    and (Haskell) "[]"
7046    and (Scala) "!Nil"
7047| class_instance list :: equal \<rightharpoonup>
7048    (Haskell) -
7049| constant "HOL.equal :: 'a list \<Rightarrow> 'a list \<Rightarrow> bool" \<rightharpoonup>
7050    (Haskell) infix 4 "=="
7051
7052setup \<open>fold (List_Code.add_literal_list) ["SML", "OCaml", "Haskell", "Scala"]\<close>
7053
7054code_reserved SML
7055  list
7056
7057code_reserved OCaml
7058  list
7059
7060
7061subsubsection \<open>Use convenient predefined operations\<close>
7062
7063code_printing
7064  constant "(@)" \<rightharpoonup>
7065    (SML) infixr 7 "@"
7066    and (OCaml) infixr 6 "@"
7067    and (Haskell) infixr 5 "++"
7068    and (Scala) infixl 7 "++"
7069| constant map \<rightharpoonup>
7070    (Haskell) "map"
7071| constant filter \<rightharpoonup>
7072    (Haskell) "filter"
7073| constant concat \<rightharpoonup>
7074    (Haskell) "concat"
7075| constant List.maps \<rightharpoonup>
7076    (Haskell) "concatMap"
7077| constant rev \<rightharpoonup>
7078    (Haskell) "reverse"
7079| constant zip \<rightharpoonup>
7080    (Haskell) "zip"
7081| constant List.null \<rightharpoonup>
7082    (Haskell) "null"
7083| constant takeWhile \<rightharpoonup>
7084    (Haskell) "takeWhile"
7085| constant dropWhile \<rightharpoonup>
7086    (Haskell) "dropWhile"
7087| constant list_all \<rightharpoonup>
7088    (Haskell) "all"
7089| constant list_ex \<rightharpoonup>
7090    (Haskell) "any"
7091
7092
7093subsubsection \<open>Implementation of sets by lists\<close>
7094
7095lemma is_empty_set [code]:
7096  "Set.is_empty (set xs) \<longleftrightarrow> List.null xs"
7097  by (simp add: Set.is_empty_def null_def)
7098
7099lemma empty_set [code]:
7100  "{} = set []"
7101  by simp
7102
7103lemma UNIV_coset [code]:
7104  "UNIV = List.coset []"
7105  by simp
7106
7107lemma compl_set [code]:
7108  "- set xs = List.coset xs"
7109  by simp
7110
7111lemma compl_coset [code]:
7112  "- List.coset xs = set xs"
7113  by simp
7114
7115lemma [code]:
7116  "x \<in> set xs \<longleftrightarrow> List.member xs x"
7117  "x \<in> List.coset xs \<longleftrightarrow> \<not> List.member xs x"
7118  by (simp_all add: member_def)
7119
7120lemma insert_code [code]:
7121  "insert x (set xs) = set (List.insert x xs)"
7122  "insert x (List.coset xs) = List.coset (removeAll x xs)"
7123  by simp_all
7124
7125lemma remove_code [code]:
7126  "Set.remove x (set xs) = set (removeAll x xs)"
7127  "Set.remove x (List.coset xs) = List.coset (List.insert x xs)"
7128  by (simp_all add: remove_def Compl_insert)
7129
7130lemma filter_set [code]:
7131  "Set.filter P (set xs) = set (filter P xs)"
7132  by auto
7133
7134lemma image_set [code]:
7135  "image f (set xs) = set (map f xs)"
7136  by simp
7137
7138lemma subset_code [code]:
7139  "set xs \<le> B \<longleftrightarrow> (\<forall>x\<in>set xs. x \<in> B)"
7140  "A \<le> List.coset ys \<longleftrightarrow> (\<forall>y\<in>set ys. y \<notin> A)"
7141  "List.coset [] \<le> set [] \<longleftrightarrow> False"
7142  by auto
7143
7144text \<open>A frequent case -- avoid intermediate sets\<close>
7145lemma [code_unfold]:
7146  "set xs \<subseteq> set ys \<longleftrightarrow> list_all (\<lambda>x. x \<in> set ys) xs"
7147  by (auto simp: list_all_iff)
7148
7149lemma Ball_set [code]:
7150  "Ball (set xs) P \<longleftrightarrow> list_all P xs"
7151  by (simp add: list_all_iff)
7152
7153lemma Bex_set [code]:
7154  "Bex (set xs) P \<longleftrightarrow> list_ex P xs"
7155  by (simp add: list_ex_iff)
7156
7157lemma card_set [code]:
7158  "card (set xs) = length (remdups xs)"
7159proof -
7160  have "card (set (remdups xs)) = length (remdups xs)"
7161    by (rule distinct_card) simp
7162  then show ?thesis by simp
7163qed
7164
7165lemma the_elem_set [code]:
7166  "the_elem (set [x]) = x"
7167  by simp
7168
7169lemma Pow_set [code]:
7170  "Pow (set []) = {{}}"
7171  "Pow (set (x # xs)) = (let A = Pow (set xs) in A \<union> insert x ` A)"
7172  by (simp_all add: Pow_insert Let_def)
7173
7174definition map_project :: "('a \<Rightarrow> 'b option) \<Rightarrow> 'a set \<Rightarrow> 'b set" where
7175  "map_project f A = {b. \<exists> a \<in> A. f a = Some b}"
7176
7177lemma [code]:
7178  "map_project f (set xs) = set (List.map_filter f xs)"
7179  by (auto simp add: map_project_def map_filter_def image_def)
7180
7181hide_const (open) map_project
7182
7183
7184text \<open>Operations on relations\<close>
7185
7186lemma product_code [code]:
7187  "Product_Type.product (set xs) (set ys) = set [(x, y). x \<leftarrow> xs, y \<leftarrow> ys]"
7188  by (auto simp add: Product_Type.product_def)
7189
7190lemma Id_on_set [code]:
7191  "Id_on (set xs) = set [(x, x). x \<leftarrow> xs]"
7192  by (auto simp add: Id_on_def)
7193
7194lemma [code]:
7195  "R `` S = List.map_project (\<lambda>(x, y). if x \<in> S then Some y else None) R"
7196unfolding map_project_def by (auto split: prod.split if_split_asm)
7197
7198lemma trancl_set_ntrancl [code]:
7199  "trancl (set xs) = ntrancl (card (set xs) - 1) (set xs)"
7200  by (simp add: finite_trancl_ntranl)
7201
7202lemma set_relcomp [code]:
7203  "set xys O set yzs = set ([(fst xy, snd yz). xy \<leftarrow> xys, yz \<leftarrow> yzs, snd xy = fst yz])"
7204  by auto (auto simp add: Bex_def image_def)
7205
7206lemma wf_set [code]:
7207  "wf (set xs) = acyclic (set xs)"
7208  by (simp add: wf_iff_acyclic_if_finite)
7209
7210
7211subsection \<open>Setup for Lifting/Transfer\<close>
7212
7213subsubsection \<open>Transfer rules for the Transfer package\<close>
7214
7215context includes lifting_syntax
7216begin
7217
7218lemma tl_transfer [transfer_rule]:
7219  "(list_all2 A ===> list_all2 A) tl tl"
7220  unfolding tl_def[abs_def] by transfer_prover
7221
7222lemma butlast_transfer [transfer_rule]:
7223  "(list_all2 A ===> list_all2 A) butlast butlast"
7224  by (rule rel_funI, erule list_all2_induct, auto)
7225
7226lemma map_rec: "map f xs = rec_list Nil (%x _ y. Cons (f x) y) xs"
7227  by (induct xs) auto
7228
7229lemma append_transfer [transfer_rule]:
7230  "(list_all2 A ===> list_all2 A ===> list_all2 A) append append"
7231  unfolding List.append_def by transfer_prover
7232
7233lemma rev_transfer [transfer_rule]:
7234  "(list_all2 A ===> list_all2 A) rev rev"
7235  unfolding List.rev_def by transfer_prover
7236
7237lemma filter_transfer [transfer_rule]:
7238  "((A ===> (=)) ===> list_all2 A ===> list_all2 A) filter filter"
7239  unfolding List.filter_def by transfer_prover
7240
7241lemma fold_transfer [transfer_rule]:
7242  "((A ===> B ===> B) ===> list_all2 A ===> B ===> B) fold fold"
7243  unfolding List.fold_def by transfer_prover
7244
7245lemma foldr_transfer [transfer_rule]:
7246  "((A ===> B ===> B) ===> list_all2 A ===> B ===> B) foldr foldr"
7247  unfolding List.foldr_def by transfer_prover
7248
7249lemma foldl_transfer [transfer_rule]:
7250  "((B ===> A ===> B) ===> B ===> list_all2 A ===> B) foldl foldl"
7251  unfolding List.foldl_def by transfer_prover
7252
7253lemma concat_transfer [transfer_rule]:
7254  "(list_all2 (list_all2 A) ===> list_all2 A) concat concat"
7255  unfolding List.concat_def by transfer_prover
7256
7257lemma drop_transfer [transfer_rule]:
7258  "((=) ===> list_all2 A ===> list_all2 A) drop drop"
7259  unfolding List.drop_def by transfer_prover
7260
7261lemma take_transfer [transfer_rule]:
7262  "((=) ===> list_all2 A ===> list_all2 A) take take"
7263  unfolding List.take_def by transfer_prover
7264
7265lemma list_update_transfer [transfer_rule]:
7266  "(list_all2 A ===> (=) ===> A ===> list_all2 A) list_update list_update"
7267  unfolding list_update_def by transfer_prover
7268
7269lemma takeWhile_transfer [transfer_rule]:
7270  "((A ===> (=)) ===> list_all2 A ===> list_all2 A) takeWhile takeWhile"
7271  unfolding takeWhile_def by transfer_prover
7272
7273lemma dropWhile_transfer [transfer_rule]:
7274  "((A ===> (=)) ===> list_all2 A ===> list_all2 A) dropWhile dropWhile"
7275  unfolding dropWhile_def by transfer_prover
7276
7277lemma zip_transfer [transfer_rule]:
7278  "(list_all2 A ===> list_all2 B ===> list_all2 (rel_prod A B)) zip zip"
7279  unfolding zip_def by transfer_prover
7280
7281lemma product_transfer [transfer_rule]:
7282  "(list_all2 A ===> list_all2 B ===> list_all2 (rel_prod A B)) List.product List.product"
7283  unfolding List.product_def by transfer_prover
7284
7285lemma product_lists_transfer [transfer_rule]:
7286  "(list_all2 (list_all2 A) ===> list_all2 (list_all2 A)) product_lists product_lists"
7287  unfolding product_lists_def by transfer_prover
7288
7289lemma insert_transfer [transfer_rule]:
7290  assumes [transfer_rule]: "bi_unique A"
7291  shows "(A ===> list_all2 A ===> list_all2 A) List.insert List.insert"
7292  unfolding List.insert_def [abs_def] by transfer_prover
7293
7294lemma find_transfer [transfer_rule]:
7295  "((A ===> (=)) ===> list_all2 A ===> rel_option A) List.find List.find"
7296  unfolding List.find_def by transfer_prover
7297
7298lemma those_transfer [transfer_rule]:
7299  "(list_all2 (rel_option P) ===> rel_option (list_all2 P)) those those"
7300  unfolding List.those_def by transfer_prover
7301
7302lemma remove1_transfer [transfer_rule]:
7303  assumes [transfer_rule]: "bi_unique A"
7304  shows "(A ===> list_all2 A ===> list_all2 A) remove1 remove1"
7305  unfolding remove1_def by transfer_prover
7306
7307lemma removeAll_transfer [transfer_rule]:
7308  assumes [transfer_rule]: "bi_unique A"
7309  shows "(A ===> list_all2 A ===> list_all2 A) removeAll removeAll"
7310  unfolding removeAll_def by transfer_prover
7311
7312lemma distinct_transfer [transfer_rule]:
7313  assumes [transfer_rule]: "bi_unique A"
7314  shows "(list_all2 A ===> (=)) distinct distinct"
7315  unfolding distinct_def by transfer_prover
7316
7317lemma remdups_transfer [transfer_rule]:
7318  assumes [transfer_rule]: "bi_unique A"
7319  shows "(list_all2 A ===> list_all2 A) remdups remdups"
7320  unfolding remdups_def by transfer_prover
7321
7322lemma remdups_adj_transfer [transfer_rule]:
7323  assumes [transfer_rule]: "bi_unique A"
7324  shows "(list_all2 A ===> list_all2 A) remdups_adj remdups_adj"
7325  proof (rule rel_funI, erule list_all2_induct)
7326  qed (auto simp: remdups_adj_Cons assms[unfolded bi_unique_def] split: list.splits)
7327
7328lemma replicate_transfer [transfer_rule]:
7329  "((=) ===> A ===> list_all2 A) replicate replicate"
7330  unfolding replicate_def by transfer_prover
7331
7332lemma length_transfer [transfer_rule]:
7333  "(list_all2 A ===> (=)) length length"
7334  unfolding size_list_overloaded_def size_list_def by transfer_prover
7335
7336lemma rotate1_transfer [transfer_rule]:
7337  "(list_all2 A ===> list_all2 A) rotate1 rotate1"
7338  unfolding rotate1_def by transfer_prover
7339
7340lemma rotate_transfer [transfer_rule]:
7341  "((=) ===> list_all2 A ===> list_all2 A) rotate rotate"
7342  unfolding rotate_def [abs_def] by transfer_prover
7343
7344lemma nths_transfer [transfer_rule]:
7345  "(list_all2 A ===> rel_set (=) ===> list_all2 A) nths nths"
7346  unfolding nths_def [abs_def] by transfer_prover
7347
7348lemma subseqs_transfer [transfer_rule]:
7349  "(list_all2 A ===> list_all2 (list_all2 A)) subseqs subseqs"
7350  unfolding subseqs_def [abs_def] by transfer_prover
7351
7352lemma partition_transfer [transfer_rule]:
7353  "((A ===> (=)) ===> list_all2 A ===> rel_prod (list_all2 A) (list_all2 A))
7354    partition partition"
7355  unfolding partition_def by transfer_prover
7356
7357lemma lists_transfer [transfer_rule]:
7358  "(rel_set A ===> rel_set (list_all2 A)) lists lists"
7359  apply (rule rel_funI, rule rel_setI)
7360  apply (erule lists.induct, simp)
7361  apply (simp only: rel_set_def list_all2_Cons1, metis lists.Cons)
7362  apply (erule lists.induct, simp)
7363  apply (simp only: rel_set_def list_all2_Cons2, metis lists.Cons)
7364  done
7365
7366lemma set_Cons_transfer [transfer_rule]:
7367  "(rel_set A ===> rel_set (list_all2 A) ===> rel_set (list_all2 A))
7368    set_Cons set_Cons"
7369  unfolding rel_fun_def rel_set_def set_Cons_def
7370  apply safe
7371  apply (simp add: list_all2_Cons1, fast)
7372  apply (simp add: list_all2_Cons2, fast)
7373  done
7374
7375lemma listset_transfer [transfer_rule]:
7376  "(list_all2 (rel_set A) ===> rel_set (list_all2 A)) listset listset"
7377  unfolding listset_def by transfer_prover
7378
7379lemma null_transfer [transfer_rule]:
7380  "(list_all2 A ===> (=)) List.null List.null"
7381  unfolding rel_fun_def List.null_def by auto
7382
7383lemma list_all_transfer [transfer_rule]:
7384  "((A ===> (=)) ===> list_all2 A ===> (=)) list_all list_all"
7385  unfolding list_all_iff [abs_def] by transfer_prover
7386
7387lemma list_ex_transfer [transfer_rule]:
7388  "((A ===> (=)) ===> list_all2 A ===> (=)) list_ex list_ex"
7389  unfolding list_ex_iff [abs_def] by transfer_prover
7390
7391lemma splice_transfer [transfer_rule]:
7392  "(list_all2 A ===> list_all2 A ===> list_all2 A) splice splice"
7393  apply (rule rel_funI, erule list_all2_induct, simp add: rel_fun_def, simp)
7394  apply (rule rel_funI)
7395  apply (erule_tac xs=x in list_all2_induct, simp, simp add: rel_fun_def)
7396  done
7397
7398lemma shuffle_transfer [transfer_rule]:
7399  "(list_all2 A ===> list_all2 A ===> rel_set (list_all2 A)) shuffle shuffle"
7400proof (intro rel_funI, goal_cases)
7401  case (1 xs xs' ys ys')
7402  thus ?case
7403  proof (induction xs ys arbitrary: xs' ys' rule: shuffle.induct)
7404    case (3 x xs y ys xs' ys')
7405    from "3.prems" obtain x' xs'' where xs': "xs' = x' # xs''" by (cases xs') auto
7406    from "3.prems" obtain y' ys'' where ys': "ys' = y' # ys''" by (cases ys') auto
7407    have [transfer_rule]: "A x x'" "A y y'" "list_all2 A xs xs''" "list_all2 A ys ys''"
7408      using "3.prems" by (simp_all add: xs' ys')
7409    have [transfer_rule]: "rel_set (list_all2 A) (shuffle xs (y # ys)) (shuffle xs'' ys')" and
7410         [transfer_rule]: "rel_set (list_all2 A) (shuffle (x # xs) ys) (shuffle xs' ys'')"
7411      using "3.prems" by (auto intro!: "3.IH" simp: xs' ys')
7412    have "rel_set (list_all2 A) ((#) x ` shuffle xs (y # ys) \<union> (#) y ` shuffle (x # xs) ys)
7413               ((#) x' ` shuffle xs'' ys' \<union> (#) y' ` shuffle xs' ys'')" by transfer_prover
7414    thus ?case by (simp add: xs' ys')
7415  qed (auto simp: rel_set_def)
7416qed
7417
7418lemma rtrancl_parametric [transfer_rule]:
7419  assumes [transfer_rule]: "bi_unique A" "bi_total A"
7420  shows "(rel_set (rel_prod A A) ===> rel_set (rel_prod A A)) rtrancl rtrancl"
7421unfolding rtrancl_def by transfer_prover
7422
7423lemma monotone_parametric [transfer_rule]:
7424  assumes [transfer_rule]: "bi_total A"
7425  shows "((A ===> A ===> (=)) ===> (B ===> B ===> (=)) ===> (A ===> B) ===> (=)) monotone monotone"
7426unfolding monotone_def[abs_def] by transfer_prover
7427
7428lemma fun_ord_parametric [transfer_rule]:
7429  assumes [transfer_rule]: "bi_total C"
7430  shows "((A ===> B ===> (=)) ===> (C ===> A) ===> (C ===> B) ===> (=)) fun_ord fun_ord"
7431unfolding fun_ord_def[abs_def] by transfer_prover
7432
7433lemma fun_lub_parametric [transfer_rule]:
7434  assumes [transfer_rule]: "bi_total A"  "bi_unique A"
7435  shows "((rel_set A ===> B) ===> rel_set (C ===> A) ===> C ===> B) fun_lub fun_lub"
7436unfolding fun_lub_def[abs_def] by transfer_prover
7437
7438end
7439
7440end
7441