src/HOL/List.thy
author paulson
Wed, 01 Sep 2004 15:03:41 +0200
changeset 15168 33a08cfc3ae5
parent 15140 322485b816ac
child 15176 2fd60846f485
permissions -rw-r--r--
new functions for sets of lists
     1 (*  Title:      HOL/List.thy
     2     ID:         $Id$
     3     Author:     Tobias Nipkow
     4 *)
     5 
     6 header {* The datatype of finite lists *}
     7 
     8 theory List
     9 imports PreList
    10 begin
    11 
    12 datatype 'a list =
    13     Nil    ("[]")
    14   | Cons 'a  "'a list"    (infixr "#" 65)
    15 
    16 consts
    17   "@" :: "'a list => 'a list => 'a list"    (infixr 65)
    18   filter:: "('a => bool) => 'a list => 'a list"
    19   concat:: "'a list list => 'a list"
    20   foldl :: "('b => 'a => 'b) => 'b => 'a list => 'b"
    21   foldr :: "('a => 'b => 'b) => 'a list => 'b => 'b"
    22   hd:: "'a list => 'a"
    23   tl:: "'a list => 'a list"
    24   last:: "'a list => 'a"
    25   butlast :: "'a list => 'a list"
    26   set :: "'a list => 'a set"
    27   list_all:: "('a => bool) => ('a list => bool)"
    28   list_all2 :: "('a => 'b => bool) => 'a list => 'b list => bool"
    29   map :: "('a=>'b) => ('a list => 'b list)"
    30   mem :: "'a => 'a list => bool"    (infixl 55)
    31   nth :: "'a list => nat => 'a"    (infixl "!" 100)
    32   list_update :: "'a list => nat => 'a => 'a list"
    33   take:: "nat => 'a list => 'a list"
    34   drop:: "nat => 'a list => 'a list"
    35   takeWhile :: "('a => bool) => 'a list => 'a list"
    36   dropWhile :: "('a => bool) => 'a list => 'a list"
    37   rev :: "'a list => 'a list"
    38   zip :: "'a list => 'b list => ('a * 'b) list"
    39   upt :: "nat => nat => nat list" ("(1[_../_'(])")
    40   remdups :: "'a list => 'a list"
    41   remove1 :: "'a => 'a list => 'a list"
    42   null:: "'a list => bool"
    43   "distinct":: "'a list => bool"
    44   replicate :: "nat => 'a => 'a list"
    45 
    46 nonterminals lupdbinds lupdbind
    47 
    48 syntax
    49   -- {* list Enumeration *}
    50   "@list" :: "args => 'a list"    ("[(_)]")
    51 
    52   -- {* Special syntax for filter *}
    53   "@filter" :: "[pttrn, 'a list, bool] => 'a list"    ("(1[_:_./ _])")
    54 
    55   -- {* list update *}
    56   "_lupdbind":: "['a, 'a] => lupdbind"    ("(2_ :=/ _)")
    57   "" :: "lupdbind => lupdbinds"    ("_")
    58   "_lupdbinds" :: "[lupdbind, lupdbinds] => lupdbinds"    ("_,/ _")
    59   "_LUpdate" :: "['a, lupdbinds] => 'a"    ("_/[(_)]" [900,0] 900)
    60 
    61   upto:: "nat => nat => nat list"    ("(1[_../_])")
    62 
    63 translations
    64   "[x, xs]" == "x#[xs]"
    65   "[x]" == "x#[]"
    66   "[x:xs . P]"== "filter (%x. P) xs"
    67 
    68   "_LUpdate xs (_lupdbinds b bs)"== "_LUpdate (_LUpdate xs b) bs"
    69   "xs[i:=x]" == "list_update xs i x"
    70 
    71   "[i..j]" == "[i..(Suc j)(]"
    72 
    73 
    74 syntax (xsymbols)
    75   "@filter" :: "[pttrn, 'a list, bool] => 'a list"("(1[_\<in>_ ./ _])")
    76 syntax (HTML output)
    77   "@filter" :: "[pttrn, 'a list, bool] => 'a list"("(1[_\<in>_ ./ _])")
    78 
    79 
    80 text {*
    81   Function @{text size} is overloaded for all datatypes. Users may
    82   refer to the list version as @{text length}. *}
    83 
    84 syntax length :: "'a list => nat"
    85 translations "length" => "size :: _ list => nat"
    86 
    87 typed_print_translation {*
    88   let
    89     fun size_tr' _ (Type ("fun", (Type ("list", _) :: _))) [t] =
    90           Syntax.const "length" $ t
    91       | size_tr' _ _ _ = raise Match;
    92   in [("size", size_tr')] end
    93 *}
    94 
    95 primrec
    96 "hd(x#xs) = x"
    97 primrec
    98 "tl([]) = []"
    99 "tl(x#xs) = xs"
   100 primrec
   101 "null([]) = True"
   102 "null(x#xs) = False"
   103 primrec
   104 "last(x#xs) = (if xs=[] then x else last xs)"
   105 primrec
   106 "butlast []= []"
   107 "butlast(x#xs) = (if xs=[] then [] else x#butlast xs)"
   108 primrec
   109 "x mem [] = False"
   110 "x mem (y#ys) = (if y=x then True else x mem ys)"
   111 primrec
   112 "set [] = {}"
   113 "set (x#xs) = insert x (set xs)"
   114 primrec
   115 list_all_Nil:"list_all P [] = True"
   116 list_all_Cons: "list_all P (x#xs) = (P(x) \<and> list_all P xs)"
   117 primrec
   118 "map f [] = []"
   119 "map f (x#xs) = f(x)#map f xs"
   120 primrec
   121 append_Nil:"[]@ys = ys"
   122 append_Cons: "(x#xs)@ys = x#(xs@ys)"
   123 primrec
   124 "rev([]) = []"
   125 "rev(x#xs) = rev(xs) @ [x]"
   126 primrec
   127 "filter P [] = []"
   128 "filter P (x#xs) = (if P x then x#filter P xs else filter P xs)"
   129 primrec
   130 foldl_Nil:"foldl f a [] = a"
   131 foldl_Cons: "foldl f a (x#xs) = foldl f (f a x) xs"
   132 primrec
   133 "foldr f [] a = a"
   134 "foldr f (x#xs) a = f x (foldr f xs a)"
   135 primrec
   136 "concat([]) = []"
   137 "concat(x#xs) = x @ concat(xs)"
   138 primrec
   139 drop_Nil:"drop n [] = []"
   140 drop_Cons: "drop n (x#xs) = (case n of 0 => x#xs | Suc(m) => drop m xs)"
   141 -- {* Warning: simpset does not contain this definition *}
   142 -- {* but separate theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
   143 primrec
   144 take_Nil:"take n [] = []"
   145 take_Cons: "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)"
   146 -- {* Warning: simpset does not contain this definition *}
   147 -- {* but separate theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
   148 primrec
   149 nth_Cons:"(x#xs)!n = (case n of 0 => x | (Suc k) => xs!k)"
   150 -- {* Warning: simpset does not contain this definition *}
   151 -- {* but separate theorems for @{text "n = 0"} and @{text "n = Suc k"} *}
   152 primrec
   153 "[][i:=v] = []"
   154 "(x#xs)[i:=v] =
   155 (case i of 0 => v # xs
   156 | Suc j => x # xs[j:=v])"
   157 primrec
   158 "takeWhile P [] = []"
   159 "takeWhile P (x#xs) = (if P x then x#takeWhile P xs else [])"
   160 primrec
   161 "dropWhile P [] = []"
   162 "dropWhile P (x#xs) = (if P x then dropWhile P xs else x#xs)"
   163 primrec
   164 "zip xs [] = []"
   165 zip_Cons: "zip xs (y#ys) = (case xs of [] => [] | z#zs => (z,y)#zip zs ys)"
   166 -- {* Warning: simpset does not contain this definition *}
   167 -- {* but separate theorems for @{text "xs = []"} and @{text "xs = z # zs"} *}
   168 primrec
   169 upt_0: "[i..0(] = []"
   170 upt_Suc: "[i..(Suc j)(] = (if i <= j then [i..j(] @ [j] else [])"
   171 primrec
   172 "distinct [] = True"
   173 "distinct (x#xs) = (x ~: set xs \<and> distinct xs)"
   174 primrec
   175 "remdups [] = []"
   176 "remdups (x#xs) = (if x : set xs then remdups xs else x # remdups xs)"
   177 primrec
   178 "remove1 x [] = []"
   179 "remove1 x (y#xs) = (if x=y then xs else y # remove1 x xs)"
   180 primrec
   181 replicate_0: "replicate 0 x = []"
   182 replicate_Suc: "replicate (Suc n) x = x # replicate n x"
   183 defs
   184  list_all2_def:
   185  "list_all2 P xs ys == length xs = length ys \<and> (\<forall>(x, y) \<in> set (zip xs ys). P x y)"
   186 
   187 
   188 subsection {* Lexicographic orderings on lists *}
   189 
   190 consts
   191 lexn :: "('a * 'a)set => nat => ('a list * 'a list)set"
   192 primrec
   193 "lexn r 0 = {}"
   194 "lexn r (Suc n) =
   195 (prod_fun (%(x,xs). x#xs) (%(x,xs). x#xs) ` (r <*lex*> lexn r n)) Int
   196 {(xs,ys). length xs = Suc n \<and> length ys = Suc n}"
   197 
   198 constdefs
   199 lex :: "('a \<times> 'a) set => ('a list \<times> 'a list) set"
   200 "lex r == \<Union>n. lexn r n"
   201 
   202 lexico :: "('a \<times> 'a) set => ('a list \<times> 'a list) set"
   203 "lexico r == inv_image (less_than <*lex*> lex r) (%xs. (length xs, xs))"
   204 
   205 sublist :: "'a list => nat set => 'a list"
   206 "sublist xs A == map fst (filter (%p. snd p : A) (zip xs [0..size xs(]))"
   207 
   208 
   209 lemma not_Cons_self [simp]: "xs \<noteq> x # xs"
   210 by (induct xs) auto
   211 
   212 lemmas not_Cons_self2 [simp] = not_Cons_self [symmetric]
   213 
   214 lemma neq_Nil_conv: "(xs \<noteq> []) = (\<exists>y ys. xs = y # ys)"
   215 by (induct xs) auto
   216 
   217 lemma length_induct:
   218 "(!!xs. \<forall>ys. length ys < length xs --> P ys ==> P xs) ==> P xs"
   219 by (rule measure_induct [of length]) rules
   220 
   221 
   222 subsection {* @{text lists}: the list-forming operator over sets *}
   223 
   224 consts lists :: "'a set => 'a list set"
   225 inductive "lists A"
   226  intros
   227   Nil [intro!]: "[]: lists A"
   228   Cons [intro!]: "[| a: A;l: lists A|] ==> a#l : lists A"
   229 
   230 inductive_cases listsE [elim!]: "x#l : lists A"
   231 
   232 lemma lists_mono [mono]: "A \<subseteq> B ==> lists A \<subseteq> lists B"
   233 by (unfold lists.defs) (blast intro!: lfp_mono)
   234 
   235 lemma lists_IntI:
   236   assumes l: "l: lists A" shows "l: lists B ==> l: lists (A Int B)" using l
   237   by induct blast+
   238 
   239 lemma lists_Int_eq [simp]: "lists (A \<inter> B) = lists A \<inter> lists B"
   240 proof (rule mono_Int [THEN equalityI])
   241   show "mono lists" by (simp add: mono_def lists_mono)
   242   show "lists A \<inter> lists B \<subseteq> lists (A \<inter> B)" by (blast intro: lists_IntI)
   243 qed
   244 
   245 lemma append_in_lists_conv [iff]:
   246      "(xs @ ys : lists A) = (xs : lists A \<and> ys : lists A)"
   247 by (induct xs) auto
   248 
   249 
   250 subsection {* @{text length} *}
   251 
   252 text {*
   253 Needs to come before @{text "@"} because of theorem @{text
   254 append_eq_append_conv}.
   255 *}
   256 
   257 lemma length_append [simp]: "length (xs @ ys) = length xs + length ys"
   258 by (induct xs) auto
   259 
   260 lemma length_map [simp]: "length (map f xs) = length xs"
   261 by (induct xs) auto
   262 
   263 lemma length_rev [simp]: "length (rev xs) = length xs"
   264 by (induct xs) auto
   265 
   266 lemma length_tl [simp]: "length (tl xs) = length xs - 1"
   267 by (cases xs) auto
   268 
   269 lemma length_0_conv [iff]: "(length xs = 0) = (xs = [])"
   270 by (induct xs) auto
   271 
   272 lemma length_greater_0_conv [iff]: "(0 < length xs) = (xs \<noteq> [])"
   273 by (induct xs) auto
   274 
   275 lemma length_Suc_conv:
   276 "(length xs = Suc n) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
   277 by (induct xs) auto
   278 
   279 lemma Suc_length_conv:
   280 "(Suc n = length xs) = (\<exists>y ys. xs = y # ys \<and> length ys = n)"
   281 apply (induct xs, simp, simp)
   282 apply blast
   283 done
   284 
   285 lemma impossible_Cons [rule_format]: 
   286   "length xs <= length ys --> xs = x # ys = False"
   287 apply (induct xs, auto)
   288 done
   289 
   290 lemma list_induct2[consumes 1]: "\<And>ys.
   291  \<lbrakk> length xs = length ys;
   292    P [] [];
   293    \<And>x xs y ys. \<lbrakk> length xs = length ys; P xs ys \<rbrakk> \<Longrightarrow> P (x#xs) (y#ys) \<rbrakk>
   294  \<Longrightarrow> P xs ys"
   295 apply(induct xs)
   296  apply simp
   297 apply(case_tac ys)
   298  apply simp
   299 apply(simp)
   300 done
   301 
   302 subsection {* @{text "@"} -- append *}
   303 
   304 lemma append_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)"
   305 by (induct xs) auto
   306 
   307 lemma append_Nil2 [simp]: "xs @ [] = xs"
   308 by (induct xs) auto
   309 
   310 lemma append_is_Nil_conv [iff]: "(xs @ ys = []) = (xs = [] \<and> ys = [])"
   311 by (induct xs) auto
   312 
   313 lemma Nil_is_append_conv [iff]: "([] = xs @ ys) = (xs = [] \<and> ys = [])"
   314 by (induct xs) auto
   315 
   316 lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])"
   317 by (induct xs) auto
   318 
   319 lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])"
   320 by (induct xs) auto
   321 
   322 lemma append_eq_append_conv [simp]:
   323  "!!ys. length xs = length ys \<or> length us = length vs
   324  ==> (xs@us = ys@vs) = (xs=ys \<and> us=vs)"
   325 apply (induct xs)
   326  apply (case_tac ys, simp, force)
   327 apply (case_tac ys, force, simp)
   328 done
   329 
   330 lemma append_eq_append_conv2: "!!ys zs ts.
   331  (xs @ ys = zs @ ts) =
   332  (EX us. xs = zs @ us & us @ ys = ts | xs @ us = zs & ys = us@ ts)"
   333 apply (induct xs)
   334  apply fastsimp
   335 apply(case_tac zs)
   336  apply simp
   337 apply fastsimp
   338 done
   339 
   340 lemma same_append_eq [iff]: "(xs @ ys = xs @ zs) = (ys = zs)"
   341 by simp
   342 
   343 lemma append1_eq_conv [iff]: "(xs @ [x] = ys @ [y]) = (xs = ys \<and> x = y)"
   344 by simp
   345 
   346 lemma append_same_eq [iff]: "(ys @ xs = zs @ xs) = (ys = zs)"
   347 by simp
   348 
   349 lemma append_self_conv2 [iff]: "(xs @ ys = ys) = (xs = [])"
   350 using append_same_eq [of _ _ "[]"] by auto
   351 
   352 lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])"
   353 using append_same_eq [of "[]"] by auto
   354 
   355 lemma hd_Cons_tl [simp]: "xs \<noteq> [] ==> hd xs # tl xs = xs"
   356 by (induct xs) auto
   357 
   358 lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)"
   359 by (induct xs) auto
   360 
   361 lemma hd_append2 [simp]: "xs \<noteq> [] ==> hd (xs @ ys) = hd xs"
   362 by (simp add: hd_append split: list.split)
   363 
   364 lemma tl_append: "tl (xs @ ys) = (case xs of [] => tl ys | z#zs => zs @ ys)"
   365 by (simp split: list.split)
   366 
   367 lemma tl_append2 [simp]: "xs \<noteq> [] ==> tl (xs @ ys) = tl xs @ ys"
   368 by (simp add: tl_append split: list.split)
   369 
   370 
   371 lemma Cons_eq_append_conv: "x#xs = ys@zs =
   372  (ys = [] & x#xs = zs | (EX ys'. x#ys' = ys & xs = ys'@zs))"
   373 by(cases ys) auto
   374 
   375 
   376 text {* Trivial rules for solving @{text "@"}-equations automatically. *}
   377 
   378 lemma eq_Nil_appendI: "xs = ys ==> xs = [] @ ys"
   379 by simp
   380 
   381 lemma Cons_eq_appendI:
   382 "[| x # xs1 = ys; xs = xs1 @ zs |] ==> x # xs = ys @ zs"
   383 by (drule sym) simp
   384 
   385 lemma append_eq_appendI:
   386 "[| xs @ xs1 = zs; ys = xs1 @ us |] ==> xs @ ys = zs @ us"
   387 by (drule sym) simp
   388 
   389 
   390 text {*
   391 Simplification procedure for all list equalities.
   392 Currently only tries to rearrange @{text "@"} to see if
   393 - both lists end in a singleton list,
   394 - or both lists end in the same list.
   395 *}
   396 
   397 ML_setup {*
   398 local
   399 
   400 val append_assoc = thm "append_assoc";
   401 val append_Nil = thm "append_Nil";
   402 val append_Cons = thm "append_Cons";
   403 val append1_eq_conv = thm "append1_eq_conv";
   404 val append_same_eq = thm "append_same_eq";
   405 
   406 fun last (cons as Const("List.list.Cons",_) $ _ $ xs) =
   407   (case xs of Const("List.list.Nil",_) => cons | _ => last xs)
   408   | last (Const("List.op @",_) $ _ $ ys) = last ys
   409   | last t = t;
   410 
   411 fun list1 (Const("List.list.Cons",_) $ _ $ Const("List.list.Nil",_)) = true
   412   | list1 _ = false;
   413 
   414 fun butlast ((cons as Const("List.list.Cons",_) $ x) $ xs) =
   415   (case xs of Const("List.list.Nil",_) => xs | _ => cons $ butlast xs)
   416   | butlast ((app as Const("List.op @",_) $ xs) $ ys) = app $ butlast ys
   417   | butlast xs = Const("List.list.Nil",fastype_of xs);
   418 
   419 val rearr_tac =
   420   simp_tac (HOL_basic_ss addsimps [append_assoc, append_Nil, append_Cons]);
   421 
   422 fun list_eq sg _ (F as (eq as Const(_,eqT)) $ lhs $ rhs) =
   423   let
   424     val lastl = last lhs and lastr = last rhs;
   425     fun rearr conv =
   426       let
   427         val lhs1 = butlast lhs and rhs1 = butlast rhs;
   428         val Type(_,listT::_) = eqT
   429         val appT = [listT,listT] ---> listT
   430         val app = Const("List.op @",appT)
   431         val F2 = eq $ (app$lhs1$lastl) $ (app$rhs1$lastr)
   432         val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (F,F2));
   433         val thm = Tactic.prove sg [] [] eq (K (rearr_tac 1));
   434       in Some ((conv RS (thm RS trans)) RS eq_reflection) end;
   435 
   436   in
   437     if list1 lastl andalso list1 lastr then rearr append1_eq_conv
   438     else if lastl aconv lastr then rearr append_same_eq
   439     else None
   440   end;
   441 
   442 in
   443 
   444 val list_eq_simproc =
   445   Simplifier.simproc (Theory.sign_of (the_context ())) "list_eq" ["(xs::'a list) = ys"] list_eq;
   446 
   447 end;
   448 
   449 Addsimprocs [list_eq_simproc];
   450 *}
   451 
   452 
   453 subsection {* @{text map} *}
   454 
   455 lemma map_ext: "(!!x. x : set xs --> f x = g x) ==> map f xs = map g xs"
   456 by (induct xs) simp_all
   457 
   458 lemma map_ident [simp]: "map (\<lambda>x. x) = (\<lambda>xs. xs)"
   459 by (rule ext, induct_tac xs) auto
   460 
   461 lemma map_append [simp]: "map f (xs @ ys) = map f xs @ map f ys"
   462 by (induct xs) auto
   463 
   464 lemma map_compose: "map (f o g) xs = map f (map g xs)"
   465 by (induct xs) (auto simp add: o_def)
   466 
   467 lemma rev_map: "rev (map f xs) = map f (rev xs)"
   468 by (induct xs) auto
   469 
   470 lemma map_eq_conv[simp]: "(map f xs = map g xs) = (!x : set xs. f x = g x)"
   471 by (induct xs) auto
   472 
   473 lemma map_cong [recdef_cong]:
   474 "xs = ys ==> (!!x. x : set ys ==> f x = g x) ==> map f xs = map g ys"
   475 -- {* a congruence rule for @{text map} *}
   476 by simp
   477 
   478 lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])"
   479 by (cases xs) auto
   480 
   481 lemma Nil_is_map_conv [iff]: "([] = map f xs) = (xs = [])"
   482 by (cases xs) auto
   483 
   484 lemma map_eq_Cons_conv[iff]:
   485  "(map f xs = y#ys) = (\<exists>z zs. xs = z#zs \<and> f z = y \<and> map f zs = ys)"
   486 by (cases xs) auto
   487 
   488 lemma Cons_eq_map_conv[iff]:
   489  "(x#xs = map f ys) = (\<exists>z zs. ys = z#zs \<and> x = f z \<and> xs = map f zs)"
   490 by (cases ys) auto
   491 
   492 lemma ex_map_conv:
   493   "(EX xs. ys = map f xs) = (ALL y : set ys. EX x. y = f x)"
   494 by(induct ys, auto)
   495 
   496 lemma map_eq_imp_length_eq:
   497   "!!xs. map f xs = map f ys ==> length xs = length ys"
   498 apply (induct ys)
   499  apply simp
   500 apply(simp (no_asm_use))
   501 apply clarify
   502 apply(simp (no_asm_use))
   503 apply fast
   504 done
   505 
   506 lemma map_inj_on:
   507  "[| map f xs = map f ys; inj_on f (set xs Un set ys) |]
   508   ==> xs = ys"
   509 apply(frule map_eq_imp_length_eq)
   510 apply(rotate_tac -1)
   511 apply(induct rule:list_induct2)
   512  apply simp
   513 apply(simp)
   514 apply (blast intro:sym)
   515 done
   516 
   517 lemma inj_on_map_eq_map:
   518  "inj_on f (set xs Un set ys) \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
   519 by(blast dest:map_inj_on)
   520 
   521 lemma map_injective:
   522  "!!xs. map f xs = map f ys ==> inj f ==> xs = ys"
   523 by (induct ys) (auto dest!:injD)
   524 
   525 lemma inj_map_eq_map[simp]: "inj f \<Longrightarrow> (map f xs = map f ys) = (xs = ys)"
   526 by(blast dest:map_injective)
   527 
   528 lemma inj_mapI: "inj f ==> inj (map f)"
   529 by (rules dest: map_injective injD intro: inj_onI)
   530 
   531 lemma inj_mapD: "inj (map f) ==> inj f"
   532 apply (unfold inj_on_def, clarify)
   533 apply (erule_tac x = "[x]" in ballE)
   534  apply (erule_tac x = "[y]" in ballE, simp, blast)
   535 apply blast
   536 done
   537 
   538 lemma inj_map[iff]: "inj (map f) = inj f"
   539 by (blast dest: inj_mapD intro: inj_mapI)
   540 
   541 lemma map_idI: "(\<And>x. x \<in> set xs \<Longrightarrow> f x = x) \<Longrightarrow> map f xs = xs"
   542 by (induct xs, auto)
   543 
   544 lemma map_fun_upd [simp]: "y \<notin> set xs \<Longrightarrow> map (f(y:=v)) xs = map f xs"
   545 by (induct xs) auto
   546 
   547 lemma map_fst_zip[simp]:
   548   "length xs = length ys \<Longrightarrow> map fst (zip xs ys) = xs"
   549 by (induct rule:list_induct2, simp_all)
   550 
   551 lemma map_snd_zip[simp]:
   552   "length xs = length ys \<Longrightarrow> map snd (zip xs ys) = ys"
   553 by (induct rule:list_induct2, simp_all)
   554 
   555 
   556 subsection {* @{text rev} *}
   557 
   558 lemma rev_append [simp]: "rev (xs @ ys) = rev ys @ rev xs"
   559 by (induct xs) auto
   560 
   561 lemma rev_rev_ident [simp]: "rev (rev xs) = xs"
   562 by (induct xs) auto
   563 
   564 lemma rev_is_Nil_conv [iff]: "(rev xs = []) = (xs = [])"
   565 by (induct xs) auto
   566 
   567 lemma Nil_is_rev_conv [iff]: "([] = rev xs) = (xs = [])"
   568 by (induct xs) auto
   569 
   570 lemma rev_is_rev_conv [iff]: "!!ys. (rev xs = rev ys) = (xs = ys)"
   571 apply (induct xs, force)
   572 apply (case_tac ys, simp, force)
   573 done
   574 
   575 lemma rev_induct [case_names Nil snoc]:
   576   "[| P []; !!x xs. P xs ==> P (xs @ [x]) |] ==> P xs"
   577 apply(subst rev_rev_ident[symmetric])
   578 apply(rule_tac list = "rev xs" in list.induct, simp_all)
   579 done
   580 
   581 ML {* val rev_induct_tac = induct_thm_tac (thm "rev_induct") *}-- "compatibility"
   582 
   583 lemma rev_exhaust [case_names Nil snoc]:
   584   "(xs = [] ==> P) ==>(!!ys y. xs = ys @ [y] ==> P) ==> P"
   585 by (induct xs rule: rev_induct) auto
   586 
   587 lemmas rev_cases = rev_exhaust
   588 
   589 
   590 subsection {* @{text set} *}
   591 
   592 lemma finite_set [iff]: "finite (set xs)"
   593 by (induct xs) auto
   594 
   595 lemma set_append [simp]: "set (xs @ ys) = (set xs \<union> set ys)"
   596 by (induct xs) auto
   597 
   598 lemma hd_in_set: "l = x#xs \<Longrightarrow> x\<in>set l"
   599 by (case_tac l, auto)
   600 
   601 lemma set_subset_Cons: "set xs \<subseteq> set (x # xs)"
   602 by auto
   603 
   604 lemma set_ConsD: "y \<in> set (x # xs) \<Longrightarrow> y=x \<or> y \<in> set xs" 
   605 by auto
   606 
   607 lemma set_empty [iff]: "(set xs = {}) = (xs = [])"
   608 by (induct xs) auto
   609 
   610 lemma set_rev [simp]: "set (rev xs) = set xs"
   611 by (induct xs) auto
   612 
   613 lemma set_map [simp]: "set (map f xs) = f`(set xs)"
   614 by (induct xs) auto
   615 
   616 lemma set_filter [simp]: "set (filter P xs) = {x. x : set xs \<and> P x}"
   617 by (induct xs) auto
   618 
   619 lemma set_upt [simp]: "set[i..j(] = {k. i \<le> k \<and> k < j}"
   620 apply (induct j, simp_all)
   621 apply (erule ssubst, auto)
   622 done
   623 
   624 lemma in_set_conv_decomp: "(x : set xs) = (\<exists>ys zs. xs = ys @ x # zs)"
   625 proof (induct xs)
   626   case Nil show ?case by simp
   627   case (Cons a xs)
   628   show ?case
   629   proof 
   630     assume "x \<in> set (a # xs)"
   631     with prems show "\<exists>ys zs. a # xs = ys @ x # zs"
   632       by (simp, blast intro: Cons_eq_appendI)
   633   next
   634     assume "\<exists>ys zs. a # xs = ys @ x # zs"
   635     then obtain ys zs where eq: "a # xs = ys @ x # zs" by blast
   636     show "x \<in> set (a # xs)" 
   637       by (cases ys, auto simp add: eq)
   638   qed
   639 qed
   640 
   641 lemma in_lists_conv_set: "(xs : lists A) = (\<forall>x \<in> set xs. x : A)"
   642 -- {* eliminate @{text lists} in favour of @{text set} *}
   643 by (induct xs) auto
   644 
   645 lemma in_listsD [dest!]: "xs \<in> lists A ==> \<forall>x\<in>set xs. x \<in> A"
   646 by (rule in_lists_conv_set [THEN iffD1])
   647 
   648 lemma in_listsI [intro!]: "\<forall>x\<in>set xs. x \<in> A ==> xs \<in> lists A"
   649 by (rule in_lists_conv_set [THEN iffD2])
   650 
   651 lemma lists_UNIV [simp]: "lists UNIV = UNIV"
   652 by auto
   653 
   654 lemma finite_list: "finite A ==> EX l. set l = A"
   655 apply (erule finite_induct, auto)
   656 apply (rule_tac x="x#l" in exI, auto)
   657 done
   658 
   659 lemma card_length: "card (set xs) \<le> length xs"
   660 by (induct xs) (auto simp add: card_insert_if)
   661 
   662 
   663 subsection{*Sets of Lists*}
   664 
   665 text{*Resembles a Cartesian product: it denotes the set of lists with head
   666   drawn from @{term A} and tail drawn from @{term Xs}.*}
   667 constdefs
   668   set_Cons :: "'a set \<Rightarrow> 'a list set \<Rightarrow> 'a list set"
   669    "set_Cons A XS == {z. \<exists>x xs. z = x#xs & x \<in> A & xs \<in> XS}"
   670 
   671 lemma [simp]: "set_Cons A {[]} = (%x. [x])`A"
   672 by (auto simp add: set_Cons_def)
   673 
   674 text{*Yields the set of lists, all of the same length as the argument and
   675 with elements drawn from the corresponding element of the argument.*}
   676 consts  listset :: "'a set list \<Rightarrow> 'a list set"
   677 primrec
   678    "listset []    = {[]}"
   679    "listset(A#As) = set_Cons A (listset As)"
   680 
   681 
   682 subsection {* @{text mem} *}
   683 
   684 lemma set_mem_eq: "(x mem xs) = (x : set xs)"
   685 by (induct xs) auto
   686 
   687 
   688 subsection {* @{text list_all} *}
   689 
   690 lemma list_all_conv: "list_all P xs = (\<forall>x \<in> set xs. P x)"
   691 by (induct xs) auto
   692 
   693 lemma list_all_append [simp]:
   694 "list_all P (xs @ ys) = (list_all P xs \<and> list_all P ys)"
   695 by (induct xs) auto
   696 
   697 
   698 subsection {* @{text filter} *}
   699 
   700 lemma filter_append [simp]: "filter P (xs @ ys) = filter P xs @ filter P ys"
   701 by (induct xs) auto
   702 
   703 lemma filter_filter [simp]: "filter P (filter Q xs) = filter (\<lambda>x. Q x \<and> P x) xs"
   704 by (induct xs) auto
   705 
   706 lemma filter_True [simp]: "\<forall>x \<in> set xs. P x ==> filter P xs = xs"
   707 by (induct xs) auto
   708 
   709 lemma filter_False [simp]: "\<forall>x \<in> set xs. \<not> P x ==> filter P xs = []"
   710 by (induct xs) auto
   711 
   712 lemma length_filter [simp]: "length (filter P xs) \<le> length xs"
   713 by (induct xs) (auto simp add: le_SucI)
   714 
   715 lemma filter_is_subset [simp]: "set (filter P xs) \<le> set xs"
   716 by auto
   717 
   718 
   719 subsection {* @{text concat} *}
   720 
   721 lemma concat_append [simp]: "concat (xs @ ys) = concat xs @ concat ys"
   722 by (induct xs) auto
   723 
   724 lemma concat_eq_Nil_conv [iff]: "(concat xss = []) = (\<forall>xs \<in> set xss. xs = [])"
   725 by (induct xss) auto
   726 
   727 lemma Nil_eq_concat_conv [iff]: "([] = concat xss) = (\<forall>xs \<in> set xss. xs = [])"
   728 by (induct xss) auto
   729 
   730 lemma set_concat [simp]: "set (concat xs) = \<Union>(set ` set xs)"
   731 by (induct xs) auto
   732 
   733 lemma map_concat: "map f (concat xs) = concat (map (map f) xs)"
   734 by (induct xs) auto
   735 
   736 lemma filter_concat: "filter p (concat xs) = concat (map (filter p) xs)"
   737 by (induct xs) auto
   738 
   739 lemma rev_concat: "rev (concat xs) = concat (map rev (rev xs))"
   740 by (induct xs) auto
   741 
   742 
   743 subsection {* @{text nth} *}
   744 
   745 lemma nth_Cons_0 [simp]: "(x # xs)!0 = x"
   746 by auto
   747 
   748 lemma nth_Cons_Suc [simp]: "(x # xs)!(Suc n) = xs!n"
   749 by auto
   750 
   751 declare nth.simps [simp del]
   752 
   753 lemma nth_append:
   754 "!!n. (xs @ ys)!n = (if n < length xs then xs!n else ys!(n - length xs))"
   755 apply (induct "xs", simp)
   756 apply (case_tac n, auto)
   757 done
   758 
   759 lemma nth_append_length [simp]: "(xs @ x # ys) ! length xs = x"
   760 by (induct "xs") auto
   761 
   762 lemma nth_append_length_plus[simp]: "(xs @ ys) ! (length xs + n) = ys ! n"
   763 by (induct "xs") auto
   764 
   765 lemma nth_map [simp]: "!!n. n < length xs ==> (map f xs)!n = f(xs!n)"
   766 apply (induct xs, simp)
   767 apply (case_tac n, auto)
   768 done
   769 
   770 lemma set_conv_nth: "set xs = {xs!i | i. i < length xs}"
   771 apply (induct_tac xs, simp, simp)
   772 apply safe
   773 apply (rule_tac x = 0 in exI, simp)
   774  apply (rule_tac x = "Suc i" in exI, simp)
   775 apply (case_tac i, simp)
   776 apply (rename_tac j)
   777 apply (rule_tac x = j in exI, simp)
   778 done
   779 
   780 lemma list_ball_nth: "[| n < length xs; !x : set xs. P x|] ==> P(xs!n)"
   781 by (auto simp add: set_conv_nth)
   782 
   783 lemma nth_mem [simp]: "n < length xs ==> xs!n : set xs"
   784 by (auto simp add: set_conv_nth)
   785 
   786 lemma all_nth_imp_all_set:
   787 "[| !i < length xs. P(xs!i); x : set xs|] ==> P x"
   788 by (auto simp add: set_conv_nth)
   789 
   790 lemma all_set_conv_all_nth:
   791 "(\<forall>x \<in> set xs. P x) = (\<forall>i. i < length xs --> P (xs ! i))"
   792 by (auto simp add: set_conv_nth)
   793 
   794 
   795 subsection {* @{text list_update} *}
   796 
   797 lemma length_list_update [simp]: "!!i. length(xs[i:=x]) = length xs"
   798 by (induct xs) (auto split: nat.split)
   799 
   800 lemma nth_list_update:
   801 "!!i j. i < length xs==> (xs[i:=x])!j = (if i = j then x else xs!j)"
   802 by (induct xs) (auto simp add: nth_Cons split: nat.split)
   803 
   804 lemma nth_list_update_eq [simp]: "i < length xs ==> (xs[i:=x])!i = x"
   805 by (simp add: nth_list_update)
   806 
   807 lemma nth_list_update_neq [simp]: "!!i j. i \<noteq> j ==> xs[i:=x]!j = xs!j"
   808 by (induct xs) (auto simp add: nth_Cons split: nat.split)
   809 
   810 lemma list_update_overwrite [simp]:
   811 "!!i. i < size xs ==> xs[i:=x, i:=y] = xs[i:=y]"
   812 by (induct xs) (auto split: nat.split)
   813 
   814 lemma list_update_id[simp]: "!!i. i < length xs ==> xs[i := xs!i] = xs"
   815 apply (induct xs, simp)
   816 apply(simp split:nat.splits)
   817 done
   818 
   819 lemma list_update_same_conv:
   820 "!!i. i < length xs ==> (xs[i := x] = xs) = (xs!i = x)"
   821 by (induct xs) (auto split: nat.split)
   822 
   823 lemma list_update_append1:
   824  "!!i. i < size xs \<Longrightarrow> (xs @ ys)[i:=x] = xs[i:=x] @ ys"
   825 apply (induct xs, simp)
   826 apply(simp split:nat.split)
   827 done
   828 
   829 lemma list_update_length [simp]:
   830  "(xs @ x # ys)[length xs := y] = (xs @ y # ys)"
   831 by (induct xs, auto)
   832 
   833 lemma update_zip:
   834 "!!i xy xs. length xs = length ys ==>
   835 (zip xs ys)[i:=xy] = zip (xs[i:=fst xy]) (ys[i:=snd xy])"
   836 by (induct ys) (auto, case_tac xs, auto split: nat.split)
   837 
   838 lemma set_update_subset_insert: "!!i. set(xs[i:=x]) <= insert x (set xs)"
   839 by (induct xs) (auto split: nat.split)
   840 
   841 lemma set_update_subsetI: "[| set xs <= A; x:A |] ==> set(xs[i := x]) <= A"
   842 by (blast dest!: set_update_subset_insert [THEN subsetD])
   843 
   844 
   845 subsection {* @{text last} and @{text butlast} *}
   846 
   847 lemma last_snoc [simp]: "last (xs @ [x]) = x"
   848 by (induct xs) auto
   849 
   850 lemma butlast_snoc [simp]: "butlast (xs @ [x]) = xs"
   851 by (induct xs) auto
   852 
   853 lemma last_ConsL: "xs = [] \<Longrightarrow> last(x#xs) = x"
   854 by(simp add:last.simps)
   855 
   856 lemma last_ConsR: "xs \<noteq> [] \<Longrightarrow> last(x#xs) = last xs"
   857 by(simp add:last.simps)
   858 
   859 lemma last_append: "last(xs @ ys) = (if ys = [] then last xs else last ys)"
   860 by (induct xs) (auto)
   861 
   862 lemma last_appendL[simp]: "ys = [] \<Longrightarrow> last(xs @ ys) = last xs"
   863 by(simp add:last_append)
   864 
   865 lemma last_appendR[simp]: "ys \<noteq> [] \<Longrightarrow> last(xs @ ys) = last ys"
   866 by(simp add:last_append)
   867 
   868 
   869 lemma length_butlast [simp]: "length (butlast xs) = length xs - 1"
   870 by (induct xs rule: rev_induct) auto
   871 
   872 lemma butlast_append:
   873 "!!ys. butlast (xs @ ys) = (if ys = [] then butlast xs else xs @ butlast ys)"
   874 by (induct xs) auto
   875 
   876 lemma append_butlast_last_id [simp]:
   877 "xs \<noteq> [] ==> butlast xs @ [last xs] = xs"
   878 by (induct xs) auto
   879 
   880 lemma in_set_butlastD: "x : set (butlast xs) ==> x : set xs"
   881 by (induct xs) (auto split: split_if_asm)
   882 
   883 lemma in_set_butlast_appendI:
   884 "x : set (butlast xs) | x : set (butlast ys) ==> x : set (butlast (xs @ ys))"
   885 by (auto dest: in_set_butlastD simp add: butlast_append)
   886 
   887 
   888 subsection {* @{text take} and @{text drop} *}
   889 
   890 lemma take_0 [simp]: "take 0 xs = []"
   891 by (induct xs) auto
   892 
   893 lemma drop_0 [simp]: "drop 0 xs = xs"
   894 by (induct xs) auto
   895 
   896 lemma take_Suc_Cons [simp]: "take (Suc n) (x # xs) = x # take n xs"
   897 by simp
   898 
   899 lemma drop_Suc_Cons [simp]: "drop (Suc n) (x # xs) = drop n xs"
   900 by simp
   901 
   902 declare take_Cons [simp del] and drop_Cons [simp del]
   903 
   904 lemma take_Suc: "xs ~= [] ==> take (Suc n) xs = hd xs # take n (tl xs)"
   905 by(clarsimp simp add:neq_Nil_conv)
   906 
   907 lemma drop_Suc: "drop (Suc n) xs = drop n (tl xs)"
   908 by(cases xs, simp_all)
   909 
   910 lemma drop_tl: "!!n. drop n (tl xs) = tl(drop n xs)"
   911 by(induct xs, simp_all add:drop_Cons drop_Suc split:nat.split)
   912 
   913 lemma nth_via_drop: "!!n. drop n xs = y#ys \<Longrightarrow> xs!n = y"
   914 apply (induct xs, simp)
   915 apply(simp add:drop_Cons nth_Cons split:nat.splits)
   916 done
   917 
   918 lemma take_Suc_conv_app_nth:
   919  "!!i. i < length xs \<Longrightarrow> take (Suc i) xs = take i xs @ [xs!i]"
   920 apply (induct xs, simp)
   921 apply (case_tac i, auto)
   922 done
   923 
   924 lemma drop_Suc_conv_tl:
   925   "!!i. i < length xs \<Longrightarrow> (xs!i) # (drop (Suc i) xs) = drop i xs"
   926 apply (induct xs, simp)
   927 apply (case_tac i, auto)
   928 done
   929 
   930 lemma length_take [simp]: "!!xs. length (take n xs) = min (length xs) n"
   931 by (induct n) (auto, case_tac xs, auto)
   932 
   933 lemma length_drop [simp]: "!!xs. length (drop n xs) = (length xs - n)"
   934 by (induct n) (auto, case_tac xs, auto)
   935 
   936 lemma take_all [simp]: "!!xs. length xs <= n ==> take n xs = xs"
   937 by (induct n) (auto, case_tac xs, auto)
   938 
   939 lemma drop_all [simp]: "!!xs. length xs <= n ==> drop n xs = []"
   940 by (induct n) (auto, case_tac xs, auto)
   941 
   942 lemma take_append [simp]:
   943 "!!xs. take n (xs @ ys) = (take n xs @ take (n - length xs) ys)"
   944 by (induct n) (auto, case_tac xs, auto)
   945 
   946 lemma drop_append [simp]:
   947 "!!xs. drop n (xs @ ys) = drop n xs @ drop (n - length xs) ys"
   948 by (induct n) (auto, case_tac xs, auto)
   949 
   950 lemma take_take [simp]: "!!xs n. take n (take m xs) = take (min n m) xs"
   951 apply (induct m, auto)
   952 apply (case_tac xs, auto)
   953 apply (case_tac na, auto)
   954 done
   955 
   956 lemma drop_drop [simp]: "!!xs. drop n (drop m xs) = drop (n + m) xs"
   957 apply (induct m, auto)
   958 apply (case_tac xs, auto)
   959 done
   960 
   961 lemma take_drop: "!!xs n. take n (drop m xs) = drop m (take (n + m) xs)"
   962 apply (induct m, auto)
   963 apply (case_tac xs, auto)
   964 done
   965 
   966 lemma drop_take: "!!m n. drop n (take m xs) = take (m-n) (drop n xs)"
   967 apply(induct xs)
   968  apply simp
   969 apply(simp add: take_Cons drop_Cons split:nat.split)
   970 done
   971 
   972 lemma append_take_drop_id [simp]: "!!xs. take n xs @ drop n xs = xs"
   973 apply (induct n, auto)
   974 apply (case_tac xs, auto)
   975 done
   976 
   977 lemma take_eq_Nil[simp]: "!!n. (take n xs = []) = (n = 0 \<or> xs = [])"
   978 apply(induct xs)
   979  apply simp
   980 apply(simp add:take_Cons split:nat.split)
   981 done
   982 
   983 lemma drop_eq_Nil[simp]: "!!n. (drop n xs = []) = (length xs <= n)"
   984 apply(induct xs)
   985 apply simp
   986 apply(simp add:drop_Cons split:nat.split)
   987 done
   988 
   989 lemma take_map: "!!xs. take n (map f xs) = map f (take n xs)"
   990 apply (induct n, auto)
   991 apply (case_tac xs, auto)
   992 done
   993 
   994 lemma drop_map: "!!xs. drop n (map f xs) = map f (drop n xs)"
   995 apply (induct n, auto)
   996 apply (case_tac xs, auto)
   997 done
   998 
   999 lemma rev_take: "!!i. rev (take i xs) = drop (length xs - i) (rev xs)"
  1000 apply (induct xs, auto)
  1001 apply (case_tac i, auto)
  1002 done
  1003 
  1004 lemma rev_drop: "!!i. rev (drop i xs) = take (length xs - i) (rev xs)"
  1005 apply (induct xs, auto)
  1006 apply (case_tac i, auto)
  1007 done
  1008 
  1009 lemma nth_take [simp]: "!!n i. i < n ==> (take n xs)!i = xs!i"
  1010 apply (induct xs, auto)
  1011 apply (case_tac n, blast)
  1012 apply (case_tac i, auto)
  1013 done
  1014 
  1015 lemma nth_drop [simp]:
  1016 "!!xs i. n + i <= length xs ==> (drop n xs)!i = xs!(n + i)"
  1017 apply (induct n, auto)
  1018 apply (case_tac xs, auto)
  1019 done
  1020 
  1021 lemma set_take_subset: "\<And>n. set(take n xs) \<subseteq> set xs"
  1022 by(induct xs)(auto simp:take_Cons split:nat.split)
  1023 
  1024 lemma set_drop_subset: "\<And>n. set(drop n xs) \<subseteq> set xs"
  1025 by(induct xs)(auto simp:drop_Cons split:nat.split)
  1026 
  1027 lemma in_set_takeD: "x : set(take n xs) \<Longrightarrow> x : set xs"
  1028 using set_take_subset by fast
  1029 
  1030 lemma in_set_dropD: "x : set(drop n xs) \<Longrightarrow> x : set xs"
  1031 using set_drop_subset by fast
  1032 
  1033 lemma append_eq_conv_conj:
  1034 "!!zs. (xs @ ys = zs) = (xs = take (length xs) zs \<and> ys = drop (length xs) zs)"
  1035 apply (induct xs, simp, clarsimp)
  1036 apply (case_tac zs, auto)
  1037 done
  1038 
  1039 lemma take_add [rule_format]: 
  1040     "\<forall>i. i+j \<le> length(xs) --> take (i+j) xs = take i xs @ take j (drop i xs)"
  1041 apply (induct xs, auto) 
  1042 apply (case_tac i, simp_all) 
  1043 done
  1044 
  1045 lemma append_eq_append_conv_if:
  1046  "!! ys\<^isub>1. (xs\<^isub>1 @ xs\<^isub>2 = ys\<^isub>1 @ ys\<^isub>2) =
  1047   (if size xs\<^isub>1 \<le> size ys\<^isub>1
  1048    then xs\<^isub>1 = take (size xs\<^isub>1) ys\<^isub>1 \<and> xs\<^isub>2 = drop (size xs\<^isub>1) ys\<^isub>1 @ ys\<^isub>2
  1049    else take (size ys\<^isub>1) xs\<^isub>1 = ys\<^isub>1 \<and> drop (size ys\<^isub>1) xs\<^isub>1 @ xs\<^isub>2 = ys\<^isub>2)"
  1050 apply(induct xs\<^isub>1)
  1051  apply simp
  1052 apply(case_tac ys\<^isub>1)
  1053 apply simp_all
  1054 done
  1055 
  1056 lemma take_hd_drop:
  1057   "!!n. n < length xs \<Longrightarrow> take n xs @ [hd (drop n xs)] = take (n+1) xs"
  1058 apply(induct xs)
  1059 apply simp
  1060 apply(simp add:drop_Cons split:nat.split)
  1061 done
  1062 
  1063 
  1064 subsection {* @{text takeWhile} and @{text dropWhile} *}
  1065 
  1066 lemma takeWhile_dropWhile_id [simp]: "takeWhile P xs @ dropWhile P xs = xs"
  1067 by (induct xs) auto
  1068 
  1069 lemma takeWhile_append1 [simp]:
  1070 "[| x:set xs; ~P(x)|] ==> takeWhile P (xs @ ys) = takeWhile P xs"
  1071 by (induct xs) auto
  1072 
  1073 lemma takeWhile_append2 [simp]:
  1074 "(!!x. x : set xs ==> P x) ==> takeWhile P (xs @ ys) = xs @ takeWhile P ys"
  1075 by (induct xs) auto
  1076 
  1077 lemma takeWhile_tail: "\<not> P x ==> takeWhile P (xs @ (x#l)) = takeWhile P xs"
  1078 by (induct xs) auto
  1079 
  1080 lemma dropWhile_append1 [simp]:
  1081 "[| x : set xs; ~P(x)|] ==> dropWhile P (xs @ ys) = (dropWhile P xs)@ys"
  1082 by (induct xs) auto
  1083 
  1084 lemma dropWhile_append2 [simp]:
  1085 "(!!x. x:set xs ==> P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys"
  1086 by (induct xs) auto
  1087 
  1088 lemma set_take_whileD: "x : set (takeWhile P xs) ==> x : set xs \<and> P x"
  1089 by (induct xs) (auto split: split_if_asm)
  1090 
  1091 lemma takeWhile_eq_all_conv[simp]:
  1092  "(takeWhile P xs = xs) = (\<forall>x \<in> set xs. P x)"
  1093 by(induct xs, auto)
  1094 
  1095 lemma dropWhile_eq_Nil_conv[simp]:
  1096  "(dropWhile P xs = []) = (\<forall>x \<in> set xs. P x)"
  1097 by(induct xs, auto)
  1098 
  1099 lemma dropWhile_eq_Cons_conv:
  1100  "(dropWhile P xs = y#ys) = (xs = takeWhile P xs @ y # ys & \<not> P y)"
  1101 by(induct xs, auto)
  1102 
  1103 
  1104 subsection {* @{text zip} *}
  1105 
  1106 lemma zip_Nil [simp]: "zip [] ys = []"
  1107 by (induct ys) auto
  1108 
  1109 lemma zip_Cons_Cons [simp]: "zip (x # xs) (y # ys) = (x, y) # zip xs ys"
  1110 by simp
  1111 
  1112 declare zip_Cons [simp del]
  1113 
  1114 lemma length_zip [simp]:
  1115 "!!xs. length (zip xs ys) = min (length xs) (length ys)"
  1116 apply (induct ys, simp)
  1117 apply (case_tac xs, auto)
  1118 done
  1119 
  1120 lemma zip_append1:
  1121 "!!xs. zip (xs @ ys) zs =
  1122 zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)"
  1123 apply (induct zs, simp)
  1124 apply (case_tac xs, simp_all)
  1125 done
  1126 
  1127 lemma zip_append2:
  1128 "!!ys. zip xs (ys @ zs) =
  1129 zip (take (length ys) xs) ys @ zip (drop (length ys) xs) zs"
  1130 apply (induct xs, simp)
  1131 apply (case_tac ys, simp_all)
  1132 done
  1133 
  1134 lemma zip_append [simp]:
  1135  "[| length xs = length us; length ys = length vs |] ==>
  1136 zip (xs@ys) (us@vs) = zip xs us @ zip ys vs"
  1137 by (simp add: zip_append1)
  1138 
  1139 lemma zip_rev:
  1140 "length xs = length ys ==> zip (rev xs) (rev ys) = rev (zip xs ys)"
  1141 by (induct rule:list_induct2, simp_all)
  1142 
  1143 lemma nth_zip [simp]:
  1144 "!!i xs. [| i < length xs; i < length ys|] ==> (zip xs ys)!i = (xs!i, ys!i)"
  1145 apply (induct ys, simp)
  1146 apply (case_tac xs)
  1147  apply (simp_all add: nth.simps split: nat.split)
  1148 done
  1149 
  1150 lemma set_zip:
  1151 "set (zip xs ys) = {(xs!i, ys!i) | i. i < min (length xs) (length ys)}"
  1152 by (simp add: set_conv_nth cong: rev_conj_cong)
  1153 
  1154 lemma zip_update:
  1155 "length xs = length ys ==> zip (xs[i:=x]) (ys[i:=y]) = (zip xs ys)[i:=(x,y)]"
  1156 by (rule sym, simp add: update_zip)
  1157 
  1158 lemma zip_replicate [simp]:
  1159 "!!j. zip (replicate i x) (replicate j y) = replicate (min i j) (x,y)"
  1160 apply (induct i, auto)
  1161 apply (case_tac j, auto)
  1162 done
  1163 
  1164 
  1165 subsection {* @{text list_all2} *}
  1166 
  1167 lemma list_all2_lengthD [intro?]: 
  1168   "list_all2 P xs ys ==> length xs = length ys"
  1169 by (simp add: list_all2_def)
  1170 
  1171 lemma list_all2_Nil [iff]: "list_all2 P [] ys = (ys = [])"
  1172 by (simp add: list_all2_def)
  1173 
  1174 lemma list_all2_Nil2[iff]: "list_all2 P xs [] = (xs = [])"
  1175 by (simp add: list_all2_def)
  1176 
  1177 lemma list_all2_Cons [iff]:
  1178 "list_all2 P (x # xs) (y # ys) = (P x y \<and> list_all2 P xs ys)"
  1179 by (auto simp add: list_all2_def)
  1180 
  1181 lemma list_all2_Cons1:
  1182 "list_all2 P (x # xs) ys = (\<exists>z zs. ys = z # zs \<and> P x z \<and> list_all2 P xs zs)"
  1183 by (cases ys) auto
  1184 
  1185 lemma list_all2_Cons2:
  1186 "list_all2 P xs (y # ys) = (\<exists>z zs. xs = z # zs \<and> P z y \<and> list_all2 P zs ys)"
  1187 by (cases xs) auto
  1188 
  1189 lemma list_all2_rev [iff]:
  1190 "list_all2 P (rev xs) (rev ys) = list_all2 P xs ys"
  1191 by (simp add: list_all2_def zip_rev cong: conj_cong)
  1192 
  1193 lemma list_all2_rev1:
  1194 "list_all2 P (rev xs) ys = list_all2 P xs (rev ys)"
  1195 by (subst list_all2_rev [symmetric]) simp
  1196 
  1197 lemma list_all2_append1:
  1198 "list_all2 P (xs @ ys) zs =
  1199 (EX us vs. zs = us @ vs \<and> length us = length xs \<and> length vs = length ys \<and>
  1200 list_all2 P xs us \<and> list_all2 P ys vs)"
  1201 apply (simp add: list_all2_def zip_append1)
  1202 apply (rule iffI)
  1203  apply (rule_tac x = "take (length xs) zs" in exI)
  1204  apply (rule_tac x = "drop (length xs) zs" in exI)
  1205  apply (force split: nat_diff_split simp add: min_def, clarify)
  1206 apply (simp add: ball_Un)
  1207 done
  1208 
  1209 lemma list_all2_append2:
  1210 "list_all2 P xs (ys @ zs) =
  1211 (EX us vs. xs = us @ vs \<and> length us = length ys \<and> length vs = length zs \<and>
  1212 list_all2 P us ys \<and> list_all2 P vs zs)"
  1213 apply (simp add: list_all2_def zip_append2)
  1214 apply (rule iffI)
  1215  apply (rule_tac x = "take (length ys) xs" in exI)
  1216  apply (rule_tac x = "drop (length ys) xs" in exI)
  1217  apply (force split: nat_diff_split simp add: min_def, clarify)
  1218 apply (simp add: ball_Un)
  1219 done
  1220 
  1221 lemma list_all2_append:
  1222   "length xs = length ys \<Longrightarrow>
  1223   list_all2 P (xs@us) (ys@vs) = (list_all2 P xs ys \<and> list_all2 P us vs)"
  1224 by (induct rule:list_induct2, simp_all)
  1225 
  1226 lemma list_all2_appendI [intro?, trans]:
  1227   "\<lbrakk> list_all2 P a b; list_all2 P c d \<rbrakk> \<Longrightarrow> list_all2 P (a@c) (b@d)"
  1228   by (simp add: list_all2_append list_all2_lengthD)
  1229 
  1230 lemma list_all2_conv_all_nth:
  1231 "list_all2 P xs ys =
  1232 (length xs = length ys \<and> (\<forall>i < length xs. P (xs!i) (ys!i)))"
  1233 by (force simp add: list_all2_def set_zip)
  1234 
  1235 lemma list_all2_trans:
  1236   assumes tr: "!!a b c. P1 a b ==> P2 b c ==> P3 a c"
  1237   shows "!!bs cs. list_all2 P1 as bs ==> list_all2 P2 bs cs ==> list_all2 P3 as cs"
  1238         (is "!!bs cs. PROP ?Q as bs cs")
  1239 proof (induct as)
  1240   fix x xs bs assume I1: "!!bs cs. PROP ?Q xs bs cs"
  1241   show "!!cs. PROP ?Q (x # xs) bs cs"
  1242   proof (induct bs)
  1243     fix y ys cs assume I2: "!!cs. PROP ?Q (x # xs) ys cs"
  1244     show "PROP ?Q (x # xs) (y # ys) cs"
  1245       by (induct cs) (auto intro: tr I1 I2)
  1246   qed simp
  1247 qed simp
  1248 
  1249 lemma list_all2_all_nthI [intro?]:
  1250   "length a = length b \<Longrightarrow> (\<And>n. n < length a \<Longrightarrow> P (a!n) (b!n)) \<Longrightarrow> list_all2 P a b"
  1251   by (simp add: list_all2_conv_all_nth)
  1252 
  1253 lemma list_all2I:
  1254   "\<forall>x \<in> set (zip a b). split P x \<Longrightarrow> length a = length b \<Longrightarrow> list_all2 P a b"
  1255   by (simp add: list_all2_def)
  1256 
  1257 lemma list_all2_nthD:
  1258   "\<lbrakk> list_all2 P xs ys; p < size xs \<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)"
  1259   by (simp add: list_all2_conv_all_nth)
  1260 
  1261 lemma list_all2_nthD2:
  1262   "\<lbrakk>list_all2 P xs ys; p < size ys\<rbrakk> \<Longrightarrow> P (xs!p) (ys!p)"
  1263   by (frule list_all2_lengthD) (auto intro: list_all2_nthD)
  1264 
  1265 lemma list_all2_map1: 
  1266   "list_all2 P (map f as) bs = list_all2 (\<lambda>x y. P (f x) y) as bs"
  1267   by (simp add: list_all2_conv_all_nth)
  1268 
  1269 lemma list_all2_map2: 
  1270   "list_all2 P as (map f bs) = list_all2 (\<lambda>x y. P x (f y)) as bs"
  1271   by (auto simp add: list_all2_conv_all_nth)
  1272 
  1273 lemma list_all2_refl [intro?]:
  1274   "(\<And>x. P x x) \<Longrightarrow> list_all2 P xs xs"
  1275   by (simp add: list_all2_conv_all_nth)
  1276 
  1277 lemma list_all2_update_cong:
  1278   "\<lbrakk> i<size xs; list_all2 P xs ys; P x y \<rbrakk> \<Longrightarrow> list_all2 P (xs[i:=x]) (ys[i:=y])"
  1279   by (simp add: list_all2_conv_all_nth nth_list_update)
  1280 
  1281 lemma list_all2_update_cong2:
  1282   "\<lbrakk>list_all2 P xs ys; P x y; i < length ys\<rbrakk> \<Longrightarrow> list_all2 P (xs[i:=x]) (ys[i:=y])"
  1283   by (simp add: list_all2_lengthD list_all2_update_cong)
  1284 
  1285 lemma list_all2_takeI [simp,intro?]:
  1286   "\<And>n ys. list_all2 P xs ys \<Longrightarrow> list_all2 P (take n xs) (take n ys)"
  1287   apply (induct xs)
  1288    apply simp
  1289   apply (clarsimp simp add: list_all2_Cons1)
  1290   apply (case_tac n)
  1291   apply auto
  1292   done
  1293 
  1294 lemma list_all2_dropI [simp,intro?]:
  1295   "\<And>n bs. list_all2 P as bs \<Longrightarrow> list_all2 P (drop n as) (drop n bs)"
  1296   apply (induct as, simp)
  1297   apply (clarsimp simp add: list_all2_Cons1)
  1298   apply (case_tac n, simp, simp)
  1299   done
  1300 
  1301 lemma list_all2_mono [intro?]:
  1302   "\<And>y. list_all2 P x y \<Longrightarrow> (\<And>x y. P x y \<Longrightarrow> Q x y) \<Longrightarrow> list_all2 Q x y"
  1303   apply (induct x, simp)
  1304   apply (case_tac y, auto)
  1305   done
  1306 
  1307 
  1308 subsection {* @{text foldl} and @{text foldr} *}
  1309 
  1310 lemma foldl_append [simp]:
  1311 "!!a. foldl f a (xs @ ys) = foldl f (foldl f a xs) ys"
  1312 by (induct xs) auto
  1313 
  1314 lemma foldr_append[simp]: "foldr f (xs @ ys) a = foldr f xs (foldr f ys a)"
  1315 by (induct xs) auto
  1316 
  1317 lemma foldr_foldl: "foldr f xs a = foldl (%x y. f y x) a (rev xs)"
  1318 by (induct xs) auto
  1319 
  1320 lemma foldl_foldr: "foldl f a xs = foldr (%x y. f y x) (rev xs) a"
  1321 by (simp add: foldr_foldl [of "%x y. f y x" "rev xs"])
  1322 
  1323 text {*
  1324 Note: @{text "n \<le> foldl (op +) n ns"} looks simpler, but is more
  1325 difficult to use because it requires an additional transitivity step.
  1326 *}
  1327 
  1328 lemma start_le_sum: "!!n::nat. m <= n ==> m <= foldl (op +) n ns"
  1329 by (induct ns) auto
  1330 
  1331 lemma elem_le_sum: "!!n::nat. n : set ns ==> n <= foldl (op +) 0 ns"
  1332 by (force intro: start_le_sum simp add: in_set_conv_decomp)
  1333 
  1334 lemma sum_eq_0_conv [iff]:
  1335 "!!m::nat. (foldl (op +) m ns = 0) = (m = 0 \<and> (\<forall>n \<in> set ns. n = 0))"
  1336 by (induct ns) auto
  1337 
  1338 
  1339 subsection {* @{text upto} *}
  1340 
  1341 lemma upt_rec: "[i..j(] = (if i<j then i#[Suc i..j(] else [])"
  1342 -- {* Does not terminate! *}
  1343 by (induct j) auto
  1344 
  1345 lemma upt_conv_Nil [simp]: "j <= i ==> [i..j(] = []"
  1346 by (subst upt_rec) simp
  1347 
  1348 lemma upt_Suc_append: "i <= j ==> [i..(Suc j)(] = [i..j(]@[j]"
  1349 -- {* Only needed if @{text upt_Suc} is deleted from the simpset. *}
  1350 by simp
  1351 
  1352 lemma upt_conv_Cons: "i < j ==> [i..j(] = i # [Suc i..j(]"
  1353 apply(rule trans)
  1354 apply(subst upt_rec)
  1355  prefer 2 apply (rule refl, simp)
  1356 done
  1357 
  1358 lemma upt_add_eq_append: "i<=j ==> [i..j+k(] = [i..j(]@[j..j+k(]"
  1359 -- {* LOOPS as a simprule, since @{text "j <= j"}. *}
  1360 by (induct k) auto
  1361 
  1362 lemma length_upt [simp]: "length [i..j(] = j - i"
  1363 by (induct j) (auto simp add: Suc_diff_le)
  1364 
  1365 lemma nth_upt [simp]: "i + k < j ==> [i..j(] ! k = i + k"
  1366 apply (induct j)
  1367 apply (auto simp add: less_Suc_eq nth_append split: nat_diff_split)
  1368 done
  1369 
  1370 lemma take_upt [simp]: "!!i. i+m <= n ==> take m [i..n(] = [i..i+m(]"
  1371 apply (induct m, simp)
  1372 apply (subst upt_rec)
  1373 apply (rule sym)
  1374 apply (subst upt_rec)
  1375 apply (simp del: upt.simps)
  1376 done
  1377 
  1378 lemma map_Suc_upt: "map Suc [m..n(] = [Suc m..n]"
  1379 by (induct n) auto
  1380 
  1381 lemma nth_map_upt: "!!i. i < n-m ==> (map f [m..n(]) ! i = f(m+i)"
  1382 apply (induct n m rule: diff_induct)
  1383 prefer 3 apply (subst map_Suc_upt[symmetric])
  1384 apply (auto simp add: less_diff_conv nth_upt)
  1385 done
  1386 
  1387 lemma nth_take_lemma:
  1388   "!!xs ys. k <= length xs ==> k <= length ys ==>
  1389      (!!i. i < k --> xs!i = ys!i) ==> take k xs = take k ys"
  1390 apply (atomize, induct k)
  1391 apply (simp_all add: less_Suc_eq_0_disj all_conj_distrib, clarify)
  1392 txt {* Both lists must be non-empty *}
  1393 apply (case_tac xs, simp)
  1394 apply (case_tac ys, clarify)
  1395  apply (simp (no_asm_use))
  1396 apply clarify
  1397 txt {* prenexing's needed, not miniscoping *}
  1398 apply (simp (no_asm_use) add: all_simps [symmetric] del: all_simps)
  1399 apply blast
  1400 done
  1401 
  1402 lemma nth_equalityI:
  1403  "[| length xs = length ys; ALL i < length xs. xs!i = ys!i |] ==> xs = ys"
  1404 apply (frule nth_take_lemma [OF le_refl eq_imp_le])
  1405 apply (simp_all add: take_all)
  1406 done
  1407 
  1408 (* needs nth_equalityI *)
  1409 lemma list_all2_antisym:
  1410   "\<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> 
  1411   \<Longrightarrow> xs = ys"
  1412   apply (simp add: list_all2_conv_all_nth) 
  1413   apply (rule nth_equalityI, blast, simp)
  1414   done
  1415 
  1416 lemma take_equalityI: "(\<forall>i. take i xs = take i ys) ==> xs = ys"
  1417 -- {* The famous take-lemma. *}
  1418 apply (drule_tac x = "max (length xs) (length ys)" in spec)
  1419 apply (simp add: le_max_iff_disj take_all)
  1420 done
  1421 
  1422 
  1423 subsection {* @{text "distinct"} and @{text remdups} *}
  1424 
  1425 lemma distinct_append [simp]:
  1426 "distinct (xs @ ys) = (distinct xs \<and> distinct ys \<and> set xs \<inter> set ys = {})"
  1427 by (induct xs) auto
  1428 
  1429 lemma set_remdups [simp]: "set (remdups xs) = set xs"
  1430 by (induct xs) (auto simp add: insert_absorb)
  1431 
  1432 lemma distinct_remdups [iff]: "distinct (remdups xs)"
  1433 by (induct xs) auto
  1434 
  1435 lemma remdups_eq_nil_iff [simp]: "(remdups x = []) = (x = [])"
  1436   by (induct_tac x, auto) 
  1437 
  1438 lemma remdups_eq_nil_right_iff [simp]: "([] = remdups x) = (x = [])"
  1439   by (induct_tac x, auto)
  1440 
  1441 lemma distinct_filter [simp]: "distinct xs ==> distinct (filter P xs)"
  1442 by (induct xs) auto
  1443 
  1444 text {*
  1445 It is best to avoid this indexed version of distinct, but sometimes
  1446 it is useful. *}
  1447 lemma distinct_conv_nth:
  1448 "distinct xs = (\<forall>i j. i < size xs \<and> j < size xs \<and> i \<noteq> j --> xs!i \<noteq> xs!j)"
  1449 apply (induct_tac xs, simp, simp)
  1450 apply (rule iffI, clarsimp)
  1451  apply (case_tac i)
  1452 apply (case_tac j, simp)
  1453 apply (simp add: set_conv_nth)
  1454  apply (case_tac j)
  1455 apply (clarsimp simp add: set_conv_nth, simp)
  1456 apply (rule conjI)
  1457  apply (clarsimp simp add: set_conv_nth)
  1458  apply (erule_tac x = 0 in allE)
  1459  apply (erule_tac x = "Suc i" in allE, simp, clarsimp)
  1460 apply (erule_tac x = "Suc i" in allE)
  1461 apply (erule_tac x = "Suc j" in allE, simp)
  1462 done
  1463 
  1464 lemma distinct_card: "distinct xs ==> card (set xs) = size xs"
  1465   by (induct xs) auto
  1466 
  1467 lemma card_distinct: "card (set xs) = size xs ==> distinct xs"
  1468 proof (induct xs)
  1469   case Nil thus ?case by simp
  1470 next
  1471   case (Cons x xs)
  1472   show ?case
  1473   proof (cases "x \<in> set xs")
  1474     case False with Cons show ?thesis by simp
  1475   next
  1476     case True with Cons.prems
  1477     have "card (set xs) = Suc (length xs)" 
  1478       by (simp add: card_insert_if split: split_if_asm)
  1479     moreover have "card (set xs) \<le> length xs" by (rule card_length)
  1480     ultimately have False by simp
  1481     thus ?thesis ..
  1482   qed
  1483 qed
  1484 
  1485 lemma inj_on_setI: "distinct(map f xs) ==> inj_on f (set xs)"
  1486 apply(induct xs)
  1487  apply simp
  1488 apply fastsimp
  1489 done
  1490 
  1491 lemma inj_on_set_conv:
  1492  "distinct xs \<Longrightarrow> inj_on f (set xs) = distinct(map f xs)"
  1493 apply(induct xs)
  1494  apply simp
  1495 apply fastsimp
  1496 done
  1497 
  1498 
  1499 subsection {* @{text remove1} *}
  1500 
  1501 lemma set_remove1_subset: "set(remove1 x xs) <= set xs"
  1502 apply(induct xs)
  1503  apply simp
  1504 apply simp
  1505 apply blast
  1506 done
  1507 
  1508 lemma [simp]: "distinct xs ==> set(remove1 x xs) = set xs - {x}"
  1509 apply(induct xs)
  1510  apply simp
  1511 apply simp
  1512 apply blast
  1513 done
  1514 
  1515 lemma notin_set_remove1[simp]: "x ~: set xs ==> x ~: set(remove1 y xs)"
  1516 apply(insert set_remove1_subset)
  1517 apply fast
  1518 done
  1519 
  1520 lemma distinct_remove1[simp]: "distinct xs ==> distinct(remove1 x xs)"
  1521 by (induct xs) simp_all
  1522 
  1523 
  1524 subsection {* @{text replicate} *}
  1525 
  1526 lemma length_replicate [simp]: "length (replicate n x) = n"
  1527 by (induct n) auto
  1528 
  1529 lemma map_replicate [simp]: "map f (replicate n x) = replicate n (f x)"
  1530 by (induct n) auto
  1531 
  1532 lemma replicate_app_Cons_same:
  1533 "(replicate n x) @ (x # xs) = x # replicate n x @ xs"
  1534 by (induct n) auto
  1535 
  1536 lemma rev_replicate [simp]: "rev (replicate n x) = replicate n x"
  1537 apply (induct n, simp)
  1538 apply (simp add: replicate_app_Cons_same)
  1539 done
  1540 
  1541 lemma replicate_add: "replicate (n + m) x = replicate n x @ replicate m x"
  1542 by (induct n) auto
  1543 
  1544 lemma hd_replicate [simp]: "n \<noteq> 0 ==> hd (replicate n x) = x"
  1545 by (induct n) auto
  1546 
  1547 lemma tl_replicate [simp]: "n \<noteq> 0 ==> tl (replicate n x) = replicate (n - 1) x"
  1548 by (induct n) auto
  1549 
  1550 lemma last_replicate [simp]: "n \<noteq> 0 ==> last (replicate n x) = x"
  1551 by (atomize (full), induct n) auto
  1552 
  1553 lemma nth_replicate[simp]: "!!i. i < n ==> (replicate n x)!i = x"
  1554 apply (induct n, simp)
  1555 apply (simp add: nth_Cons split: nat.split)
  1556 done
  1557 
  1558 lemma set_replicate_Suc: "set (replicate (Suc n) x) = {x}"
  1559 by (induct n) auto
  1560 
  1561 lemma set_replicate [simp]: "n \<noteq> 0 ==> set (replicate n x) = {x}"
  1562 by (fast dest!: not0_implies_Suc intro!: set_replicate_Suc)
  1563 
  1564 lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})"
  1565 by auto
  1566 
  1567 lemma in_set_replicateD: "x : set (replicate n y) ==> x = y"
  1568 by (simp add: set_replicate_conv_if split: split_if_asm)
  1569 
  1570 
  1571 subsection {* Lexicographic orderings on lists *}
  1572 
  1573 lemma wf_lexn: "wf r ==> wf (lexn r n)"
  1574 apply (induct_tac n, simp, simp)
  1575 apply(rule wf_subset)
  1576  prefer 2 apply (rule Int_lower1)
  1577 apply(rule wf_prod_fun_image)
  1578  prefer 2 apply (rule inj_onI, auto)
  1579 done
  1580 
  1581 lemma lexn_length:
  1582      "!!xs ys. (xs, ys) : lexn r n ==> length xs = n \<and> length ys = n"
  1583 by (induct n) auto
  1584 
  1585 lemma wf_lex [intro!]: "wf r ==> wf (lex r)"
  1586 apply (unfold lex_def)
  1587 apply (rule wf_UN)
  1588 apply (blast intro: wf_lexn, clarify)
  1589 apply (rename_tac m n)
  1590 apply (subgoal_tac "m \<noteq> n")
  1591  prefer 2 apply blast
  1592 apply (blast dest: lexn_length not_sym)
  1593 done
  1594 
  1595 lemma lexn_conv:
  1596 "lexn r n =
  1597 {(xs,ys). length xs = n \<and> length ys = n \<and>
  1598 (\<exists>xys x y xs' ys'. xs= xys @ x#xs' \<and> ys= xys @ y # ys' \<and> (x, y):r)}"
  1599 apply (induct_tac n, simp, blast)
  1600 apply (simp add: image_Collect lex_prod_def, safe, blast)
  1601  apply (rule_tac x = "ab # xys" in exI, simp)
  1602 apply (case_tac xys, simp_all, blast)
  1603 done
  1604 
  1605 lemma lex_conv:
  1606 "lex r =
  1607 {(xs,ys). length xs = length ys \<and>
  1608 (\<exists>xys x y xs' ys'. xs = xys @ x # xs' \<and> ys = xys @ y # ys' \<and> (x, y):r)}"
  1609 by (force simp add: lex_def lexn_conv)
  1610 
  1611 lemma wf_lexico [intro!]: "wf r ==> wf (lexico r)"
  1612 by (unfold lexico_def) blast
  1613 
  1614 lemma lexico_conv:
  1615 "lexico r = {(xs,ys). length xs < length ys |
  1616 length xs = length ys \<and> (xs, ys) : lex r}"
  1617 by (simp add: lexico_def diag_def lex_prod_def measure_def inv_image_def)
  1618 
  1619 lemma Nil_notin_lex [iff]: "([], ys) \<notin> lex r"
  1620 by (simp add: lex_conv)
  1621 
  1622 lemma Nil2_notin_lex [iff]: "(xs, []) \<notin> lex r"
  1623 by (simp add:lex_conv)
  1624 
  1625 lemma Cons_in_lex [iff]:
  1626 "((x # xs, y # ys) : lex r) =
  1627 ((x, y) : r \<and> length xs = length ys | x = y \<and> (xs, ys) : lex r)"
  1628 apply (simp add: lex_conv)
  1629 apply (rule iffI)
  1630  prefer 2 apply (blast intro: Cons_eq_appendI, clarify)
  1631 apply (case_tac xys, simp, simp)
  1632 apply blast
  1633 done
  1634 
  1635 
  1636 subsection {* @{text sublist} --- a generalization of @{text nth} to sets *}
  1637 
  1638 lemma sublist_empty [simp]: "sublist xs {} = []"
  1639 by (auto simp add: sublist_def)
  1640 
  1641 lemma sublist_nil [simp]: "sublist [] A = []"
  1642 by (auto simp add: sublist_def)
  1643 
  1644 lemma sublist_shift_lemma:
  1645      "map fst [p:zip xs [i..i + length xs(] . snd p : A] =
  1646       map fst [p:zip xs [0..length xs(] . snd p + i : A]"
  1647 by (induct xs rule: rev_induct) (simp_all add: add_commute)
  1648 
  1649 lemma sublist_append:
  1650      "sublist (l @ l') A = sublist l A @ sublist l' {j. j + length l : A}"
  1651 apply (unfold sublist_def)
  1652 apply (induct l' rule: rev_induct, simp)
  1653 apply (simp add: upt_add_eq_append[of 0] zip_append sublist_shift_lemma)
  1654 apply (simp add: add_commute)
  1655 done
  1656 
  1657 lemma sublist_Cons:
  1658 "sublist (x # l) A = (if 0:A then [x] else []) @ sublist l {j. Suc j : A}"
  1659 apply (induct l rule: rev_induct)
  1660  apply (simp add: sublist_def)
  1661 apply (simp del: append_Cons add: append_Cons[symmetric] sublist_append)
  1662 done
  1663 
  1664 lemma sublist_singleton [simp]: "sublist [x] A = (if 0 : A then [x] else [])"
  1665 by (simp add: sublist_Cons)
  1666 
  1667 lemma sublist_upt_eq_take [simp]: "sublist l {..<n} = take n l"
  1668 apply (induct l rule: rev_induct, simp)
  1669 apply (simp split: nat_diff_split add: sublist_append)
  1670 done
  1671 
  1672 
  1673 lemma take_Cons':
  1674      "take n (x # xs) = (if n = 0 then [] else x # take (n - 1) xs)"
  1675 by (cases n) simp_all
  1676 
  1677 lemma drop_Cons':
  1678      "drop n (x # xs) = (if n = 0 then x # xs else drop (n - 1) xs)"
  1679 by (cases n) simp_all
  1680 
  1681 lemma nth_Cons': "(x # xs)!n = (if n = 0 then x else xs!(n - 1))"
  1682 by (cases n) simp_all
  1683 
  1684 lemmas [simp] = take_Cons'[of "number_of v",standard]
  1685                 drop_Cons'[of "number_of v",standard]
  1686                 nth_Cons'[of _ _ "number_of v",standard]
  1687 
  1688 
  1689 lemma "card (set xs) = size xs \<Longrightarrow> distinct xs"
  1690 proof (induct xs)
  1691   case Nil thus ?case by simp
  1692 next
  1693   case (Cons x xs)
  1694   show ?case
  1695   proof (cases "x \<in> set xs")
  1696     case False with Cons show ?thesis by simp
  1697   next
  1698     case True with Cons.prems
  1699     have "card (set xs) = Suc (length xs)" 
  1700       by (simp add: card_insert_if split: split_if_asm)
  1701     moreover have "card (set xs) \<le> length xs" by (rule card_length)
  1702     ultimately have False by simp
  1703     thus ?thesis ..
  1704   qed
  1705 qed
  1706 
  1707 subsection {* Characters and strings *}
  1708 
  1709 datatype nibble =
  1710     Nibble0 | Nibble1 | Nibble2 | Nibble3 | Nibble4 | Nibble5 | Nibble6 | Nibble7
  1711   | Nibble8 | Nibble9 | NibbleA | NibbleB | NibbleC | NibbleD | NibbleE | NibbleF
  1712 
  1713 datatype char = Char nibble nibble
  1714   -- "Note: canonical order of character encoding coincides with standard term ordering"
  1715 
  1716 types string = "char list"
  1717 
  1718 syntax
  1719   "_Char" :: "xstr => char"    ("CHR _")
  1720   "_String" :: "xstr => string"    ("_")
  1721 
  1722 parse_ast_translation {*
  1723   let
  1724     val constants = Syntax.Appl o map Syntax.Constant;
  1725 
  1726     fun mk_nib n = "Nibble" ^ chr (n + (if n <= 9 then ord "0" else ord "A" - 10));
  1727     fun mk_char c =
  1728       if Symbol.is_ascii c andalso Symbol.is_printable c then
  1729         constants ["Char", mk_nib (ord c div 16), mk_nib (ord c mod 16)]
  1730       else error ("Printable ASCII character expected: " ^ quote c);
  1731 
  1732     fun mk_string [] = Syntax.Constant "Nil"
  1733       | mk_string (c :: cs) = Syntax.Appl [Syntax.Constant "Cons", mk_char c, mk_string cs];
  1734 
  1735     fun char_ast_tr [Syntax.Variable xstr] =
  1736         (case Syntax.explode_xstr xstr of
  1737           [c] => mk_char c
  1738         | _ => error ("Single character expected: " ^ xstr))
  1739       | char_ast_tr asts = raise AST ("char_ast_tr", asts);
  1740 
  1741     fun string_ast_tr [Syntax.Variable xstr] =
  1742         (case Syntax.explode_xstr xstr of
  1743           [] => constants [Syntax.constrainC, "Nil", "string"]
  1744         | cs => mk_string cs)
  1745       | string_ast_tr asts = raise AST ("string_tr", asts);
  1746   in [("_Char", char_ast_tr), ("_String", string_ast_tr)] end;
  1747 *}
  1748 
  1749 ML {*
  1750 fun int_of_nibble h =
  1751   if "0" <= h andalso h <= "9" then ord h - ord "0"
  1752   else if "A" <= h andalso h <= "F" then ord h - ord "A" + 10
  1753   else raise Match;
  1754 
  1755 fun nibble_of_int i =
  1756   if i <= 9 then chr (ord "0" + i) else chr (ord "A" + i - 10);
  1757 *}
  1758 
  1759 print_ast_translation {*
  1760   let
  1761     fun dest_nib (Syntax.Constant c) =
  1762         (case explode c of
  1763           ["N", "i", "b", "b", "l", "e", h] => int_of_nibble h
  1764         | _ => raise Match)
  1765       | dest_nib _ = raise Match;
  1766 
  1767     fun dest_chr c1 c2 =
  1768       let val c = chr (dest_nib c1 * 16 + dest_nib c2)
  1769       in if Symbol.is_printable c then c else raise Match end;
  1770 
  1771     fun dest_char (Syntax.Appl [Syntax.Constant "Char", c1, c2]) = dest_chr c1 c2
  1772       | dest_char _ = raise Match;
  1773 
  1774     fun xstr cs = Syntax.Appl [Syntax.Constant "_xstr", Syntax.Variable (Syntax.implode_xstr cs)];
  1775 
  1776     fun char_ast_tr' [c1, c2] = Syntax.Appl [Syntax.Constant "_Char", xstr [dest_chr c1 c2]]
  1777       | char_ast_tr' _ = raise Match;
  1778 
  1779     fun list_ast_tr' [args] = Syntax.Appl [Syntax.Constant "_String",
  1780             xstr (map dest_char (Syntax.unfold_ast "_args" args))]
  1781       | list_ast_tr' ts = raise Match;
  1782   in [("Char", char_ast_tr'), ("@list", list_ast_tr')] end;
  1783 *}
  1784 
  1785 subsection {* Code generator setup *}
  1786 
  1787 ML {*
  1788 local
  1789 
  1790 fun list_codegen thy gr dep b t =
  1791   let val (gr', ps) = foldl_map (Codegen.invoke_codegen thy dep false)
  1792     (gr, HOLogic.dest_list t)
  1793   in Some (gr', Pretty.list "[" "]" ps) end handle TERM _ => None;
  1794 
  1795 fun dest_nibble (Const (s, _)) = int_of_nibble (unprefix "List.nibble.Nibble" s)
  1796   | dest_nibble _ = raise Match;
  1797 
  1798 fun char_codegen thy gr dep b (Const ("List.char.Char", _) $ c1 $ c2) =
  1799     (let val c = chr (dest_nibble c1 * 16 + dest_nibble c2)
  1800      in if Symbol.is_printable c then Some (gr, Pretty.quote (Pretty.str c))
  1801        else None
  1802      end handle LIST _ => None | Match => None)
  1803   | char_codegen thy gr dep b _ = None;
  1804 
  1805 in
  1806 
  1807 val list_codegen_setup =
  1808   [Codegen.add_codegen "list_codegen" list_codegen,
  1809    Codegen.add_codegen "char_codegen" char_codegen];
  1810 
  1811 end;
  1812 
  1813 val term_of_list = HOLogic.mk_list;
  1814 
  1815 fun gen_list' aG i j = frequency
  1816   [(i, fn () => aG j :: gen_list' aG (i-1) j), (1, fn () => [])] ()
  1817 and gen_list aG i = gen_list' aG i i;
  1818 
  1819 val nibbleT = Type ("List.nibble", []);
  1820 
  1821 fun term_of_char c =
  1822   Const ("List.char.Char", nibbleT --> nibbleT --> Type ("List.char", [])) $
  1823     Const ("List.nibble.Nibble" ^ nibble_of_int (ord c div 16), nibbleT) $
  1824     Const ("List.nibble.Nibble" ^ nibble_of_int (ord c mod 16), nibbleT);
  1825 
  1826 fun gen_char i = chr (random_range (ord "a") (Int.min (ord "a" + i, ord "z")));
  1827 *}
  1828 
  1829 types_code
  1830   "list" ("_ list")
  1831   "char" ("string")
  1832 
  1833 consts_code "Cons" ("(_ ::/ _)")
  1834 
  1835 setup list_codegen_setup
  1836 
  1837 end