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