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